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