src/HOL/Tools/cnf_funcs.ML
author bulwahn
Fri, 18 Mar 2011 18:19:42 +0100
changeset 42025 cb5b1e85b32e
parent 41447 537b290bbe38
child 42335 cb8aa792d138
permissions -rw-r--r--
adding eval option to quickcheck
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
     1
(*  Title:      HOL/Tools/cnf_funcs.ML
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
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    42
  val clause2raw_thm: thm -> thm
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    43
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    44
  val weakening_tac: int -> tactic  (* removes the first hypothesis of a subgoal *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    45
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    46
  val make_cnf_thm: theory -> term -> thm
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    47
  val make_cnfx_thm: theory -> term -> thm
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    48
  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
    49
  val cnfx_rewrite_tac: Proof.context -> int -> tactic
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    50
    (* converts all prems of a subgoal to (almost) definitional CNF *)
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
    51
end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    52
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    53
structure cnf : CNF =
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    54
struct
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    55
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    56
val clause2raw_notE      = @{lemma "[| P; ~P |] ==> False" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    57
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
    58
val clause2raw_not_not   = @{lemma "P ==> ~~P" by auto};
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    59
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    60
val iff_refl             = @{lemma "(P::bool) = P" by auto};
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    61
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
    62
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
    63
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
    64
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    65
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
    66
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
    67
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
    68
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
    69
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
    70
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
    71
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
    72
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
    73
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
    74
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    75
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
    76
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
    77
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
    78
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
    79
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
    80
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
    81
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
    82
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
    83
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    84
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
    85
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
    86
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    87
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
    88
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
    89
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
    90
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
    91
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    92
val weakening_thm        = @{lemma "[| P; Q |] ==> Q" by auto};
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    93
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 29265
diff changeset
    94
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
    95
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    96
fun is_atom (Const (@{const_name False}, _)) = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    97
  | is_atom (Const (@{const_name True}, _)) = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    98
  | is_atom (Const (@{const_name HOL.conj}, _) $ _ $ _) = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    99
  | is_atom (Const (@{const_name HOL.disj}, _) $ _ $ _) = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   100
  | is_atom (Const (@{const_name HOL.implies}, _) $ _ $ _) = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   101
  | is_atom (Const (@{const_name HOL.eq}, Type ("fun", @{typ bool} :: _)) $ _ $ _) = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   102
  | is_atom (Const (@{const_name Not}, _) $ _) = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   103
  | is_atom _ = true;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   104
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38553
diff changeset
   105
fun is_literal (Const (@{const_name Not}, _) $ x) = is_atom x
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   106
  | is_literal x = is_atom x;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   107
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   108
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
   109
  | is_clause x = is_literal x;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   110
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   111
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   112
(* 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
   113
(*      and the atom's negation                                              *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   114
(* ------------------------------------------------------------------------- *)
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
fun clause_is_trivial c =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   117
  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   118
    fun dual (Const (@{const_name Not}, _) $ x) = x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   119
      | dual x = HOLogic.Not $ x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   120
    fun has_duals [] = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   121
      | has_duals (x::xs) = member (op =) xs (dual x) orelse has_duals xs
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   122
  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   123
    has_duals (HOLogic.disjuncts c)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   124
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   125
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   126
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   127
(* 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
   128
(*        [...] |- x1 | ... | xn                                             *)
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   129
(*      (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
   130
(*        [..., x1', ..., xn'] |- False ,                                    *)
e6fe74eebda3 faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
webertj
parents: 19236
diff changeset
   131
(*      where each xi' is the negation normal form of ~xi                    *)
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   132
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   133
20440
e6fe74eebda3 faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
webertj
parents: 19236
diff changeset
   134
fun clause2raw_thm clause =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   135
  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   136
    (* eliminates negated disjunctions from the i-th premise, possibly *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   137
    (* adding new premises, then continues with the (i+1)-th premise   *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   138
    (* int -> Thm.thm -> Thm.thm *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   139
    fun not_disj_to_prem i thm =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   140
      if i > nprems_of thm then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   141
        thm
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   142
      else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   143
        not_disj_to_prem (i+1) (Seq.hd (REPEAT_DETERM (rtac clause2raw_not_disj i) thm))
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   144
    (* moves all premises to hyps, i.e. "[...] |- A1 ==> ... ==> An ==> B" *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   145
    (* becomes "[..., A1, ..., An] |- B"                                   *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   146
    (* Thm.thm -> Thm.thm *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   147
    fun prems_to_hyps thm =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   148
      fold (fn cprem => fn thm' =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   149
        Thm.implies_elim thm' (Thm.assume cprem)) (cprems_of thm) thm
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   150
  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   151
    (* [...] |- ~(x1 | ... | xn) ==> False *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   152
    (clause2raw_notE OF [clause])
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
    |> not_disj_to_prem 1
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
    |> Seq.hd o TRYALL (rtac clause2raw_not_not)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   157
    (* [..., x1', ..., xn'] |- False *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   158
    |> prems_to_hyps
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   159
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   160
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   161
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   162
(* inst_thm: instantiates a theorem with a list of terms                     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   163
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   164
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   165
fun inst_thm thy ts thm =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   166
  instantiate' [] (map (SOME o cterm_of thy) ts) thm;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   167
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   168
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   169
(*                         Naive CNF transformation                          *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   170
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   171
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   172
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   173
(* 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
   174
(*      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
   175
(*      of t; implications ("-->") and equivalences ("=" on bool) are        *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   176
(*      eliminated (possibly causing an exponential blowup)                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   177
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   178
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   179
fun make_nnf_thm thy (Const (@{const_name HOL.conj}, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   180
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   181
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   182
        val thm2 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   183
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   184
        conj_cong OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   185
      end
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   186
  | make_nnf_thm thy (Const (@{const_name HOL.disj}, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   187
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   188
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   189
        val thm2 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   190
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   191
        disj_cong OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   192
      end
38786
e46e7a9cb622 formerly unnamed infix impliciation now named HOL.implies
haftmann
parents: 38558
diff changeset
   193
  | make_nnf_thm thy (Const (@{const_name HOL.implies}, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   194
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   195
        val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   196
        val thm2 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   197
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   198
        make_nnf_imp OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   199
      end
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38795
diff changeset
   200
  | 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
   201
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   202
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   203
        val thm2 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   204
        val thm3 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   205
        val thm4 = make_nnf_thm thy (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   206
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   207
        make_nnf_iff OF [thm1, thm2, thm3, thm4]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   208
      end
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38553
diff changeset
   209
  | make_nnf_thm thy (Const (@{const_name Not}, _) $ Const (@{const_name False}, _)) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   210
      make_nnf_not_false
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38553
diff changeset
   211
  | make_nnf_thm thy (Const (@{const_name Not}, _) $ Const (@{const_name True}, _)) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   212
      make_nnf_not_true
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   213
  | 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
   214
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   215
        val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   216
        val thm2 = make_nnf_thm thy (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   217
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   218
        make_nnf_not_conj OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   219
      end
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   220
  | 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
   221
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   222
        val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   223
        val thm2 = make_nnf_thm thy (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   224
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   225
        make_nnf_not_disj OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   226
      end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   227
  | make_nnf_thm thy
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   228
      (Const (@{const_name Not}, _) $
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   229
        (Const (@{const_name HOL.implies}, _) $ x $ y)) =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   230
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   231
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   232
        val thm2 = make_nnf_thm thy (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   233
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   234
        make_nnf_not_imp OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   235
      end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   236
  | make_nnf_thm thy
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   237
      (Const (@{const_name Not}, _) $
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   238
        (Const (@{const_name HOL.eq}, Type ("fun", @{typ bool} :: _)) $ x $ y)) =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   239
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   240
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   241
        val thm2 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   242
        val thm3 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   243
        val thm4 = make_nnf_thm thy (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   244
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   245
        make_nnf_not_iff OF [thm1, thm2, thm3, thm4]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   246
      end
38558
32ad17fe2b9c tuned quotes
haftmann
parents: 38553
diff changeset
   247
  | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name Not}, _) $ x)) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   248
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   249
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   250
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   251
        make_nnf_not_not OF [thm1]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   252
      end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   253
  | 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
   254
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   255
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   256
(* 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
   257
(*      t, but simplified wrt. the following theorems:                       *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   258
(*        (True & x) = x                                                     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   259
(*        (x & True) = x                                                     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   260
(*        (False & x) = False                                                *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   261
(*        (x & False) = False                                                *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   262
(*        (True | x) = True                                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   263
(*        (x | True) = True                                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   264
(*        (False | x) = x                                                    *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   265
(*        (x | False) = x                                                    *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   266
(*      No simplification is performed below connectives other than & and |. *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   267
(*      Optimization: The right-hand side of a conjunction (disjunction) is  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   268
(*      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
   269
(*      (True, respectively).                                                *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   270
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   271
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   272
(* Theory.theory -> Term.term -> Thm.thm *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   273
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   274
fun simp_True_False_thm thy (Const (@{const_name HOL.conj}, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   275
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   276
        val thm1 = simp_True_False_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   277
        val x'= (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   278
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   279
        if x' = HOLogic.false_const then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   280
          simp_TF_conj_False_l OF [thm1]  (* (x & y) = False *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   281
        else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   282
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   283
            val thm2 = simp_True_False_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   284
            val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm2
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   285
          in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   286
            if x' = HOLogic.true_const then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   287
              simp_TF_conj_True_l OF [thm1, thm2]  (* (x & y) = y' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   288
            else if y' = HOLogic.false_const then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   289
              simp_TF_conj_False_r OF [thm2]  (* (x & y) = False *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   290
            else if y' = HOLogic.true_const then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   291
              simp_TF_conj_True_r OF [thm1, thm2]  (* (x & y) = x' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   292
            else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   293
              conj_cong OF [thm1, thm2]  (* (x & y) = (x' & y') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   294
          end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   295
      end
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
   296
  | simp_True_False_thm thy (Const (@{const_name HOL.disj}, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   297
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   298
        val thm1 = simp_True_False_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   299
        val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   300
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   301
        if x' = HOLogic.true_const then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   302
          simp_TF_disj_True_l OF [thm1]  (* (x | y) = True *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   303
        else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   304
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   305
            val thm2 = simp_True_False_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   306
            val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm2
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   307
          in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   308
            if x' = HOLogic.false_const then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   309
              simp_TF_disj_False_l OF [thm1, thm2]  (* (x | y) = y' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   310
            else if y' = HOLogic.true_const then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   311
              simp_TF_disj_True_r OF [thm2]  (* (x | y) = True *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   312
            else if y' = HOLogic.false_const then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   313
              simp_TF_disj_False_r OF [thm1, thm2]  (* (x | y) = x' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   314
            else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   315
              disj_cong OF [thm1, thm2]  (* (x | y) = (x' | y') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   316
          end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   317
      end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   318
  | 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
   319
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   320
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   321
(* 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
   322
(*      is in conjunction normal form.  May cause an exponential blowup      *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   323
(*      in the length of the term.                                           *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   324
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   325
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   326
fun make_cnf_thm thy t =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   327
  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   328
    fun make_cnf_thm_from_nnf (Const (@{const_name HOL.conj}, _) $ x $ y) =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   329
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   330
            val thm1 = make_cnf_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   331
            val thm2 = make_cnf_thm_from_nnf y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   332
          in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   333
            conj_cong OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   334
          end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   335
      | make_cnf_thm_from_nnf (Const (@{const_name HOL.disj}, _) $ x $ y) =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   336
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   337
            (* produces a theorem "(x' | y') = t'", where x', y', and t' are in CNF *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   338
            fun make_cnf_disj_thm (Const (@{const_name HOL.conj}, _) $ x1 $ x2) y' =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   339
                  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   340
                    val thm1 = make_cnf_disj_thm x1 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   341
                    val thm2 = make_cnf_disj_thm x2 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   342
                  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   343
                    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
   344
                  end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   345
              | make_cnf_disj_thm x' (Const (@{const_name HOL.conj}, _) $ y1 $ y2) =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   346
                  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   347
                    val thm1 = make_cnf_disj_thm x' y1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   348
                    val thm2 = make_cnf_disj_thm x' y2
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   349
                  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   350
                    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
   351
                  end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   352
              | make_cnf_disj_thm x' y' =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   353
                  inst_thm thy [HOLogic.mk_disj (x', y')] iff_refl  (* (x' | y') = (x' | y') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   354
            val thm1 = make_cnf_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   355
            val thm2 = make_cnf_thm_from_nnf y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   356
            val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   357
            val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm2
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   358
            val disj_thm = disj_cong OF [thm1, thm2]  (* (x | y) = (x' | y') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   359
          in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   360
            iff_trans OF [disj_thm, make_cnf_disj_thm x' y']
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   361
          end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   362
      | make_cnf_thm_from_nnf t = inst_thm thy [t] iff_refl
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   363
    (* convert 't' to NNF first *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   364
    val nnf_thm = make_nnf_thm thy t
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   365
    val nnf = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) nnf_thm
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   366
    (* then simplify wrt. True/False (this should preserve NNF) *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   367
    val simp_thm = simp_True_False_thm thy nnf
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   368
    val simp = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) simp_thm
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   369
    (* finally, convert to CNF (this should preserve the simplification) *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   370
    val cnf_thm = make_cnf_thm_from_nnf simp
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   371
  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   372
    iff_trans OF [iff_trans OF [nnf_thm, simp_thm], cnf_thm]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   373
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   374
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   375
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   376
(*            CNF transformation by introducing new literals                 *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   377
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   378
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   379
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   380
(* 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
   381
(*      t' is almost in conjunction normal form, except that conjunctions    *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   382
(*      and existential quantifiers may be nested.  (Use e.g. 'REPEAT_DETERM *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   383
(*      (etac exE i ORELSE etac conjE i)' afterwards to normalize.)  May     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   384
(*      introduce new (existentially bound) literals.  Note: the current     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   385
(*      implementation calls 'make_nnf_thm', causing an exponential blowup   *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   386
(*      in the case of nested equivalences.                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   387
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   388
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   389
fun make_cnfx_thm thy t =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   390
  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   391
    val var_id = Unsynchronized.ref 0  (* properly initialized below *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   392
    fun new_free () =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   393
      Free ("cnfx_" ^ string_of_int (Unsynchronized.inc var_id), HOLogic.boolT)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   394
    fun make_cnfx_thm_from_nnf (Const (@{const_name HOL.conj}, _) $ x $ y) : thm =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   395
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   396
            val thm1 = make_cnfx_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   397
            val thm2 = make_cnfx_thm_from_nnf y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   398
          in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   399
            conj_cong OF [thm1, thm2]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   400
          end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   401
      | make_cnfx_thm_from_nnf (Const (@{const_name HOL.disj}, _) $ x $ y) =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   402
          if is_clause x andalso is_clause y then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   403
            inst_thm thy [HOLogic.mk_disj (x, y)] iff_refl
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   404
          else if is_literal y orelse is_literal x then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   405
            let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   406
              (* produces a theorem "(x' | y') = t'", where x', y', and t' are *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   407
              (* almost in CNF, and x' or y' is a literal                      *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   408
              fun make_cnfx_disj_thm (Const (@{const_name HOL.conj}, _) $ x1 $ x2) y' =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   409
                    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   410
                      val thm1 = make_cnfx_disj_thm x1 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   411
                      val thm2 = make_cnfx_disj_thm x2 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   412
                    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   413
                      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
   414
                    end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   415
                | make_cnfx_disj_thm x' (Const (@{const_name HOL.conj}, _) $ y1 $ y2) =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   416
                    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   417
                      val thm1 = make_cnfx_disj_thm x' y1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   418
                      val thm2 = make_cnfx_disj_thm x' y2
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   419
                    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   420
                      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
   421
                    end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   422
                | make_cnfx_disj_thm (@{term "Ex::(bool => bool) => bool"} $ x') y' =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   423
                    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   424
                      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
   425
                      val var = new_free ()
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   426
                      val thm2 = make_cnfx_disj_thm (betapply (x', var)) y'  (* (x' | y') = body' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   427
                      val thm3 = Thm.forall_intr (cterm_of thy var) thm2     (* !!v. (x' | y') = body' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   428
                      val thm4 = Thm.strip_shyps (thm3 COMP allI)            (* ALL v. (x' | y') = body' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   429
                      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
   430
                    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   431
                      iff_trans OF [thm1, thm5]  (* ((Ex x') | y') = (Ex v. body') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   432
                    end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   433
                | make_cnfx_disj_thm x' (@{term "Ex::(bool => bool) => bool"} $ y') =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   434
                    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   435
                      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
   436
                      val var = new_free ()
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   437
                      val thm2 = make_cnfx_disj_thm x' (betapply (y', var))  (* (x' | y') = body' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   438
                      val thm3 = Thm.forall_intr (cterm_of thy var) thm2     (* !!v. (x' | y') = body' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   439
                      val thm4 = Thm.strip_shyps (thm3 COMP allI)            (* ALL v. (x' | y') = body' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   440
                      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
   441
                    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   442
                      iff_trans OF [thm1, thm5]  (* (x' | (Ex y')) = (EX v. body') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   443
                    end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   444
                | make_cnfx_disj_thm x' y' =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   445
                    inst_thm thy [HOLogic.mk_disj (x', y')] iff_refl  (* (x' | y') = (x' | y') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   446
              val thm1 = make_cnfx_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   447
              val thm2 = make_cnfx_thm_from_nnf y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   448
              val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   449
              val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm2
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   450
              val disj_thm = disj_cong OF [thm1, thm2]  (* (x | y) = (x' | y') *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   451
            in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   452
              iff_trans OF [disj_thm, make_cnfx_disj_thm x' y']
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   453
            end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   454
          else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   455
            let  (* neither 'x' nor 'y' is a literal: introduce a fresh variable *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   456
              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
   457
              val var = new_free ()
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   458
              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
   459
              val thm2 = make_cnfx_thm_from_nnf body              (* (x | v) & (y | ~v) = body' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   460
              val thm3 = Thm.forall_intr (cterm_of thy var) thm2  (* !!v. (x | v) & (y | ~v) = body' *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   461
              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
   462
              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
   463
            in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   464
              iff_trans OF [thm1, thm5]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   465
            end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   466
      | make_cnfx_thm_from_nnf t = inst_thm thy [t] iff_refl
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   467
    (* convert 't' to NNF first *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   468
    val nnf_thm = make_nnf_thm thy t
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   469
    val nnf = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) nnf_thm
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   470
    (* then simplify wrt. True/False (this should preserve NNF) *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   471
    val simp_thm = simp_True_False_thm thy nnf
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   472
    val simp = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) simp_thm
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   473
    (* 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
   474
    val _ = (var_id := fold (fn free => fn max =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   475
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   476
        val (name, _) = dest_Free free
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   477
        val idx =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   478
          if String.isPrefix "cnfx_" name then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   479
            (Int.fromString o String.extract) (name, String.size "cnfx_", NONE)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   480
          else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   481
            NONE
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   482
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   483
        Int.max (max, the_default 0 idx)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   484
      end) (OldTerm.term_frees simp) 0)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   485
    (* finally, convert to definitional CNF (this should preserve the simplification) *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   486
    val cnfx_thm = make_cnfx_thm_from_nnf simp
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   487
  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   488
    iff_trans OF [iff_trans OF [nnf_thm, simp_thm], cnfx_thm]
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   489
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   490
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   491
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   492
(*                                  Tactics                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   493
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   494
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   495
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   496
(* weakening_tac: removes the first hypothesis of the 'i'-th subgoal         *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   497
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   498
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   499
fun weakening_tac i =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   500
  dtac weakening_thm i THEN atac (i+1);
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   501
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   502
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   503
(* 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
   504
(*      (possibly causing an exponential blowup in the length of each        *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   505
(*      premise)                                                             *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   506
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   507
32232
6c394343360f proper context for SAT tactics;
wenzelm
parents: 32231
diff changeset
   508
fun cnf_rewrite_tac ctxt i =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   509
  (* cut the CNF formulas as new premises *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   510
  Subgoal.FOCUS (fn {prems, ...} =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   511
    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   512
      val thy = ProofContext.theory_of ctxt
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   513
      val cnf_thms = map (make_cnf_thm thy o HOLogic.dest_Trueprop o Thm.prop_of) prems
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   514
      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
   515
    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   516
      cut_facts_tac cut_thms 1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   517
    end) ctxt i
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   518
  (* remove the original premises *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   519
  THEN SELECT_GOAL (fn thm =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   520
    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   521
      val n = Logic.count_prems ((Term.strip_all_body o fst o Logic.dest_implies o prop_of) thm)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   522
    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   523
      PRIMITIVE (funpow (n div 2) (Seq.hd o weakening_tac 1)) thm
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   524
    end) i;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   525
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   526
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   527
(* 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
   528
(*      (possibly introducing new literals)                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   529
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   530
32232
6c394343360f proper context for SAT tactics;
wenzelm
parents: 32231
diff changeset
   531
fun cnfx_rewrite_tac ctxt i =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   532
  (* cut the CNF formulas as new premises *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   533
  Subgoal.FOCUS (fn {prems, ...} =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   534
    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   535
      val thy = ProofContext.theory_of ctxt;
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   536
      val cnfx_thms = map (make_cnfx_thm thy o HOLogic.dest_Trueprop o prop_of) prems
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   537
      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
   538
    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   539
      cut_facts_tac cut_thms 1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   540
    end) ctxt i
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   541
  (* remove the original premises *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   542
  THEN SELECT_GOAL (fn thm =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   543
    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   544
      val n = Logic.count_prems ((Term.strip_all_body o fst o Logic.dest_implies o prop_of) thm)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   545
    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   546
      PRIMITIVE (funpow (n div 2) (Seq.hd o weakening_tac 1)) thm
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   547
    end) i;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   548
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   549
end;