src/HOL/Tools/cnf.ML
author wenzelm
Fri, 06 Mar 2015 23:56:43 +0100
changeset 59642 929984c529d3
parent 59621 291934bac95e
child 60696 8304fb4fb823
permissions -rw-r--r--
clarified context;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55239
97921d23ebe3 more standard file/module names;
wenzelm
parents: 45740
diff changeset
     1
(*  Title:      HOL/Tools/cnf.ML
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
     2
    Author:     Alwen Tiu, QSL Team, LORIA (http://qsl.loria.fr)
29265
5b4247055bd7 moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents: 26341
diff changeset
     3
    Author:     Tjark Weber, TU Muenchen
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
     4
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
     5
FIXME: major overlaps with the code in meson.ML
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
     6
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
     7
Functions and tactics to transform a formula into Conjunctive Normal
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
     8
Form (CNF).
24958
ff15f76741bd rationalized redundant code
paulson
parents: 21576
diff changeset
     9
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    10
A formula in CNF is of the following form:
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    11
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    12
    (x11 | x12 | ... | x1n) & ... & (xm1 | xm2 | ... | xmk)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    13
    False
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    14
    True
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    15
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    16
where each xij is a literal (a positive or negative atomic Boolean
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    17
term), i.e. the formula is a conjunction of disjunctions of literals,
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    18
or "False", or "True".
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
    19
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    20
A (non-empty) disjunction of literals is referred to as "clause".
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    21
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    22
For the purpose of SAT proof reconstruction, we also make use of
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    23
another representation of clauses, which we call the "raw clauses".
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    24
Raw clauses are of the form
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    25
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    26
    [..., x1', x2', ..., xn'] |- False ,
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    27
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    28
where each xi is a literal, and each xi' is the negation normal form
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    29
of ~xi.
19236
150e8b0fb991 clauses now use (meta-)hyps instead of (meta-)implications; significant speedup
webertj
parents: 17809
diff changeset
    30
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    31
Literals are successively removed from the hyps of raw clauses by
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32740
diff changeset
    32
resolution during SAT proof reconstruction.
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    33
*)
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    34
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    35
signature CNF =
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    36
sig
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    37
  val is_atom: term -> bool
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    38
  val is_literal: term -> bool
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    39
  val is_clause: term -> bool
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    40
  val clause_is_trivial: term -> bool
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
    41
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
    42
  val clause2raw_thm: Proof.context -> thm -> thm
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
    43
  val make_nnf_thm: theory -> term -> thm
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    44
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58839
diff changeset
    45
  val weakening_tac: Proof.context -> int -> tactic  (* removes the first hypothesis of a subgoal *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    46
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
    47
  val make_cnf_thm: Proof.context -> term -> thm
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
    48
  val make_cnfx_thm: Proof.context -> term -> thm
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    49
  val cnf_rewrite_tac: Proof.context -> int -> tactic  (* converts all prems of a subgoal to CNF *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    50
  val cnfx_rewrite_tac: Proof.context -> int -> tactic
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    51
    (* converts all prems of a subgoal to (almost) definitional CNF *)
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
    52
end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    53
55239
97921d23ebe3 more standard file/module names;
wenzelm
parents: 45740
diff changeset
    54
structure CNF : CNF =
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    55
struct
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    56
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    57
val clause2raw_notE      = @{lemma "[| P; ~P |] ==> False" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    58
val clause2raw_not_disj  = @{lemma "[| ~P; ~Q |] ==> ~(P | Q)" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    59
val clause2raw_not_not   = @{lemma "P ==> ~~P" by auto};
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    60
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    61
val iff_refl             = @{lemma "(P::bool) = P" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    62
val iff_trans            = @{lemma "[| (P::bool) = Q; Q = R |] ==> P = R" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    63
val conj_cong            = @{lemma "[| P = P'; Q = Q' |] ==> (P & Q) = (P' & Q')" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    64
val disj_cong            = @{lemma "[| P = P'; Q = Q' |] ==> (P | Q) = (P' | Q')" by auto};
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    65
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    66
val make_nnf_imp         = @{lemma "[| (~P) = P'; Q = Q' |] ==> (P --> Q) = (P' | Q')" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    67
val make_nnf_iff         = @{lemma "[| P = P'; (~P) = NP; Q = Q'; (~Q) = NQ |] ==> (P = Q) = ((P' | NQ) & (NP | Q'))" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    68
val make_nnf_not_false   = @{lemma "(~False) = True" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    69
val make_nnf_not_true    = @{lemma "(~True) = False" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    70
val make_nnf_not_conj    = @{lemma "[| (~P) = P'; (~Q) = Q' |] ==> (~(P & Q)) = (P' | Q')" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    71
val make_nnf_not_disj    = @{lemma "[| (~P) = P'; (~Q) = Q' |] ==> (~(P | Q)) = (P' & Q')" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    72
val make_nnf_not_imp     = @{lemma "[| P = P'; (~Q) = Q' |] ==> (~(P --> Q)) = (P' & Q')" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    73
val make_nnf_not_iff     = @{lemma "[| P = P'; (~P) = NP; Q = Q'; (~Q) = NQ |] ==> (~(P = Q)) = ((P' | Q') & (NP | NQ))" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    74
val make_nnf_not_not     = @{lemma "P = P' ==> (~~P) = P'" by auto};
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    75
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    76
val simp_TF_conj_True_l  = @{lemma "[| P = True; Q = Q' |] ==> (P & Q) = Q'" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    77
val simp_TF_conj_True_r  = @{lemma "[| P = P'; Q = True |] ==> (P & Q) = P'" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    78
val simp_TF_conj_False_l = @{lemma "P = False ==> (P & Q) = False" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    79
val simp_TF_conj_False_r = @{lemma "Q = False ==> (P & Q) = False" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    80
val simp_TF_disj_True_l  = @{lemma "P = True ==> (P | Q) = True" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    81
val simp_TF_disj_True_r  = @{lemma "Q = True ==> (P | Q) = True" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    82
val simp_TF_disj_False_l = @{lemma "[| P = False; Q = Q' |] ==> (P | Q) = Q'" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    83
val simp_TF_disj_False_r = @{lemma "[| P = P'; Q = False |] ==> (P | Q) = P'" by auto};
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    84
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    85
val make_cnf_disj_conj_l = @{lemma "[| (P | R) = PR; (Q | R) = QR |] ==> ((P & Q) | R) = (PR & QR)" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    86
val make_cnf_disj_conj_r = @{lemma "[| (P | Q) = PQ; (P | R) = PR |] ==> (P | (Q & R)) = (PQ & PR)" by auto};
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    87
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    88
val make_cnfx_disj_ex_l  = @{lemma "((EX (x::bool). P x) | Q) = (EX x. P x | Q)" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    89
val make_cnfx_disj_ex_r  = @{lemma "(P | (EX (x::bool). Q x)) = (EX x. P | Q x)" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    90
val make_cnfx_newlit     = @{lemma "(P | Q) = (EX x. (P | x) & (Q | ~x))" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    91
val make_cnfx_ex_cong    = @{lemma "(ALL (x::bool). P x = Q x) ==> (EX x. P x) = (EX x. Q x)" by auto};
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    92
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    93
val weakening_thm        = @{lemma "[| P; Q |] ==> Q" by auto};
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    94
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    95
val cnftac_eq_imp        = @{lemma "[| P = Q; P |] ==> Q" by auto};
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    96
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    97
fun is_atom (Const (@{const_name False}, _)) = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    98
  | is_atom (Const (@{const_name True}, _)) = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    99
  | is_atom (Const (@{const_name HOL.conj}, _) $ _ $ _) = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   100
  | is_atom (Const (@{const_name HOL.disj}, _) $ _ $ _) = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   101
  | is_atom (Const (@{const_name HOL.implies}, _) $ _ $ _) = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   102
  | is_atom (Const (@{const_name HOL.eq}, Type ("fun", @{typ bool} :: _)) $ _ $ _) = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   103
  | is_atom (Const (@{const_name Not}, _) $ _) = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   104
  | is_atom _ = true;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   105
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38553
diff changeset
   106
fun is_literal (Const (@{const_name Not}, _) $ x) = is_atom x
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   107
  | is_literal x = is_atom x;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   108
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   109
fun is_clause (Const (@{const_name HOL.disj}, _) $ x $ y) = is_clause x andalso is_clause y
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   110
  | is_clause x = is_literal x;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   111
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   112
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   113
(* clause_is_trivial: a clause is trivially true if it contains both an atom *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   114
(*      and the atom's negation                                              *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   115
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   116
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   117
fun clause_is_trivial c =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   118
  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   119
    fun dual (Const (@{const_name Not}, _) $ x) = x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   120
      | dual x = HOLogic.Not $ x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   121
    fun has_duals [] = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   122
      | has_duals (x::xs) = member (op =) xs (dual x) orelse has_duals xs
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   123
  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   124
    has_duals (HOLogic.disjuncts c)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   125
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   126
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   127
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   128
(* clause2raw_thm: translates a clause into a raw clause, i.e.               *)
20440
e6fe74eebda3 faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
webertj
parents: 19236
diff changeset
   129
(*        [...] |- x1 | ... | xn                                             *)
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   130
(*      (where each xi is a literal) is translated to                        *)
20440
e6fe74eebda3 faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
webertj
parents: 19236
diff changeset
   131
(*        [..., x1', ..., xn'] |- False ,                                    *)
e6fe74eebda3 faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
webertj
parents: 19236
diff changeset
   132
(*      where each xi' is the negation normal form of ~xi                    *)
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   133
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   134
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
   135
fun clause2raw_thm ctxt clause =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   136
  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   137
    (* eliminates negated disjunctions from the i-th premise, possibly *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   138
    (* adding new premises, then continues with the (i+1)-th premise   *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   139
    (* int -> Thm.thm -> Thm.thm *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   140
    fun not_disj_to_prem i thm =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   141
      if i > Thm.nprems_of thm then
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   142
        thm
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   143
      else
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
   144
        not_disj_to_prem (i+1)
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
   145
          (Seq.hd (REPEAT_DETERM (resolve_tac ctxt [clause2raw_not_disj] i) thm))
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   146
    (* moves all premises to hyps, i.e. "[...] |- A1 ==> ... ==> An ==> B" *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   147
    (* becomes "[..., A1, ..., An] |- B"                                   *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   148
    (* Thm.thm -> Thm.thm *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   149
    fun prems_to_hyps thm =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   150
      fold (fn cprem => fn thm' =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   151
        Thm.implies_elim thm' (Thm.assume cprem)) (cprems_of thm) thm
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   152
  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   153
    (* [...] |- ~(x1 | ... | xn) ==> False *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   154
    (clause2raw_notE OF [clause])
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   155
    (* [...] |- ~x1 ==> ... ==> ~xn ==> False *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   156
    |> not_disj_to_prem 1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   157
    (* [...] |- x1' ==> ... ==> xn' ==> False *)
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
   158
    |> Seq.hd o TRYALL (resolve_tac ctxt [clause2raw_not_not])
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   159
    (* [..., x1', ..., xn'] |- False *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   160
    |> prems_to_hyps
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   161
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   162
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   163
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   164
(* inst_thm: instantiates a theorem with a list of terms                     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   165
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   166
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   167
fun inst_thm thy ts thm =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   168
  instantiate' [] (map (SOME o Thm.global_cterm_of thy) ts) thm;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   169
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   170
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   171
(*                         Naive CNF transformation                          *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   172
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   173
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   174
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   175
(* make_nnf_thm: produces a theorem of the form t = t', where t' is the      *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   176
(*      negation normal form (i.e. negation only occurs in front of atoms)   *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   177
(*      of t; implications ("-->") and equivalences ("=" on bool) are        *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   178
(*      eliminated (possibly causing an exponential blowup)                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   179
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   180
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   181
fun make_nnf_thm thy (Const (@{const_name HOL.conj}, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   182
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   183
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   184
        val thm2 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   185
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   186
        conj_cong OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   187
      end
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   188
  | make_nnf_thm thy (Const (@{const_name HOL.disj}, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   189
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   190
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   191
        val thm2 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   192
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   193
        disj_cong OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   194
      end
38786
e46e7a9cb622 formerly unnamed infix impliciation now named HOL.implies
haftmann
parents: 38558
diff changeset
   195
  | make_nnf_thm thy (Const (@{const_name HOL.implies}, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   196
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   197
        val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   198
        val thm2 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   199
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   200
        make_nnf_imp OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   201
      end
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38795
diff changeset
   202
  | make_nnf_thm thy (Const (@{const_name HOL.eq}, Type ("fun", @{typ bool} :: _)) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   203
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   204
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   205
        val thm2 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   206
        val thm3 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   207
        val thm4 = make_nnf_thm thy (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   208
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   209
        make_nnf_iff OF [thm1, thm2, thm3, thm4]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   210
      end
56506
wenzelm
parents: 55239
diff changeset
   211
  | make_nnf_thm _ (Const (@{const_name Not}, _) $ Const (@{const_name False}, _)) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   212
      make_nnf_not_false
56506
wenzelm
parents: 55239
diff changeset
   213
  | make_nnf_thm _ (Const (@{const_name Not}, _) $ Const (@{const_name True}, _)) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   214
      make_nnf_not_true
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   215
  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.conj}, _) $ x $ y)) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   216
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   217
        val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   218
        val thm2 = make_nnf_thm thy (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   219
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   220
        make_nnf_not_conj OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   221
      end
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   222
  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.disj}, _) $ x $ y)) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   223
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   224
        val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   225
        val thm2 = make_nnf_thm thy (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   226
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   227
        make_nnf_not_disj OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   228
      end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   229
  | make_nnf_thm thy
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   230
      (Const (@{const_name Not}, _) $
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   231
        (Const (@{const_name HOL.implies}, _) $ x $ y)) =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   232
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   233
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   234
        val thm2 = make_nnf_thm thy (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   235
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   236
        make_nnf_not_imp OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   237
      end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   238
  | make_nnf_thm thy
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   239
      (Const (@{const_name Not}, _) $
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   240
        (Const (@{const_name HOL.eq}, Type ("fun", @{typ bool} :: _)) $ x $ y)) =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   241
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   242
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   243
        val thm2 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   244
        val thm3 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   245
        val thm4 = make_nnf_thm thy (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   246
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   247
        make_nnf_not_iff OF [thm1, thm2, thm3, thm4]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   248
      end
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38553
diff changeset
   249
  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name Not}, _) $ x)) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   250
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   251
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   252
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   253
        make_nnf_not_not OF [thm1]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   254
      end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   255
  | make_nnf_thm thy t = inst_thm thy [t] iff_refl;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   256
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   257
val meta_eq_to_obj_eq = @{thm meta_eq_to_obj_eq}
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   258
val eq_reflection = @{thm eq_reflection}
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   259
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   260
fun make_under_quantifiers ctxt make t =
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   261
  let
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   262
    fun conv ctxt ct =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   263
      (case Thm.term_of ct of
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 44121
diff changeset
   264
        Const _ $ Abs _ => Conv.comb_conv (conv ctxt) ct
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   265
      | Abs _ => Conv.abs_conv (conv o snd) ctxt ct
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   266
      | Const _ => Conv.all_conv ct
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   267
      | t => make t RS eq_reflection)
59642
929984c529d3 clarified context;
wenzelm
parents: 59621
diff changeset
   268
  in conv ctxt (Thm.cterm_of ctxt t) RS meta_eq_to_obj_eq end
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   269
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   270
fun make_nnf_thm_under_quantifiers ctxt =
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42335
diff changeset
   271
  make_under_quantifiers ctxt (make_nnf_thm (Proof_Context.theory_of ctxt))
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   272
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   273
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   274
(* simp_True_False_thm: produces a theorem t = t', where t' is equivalent to *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   275
(*      t, but simplified wrt. the following theorems:                       *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   276
(*        (True & x) = x                                                     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   277
(*        (x & True) = x                                                     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   278
(*        (False & x) = False                                                *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   279
(*        (x & False) = False                                                *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   280
(*        (True | x) = True                                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   281
(*        (x | True) = True                                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   282
(*        (False | x) = x                                                    *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   283
(*        (x | False) = x                                                    *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   284
(*      No simplification is performed below connectives other than & and |. *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   285
(*      Optimization: The right-hand side of a conjunction (disjunction) is  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   286
(*      simplified only if the left-hand side does not simplify to False     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   287
(*      (True, respectively).                                                *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   288
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   289
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   290
(* Theory.theory -> Term.term -> Thm.thm *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   291
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   292
fun simp_True_False_thm thy (Const (@{const_name HOL.conj}, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   293
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   294
        val thm1 = simp_True_False_thm thy x
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   295
        val x'= (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm1
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   296
      in
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45511
diff changeset
   297
        if x' = @{term False} then
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   298
          simp_TF_conj_False_l OF [thm1]  (* (x & y) = False *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   299
        else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   300
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   301
            val thm2 = simp_True_False_thm thy y
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   302
            val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   303
          in
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45511
diff changeset
   304
            if x' = @{term True} then
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   305
              simp_TF_conj_True_l OF [thm1, thm2]  (* (x & y) = y' *)
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45511
diff changeset
   306
            else if y' = @{term False} then
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   307
              simp_TF_conj_False_r OF [thm2]  (* (x & y) = False *)
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45511
diff changeset
   308
            else if y' = @{term True} then
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   309
              simp_TF_conj_True_r OF [thm1, thm2]  (* (x & y) = x' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   310
            else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   311
              conj_cong OF [thm1, thm2]  (* (x & y) = (x' & y') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   312
          end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   313
      end
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   314
  | simp_True_False_thm thy (Const (@{const_name HOL.disj}, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   315
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   316
        val thm1 = simp_True_False_thm thy x
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   317
        val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm1
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   318
      in
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45511
diff changeset
   319
        if x' = @{term True} then
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   320
          simp_TF_disj_True_l OF [thm1]  (* (x | y) = True *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   321
        else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   322
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   323
            val thm2 = simp_True_False_thm thy y
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   324
            val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   325
          in
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45511
diff changeset
   326
            if x' = @{term False} then
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   327
              simp_TF_disj_False_l OF [thm1, thm2]  (* (x | y) = y' *)
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45511
diff changeset
   328
            else if y' = @{term True} then
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   329
              simp_TF_disj_True_r OF [thm2]  (* (x | y) = True *)
45740
132a3e1c0fe5 more antiquotations;
wenzelm
parents: 45511
diff changeset
   330
            else if y' = @{term False} then
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   331
              simp_TF_disj_False_r OF [thm1, thm2]  (* (x | y) = x' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   332
            else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   333
              disj_cong OF [thm1, thm2]  (* (x | y) = (x' | y') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   334
          end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   335
      end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   336
  | simp_True_False_thm thy t = inst_thm thy [t] iff_refl;  (* t = t *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   337
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   338
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   339
(* make_cnf_thm: given any HOL term 't', produces a theorem t = t', where t' *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   340
(*      is in conjunction normal form.  May cause an exponential blowup      *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   341
(*      in the length of the term.                                           *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   342
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   343
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   344
fun make_cnf_thm ctxt t =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   345
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42335
diff changeset
   346
    val thy = Proof_Context.theory_of ctxt
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   347
    fun make_cnf_thm_from_nnf (Const (@{const_name HOL.conj}, _) $ x $ y) =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   348
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   349
            val thm1 = make_cnf_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   350
            val thm2 = make_cnf_thm_from_nnf y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   351
          in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   352
            conj_cong OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   353
          end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   354
      | make_cnf_thm_from_nnf (Const (@{const_name HOL.disj}, _) $ x $ y) =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   355
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   356
            (* produces a theorem "(x' | y') = t'", where x', y', and t' are in CNF *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   357
            fun make_cnf_disj_thm (Const (@{const_name HOL.conj}, _) $ x1 $ x2) y' =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   358
                  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   359
                    val thm1 = make_cnf_disj_thm x1 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   360
                    val thm2 = make_cnf_disj_thm x2 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   361
                  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   362
                    make_cnf_disj_conj_l OF [thm1, thm2]  (* ((x1 & x2) | y') = ((x1 | y')' & (x2 | y')') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   363
                  end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   364
              | make_cnf_disj_thm x' (Const (@{const_name HOL.conj}, _) $ y1 $ y2) =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   365
                  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   366
                    val thm1 = make_cnf_disj_thm x' y1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   367
                    val thm2 = make_cnf_disj_thm x' y2
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   368
                  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   369
                    make_cnf_disj_conj_r OF [thm1, thm2]  (* (x' | (y1 & y2)) = ((x' | y1)' & (x' | y2)') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   370
                  end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   371
              | make_cnf_disj_thm x' y' =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   372
                  inst_thm thy [HOLogic.mk_disj (x', y')] iff_refl  (* (x' | y') = (x' | y') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   373
            val thm1 = make_cnf_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   374
            val thm2 = make_cnf_thm_from_nnf y
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   375
            val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm1
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   376
            val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   377
            val disj_thm = disj_cong OF [thm1, thm2]  (* (x | y) = (x' | y') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   378
          in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   379
            iff_trans OF [disj_thm, make_cnf_disj_thm x' y']
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   380
          end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   381
      | make_cnf_thm_from_nnf t = inst_thm thy [t] iff_refl
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   382
    (* convert 't' to NNF first *)
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   383
    val nnf_thm = make_nnf_thm_under_quantifiers ctxt t
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   384
(*###
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   385
    val nnf_thm = make_nnf_thm thy t
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   386
*)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   387
    val nnf = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) nnf_thm
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   388
    (* then simplify wrt. True/False (this should preserve NNF) *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   389
    val simp_thm = simp_True_False_thm thy nnf
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   390
    val simp = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) simp_thm
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   391
    (* finally, convert to CNF (this should preserve the simplification) *)
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   392
    val cnf_thm = make_under_quantifiers ctxt make_cnf_thm_from_nnf simp
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   393
(* ###
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   394
    val cnf_thm = make_cnf_thm_from_nnf simp
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   395
*)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   396
  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   397
    iff_trans OF [iff_trans OF [nnf_thm, simp_thm], cnf_thm]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   398
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   399
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   400
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   401
(*            CNF transformation by introducing new literals                 *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   402
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   403
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   404
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   405
(* make_cnfx_thm: given any HOL term 't', produces a theorem t = t', where   *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   406
(*      t' is almost in conjunction normal form, except that conjunctions    *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   407
(*      and existential quantifiers may be nested.  (Use e.g. 'REPEAT_DETERM *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   408
(*      (etac exE i ORELSE etac conjE i)' afterwards to normalize.)  May     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   409
(*      introduce new (existentially bound) literals.  Note: the current     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   410
(*      implementation calls 'make_nnf_thm', causing an exponential blowup   *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   411
(*      in the case of nested equivalences.                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   412
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   413
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   414
fun make_cnfx_thm ctxt t =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   415
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42335
diff changeset
   416
    val thy = Proof_Context.theory_of ctxt
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   417
    val var_id = Unsynchronized.ref 0  (* properly initialized below *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   418
    fun new_free () =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   419
      Free ("cnfx_" ^ string_of_int (Unsynchronized.inc var_id), HOLogic.boolT)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   420
    fun make_cnfx_thm_from_nnf (Const (@{const_name HOL.conj}, _) $ x $ y) : thm =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   421
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   422
            val thm1 = make_cnfx_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   423
            val thm2 = make_cnfx_thm_from_nnf y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   424
          in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   425
            conj_cong OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   426
          end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   427
      | make_cnfx_thm_from_nnf (Const (@{const_name HOL.disj}, _) $ x $ y) =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   428
          if is_clause x andalso is_clause y then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   429
            inst_thm thy [HOLogic.mk_disj (x, y)] iff_refl
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   430
          else if is_literal y orelse is_literal x then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   431
            let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   432
              (* produces a theorem "(x' | y') = t'", where x', y', and t' are *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   433
              (* almost in CNF, and x' or y' is a literal                      *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   434
              fun make_cnfx_disj_thm (Const (@{const_name HOL.conj}, _) $ x1 $ x2) y' =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   435
                    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   436
                      val thm1 = make_cnfx_disj_thm x1 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   437
                      val thm2 = make_cnfx_disj_thm x2 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   438
                    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   439
                      make_cnf_disj_conj_l OF [thm1, thm2]  (* ((x1 & x2) | y') = ((x1 | y')' & (x2 | y')') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   440
                    end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   441
                | make_cnfx_disj_thm x' (Const (@{const_name HOL.conj}, _) $ y1 $ y2) =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   442
                    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   443
                      val thm1 = make_cnfx_disj_thm x' y1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   444
                      val thm2 = make_cnfx_disj_thm x' y2
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   445
                    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   446
                      make_cnf_disj_conj_r OF [thm1, thm2]  (* (x' | (y1 & y2)) = ((x' | y1)' & (x' | y2)') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   447
                    end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   448
                | make_cnfx_disj_thm (@{term "Ex::(bool => bool) => bool"} $ x') y' =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   449
                    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   450
                      val thm1 = inst_thm thy [x', y'] make_cnfx_disj_ex_l   (* ((Ex x') | y') = (Ex (x' | y')) *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   451
                      val var = new_free ()
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   452
                      val thm2 = make_cnfx_disj_thm (betapply (x', var)) y'  (* (x' | y') = body' *)
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   453
                      val thm3 = Thm.forall_intr (Thm.global_cterm_of thy var) thm2 (* !!v. (x' | y') = body' *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   454
                      val thm4 = Thm.strip_shyps (thm3 COMP allI)            (* ALL v. (x' | y') = body' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   455
                      val thm5 = Thm.strip_shyps (thm4 RS make_cnfx_ex_cong) (* (EX v. (x' | y')) = (EX v. body') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   456
                    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   457
                      iff_trans OF [thm1, thm5]  (* ((Ex x') | y') = (Ex v. body') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   458
                    end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   459
                | make_cnfx_disj_thm x' (@{term "Ex::(bool => bool) => bool"} $ y') =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   460
                    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   461
                      val thm1 = inst_thm thy [x', y'] make_cnfx_disj_ex_r   (* (x' | (Ex y')) = (Ex (x' | y')) *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   462
                      val var = new_free ()
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   463
                      val thm2 = make_cnfx_disj_thm x' (betapply (y', var))  (* (x' | y') = body' *)
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   464
                      val thm3 = Thm.forall_intr (Thm.global_cterm_of thy var) thm2 (* !!v. (x' | y') = body' *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   465
                      val thm4 = Thm.strip_shyps (thm3 COMP allI)            (* ALL v. (x' | y') = body' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   466
                      val thm5 = Thm.strip_shyps (thm4 RS make_cnfx_ex_cong) (* (EX v. (x' | y')) = (EX v. body') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   467
                    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   468
                      iff_trans OF [thm1, thm5]  (* (x' | (Ex y')) = (EX v. body') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   469
                    end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   470
                | make_cnfx_disj_thm x' y' =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   471
                    inst_thm thy [HOLogic.mk_disj (x', y')] iff_refl  (* (x' | y') = (x' | y') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   472
              val thm1 = make_cnfx_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   473
              val thm2 = make_cnfx_thm_from_nnf y
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   474
              val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm1
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   475
              val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   476
              val disj_thm = disj_cong OF [thm1, thm2]  (* (x | y) = (x' | y') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   477
            in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   478
              iff_trans OF [disj_thm, make_cnfx_disj_thm x' y']
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   479
            end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   480
          else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   481
            let  (* neither 'x' nor 'y' is a literal: introduce a fresh variable *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   482
              val thm1 = inst_thm thy [x, y] make_cnfx_newlit     (* (x | y) = EX v. (x | v) & (y | ~v) *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   483
              val var = new_free ()
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   484
              val body = HOLogic.mk_conj (HOLogic.mk_disj (x, var), HOLogic.mk_disj (y, HOLogic.Not $ var))
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   485
              val thm2 = make_cnfx_thm_from_nnf body              (* (x | v) & (y | ~v) = body' *)
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59582
diff changeset
   486
              val thm3 = Thm.forall_intr (Thm.global_cterm_of thy var) thm2 (* !!v. (x | v) & (y | ~v) = body' *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   487
              val thm4 = Thm.strip_shyps (thm3 COMP allI)         (* ALL v. (x | v) & (y | ~v) = body' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   488
              val thm5 = Thm.strip_shyps (thm4 RS make_cnfx_ex_cong)  (* (EX v. (x | v) & (y | ~v)) = (EX v. body') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   489
            in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   490
              iff_trans OF [thm1, thm5]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   491
            end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   492
      | make_cnfx_thm_from_nnf t = inst_thm thy [t] iff_refl
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   493
    (* convert 't' to NNF first *)
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   494
    val nnf_thm = make_nnf_thm_under_quantifiers ctxt t
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   495
(* ###
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   496
    val nnf_thm = make_nnf_thm thy t
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   497
*)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   498
    val nnf = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) nnf_thm
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   499
    (* then simplify wrt. True/False (this should preserve NNF) *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   500
    val simp_thm = simp_True_False_thm thy nnf
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   501
    val simp = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) simp_thm
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   502
    (* initialize var_id, in case the term already contains variables of the form "cnfx_<int>" *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   503
    val _ = (var_id := fold (fn free => fn max =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   504
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   505
        val (name, _) = dest_Free free
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   506
        val idx =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   507
          if String.isPrefix "cnfx_" name then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   508
            (Int.fromString o String.extract) (name, String.size "cnfx_", NONE)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   509
          else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   510
            NONE
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   511
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   512
        Int.max (max, the_default 0 idx)
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
   513
      end) (Misc_Legacy.term_frees simp) 0)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   514
    (* finally, convert to definitional CNF (this should preserve the simplification) *)
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   515
    val cnfx_thm = make_under_quantifiers ctxt make_cnfx_thm_from_nnf simp
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   516
(*###
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   517
    val cnfx_thm = make_cnfx_thm_from_nnf simp
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   518
*)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   519
  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   520
    iff_trans OF [iff_trans OF [nnf_thm, simp_thm], cnfx_thm]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   521
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   522
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   523
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   524
(*                                  Tactics                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   525
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   526
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   527
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   528
(* weakening_tac: removes the first hypothesis of the 'i'-th subgoal         *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   529
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   530
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58839
diff changeset
   531
fun weakening_tac ctxt i =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
   532
  dresolve_tac ctxt [weakening_thm] i THEN assume_tac ctxt (i+1);
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   533
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   534
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   535
(* cnf_rewrite_tac: converts all premises of the 'i'-th subgoal to CNF       *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   536
(*      (possibly causing an exponential blowup in the length of each        *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   537
(*      premise)                                                             *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   538
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   539
32232
6c394343360f proper context for SAT tactics;
wenzelm
parents: 32231
diff changeset
   540
fun cnf_rewrite_tac ctxt i =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   541
  (* cut the CNF formulas as new premises *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   542
  Subgoal.FOCUS (fn {prems, ...} =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   543
    let
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   544
      val cnf_thms = map (make_cnf_thm ctxt o HOLogic.dest_Trueprop o Thm.prop_of) prems
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   545
      val cut_thms = map (fn (th, pr) => cnftac_eq_imp OF [th, pr]) (cnf_thms ~~ prems)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   546
    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   547
      cut_facts_tac cut_thms 1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   548
    end) ctxt i
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   549
  (* remove the original premises *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   550
  THEN SELECT_GOAL (fn thm =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   551
    let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   552
      val n = Logic.count_prems ((Term.strip_all_body o fst o Logic.dest_implies o Thm.prop_of) thm)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   553
    in
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58839
diff changeset
   554
      PRIMITIVE (funpow (n div 2) (Seq.hd o weakening_tac ctxt 1)) thm
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   555
    end) i;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   556
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   557
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   558
(* cnfx_rewrite_tac: converts all premises of the 'i'-th subgoal to CNF      *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   559
(*      (possibly introducing new literals)                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   560
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   561
32232
6c394343360f proper context for SAT tactics;
wenzelm
parents: 32231
diff changeset
   562
fun cnfx_rewrite_tac ctxt i =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   563
  (* cut the CNF formulas as new premises *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   564
  Subgoal.FOCUS (fn {prems, ...} =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   565
    let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   566
      val cnfx_thms = map (make_cnfx_thm ctxt o HOLogic.dest_Trueprop o Thm.prop_of) prems
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   567
      val cut_thms = map (fn (th, pr) => cnftac_eq_imp OF [th, pr]) (cnfx_thms ~~ prems)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   568
    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   569
      cut_facts_tac cut_thms 1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   570
    end) ctxt i
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   571
  (* remove the original premises *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   572
  THEN SELECT_GOAL (fn thm =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   573
    let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   574
      val n = Logic.count_prems ((Term.strip_all_body o fst o Logic.dest_implies o Thm.prop_of) thm)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   575
    in
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58839
diff changeset
   576
      PRIMITIVE (funpow (n div 2) (Seq.hd o weakening_tac ctxt 1)) thm
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   577
    end) i;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   578
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   579
end;