| author | wenzelm | 
| Mon, 04 Nov 2024 21:25:26 +0100 | |
| changeset 81344 | 1b9ea66810ff | 
| parent 77896 | a9626bcb0c3b | 
| child 81945 | 23957ebffaf7 | 
| permissions | -rw-r--r-- | 
| 55239 | 1  | 
(* Title: HOL/Tools/sat.ML  | 
| 17618 | 2  | 
Author: Stephan Merz and Alwen Tiu, QSL Team, LORIA (http://qsl.loria.fr)  | 
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
27115 
diff
changeset
 | 
3  | 
Author: Tjark Weber, TU Muenchen  | 
| 17618 | 4  | 
|
5  | 
Proof reconstruction from SAT solvers.  | 
|
6  | 
||
7  | 
Description:  | 
|
8  | 
This file defines several tactics to invoke a proof-producing  | 
|
9  | 
SAT solver on a propositional goal in clausal form.  | 
|
10  | 
||
11  | 
We use a sequent presentation of clauses to speed up resolution  | 
|
| 17695 | 12  | 
proof reconstruction.  | 
13  | 
We call such clauses "raw clauses", which are of the form  | 
|
| 
20440
 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 
webertj 
parents: 
20371 
diff
changeset
 | 
14  | 
[x1, ..., xn, P] |- False  | 
| 
19236
 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 
webertj 
parents: 
17843 
diff
changeset
 | 
15  | 
(note the use of |- instead of ==>, i.e. of Isabelle's (meta-)hyps here),  | 
| 55239 | 16  | 
where each xi is a literal (see also comments in cnf.ML).  | 
| 17618 | 17  | 
|
| 
19236
 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 
webertj 
parents: 
17843 
diff
changeset
 | 
18  | 
This does not work for goals containing schematic variables!  | 
| 
 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 
webertj 
parents: 
17843 
diff
changeset
 | 
19  | 
|
| 
20039
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
20  | 
The tactic produces a clause representation of the given goal  | 
| 
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
21  | 
in DIMACS format and invokes a SAT solver, which should return  | 
| 
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
22  | 
a proof consisting of a sequence of resolution steps, indicating  | 
| 
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
23  | 
the two input clauses, and resulting in new clauses, leading to  | 
| 
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
24  | 
the empty clause (i.e. "False"). The tactic replays this proof  | 
| 
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
25  | 
in Isabelle and thus solves the overall goal.  | 
| 17618 | 26  | 
|
| 
20440
 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 
webertj 
parents: 
20371 
diff
changeset
 | 
27  | 
There are three SAT tactics available. They differ in the CNF transformation  | 
| 
20039
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
28  | 
used. "sat_tac" uses naive CNF transformation to transform the theorem to be  | 
| 
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
29  | 
proved before giving it to the SAT solver. The naive transformation in the  | 
| 
20440
 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 
webertj 
parents: 
20371 
diff
changeset
 | 
30  | 
worst case can lead to an exponential blow up in formula size. Another  | 
| 
20039
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
31  | 
tactic, "satx_tac", uses "definitional CNF transformation" which attempts to  | 
| 
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
32  | 
produce a formula of linear size increase compared to the input formula, at  | 
| 55239 | 33  | 
the cost of possibly introducing new variables. See cnf.ML for more  | 
| 
20440
 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 
webertj 
parents: 
20371 
diff
changeset
 | 
34  | 
comments on the CNF transformation. "rawsat_tac" should be used with  | 
| 
 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 
webertj 
parents: 
20371 
diff
changeset
 | 
35  | 
caution: no CNF transformation is performed, and the tactic's behavior is  | 
| 
 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 
webertj 
parents: 
20371 
diff
changeset
 | 
36  | 
undefined if the subgoal is not already given as [| C1; ...; Cn |] ==> False,  | 
| 
 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 
webertj 
parents: 
20371 
diff
changeset
 | 
37  | 
where each Ci is a disjunction.  | 
| 17618 | 38  | 
|
| 
20039
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
39  | 
The SAT solver to be used can be set via the "solver" reference. See  | 
| 
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
40  | 
sat_solvers.ML for possible values, and etc/settings for required (solver-  | 
| 
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
41  | 
dependent) configuration settings. To replay SAT proofs in Isabelle, you  | 
| 
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
42  | 
must of course use a proof-producing SAT solver in the first place.  | 
| 
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
43  | 
|
| 52059 | 44  | 
Proofs are replayed only if "quick_and_dirty" is false. If  | 
45  | 
"quick_and_dirty" is true, the theorem (in case the SAT solver claims its  | 
|
| 
20039
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
46  | 
negation to be unsatisfiable) is proved via an oracle.  | 
| 17618 | 47  | 
*)  | 
48  | 
||
49  | 
signature SAT =  | 
|
50  | 
sig  | 
|
| 55240 | 51  | 
val trace: bool Config.T (* print trace messages *)  | 
52  | 
val solver: string Config.T (* name of SAT solver to be used *)  | 
|
| 41447 | 53  | 
val counter: int Unsynchronized.ref (* output: number of resolution steps during last proof replay *)  | 
54  | 
val rawsat_thm: Proof.context -> cterm list -> thm  | 
|
55  | 
val rawsat_tac: Proof.context -> int -> tactic  | 
|
56  | 
val sat_tac: Proof.context -> int -> tactic  | 
|
57  | 
val satx_tac: Proof.context -> int -> tactic  | 
|
| 17618 | 58  | 
end  | 
59  | 
||
| 55239 | 60  | 
structure SAT : SAT =  | 
| 
17622
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
61  | 
struct  | 
| 17618 | 62  | 
|
| 67149 | 63  | 
val trace = Attrib.setup_config_bool \<^binding>\<open>sat_trace\<close> (K false);  | 
| 
17622
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
64  | 
|
| 55240 | 65  | 
fun cond_tracing ctxt msg =  | 
66  | 
if Config.get ctxt trace then tracing (msg ()) else ();  | 
|
67  | 
||
| 67149 | 68  | 
val solver = Attrib.setup_config_string \<^binding>\<open>sat_solver\<close> (K "cdclite");  | 
| 32740 | 69  | 
(*see HOL/Tools/sat_solver.ML for possible values*)  | 
| 
20039
 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 
webertj 
parents: 
19976 
diff
changeset
 | 
70  | 
|
| 32740 | 71  | 
val counter = Unsynchronized.ref 0;  | 
| 
17622
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
72  | 
|
| 
20278
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
73  | 
(* ------------------------------------------------------------------------- *)  | 
| 
21768
 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 
webertj 
parents: 
21756 
diff
changeset
 | 
74  | 
(* lit_ord: an order on integers that considers their absolute values only, *)  | 
| 
 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 
webertj 
parents: 
21756 
diff
changeset
 | 
75  | 
(* thereby treating integers that represent the same atom (positively *)  | 
| 
 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 
webertj 
parents: 
21756 
diff
changeset
 | 
76  | 
(* or negatively) as equal *)  | 
| 
 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 
webertj 
parents: 
21756 
diff
changeset
 | 
77  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 
webertj 
parents: 
21756 
diff
changeset
 | 
78  | 
|
| 41447 | 79  | 
fun lit_ord (i, j) = int_ord (abs i, abs j);  | 
| 
21768
 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 
webertj 
parents: 
21756 
diff
changeset
 | 
80  | 
|
| 
 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 
webertj 
parents: 
21756 
diff
changeset
 | 
81  | 
(* ------------------------------------------------------------------------- *)  | 
| 
20278
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
82  | 
(* CLAUSE: during proof reconstruction, three kinds of clauses are *)  | 
| 
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
83  | 
(* distinguished: *)  | 
| 
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
84  | 
(* 1. NO_CLAUSE: clause not proved (yet) *)  | 
| 
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
85  | 
(* 2. ORIG_CLAUSE: a clause as it occurs in the original problem *)  | 
| 
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
86  | 
(* 3. RAW_CLAUSE: a raw clause, with additional precomputed information *)  | 
| 
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
87  | 
(* (a mapping from int's to its literals) for faster proof *)  | 
| 
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
88  | 
(* reconstruction *)  | 
| 
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
89  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
90  | 
|
| 41447 | 91  | 
datatype CLAUSE =  | 
92  | 
NO_CLAUSE  | 
|
93  | 
| ORIG_CLAUSE of thm  | 
|
94  | 
| RAW_CLAUSE of thm * (int * cterm) list;  | 
|
| 
20278
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
95  | 
|
| 
17622
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
96  | 
(* ------------------------------------------------------------------------- *)  | 
| 
17809
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
97  | 
(* resolve_raw_clauses: given a non-empty list of raw clauses, we fold *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
98  | 
(* resolution over the list (starting with its head), i.e. with two raw *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
99  | 
(* clauses *)  | 
| 
20440
 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 
webertj 
parents: 
20371 
diff
changeset
 | 
100  | 
(* [P, x1, ..., a, ..., xn] |- False *)  | 
| 
17622
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
101  | 
(* and *)  | 
| 
20440
 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 
webertj 
parents: 
20371 
diff
changeset
 | 
102  | 
(* [Q, y1, ..., a', ..., ym] |- False *)  | 
| 
19976
 
aa35f8e27c73
comments fixed, minor optimization wrt. certifying terms
 
webertj 
parents: 
19553 
diff
changeset
 | 
103  | 
(* (where a and a' are dual to each other), we convert the first clause *)  | 
| 
 
aa35f8e27c73
comments fixed, minor optimization wrt. certifying terms
 
webertj 
parents: 
19553 
diff
changeset
 | 
104  | 
(* to *)  | 
| 
20440
 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 
webertj 
parents: 
20371 
diff
changeset
 | 
105  | 
(* [P, x1, ..., xn] |- a ==> False , *)  | 
| 
19976
 
aa35f8e27c73
comments fixed, minor optimization wrt. certifying terms
 
webertj 
parents: 
19553 
diff
changeset
 | 
106  | 
(* the second clause to *)  | 
| 
20440
 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 
webertj 
parents: 
20371 
diff
changeset
 | 
107  | 
(* [Q, y1, ..., ym] |- a' ==> False *)  | 
| 
19976
 
aa35f8e27c73
comments fixed, minor optimization wrt. certifying terms
 
webertj 
parents: 
19553 
diff
changeset
 | 
108  | 
(* and then perform resolution with *)  | 
| 
 
aa35f8e27c73
comments fixed, minor optimization wrt. certifying terms
 
webertj 
parents: 
19553 
diff
changeset
 | 
109  | 
(* [| ?P ==> False; ~?P ==> False |] ==> False *)  | 
| 
 
aa35f8e27c73
comments fixed, minor optimization wrt. certifying terms
 
webertj 
parents: 
19553 
diff
changeset
 | 
110  | 
(* to produce *)  | 
| 
20440
 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 
webertj 
parents: 
20371 
diff
changeset
 | 
111  | 
(* [P, Q, x1, ..., xn, y1, ..., ym] |- False *)  | 
| 
21768
 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 
webertj 
parents: 
21756 
diff
changeset
 | 
112  | 
(* Each clause is accompanied with an association list mapping integers *)  | 
| 
 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 
webertj 
parents: 
21756 
diff
changeset
 | 
113  | 
(* (positive for positive literals, negative for negative literals, and *)  | 
| 
 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 
webertj 
parents: 
21756 
diff
changeset
 | 
114  | 
(* the same absolute value for dual literals) to the actual literals as *)  | 
| 
 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 
webertj 
parents: 
21756 
diff
changeset
 | 
115  | 
(* cterms. *)  | 
| 
17622
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
116  | 
(* ------------------------------------------------------------------------- *)  | 
| 17618 | 117  | 
|
| 55236 | 118  | 
fun resolve_raw_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: 
32010 
diff
changeset
 | 
119  | 
      raise THM ("Proof reconstruction failed (empty list of resolvents)!", 0, [])
 | 
| 55236 | 120  | 
| resolve_raw_clauses ctxt (c::cs) =  | 
| 41447 | 121  | 
let  | 
122  | 
(* merges two sorted lists wrt. 'lit_ord', suppressing duplicates *)  | 
|
123  | 
fun merge xs [] = xs  | 
|
124  | 
| merge [] ys = ys  | 
|
125  | 
| merge (x :: xs) (y :: ys) =  | 
|
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58839 
diff
changeset
 | 
126  | 
(case lit_ord (apply2 fst (x, y)) of  | 
| 41447 | 127  | 
LESS => x :: merge xs (y :: ys)  | 
128  | 
| EQUAL => x :: merge xs ys  | 
|
129  | 
| GREATER => y :: merge (x :: xs) ys)  | 
|
| 
21768
 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 
webertj 
parents: 
21756 
diff
changeset
 | 
130  | 
|
| 41447 | 131  | 
(* find out which two hyps are used in the resolution *)  | 
132  | 
fun find_res_hyps ([], _) _ =  | 
|
133  | 
              raise THM ("Proof reconstruction failed (no literal for resolution)!", 0, [])
 | 
|
134  | 
| find_res_hyps (_, []) _ =  | 
|
135  | 
              raise THM ("Proof reconstruction failed (no literal for resolution)!", 0, [])
 | 
|
136  | 
| find_res_hyps (h1 :: hyps1, h2 :: hyps2) acc =  | 
|
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58839 
diff
changeset
 | 
137  | 
(case lit_ord (apply2 fst (h1, h2)) of  | 
| 41447 | 138  | 
LESS => find_res_hyps (hyps1, h2 :: hyps2) (h1 :: acc)  | 
139  | 
| EQUAL =>  | 
|
140  | 
let  | 
|
141  | 
val (i1, chyp1) = h1  | 
|
142  | 
val (i2, chyp2) = h2  | 
|
143  | 
in  | 
|
144  | 
if i1 = ~ i2 then  | 
|
145  | 
(i1 < 0, chyp1, chyp2, rev acc @ merge hyps1 hyps2)  | 
|
146  | 
else (* i1 = i2 *)  | 
|
147  | 
find_res_hyps (hyps1, hyps2) (h1 :: acc)  | 
|
148  | 
end  | 
|
149  | 
| GREATER => find_res_hyps (h1 :: hyps1, hyps2) (h2 :: acc))  | 
|
| 
19976
 
aa35f8e27c73
comments fixed, minor optimization wrt. certifying terms
 
webertj 
parents: 
19553 
diff
changeset
 | 
150  | 
|
| 41447 | 151  | 
fun resolution (c1, hyps1) (c2, hyps2) =  | 
152  | 
let  | 
|
153  | 
val _ =  | 
|
| 55240 | 154  | 
cond_tracing ctxt (fn () =>  | 
| 61268 | 155  | 
"Resolving clause: " ^ Thm.string_of_thm ctxt c1 ^  | 
| 
77863
 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 
wenzelm 
parents: 
77808 
diff
changeset
 | 
156  | 
" (hyps: " ^ ML_Syntax.print_list (Syntax.string_of_term ctxt) (Thm.hyps_of c1) ^  | 
| 61268 | 157  | 
")\nwith clause: " ^ Thm.string_of_thm ctxt c2 ^  | 
| 
77863
 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 
wenzelm 
parents: 
77808 
diff
changeset
 | 
158  | 
" (hyps: " ^ ML_Syntax.print_list (Syntax.string_of_term ctxt) (Thm.hyps_of c2) ^ ")")  | 
| 
17809
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
159  | 
|
| 41447 | 160  | 
(* the two literals used for resolution *)  | 
161  | 
val (hyp1_is_neg, hyp1, hyp2, new_hyps) = find_res_hyps (hyps1, hyps2) []  | 
|
| 
19236
 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 
webertj 
parents: 
17843 
diff
changeset
 | 
162  | 
|
| 41447 | 163  | 
val c1' = Thm.implies_intr hyp1 c1 (* Gamma1 |- hyp1 ==> False *)  | 
164  | 
val c2' = Thm.implies_intr hyp2 c2 (* Gamma2 |- hyp2 ==> False *)  | 
|
| 17618 | 165  | 
|
| 74610 | 166  | 
val res_thm =  | 
| 41447 | 167  | 
let  | 
| 74610 | 168  | 
val P =  | 
| 41447 | 169  | 
snd (Thm.dest_comb (if hyp1_is_neg then hyp2 else hyp1)) (* strip Trueprop *)  | 
170  | 
in  | 
|
| 74610 | 171  | 
\<^instantiate>\<open>P in  | 
172  | 
lemma \<open>(P \<Longrightarrow> False) \<Longrightarrow> (\<not> P \<Longrightarrow> False) \<Longrightarrow> False\<close> by (rule case_split)\<close>  | 
|
| 41447 | 173  | 
end  | 
| 
19236
 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 
webertj 
parents: 
17843 
diff
changeset
 | 
174  | 
|
| 41447 | 175  | 
val _ =  | 
| 55240 | 176  | 
cond_tracing ctxt  | 
| 61268 | 177  | 
(fn () => "Resolution theorem: " ^ Thm.string_of_thm ctxt res_thm)  | 
| 
19236
 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 
webertj 
parents: 
17843 
diff
changeset
 | 
178  | 
|
| 41447 | 179  | 
(* Gamma1, Gamma2 |- False *)  | 
180  | 
val c_new =  | 
|
181  | 
Thm.implies_elim  | 
|
182  | 
(Thm.implies_elim res_thm (if hyp1_is_neg then c2' else c1'))  | 
|
183  | 
(if hyp1_is_neg then c1' else c2')  | 
|
| 
20278
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
184  | 
|
| 41447 | 185  | 
val _ =  | 
| 55240 | 186  | 
cond_tracing ctxt (fn () =>  | 
| 61268 | 187  | 
"Resulting clause: " ^ Thm.string_of_thm ctxt c_new ^  | 
| 55240 | 188  | 
" (hyps: " ^  | 
| 
77863
 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 
wenzelm 
parents: 
77808 
diff
changeset
 | 
189  | 
ML_Syntax.print_list (Syntax.string_of_term ctxt) (Thm.hyps_of c_new) ^ ")")  | 
| 55240 | 190  | 
|
| 41447 | 191  | 
val _ = Unsynchronized.inc counter  | 
192  | 
in  | 
|
193  | 
(c_new, new_hyps)  | 
|
194  | 
end  | 
|
195  | 
in  | 
|
196  | 
fold resolution cs c  | 
|
197  | 
end;  | 
|
| 17618 | 198  | 
|
| 
17622
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
199  | 
(* ------------------------------------------------------------------------- *)  | 
| 
17809
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
200  | 
(* replay_proof: replays the resolution proof returned by the SAT solver; *)  | 
| 56853 | 201  | 
(* cf. SAT_Solver.proof for details of the proof format. Updates the *)  | 
| 
17809
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
202  | 
(* 'clauses' array with derived clauses, and returns the derived clause *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
203  | 
(* at index 'empty_id' (which should just be "False" if proof *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
204  | 
(* reconstruction was successful, with the used clauses as hyps). *)  | 
| 
20278
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
205  | 
(* 'atom_table' must contain an injective mapping from all atoms that *)  | 
| 
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
206  | 
(* occur (as part of a literal) in 'clauses' to positive integers. *)  | 
| 
17622
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
207  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
208  | 
|
| 55236 | 209  | 
fun replay_proof ctxt atom_table clauses (clause_table, empty_id) =  | 
| 41447 | 210  | 
let  | 
211  | 
fun index_of_literal chyp =  | 
|
212  | 
(case (HOLogic.dest_Trueprop o Thm.term_of) chyp of  | 
|
| 67149 | 213  | 
(Const (\<^const_name>\<open>Not\<close>, _) $ atom) =>  | 
| 41447 | 214  | 
SOME (~ (the (Termtab.lookup atom_table atom)))  | 
215  | 
| atom => SOME (the (Termtab.lookup atom_table atom)))  | 
|
216  | 
handle TERM _ => NONE; (* 'chyp' is not a literal *)  | 
|
| 
20278
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
217  | 
|
| 41447 | 218  | 
fun prove_clause id =  | 
| 77888 | 219  | 
(case Array.nth clauses id of  | 
| 41447 | 220  | 
RAW_CLAUSE clause => clause  | 
221  | 
| ORIG_CLAUSE thm =>  | 
|
222  | 
(* convert the original clause *)  | 
|
223  | 
let  | 
|
| 55240 | 224  | 
val _ = cond_tracing ctxt (fn () => "Using original clause #" ^ string_of_int id)  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59352 
diff
changeset
 | 
225  | 
val raw = CNF.clause2raw_thm ctxt thm  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58839 
diff
changeset
 | 
226  | 
val hyps = sort (lit_ord o apply2 fst) (map_filter (fn chyp =>  | 
| 60949 | 227  | 
Option.map (rpair chyp) (index_of_literal chyp)) (Thm.chyps_of raw))  | 
| 41447 | 228  | 
val clause = (raw, hyps)  | 
| 77896 | 229  | 
val _ = Array.upd clauses id (RAW_CLAUSE clause)  | 
| 41447 | 230  | 
in  | 
231  | 
clause  | 
|
232  | 
end  | 
|
233  | 
| NO_CLAUSE =>  | 
|
234  | 
(* prove the clause, using information from 'clause_table' *)  | 
|
235  | 
let  | 
|
| 55240 | 236  | 
val _ = cond_tracing ctxt (fn () => "Proving clause #" ^ string_of_int id ^ " ...")  | 
| 41447 | 237  | 
val ids = the (Inttab.lookup clause_table id)  | 
| 55236 | 238  | 
val clause = resolve_raw_clauses ctxt (map prove_clause ids)  | 
| 77896 | 239  | 
val _ = Array.upd clauses id (RAW_CLAUSE clause)  | 
| 41447 | 240  | 
val _ =  | 
| 55240 | 241  | 
cond_tracing ctxt  | 
242  | 
(fn () => "Replay chain successful; clause stored at #" ^ string_of_int id)  | 
|
| 41447 | 243  | 
in  | 
244  | 
clause  | 
|
245  | 
end)  | 
|
| 17618 | 246  | 
|
| 41447 | 247  | 
val _ = counter := 0  | 
248  | 
val empty_clause = fst (prove_clause empty_id)  | 
|
249  | 
val _ =  | 
|
| 55240 | 250  | 
cond_tracing ctxt (fn () =>  | 
251  | 
"Proof reconstruction successful; " ^  | 
|
252  | 
string_of_int (!counter) ^ " resolution step(s) total.")  | 
|
| 41447 | 253  | 
in  | 
254  | 
empty_clause  | 
|
255  | 
end;  | 
|
| 
17622
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
256  | 
|
| 
20278
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
257  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
258  | 
(* string_of_prop_formula: return a human-readable string representation of *)  | 
| 
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
259  | 
(* a 'prop_formula' (just for tracing) *)  | 
| 
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
260  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 
webertj 
parents: 
20170 
diff
changeset
 | 
261  | 
|
| 41471 | 262  | 
fun string_of_prop_formula Prop_Logic.True = "True"  | 
263  | 
| string_of_prop_formula Prop_Logic.False = "False"  | 
|
264  | 
| string_of_prop_formula (Prop_Logic.BoolVar i) = "x" ^ string_of_int i  | 
|
265  | 
| string_of_prop_formula (Prop_Logic.Not fm) = "~" ^ string_of_prop_formula fm  | 
|
266  | 
| string_of_prop_formula (Prop_Logic.Or (fm1, fm2)) =  | 
|
| 41447 | 267  | 
      "(" ^ string_of_prop_formula fm1 ^ " v " ^ string_of_prop_formula fm2 ^ ")"
 | 
| 41471 | 268  | 
| string_of_prop_formula (Prop_Logic.And (fm1, fm2)) =  | 
| 41447 | 269  | 
      "(" ^ string_of_prop_formula fm1 ^ " & " ^ string_of_prop_formula fm2 ^ ")";
 | 
| 
17622
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
270  | 
|
| 
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
271  | 
(* ------------------------------------------------------------------------- *)  | 
| 21267 | 272  | 
(* rawsat_thm: run external SAT solver with the given clauses. Reconstructs *)  | 
273  | 
(* a proof from the resulting proof trace of the SAT solver. The *)  | 
|
274  | 
(* theorem returned is just "False" (with some of the given clauses as *)  | 
|
275  | 
(* hyps). *)  | 
|
276  | 
(* ------------------------------------------------------------------------- *)  | 
|
277  | 
||
| 
32432
 
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
278  | 
fun rawsat_thm ctxt clauses =  | 
| 41447 | 279  | 
let  | 
280  | 
(* remove premises that equal "True" *)  | 
|
281  | 
val clauses' = filter (fn clause =>  | 
|
| 67149 | 282  | 
(not_equal \<^term>\<open>True\<close> o HOLogic.dest_Trueprop o Thm.term_of) clause  | 
| 41447 | 283  | 
        handle TERM ("dest_Trueprop", _) => true) clauses
 | 
284  | 
(* remove non-clausal premises -- of course this shouldn't actually *)  | 
|
285  | 
(* remove anything as long as 'rawsat_tac' is only called after the *)  | 
|
286  | 
(* premises have been converted to clauses *)  | 
|
287  | 
val clauses'' = filter (fn clause =>  | 
|
| 55239 | 288  | 
((CNF.is_clause o HOLogic.dest_Trueprop o Thm.term_of) clause  | 
| 41447 | 289  | 
        handle TERM ("dest_Trueprop", _) => false)
 | 
| 
59352
 
63c02d051661
tuned warnings: observe Context_Position.is_visible;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
290  | 
orelse  | 
| 
 
63c02d051661
tuned warnings: observe Context_Position.is_visible;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
291  | 
(if Context_Position.is_visible ctxt then  | 
| 
 
63c02d051661
tuned warnings: observe Context_Position.is_visible;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
292  | 
          warning ("Ignoring non-clausal premise " ^ Syntax.string_of_term ctxt (Thm.term_of clause))
 | 
| 
 
63c02d051661
tuned warnings: observe Context_Position.is_visible;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
293  | 
else (); false)) clauses'  | 
| 41447 | 294  | 
(* remove trivial clauses -- this is necessary because zChaff removes *)  | 
295  | 
(* trivial clauses during preprocessing, and otherwise our clause *)  | 
|
296  | 
(* numbering would be off *)  | 
|
297  | 
val nontrivial_clauses =  | 
|
| 55239 | 298  | 
filter (not o CNF.clause_is_trivial o HOLogic.dest_Trueprop o Thm.term_of) clauses''  | 
| 41447 | 299  | 
(* sort clauses according to the term order -- an optimization, *)  | 
300  | 
(* useful because forming the union of hypotheses, as done by *)  | 
|
301  | 
(* Conjunction.intr_balanced and fold Thm.weaken below, is quadratic for *)  | 
|
302  | 
(* terms sorted in descending order, while only linear for terms *)  | 
|
303  | 
(* sorted in ascending order *)  | 
|
| 67559 | 304  | 
val sorted_clauses = sort Thm.fast_term_ord nontrivial_clauses  | 
| 41447 | 305  | 
val _ =  | 
| 55240 | 306  | 
cond_tracing ctxt (fn () =>  | 
307  | 
"Sorted non-trivial clauses:\n" ^  | 
|
308  | 
cat_lines (map (Syntax.string_of_term ctxt o Thm.term_of) sorted_clauses))  | 
|
| 41471 | 309  | 
(* translate clauses from HOL terms to Prop_Logic.prop_formula *)  | 
| 41447 | 310  | 
val (fms, atom_table) =  | 
| 41471 | 311  | 
fold_map (Prop_Logic.prop_formula_of_term o HOLogic.dest_Trueprop o Thm.term_of)  | 
| 41447 | 312  | 
sorted_clauses Termtab.empty  | 
313  | 
val _ =  | 
|
| 55240 | 314  | 
cond_tracing ctxt  | 
315  | 
(fn () => "Invoking SAT solver on clauses:\n" ^ cat_lines (map string_of_prop_formula fms))  | 
|
| 41471 | 316  | 
val fm = Prop_Logic.all fms  | 
| 41447 | 317  | 
fun make_quick_and_dirty_thm () =  | 
318  | 
let  | 
|
319  | 
val _ =  | 
|
| 55240 | 320  | 
cond_tracing ctxt  | 
321  | 
(fn () => "quick_and_dirty is set: proof reconstruction skipped, using oracle instead")  | 
|
| 67149 | 322  | 
val False_thm = Skip_Proof.make_thm_cterm \<^cprop>\<open>False\<close>  | 
| 41447 | 323  | 
in  | 
324  | 
(* 'fold Thm.weaken (rev sorted_clauses)' is linear, while 'fold *)  | 
|
325  | 
(* Thm.weaken sorted_clauses' would be quadratic, since we sorted *)  | 
|
326  | 
(* clauses in ascending order (which is linear for *)  | 
|
327  | 
(* 'Conjunction.intr_balanced', used below) *)  | 
|
328  | 
fold Thm.weaken (rev sorted_clauses) False_thm  | 
|
329  | 
end  | 
|
330  | 
in  | 
|
331  | 
case  | 
|
| 56817 | 332  | 
let  | 
333  | 
val the_solver = Config.get ctxt solver  | 
|
334  | 
val _ = cond_tracing ctxt (fn () => "Invoking solver " ^ the_solver)  | 
|
| 56853 | 335  | 
in SAT_Solver.invoke_solver the_solver fm end  | 
| 41447 | 336  | 
of  | 
| 56853 | 337  | 
SAT_Solver.UNSATISFIABLE (SOME (clause_table, empty_id)) =>  | 
| 55240 | 338  | 
(cond_tracing ctxt (fn () =>  | 
339  | 
"Proof trace from SAT solver:\n" ^  | 
|
340  | 
"clauses: " ^ ML_Syntax.print_list  | 
|
341  | 
(ML_Syntax.print_pair ML_Syntax.print_int (ML_Syntax.print_list ML_Syntax.print_int))  | 
|
342  | 
(Inttab.dest clause_table) ^ "\n" ^  | 
|
343  | 
"empty clause: " ^ string_of_int empty_id);  | 
|
| 52059 | 344  | 
if Config.get ctxt quick_and_dirty then  | 
| 41447 | 345  | 
make_quick_and_dirty_thm ()  | 
346  | 
else  | 
|
347  | 
let  | 
|
348  | 
(* optimization: convert the given clauses to "[c_1 && ... && c_n] |- c_i"; *)  | 
|
349  | 
(* this avoids accumulation of hypotheses during resolution *)  | 
|
350  | 
(* [c_1, ..., c_n] |- c_1 && ... && c_n *)  | 
|
351  | 
val clauses_thm = Conjunction.intr_balanced (map Thm.assume sorted_clauses)  | 
|
352  | 
(* [c_1 && ... && c_n] |- c_1 && ... && c_n *)  | 
|
| 77872 | 353  | 
val clauses_cprop = Thm.cprop_of clauses_thm  | 
| 41447 | 354  | 
(* [[c_1 && ... && c_n] |- c_1, ..., [c_1 && ... && c_n] |- c_n] *)  | 
| 77872 | 355  | 
val cnf_clauses =  | 
356  | 
Conjunction.elim_balanced (length sorted_clauses) (Thm.assume clauses_cprop)  | 
|
| 41447 | 357  | 
(* initialize the clause array with the given clauses *)  | 
| 52049 | 358  | 
val max_idx = fst (the (Inttab.max clause_table))  | 
| 41447 | 359  | 
val clause_arr = Array.array (max_idx + 1, NO_CLAUSE)  | 
360  | 
val _ =  | 
|
| 77896 | 361  | 
fold (fn thm => fn i => (Array.upd clause_arr i (ORIG_CLAUSE thm); i + 1))  | 
| 41447 | 362  | 
cnf_clauses 0  | 
363  | 
(* replay the proof to derive the empty clause *)  | 
|
364  | 
(* [c_1 && ... && c_n] |- False *)  | 
|
| 55236 | 365  | 
val raw_thm = replay_proof ctxt atom_table clause_arr (clause_table, empty_id)  | 
| 41447 | 366  | 
in  | 
367  | 
(* [c_1, ..., c_n] |- False *)  | 
|
| 77872 | 368  | 
Thm.implies_elim (Thm.implies_intr clauses_cprop raw_thm) clauses_thm  | 
| 41447 | 369  | 
end)  | 
| 56853 | 370  | 
| SAT_Solver.UNSATISFIABLE NONE =>  | 
| 52059 | 371  | 
if Config.get ctxt quick_and_dirty then  | 
| 
59352
 
63c02d051661
tuned warnings: observe Context_Position.is_visible;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
372  | 
(if Context_Position.is_visible ctxt then  | 
| 
 
63c02d051661
tuned warnings: observe Context_Position.is_visible;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
373  | 
warning "SAT solver claims the formula to be unsatisfiable, but did not provide a proof"  | 
| 
 
63c02d051661
tuned warnings: observe Context_Position.is_visible;
 
wenzelm 
parents: 
59058 
diff
changeset
 | 
374  | 
else ();  | 
| 41447 | 375  | 
make_quick_and_dirty_thm ())  | 
376  | 
else  | 
|
377  | 
          raise THM ("SAT solver claims the formula to be unsatisfiable, but did not provide a proof", 0, [])
 | 
|
| 56853 | 378  | 
| SAT_Solver.SATISFIABLE assignment =>  | 
| 41447 | 379  | 
let  | 
380  | 
val msg =  | 
|
381  | 
"SAT solver found a countermodel:\n" ^  | 
|
382  | 
(commas o map (fn (term, idx) =>  | 
|
| 69593 | 383  | 
Syntax.string_of_term_global \<^theory> term ^ ": " ^  | 
| 41447 | 384  | 
(case assignment idx of NONE => "arbitrary"  | 
385  | 
| SOME true => "true" | SOME false => "false")))  | 
|
386  | 
(Termtab.dest atom_table)  | 
|
387  | 
in  | 
|
388  | 
raise THM (msg, 0, [])  | 
|
389  | 
end  | 
|
| 56853 | 390  | 
| SAT_Solver.UNKNOWN =>  | 
| 41447 | 391  | 
        raise THM ("SAT solver failed to decide the formula", 0, [])
 | 
392  | 
end;  | 
|
| 17618 | 393  | 
|
| 
17622
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
394  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
395  | 
(* Tactics *)  | 
| 
 
5d03a69481b6
code reformatted and restructured, many minor modifications
 
webertj 
parents: 
17618 
diff
changeset
 | 
396  | 
(* ------------------------------------------------------------------------- *)  | 
| 17618 | 397  | 
|
| 
17809
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
398  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
399  | 
(* rawsat_tac: solves the i-th subgoal of the proof state; this subgoal *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
400  | 
(* should be of the form *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
401  | 
(* [| c1; c2; ...; ck |] ==> False *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
402  | 
(* where each cj is a non-empty clause (i.e. a disjunction of literals) *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
403  | 
(* or "True" *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
404  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
405  | 
|
| 32232 | 406  | 
fun rawsat_tac ctxt i =  | 
| 
32432
 
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
 
wenzelm 
parents: 
32283 
diff
changeset
 | 
407  | 
  Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
 | 
| 59582 | 408  | 
resolve_tac ctxt' [rawsat_thm ctxt' (map Thm.cprop_of prems)] 1) ctxt i;  | 
| 17618 | 409  | 
|
| 
17809
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
410  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
411  | 
(* pre_cnf_tac: converts the i-th subgoal *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
412  | 
(* [| A1 ; ... ; An |] ==> B *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
413  | 
(* to *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
414  | 
(* [| A1; ... ; An ; ~B |] ==> False *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
415  | 
(* (handling meta-logical connectives in B properly before negating), *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
416  | 
(* then replaces meta-logical connectives in the premises (i.e. "==>", *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
417  | 
(* "!!" and "==") by connectives of the HOL object-logic (i.e. by *)  | 
| 
19553
 
9d15911f1893
pre_cnf_tac: beta-eta-normalization restricted to the current subgoal
 
webertj 
parents: 
19534 
diff
changeset
 | 
418  | 
(* "-->", "!", and "="), then performs beta-eta-normalization on the *)  | 
| 
 
9d15911f1893
pre_cnf_tac: beta-eta-normalization restricted to the current subgoal
 
webertj 
parents: 
19534 
diff
changeset
 | 
419  | 
(* subgoal *)  | 
| 
17809
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
420  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
421  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52059 
diff
changeset
 | 
422  | 
fun pre_cnf_tac ctxt =  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59352 
diff
changeset
 | 
423  | 
  resolve_tac ctxt @{thms ccontr} THEN'
 | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52059 
diff
changeset
 | 
424  | 
Object_Logic.atomize_prems_tac ctxt THEN'  | 
| 41447 | 425  | 
CONVERSION Drule.beta_eta_conversion;  | 
| 
17809
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
426  | 
|
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
427  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
428  | 
(* cnfsat_tac: checks if the empty clause "False" occurs among the premises; *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
429  | 
(* if not, eliminates conjunctions (i.e. each clause of the CNF formula *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
430  | 
(* becomes a separate premise), then applies 'rawsat_tac' to solve the *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
431  | 
(* subgoal *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
432  | 
(* ------------------------------------------------------------------------- *)  | 
| 17697 | 433  | 
|
| 32232 | 434  | 
fun cnfsat_tac ctxt i =  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59352 
diff
changeset
 | 
435  | 
(eresolve_tac ctxt [FalseE] i) ORELSE  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59352 
diff
changeset
 | 
436  | 
(REPEAT_DETERM (eresolve_tac ctxt [conjE] i) THEN rawsat_tac ctxt i);  | 
| 17618 | 437  | 
|
| 
17809
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
438  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
439  | 
(* cnfxsat_tac: checks if the empty clause "False" occurs among the *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
440  | 
(* premises; if not, eliminates conjunctions (i.e. each clause of the *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
441  | 
(* CNF formula becomes a separate premise) and existential quantifiers, *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
442  | 
(* then applies 'rawsat_tac' to solve the subgoal *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
443  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
444  | 
|
| 32232 | 445  | 
fun cnfxsat_tac ctxt i =  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59352 
diff
changeset
 | 
446  | 
(eresolve_tac ctxt [FalseE] i) ORELSE  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59352 
diff
changeset
 | 
447  | 
(REPEAT_DETERM (eresolve_tac ctxt [conjE] i ORELSE eresolve_tac ctxt [exE] i) THEN  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59352 
diff
changeset
 | 
448  | 
rawsat_tac ctxt i);  | 
| 17618 | 449  | 
|
| 
17809
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
450  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
451  | 
(* sat_tac: tactic for calling an external SAT solver, taking as input an *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
452  | 
(* arbitrary formula. The input is translated to CNF, possibly causing *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
453  | 
(* an exponential blowup. *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
454  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
455  | 
|
| 32232 | 456  | 
fun sat_tac ctxt i =  | 
| 55239 | 457  | 
pre_cnf_tac ctxt i THEN CNF.cnf_rewrite_tac ctxt i THEN cnfsat_tac ctxt i;  | 
| 
17809
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
458  | 
|
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
459  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
460  | 
(* satx_tac: tactic for calling an external SAT solver, taking as input an *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
461  | 
(* arbitrary formula. The input is translated to CNF, possibly *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
462  | 
(* introducing new literals. *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
463  | 
(* ------------------------------------------------------------------------- *)  | 
| 
 
195045659c06
Tactics sat and satx reimplemented, several improvements
 
webertj 
parents: 
17697 
diff
changeset
 | 
464  | 
|
| 32232 | 465  | 
fun satx_tac ctxt i =  | 
| 55239 | 466  | 
pre_cnf_tac ctxt i THEN CNF.cnfx_rewrite_tac ctxt i THEN cnfxsat_tac ctxt i;  | 
| 17618 | 467  | 
|
| 23533 | 468  | 
end;  |