| author | wenzelm | 
| Tue, 01 Jun 2010 11:37:24 +0200 | |
| changeset 37220 | d416e49b3926 | 
| parent 36603 | d5d6111761a6 | 
| child 37388 | 793618618f78 | 
| 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  | 
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
 | 
3  | 
|
| 9869 | 4  | 
The MESON resolution proof procedure for HOL.  | 
| 29267 | 5  | 
When making clauses, avoids using the rewriter -- instead uses RS recursively.  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
7  | 
|
| 24300 | 8  | 
signature MESON =  | 
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
9  | 
sig  | 
| 32955 | 10  | 
val trace: bool Unsynchronized.ref  | 
| 24300 | 11  | 
val term_pair_of: indexname * (typ * 'a) -> term * 'a  | 
12  | 
val first_order_resolve: thm -> thm -> thm  | 
|
13  | 
val flexflex_first_order: thm -> thm  | 
|
14  | 
val size_of_subgoals: thm -> int  | 
|
| 
26562
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
15  | 
val too_many_clauses: Proof.context option -> term -> bool  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
16  | 
val make_cnf: thm list -> thm -> Proof.context -> thm list * Proof.context  | 
| 24300 | 17  | 
val finish_cnf: thm list -> thm list  | 
| 32262 | 18  | 
val make_nnf: Proof.context -> thm -> thm  | 
19  | 
val skolemize: Proof.context -> thm -> thm  | 
|
| 24300 | 20  | 
val is_fol_term: theory -> term -> bool  | 
| 
35869
 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 
blanchet 
parents: 
35625 
diff
changeset
 | 
21  | 
val make_clauses_unsorted: thm list -> thm list  | 
| 24300 | 22  | 
val make_clauses: thm list -> thm list  | 
23  | 
val make_horns: thm list -> thm list  | 
|
24  | 
val best_prolog_tac: (thm -> int) -> thm list -> tactic  | 
|
25  | 
val depth_prolog_tac: thm list -> tactic  | 
|
26  | 
val gocls: thm list -> thm list  | 
|
| 32262 | 27  | 
val skolemize_prems_tac: Proof.context -> thm list -> int -> tactic  | 
28  | 
val MESON: (thm list -> thm list) -> (thm list -> tactic) -> Proof.context -> int -> tactic  | 
|
29  | 
val best_meson_tac: (thm -> int) -> Proof.context -> int -> tactic  | 
|
30  | 
val safe_best_meson_tac: Proof.context -> int -> tactic  | 
|
31  | 
val depth_meson_tac: Proof.context -> int -> tactic  | 
|
| 24300 | 32  | 
val prolog_step_tac': thm list -> int -> tactic  | 
33  | 
val iter_deepen_prolog_tac: thm list -> tactic  | 
|
| 32262 | 34  | 
val iter_deepen_meson_tac: Proof.context -> thm list -> int -> tactic  | 
| 24300 | 35  | 
val make_meta_clause: thm -> thm  | 
36  | 
val make_meta_clauses: thm list -> thm list  | 
|
| 32262 | 37  | 
val meson_tac: Proof.context -> thm list -> int -> tactic  | 
| 24300 | 38  | 
val negate_head: thm -> thm  | 
39  | 
val select_literal: int -> thm -> thm  | 
|
| 32262 | 40  | 
val skolemize_tac: Proof.context -> int -> tactic  | 
41  | 
val setup: theory -> theory  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
42  | 
end  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
43  | 
|
| 24300 | 44  | 
structure Meson: MESON =  | 
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
45  | 
struct  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
46  | 
|
| 32955 | 47  | 
val trace = Unsynchronized.ref false;  | 
48  | 
fun trace_msg msg = if ! trace then tracing (msg ()) else ();  | 
|
49  | 
||
| 
26562
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
50  | 
val max_clauses_default = 60;  | 
| 36001 | 51  | 
val (max_clauses, setup) = Attrib.config_int "max_clauses" (K max_clauses_default);  | 
| 
26562
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
52  | 
|
| 31454 | 53  | 
val disj_forward = @{thm disj_forward};
 | 
54  | 
val disj_forward2 = @{thm disj_forward2};
 | 
|
55  | 
val make_pos_rule = @{thm make_pos_rule};
 | 
|
56  | 
val make_pos_rule' = @{thm make_pos_rule'};
 | 
|
57  | 
val make_pos_goal = @{thm make_pos_goal};
 | 
|
58  | 
val make_neg_rule = @{thm make_neg_rule};
 | 
|
59  | 
val make_neg_rule' = @{thm make_neg_rule'};
 | 
|
60  | 
val make_neg_goal = @{thm make_neg_goal};
 | 
|
61  | 
val conj_forward = @{thm conj_forward};
 | 
|
62  | 
val all_forward = @{thm all_forward};
 | 
|
63  | 
val ex_forward = @{thm ex_forward};
 | 
|
64  | 
val choice = @{thm choice};
 | 
|
65  | 
||
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
66  | 
val not_conjD = thm "meson_not_conjD";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
67  | 
val not_disjD = thm "meson_not_disjD";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
68  | 
val not_notD = thm "meson_not_notD";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
69  | 
val not_allD = thm "meson_not_allD";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
70  | 
val not_exD = thm "meson_not_exD";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
71  | 
val imp_to_disjD = thm "meson_imp_to_disjD";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
72  | 
val not_impD = thm "meson_not_impD";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
73  | 
val iff_to_disjD = thm "meson_iff_to_disjD";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
74  | 
val not_iffD = thm "meson_not_iffD";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
75  | 
val conj_exD1 = thm "meson_conj_exD1";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
76  | 
val conj_exD2 = thm "meson_conj_exD2";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
77  | 
val disj_exD = thm "meson_disj_exD";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
78  | 
val disj_exD1 = thm "meson_disj_exD1";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
79  | 
val disj_exD2 = thm "meson_disj_exD2";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
80  | 
val disj_assoc = thm "meson_disj_assoc";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
81  | 
val disj_comm = thm "meson_disj_comm";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
82  | 
val disj_FalseD1 = thm "meson_disj_FalseD1";  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
83  | 
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
 | 
84  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
85  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
86  | 
(**** Operators for forward proof ****)  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
87  | 
|
| 
20417
 
c611b1412056
better skolemization, using first-order resolution rather than hoping for the right result
 
paulson 
parents: 
20361 
diff
changeset
 | 
88  | 
|
| 
 
c611b1412056
better skolemization, using first-order resolution rather than hoping for the right result
 
paulson 
parents: 
20361 
diff
changeset
 | 
89  | 
(** First-order Resolution **)  | 
| 
 
c611b1412056
better skolemization, using first-order resolution rather than hoping for the right result
 
paulson 
parents: 
20361 
diff
changeset
 | 
90  | 
|
| 
 
c611b1412056
better skolemization, using first-order resolution rather than hoping for the right result
 
paulson 
parents: 
20361 
diff
changeset
 | 
91  | 
fun typ_pair_of (ix, (sort,ty)) = (TVar (ix,sort), ty);  | 
| 
 
c611b1412056
better skolemization, using first-order resolution rather than hoping for the right result
 
paulson 
parents: 
20361 
diff
changeset
 | 
92  | 
fun term_pair_of (ix, (ty,t)) = (Var (ix,ty), t);  | 
| 
 
c611b1412056
better skolemization, using first-order resolution rather than hoping for the right result
 
paulson 
parents: 
20361 
diff
changeset
 | 
93  | 
|
| 
 
c611b1412056
better skolemization, using first-order resolution rather than hoping for the right result
 
paulson 
parents: 
20361 
diff
changeset
 | 
94  | 
(*FIXME: currently does not "rename variables apart"*)  | 
| 
 
c611b1412056
better skolemization, using first-order resolution rather than hoping for the right result
 
paulson 
parents: 
20361 
diff
changeset
 | 
95  | 
fun first_order_resolve thA thB =  | 
| 32262 | 96  | 
(case  | 
97  | 
try (fn () =>  | 
|
98  | 
let val thy = theory_of_thm thA  | 
|
99  | 
val tmA = concl_of thA  | 
|
100  | 
          val Const("==>",_) $ tmB $ _ = prop_of thB
 | 
|
101  | 
val (tyenv, tenv) = Pattern.first_order_match thy (tmB, tmA) (Vartab.empty, Vartab.empty)  | 
|
102  | 
val ct_pairs = map (pairself (cterm_of thy) o term_pair_of) (Vartab.dest tenv)  | 
|
103  | 
in thA RS (cterm_instantiate ct_pairs thB) end) () of  | 
|
104  | 
SOME th => th  | 
|
105  | 
  | NONE => raise THM ("first_order_resolve", 0, [thA, thB]));
 | 
|
| 
18175
 
7858b777569a
new version of "tryres" allowing multiple unifiers (apparently needed for
 
paulson 
parents: 
18141 
diff
changeset
 | 
106  | 
|
| 24300 | 107  | 
fun flexflex_first_order th =  | 
| 
23440
 
37860871f241
Added flexflex_first_order and tidied first_order_resolution
 
paulson 
parents: 
23262 
diff
changeset
 | 
108  | 
case (tpairs_of th) of  | 
| 
 
37860871f241
Added flexflex_first_order and tidied first_order_resolution
 
paulson 
parents: 
23262 
diff
changeset
 | 
109  | 
[] => th  | 
| 
 
37860871f241
Added flexflex_first_order and tidied first_order_resolution
 
paulson 
parents: 
23262 
diff
changeset
 | 
110  | 
| pairs =>  | 
| 24300 | 111  | 
let val thy = theory_of_thm th  | 
| 32032 | 112  | 
val (tyenv, tenv) =  | 
113  | 
fold (Pattern.first_order_match thy) pairs (Vartab.empty, Vartab.empty)  | 
|
| 24300 | 114  | 
val t_pairs = map term_pair_of (Vartab.dest tenv)  | 
115  | 
val th' = Thm.instantiate ([], map (pairself (cterm_of thy)) t_pairs) th  | 
|
116  | 
in th' end  | 
|
117  | 
handle THM _ => th;  | 
|
| 
23440
 
37860871f241
Added flexflex_first_order and tidied first_order_resolution
 
paulson 
parents: 
23262 
diff
changeset
 | 
118  | 
|
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
119  | 
(*Forward proof while preserving bound variables names*)  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
120  | 
fun rename_bvs_RS th rl =  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
121  | 
let val th' = th RS rl  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
122  | 
in Thm.rename_boundvars (concl_of th') (concl_of th) th' end;  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
123  | 
|
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
124  | 
(*raises exception if no rules apply*)  | 
| 24300 | 125  | 
fun tryres (th, rls) =  | 
| 
18141
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18024 
diff
changeset
 | 
126  | 
  let fun tryall [] = raise THM("tryres", 0, th::rls)
 | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
127  | 
| tryall (rl::rls) = (rename_bvs_RS th rl handle THM _ => tryall rls)  | 
| 
18141
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18024 
diff
changeset
 | 
128  | 
in tryall rls end;  | 
| 24300 | 129  | 
|
| 
21050
 
d0447a511edb
More robust error handling in make_nnf and forward_res
 
paulson 
parents: 
21046 
diff
changeset
 | 
130  | 
(*Permits forward proof from rules that discharge assumptions. The supplied proof state st,  | 
| 
 
d0447a511edb
More robust error handling in make_nnf and forward_res
 
paulson 
parents: 
21046 
diff
changeset
 | 
131  | 
e.g. from conj_forward, should have the form  | 
| 
 
d0447a511edb
More robust error handling in make_nnf and forward_res
 
paulson 
parents: 
21046 
diff
changeset
 | 
132  | 
"[| P' ==> ?P; Q' ==> ?Q |] ==> ?P & ?Q"  | 
| 
 
d0447a511edb
More robust error handling in make_nnf and forward_res
 
paulson 
parents: 
21046 
diff
changeset
 | 
133  | 
and the effect should be to instantiate ?P and ?Q with normalized versions of P' and Q'.*)  | 
| 32262 | 134  | 
fun forward_res ctxt nf st =  | 
| 
21050
 
d0447a511edb
More robust error handling in make_nnf and forward_res
 
paulson 
parents: 
21046 
diff
changeset
 | 
135  | 
let fun forward_tacf [prem] = rtac (nf prem) 1  | 
| 24300 | 136  | 
| forward_tacf prems =  | 
| 
32091
 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 
wenzelm 
parents: 
32032 
diff
changeset
 | 
137  | 
error (cat_lines  | 
| 
 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 
wenzelm 
parents: 
32032 
diff
changeset
 | 
138  | 
              ("Bad proof state in forward_res, please inform lcp@cl.cam.ac.uk:" ::
 | 
| 32262 | 139  | 
Display.string_of_thm ctxt st ::  | 
140  | 
"Premises:" :: map (Display.string_of_thm ctxt) prems))  | 
|
| 
21050
 
d0447a511edb
More robust error handling in make_nnf and forward_res
 
paulson 
parents: 
21046 
diff
changeset
 | 
141  | 
in  | 
| 
32231
 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
142  | 
case Seq.pull (ALLGOALS (OldGoals.METAHYPS forward_tacf) st)  | 
| 
21050
 
d0447a511edb
More robust error handling in make_nnf and forward_res
 
paulson 
parents: 
21046 
diff
changeset
 | 
143  | 
of SOME(th,_) => th  | 
| 
 
d0447a511edb
More robust error handling in make_nnf and forward_res
 
paulson 
parents: 
21046 
diff
changeset
 | 
144  | 
     | NONE => raise THM("forward_res", 0, [st])
 | 
| 
 
d0447a511edb
More robust error handling in make_nnf and forward_res
 
paulson 
parents: 
21046 
diff
changeset
 | 
145  | 
end;  | 
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
146  | 
|
| 
20134
 
73cb53843190
has_consts renamed to has_conn, now actually parses the first-order formula
 
paulson 
parents: 
20119 
diff
changeset
 | 
147  | 
(*Are any of the logical connectives in "bs" present in the term?*)  | 
| 
 
73cb53843190
has_consts renamed to has_conn, now actually parses the first-order formula
 
paulson 
parents: 
20119 
diff
changeset
 | 
148  | 
fun has_conns bs =  | 
| 
 
73cb53843190
has_consts renamed to has_conn, now actually parses the first-order formula
 
paulson 
parents: 
20119 
diff
changeset
 | 
149  | 
let fun has (Const(a,_)) = false  | 
| 
 
73cb53843190
has_consts renamed to has_conn, now actually parses the first-order formula
 
paulson 
parents: 
20119 
diff
changeset
 | 
150  | 
        | has (Const("Trueprop",_) $ p) = has p
 | 
| 
 
73cb53843190
has_consts renamed to has_conn, now actually parses the first-order formula
 
paulson 
parents: 
20119 
diff
changeset
 | 
151  | 
        | has (Const("Not",_) $ p) = has p
 | 
| 
 
73cb53843190
has_consts renamed to has_conn, now actually parses the first-order formula
 
paulson 
parents: 
20119 
diff
changeset
 | 
152  | 
        | has (Const("op |",_) $ p $ q) = member (op =) bs "op |" orelse has p orelse has q
 | 
| 
 
73cb53843190
has_consts renamed to has_conn, now actually parses the first-order formula
 
paulson 
parents: 
20119 
diff
changeset
 | 
153  | 
        | has (Const("op &",_) $ p $ q) = member (op =) bs "op &" orelse has p orelse has q
 | 
| 
 
73cb53843190
has_consts renamed to has_conn, now actually parses the first-order formula
 
paulson 
parents: 
20119 
diff
changeset
 | 
154  | 
        | has (Const("All",_) $ Abs(_,_,p)) = member (op =) bs "All" orelse has p
 | 
| 
 
73cb53843190
has_consts renamed to has_conn, now actually parses the first-order formula
 
paulson 
parents: 
20119 
diff
changeset
 | 
155  | 
        | has (Const("Ex",_) $ Abs(_,_,p)) = member (op =) bs "Ex" orelse has p
 | 
| 24300 | 156  | 
| has _ = false  | 
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
157  | 
in has end;  | 
| 24300 | 158  | 
|
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
159  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
160  | 
(**** Clause handling ****)  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
161  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
162  | 
fun literals (Const("Trueprop",_) $ P) = literals P
 | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
163  | 
  | literals (Const("op |",_) $ P $ Q) = literals P @ literals Q
 | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
164  | 
  | literals (Const("Not",_) $ P) = [(false,P)]
 | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
165  | 
| literals P = [(true,P)];  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
166  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
167  | 
(*number of literals in a term*)  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
168  | 
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
 | 
169  | 
|
| 18389 | 170  | 
|
171  | 
(*** Tautology Checking ***)  | 
|
172  | 
||
| 24300 | 173  | 
fun signed_lits_aux (Const ("op |", _) $ P $ Q) (poslits, neglits) =
 | 
| 18389 | 174  | 
signed_lits_aux Q (signed_lits_aux P (poslits, neglits))  | 
175  | 
  | signed_lits_aux (Const("Not",_) $ P) (poslits, neglits) = (poslits, P::neglits)
 | 
|
176  | 
| signed_lits_aux P (poslits, neglits) = (P::poslits, neglits);  | 
|
| 24300 | 177  | 
|
| 18389 | 178  | 
fun signed_lits th = signed_lits_aux (HOLogic.dest_Trueprop (concl_of th)) ([],[]);  | 
179  | 
||
180  | 
(*Literals like X=X are tautologous*)  | 
|
181  | 
fun taut_poslit (Const("op =",_) $ t $ u) = t aconv u
 | 
|
182  | 
  | taut_poslit (Const("True",_)) = true
 | 
|
183  | 
| taut_poslit _ = false;  | 
|
184  | 
||
185  | 
fun is_taut th =  | 
|
186  | 
let val (poslits,neglits) = signed_lits th  | 
|
187  | 
in exists taut_poslit poslits  | 
|
188  | 
orelse  | 
|
| 20073 | 189  | 
exists (member (op aconv) neglits) (HOLogic.false_const :: poslits)  | 
| 19894 | 190  | 
end  | 
| 24300 | 191  | 
handle TERM _ => false; (*probably dest_Trueprop on a weird theorem*)  | 
| 18389 | 192  | 
|
193  | 
||
194  | 
(*** To remove trivial negated equality literals from clauses ***)  | 
|
195  | 
||
196  | 
(*They are typically functional reflexivity axioms and are the converses of  | 
|
197  | 
injectivity equivalences*)  | 
|
| 24300 | 198  | 
|
| 18389 | 199  | 
val not_refl_disj_D = thm"meson_not_refl_disj_D";  | 
200  | 
||
| 20119 | 201  | 
(*Is either term a Var that does not properly occur in the other term?*)  | 
202  | 
fun eliminable (t as Var _, u) = t aconv u orelse not (Logic.occs(t,u))  | 
|
203  | 
| eliminable (u, t as Var _) = t aconv u orelse not (Logic.occs(t,u))  | 
|
204  | 
| eliminable _ = false;  | 
|
205  | 
||
| 18389 | 206  | 
fun refl_clause_aux 0 th = th  | 
207  | 
| refl_clause_aux n th =  | 
|
208  | 
case HOLogic.dest_Trueprop (concl_of th) of  | 
|
| 24300 | 209  | 
          (Const ("op |", _) $ (Const ("op |", _) $ _ $ _) $ _) =>
 | 
| 18389 | 210  | 
refl_clause_aux n (th RS disj_assoc) (*isolate an atom as first disjunct*)  | 
| 24300 | 211  | 
        | (Const ("op |", _) $ (Const("Not",_) $ (Const("op =",_) $ t $ u)) $ _) =>
 | 
212  | 
if eliminable(t,u)  | 
|
213  | 
then refl_clause_aux (n-1) (th RS not_refl_disj_D) (*Var inequation: delete*)  | 
|
214  | 
else refl_clause_aux (n-1) (th RS disj_comm) (*not between Vars: ignore*)  | 
|
215  | 
        | (Const ("op |", _) $ _ $ _) => refl_clause_aux n (th RS disj_comm)
 | 
|
216  | 
| _ => (*not a disjunction*) th;  | 
|
| 18389 | 217  | 
|
| 24300 | 218  | 
fun notequal_lits_count (Const ("op |", _) $ P $ Q) =
 | 
| 18389 | 219  | 
notequal_lits_count P + notequal_lits_count Q  | 
220  | 
  | notequal_lits_count (Const("Not",_) $ (Const("op =",_) $ _ $ _)) = 1
 | 
|
221  | 
| notequal_lits_count _ = 0;  | 
|
222  | 
||
223  | 
(*Simplify a clause by applying reflexivity to its negated equality literals*)  | 
|
| 24300 | 224  | 
fun refl_clause th =  | 
| 18389 | 225  | 
let val neqs = notequal_lits_count (HOLogic.dest_Trueprop (concl_of th))  | 
| 19894 | 226  | 
in zero_var_indexes (refl_clause_aux neqs th) end  | 
| 24300 | 227  | 
handle TERM _ => th; (*probably dest_Trueprop on a weird theorem*)  | 
| 18389 | 228  | 
|
229  | 
||
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
230  | 
(*** Removal of duplicate literals ***)  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
231  | 
|
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
232  | 
(*Forward proof, passing extra assumptions as theorems to the tactic*)  | 
| 32262 | 233  | 
fun forward_res2 ctxt nf hyps st =  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
234  | 
case Seq.pull  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
235  | 
(REPEAT  | 
| 
32231
 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
236  | 
(OldGoals.METAHYPS (fn major::minors => rtac (nf (minors@hyps) major) 1) 1)  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
237  | 
st)  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
238  | 
of SOME(th,_) => th  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
239  | 
   | NONE => raise THM("forward_res2", 0, [st]);
 | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
240  | 
|
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
241  | 
(*Remove duplicates in P|Q by assuming ~P in Q  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
242  | 
rls (initially []) accumulates assumptions of the form P==>False*)  | 
| 32262 | 243  | 
fun nodups_aux ctxt rls th = nodups_aux ctxt rls (th RS disj_assoc)  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
244  | 
handle THM _ => tryres(th,rls)  | 
| 32262 | 245  | 
handle THM _ => tryres(forward_res2 ctxt (nodups_aux ctxt) rls (th RS disj_forward2),  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
246  | 
[disj_FalseD1, disj_FalseD2, asm_rl])  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
247  | 
handle THM _ => th;  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
248  | 
|
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
249  | 
(*Remove duplicate literals, if there are any*)  | 
| 32262 | 250  | 
fun nodups ctxt th =  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
251  | 
if has_duplicates (op =) (literals (prop_of th))  | 
| 32262 | 252  | 
then nodups_aux ctxt [] th  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
253  | 
else th;  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
254  | 
|
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
255  | 
|
| 18389 | 256  | 
(*** The basic CNF transformation ***)  | 
257  | 
||
| 
26562
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
258  | 
fun too_many_clauses ctxto t =  | 
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
259  | 
let  | 
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
260  | 
val max_cl = case ctxto of SOME ctxt => Config.get ctxt max_clauses  | 
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
261  | 
| NONE => max_clauses_default  | 
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
262  | 
|
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
263  | 
fun sum x y = if x < max_cl andalso y < max_cl then x+y else max_cl;  | 
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
264  | 
fun prod x y = if x < max_cl andalso y < max_cl then x*y else max_cl;  | 
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
265  | 
|
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
266  | 
(*Estimate the number of clauses in order to detect infeasible theorems*)  | 
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
267  | 
  fun signed_nclauses b (Const("Trueprop",_) $ t) = signed_nclauses b t
 | 
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
268  | 
    | signed_nclauses b (Const("Not",_) $ t) = signed_nclauses (not b) t
 | 
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
269  | 
    | signed_nclauses b (Const("op &",_) $ t $ u) =
 | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
270  | 
if b then sum (signed_nclauses b t) (signed_nclauses b u)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
271  | 
else prod (signed_nclauses b t) (signed_nclauses b u)  | 
| 
26562
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
272  | 
    | signed_nclauses b (Const("op |",_) $ t $ u) =
 | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
273  | 
if b then prod (signed_nclauses b t) (signed_nclauses b u)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
274  | 
else sum (signed_nclauses b t) (signed_nclauses b u)  | 
| 
26562
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
275  | 
    | signed_nclauses b (Const("op -->",_) $ t $ u) =
 | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
276  | 
if b then prod (signed_nclauses (not b) t) (signed_nclauses b u)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
277  | 
else sum (signed_nclauses (not b) t) (signed_nclauses b u)  | 
| 
26562
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
278  | 
    | signed_nclauses b (Const("op =", Type ("fun", [T, _])) $ t $ u) =
 | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
279  | 
if T = HOLogic.boolT then (*Boolean equality is if-and-only-if*)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
280  | 
if b then sum (prod (signed_nclauses (not b) t) (signed_nclauses b u))  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
281  | 
(prod (signed_nclauses (not b) u) (signed_nclauses b t))  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
282  | 
else sum (prod (signed_nclauses b t) (signed_nclauses b u))  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
283  | 
(prod (signed_nclauses (not b) t) (signed_nclauses (not b) u))  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
284  | 
else 1  | 
| 
26562
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
285  | 
    | signed_nclauses b (Const("Ex", _) $ Abs (_,_,t)) = signed_nclauses b t
 | 
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
286  | 
    | signed_nclauses b (Const("All",_) $ Abs (_,_,t)) = signed_nclauses b t
 | 
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
287  | 
| signed_nclauses _ _ = 1; (* literal *)  | 
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
288  | 
in  | 
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
289  | 
signed_nclauses true t >= max_cl  | 
| 
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
26424 
diff
changeset
 | 
290  | 
end;  | 
| 19894 | 291  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
292  | 
(*Replaces universally quantified variables by FREE variables -- because  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
293  | 
assumptions may not contain scheme variables. Later, generalize using Variable.export. *)  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
294  | 
local  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
295  | 
val spec_var = Thm.dest_arg (Thm.dest_arg (#2 (Thm.dest_implies (Thm.cprop_of spec))));  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
296  | 
val spec_varT = #T (Thm.rep_cterm spec_var);  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
297  | 
  fun name_of (Const ("All", _) $ Abs(x,_,_)) = x | name_of _ = Name.uu;
 | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
298  | 
in  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
299  | 
fun freeze_spec th ctxt =  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
300  | 
let  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
301  | 
val cert = Thm.cterm_of (ProofContext.theory_of ctxt);  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
302  | 
val ([x], ctxt') = Variable.variant_fixes [name_of (HOLogic.dest_Trueprop (concl_of th))] ctxt;  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
303  | 
val spec' = Thm.instantiate ([], [(spec_var, cert (Free (x, spec_varT)))]) spec;  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
304  | 
in (th RS spec', ctxt') end  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
305  | 
end;  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
306  | 
|
| 
15998
 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 
paulson 
parents: 
15965 
diff
changeset
 | 
307  | 
(*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: 
15965 
diff
changeset
 | 
308  | 
and then normalized via function nf. The normal form is given to resolve_tac,  | 
| 
22515
 
f4061faa5fda
"generalize" now replaces ugly mes_XXX generated symbols by 1-letter identifiers.
 
paulson 
parents: 
22381 
diff
changeset
 | 
309  | 
instantiate a Boolean variable created by resolution with disj_forward. Since  | 
| 
 
f4061faa5fda
"generalize" now replaces ugly mes_XXX generated symbols by 1-letter identifiers.
 
paulson 
parents: 
22381 
diff
changeset
 | 
310  | 
(nf prem) returns a LIST of theorems, we can backtrack to get all combinations.*)  | 
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
311  | 
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
 | 
312  | 
|
| 24300 | 313  | 
(*Any need to extend this list with  | 
| 26424 | 314  | 
"HOL.type_class","HOL.eq_class","Pure.term"?*)  | 
| 24300 | 315  | 
val has_meta_conn =  | 
| 29684 | 316  | 
exists_Const (member (op =) ["==", "==>", "=simp=>", "all", "prop"] o #1);  | 
| 
20417
 
c611b1412056
better skolemization, using first-order resolution rather than hoping for the right result
 
paulson 
parents: 
20361 
diff
changeset
 | 
317  | 
|
| 24300 | 318  | 
fun apply_skolem_ths (th, rls) =  | 
| 
20417
 
c611b1412056
better skolemization, using first-order resolution rather than hoping for the right result
 
paulson 
parents: 
20361 
diff
changeset
 | 
319  | 
  let fun tryall [] = raise THM("apply_skolem_ths", 0, th::rls)
 | 
| 
 
c611b1412056
better skolemization, using first-order resolution rather than hoping for the right result
 
paulson 
parents: 
20361 
diff
changeset
 | 
320  | 
| tryall (rl::rls) = (first_order_resolve th rl handle THM _ => tryall rls)  | 
| 
 
c611b1412056
better skolemization, using first-order resolution rather than hoping for the right result
 
paulson 
parents: 
20361 
diff
changeset
 | 
321  | 
in tryall rls end;  | 
| 
22515
 
f4061faa5fda
"generalize" now replaces ugly mes_XXX generated symbols by 1-letter identifiers.
 
paulson 
parents: 
22381 
diff
changeset
 | 
322  | 
|
| 
15998
 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 
paulson 
parents: 
15965 
diff
changeset
 | 
323  | 
(*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: 
15965 
diff
changeset
 | 
324  | 
Strips universal quantifiers and breaks up conjunctions.  | 
| 
 
bc916036cf84
new cnf function taking Skolemization theorems as an extra argument
 
paulson 
parents: 
15965 
diff
changeset
 | 
325  | 
Eliminates existential quantifiers using skoths: Skolemization theorems.*)  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
326  | 
fun cnf skoths ctxt (th,ths) =  | 
| 33222 | 327  | 
let val ctxtr = Unsynchronized.ref ctxt (* FIXME ??? *)  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
328  | 
fun cnf_aux (th,ths) =  | 
| 24300 | 329  | 
if not (can HOLogic.dest_Trueprop (prop_of th)) then ths (*meta-level: ignore*)  | 
330  | 
else if not (has_conns ["All","Ex","op &"] (prop_of th))  | 
|
| 32262 | 331  | 
then nodups ctxt th :: ths (*no work to do, terminate*)  | 
| 24300 | 332  | 
else case head_of (HOLogic.dest_Trueprop (concl_of th)) of  | 
333  | 
            Const ("op &", _) => (*conjunction*)
 | 
|
334  | 
cnf_aux (th RS conjunct1, cnf_aux (th RS conjunct2, ths))  | 
|
335  | 
          | Const ("All", _) => (*universal quantifier*)
 | 
|
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
336  | 
let val (th',ctxt') = freeze_spec th (!ctxtr)  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
337  | 
in ctxtr := ctxt'; cnf_aux (th', ths) end  | 
| 24300 | 338  | 
          | Const ("Ex", _) =>
 | 
339  | 
(*existential quantifier: Insert Skolem functions*)  | 
|
340  | 
cnf_aux (apply_skolem_ths (th,skoths), ths)  | 
|
341  | 
          | Const ("op |", _) =>
 | 
|
342  | 
(*Disjunction of P, Q: Create new goal of proving ?P | ?Q and solve it using  | 
|
343  | 
all combinations of converting P, Q to CNF.*)  | 
|
344  | 
let val tac =  | 
|
| 32262 | 345  | 
OldGoals.METAHYPS (resop cnf_nil) 1 THEN  | 
| 
32231
 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
346  | 
(fn st' => st' |> OldGoals.METAHYPS (resop cnf_nil) 1)  | 
| 24300 | 347  | 
in Seq.list_of (tac (th RS disj_forward)) @ ths end  | 
| 32262 | 348  | 
| _ => nodups ctxt th :: ths (*no work to do*)  | 
| 19154 | 349  | 
and cnf_nil th = cnf_aux (th,[])  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
350  | 
val cls =  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
351  | 
if too_many_clauses (SOME ctxt) (concl_of th)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
352  | 
then (trace_msg (fn () => "cnf is ignoring: " ^ Display.string_of_thm ctxt th); ths)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32955 
diff
changeset
 | 
353  | 
else cnf_aux (th,ths)  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
354  | 
in (cls, !ctxtr) end;  | 
| 
22515
 
f4061faa5fda
"generalize" now replaces ugly mes_XXX generated symbols by 1-letter identifiers.
 
paulson 
parents: 
22381 
diff
changeset
 | 
355  | 
|
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
356  | 
fun make_cnf skoths th ctxt = cnf skoths ctxt (th, []);  | 
| 
20417
 
c611b1412056
better skolemization, using first-order resolution rather than hoping for the right result
 
paulson 
parents: 
20361 
diff
changeset
 | 
357  | 
|
| 
 
c611b1412056
better skolemization, using first-order resolution rather than hoping for the right result
 
paulson 
parents: 
20361 
diff
changeset
 | 
358  | 
(*Generalization, removal of redundant equalities, removal of tautologies.*)  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
359  | 
fun finish_cnf ths = filter (not o is_taut) (map refl_clause ths);  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
360  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
361  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
362  | 
(**** Generation of contrapositives ****)  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
363  | 
|
| 24300 | 364  | 
fun is_left (Const ("Trueprop", _) $
 | 
| 
21102
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
365  | 
               (Const ("op |", _) $ (Const ("op |", _) $ _ $ _) $ _)) = true
 | 
| 
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
366  | 
| is_left _ = false;  | 
| 24300 | 367  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
368  | 
(*Associate disjuctions to right -- make leftmost disjunct a LITERAL*)  | 
| 24300 | 369  | 
fun assoc_right th =  | 
| 
21102
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
370  | 
if is_left (prop_of th) then assoc_right (th RS disj_assoc)  | 
| 
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
371  | 
else th;  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
372  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
373  | 
(*Must check for negative literal first!*)  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
374  | 
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
 | 
375  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
376  | 
(*For ordinary resolution. *)  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
377  | 
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
 | 
378  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
379  | 
(*Create a goal or support clause, conclusing False*)  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
380  | 
fun make_goal th = (*Must check for negative literal first!*)  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
381  | 
make_goal (tryres(th, clause_rules))  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
382  | 
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
 | 
383  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
384  | 
(*Sort clauses by number of literals*)  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
385  | 
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
 | 
386  | 
|
| 18389 | 387  | 
fun sort_clauses ths = sort (make_ord fewerlits) ths;  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
388  | 
|
| 15581 | 389  | 
(*True if the given type contains bool anywhere*)  | 
390  | 
fun has_bool (Type("bool",_)) = true
 | 
|
391  | 
| has_bool (Type(_, Ts)) = exists has_bool Ts  | 
|
392  | 
| has_bool _ = false;  | 
|
| 24300 | 393  | 
|
394  | 
(*Is the string the name of a connective? Really only | and Not can remain,  | 
|
395  | 
since this code expects to be called on a clause form.*)  | 
|
| 19875 | 396  | 
val is_conn = member (op =)  | 
| 24300 | 397  | 
["Trueprop", "op &", "op |", "op -->", "Not",  | 
| 15613 | 398  | 
"All", "Ex", "Ball", "Bex"];  | 
399  | 
||
| 24300 | 400  | 
(*True if the term contains a function--not a logical connective--where the type  | 
| 
20524
 
1493053fc111
Tweaks to is_fol_term, the first-order test. We don't count "=" as a connective
 
paulson 
parents: 
20417 
diff
changeset
 | 
401  | 
of any argument contains bool.*)  | 
| 24300 | 402  | 
val has_bool_arg_const =  | 
| 15613 | 403  | 
exists_Const  | 
404  | 
(fn (c,T) => not(is_conn c) andalso exists (has_bool) (binder_types T));  | 
|
| 
22381
 
cb145d434284
The first-order test now tests for the obscure case of a polymorphic constant like 1 being
 
paulson 
parents: 
22360 
diff
changeset
 | 
405  | 
|
| 24300 | 406  | 
(*A higher-order instance of a first-order constant? Example is the definition of  | 
| 
34974
 
18b41bba42b5
new theory Algebras.thy for generic algebraic structures
 
haftmann 
parents: 
33832 
diff
changeset
 | 
407  | 
one, 1, at a function type in theory SetsAndFunctions.*)  | 
| 24300 | 408  | 
fun higher_inst_const thy (c,T) =  | 
| 
22381
 
cb145d434284
The first-order test now tests for the obscure case of a polymorphic constant like 1 being
 
paulson 
parents: 
22360 
diff
changeset
 | 
409  | 
case binder_types T of  | 
| 
 
cb145d434284
The first-order test now tests for the obscure case of a polymorphic constant like 1 being
 
paulson 
parents: 
22360 
diff
changeset
 | 
410  | 
[] => false (*not a function type, OK*)  | 
| 
 
cb145d434284
The first-order test now tests for the obscure case of a polymorphic constant like 1 being
 
paulson 
parents: 
22360 
diff
changeset
 | 
411  | 
| Ts => length (binder_types (Sign.the_const_type thy c)) <> length Ts;  | 
| 
 
cb145d434284
The first-order test now tests for the obscure case of a polymorphic constant like 1 being
 
paulson 
parents: 
22360 
diff
changeset
 | 
412  | 
|
| 
24742
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24300 
diff
changeset
 | 
413  | 
(*Returns false if any Vars in the theorem mention type bool.  | 
| 
21102
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
414  | 
Also rejects functions whose arguments are Booleans or other functions.*)  | 
| 
22381
 
cb145d434284
The first-order test now tests for the obscure case of a polymorphic constant like 1 being
 
paulson 
parents: 
22360 
diff
changeset
 | 
415  | 
fun is_fol_term thy t =  | 
| 
 
cb145d434284
The first-order test now tests for the obscure case of a polymorphic constant like 1 being
 
paulson 
parents: 
22360 
diff
changeset
 | 
416  | 
Term.is_first_order ["all","All","Ex"] t andalso  | 
| 29267 | 417  | 
not (exists_subterm (fn Var (_, T) => has_bool T | _ => false) t orelse  | 
| 24300 | 418  | 
has_bool_arg_const t orelse  | 
419  | 
exists_Const (higher_inst_const thy) t orelse  | 
|
420  | 
has_meta_conn t);  | 
|
| 
19204
 
b8f7de7ef629
Tidying, and getting rid of SELECT_GOAL (as it does something different now)
 
paulson 
parents: 
19154 
diff
changeset
 | 
421  | 
|
| 
21102
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
422  | 
fun rigid t = not (is_Var (head_of t));  | 
| 
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
423  | 
|
| 
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
424  | 
fun ok4horn (Const ("Trueprop",_) $ (Const ("op |", _) $ t $ _)) = rigid t
 | 
| 
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
425  | 
  | ok4horn (Const ("Trueprop",_) $ t) = rigid t
 | 
| 
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
426  | 
| ok4horn _ = false;  | 
| 
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
427  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
428  | 
(*Create a meta-level Horn clause*)  | 
| 24300 | 429  | 
fun make_horn crules th =  | 
430  | 
if ok4horn (concl_of th)  | 
|
| 
21102
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
431  | 
then make_horn crules (tryres(th,crules)) handle THM _ => th  | 
| 
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
432  | 
else th;  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
433  | 
|
| 16563 | 434  | 
(*Generate Horn clauses for all contrapositives of a clause. The input, th,  | 
435  | 
is a HOL disjunction.*)  | 
|
| 33339 | 436  | 
fun add_contras crules th hcs =  | 
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
437  | 
let fun rots (0,th) = hcs  | 
| 24300 | 438  | 
| rots (k,th) = zero_var_indexes (make_horn crules th) ::  | 
439  | 
rots(k-1, assoc_right (th RS disj_comm))  | 
|
| 15862 | 440  | 
in case nliterals(prop_of th) of  | 
| 24300 | 441  | 
1 => th::hcs  | 
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
442  | 
| n => rots(n, assoc_right th)  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
443  | 
end;  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
444  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
445  | 
(*Use "theorem naming" to label the clauses*)  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
446  | 
fun name_thms label =  | 
| 33339 | 447  | 
let fun name1 th (k, ths) =  | 
| 
27865
 
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
 
wenzelm 
parents: 
27239 
diff
changeset
 | 
448  | 
(k-1, Thm.put_name_hint (label ^ string_of_int k) th :: ths)  | 
| 33339 | 449  | 
in fn ths => #2 (fold_rev name1 ths (length ths, [])) end;  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
450  | 
|
| 16563 | 451  | 
(*Is the given disjunction an all-negative support clause?*)  | 
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
452  | 
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
 | 
453  | 
|
| 33317 | 454  | 
val neg_clauses = filter is_negative;  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
455  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
456  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
457  | 
(***** MESON PROOF PROCEDURE *****)  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
458  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
459  | 
fun rhyps (Const("==>",_) $ (Const("Trueprop",_) $ A) $ phi,
 | 
| 24300 | 460  | 
As) = rhyps(phi, A::As)  | 
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
461  | 
| rhyps (_, As) = As;  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
462  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
463  | 
(** 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
 | 
464  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
465  | 
(*The stringtree detects repeated assumptions.*)  | 
| 33245 | 466  | 
fun ins_term t net = 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
 | 
467  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
468  | 
(*detects repetitions in a list of terms*)  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
469  | 
fun has_reps [] = false  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
470  | 
| has_reps [_] = false  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
471  | 
| has_reps [t,u] = (t aconv u)  | 
| 33245 | 472  | 
| has_reps ts = (fold ins_term ts Net.empty; false) handle Net.INSERT => true;  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
473  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
474  | 
(*Like TRYALL eq_assume_tac, but avoids expensive THEN calls*)  | 
| 18508 | 475  | 
fun TRYING_eq_assume_tac 0 st = Seq.single st  | 
476  | 
| TRYING_eq_assume_tac i st =  | 
|
| 31945 | 477  | 
TRYING_eq_assume_tac (i-1) (Thm.eq_assumption i st)  | 
| 18508 | 478  | 
handle THM _ => TRYING_eq_assume_tac (i-1) st;  | 
479  | 
||
480  | 
fun TRYALL_eq_assume_tac st = TRYING_eq_assume_tac (nprems_of st) st;  | 
|
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
481  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
482  | 
(*Loop checking: FAIL if trying to prove the same thing twice  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
483  | 
-- if *ANY* subgoal has repeated literals*)  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
484  | 
fun check_tac st =  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
485  | 
if exists (fn prem => has_reps (rhyps(prem,[]))) (prems_of st)  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
486  | 
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
 | 
487  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
488  | 
|
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
489  | 
(* net_resolve_tac actually made it slower... *)  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
490  | 
fun prolog_step_tac horns i =  | 
| 
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
491  | 
(assume_tac i APPEND resolve_tac horns i) THEN check_tac THEN  | 
| 18508 | 492  | 
TRYALL_eq_assume_tac;  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
493  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
494  | 
(*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*)  | 
| 33339 | 495  | 
fun addconcl prem sz = size_of_term (Logic.strip_assums_concl prem) + sz;  | 
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
496  | 
|
| 33339 | 497  | 
fun size_of_subgoals st = fold_rev addconcl (prems_of st) 0;  | 
| 
15579
 
32bee18c675f
Tools/meson.ML: signature, structure and "open" rather than "local"
 
paulson 
parents: 
15574 
diff
changeset
 | 
498  | 
|
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
499  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
500  | 
(*Negation Normal Form*)  | 
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
501  | 
val nnf_rls = [imp_to_disjD, iff_to_disjD, not_conjD, not_disjD,  | 
| 9869 | 502  | 
not_impD, not_iffD, not_allD, not_exD, not_notD];  | 
| 15581 | 503  | 
|
| 
21102
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
504  | 
fun ok4nnf (Const ("Trueprop",_) $ (Const ("Not", _) $ t)) = rigid t
 | 
| 
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
505  | 
  | ok4nnf (Const ("Trueprop",_) $ t) = rigid t
 | 
| 
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
506  | 
| ok4nnf _ = false;  | 
| 
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
507  | 
|
| 32262 | 508  | 
fun make_nnf1 ctxt th =  | 
| 24300 | 509  | 
if ok4nnf (concl_of th)  | 
| 32262 | 510  | 
then make_nnf1 ctxt (tryres(th, nnf_rls))  | 
| 
28174
 
626f0a79a4b9
 more careful exception handling in order to prevent backtracking; miscellaneous tidying up.
 
paulson 
parents: 
27865 
diff
changeset
 | 
511  | 
    handle THM ("tryres", _, _) =>
 | 
| 32262 | 512  | 
forward_res ctxt (make_nnf1 ctxt)  | 
| 9869 | 513  | 
(tryres(th, [conj_forward,disj_forward,all_forward,ex_forward]))  | 
| 
28174
 
626f0a79a4b9
 more careful exception handling in order to prevent backtracking; miscellaneous tidying up.
 
paulson 
parents: 
27865 
diff
changeset
 | 
514  | 
    handle THM ("tryres", _, _) => th
 | 
| 
21102
 
7f2ebe5c5b72
Conversion to clause form now tolerates Boolean variables without looping.
 
paulson 
parents: 
21095 
diff
changeset
 | 
515  | 
else th;  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
516  | 
|
| 24300 | 517  | 
(*The simplification removes defined quantifiers and occurrences of True and False.  | 
| 20018 | 518  | 
nnf_ss also includes the one-point simprocs,  | 
| 
18405
 
afb1a52a7011
removal of some redundancies (e.g. one-point-rules) in clause production
 
paulson 
parents: 
18389 
diff
changeset
 | 
519  | 
which are needed to avoid the various one-point theorems from generating junk clauses.*)  | 
| 19894 | 520  | 
val nnf_simps =  | 
| 35410 | 521  | 
  [@{thm simp_implies_def}, @{thm Ex1_def}, @{thm Ball_def},@{thm  Bex_def}, @{thm if_True},
 | 
522  | 
    @{thm if_False}, @{thm if_cancel}, @{thm if_eq_cancel}, @{thm cases_simp}];
 | 
|
| 19894 | 523  | 
val nnf_extra_simps =  | 
| 35410 | 524  | 
  @{thms split_ifs} @ @{thms ex_simps} @ @{thms all_simps} @ @{thms simp_thms};
 | 
| 
18405
 
afb1a52a7011
removal of some redundancies (e.g. one-point-rules) in clause production
 
paulson 
parents: 
18389 
diff
changeset
 | 
525  | 
|
| 
 
afb1a52a7011
removal of some redundancies (e.g. one-point-rules) in clause production
 
paulson 
parents: 
18389 
diff
changeset
 | 
526  | 
val nnf_ss =  | 
| 24300 | 527  | 
HOL_basic_ss addsimps nnf_extra_simps  | 
| 24040 | 528  | 
    addsimprocs [defALL_regroup,defEX_regroup, @{simproc neq}, @{simproc let_simp}];
 | 
| 15872 | 529  | 
|
| 32262 | 530  | 
fun make_nnf ctxt th = case prems_of th of  | 
| 
21050
 
d0447a511edb
More robust error handling in make_nnf and forward_res
 
paulson 
parents: 
21046 
diff
changeset
 | 
531  | 
[] => th |> rewrite_rule (map safe_mk_meta_eq nnf_simps)  | 
| 24300 | 532  | 
|> simplify nnf_ss  | 
| 32262 | 533  | 
|> make_nnf1 ctxt  | 
| 
21050
 
d0447a511edb
More robust error handling in make_nnf and forward_res
 
paulson 
parents: 
21046 
diff
changeset
 | 
534  | 
  | _ => raise THM ("make_nnf: premises in argument", 0, [th]);
 | 
| 15581 | 535  | 
|
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15946 
diff
changeset
 | 
536  | 
(*Pull existential quantifiers to front. This accomplishes Skolemization for  | 
| 
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15946 
diff
changeset
 | 
537  | 
clauses that arise from a subgoal.*)  | 
| 32262 | 538  | 
fun skolemize1 ctxt th =  | 
| 
20134
 
73cb53843190
has_consts renamed to has_conn, now actually parses the first-order formula
 
paulson 
parents: 
20119 
diff
changeset
 | 
539  | 
if not (has_conns ["Ex"] (prop_of th)) then th  | 
| 32262 | 540  | 
else (skolemize1 ctxt (tryres(th, [choice, conj_exD1, conj_exD2,  | 
| 
15679
 
28eb0fe50533
Integrating the reconstruction files into the building of HOL
 
quigley 
parents: 
15653 
diff
changeset
 | 
541  | 
disj_exD, disj_exD1, disj_exD2])))  | 
| 
28174
 
626f0a79a4b9
 more careful exception handling in order to prevent backtracking; miscellaneous tidying up.
 
paulson 
parents: 
27865 
diff
changeset
 | 
542  | 
    handle THM ("tryres", _, _) =>
 | 
| 32262 | 543  | 
skolemize1 ctxt (forward_res ctxt (skolemize1 ctxt)  | 
| 9869 | 544  | 
(tryres (th, [conj_forward, disj_forward, all_forward])))  | 
| 
28174
 
626f0a79a4b9
 more careful exception handling in order to prevent backtracking; miscellaneous tidying up.
 
paulson 
parents: 
27865 
diff
changeset
 | 
545  | 
    handle THM ("tryres", _, _) => 
 | 
| 32262 | 546  | 
forward_res ctxt (skolemize1 ctxt) (rename_bvs_RS th ex_forward);  | 
| 29684 | 547  | 
|
| 32262 | 548  | 
fun skolemize ctxt th = skolemize1 ctxt (make_nnf ctxt th);  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
549  | 
|
| 32262 | 550  | 
fun skolemize_nnf_list _ [] = []  | 
551  | 
| skolemize_nnf_list ctxt (th::ths) =  | 
|
552  | 
skolemize ctxt th :: skolemize_nnf_list ctxt ths  | 
|
| 
25710
 
4cdf7de81e1b
Replaced refs by config params; finer critical section in mets method
 
paulson 
parents: 
25694 
diff
changeset
 | 
553  | 
handle THM _ => (*RS can fail if Unify.search_bound is too small*)  | 
| 32955 | 554  | 
(trace_msg (fn () => "Failed to Skolemize " ^ Display.string_of_thm ctxt th);  | 
| 32262 | 555  | 
skolemize_nnf_list ctxt ths);  | 
| 
25694
 
cbb59ba6bf0c
Skolemization now catches exception THM, which may be raised if unification fails.
 
paulson 
parents: 
24937 
diff
changeset
 | 
556  | 
|
| 33339 | 557  | 
fun add_clauses th cls =  | 
| 
36603
 
d5d6111761a6
renamed Variable.thm_context to Variable.global_thm_context to emphasize that this is not the real thing;
 
wenzelm 
parents: 
36001 
diff
changeset
 | 
558  | 
let val ctxt0 = Variable.global_thm_context th  | 
| 33339 | 559  | 
val (cnfs, ctxt) = make_cnf [] th ctxt0  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24827 
diff
changeset
 | 
560  | 
in Variable.export ctxt ctxt0 cnfs @ cls end;  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
561  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
562  | 
(*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
 | 
563  | 
The resulting clauses are HOL disjunctions.*)  | 
| 
35869
 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 
blanchet 
parents: 
35625 
diff
changeset
 | 
564  | 
fun make_clauses_unsorted ths = fold_rev add_clauses ths [];  | 
| 
 
cac366550624
start work on direct proof reconstruction for Sledgehammer
 
blanchet 
parents: 
35625 
diff
changeset
 | 
565  | 
val make_clauses = sort_clauses o make_clauses_unsorted;  | 
| 
15773
 
f14ae2432710
Completed integration of reconstruction code.  Now finds and displays proofs when used with modified version
 
quigley 
parents: 
15736 
diff
changeset
 | 
566  | 
|
| 16563 | 567  | 
(*Convert a list of clauses (disjunctions) to Horn clauses (contrapositives)*)  | 
| 9869 | 568  | 
fun make_horns ths =  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
569  | 
name_thms "Horn#"  | 
| 33339 | 570  | 
(distinct Thm.eq_thm_prop (fold_rev (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
 | 
571  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
572  | 
(*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
 | 
573  | 
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
 | 
574  | 
|
| 9869 | 575  | 
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
 | 
576  | 
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
 | 
577  | 
|
| 9869 | 578  | 
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
 | 
579  | 
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
 | 
580  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
581  | 
(*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
 | 
582  | 
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
 | 
583  | 
|
| 32262 | 584  | 
fun skolemize_prems_tac ctxt prems =  | 
585  | 
cut_facts_tac (skolemize_nnf_list ctxt prems) THEN'  | 
|
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
586  | 
REPEAT o (etac exE);  | 
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
587  | 
|
| 
22546
 
c40d7ab8cbc5
MESON tactical takes an additional argument: the clausification function.
 
paulson 
parents: 
22515 
diff
changeset
 | 
588  | 
(*Basis of all meson-tactics. Supplies cltac with clauses: HOL disjunctions.  | 
| 
 
c40d7ab8cbc5
MESON tactical takes an additional argument: the clausification function.
 
paulson 
parents: 
22515 
diff
changeset
 | 
589  | 
Function mkcl converts theorems to clauses.*)  | 
| 32262 | 590  | 
fun MESON mkcl cltac ctxt i st =  | 
| 16588 | 591  | 
SELECT_GOAL  | 
| 35625 | 592  | 
(EVERY [Object_Logic.atomize_prems_tac 1,  | 
| 23552 | 593  | 
rtac ccontr 1,  | 
| 32283 | 594  | 
            Subgoal.FOCUS (fn {context = ctxt', prems = negs, ...} =>
 | 
| 32262 | 595  | 
EVERY1 [skolemize_prems_tac ctxt negs,  | 
| 32283 | 596  | 
Subgoal.FOCUS (cltac o mkcl o #prems) ctxt']) ctxt 1]) i st  | 
| 24300 | 597  | 
handle THM _ => no_tac st; (*probably from make_meta_clause, not first-order*)  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
598  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
599  | 
(** Best-first search versions **)  | 
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
600  | 
|
| 16563 | 601  | 
(*ths is a list of additional clauses (HOL disjunctions) to use.*)  | 
| 9869 | 602  | 
fun best_meson_tac sizef =  | 
| 24300 | 603  | 
MESON make_clauses  | 
| 
22546
 
c40d7ab8cbc5
MESON tactical takes an additional argument: the clausification function.
 
paulson 
parents: 
22515 
diff
changeset
 | 
604  | 
(fn cls =>  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
605  | 
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
 | 
606  | 
(has_fewer_prems 1, sizef)  | 
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
607  | 
(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
 | 
608  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
609  | 
(*First, breaks the goal into independent units*)  | 
| 32262 | 610  | 
fun safe_best_meson_tac ctxt =  | 
611  | 
SELECT_GOAL (TRY (safe_tac (claset_of ctxt)) THEN  | 
|
612  | 
TRYALL (best_meson_tac size_of_subgoals ctxt));  | 
|
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
613  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
614  | 
(** Depth-first search version **)  | 
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
615  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
616  | 
val depth_meson_tac =  | 
| 
22546
 
c40d7ab8cbc5
MESON tactical takes an additional argument: the clausification function.
 
paulson 
parents: 
22515 
diff
changeset
 | 
617  | 
MESON make_clauses  | 
| 
 
c40d7ab8cbc5
MESON tactical takes an additional argument: the clausification function.
 
paulson 
parents: 
22515 
diff
changeset
 | 
618  | 
(fn cls => EVERY [resolve_tac (gocls cls) 1, depth_prolog_tac (make_horns cls)]);  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
619  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
620  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
621  | 
(** Iterative deepening version **)  | 
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
622  | 
|
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
623  | 
(*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
 | 
624  | 
having only one eq_assume_tac speeds it up!*)  | 
| 9869 | 625  | 
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
 | 
626  | 
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
 | 
627  | 
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
 | 
628  | 
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
 | 
629  | 
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
 | 
630  | 
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
 | 
631  | 
((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
 | 
632  | 
end;  | 
| 
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
633  | 
|
| 9869 | 634  | 
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
 | 
635  | 
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
 | 
636  | 
|
| 32262 | 637  | 
fun iter_deepen_meson_tac ctxt ths = ctxt |> MESON make_clauses  | 
| 
32091
 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 
wenzelm 
parents: 
32032 
diff
changeset
 | 
638  | 
(fn cls =>  | 
| 
 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 
wenzelm 
parents: 
32032 
diff
changeset
 | 
639  | 
(case (gocls (cls @ ths)) of  | 
| 
 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 
wenzelm 
parents: 
32032 
diff
changeset
 | 
640  | 
[] => no_tac (*no goal clauses*)  | 
| 
 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 
wenzelm 
parents: 
32032 
diff
changeset
 | 
641  | 
| goes =>  | 
| 
 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 
wenzelm 
parents: 
32032 
diff
changeset
 | 
642  | 
let  | 
| 
 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 
wenzelm 
parents: 
32032 
diff
changeset
 | 
643  | 
val horns = make_horns (cls @ ths)  | 
| 32955 | 644  | 
val _ = trace_msg (fn () =>  | 
| 
32091
 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 
wenzelm 
parents: 
32032 
diff
changeset
 | 
645  | 
            cat_lines ("meson method called:" ::
 | 
| 32262 | 646  | 
map (Display.string_of_thm ctxt) (cls @ ths) @  | 
647  | 
["clauses:"] @ map (Display.string_of_thm ctxt) horns))  | 
|
| 
32091
 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 
wenzelm 
parents: 
32032 
diff
changeset
 | 
648  | 
in THEN_ITER_DEEPEN (resolve_tac goes 1) (has_fewer_prems 1) (prolog_step_tac' horns) end));  | 
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
649  | 
|
| 32262 | 650  | 
fun meson_tac ctxt ths =  | 
651  | 
SELECT_GOAL (TRY (safe_tac (claset_of ctxt)) THEN TRYALL (iter_deepen_meson_tac ctxt ths));  | 
|
| 9869 | 652  | 
|
653  | 
||
| 14813 | 654  | 
(**** Code to support ordinary resolution, rather than Model Elimination ****)  | 
| 14744 | 655  | 
|
| 24300 | 656  | 
(*Convert a list of clauses (disjunctions) to meta-level clauses (==>),  | 
| 15008 | 657  | 
with no contrapositives, for ordinary resolution.*)  | 
| 14744 | 658  | 
|
659  | 
(*Rules to convert the head literal into a negated assumption. If the head  | 
|
660  | 
literal is already negated, then using notEfalse instead of notEfalse'  | 
|
661  | 
prevents a double negation.*)  | 
|
| 27239 | 662  | 
val notEfalse = read_instantiate @{context} [(("R", 0), "False")] notE;
 | 
| 14744 | 663  | 
val notEfalse' = rotate_prems 1 notEfalse;  | 
664  | 
||
| 24300 | 665  | 
fun negated_asm_of_head th =  | 
| 14744 | 666  | 
th RS notEfalse handle THM _ => th RS notEfalse';  | 
667  | 
||
| 
26066
 
19df083a2bbf
make_meta_clause bugfix: now works for higher-order clauses like LeastI_ex
 
paulson 
parents: 
25710 
diff
changeset
 | 
668  | 
(*Converting one theorem from a disjunction to a meta-level clause*)  | 
| 
 
19df083a2bbf
make_meta_clause bugfix: now works for higher-order clauses like LeastI_ex
 
paulson 
parents: 
25710 
diff
changeset
 | 
669  | 
fun make_meta_clause th =  | 
| 33832 | 670  | 
let val (fth,thaw) = Drule.legacy_freeze_thaw_robust th  | 
| 
26066
 
19df083a2bbf
make_meta_clause bugfix: now works for higher-order clauses like LeastI_ex
 
paulson 
parents: 
25710 
diff
changeset
 | 
671  | 
in  | 
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35625 
diff
changeset
 | 
672  | 
(zero_var_indexes o Thm.varifyT_global o thaw 0 o  | 
| 
26066
 
19df083a2bbf
make_meta_clause bugfix: now works for higher-order clauses like LeastI_ex
 
paulson 
parents: 
25710 
diff
changeset
 | 
673  | 
negated_asm_of_head o make_horn resolution_clause_rules) fth  | 
| 
 
19df083a2bbf
make_meta_clause bugfix: now works for higher-order clauses like LeastI_ex
 
paulson 
parents: 
25710 
diff
changeset
 | 
674  | 
end;  | 
| 24300 | 675  | 
|
| 14744 | 676  | 
fun make_meta_clauses ths =  | 
677  | 
name_thms "MClause#"  | 
|
| 
22360
 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 
wenzelm 
parents: 
22130 
diff
changeset
 | 
678  | 
(distinct Thm.eq_thm_prop (map make_meta_clause ths));  | 
| 14744 | 679  | 
|
680  | 
(*Permute a rule's premises to move the i-th premise to the last position.*)  | 
|
681  | 
fun make_last i th =  | 
|
| 24300 | 682  | 
let val n = nprems_of th  | 
683  | 
in if 1 <= i andalso i <= n  | 
|
| 14744 | 684  | 
then Thm.permute_prems (i-1) 1 th  | 
| 15118 | 685  | 
      else raise THM("select_literal", i, [th])
 | 
| 14744 | 686  | 
end;  | 
687  | 
||
688  | 
(*Maps a rule that ends "... ==> P ==> False" to "... ==> ~P" while suppressing  | 
|
689  | 
double-negations.*)  | 
|
| 35410 | 690  | 
val negate_head = rewrite_rule [@{thm atomize_not}, not_not RS eq_reflection];
 | 
| 14744 | 691  | 
|
692  | 
(*Maps the clause [P1,...Pn]==>False to [P1,...,P(i-1),P(i+1),...Pn] ==> ~P*)  | 
|
693  | 
fun select_literal i cl = negate_head (make_last i cl);  | 
|
694  | 
||
| 18508 | 695  | 
|
| 14813 | 696  | 
(*Top-level Skolemization. Allows part of the conversion to clauses to be  | 
| 24300 | 697  | 
expressed as a tactic (or Isar method). Each assumption of the selected  | 
| 14813 | 698  | 
goal is converted to NNF and then its existential quantifiers are pulled  | 
| 24300 | 699  | 
to the front. Finally, all existential quantifiers are eliminated,  | 
| 14813 | 700  | 
leaving !!-quantified variables. Perhaps Safe_tac should follow, but it  | 
701  | 
might generate many subgoals.*)  | 
|
| 
18194
 
940515d2fa26
-- removed "check_is_fol" from "make_nnf" so that the NNF procedure doesn't check whether a thm is FOL.
 
mengj 
parents: 
18175 
diff
changeset
 | 
702  | 
|
| 32262 | 703  | 
fun skolemize_tac ctxt = SUBGOAL (fn (goal, i) =>  | 
704  | 
let val ts = Logic.strip_assums_hyp goal  | 
|
| 24300 | 705  | 
in  | 
| 32262 | 706  | 
EVERY'  | 
707  | 
[OldGoals.METAHYPS (fn hyps =>  | 
|
708  | 
(cut_facts_tac (skolemize_nnf_list ctxt hyps) 1  | 
|
709  | 
THEN REPEAT (etac exE 1))),  | 
|
710  | 
REPEAT_DETERM_N (length ts) o (etac thin_rl)] i  | 
|
711  | 
end);  | 
|
| 
18194
 
940515d2fa26
-- removed "check_is_fol" from "make_nnf" so that the NNF procedure doesn't check whether a thm is FOL.
 
mengj 
parents: 
18175 
diff
changeset
 | 
712  | 
|
| 
9840
 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 
paulson 
parents:  
diff
changeset
 | 
713  | 
end;  |