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