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