| author | wenzelm | 
| Tue, 09 Oct 2007 17:10:43 +0200 | |
| changeset 24930 | cc2e0e8c81af | 
| parent 23590 | ad95084a5c63 | 
| child 26928 | ca87aff1ad2d | 
| permissions | -rw-r--r-- | 
| 17618 | 1 | (* Title: HOL/Tools/sat_funcs.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Stephan Merz and Alwen Tiu, QSL Team, LORIA (http://qsl.loria.fr) | |
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 4 | Author: Tjark Weber | 
| 19236 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 5 | Copyright 2005-2006 | 
| 17618 | 6 | |
| 7 | ||
| 8 | Proof reconstruction from SAT solvers. | |
| 9 | ||
| 10 | Description: | |
| 11 | This file defines several tactics to invoke a proof-producing | |
| 12 | SAT solver on a propositional goal in clausal form. | |
| 13 | ||
| 14 | We use a sequent presentation of clauses to speed up resolution | |
| 17695 | 15 | proof reconstruction. | 
| 16 | 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 | 17 | [x1, ..., xn, P] |- False | 
| 19236 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 18 | (note the use of |- instead of ==>, i.e. of Isabelle's (meta-)hyps here), | 
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 19 | where each xi is a literal (see also comments in cnf_funcs.ML). | 
| 17618 | 20 | |
| 19236 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 21 | 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 | 22 | |
| 20039 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 webertj parents: 
19976diff
changeset | 23 | The tactic produces a clause representation of the given goal | 
| 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 webertj parents: 
19976diff
changeset | 24 | 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 | 25 | a proof consisting of a sequence of resolution steps, indicating | 
| 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 webertj parents: 
19976diff
changeset | 26 | 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 | 27 | 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 | 28 | in Isabelle and thus solves the overall goal. | 
| 17618 | 29 | |
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 30 | 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 | 31 | 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 | 32 | 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 | 33 | 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 | 34 | tactic, "satx_tac", uses "definitional CNF transformation" which attempts to | 
| 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 webertj parents: 
19976diff
changeset | 35 | produce a formula of linear size increase compared to the input formula, at | 
| 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 webertj parents: 
19976diff
changeset | 36 | the cost of possibly introducing new variables. See cnf_funcs.ML for more | 
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 37 | 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 | 38 | 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 | 39 | 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 | 40 | where each Ci is a disjunction. | 
| 17618 | 41 | |
| 20039 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 webertj parents: 
19976diff
changeset | 42 | 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 | 43 | 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 | 44 | dependent) configuration settings. To replay SAT proofs in Isabelle, you | 
| 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 webertj parents: 
19976diff
changeset | 45 | 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 | 46 | |
| 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 webertj parents: 
19976diff
changeset | 47 | Proofs are replayed only if "!quick_and_dirty" is false. If | 
| 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 webertj parents: 
19976diff
changeset | 48 | "!quick_and_dirty" is true, the theorem (in case the SAT solver claims its | 
| 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 webertj parents: 
19976diff
changeset | 49 | negation to be unsatisfiable) is proved via an oracle. | 
| 17618 | 50 | *) | 
| 51 | ||
| 52 | signature SAT = | |
| 53 | sig | |
| 21267 | 54 | val trace_sat : bool ref (* input: print trace messages *) | 
| 55 | val solver : string ref (* input: name of SAT solver to be used *) | |
| 56 | val counter : int ref (* output: number of resolution steps during last proof replay *) | |
| 57 | val rawsat_thm : cterm list -> thm | |
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 58 | val rawsat_tac : int -> Tactical.tactic | 
| 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 59 | val sat_tac : int -> Tactical.tactic | 
| 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 60 | val satx_tac : int -> Tactical.tactic | 
| 17618 | 61 | end | 
| 62 | ||
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 63 | functor SATFunc (structure cnf : CNF) : SAT = | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 64 | struct | 
| 17618 | 65 | |
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 66 | val trace_sat = ref false; | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 67 | |
| 20170 | 68 | val solver = ref "zchaff_with_proofs"; (* see HOL/Tools/sat_solver.ML for possible values *) | 
| 20039 
4293f932fe83
"solver" reference added to make the SAT solver configurable
 webertj parents: 
19976diff
changeset | 69 | |
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 70 | val counter = ref 0; | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 71 | |
| 21267 | 72 | val resolution_thm = (* "[| ?P ==> False; ~ ?P ==> False |] ==> False" *) | 
| 19236 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 73 | let | 
| 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 74 | val cterm = cterm_of (the_context ()) | 
| 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 75 | 		val Q     = Var (("Q", 0), HOLogic.boolT)
 | 
| 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 76 | val False = HOLogic.false_const | 
| 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 77 | in | 
| 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 78 | Thm.instantiate ([], [(cterm Q, cterm False)]) case_split_thm | 
| 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 79 | end; | 
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 80 | |
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 81 | val cP = cterm_of (theory_of_thm resolution_thm) (Var (("P", 0), HOLogic.boolT));
 | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 82 | |
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 83 | (* ------------------------------------------------------------------------- *) | 
| 21768 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 84 | (* 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 | 85 | (* thereby treating integers that represent the same atom (positively *) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 86 | (* or negatively) as equal *) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 87 | (* ------------------------------------------------------------------------- *) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 88 | |
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 89 | fun lit_ord (i, j) = | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 90 | int_ord (abs i, abs j); | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 91 | |
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 92 | (* ------------------------------------------------------------------------- *) | 
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 93 | (* 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 | 94 | (* distinguished: *) | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 95 | (* 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 | 96 | (* 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 | 97 | (* 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 | 98 | (* (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 | 99 | (* reconstruction *) | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 100 | (* ------------------------------------------------------------------------- *) | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 101 | |
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 102 | datatype CLAUSE = NO_CLAUSE | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 103 | | ORIG_CLAUSE of Thm.thm | 
| 21768 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 104 | | RAW_CLAUSE of Thm.thm * (int * Thm.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 | 105 | |
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 106 | (* ------------------------------------------------------------------------- *) | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 107 | (* 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 | 108 | (* 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 | 109 | (* clauses *) | 
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 110 | (* [P, x1, ..., a, ..., xn] |- False *) | 
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 111 | (* and *) | 
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 112 | (* [Q, y1, ..., a', ..., ym] |- False *) | 
| 19976 
aa35f8e27c73
comments fixed, minor optimization wrt. certifying terms
 webertj parents: 
19553diff
changeset | 113 | (* (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 | 114 | (* to *) | 
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 115 | (* [P, x1, ..., xn] |- a ==> False , *) | 
| 19976 
aa35f8e27c73
comments fixed, minor optimization wrt. certifying terms
 webertj parents: 
19553diff
changeset | 116 | (* the second clause to *) | 
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 117 | (* [Q, y1, ..., ym] |- a' ==> False *) | 
| 19976 
aa35f8e27c73
comments fixed, minor optimization wrt. certifying terms
 webertj parents: 
19553diff
changeset | 118 | (* and then perform resolution with *) | 
| 
aa35f8e27c73
comments fixed, minor optimization wrt. certifying terms
 webertj parents: 
19553diff
changeset | 119 | (* [| ?P ==> False; ~?P ==> False |] ==> False *) | 
| 
aa35f8e27c73
comments fixed, minor optimization wrt. certifying terms
 webertj parents: 
19553diff
changeset | 120 | (* to produce *) | 
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 121 | (* [P, Q, x1, ..., xn, y1, ..., ym] |- False *) | 
| 21768 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 122 | (* Each clause is accompanied with an association list mapping integers *) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 123 | (* (positive for positive literals, negative for negative literals, and *) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 124 | (* 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 | 125 | (* cterms. *) | 
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 126 | (* ------------------------------------------------------------------------- *) | 
| 17618 | 127 | |
| 21768 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 128 | (* (Thm.thm * (int * Thm.cterm) list) list -> Thm.thm * (int * Thm.cterm) list *) | 
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 129 | |
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 130 | fun resolve_raw_clauses [] = | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 131 | 	raise THM ("Proof reconstruction failed (empty list of resolvents)!", 0, [])
 | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 132 | | resolve_raw_clauses (c::cs) = | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 133 | let | 
| 21768 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 134 | (* merges two sorted lists wrt. 'lit_ord', suppressing duplicates *) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 135 | fun merge xs [] = xs | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 136 | | merge [] ys = ys | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 137 | | merge (x::xs) (y::ys) = | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 138 | (case (lit_ord o pairself fst) (x, y) of | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 139 | LESS => x :: merge xs (y::ys) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 140 | | EQUAL => x :: merge xs ys | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 141 | | GREATER => y :: merge (x::xs) ys) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 142 | |
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 143 | (* find out which two hyps are used in the resolution *) | 
| 21768 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 144 | (* (int * Thm.cterm) list * (int * Thm.cterm) list -> (int * Thm.cterm) list -> bool * Thm.cterm * Thm.cterm * (int * Thm.cterm) list *) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 145 | fun find_res_hyps ([], _) _ = | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 146 | 			raise THM ("Proof reconstruction failed (no literal for resolution)!", 0, [])
 | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 147 | | find_res_hyps (_, []) _ = | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 148 | 			raise THM ("Proof reconstruction failed (no literal for resolution)!", 0, [])
 | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 149 | | find_res_hyps (h1 :: hyps1, h2 :: hyps2) acc = | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 150 | (case (lit_ord o pairself fst) (h1, h2) of | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 151 | LESS => find_res_hyps (hyps1, h2 :: hyps2) (h1 :: acc) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 152 | | EQUAL => let | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 153 | val (i1, chyp1) = h1 | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 154 | val (i2, chyp2) = h2 | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 155 | in | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 156 | if i1 = ~ i2 then | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 157 | (i1 < 0, chyp1, chyp2, rev acc @ merge hyps1 hyps2) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 158 | else (* i1 = i2 *) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 159 | find_res_hyps (hyps1, hyps2) (h1 :: acc) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 160 | end | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 161 | | GREATER => find_res_hyps (h1 :: hyps1, hyps2) (h2 :: acc)) | 
| 19976 
aa35f8e27c73
comments fixed, minor optimization wrt. certifying terms
 webertj parents: 
19553diff
changeset | 162 | |
| 21768 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 163 | (* Thm.thm * (int * Thm.cterm) list -> Thm.thm * (int * Thm.cterm) list -> Thm.thm * (int * Thm.cterm) list *) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 164 | fun resolution (c1, hyps1) (c2, hyps2) = | 
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 165 | let | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 166 | val _ = if !trace_sat then | 
| 21756 | 167 | 					tracing ("Resolving clause: " ^ string_of_thm c1 ^ " (hyps: " ^ ML_Syntax.print_list (Sign.string_of_term (theory_of_thm c1)) (#hyps (rep_thm c1))
 | 
| 168 | ^ ")\nwith clause: " ^ string_of_thm c2 ^ " (hyps: " ^ ML_Syntax.print_list (Sign.string_of_term (theory_of_thm c2)) (#hyps (rep_thm c2)) ^ ")") | |
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 169 | else () | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 170 | |
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 171 | (* the two literals used for resolution *) | 
| 21768 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 172 | 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 | 173 | |
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 174 | val c1' = Thm.implies_intr hyp1 c1 (* Gamma1 |- hyp1 ==> False *) | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 175 | val c2' = Thm.implies_intr hyp2 c2 (* Gamma2 |- hyp2 ==> False *) | 
| 17618 | 176 | |
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 177 | val res_thm = (* |- (lit ==> False) ==> (~lit ==> False) ==> False *) | 
| 19236 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 178 | let | 
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 179 | val cLit = snd (Thm.dest_comb (if hyp1_is_neg then hyp2 else hyp1)) (* strip Trueprop *) | 
| 19236 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 180 | in | 
| 19976 
aa35f8e27c73
comments fixed, minor optimization wrt. certifying terms
 webertj parents: 
19553diff
changeset | 181 | Thm.instantiate ([], [(cP, cLit)]) resolution_thm | 
| 19236 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 182 | end | 
| 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 183 | |
| 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 184 | val _ = if !trace_sat then | 
| 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 185 | 					tracing ("Resolution theorem: " ^ string_of_thm res_thm)
 | 
| 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 186 | else () | 
| 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 187 | |
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 188 | (* Gamma1, Gamma2 |- False *) | 
| 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 189 | val c_new = Thm.implies_elim | 
| 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 190 | (Thm.implies_elim res_thm (if hyp1_is_neg then c2' else c1')) | 
| 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 191 | (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 | 192 | |
| 19236 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 193 | val _ = if !trace_sat then | 
| 21756 | 194 | 					tracing ("Resulting clause: " ^ string_of_thm c_new ^ " (hyps: " ^ ML_Syntax.print_list (Sign.string_of_term (theory_of_thm c_new)) (#hyps (rep_thm c_new)) ^ ")")
 | 
| 19236 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 195 | else () | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 196 | val _ = inc counter | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 197 | in | 
| 21768 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 198 | (c_new, new_hyps) | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 199 | end | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 200 | in | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 201 | fold resolution cs c | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 202 | end; | 
| 17618 | 203 | |
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 204 | (* ------------------------------------------------------------------------- *) | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 205 | (* replay_proof: replays the resolution proof returned by the SAT solver; *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 206 | (* cf. SatSolver.proof for details of the proof format. Updates the *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 207 | (* 'clauses' array with derived clauses, and returns the derived clause *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 208 | (* at index 'empty_id' (which should just be "False" if proof *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 209 | (* 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 | 210 | (* '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 | 211 | (* occur (as part of a literal) in 'clauses' to positive integers. *) | 
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 212 | (* ------------------------------------------------------------------------- *) | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 213 | |
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 214 | (* int Termtab.table -> CLAUSE Array.array -> SatSolver.proof -> Thm.thm *) | 
| 17618 | 215 | |
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 216 | fun replay_proof atom_table clauses (clause_table, empty_id) = | 
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 217 | let | 
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 218 | (* Thm.cterm -> int option *) | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 219 | fun index_of_literal chyp = ( | 
| 21586 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 220 | case (HOLogic.dest_Trueprop o Thm.term_of) chyp of | 
| 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 | 		  (Const ("Not", _) $ atom) =>
 | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 222 | SOME (~(valOf (Termtab.lookup atom_table atom))) | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 223 | | atom => | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 224 | SOME (valOf (Termtab.lookup atom_table atom)) | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 225 | ) handle TERM _ => NONE; (* 'chyp' is not a literal *) | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 226 | |
| 21768 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 227 | (* int -> Thm.thm * (int * Thm.cterm) list *) | 
| 17623 
ae4af66b3072
replay_proof optimized: now performs backwards proof search
 webertj parents: 
17622diff
changeset | 228 | fun prove_clause id = | 
| 
ae4af66b3072
replay_proof optimized: now performs backwards proof search
 webertj parents: 
17622diff
changeset | 229 | case Array.sub (clauses, id) of | 
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 230 | RAW_CLAUSE clause => | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 231 | clause | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 232 | | ORIG_CLAUSE thm => | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 233 | (* convert the original clause *) | 
| 17623 
ae4af66b3072
replay_proof optimized: now performs backwards proof search
 webertj parents: 
17622diff
changeset | 234 | let | 
| 21768 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 235 | 				val _      = if !trace_sat then tracing ("Using original clause #" ^ string_of_int id) else ()
 | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 236 | val raw = cnf.clause2raw_thm thm | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 237 | val hyps = sort (lit_ord o pairself fst) (map_filter (fn chyp => | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 238 | Option.map (rpair chyp) (index_of_literal chyp)) (#hyps (Thm.crep_thm raw))) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 239 | val clause = (raw, hyps) | 
| 
69165d27b55b
ordered lists instead of tables for resolving hyps; speedup
 webertj parents: 
21756diff
changeset | 240 | val _ = Array.update (clauses, id, RAW_CLAUSE clause) | 
| 17623 
ae4af66b3072
replay_proof optimized: now performs backwards proof search
 webertj parents: 
17622diff
changeset | 241 | in | 
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 242 | clause | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 243 | end | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 244 | | NO_CLAUSE => | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 245 | (* prove the clause, using information from 'clause_table' *) | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 246 | let | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 247 | 				val _      = if !trace_sat then tracing ("Proving clause #" ^ string_of_int id ^ " ...") else ()
 | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 248 | val ids = valOf (Inttab.lookup clause_table id) | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 249 | val clause = resolve_raw_clauses (map prove_clause ids) | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 250 | val _ = Array.update (clauses, id, RAW_CLAUSE clause) | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 251 | 				val _      = if !trace_sat then tracing ("Replay chain successful; clause stored at #" ^ string_of_int id) else ()
 | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 252 | in | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 253 | clause | 
| 17623 
ae4af66b3072
replay_proof optimized: now performs backwards proof search
 webertj parents: 
17622diff
changeset | 254 | end | 
| 17618 | 255 | |
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 256 | val _ = counter := 0 | 
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 257 | val empty_clause = fst (prove_clause empty_id) | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 258 | 	val _            = if !trace_sat then tracing ("Proof reconstruction successful; " ^ string_of_int (!counter) ^ " resolution step(s) total.") else ()
 | 
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 259 | in | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 260 | empty_clause | 
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 261 | end; | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 262 | |
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 263 | (* ------------------------------------------------------------------------- *) | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 264 | (* 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 | 265 | (* 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 | 266 | (* ------------------------------------------------------------------------- *) | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 267 | |
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 268 | (* PropLogic.prop_formula -> string *) | 
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 269 | |
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 270 | fun string_of_prop_formula PropLogic.True = "True" | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 271 | | string_of_prop_formula PropLogic.False = "False" | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 272 | | string_of_prop_formula (PropLogic.BoolVar i) = "x" ^ string_of_int i | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 273 | | string_of_prop_formula (PropLogic.Not fm) = "~" ^ string_of_prop_formula fm | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 274 |   | string_of_prop_formula (PropLogic.Or (fm1, fm2))  = "(" ^ string_of_prop_formula fm1 ^ " v " ^ string_of_prop_formula fm2 ^ ")"
 | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 275 |   | string_of_prop_formula (PropLogic.And (fm1, fm2)) = "(" ^ string_of_prop_formula fm1 ^ " & " ^ string_of_prop_formula fm2 ^ ")";
 | 
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 276 | |
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 277 | (* ------------------------------------------------------------------------- *) | 
| 21267 | 278 | (* take_prefix: *) | 
| 279 | (* take_prefix n [x_1, ..., x_k] = ([x_1, ..., x_n], [x_n+1, ..., x_k]) *) | |
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 280 | (* ------------------------------------------------------------------------- *) | 
| 17618 | 281 | |
| 21267 | 282 | (* int -> 'a list -> 'a list * 'a list *) | 
| 283 | ||
| 284 | fun take_prefix n xs = | |
| 285 | let | |
| 286 | fun take 0 (rxs, xs) = (rev rxs, xs) | |
| 287 | | take _ (rxs, []) = (rev rxs, []) | |
| 288 | | take n (rxs, x :: xs) = take (n-1) (x :: rxs, xs) | |
| 289 | in | |
| 290 | take n ([], xs) | |
| 291 | end; | |
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 292 | |
| 21267 | 293 | (* ------------------------------------------------------------------------- *) | 
| 294 | (* rawsat_thm: run external SAT solver with the given clauses. Reconstructs *) | |
| 295 | (* a proof from the resulting proof trace of the SAT solver. The *) | |
| 296 | (* theorem returned is just "False" (with some of the given clauses as *) | |
| 297 | (* hyps). *) | |
| 298 | (* ------------------------------------------------------------------------- *) | |
| 299 | ||
| 300 | (* Thm.cterm list -> Thm.thm *) | |
| 301 | ||
| 302 | fun rawsat_thm clauses = | |
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 303 | let | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 304 | (* remove premises that equal "True" *) | 
| 21267 | 305 | val clauses' = filter (fn clause => | 
| 21586 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 306 | (not_equal HOLogic.true_const o HOLogic.dest_Trueprop o Thm.term_of) clause | 
| 21267 | 307 | 			handle TERM ("dest_Trueprop", _) => true) clauses
 | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 308 | (* remove non-clausal premises -- of course this shouldn't actually *) | 
| 21267 | 309 | (* remove anything as long as 'rawsat_tac' is only called after the *) | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 310 | (* premises have been converted to clauses *) | 
| 21267 | 311 | val clauses'' = filter (fn clause => | 
| 21586 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 312 | ((cnf.is_clause o HOLogic.dest_Trueprop o Thm.term_of) clause | 
| 21267 | 313 | 			handle TERM ("dest_Trueprop", _) => false)
 | 
| 314 | orelse ( | |
| 315 | 			warning ("Ignoring non-clausal premise " ^ string_of_cterm clause);
 | |
| 316 | false)) clauses' | |
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 317 | (* remove trivial clauses -- this is necessary because zChaff removes *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 318 | (* trivial clauses during preprocessing, and otherwise our clause *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 319 | (* numbering would be off *) | 
| 21586 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 320 | val nontrivial_clauses = filter (not o cnf.clause_is_trivial o HOLogic.dest_Trueprop o Thm.term_of) clauses'' | 
| 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 321 | (* sort clauses according to the term order -- an optimization, *) | 
| 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 322 | (* useful because forming the union of hypotheses, as done by *) | 
| 23533 | 323 | (* Conjunction.intr_balanced and fold Thm.weaken below, is quadratic for *) | 
| 21586 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 324 | (* terms sorted in descending order, while only linear for terms *) | 
| 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 325 | (* sorted in ascending order *) | 
| 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 326 | val sorted_clauses = sort (Term.fast_term_ord o pairself Thm.term_of) nontrivial_clauses | 
| 21267 | 327 | val _ = if !trace_sat then | 
| 21586 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 328 | 			tracing ("Sorted non-trivial clauses:\n" ^ space_implode "\n" (map string_of_cterm sorted_clauses))
 | 
| 19534 | 329 | else () | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 330 | (* translate clauses from HOL terms to PropLogic.prop_formula *) | 
| 21586 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 331 | val (fms, atom_table) = fold_map (PropLogic.prop_formula_of_term o HOLogic.dest_Trueprop o Thm.term_of) sorted_clauses Termtab.empty | 
| 21267 | 332 | val _ = if !trace_sat then | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 333 | 			tracing ("Invoking SAT solver on clauses:\n" ^ space_implode "\n" (map string_of_prop_formula fms))
 | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 334 | else () | 
| 21586 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 335 | val fm = PropLogic.all fms | 
| 17842 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 336 | (* unit -> Thm.thm *) | 
| 21267 | 337 | fun make_quick_and_dirty_thm () = | 
| 338 | let | |
| 339 | val _ = if !trace_sat then | |
| 340 | tracing "'quick_and_dirty' is set: proof reconstruction skipped, using oracle instead." | |
| 341 | else () | |
| 342 | val False_thm = SkipProof.make_thm (the_context ()) (HOLogic.Trueprop $ HOLogic.false_const) | |
| 343 | in | |
| 21586 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 344 | (* 'fold Thm.weaken (rev sorted_clauses)' is linear, while 'fold *) | 
| 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 345 | (* Thm.weaken sorted_clauses' would be quadratic, since we sorted *) | 
| 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 346 | (* clauses in ascending order (which is linear for *) | 
| 23533 | 347 | (* 'Conjunction.intr_balanced', used below) *) | 
| 21586 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 348 | fold Thm.weaken (rev sorted_clauses) False_thm | 
| 21267 | 349 | end | 
| 17618 | 350 | in | 
| 21268 | 351 | 	case (tracing ("Invoking solver " ^ (!solver)); SatSolver.invoke_solver (!solver) fm) of
 | 
| 17842 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 352 | SatSolver.UNSATISFIABLE (SOME (clause_table, empty_id)) => ( | 
| 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 353 | if !trace_sat then | 
| 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 354 | 			tracing ("Proof trace from SAT solver:\n" ^
 | 
| 21756 | 355 | "clauses: " ^ ML_Syntax.print_list (ML_Syntax.print_pair Int.toString (ML_Syntax.print_list Int.toString)) (Inttab.dest clause_table) ^ "\n" ^ | 
| 21586 
8da782143bde
clauses sorted according to term order (significant speedup in some cases)
 webertj parents: 
21474diff
changeset | 356 | "empty clause: " ^ Int.toString empty_id) | 
| 17842 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 357 | else (); | 
| 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 358 | if !quick_and_dirty then | 
| 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 359 | make_quick_and_dirty_thm () | 
| 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 360 | else | 
| 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 361 | let | 
| 21267 | 362 | (* optimization: convert the given clauses to "[c_1 && ... && c_n] |- c_i"; *) | 
| 363 | (* this avoids accumulation of hypotheses during resolution *) | |
| 364 | (* [c_1, ..., c_n] |- c_1 && ... && c_n *) | |
| 23533 | 365 | val clauses_thm = Conjunction.intr_balanced (map Thm.assume sorted_clauses) | 
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 366 | (* [c_1 && ... && c_n] |- c_1 && ... && c_n *) | 
| 21267 | 367 | val cnf_cterm = cprop_of clauses_thm | 
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 368 | val cnf_thm = Thm.assume cnf_cterm | 
| 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 369 | (* [[c_1 && ... && c_n] |- c_1, ..., [c_1 && ... && c_n] |- c_n] *) | 
| 23533 | 370 | val cnf_clauses = Conjunction.elim_balanced (length sorted_clauses) cnf_thm | 
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 371 | (* initialize the clause array with the given clauses *) | 
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 372 | val max_idx = valOf (Inttab.max_key clause_table) | 
| 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 373 | val clause_arr = Array.array (max_idx + 1, NO_CLAUSE) | 
| 21267 | 374 | val _ = fold (fn thm => fn idx => (Array.update (clause_arr, idx, ORIG_CLAUSE thm); idx+1)) cnf_clauses 0 | 
| 19236 
150e8b0fb991
clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
 webertj parents: 
17843diff
changeset | 375 | (* replay the proof to derive the empty clause *) | 
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 376 | (* [c_1 && ... && c_n] |- False *) | 
| 21268 | 377 | val raw_thm = replay_proof atom_table clause_arr (clause_table, empty_id) | 
| 17842 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 378 | in | 
| 20440 
e6fe74eebda3
faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
 webertj parents: 
20371diff
changeset | 379 | (* [c_1, ..., c_n] |- False *) | 
| 21267 | 380 | Thm.implies_elim (Thm.implies_intr cnf_cterm raw_thm) clauses_thm | 
| 17842 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 381 | end) | 
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 382 | | SatSolver.UNSATISFIABLE NONE => | 
| 17842 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 383 | if !quick_and_dirty then ( | 
| 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 384 | warning "SAT solver claims the formula to be unsatisfiable, but did not provide a proof"; | 
| 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 385 | make_quick_and_dirty_thm () | 
| 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 386 | ) else | 
| 
e661a78472f0
no proof reconstruction when quick_and_dirty is set
 webertj parents: 
17809diff
changeset | 387 | 			raise THM ("SAT solver claims the formula to be unsatisfiable, but did not provide a proof", 0, [])
 | 
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 388 | | SatSolver.SATISFIABLE assignment => | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 389 | let | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 390 | val msg = "SAT solver found a countermodel:\n" | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 391 | ^ (commas | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 392 | o map (fn (term, idx) => | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 393 | Sign.string_of_term (the_context ()) term ^ ": " | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 394 | ^ (case assignment idx of NONE => "arbitrary" | SOME true => "true" | SOME false => "false"))) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 395 | (Termtab.dest atom_table) | 
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 396 | in | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 397 | raise THM (msg, 0, []) | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 398 | end | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 399 | | SatSolver.UNKNOWN => | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 400 | 		raise THM ("SAT solver failed to decide the formula", 0, [])
 | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 401 | end; | 
| 17618 | 402 | |
| 17622 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 403 | (* ------------------------------------------------------------------------- *) | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 404 | (* Tactics *) | 
| 
5d03a69481b6
code reformatted and restructured, many minor modifications
 webertj parents: 
17618diff
changeset | 405 | (* ------------------------------------------------------------------------- *) | 
| 17618 | 406 | |
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 407 | (* ------------------------------------------------------------------------- *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 408 | (* 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 | 409 | (* should be of the form *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 410 | (* [| c1; c2; ...; ck |] ==> False *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 411 | (* 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 | 412 | (* or "True" *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 413 | (* ------------------------------------------------------------------------- *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 414 | |
| 21267 | 415 | fun rawsat_tac i = METAHYPS (fn prems => rtac (rawsat_thm (map cprop_of prems)) 1) i; | 
| 17618 | 416 | |
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 417 | (* ------------------------------------------------------------------------- *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 418 | (* pre_cnf_tac: converts the i-th subgoal *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 419 | (* [| A1 ; ... ; An |] ==> B *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 420 | (* to *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 421 | (* [| A1; ... ; An ; ~B |] ==> False *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 422 | (* (handling meta-logical connectives in B properly before negating), *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 423 | (* then replaces meta-logical connectives in the premises (i.e. "==>", *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 424 | (* "!!" 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 | 425 | (* "-->", "!", and "="), then performs beta-eta-normalization on the *) | 
| 
9d15911f1893
pre_cnf_tac: beta-eta-normalization restricted to the current subgoal
 webertj parents: 
19534diff
changeset | 426 | (* subgoal *) | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 427 | (* ------------------------------------------------------------------------- *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 428 | |
| 23533 | 429 | val pre_cnf_tac = | 
| 430 | rtac ccontr THEN' | |
| 23590 
ad95084a5c63
renamed ObjectLogic.atomize_tac to ObjectLogic.atomize_prems_tac;
 wenzelm parents: 
23533diff
changeset | 431 | ObjectLogic.atomize_prems_tac THEN' | 
| 23533 | 432 | CONVERSION Drule.beta_eta_conversion; | 
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 433 | |
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 434 | (* ------------------------------------------------------------------------- *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 435 | (* cnfsat_tac: checks if the empty clause "False" occurs among the premises; *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 436 | (* if not, eliminates conjunctions (i.e. each clause of the CNF formula *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 437 | (* becomes a separate premise), then applies 'rawsat_tac' to solve the *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 438 | (* subgoal *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 439 | (* ------------------------------------------------------------------------- *) | 
| 17697 | 440 | |
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 441 | fun cnfsat_tac i = | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 442 | (etac FalseE i) ORELSE (REPEAT_DETERM (etac conjE i) THEN rawsat_tac i); | 
| 17618 | 443 | |
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 444 | (* ------------------------------------------------------------------------- *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 445 | (* cnfxsat_tac: checks if the empty clause "False" occurs among the *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 446 | (* premises; if not, eliminates conjunctions (i.e. each clause of the *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 447 | (* CNF formula becomes a separate premise) and existential quantifiers, *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 448 | (* then applies 'rawsat_tac' to solve the subgoal *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 449 | (* ------------------------------------------------------------------------- *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 450 | |
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 451 | fun cnfxsat_tac i = | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 452 | (etac FalseE i) ORELSE | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 453 | (REPEAT_DETERM (etac conjE i ORELSE etac exE i) THEN rawsat_tac i); | 
| 17618 | 454 | |
| 17809 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 455 | (* ------------------------------------------------------------------------- *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 456 | (* 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 | 457 | (* arbitrary formula. The input is translated to CNF, possibly causing *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 458 | (* an exponential blowup. *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 459 | (* ------------------------------------------------------------------------- *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 460 | |
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 461 | fun sat_tac i = | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 462 | pre_cnf_tac i THEN cnf.cnf_rewrite_tac i THEN cnfsat_tac i; | 
| 17809 
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 | (* ------------------------------------------------------------------------- *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 465 | (* 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 | 466 | (* arbitrary formula. The input is translated to CNF, possibly *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 467 | (* introducing new literals. *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 468 | (* ------------------------------------------------------------------------- *) | 
| 
195045659c06
Tactics sat and satx reimplemented, several improvements
 webertj parents: 
17697diff
changeset | 469 | |
| 20278 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 470 | fun satx_tac i = | 
| 
28be10991666
proof reconstruction now uses its own data structure to manage hyps; should be both faster and more robust
 webertj parents: 
20170diff
changeset | 471 | pre_cnf_tac i THEN cnf.cnfx_rewrite_tac i THEN cnfxsat_tac i; | 
| 17618 | 472 | |
| 23533 | 473 | end; |