src/HOL/Tools/cnf.ML
author wenzelm
Mon, 25 May 2020 22:37:14 +0200
changeset 71892 dff81ce866d4
parent 70486 1dc3514c1719
child 80637 e2f0265f64e3
permissions -rw-r--r--
obsolete;
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
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
    57
fun is_atom (Const (\<^const_name>\<open>False\<close>, _)) = false
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
    58
  | is_atom (Const (\<^const_name>\<open>True\<close>, _)) = false
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
    59
  | is_atom (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ _ $ _) = false
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
    60
  | is_atom (Const (\<^const_name>\<open>HOL.disj\<close>, _) $ _ $ _) = false
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
    61
  | is_atom (Const (\<^const_name>\<open>HOL.implies\<close>, _) $ _ $ _) = false
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
    62
  | is_atom (Const (\<^const_name>\<open>HOL.eq\<close>, Type ("fun", \<^typ>\<open>bool\<close> :: _)) $ _ $ _) = false
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
    63
  | is_atom (Const (\<^const_name>\<open>Not\<close>, _) $ _) = false
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    64
  | is_atom _ = true;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    65
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
    66
fun is_literal (Const (\<^const_name>\<open>Not\<close>, _) $ x) = is_atom x
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    67
  | is_literal x = is_atom x;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    68
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
    69
fun is_clause (Const (\<^const_name>\<open>HOL.disj\<close>, _) $ x $ y) = is_clause x andalso is_clause y
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    70
  | is_clause x = is_literal x;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    71
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
    72
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
    73
(* 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
    74
(*      and the atom's negation                                              *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
    75
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
    76
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
    77
fun clause_is_trivial c =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    78
  let
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
    79
    fun dual (Const (\<^const_name>\<open>Not\<close>, _) $ x) = x
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    80
      | dual x = HOLogic.Not $ x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    81
    fun has_duals [] = false
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    82
      | has_duals (x::xs) = member (op =) xs (dual x) orelse has_duals xs
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    83
  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    84
    has_duals (HOLogic.disjuncts c)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    85
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    86
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
    87
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
    88
(* 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
    89
(*        [...] |- x1 | ... | xn                                             *)
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
    90
(*      (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
    91
(*        [..., x1', ..., xn'] |- False ,                                    *)
e6fe74eebda3 faster clause representation (again): full CNF formula as a hypothesis, instead of separate clauses
webertj
parents: 19236
diff changeset
    92
(*      where each xi' is the negation normal form of ~xi                    *)
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
    93
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
    94
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
    95
fun clause2raw_thm ctxt clause =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    96
  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    97
    (* eliminates negated disjunctions from the i-th premise, possibly *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    98
    (* adding new premises, then continues with the (i+1)-th premise   *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
    99
    fun not_disj_to_prem i thm =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   100
      if i > Thm.nprems_of thm then
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   101
        thm
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   102
      else
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58963
diff changeset
   103
        not_disj_to_prem (i+1)
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   104
          (Seq.hd (REPEAT_DETERM (resolve_tac ctxt @{thms cnf.clause2raw_not_disj} i) thm))
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   105
    (* moves all premises to hyps, i.e. "[...] |- A1 ==> ... ==> An ==> B" *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   106
    (* becomes "[..., A1, ..., An] |- B"                                   *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   107
    fun prems_to_hyps thm =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   108
      fold (fn cprem => fn thm' =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   109
        Thm.implies_elim thm' (Thm.assume cprem)) (cprems_of thm) thm
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   110
  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   111
    (* [...] |- ~(x1 | ... | xn) ==> False *)
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   112
    (@{thm cnf.clause2raw_notE} OF [clause])
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   113
    (* [...] |- ~x1 ==> ... ==> ~xn ==> False *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   114
    |> not_disj_to_prem 1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   115
    (* [...] |- x1' ==> ... ==> xn' ==> False *)
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   116
    |> Seq.hd o TRYALL (resolve_tac ctxt @{thms cnf.clause2raw_not_not})
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   117
    (* [..., x1', ..., xn'] |- False *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   118
    |> prems_to_hyps
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   119
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   120
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   121
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   122
(* inst_thm: instantiates a theorem with a list of terms                     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   123
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   124
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   125
fun inst_thm thy ts thm =
60801
7664e0916eec tuned signature;
wenzelm
parents: 60759
diff changeset
   126
  Thm.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
   127
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   128
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   129
(*                         Naive CNF transformation                          *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   130
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   131
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   132
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   133
(* 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
   134
(*      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
   135
(*      of t; implications ("-->") and equivalences ("=" on bool) are        *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   136
(*      eliminated (possibly causing an exponential blowup)                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   137
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   138
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   139
fun make_nnf_thm thy (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   140
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   141
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   142
        val thm2 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   143
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   144
        @{thm cnf.conj_cong} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   145
      end
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   146
  | make_nnf_thm thy (Const (\<^const_name>\<open>HOL.disj\<close>, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   147
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   148
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   149
        val thm2 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   150
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   151
        @{thm cnf.disj_cong} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   152
      end
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   153
  | make_nnf_thm thy (Const (\<^const_name>\<open>HOL.implies\<close>, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   154
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   155
        val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   156
        val thm2 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   157
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   158
        @{thm cnf.make_nnf_imp} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   159
      end
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   160
  | make_nnf_thm thy (Const (\<^const_name>\<open>HOL.eq\<close>, Type ("fun", \<^typ>\<open>bool\<close> :: _)) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   161
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   162
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   163
        val thm2 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   164
        val thm3 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   165
        val thm4 = make_nnf_thm thy (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   166
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   167
        @{thm cnf.make_nnf_iff} OF [thm1, thm2, thm3, thm4]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   168
      end
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   169
  | make_nnf_thm _ (Const (\<^const_name>\<open>Not\<close>, _) $ Const (\<^const_name>\<open>False\<close>, _)) =
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   170
      @{thm cnf.make_nnf_not_false}
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   171
  | make_nnf_thm _ (Const (\<^const_name>\<open>Not\<close>, _) $ Const (\<^const_name>\<open>True\<close>, _)) =
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   172
      @{thm cnf.make_nnf_not_true}
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   173
  | make_nnf_thm thy (Const (\<^const_name>\<open>Not\<close>, _) $ (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ x $ y)) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   174
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   175
        val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   176
        val thm2 = make_nnf_thm thy (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   177
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   178
        @{thm cnf.make_nnf_not_conj} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   179
      end
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   180
  | make_nnf_thm thy (Const (\<^const_name>\<open>Not\<close>, _) $ (Const (\<^const_name>\<open>HOL.disj\<close>, _) $ x $ y)) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   181
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   182
        val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   183
        val thm2 = make_nnf_thm thy (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   184
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   185
        @{thm cnf.make_nnf_not_disj} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   186
      end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   187
  | make_nnf_thm thy
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   188
      (Const (\<^const_name>\<open>Not\<close>, _) $ (Const (\<^const_name>\<open>HOL.implies\<close>, _) $ 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 (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   192
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   193
        @{thm cnf.make_nnf_not_imp} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   194
      end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   195
  | make_nnf_thm thy
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   196
      (Const (\<^const_name>\<open>Not\<close>, _) $
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   197
        (Const (\<^const_name>\<open>HOL.eq\<close>, Type ("fun", \<^typ>\<open>bool\<close> :: _)) $ x $ y)) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   198
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   199
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   200
        val thm2 = make_nnf_thm thy (HOLogic.Not $ x)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   201
        val thm3 = make_nnf_thm thy y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   202
        val thm4 = make_nnf_thm thy (HOLogic.Not $ y)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   203
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   204
        @{thm cnf.make_nnf_not_iff} OF [thm1, thm2, thm3, thm4]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   205
      end
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   206
  | make_nnf_thm thy (Const (\<^const_name>\<open>Not\<close>, _) $ (Const (\<^const_name>\<open>Not\<close>, _) $ x)) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   207
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   208
        val thm1 = make_nnf_thm thy x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   209
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   210
        @{thm cnf.make_nnf_not_not} OF [thm1]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   211
      end
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   212
  | make_nnf_thm thy t = inst_thm thy [t] @{thm cnf.iff_refl};
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   213
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   214
fun make_under_quantifiers ctxt make t =
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   215
  let
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   216
    fun conv ctxt ct =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   217
      (case Thm.term_of ct of
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 44121
diff changeset
   218
        Const _ $ Abs _ => Conv.comb_conv (conv ctxt) ct
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   219
      | Abs _ => Conv.abs_conv (conv o snd) ctxt ct
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   220
      | Const _ => Conv.all_conv ct
67710
cc2db3239932 added HOLogic.mk_obj_eq convenience and eliminated some clones;
wenzelm
parents: 67149
diff changeset
   221
      | t => make t RS @{thm eq_reflection})
cc2db3239932 added HOLogic.mk_obj_eq convenience and eliminated some clones;
wenzelm
parents: 67149
diff changeset
   222
  in HOLogic.mk_obj_eq (conv ctxt (Thm.cterm_of ctxt t)) end
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   223
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   224
fun make_nnf_thm_under_quantifiers ctxt =
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42335
diff changeset
   225
  make_under_quantifiers ctxt (make_nnf_thm (Proof_Context.theory_of ctxt))
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   226
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   227
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   228
(* 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
   229
(*      t, but simplified wrt. the following theorems:                       *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   230
(*        (True & x) = x                                                     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   231
(*        (x & True) = x                                                     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   232
(*        (False & x) = False                                                *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   233
(*        (x & False) = False                                                *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   234
(*        (True | x) = True                                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   235
(*        (x | True) = True                                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   236
(*        (False | x) = x                                                    *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   237
(*        (x | False) = x                                                    *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   238
(*      No simplification is performed below connectives other than & and |. *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   239
(*      Optimization: The right-hand side of a conjunction (disjunction) is  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   240
(*      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
   241
(*      (True, respectively).                                                *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   242
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   243
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   244
fun simp_True_False_thm thy (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   245
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   246
        val thm1 = simp_True_False_thm thy x
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   247
        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
   248
      in
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   249
        if x' = \<^term>\<open>False\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   250
          @{thm cnf.simp_TF_conj_False_l} OF [thm1]  (* (x & y) = False *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   251
        else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   252
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   253
            val thm2 = simp_True_False_thm thy y
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   254
            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
   255
          in
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   256
            if x' = \<^term>\<open>True\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   257
              @{thm cnf.simp_TF_conj_True_l} OF [thm1, thm2]  (* (x & y) = y' *)
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   258
            else if y' = \<^term>\<open>False\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   259
              @{thm cnf.simp_TF_conj_False_r} OF [thm2]  (* (x & y) = False *)
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   260
            else if y' = \<^term>\<open>True\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   261
              @{thm cnf.simp_TF_conj_True_r} OF [thm1, thm2]  (* (x & y) = x' *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   262
            else
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   263
              @{thm cnf.conj_cong} OF [thm1, thm2]  (* (x & y) = (x' & y') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   264
          end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   265
      end
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   266
  | simp_True_False_thm thy (Const (\<^const_name>\<open>HOL.disj\<close>, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   267
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   268
        val thm1 = simp_True_False_thm thy x
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   269
        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
   270
      in
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   271
        if x' = \<^term>\<open>True\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   272
          @{thm cnf.simp_TF_disj_True_l} OF [thm1]  (* (x | y) = True *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   273
        else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   274
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   275
            val thm2 = simp_True_False_thm thy y
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   276
            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
   277
          in
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   278
            if x' = \<^term>\<open>False\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   279
              @{thm cnf.simp_TF_disj_False_l} OF [thm1, thm2]  (* (x | y) = y' *)
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   280
            else if y' = \<^term>\<open>True\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   281
              @{thm cnf.simp_TF_disj_True_r} OF [thm2]  (* (x | y) = True *)
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   282
            else if y' = \<^term>\<open>False\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   283
              @{thm cnf.simp_TF_disj_False_r} OF [thm1, thm2]  (* (x | y) = x' *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   284
            else
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   285
              @{thm cnf.disj_cong} OF [thm1, thm2]  (* (x | y) = (x' | y') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   286
          end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   287
      end
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   288
  | simp_True_False_thm thy t = inst_thm thy [t] @{thm cnf.iff_refl};  (* t = t *)
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
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   291
(* 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
   292
(*      is in conjunction normal form.  May cause an exponential blowup      *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   293
(*      in the length of the term.                                           *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   294
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   295
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   296
fun make_cnf_thm ctxt t =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   297
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42335
diff changeset
   298
    val thy = Proof_Context.theory_of ctxt
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   299
    fun make_cnf_thm_from_nnf (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   300
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   301
            val thm1 = make_cnf_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   302
            val thm2 = make_cnf_thm_from_nnf y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   303
          in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   304
            @{thm cnf.conj_cong} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   305
          end
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   306
      | make_cnf_thm_from_nnf (Const (\<^const_name>\<open>HOL.disj\<close>, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   307
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   308
            (* produces a theorem "(x' | y') = t'", where x', y', and t' are in CNF *)
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   309
            fun make_cnf_disj_thm (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ x1 $ x2) y' =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   310
                  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   311
                    val thm1 = make_cnf_disj_thm x1 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   312
                    val thm2 = make_cnf_disj_thm x2 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   313
                  in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   314
                    @{thm cnf.make_cnf_disj_conj_l} OF [thm1, thm2]  (* ((x1 & x2) | y') = ((x1 | y')' & (x2 | y')') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   315
                  end
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   316
              | make_cnf_disj_thm x' (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ y1 $ y2) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   317
                  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   318
                    val thm1 = make_cnf_disj_thm x' y1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   319
                    val thm2 = make_cnf_disj_thm x' y2
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   320
                  in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   321
                    @{thm cnf.make_cnf_disj_conj_r} OF [thm1, thm2]  (* (x' | (y1 & y2)) = ((x' | y1)' & (x' | y2)') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   322
                  end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   323
              | make_cnf_disj_thm x' y' =
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   324
                  inst_thm thy [HOLogic.mk_disj (x', y')] @{thm cnf.iff_refl}  (* (x' | y') = (x' | y') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   325
            val thm1 = make_cnf_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   326
            val thm2 = make_cnf_thm_from_nnf y
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   327
            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
   328
            val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   329
            val disj_thm = @{thm cnf.disj_cong} OF [thm1, thm2]  (* (x | y) = (x' | y') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   330
          in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   331
            @{thm cnf.iff_trans} OF [disj_thm, make_cnf_disj_thm x' y']
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   332
          end
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   333
      | make_cnf_thm_from_nnf t = inst_thm thy [t] @{thm cnf.iff_refl}
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   334
    (* convert 't' to NNF first *)
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   335
    val nnf_thm = make_nnf_thm_under_quantifiers ctxt t
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   336
(*###
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   337
    val nnf_thm = make_nnf_thm thy t
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   338
*)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   339
    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
   340
    (* then simplify wrt. True/False (this should preserve NNF) *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   341
    val simp_thm = simp_True_False_thm thy nnf
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   342
    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
   343
    (* finally, convert to CNF (this should preserve the simplification) *)
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   344
    val cnf_thm = make_under_quantifiers ctxt make_cnf_thm_from_nnf simp
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   345
(* ###
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   346
    val cnf_thm = make_cnf_thm_from_nnf simp
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   347
*)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   348
  in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   349
    @{thm cnf.iff_trans} OF [@{thm cnf.iff_trans} OF [nnf_thm, simp_thm], cnf_thm]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   350
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   351
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   352
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   353
(*            CNF transformation by introducing new literals                 *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   354
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   355
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   356
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   357
(* 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
   358
(*      t' is almost in conjunction normal form, except that conjunctions    *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   359
(*      and existential quantifiers may be nested.  (Use e.g. 'REPEAT_DETERM *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   360
(*      (etac exE i ORELSE etac conjE i)' afterwards to normalize.)  May     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   361
(*      introduce new (existentially bound) literals.  Note: the current     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   362
(*      implementation calls 'make_nnf_thm', causing an exponential blowup   *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   363
(*      in the case of nested equivalences.                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   364
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   365
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   366
fun make_cnfx_thm ctxt t =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   367
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42335
diff changeset
   368
    val thy = Proof_Context.theory_of ctxt
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   369
    val var_id = Unsynchronized.ref 0  (* properly initialized below *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   370
    fun new_free () =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   371
      Free ("cnfx_" ^ string_of_int (Unsynchronized.inc var_id), HOLogic.boolT)
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   372
    fun make_cnfx_thm_from_nnf (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ x $ y) : thm =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   373
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   374
            val thm1 = make_cnfx_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   375
            val thm2 = make_cnfx_thm_from_nnf y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   376
          in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   377
            @{thm cnf.conj_cong} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   378
          end
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   379
      | make_cnfx_thm_from_nnf (Const (\<^const_name>\<open>HOL.disj\<close>, _) $ x $ y) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   380
          if is_clause x andalso is_clause y then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   381
            inst_thm thy [HOLogic.mk_disj (x, y)] @{thm cnf.iff_refl}
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   382
          else if is_literal y orelse is_literal x then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   383
            let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   384
              (* produces a theorem "(x' | y') = t'", where x', y', and t' are *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   385
              (* almost in CNF, and x' or y' is a literal                      *)
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   386
              fun make_cnfx_disj_thm (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ x1 $ x2) y' =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   387
                    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   388
                      val thm1 = make_cnfx_disj_thm x1 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   389
                      val thm2 = make_cnfx_disj_thm x2 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   390
                    in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   391
                      @{thm cnf.make_cnf_disj_conj_l} OF [thm1, thm2]  (* ((x1 & x2) | y') = ((x1 | y')' & (x2 | y')') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   392
                    end
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   393
                | make_cnfx_disj_thm x' (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ y1 $ y2) =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   394
                    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   395
                      val thm1 = make_cnfx_disj_thm x' y1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   396
                      val thm2 = make_cnfx_disj_thm x' y2
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   397
                    in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   398
                      @{thm cnf.make_cnf_disj_conj_r} OF [thm1, thm2]  (* (x' | (y1 & y2)) = ((x' | y1)' & (x' | y2)') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   399
                    end
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   400
                | make_cnfx_disj_thm (\<^term>\<open>Ex :: (bool \<Rightarrow> bool) \<Rightarrow> bool\<close> $ x') y' =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   401
                    let
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   402
                      val thm1 = inst_thm thy [x', y'] @{thm cnf.make_cnfx_disj_ex_l}   (* ((Ex x') | y') = (Ex (x' | y')) *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   403
                      val var = new_free ()
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   404
                      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
   405
                      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
   406
                      val thm4 = Thm.strip_shyps (thm3 COMP allI)            (* ALL v. (x' | y') = body' *)
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   407
                      val thm5 = Thm.strip_shyps (thm4 RS @{thm cnf.make_cnfx_ex_cong}) (* (EX v. (x' | y')) = (EX v. body') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   408
                    in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   409
                      @{thm cnf.iff_trans} OF [thm1, thm5]  (* ((Ex x') | y') = (Ex v. body') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   410
                    end
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   411
                | make_cnfx_disj_thm x' (\<^term>\<open>Ex :: (bool \<Rightarrow> bool) \<Rightarrow> bool\<close> $ y') =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   412
                    let
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   413
                      val thm1 = inst_thm thy [x', y'] @{thm cnf.make_cnfx_disj_ex_r}   (* (x' | (Ex y')) = (Ex (x' | y')) *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   414
                      val var = new_free ()
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   415
                      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
   416
                      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
   417
                      val thm4 = Thm.strip_shyps (thm3 COMP allI)            (* ALL v. (x' | y') = body' *)
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   418
                      val thm5 = Thm.strip_shyps (thm4 RS @{thm cnf.make_cnfx_ex_cong}) (* (EX v. (x' | y')) = (EX v. body') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   419
                    in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   420
                      @{thm cnf.iff_trans} OF [thm1, thm5]  (* (x' | (Ex y')) = (EX v. body') *)
41447
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 x' y' =
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   423
                    inst_thm thy [HOLogic.mk_disj (x', y')] @{thm cnf.iff_refl}  (* (x' | y') = (x' | y') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   424
              val thm1 = make_cnfx_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   425
              val thm2 = make_cnfx_thm_from_nnf y
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   426
              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
   427
              val y' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm2
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   428
              val disj_thm = @{thm cnf.disj_cong} OF [thm1, thm2]  (* (x | y) = (x' | y') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   429
            in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   430
              @{thm cnf.iff_trans} OF [disj_thm, make_cnfx_disj_thm x' y']
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   431
            end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   432
          else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   433
            let  (* neither 'x' nor 'y' is a literal: introduce a fresh variable *)
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   434
              val thm1 = inst_thm thy [x, y] @{thm cnf.make_cnfx_newlit}     (* (x | y) = EX v. (x | v) & (y | ~v) *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   435
              val var = new_free ()
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   436
              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
   437
              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
   438
              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
   439
              val thm4 = Thm.strip_shyps (thm3 COMP allI)         (* ALL v. (x | v) & (y | ~v) = body' *)
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   440
              val thm5 = Thm.strip_shyps (thm4 RS @{thm cnf.make_cnfx_ex_cong})  (* (EX v. (x | v) & (y | ~v)) = (EX v. body') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   441
            in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   442
              @{thm cnf.iff_trans} OF [thm1, thm5]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   443
            end
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   444
      | make_cnfx_thm_from_nnf t = inst_thm thy [t] @{thm cnf.iff_refl}
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   445
    (* convert 't' to NNF first *)
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   446
    val nnf_thm = make_nnf_thm_under_quantifiers ctxt t
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   447
(* ###
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   448
    val nnf_thm = make_nnf_thm thy t
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   449
*)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   450
    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
   451
    (* then simplify wrt. True/False (this should preserve NNF) *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   452
    val simp_thm = simp_True_False_thm thy nnf
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   453
    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
   454
    (* 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
   455
    val _ = (var_id := fold (fn free => fn max =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   456
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   457
        val (name, _) = dest_Free free
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   458
        val idx =
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   459
          if String.isPrefix "cnfx_" name then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   460
            (Int.fromString o String.extract) (name, String.size "cnfx_", NONE)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   461
          else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   462
            NONE
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   463
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   464
        Int.max (max, the_default 0 idx)
44121
44adaa6db327 old term operations are legacy;
wenzelm
parents: 42361
diff changeset
   465
      end) (Misc_Legacy.term_frees simp) 0)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   466
    (* finally, convert to definitional CNF (this should preserve the simplification) *)
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   467
    val cnfx_thm = make_under_quantifiers ctxt make_cnfx_thm_from_nnf simp
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   468
(*###
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   469
    val cnfx_thm = make_cnfx_thm_from_nnf simp
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   470
*)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   471
  in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   472
    @{thm cnf.iff_trans} OF [@{thm cnf.iff_trans} OF [nnf_thm, simp_thm], cnfx_thm]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   473
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   474
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   475
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   476
(*                                  Tactics                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   477
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   478
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   479
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   480
(* weakening_tac: removes the first hypothesis of the 'i'-th subgoal         *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   481
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   482
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58839
diff changeset
   483
fun weakening_tac ctxt i =
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   484
  dresolve_tac ctxt @{thms cnf.weakening_thm} i THEN assume_tac ctxt (i+1);
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   485
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   486
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   487
(* 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
   488
(*      (possibly causing an exponential blowup in the length of each        *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   489
(*      premise)                                                             *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   490
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   491
32232
6c394343360f proper context for SAT tactics;
wenzelm
parents: 32231
diff changeset
   492
fun cnf_rewrite_tac ctxt i =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   493
  (* cut the CNF formulas as new premises *)
60696
8304fb4fb823 clarified context;
wenzelm
parents: 59642
diff changeset
   494
  Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   495
    let
60696
8304fb4fb823 clarified context;
wenzelm
parents: 59642
diff changeset
   496
      val cnf_thms = map (make_cnf_thm ctxt' o HOLogic.dest_Trueprop o Thm.prop_of) prems
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   497
      val cut_thms = map (fn (th, pr) => @{thm cnf.cnftac_eq_imp} OF [th, pr]) (cnf_thms ~~ prems)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   498
    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   499
      cut_facts_tac cut_thms 1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   500
    end) ctxt i
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   501
  (* remove the original premises *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   502
  THEN SELECT_GOAL (fn thm =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   503
    let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   504
      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
   505
    in
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58839
diff changeset
   506
      PRIMITIVE (funpow (n div 2) (Seq.hd o weakening_tac ctxt 1)) thm
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   507
    end) i;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   508
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   509
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   510
(* 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
   511
(*      (possibly introducing new literals)                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   512
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   513
32232
6c394343360f proper context for SAT tactics;
wenzelm
parents: 32231
diff changeset
   514
fun cnfx_rewrite_tac ctxt i =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   515
  (* cut the CNF formulas as new premises *)
60696
8304fb4fb823 clarified context;
wenzelm
parents: 59642
diff changeset
   516
  Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   517
    let
60696
8304fb4fb823 clarified context;
wenzelm
parents: 59642
diff changeset
   518
      val cnfx_thms = map (make_cnfx_thm ctxt' o HOLogic.dest_Trueprop o Thm.prop_of) prems
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   519
      val cut_thms = map (fn (th, pr) => @{thm cnf.cnftac_eq_imp} OF [th, pr]) (cnfx_thms ~~ prems)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   520
    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   521
      cut_facts_tac cut_thms 1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   522
    end) ctxt i
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   523
  (* remove the original premises *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   524
  THEN SELECT_GOAL (fn thm =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   525
    let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   526
      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
   527
    in
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58839
diff changeset
   528
      PRIMITIVE (funpow (n div 2) (Seq.hd o weakening_tac ctxt 1)) thm
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   529
    end) i;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   530
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   531
end;