| author | wenzelm | 
| Tue, 08 Nov 2005 10:43:08 +0100 | |
| changeset 18117 | 61a430a67d7c | 
| parent 18024 | 853e8219732a | 
| child 18141 | 89e2e8bed08f | 
| 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 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 14 | signature BASIC_MESON = | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 15 | sig | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 16 | val size_of_subgoals : thm -> int | 
| 15998 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 17 | val make_cnf : thm list -> thm -> thm list | 
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 18 | val make_nnf : thm -> thm | 
| 17849 | 19 | val make_nnf1 : thm -> thm | 
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 20 | val skolemize : thm -> thm | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 21 | val make_clauses : thm list -> thm list | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 22 | val make_horns : thm list -> thm list | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 23 | val best_prolog_tac : (thm -> int) -> thm list -> tactic | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 24 | val depth_prolog_tac : thm list -> tactic | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 25 | val gocls : thm list -> thm list | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 26 | val skolemize_prems_tac : thm list -> int -> tactic | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 27 | val MESON : (thm list -> tactic) -> int -> tactic | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 28 | val best_meson_tac : (thm -> int) -> int -> tactic | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 29 | val safe_best_meson_tac : int -> tactic | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 30 | val depth_meson_tac : int -> tactic | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 31 | val prolog_step_tac' : thm list -> int -> tactic | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 32 | val iter_deepen_prolog_tac : thm list -> tactic | 
| 16563 | 33 | val iter_deepen_meson_tac : thm list -> int -> tactic | 
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 34 | val meson_tac : int -> tactic | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 35 | val negate_head : thm -> thm | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 36 | val select_literal : int -> thm -> thm | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 37 | val skolemize_tac : int -> tactic | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 38 | val make_clauses_tac : int -> tactic | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 39 | end | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 40 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 41 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 42 | structure Meson = | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 43 | struct | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 44 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 45 | val not_conjD = thm "meson_not_conjD"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 46 | val not_disjD = thm "meson_not_disjD"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 47 | val not_notD = thm "meson_not_notD"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 48 | val not_allD = thm "meson_not_allD"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 49 | val not_exD = thm "meson_not_exD"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 50 | val imp_to_disjD = thm "meson_imp_to_disjD"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 51 | val not_impD = thm "meson_not_impD"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 52 | val iff_to_disjD = thm "meson_iff_to_disjD"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 53 | val not_iffD = thm "meson_not_iffD"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 54 | val conj_exD1 = thm "meson_conj_exD1"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 55 | val conj_exD2 = thm "meson_conj_exD2"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 56 | val disj_exD = thm "meson_disj_exD"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 57 | val disj_exD1 = thm "meson_disj_exD1"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 58 | val disj_exD2 = thm "meson_disj_exD2"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 59 | val disj_assoc = thm "meson_disj_assoc"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 60 | val disj_comm = thm "meson_disj_comm"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 61 | val disj_FalseD1 = thm "meson_disj_FalseD1"; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 62 | 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 | 63 | |
| 16563 | 64 | val depth_limit = ref 2000; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 65 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 66 | (**** Operators for forward proof ****) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 67 | |
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 68 | (*raises exception if no rules apply -- unlike RL*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 69 | fun tryres (th, rl::rls) = (th RS rl handle THM _ => tryres(th,rls)) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 70 |   | tryres (th, []) = raise THM("tryres", 0, [th]);
 | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 71 | |
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 72 | (*Permits forward proof from rules that discharge assumptions*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 73 | fun forward_res nf st = | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 74 | case Seq.pull (ALLGOALS (METAHYPS (fn [prem] => rtac (nf prem) 1)) st) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 75 | of SOME(th,_) => th | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 76 |    | NONE => raise THM("forward_res", 0, [st]);
 | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 77 | |
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 78 | |
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 79 | (*Are any of the constants in "bs" present in the term?*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 80 | fun has_consts bs = | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 81 | let fun has (Const(a,_)) = a mem bs | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 82 | 	| has (Const ("Hilbert_Choice.Eps",_) $ _) = false
 | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 83 | (*ignore constants within @-terms*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 84 | | has (f$u) = has f orelse has u | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 85 | | has (Abs(_,_,t)) = has t | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 86 | | has _ = false | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 87 | in has end; | 
| 17716 | 88 | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 89 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 90 | (**** Clause handling ****) | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 91 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 92 | fun literals (Const("Trueprop",_) $ P) = literals P
 | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 93 |   | literals (Const("op |",_) $ P $ Q) = literals P @ literals Q
 | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 94 |   | literals (Const("Not",_) $ P) = [(false,P)]
 | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 95 | | literals P = [(true,P)]; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 96 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 97 | (*number of literals in a term*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 98 | val nliterals = length o literals; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 99 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 100 | (*to detect, and remove, tautologous clauses*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 101 | fun taut_lits [] = false | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 102 | | taut_lits ((flg,t)::ts) = (not flg,t) mem ts orelse taut_lits ts; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 103 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 104 | (*Include False as a literal: an occurrence of ~False is a tautology*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 105 | fun is_taut th = taut_lits ((true, HOLogic.false_const) :: | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 106 | literals (prop_of th)); | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 107 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 108 | (*Generation of unique names -- maxidx cannot be relied upon to increase! | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 109 | Cannot rely on "variant", since variables might coincide when literals | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 110 | are joined to make a clause... | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 111 | 19 chooses "U" as the first variable name*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 112 | val name_ref = ref 19; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 113 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 114 | (*Replaces universally quantified variables by FREE variables -- because | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 115 | assumptions may not contain scheme variables. Later, call "generalize". *) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 116 | fun freeze_spec th = | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 117 | let val sth = th RS spec | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 118 | val newname = (name_ref := !name_ref + 1; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 119 | radixstring(26, "A", !name_ref)) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 120 |   in  read_instantiate [("x", newname)] sth  end;
 | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 121 | |
| 15998 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 122 | (*Used with METAHYPS below. There is one assumption, which gets bound to prem | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 123 | and then normalized via function nf. The normal form is given to resolve_tac, | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 124 | presumably to instantiate a Boolean variable.*) | 
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 125 | fun resop nf [prem] = resolve_tac (nf prem) 1; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 126 | |
| 15998 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 127 | (*Conjunctive normal form, adding clauses from th in front of ths (for foldr). | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 128 | Detects tautologies early, with "seen" holding the other literals of a clause. | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 129 | Strips universal quantifiers and breaks up conjunctions. | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 130 | Eliminates existential quantifiers using skoths: Skolemization theorems.*) | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 131 | fun cnf skoths (th,ths) = | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 132 | let fun cnf_aux seen (th,ths) = | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 133 | if taut_lits (literals(prop_of th) @ seen) | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 134 | then ths (*tautology ignored*) | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 135 | else if not (has_consts ["All","Ex","op &"] (prop_of th)) | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 136 | then th::ths (*no work to do, terminate*) | 
| 16588 | 137 | else case head_of (HOLogic.dest_Trueprop (concl_of th)) of | 
| 138 | 	    Const ("op &", _) => (*conjunction*)
 | |
| 139 | cnf_aux seen (th RS conjunct1, | |
| 140 | cnf_aux seen (th RS conjunct2, ths)) | |
| 141 | 	  | Const ("All", _) => (*universal quantifier*)
 | |
| 142 | cnf_aux seen (freeze_spec th, ths) | |
| 143 | 	  | Const ("Ex", _) => 
 | |
| 144 | (*existential quantifier: Insert Skolem functions*) | |
| 15998 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 145 | cnf_aux seen (tryres (th,skoths), ths) | 
| 16588 | 146 | 	  | Const ("op |", _) => (*disjunction*)
 | 
| 147 | let val tac = | |
| 148 | (METAHYPS (resop (cnf_nil seen)) 1) THEN | |
| 149 | (fn st' => st' |> | |
| 150 | METAHYPS | |
| 151 | (resop (cnf_nil (literals (concl_of st') @ seen))) 1) | |
| 152 | in Seq.list_of (tac (th RS disj_forward)) @ ths end | |
| 153 | | _ => (*no work to do*) th::ths | |
| 15998 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 154 | and cnf_nil seen th = (cnf_aux seen (th,[])) | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 155 | in | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 156 | name_ref := 19; (*It's safe to reset this in a top-level call to cnf*) | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 157 | (cnf skoths (th RS conjunct1, cnf skoths (th RS conjunct2, ths)) | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 158 | handle THM _ => (*not a conjunction*) cnf_aux [] (th, ths)) | 
| 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 159 | end; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 160 | |
| 16012 | 161 | (*Convert all suitable free variables to schematic variables, | 
| 162 | but don't discharge assumptions.*) | |
| 16173 | 163 | fun generalize th = Thm.varifyT (forall_elim_vars 0 (forall_intr_frees th)); | 
| 16012 | 164 | |
| 165 | fun make_cnf skoths th = map generalize (cnf skoths (th, [])); | |
| 15998 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 166 | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 167 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 168 | (**** Removal of duplicate literals ****) | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 169 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 170 | (*Forward proof, passing extra assumptions as theorems to the tactic*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 171 | fun forward_res2 nf hyps st = | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 172 | case Seq.pull | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 173 | (REPEAT | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 174 | (METAHYPS (fn major::minors => rtac (nf (minors@hyps) major) 1) 1) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 175 | st) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 176 | of SOME(th,_) => th | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 177 |    | NONE => raise THM("forward_res2", 0, [st]);
 | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 178 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 179 | (*Remove duplicates in P|Q by assuming ~P in Q | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 180 | rls (initially []) accumulates assumptions of the form P==>False*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 181 | fun nodups_aux rls th = nodups_aux rls (th RS disj_assoc) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 182 | handle THM _ => tryres(th,rls) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 183 | handle THM _ => tryres(forward_res2 nodups_aux rls (th RS disj_forward2), | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 184 | [disj_FalseD1, disj_FalseD2, asm_rl]) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 185 | handle THM _ => th; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 186 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 187 | (*Remove duplicate literals, if there are any*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 188 | fun nodups th = | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 189 | if null(findrep(literals(prop_of th))) then th | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 190 | else nodups_aux [] th; | 
| 9840 
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 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 193 | (**** Generation of contrapositives ****) | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 194 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 195 | (*Associate disjuctions to right -- make leftmost disjunct a LITERAL*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 196 | fun assoc_right th = assoc_right (th RS disj_assoc) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 197 | handle THM _ => th; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 198 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 199 | (*Must check for negative literal first!*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 200 | val clause_rules = [disj_assoc, make_neg_rule, make_pos_rule]; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 201 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 202 | (*For ordinary resolution. *) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 203 | val resolution_clause_rules = [disj_assoc, make_neg_rule', make_pos_rule']; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 204 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 205 | (*Create a goal or support clause, conclusing False*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 206 | fun make_goal th = (*Must check for negative literal first!*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 207 | make_goal (tryres(th, clause_rules)) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 208 | handle THM _ => tryres(th, [make_neg_goal, make_pos_goal]); | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 209 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 210 | (*Sort clauses by number of literals*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 211 | fun fewerlits(th1,th2) = nliterals(prop_of th1) < nliterals(prop_of th2); | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 212 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 213 | (*TAUTOLOGY CHECK SHOULD NOT BE NECESSARY!*) | 
| 16012 | 214 | fun sort_clauses ths = | 
| 215 | sort (make_ord fewerlits) (List.filter (not o is_taut) ths); | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 216 | |
| 15581 | 217 | (*True if the given type contains bool anywhere*) | 
| 218 | fun has_bool (Type("bool",_)) = true
 | |
| 219 | | has_bool (Type(_, Ts)) = exists has_bool Ts | |
| 220 | | has_bool _ = false; | |
| 221 | ||
| 15613 | 222 | (*Is the string the name of a connective? It doesn't matter if this list is | 
| 223 | incomplete, since when actually called, the only connectives likely to | |
| 224 | remain are & | Not.*) | |
| 225 | fun is_conn c = c mem_string | |
| 17404 
d16c3a62c396
the experimental tagging system, and the usual tidying
 paulson parents: 
16801diff
changeset | 226 | ["Trueprop", "HOL.tag", "op &", "op |", "op -->", "op =", "Not", | 
| 15613 | 227 | "All", "Ex", "Ball", "Bex"]; | 
| 228 | ||
| 229 | (*True if the term contains a function where the type of any argument contains | |
| 230 | bool.*) | |
| 231 | val has_bool_arg_const = | |
| 232 | exists_Const | |
| 233 | (fn (c,T) => not(is_conn c) andalso exists (has_bool) (binder_types T)); | |
| 15908 | 234 | |
| 235 | val has_meta_conn = | |
| 18024 | 236 | exists_Const (fn (c,_) => c mem_string ["==", "==>", "all", "prop"]); | 
| 15613 | 237 | |
| 16588 | 238 | (*Raises an exception if any Vars in the theorem mention type bool; they | 
| 239 | could cause make_horn to loop! Also rejects functions whose arguments are | |
| 240 | Booleans or other functions.*) | |
| 241 | fun check_is_fol th = | |
| 15613 | 242 |   let val {prop,...} = rep_thm th
 | 
| 15908 | 243 | in if exists (has_bool o fastype_of) (term_vars prop) orelse | 
| 16588 | 244 | not (Term.is_first_order ["all","All","Ex"] prop) orelse | 
| 15908 | 245 | has_bool_arg_const prop orelse | 
| 246 | has_meta_conn prop | |
| 16588 | 247 |   then raise THM ("check_is_fol", 0, [th]) else th
 | 
| 15613 | 248 | end; | 
| 15581 | 249 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 250 | (*Create a meta-level Horn clause*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 251 | fun make_horn crules th = make_horn crules (tryres(th,crules)) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 252 | handle THM _ => th; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 253 | |
| 16563 | 254 | (*Generate Horn clauses for all contrapositives of a clause. The input, th, | 
| 255 | is a HOL disjunction.*) | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 256 | fun add_contras crules (th,hcs) = | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 257 | let fun rots (0,th) = hcs | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 258 | | rots (k,th) = zero_var_indexes (make_horn crules th) :: | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 259 | rots(k-1, assoc_right (th RS disj_comm)) | 
| 15862 | 260 | in case nliterals(prop_of th) of | 
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 261 | 1 => th::hcs | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 262 | | n => rots(n, assoc_right th) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 263 | end; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 264 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 265 | (*Use "theorem naming" to label the clauses*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 266 | fun name_thms label = | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 267 | let fun name1 (th, (k,ths)) = | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 268 | (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 | 269 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 270 | in fn ths => #2 (foldr name1 (length ths, []) ths) end; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 271 | |
| 16563 | 272 | (*Is the given disjunction an all-negative support clause?*) | 
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 273 | fun is_negative th = forall (not o #1) (literals (prop_of th)); | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 274 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 275 | val neg_clauses = List.filter is_negative; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 276 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 277 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 278 | (***** MESON PROOF PROCEDURE *****) | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 279 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 280 | fun rhyps (Const("==>",_) $ (Const("Trueprop",_) $ A) $ phi,
 | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 281 | As) = rhyps(phi, A::As) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 282 | | rhyps (_, As) = As; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 283 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 284 | (** Detecting repeated assumptions in a subgoal **) | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 285 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 286 | (*The stringtree detects repeated assumptions.*) | 
| 16801 | 287 | fun ins_term (net,t) = Net.insert_term (op aconv) (t,t) net; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 288 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 289 | (*detects repetitions in a list of terms*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 290 | fun has_reps [] = false | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 291 | | has_reps [_] = false | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 292 | | has_reps [t,u] = (t aconv u) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 293 | | has_reps ts = (Library.foldl ins_term (Net.empty, ts); false) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 294 | handle INSERT => true; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 295 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 296 | (*Like TRYALL eq_assume_tac, but avoids expensive THEN calls*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 297 | fun TRYALL_eq_assume_tac 0 st = Seq.single st | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 298 | | TRYALL_eq_assume_tac i st = | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 299 | TRYALL_eq_assume_tac (i-1) (eq_assumption i st) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 300 | 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 | 301 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 302 | (*Loop checking: FAIL if trying to prove the same thing twice | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 303 | -- if *ANY* subgoal has repeated literals*) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 304 | fun check_tac st = | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 305 | if exists (fn prem => has_reps (rhyps(prem,[]))) (prems_of st) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 306 | then Seq.empty else Seq.single st; | 
| 9840 
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 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 309 | (* net_resolve_tac actually made it slower... *) | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 310 | fun prolog_step_tac horns i = | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 311 | (assume_tac i APPEND resolve_tac horns i) THEN check_tac THEN | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 312 | TRYALL eq_assume_tac; | 
| 9840 
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 | (*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*) | 
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 315 | fun addconcl(prem,sz) = size_of_term(Logic.strip_assums_concl prem) + sz | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 316 | |
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 317 | fun size_of_subgoals st = foldr addconcl 0 (prems_of st); | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 318 | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 319 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 320 | (*Negation Normal Form*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 321 | val nnf_rls = [imp_to_disjD, iff_to_disjD, not_conjD, not_disjD, | 
| 9869 | 322 | not_impD, not_iffD, not_allD, not_exD, not_notD]; | 
| 15581 | 323 | |
| 324 | fun make_nnf1 th = make_nnf1 (tryres(th, nnf_rls)) | |
| 9869 | 325 | handle THM _ => | 
| 15581 | 326 | forward_res make_nnf1 | 
| 9869 | 327 | (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 | 328 | handle THM _ => th; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 329 | |
| 17404 
d16c3a62c396
the experimental tagging system, and the usual tidying
 paulson parents: 
16801diff
changeset | 330 | (*The simplification removes defined quantifiers and occurrences of True and False, as well as tags applied to True and False.*) | 
| 
d16c3a62c396
the experimental tagging system, and the usual tidying
 paulson parents: 
16801diff
changeset | 331 | val tag_True = thm "tag_True"; | 
| 
d16c3a62c396
the experimental tagging system, and the usual tidying
 paulson parents: 
16801diff
changeset | 332 | val tag_False = thm "tag_False"; | 
| 
d16c3a62c396
the experimental tagging system, and the usual tidying
 paulson parents: 
16801diff
changeset | 333 | val nnf_simps = [Ex1_def,Ball_def,Bex_def,tag_True,tag_False] | 
| 
d16c3a62c396
the experimental tagging system, and the usual tidying
 paulson parents: 
16801diff
changeset | 334 | val nnf_ss = HOL_basic_ss addsimps nnf_simps@simp_thms; | 
| 15872 | 335 | |
| 336 | fun make_nnf th = th |> simplify nnf_ss | |
| 16588 | 337 | |> check_is_fol |> make_nnf1 | 
| 15581 | 338 | |
| 15965 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 paulson parents: 
15946diff
changeset | 339 | (*Pull existential quantifiers to front. This accomplishes Skolemization for | 
| 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 paulson parents: 
15946diff
changeset | 340 | clauses that arise from a subgoal.*) | 
| 9869 | 341 | fun skolemize th = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 342 | if not (has_consts ["Ex"] (prop_of th)) then th | 
| 15773 
f14ae2432710
Completed integration of reconstruction code.  Now finds and displays proofs when used with modified version
 quigley parents: 
15736diff
changeset | 343 | else (skolemize (tryres(th, [choice, conj_exD1, conj_exD2, | 
| 15679 
28eb0fe50533
Integrating the reconstruction files into the building of HOL
 quigley parents: 
15653diff
changeset | 344 | disj_exD, disj_exD1, disj_exD2]))) | 
| 9869 | 345 | handle THM _ => | 
| 346 | skolemize (forward_res skolemize | |
| 347 | (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 | 348 | 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 | 349 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 350 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 351 | (*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 | 352 | The resulting clauses are HOL disjunctions.*) | 
| 9869 | 353 | fun make_clauses ths = | 
| 15998 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 paulson parents: 
15965diff
changeset | 354 | (sort_clauses (map (generalize o nodups) (foldr (cnf[]) [] ths))); | 
| 15773 
f14ae2432710
Completed integration of reconstruction code.  Now finds and displays proofs when used with modified version
 quigley parents: 
15736diff
changeset | 355 | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 356 | |
| 16563 | 357 | (*Convert a list of clauses (disjunctions) to Horn clauses (contrapositives)*) | 
| 9869 | 358 | fun make_horns ths = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 359 | name_thms "Horn#" | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 360 | (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 | 361 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 362 | (*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 | 363 | 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 | 364 | |
| 9869 | 365 | 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 | 366 | 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 | 367 | |
| 9869 | 368 | 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 | 369 | 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 | 370 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 371 | (*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 | 372 | 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 | 373 | |
| 15008 | 374 | fun skolemize_prems_tac prems = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 375 | 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 | 376 | REPEAT o (etac exE); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 377 | |
| 16588 | 378 | (*Basis of all meson-tactics. Supplies cltac with clauses: HOL disjunctions*) | 
| 379 | fun MESON cltac i st = | |
| 380 | SELECT_GOAL | |
| 381 | (EVERY1 [rtac ccontr, | |
| 382 | METAHYPS (fn negs => | |
| 383 | EVERY1 [skolemize_prems_tac negs, | |
| 384 | METAHYPS (cltac o make_clauses)])]) i st | |
| 385 | handle THM _ => no_tac st; (*probably from check_is_fol*) | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 386 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 387 | (** Best-first search versions **) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 388 | |
| 16563 | 389 | (*ths is a list of additional clauses (HOL disjunctions) to use.*) | 
| 9869 | 390 | fun best_meson_tac sizef = | 
| 391 | MESON (fn cls => | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 392 | 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 | 393 | (has_fewer_prems 1, sizef) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 394 | (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 | 395 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 396 | (*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 | 397 | val safe_best_meson_tac = | 
| 9869 | 398 | 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 | 399 | 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 | 400 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 401 | (** Depth-first search version **) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 402 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 403 | val depth_meson_tac = | 
| 9869 | 404 | 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 | 405 | 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 | 406 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 407 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 408 | (** Iterative deepening version **) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 409 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 410 | (*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 | 411 | having only one eq_assume_tac speeds it up!*) | 
| 9869 | 412 | 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 | 413 | 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 | 414 | 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 | 415 | 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 | 416 | 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 | 417 | 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 | 418 | ((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 | 419 | end; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 420 | |
| 9869 | 421 | 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 | 422 | 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 | 423 | |
| 16563 | 424 | fun iter_deepen_meson_tac ths = | 
| 9869 | 425 | MESON (fn cls => | 
| 16563 | 426 | case (gocls (cls@ths)) of | 
| 427 | [] => no_tac (*no goal clauses*) | |
| 428 | | goes => | |
| 429 | (THEN_ITER_DEEPEN (resolve_tac goes 1) | |
| 430 | (has_fewer_prems 1) | |
| 431 | (prolog_step_tac' (make_horns (cls@ths))))); | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 432 | |
| 16563 | 433 | fun meson_claset_tac ths cs = | 
| 434 | SELECT_GOAL (TRY (safe_tac cs) THEN TRYALL (iter_deepen_meson_tac ths)); | |
| 9869 | 435 | |
| 16563 | 436 | val meson_tac = CLASET' (meson_claset_tac []); | 
| 9869 | 437 | |
| 438 | ||
| 14813 | 439 | (**** Code to support ordinary resolution, rather than Model Elimination ****) | 
| 14744 | 440 | |
| 15008 | 441 | (*Convert a list of clauses (disjunctions) to meta-level clauses (==>), | 
| 442 | with no contrapositives, for ordinary resolution.*) | |
| 14744 | 443 | |
| 444 | (*Rules to convert the head literal into a negated assumption. If the head | |
| 445 | literal is already negated, then using notEfalse instead of notEfalse' | |
| 446 | prevents a double negation.*) | |
| 447 | val notEfalse = read_instantiate [("R","False")] notE;
 | |
| 448 | val notEfalse' = rotate_prems 1 notEfalse; | |
| 449 | ||
| 15448 | 450 | fun negated_asm_of_head th = | 
| 14744 | 451 | th RS notEfalse handle THM _ => th RS notEfalse'; | 
| 452 | ||
| 453 | (*Converting one clause*) | |
| 15581 | 454 | fun make_meta_clause th = | 
| 16588 | 455 | negated_asm_of_head (make_horn resolution_clause_rules (check_is_fol th)); | 
| 14744 | 456 | |
| 457 | fun make_meta_clauses ths = | |
| 458 | name_thms "MClause#" | |
| 459 | (gen_distinct Drule.eq_thm_prop (map make_meta_clause ths)); | |
| 460 | ||
| 461 | (*Permute a rule's premises to move the i-th premise to the last position.*) | |
| 462 | fun make_last i th = | |
| 463 | let val n = nprems_of th | |
| 464 | in if 1 <= i andalso i <= n | |
| 465 | then Thm.permute_prems (i-1) 1 th | |
| 15118 | 466 |       else raise THM("select_literal", i, [th])
 | 
| 14744 | 467 | end; | 
| 468 | ||
| 469 | (*Maps a rule that ends "... ==> P ==> False" to "... ==> ~P" while suppressing | |
| 470 | double-negations.*) | |
| 471 | val negate_head = rewrite_rule [atomize_not, not_not RS eq_reflection]; | |
| 472 | ||
| 473 | (*Maps the clause [P1,...Pn]==>False to [P1,...,P(i-1),P(i+1),...Pn] ==> ~P*) | |
| 474 | fun select_literal i cl = negate_head (make_last i cl); | |
| 475 | ||
| 14813 | 476 | (*Top-level Skolemization. Allows part of the conversion to clauses to be | 
| 477 | expressed as a tactic (or Isar method). Each assumption of the selected | |
| 478 | goal is converted to NNF and then its existential quantifiers are pulled | |
| 479 | to the front. Finally, all existential quantifiers are eliminated, | |
| 480 | leaving !!-quantified variables. Perhaps Safe_tac should follow, but it | |
| 481 | might generate many subgoals.*) | |
| 482 | val skolemize_tac = | |
| 483 | SUBGOAL | |
| 484 | (fn (prop,_) => | |
| 485 | let val ts = Logic.strip_assums_hyp prop | |
| 486 | in EVERY1 | |
| 487 | [METAHYPS | |
| 15773 
f14ae2432710
Completed integration of reconstruction code.  Now finds and displays proofs when used with modified version
 quigley parents: 
15736diff
changeset | 488 | (fn hyps => (cut_facts_tac (map (skolemize o make_nnf) hyps) 1 | 
| 14813 | 489 | THEN REPEAT (etac exE 1))), | 
| 490 | REPEAT_DETERM_N (length ts) o (etac thin_rl)] | |
| 491 | end); | |
| 492 | ||
| 15118 | 493 | (*Top-level conversion to meta-level clauses. Each clause has | 
| 494 | leading !!-bound universal variables, to express generality. To get | |
| 495 | disjunctions instead of meta-clauses, remove "make_meta_clauses" below.*) | |
| 15008 | 496 | val make_clauses_tac = | 
| 497 | SUBGOAL | |
| 498 | (fn (prop,_) => | |
| 499 | let val ts = Logic.strip_assums_hyp prop | |
| 500 | in EVERY1 | |
| 501 | [METAHYPS | |
| 502 | (fn hyps => | |
| 15151 | 503 | (Method.insert_tac | 
| 15118 | 504 | (map forall_intr_vars | 
| 505 | (make_meta_clauses (make_clauses hyps))) 1)), | |
| 15008 | 506 | REPEAT_DETERM_N (length ts) o (etac thin_rl)] | 
| 507 | end); | |
| 16563 | 508 | |
| 509 | ||
| 510 | (*** setup the special skoklemization methods ***) | |
| 9869 | 511 | |
| 16563 | 512 | (*No CHANGED_PROP here, since these always appear in the preamble*) | 
| 9869 | 513 | |
| 16563 | 514 | val skolemize_meth = Method.SIMPLE_METHOD' HEADGOAL skolemize_tac; | 
| 515 | ||
| 516 | val make_clauses_meth = Method.SIMPLE_METHOD' HEADGOAL make_clauses_tac; | |
| 14890 | 517 | |
| 16563 | 518 | val skolemize_setup = | 
| 9869 | 519 | [Method.add_methods | 
| 16563 | 520 |   [("skolemize", Method.no_args skolemize_meth, 
 | 
| 15008 | 521 | "Skolemization into existential quantifiers"), | 
| 522 |    ("make_clauses", Method.no_args make_clauses_meth, 
 | |
| 15118 | 523 | "Conversion to !!-quantified meta-level clauses")]]; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 524 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 525 | end; | 
| 9869 | 526 | |
| 15579 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 527 | structure BasicMeson: BASIC_MESON = Meson; | 
| 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 paulson parents: 
15574diff
changeset | 528 | open BasicMeson; |