src/HOL/Tools/cnf.ML
author wenzelm
Wed, 22 Jan 2025 22:37:38 +0100
changeset 81955 33616e13e172
parent 81954 6f2bcdfa9a19
permissions -rw-r--r--
tuned signature: more operations;
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
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
    43
  val make_nnf_thm: Proof.context -> 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
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
    57
fun is_atom \<^Const_>\<open>False\<close> = false
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
    58
  | is_atom \<^Const_>\<open>True\<close> = false
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
    59
  | is_atom \<^Const_>\<open>conj for _ _\<close> = false
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
    60
  | is_atom \<^Const_>\<open>disj for _ _\<close> = false
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
    61
  | is_atom \<^Const_>\<open>implies for _ _\<close> = false
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
    62
  | is_atom \<^Const_>\<open>HOL.eq \<^Type>\<open>bool\<close> for _ _\<close> = false
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
    63
  | is_atom \<^Const_>\<open>Not for _\<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
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
    66
fun is_literal \<^Const_>\<open>Not for x\<close> = 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
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
    69
fun is_clause \<^Const_>\<open>disj for x y\<close> = 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
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
    79
    fun dual \<^Const_>\<open>Not for x\<close> = x
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
    80
      | dual x = \<^Const>\<open>Not for x\<close>
41447
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
  in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   108
    (* [...] |- ~(x1 | ... | xn) ==> False *)
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   109
    (@{thm cnf.clause2raw_notE} OF [clause])
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   110
    (* [...] |- ~x1 ==> ... ==> ~xn ==> False *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   111
    |> not_disj_to_prem 1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   112
    (* [...] |- x1' ==> ... ==> xn' ==> False *)
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   113
    |> Seq.hd o TRYALL (resolve_tac ctxt @{thms cnf.clause2raw_not_not})
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   114
    (* [..., x1', ..., xn'] |- False *)
81955
33616e13e172 tuned signature: more operations;
wenzelm
parents: 81954
diff changeset
   115
    |> Thm.assume_prems ~1
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   116
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   117
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   118
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   119
(* inst_thm: instantiates a theorem with a list of terms                     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   120
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   121
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   122
fun inst_thm ctxt ts thm =
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   123
  Thm.instantiate' [] (map (SOME o Thm.cterm_of ctxt) ts) thm;
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
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   126
(*                         Naive CNF transformation                          *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   127
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   128
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   129
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   130
(* 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
   131
(*      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
   132
(*      of t; implications ("-->") and equivalences ("=" on bool) are        *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   133
(*      eliminated (possibly causing an exponential blowup)                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   134
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   135
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   136
fun make_nnf_thm ctxt \<^Const_>\<open>conj for x y\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   137
      let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   138
        val thm1 = make_nnf_thm ctxt x
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   139
        val thm2 = make_nnf_thm ctxt y
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   140
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   141
        @{thm cnf.conj_cong} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   142
      end
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   143
  | make_nnf_thm ctxt \<^Const_>\<open>disj for x y\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   144
      let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   145
        val thm1 = make_nnf_thm ctxt x
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   146
        val thm2 = make_nnf_thm ctxt y
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   147
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   148
        @{thm cnf.disj_cong} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   149
      end
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   150
  | make_nnf_thm ctxt \<^Const_>\<open>implies for x y\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   151
      let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   152
        val thm1 = make_nnf_thm ctxt \<^Const>\<open>Not for x\<close>
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   153
        val thm2 = make_nnf_thm ctxt y
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   154
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   155
        @{thm cnf.make_nnf_imp} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   156
      end
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   157
  | make_nnf_thm ctxt \<^Const_>\<open>HOL.eq \<^Type>\<open>bool\<close> for x y\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   158
      let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   159
        val thm1 = make_nnf_thm ctxt x
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   160
        val thm2 = make_nnf_thm ctxt \<^Const>\<open>Not for x\<close>
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   161
        val thm3 = make_nnf_thm ctxt y
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   162
        val thm4 = make_nnf_thm ctxt \<^Const>\<open>Not for y\<close>
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   163
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   164
        @{thm cnf.make_nnf_iff} OF [thm1, thm2, thm3, thm4]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   165
      end
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   166
  | make_nnf_thm _ \<^Const_>\<open>Not for \<^Const_>\<open>False\<close>\<close> =
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   167
      @{thm cnf.make_nnf_not_false}
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   168
  | make_nnf_thm _ \<^Const_>\<open>Not for \<^Const_>\<open>True\<close>\<close> =
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   169
      @{thm cnf.make_nnf_not_true}
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   170
  | make_nnf_thm ctxt \<^Const_>\<open>Not for \<^Const_>\<open>conj for x y\<close>\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   171
      let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   172
        val thm1 = make_nnf_thm ctxt \<^Const>\<open>Not for x\<close>
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   173
        val thm2 = make_nnf_thm ctxt \<^Const>\<open>Not for y\<close>
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   174
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   175
        @{thm cnf.make_nnf_not_conj} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   176
      end
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   177
  | make_nnf_thm ctxt \<^Const_>\<open>Not for \<^Const_>\<open>disj for x y\<close>\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   178
      let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   179
        val thm1 = make_nnf_thm ctxt \<^Const>\<open>Not for x\<close>
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   180
        val thm2 = make_nnf_thm ctxt \<^Const>\<open>Not for y\<close>
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   181
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   182
        @{thm cnf.make_nnf_not_disj} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   183
      end
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   184
  | make_nnf_thm ctxt \<^Const_>\<open>Not for \<^Const_>\<open>implies for x y\<close>\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   185
      let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   186
        val thm1 = make_nnf_thm ctxt x
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   187
        val thm2 = make_nnf_thm ctxt \<^Const>\<open>Not for y\<close>
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   188
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   189
        @{thm cnf.make_nnf_not_imp} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   190
      end
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   191
  | make_nnf_thm ctxt \<^Const_>\<open>Not for \<^Const_>\<open>HOL.eq \<^Type>\<open>bool\<close> for x y\<close>\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   192
      let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   193
        val thm1 = make_nnf_thm ctxt x
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   194
        val thm2 = make_nnf_thm ctxt \<^Const>\<open>Not for x\<close>
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   195
        val thm3 = make_nnf_thm ctxt y
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   196
        val thm4 = make_nnf_thm ctxt \<^Const>\<open>Not for y\<close>
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   197
      in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   198
        @{thm cnf.make_nnf_not_iff} OF [thm1, thm2, thm3, thm4]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   199
      end
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   200
  | make_nnf_thm ctxt \<^Const_>\<open>Not for \<^Const_>\<open>Not for x\<close>\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   201
      let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   202
        val thm1 = make_nnf_thm ctxt x
41447
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_not} OF [thm1]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   205
      end
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   206
  | make_nnf_thm ctxt t = inst_thm ctxt [t] @{thm cnf.iff_refl};
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   207
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   208
fun make_under_quantifiers ctxt make t =
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   209
  let
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   210
    fun conv ctxt ct =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   211
      (case Thm.term_of ct of
45511
9b0f8ca4388e continued implementation of lambda-lifting in Metis
blanchet
parents: 44121
diff changeset
   212
        Const _ $ Abs _ => Conv.comb_conv (conv ctxt) ct
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   213
      | Abs _ => Conv.abs_conv (conv o snd) ctxt ct
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   214
      | Const _ => Conv.all_conv ct
67710
cc2db3239932 added HOLogic.mk_obj_eq convenience and eliminated some clones;
wenzelm
parents: 67149
diff changeset
   215
      | t => make t RS @{thm eq_reflection})
cc2db3239932 added HOLogic.mk_obj_eq convenience and eliminated some clones;
wenzelm
parents: 67149
diff changeset
   216
  in HOLogic.mk_obj_eq (conv ctxt (Thm.cterm_of ctxt t)) end
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   217
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   218
fun make_nnf_thm_under_quantifiers ctxt =
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   219
  make_under_quantifiers ctxt (make_nnf_thm ctxt)
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   220
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   221
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   222
(* 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
   223
(*      t, but simplified wrt. the following theorems:                       *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   224
(*        (True & x) = x                                                     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   225
(*        (x & True) = x                                                     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   226
(*        (False & x) = False                                                *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   227
(*        (x & False) = False                                                *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   228
(*        (True | x) = True                                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   229
(*        (x | True) = True                                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   230
(*        (False | x) = x                                                    *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   231
(*        (x | False) = x                                                    *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   232
(*      No simplification is performed below connectives other than & and |. *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   233
(*      Optimization: The right-hand side of a conjunction (disjunction) is  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   234
(*      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
   235
(*      (True, respectively).                                                *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   236
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   237
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   238
fun simp_True_False_thm ctxt \<^Const_>\<open>conj for x y\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   239
      let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   240
        val thm1 = simp_True_False_thm ctxt x
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   241
        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
   242
      in
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   243
        if x' = \<^Const>\<open>False\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   244
          @{thm cnf.simp_TF_conj_False_l} OF [thm1]  (* (x & y) = False *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   245
        else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   246
          let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   247
            val thm2 = simp_True_False_thm ctxt y
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   248
            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
   249
          in
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   250
            if x' = \<^Const>\<open>True\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   251
              @{thm cnf.simp_TF_conj_True_l} OF [thm1, thm2]  (* (x & y) = y' *)
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   252
            else if y' = \<^Const>\<open>False\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   253
              @{thm cnf.simp_TF_conj_False_r} OF [thm2]  (* (x & y) = False *)
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   254
            else if y' = \<^Const>\<open>True\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   255
              @{thm cnf.simp_TF_conj_True_r} OF [thm1, thm2]  (* (x & y) = x' *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   256
            else
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   257
              @{thm cnf.conj_cong} OF [thm1, thm2]  (* (x & y) = (x' & y') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   258
          end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   259
      end
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   260
  | simp_True_False_thm ctxt \<^Const_>\<open>disj for x y\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   261
      let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   262
        val thm1 = simp_True_False_thm ctxt x
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   263
        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
   264
      in
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   265
        if x' = \<^Const>\<open>True\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   266
          @{thm cnf.simp_TF_disj_True_l} OF [thm1]  (* (x | y) = True *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   267
        else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   268
          let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   269
            val thm2 = simp_True_False_thm ctxt y
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   270
            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
   271
          in
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   272
            if x' = \<^Const>\<open>False\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   273
              @{thm cnf.simp_TF_disj_False_l} OF [thm1, thm2]  (* (x | y) = y' *)
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   274
            else if y' = \<^Const>\<open>True\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   275
              @{thm cnf.simp_TF_disj_True_r} OF [thm2]  (* (x | y) = True *)
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   276
            else if y' = \<^Const>\<open>False\<close> then
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   277
              @{thm cnf.simp_TF_disj_False_r} OF [thm1, thm2]  (* (x | y) = x' *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   278
            else
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   279
              @{thm cnf.disj_cong} OF [thm1, thm2]  (* (x | y) = (x' | y') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   280
          end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   281
      end
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   282
  | simp_True_False_thm ctxt t = inst_thm ctxt [t] @{thm cnf.iff_refl};  (* t = t *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   283
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   284
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   285
(* 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
   286
(*      is in conjunction normal form.  May cause an exponential blowup      *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   287
(*      in the length of the term.                                           *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   288
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   289
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   290
fun make_cnf_thm ctxt t =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   291
  let
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   292
    fun make_cnf_thm_from_nnf \<^Const_>\<open>conj for x y\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   293
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   294
            val thm1 = make_cnf_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   295
            val thm2 = make_cnf_thm_from_nnf y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   296
          in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   297
            @{thm cnf.conj_cong} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   298
          end
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   299
      | make_cnf_thm_from_nnf \<^Const_>\<open>disj for x y\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   300
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   301
            (* produces a theorem "(x' | y') = t'", where x', y', and t' are in CNF *)
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   302
            fun make_cnf_disj_thm \<^Const_>\<open>conj for x1 x2\<close> y' =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   303
                  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   304
                    val thm1 = make_cnf_disj_thm x1 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   305
                    val thm2 = make_cnf_disj_thm x2 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   306
                  in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   307
                    @{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
   308
                  end
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   309
              | make_cnf_disj_thm x' \<^Const_>\<open>conj for y1 y2\<close> =
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 x' y1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   312
                    val thm2 = make_cnf_disj_thm x' y2
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_r} OF [thm1, thm2]  (* (x' | (y1 & y2)) = ((x' | y1)' & (x' | y2)') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   315
                  end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   316
              | make_cnf_disj_thm x' y' =
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   317
                  inst_thm ctxt [\<^Const>\<open>disj for x' y'\<close>] @{thm cnf.iff_refl}  (* (x' | y') = (x' | y') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   318
            val thm1 = make_cnf_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   319
            val thm2 = make_cnf_thm_from_nnf y
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   320
            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
   321
            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
   322
            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
   323
          in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   324
            @{thm cnf.iff_trans} OF [disj_thm, make_cnf_disj_thm x' y']
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   325
          end
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   326
      | make_cnf_thm_from_nnf t = inst_thm ctxt [t] @{thm cnf.iff_refl}
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   327
    (* convert 't' to NNF first *)
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   328
    val nnf_thm = make_nnf_thm_under_quantifiers ctxt t
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   329
(*###
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   330
    val nnf_thm = make_nnf_thm ctxt t
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   331
*)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   332
    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
   333
    (* then simplify wrt. True/False (this should preserve NNF) *)
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   334
    val simp_thm = simp_True_False_thm ctxt nnf
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   335
    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
   336
    (* finally, convert to CNF (this should preserve the simplification) *)
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   337
    val cnf_thm = make_under_quantifiers ctxt make_cnf_thm_from_nnf simp
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   338
(* ###
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   339
    val cnf_thm = make_cnf_thm_from_nnf simp
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   340
*)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   341
  in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   342
    @{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
   343
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   344
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   345
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   346
(*            CNF transformation by introducing new literals                 *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   347
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   348
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   349
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   350
(* 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
   351
(*      t' is almost in conjunction normal form, except that conjunctions    *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   352
(*      and existential quantifiers may be nested.  (Use e.g. 'REPEAT_DETERM *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   353
(*      (etac exE i ORELSE etac conjE i)' afterwards to normalize.)  May     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   354
(*      introduce new (existentially bound) literals.  Note: the current     *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   355
(*      implementation calls 'make_nnf_thm', causing an exponential blowup   *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   356
(*      in the case of nested equivalences.                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   357
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   358
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   359
fun make_cnfx_thm ctxt t =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   360
  let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   361
    val var_id = Unsynchronized.ref 0  (* properly initialized below *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   362
    fun new_free () =
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   363
      Free ("cnfx_" ^ string_of_int (Unsynchronized.inc var_id), \<^Type>\<open>bool\<close>)
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   364
    fun make_cnfx_thm_from_nnf \<^Const_>\<open>conj for x y\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   365
          let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   366
            val thm1 = make_cnfx_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   367
            val thm2 = make_cnfx_thm_from_nnf y
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   368
          in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   369
            @{thm cnf.conj_cong} OF [thm1, thm2]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   370
          end
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   371
      | make_cnfx_thm_from_nnf \<^Const_>\<open>disj for x y\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   372
          if is_clause x andalso is_clause y then
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   373
            inst_thm ctxt [\<^Const>\<open>disj for x y\<close>] @{thm cnf.iff_refl}
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   374
          else if is_literal y orelse is_literal x then
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   375
            let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   376
              (* produces a theorem "(x' | y') = t'", where x', y', and t' are *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   377
              (* almost in CNF, and x' or y' is a literal                      *)
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   378
              fun make_cnfx_disj_thm \<^Const_>\<open>conj for x1 x2\<close> y' =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   379
                    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   380
                      val thm1 = make_cnfx_disj_thm x1 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   381
                      val thm2 = make_cnfx_disj_thm x2 y'
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   382
                    in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   383
                      @{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
   384
                    end
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   385
                | make_cnfx_disj_thm x' \<^Const_>\<open>conj for y1 y2\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   386
                    let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   387
                      val thm1 = make_cnfx_disj_thm x' y1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   388
                      val thm2 = make_cnfx_disj_thm x' y2
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   389
                    in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   390
                      @{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
   391
                    end
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   392
                | make_cnfx_disj_thm \<^Const_>\<open>Ex \<^Type>\<open>bool\<close> for x'\<close> y' =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   393
                    let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   394
                      val thm1 = inst_thm ctxt [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
   395
                      val var = new_free ()
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   396
                      val thm2 = make_cnfx_disj_thm (betapply (x', var)) y'  (* (x' | y') = body' *)
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   397
                      val thm3 = Thm.forall_intr (Thm.cterm_of ctxt var) thm2 (* !!v. (x' | y') = body' *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   398
                      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
   399
                      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
   400
                    in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   401
                      @{thm cnf.iff_trans} OF [thm1, thm5]  (* ((Ex x') | y') = (Ex v. body') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   402
                    end
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   403
                | make_cnfx_disj_thm x' \<^Const_>\<open>Ex \<^Type>\<open>bool\<close> for y'\<close> =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   404
                    let
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   405
                      val thm1 = inst_thm ctxt [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
   406
                      val var = new_free ()
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   407
                      val thm2 = make_cnfx_disj_thm x' (betapply (y', var))  (* (x' | y') = body' *)
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   408
                      val thm3 = Thm.forall_intr (Thm.cterm_of ctxt var) thm2 (* !!v. (x' | y') = body' *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   409
                      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
   410
                      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
   411
                    in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   412
                      @{thm cnf.iff_trans} OF [thm1, thm5]  (* (x' | (Ex y')) = (EX v. body') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   413
                    end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   414
                | make_cnfx_disj_thm x' y' =
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   415
                    inst_thm ctxt [\<^Const>\<open>disj for x' y'\<close>] @{thm cnf.iff_refl}  (* (x' | y') = (x' | y') *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   416
              val thm1 = make_cnfx_thm_from_nnf x
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   417
              val thm2 = make_cnfx_thm_from_nnf y
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   418
              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
   419
              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
   420
              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
   421
            in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   422
              @{thm cnf.iff_trans} OF [disj_thm, make_cnfx_disj_thm x' y']
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   423
            end
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   424
          else
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   425
            let  (* neither 'x' nor 'y' is a literal: introduce a fresh variable *)
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   426
              val thm1 = inst_thm ctxt [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
   427
              val var = new_free ()
80637
e2f0265f64e3 tuned: more antiquotations;
wenzelm
parents: 70486
diff changeset
   428
              val body = \<^Const>\<open>conj for \<^Const>\<open>disj for x var\<close> \<^Const>\<open>disj for y \<^Const>\<open>Not for var\<close>\<close>\<close>
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   429
              val thm2 = make_cnfx_thm_from_nnf body              (* (x | v) & (y | ~v) = body' *)
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   430
              val thm3 = Thm.forall_intr (Thm.cterm_of ctxt var) thm2 (* !!v. (x | v) & (y | ~v) = body' *)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   431
              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
   432
              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
   433
            in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   434
              @{thm cnf.iff_trans} OF [thm1, thm5]
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   435
            end
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   436
      | make_cnfx_thm_from_nnf t = inst_thm ctxt [t] @{thm cnf.iff_refl}
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   437
    (* convert 't' to NNF first *)
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   438
    val nnf_thm = make_nnf_thm_under_quantifiers ctxt t
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   439
(* ###
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   440
    val nnf_thm = make_nnf_thm ctxt t
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   441
*)
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   442
    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
   443
    (* then simplify wrt. True/False (this should preserve NNF) *)
80638
21637b691fab clarified signature: prefer local context;
wenzelm
parents: 80637
diff changeset
   444
    val simp_thm = simp_True_False_thm ctxt nnf
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   445
    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
   446
    (* initialize var_id, in case the term already contains variables of the form "cnfx_<int>" *)
80639
3322b6ae6b19 tuned: more standard Isabelle/ML;
wenzelm
parents: 80638
diff changeset
   447
    val _ = (var_id := Frees.fold (fn ((name, _), _) => fn max =>
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   448
      let
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   449
        val idx =
80639
3322b6ae6b19 tuned: more standard Isabelle/ML;
wenzelm
parents: 80638
diff changeset
   450
          (case try (unprefix "cnfx_") name of
3322b6ae6b19 tuned: more standard Isabelle/ML;
wenzelm
parents: 80638
diff changeset
   451
            SOME s => Int.fromString s
3322b6ae6b19 tuned: more standard Isabelle/ML;
wenzelm
parents: 80638
diff changeset
   452
          | NONE => NONE)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   453
      in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   454
        Int.max (max, the_default 0 idx)
80639
3322b6ae6b19 tuned: more standard Isabelle/ML;
wenzelm
parents: 80638
diff changeset
   455
      end) (Frees.build (Frees.add_frees simp)) 0)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   456
    (* finally, convert to definitional CNF (this should preserve the simplification) *)
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   457
    val cnfx_thm = make_under_quantifiers ctxt make_cnfx_thm_from_nnf simp
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   458
(*###
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   459
    val cnfx_thm = make_cnfx_thm_from_nnf simp
42335
cb8aa792d138 experiment with definitional CNF
blanchet
parents: 41447
diff changeset
   460
*)
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   461
  in
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   462
    @{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
   463
  end;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   464
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   465
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   466
(*                                  Tactics                                  *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   467
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   468
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   469
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   470
(* weakening_tac: removes the first hypothesis of the 'i'-th subgoal         *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   471
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   472
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58839
diff changeset
   473
fun weakening_tac ctxt i =
70486
1dc3514c1719 prefer named lemmas -- more compact proofterms;
wenzelm
parents: 70484
diff changeset
   474
  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
   475
17809
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   476
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   477
(* 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
   478
(*      (possibly causing an exponential blowup in the length of each        *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   479
(*      premise)                                                             *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   480
(* ------------------------------------------------------------------------- *)
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   481
32232
6c394343360f proper context for SAT tactics;
wenzelm
parents: 32231
diff changeset
   482
fun cnf_rewrite_tac ctxt i =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   483
  (* cut the CNF formulas as new premises *)
60696
8304fb4fb823 clarified context;
wenzelm
parents: 59642
diff changeset
   484
  Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   485
    let
60696
8304fb4fb823 clarified context;
wenzelm
parents: 59642
diff changeset
   486
      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
   487
      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
   488
    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   489
      cut_facts_tac cut_thms 1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   490
    end) ctxt i
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   491
  (* remove the original premises *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   492
  THEN SELECT_GOAL (fn thm =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   493
    let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   494
      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
   495
    in
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58839
diff changeset
   496
      PRIMITIVE (funpow (n div 2) (Seq.hd o weakening_tac ctxt 1)) thm
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   497
    end) i;
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
(* ------------------------------------------------------------------------- *)
195045659c06 Tactics sat and satx reimplemented, several improvements
webertj
parents: 17618
diff changeset
   500
(* 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
   501
(*      (possibly introducing new literals)                                  *)
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
32232
6c394343360f proper context for SAT tactics;
wenzelm
parents: 32231
diff changeset
   504
fun cnfx_rewrite_tac ctxt i =
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   505
  (* cut the CNF formulas as new premises *)
60696
8304fb4fb823 clarified context;
wenzelm
parents: 59642
diff changeset
   506
  Subgoal.FOCUS (fn {prems, context = ctxt', ...} =>
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   507
    let
60696
8304fb4fb823 clarified context;
wenzelm
parents: 59642
diff changeset
   508
      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
   509
      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
   510
    in
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   511
      cut_facts_tac cut_thms 1
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   512
    end) ctxt i
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   513
  (* remove the original premises *)
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   514
  THEN SELECT_GOAL (fn thm =>
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   515
    let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   516
      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
   517
    in
58963
26bf09b95dda proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents: 58839
diff changeset
   518
      PRIMITIVE (funpow (n div 2) (Seq.hd o weakening_tac ctxt 1)) thm
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   519
    end) i;
17618
1330157e156a new sat tactic imports resolution proofs from zChaff
webertj
parents:
diff changeset
   520
41447
537b290bbe38 tuned whitespace, indentation, comments;
wenzelm
parents: 39035
diff changeset
   521
end;