| author | paulson | 
| Mon, 22 May 2000 12:29:02 +0200 | |
| changeset 8913 | 0bc13d5e60b8 | 
| parent 8644 | c47735e7bd1c | 
| child 8955 | 714497ad2348 | 
| permissions | -rw-r--r-- | 
| 1465 | 1 | (* Title: HOL/simpdata.ML | 
| 923 | 2 | ID: $Id$ | 
| 1465 | 3 | Author: Tobias Nipkow | 
| 923 | 4 | Copyright 1991 University of Cambridge | 
| 5 | ||
| 5304 | 6 | Instantiation of the generic simplifier for HOL. | 
| 923 | 7 | *) | 
| 8 | ||
| 1984 | 9 | section "Simplifier"; | 
| 10 | ||
| 6514 | 11 | (*** Addition of rules to simpsets and clasets simultaneously ***) (* FIXME move to Provers/clasimp.ML? *) | 
| 1984 | 12 | |
| 5190 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 13 | infix 4 addIffs delIffs; | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 14 | |
| 1984 | 15 | (*Takes UNCONDITIONAL theorems of the form A<->B to | 
| 2031 | 16 | the Safe Intr rule B==>A and | 
| 17 | the Safe Destruct rule A==>B. | |
| 1984 | 18 | Also ~A goes to the Safe Elim rule A ==> ?R | 
| 19 | Failing other cases, A is added as a Safe Intr rule*) | |
| 20 | local | |
| 21 | val iff_const = HOLogic.eq_const HOLogic.boolT; | |
| 22 | ||
| 5190 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 23 | fun addIff ((cla, simp), th) = | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 24 | (case HOLogic.dest_Trueprop (#prop (rep_thm th)) of | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 25 |                 (Const("Not", _) $ A) =>
 | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 26 | cla addSEs [zero_var_indexes (th RS notE)] | 
| 2031 | 27 | | (con $ _ $ _) => | 
| 5190 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 28 | if con = iff_const | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 29 | then cla addSIs [zero_var_indexes (th RS iffD2)] | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 30 | addSDs [zero_var_indexes (th RS iffD1)] | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 31 | else cla addSIs [th] | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 32 | | _ => cla addSIs [th], | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 33 | simp addsimps [th]) | 
| 6968 | 34 |       handle TERM _ => error ("AddIffs: theorem must be unconditional\n" ^ 
 | 
| 5190 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 35 | string_of_thm th); | 
| 1984 | 36 | |
| 5190 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 37 | fun delIff ((cla, simp), th) = | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 38 | (case HOLogic.dest_Trueprop (#prop (rep_thm th)) of | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 39 |                 (Const ("Not", _) $ A) =>
 | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 40 | cla delrules [zero_var_indexes (th RS notE)] | 
| 2031 | 41 | | (con $ _ $ _) => | 
| 5190 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 42 | if con = iff_const | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 43 | then cla delrules [zero_var_indexes (th RS iffD2), | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 44 | make_elim (zero_var_indexes (th RS iffD1))] | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 45 | else cla delrules [th] | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 46 | | _ => cla delrules [th], | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 47 | simp delsimps [th]) | 
| 6968 | 48 |       handle TERM _ => (warning("DelIffs: ignoring conditional theorem\n" ^ 
 | 
| 5190 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 49 | string_of_thm th); (cla, simp)); | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 50 | |
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 51 | fun store_clasimp (cla, simp) = (claset_ref () := cla; simpset_ref () := simp) | 
| 1984 | 52 | in | 
| 5190 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 53 | val op addIffs = foldl addIff; | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 54 | val op delIffs = foldl delIff; | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 55 | fun AddIffs thms = store_clasimp ((claset (), simpset ()) addIffs thms); | 
| 
4ae031622592
Added functions addIffs and delIffs which operate on clasimpsets.
 berghofe parents: 
4930diff
changeset | 56 | fun DelIffs thms = store_clasimp ((claset (), simpset ()) delIffs thms); | 
| 1984 | 57 | end; | 
| 58 | ||
| 5304 | 59 | |
| 7357 | 60 | val [prem] = goal (the_context ()) "x==y ==> x=y"; | 
| 7031 | 61 | by (rewtac prem); | 
| 62 | by (rtac refl 1); | |
| 63 | qed "meta_eq_to_obj_eq"; | |
| 4640 | 64 | |
| 923 | 65 | local | 
| 66 | ||
| 7357 | 67 | fun prover s = prove_goal (the_context ()) s (fn _ => [(Blast_tac 1)]); | 
| 923 | 68 | |
| 2134 | 69 | in | 
| 70 | ||
| 5552 | 71 | (*Make meta-equalities. The operator below is Trueprop*) | 
| 72 | ||
| 6128 | 73 | fun mk_meta_eq r = r RS eq_reflection; | 
| 74 | ||
| 75 | val Eq_TrueI = mk_meta_eq(prover "P --> (P = True)" RS mp); | |
| 76 | val Eq_FalseI = mk_meta_eq(prover "~P --> (P = False)" RS mp); | |
| 5304 | 77 | |
| 6128 | 78 | fun mk_eq th = case concl_of th of | 
| 79 |         Const("==",_)$_$_       => th
 | |
| 80 |     |   _$(Const("op =",_)$_$_) => mk_meta_eq th
 | |
| 81 |     |   _$(Const("Not",_)$_)    => th RS Eq_FalseI
 | |
| 82 | | _ => th RS Eq_TrueI; | |
| 83 | (* last 2 lines requires all formulae to be of the from Trueprop(.) *) | |
| 5304 | 84 | |
| 6128 | 85 | fun mk_eq_True r = Some(r RS meta_eq_to_obj_eq RS Eq_TrueI); | 
| 5552 | 86 | |
| 6128 | 87 | fun mk_meta_cong rl = | 
| 88 | standard(mk_meta_eq(replicate (nprems_of rl) meta_eq_to_obj_eq MRS rl)) | |
| 89 | handle THM _ => | |
| 90 |   error("Premises and conclusion of congruence rules must be =-equalities");
 | |
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3842diff
changeset | 91 | |
| 5975 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 92 | val not_not = prover "(~ ~ P) = P"; | 
| 923 | 93 | |
| 5975 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 94 | val simp_thms = [not_not] @ map prover | 
| 2082 | 95 | [ "(x=x) = True", | 
| 5975 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 96 | "(~True) = False", "(~False) = True", | 
| 2082 | 97 | "(~P) ~= P", "P ~= (~P)", "(P ~= Q) = (P = (~Q))", | 
| 4640 | 98 | "(True=P) = P", "(P=True) = P", "(False=P) = (~P)", "(P=False) = (~P)", | 
| 2082 | 99 | "(True --> P) = P", "(False --> P) = True", | 
| 100 | "(P --> True) = True", "(P --> P) = True", | |
| 101 | "(P --> False) = (~P)", "(P --> ~P) = (~P)", | |
| 102 | "(P & True) = P", "(True & P) = P", | |
| 2800 | 103 | "(P & False) = False", "(False & P) = False", | 
| 104 | "(P & P) = P", "(P & (P & Q)) = (P & Q)", | |
| 3913 | 105 | "(P & ~P) = False", "(~P & P) = False", | 
| 2082 | 106 | "(P | True) = True", "(True | P) = True", | 
| 2800 | 107 | "(P | False) = P", "(False | P) = P", | 
| 108 | "(P | P) = P", "(P | (P | Q)) = (P | Q)", | |
| 3913 | 109 | "(P | ~P) = True", "(~P | P) = True", | 
| 2082 | 110 | "((~P) = (~Q)) = (P=Q)", | 
| 3842 | 111 | "(!x. P) = P", "(? x. P) = P", "? x. x=t", "? x. t=x", | 
| 4351 | 112 | (*two needed for the one-point-rule quantifier simplification procs*) | 
| 113 | "(? x. x=t & P(x)) = P(t)", (*essential for termination!!*) | |
| 114 | "(! x. t=x --> P(x)) = P(t)" ]; (*covers a stray case*) | |
| 923 | 115 | |
| 5552 | 116 | (* Add congruence rules for = (instead of ==) *) | 
| 4351 | 117 | |
| 5552 | 118 | (* ###FIXME: Move to simplifier, | 
| 119 | taking mk_meta_cong as input, eliminating addeqcongs and deleqcongs *) | |
| 120 | infix 4 addcongs delcongs; | |
| 4640 | 121 | fun ss addcongs congs = ss addeqcongs (map mk_meta_cong congs); | 
| 122 | fun ss delcongs congs = ss deleqcongs (map mk_meta_cong congs); | |
| 4086 | 123 | fun Addcongs congs = (simpset_ref() := simpset() addcongs congs); | 
| 124 | fun Delcongs congs = (simpset_ref() := simpset() delcongs congs); | |
| 1264 | 125 | |
| 8641 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 126 | val cong_add_global = Simplifier.change_global_ss (op addcongs); | 
| 8644 | 127 | val cong_del_global = Simplifier.change_global_ss (op delcongs); | 
| 8641 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 128 | val cong_add_local = Simplifier.change_local_ss (op addcongs); | 
| 8644 | 129 | val cong_del_local = Simplifier.change_local_ss (op delcongs); | 
| 8641 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 130 | |
| 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 131 | val cong_attrib_setup = | 
| 8644 | 132 |  [Attrib.add_attributes [("cong",
 | 
| 133 | (Attrib.add_del_args cong_add_global cong_del_global, | |
| 134 | Attrib.add_del_args cong_add_local cong_del_local), | |
| 135 | "declare Simplifier congruence rules")]]; | |
| 8641 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 136 | |
| 5552 | 137 | |
| 1922 | 138 | val imp_cong = impI RSN | 
| 7357 | 139 | (2, prove_goal (the_context ()) "(P=P')--> (P'--> (Q=Q'))--> ((P-->Q) = (P'-->Q'))" | 
| 7031 | 140 | (fn _=> [(Blast_tac 1)]) RS mp RS mp); | 
| 1922 | 141 | |
| 1948 
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
 paulson parents: 
1922diff
changeset | 142 | (*Miniscoping: pushing in existential quantifiers*) | 
| 7648 | 143 | val ex_simps = map prover | 
| 3842 | 144 | ["(EX x. P x & Q) = ((EX x. P x) & Q)", | 
| 145 | "(EX x. P & Q x) = (P & (EX x. Q x))", | |
| 146 | "(EX x. P x | Q) = ((EX x. P x) | Q)", | |
| 147 | "(EX x. P | Q x) = (P | (EX x. Q x))", | |
| 148 | "(EX x. P x --> Q) = ((ALL x. P x) --> Q)", | |
| 149 | "(EX x. P --> Q x) = (P --> (EX x. Q x))"]; | |
| 1948 
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
 paulson parents: 
1922diff
changeset | 150 | |
| 
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
 paulson parents: 
1922diff
changeset | 151 | (*Miniscoping: pushing in universal quantifiers*) | 
| 
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
 paulson parents: 
1922diff
changeset | 152 | val all_simps = map prover | 
| 3842 | 153 | ["(ALL x. P x & Q) = ((ALL x. P x) & Q)", | 
| 154 | "(ALL x. P & Q x) = (P & (ALL x. Q x))", | |
| 155 | "(ALL x. P x | Q) = ((ALL x. P x) | Q)", | |
| 156 | "(ALL x. P | Q x) = (P | (ALL x. Q x))", | |
| 157 | "(ALL x. P x --> Q) = ((EX x. P x) --> Q)", | |
| 158 | "(ALL x. P --> Q x) = (P --> (ALL x. Q x))"]; | |
| 1948 
78e5bfcbc1e9
Added miniscoping to the simplifier: quantifiers are now pushed in
 paulson parents: 
1922diff
changeset | 159 | |
| 923 | 160 | |
| 2022 | 161 | (* elimination of existential quantifiers in assumptions *) | 
| 923 | 162 | |
| 163 | val ex_all_equiv = | |
| 7357 | 164 | let val lemma1 = prove_goal (the_context ()) | 
| 923 | 165 | "(? x. P(x) ==> PROP Q) ==> (!!x. P(x) ==> PROP Q)" | 
| 166 | (fn prems => [resolve_tac prems 1, etac exI 1]); | |
| 7357 | 167 | val lemma2 = prove_goalw (the_context ()) [Ex_def] | 
| 923 | 168 | "(!!x. P(x) ==> PROP Q) ==> (? x. P(x) ==> PROP Q)" | 
| 7031 | 169 | (fn prems => [(REPEAT(resolve_tac prems 1))]) | 
| 923 | 170 | in equal_intr lemma1 lemma2 end; | 
| 171 | ||
| 172 | end; | |
| 173 | ||
| 7648 | 174 | bind_thms ("ex_simps", ex_simps);
 | 
| 175 | bind_thms ("all_simps", all_simps);
 | |
| 7711 
4dae7a4fceaf
Rule not_not is now stored in theory (needed by Inductive).
 berghofe parents: 
7648diff
changeset | 176 | bind_thm ("not_not", not_not);
 | 
| 7648 | 177 | |
| 3654 | 178 | (* Elimination of True from asumptions: *) | 
| 179 | ||
| 7357 | 180 | val True_implies_equals = prove_goal (the_context ()) | 
| 3654 | 181 | "(True ==> PROP P) == PROP P" | 
| 7031 | 182 | (fn _ => [rtac equal_intr_rule 1, atac 2, | 
| 3654 | 183 | METAHYPS (fn prems => resolve_tac prems 1) 1, | 
| 184 | rtac TrueI 1]); | |
| 185 | ||
| 7357 | 186 | fun prove nm thm = qed_goal nm (the_context ()) thm (fn _ => [(Blast_tac 1)]); | 
| 923 | 187 | |
| 7623 | 188 | prove "eq_commute" "(a=b)=(b=a)"; | 
| 189 | prove "eq_left_commute" "(P=(Q=R)) = (Q=(P=R))"; | |
| 190 | prove "eq_assoc" "((P=Q)=R) = (P=(Q=R))"; | |
| 191 | val eq_ac = [eq_commute, eq_left_commute, eq_assoc]; | |
| 192 | ||
| 923 | 193 | prove "conj_commute" "(P&Q) = (Q&P)"; | 
| 194 | prove "conj_left_commute" "(P&(Q&R)) = (Q&(P&R))"; | |
| 195 | val conj_comms = [conj_commute, conj_left_commute]; | |
| 2134 | 196 | prove "conj_assoc" "((P&Q)&R) = (P&(Q&R))"; | 
| 923 | 197 | |
| 1922 | 198 | prove "disj_commute" "(P|Q) = (Q|P)"; | 
| 199 | prove "disj_left_commute" "(P|(Q|R)) = (Q|(P|R))"; | |
| 200 | val disj_comms = [disj_commute, disj_left_commute]; | |
| 2134 | 201 | prove "disj_assoc" "((P|Q)|R) = (P|(Q|R))"; | 
| 1922 | 202 | |
| 923 | 203 | prove "conj_disj_distribL" "(P&(Q|R)) = (P&Q | P&R)"; | 
| 204 | prove "conj_disj_distribR" "((P|Q)&R) = (P&R | Q&R)"; | |
| 1485 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 nipkow parents: 
1465diff
changeset | 205 | |
| 1892 | 206 | prove "disj_conj_distribL" "(P|(Q&R)) = ((P|Q) & (P|R))"; | 
| 207 | prove "disj_conj_distribR" "((P&Q)|R) = ((P|R) & (Q|R))"; | |
| 208 | ||
| 2134 | 209 | prove "imp_conjR" "(P --> (Q&R)) = ((P-->Q) & (P-->R))"; | 
| 210 | prove "imp_conjL" "((P&Q) -->R) = (P --> (Q --> R))"; | |
| 211 | prove "imp_disjL" "((P|Q) --> R) = ((P-->R)&(Q-->R))"; | |
| 1892 | 212 | |
| 3448 | 213 | (*These two are specialized, but imp_disj_not1 is useful in Auth/Yahalom.ML*) | 
| 8114 | 214 | prove "imp_disj_not1" "(P --> Q | R) = (~Q --> P --> R)"; | 
| 215 | prove "imp_disj_not2" "(P --> Q | R) = (~R --> P --> Q)"; | |
| 3448 | 216 | |
| 3904 | 217 | prove "imp_disj1" "((P-->Q)|R) = (P--> Q|R)"; | 
| 218 | prove "imp_disj2" "(Q|(P-->R)) = (P--> Q|R)"; | |
| 219 | ||
| 1485 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 nipkow parents: 
1465diff
changeset | 220 | prove "de_Morgan_disj" "(~(P | Q)) = (~P & ~Q)"; | 
| 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 nipkow parents: 
1465diff
changeset | 221 | prove "de_Morgan_conj" "(~(P & Q)) = (~P | ~Q)"; | 
| 3446 
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
 paulson parents: 
3282diff
changeset | 222 | prove "not_imp" "(~(P --> Q)) = (P & ~Q)"; | 
| 1922 | 223 | prove "not_iff" "(P~=Q) = (P = (~Q))"; | 
| 4743 | 224 | prove "disj_not1" "(~P | Q) = (P --> Q)"; | 
| 225 | prove "disj_not2" "(P | ~Q) = (Q --> P)"; (* changes orientation :-( *) | |
| 5975 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 226 | prove "imp_conv_disj" "(P --> Q) = ((~P) | Q)"; | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 227 | |
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 228 | prove "iff_conv_conj_imp" "(P = Q) = ((P --> Q) & (Q --> P))"; | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 229 | |
| 1485 
240cc98b94a7
Added qed_spec_mp to avoid renaming of bound vars in 'th RS spec'
 nipkow parents: 
1465diff
changeset | 230 | |
| 4830 | 231 | (*Avoids duplication of subgoals after split_if, when the true and false | 
| 2134 | 232 | cases boil down to the same thing.*) | 
| 233 | prove "cases_simp" "((P --> Q) & (~P --> Q)) = Q"; | |
| 234 | ||
| 3842 | 235 | prove "not_all" "(~ (! x. P(x))) = (? x.~P(x))"; | 
| 1922 | 236 | prove "imp_all" "((! x. P x) --> Q) = (? x. P x --> Q)"; | 
| 3842 | 237 | prove "not_ex" "(~ (? x. P(x))) = (! x.~P(x))"; | 
| 1922 | 238 | prove "imp_ex" "((? x. P x) --> Q) = (! x. P x --> Q)"; | 
| 1660 | 239 | |
| 1655 | 240 | prove "ex_disj_distrib" "(? x. P(x) | Q(x)) = ((? x. P(x)) | (? x. Q(x)))"; | 
| 241 | prove "all_conj_distrib" "(!x. P(x) & Q(x)) = ((! x. P(x)) & (! x. Q(x)))"; | |
| 242 | ||
| 2134 | 243 | (* '&' congruence rule: not included by default! | 
| 244 | May slow rewrite proofs down by as much as 50% *) | |
| 245 | ||
| 7357 | 246 | let val th = prove_goal (the_context ()) | 
| 2134 | 247 | "(P=P')--> (P'--> (Q=Q'))--> ((P&Q) = (P'&Q'))" | 
| 7031 | 248 | (fn _=> [(Blast_tac 1)]) | 
| 2134 | 249 | in  bind_thm("conj_cong",standard (impI RSN (2, th RS mp RS mp)))  end;
 | 
| 250 | ||
| 7357 | 251 | let val th = prove_goal (the_context ()) | 
| 2134 | 252 | "(Q=Q')--> (Q'--> (P=P'))--> ((P&Q) = (P'&Q'))" | 
| 7031 | 253 | (fn _=> [(Blast_tac 1)]) | 
| 2134 | 254 | in  bind_thm("rev_conj_cong",standard (impI RSN (2, th RS mp RS mp)))  end;
 | 
| 255 | ||
| 256 | (* '|' congruence rule: not included by default! *) | |
| 257 | ||
| 7357 | 258 | let val th = prove_goal (the_context ()) | 
| 2134 | 259 | "(P=P')--> (~P'--> (Q=Q'))--> ((P|Q) = (P'|Q'))" | 
| 7031 | 260 | (fn _=> [(Blast_tac 1)]) | 
| 2134 | 261 | in  bind_thm("disj_cong",standard (impI RSN (2, th RS mp RS mp)))  end;
 | 
| 262 | ||
| 263 | prove "eq_sym_conv" "(x=y) = (y=x)"; | |
| 264 | ||
| 5278 | 265 | |
| 266 | (** if-then-else rules **) | |
| 267 | ||
| 7031 | 268 | Goalw [if_def] "(if True then x else y) = x"; | 
| 269 | by (Blast_tac 1); | |
| 270 | qed "if_True"; | |
| 2134 | 271 | |
| 7031 | 272 | Goalw [if_def] "(if False then x else y) = y"; | 
| 273 | by (Blast_tac 1); | |
| 274 | qed "if_False"; | |
| 2134 | 275 | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7031diff
changeset | 276 | Goalw [if_def] "P ==> (if P then x else y) = x"; | 
| 7031 | 277 | by (Blast_tac 1); | 
| 278 | qed "if_P"; | |
| 5304 | 279 | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7031diff
changeset | 280 | Goalw [if_def] "~P ==> (if P then x else y) = y"; | 
| 7031 | 281 | by (Blast_tac 1); | 
| 282 | qed "if_not_P"; | |
| 2134 | 283 | |
| 7031 | 284 | Goal "P(if Q then x else y) = ((Q --> P(x)) & (~Q --> P(y)))"; | 
| 285 | by (res_inst_tac [("Q","Q")] (excluded_middle RS disjE) 1);
 | |
| 286 | by (stac if_P 2); | |
| 287 | by (stac if_not_P 1); | |
| 288 | by (ALLGOALS (Blast_tac)); | |
| 289 | qed "split_if"; | |
| 290 | ||
| 4830 | 291 | (* for backwards compatibility: *) | 
| 292 | val expand_if = split_if; | |
| 4205 
96632970d203
simpdata.ML: renamed split_prem_tac to split_asm_tac, added split_if_asm
 oheimb parents: 
4189diff
changeset | 293 | |
| 7031 | 294 | Goal "P(if Q then x else y) = (~((Q & ~P x) | (~Q & ~P y)))"; | 
| 295 | by (stac split_if 1); | |
| 296 | by (Blast_tac 1); | |
| 297 | qed "split_if_asm"; | |
| 2134 | 298 | |
| 7031 | 299 | Goal "(if c then x else x) = x"; | 
| 300 | by (stac split_if 1); | |
| 301 | by (Blast_tac 1); | |
| 302 | qed "if_cancel"; | |
| 5304 | 303 | |
| 7031 | 304 | Goal "(if x = y then y else x) = x"; | 
| 305 | by (stac split_if 1); | |
| 306 | by (Blast_tac 1); | |
| 307 | qed "if_eq_cancel"; | |
| 5304 | 308 | |
| 4769 
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
 paulson parents: 
4744diff
changeset | 309 | (*This form is useful for expanding IFs on the RIGHT of the ==> symbol*) | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7031diff
changeset | 310 | Goal "(if P then Q else R) = ((P-->Q) & (~P-->R))"; | 
| 7031 | 311 | by (rtac split_if 1); | 
| 312 | qed "if_bool_eq_conj"; | |
| 4769 
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
 paulson parents: 
4744diff
changeset | 313 | |
| 
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
 paulson parents: 
4744diff
changeset | 314 | (*And this form is useful for expanding IFs on the LEFT*) | 
| 7031 | 315 | Goal "(if P then Q else R) = ((P&Q) | (~P&R))"; | 
| 316 | by (stac split_if 1); | |
| 317 | by (Blast_tac 1); | |
| 318 | qed "if_bool_eq_disj"; | |
| 2134 | 319 | |
| 4351 | 320 | |
| 321 | (*** make simplification procedures for quantifier elimination ***) | |
| 322 | ||
| 323 | structure Quantifier1 = Quantifier1Fun( | |
| 324 | struct | |
| 325 | (*abstract syntax*) | |
| 326 |   fun dest_eq((c as Const("op =",_)) $ s $ t) = Some(c,s,t)
 | |
| 327 | | dest_eq _ = None; | |
| 328 |   fun dest_conj((c as Const("op &",_)) $ s $ t) = Some(c,s,t)
 | |
| 329 | | dest_conj _ = None; | |
| 330 | val conj = HOLogic.conj | |
| 331 | val imp = HOLogic.imp | |
| 332 | (*rules*) | |
| 333 | val iff_reflection = eq_reflection | |
| 334 | val iffI = iffI | |
| 335 | val sym = sym | |
| 336 | val conjI= conjI | |
| 337 | val conjE= conjE | |
| 338 | val impI = impI | |
| 339 | val impE = impE | |
| 340 | val mp = mp | |
| 341 | val exI = exI | |
| 342 | val exE = exE | |
| 343 | val allI = allI | |
| 344 | val allE = allE | |
| 345 | end); | |
| 346 | ||
| 4320 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 nipkow parents: 
4205diff
changeset | 347 | local | 
| 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 nipkow parents: 
4205diff
changeset | 348 | val ex_pattern = | 
| 7357 | 349 |   Thm.read_cterm (Theory.sign_of (the_context ())) ("EX x. P(x) & Q(x)",HOLogic.boolT)
 | 
| 3913 | 350 | |
| 4320 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 nipkow parents: 
4205diff
changeset | 351 | val all_pattern = | 
| 7357 | 352 |   Thm.read_cterm (Theory.sign_of (the_context ())) ("ALL x. P(x) & P'(x) --> Q(x)",HOLogic.boolT)
 | 
| 4320 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 nipkow parents: 
4205diff
changeset | 353 | |
| 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 nipkow parents: 
4205diff
changeset | 354 | in | 
| 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 nipkow parents: 
4205diff
changeset | 355 | val defEX_regroup = | 
| 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 nipkow parents: 
4205diff
changeset | 356 | mk_simproc "defined EX" [ex_pattern] Quantifier1.rearrange_ex; | 
| 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 nipkow parents: 
4205diff
changeset | 357 | val defALL_regroup = | 
| 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 nipkow parents: 
4205diff
changeset | 358 | mk_simproc "defined ALL" [all_pattern] Quantifier1.rearrange_all; | 
| 
24d9e6639cd4
Moved the quantifier elimination simp procs into Provers.
 nipkow parents: 
4205diff
changeset | 359 | end; | 
| 3913 | 360 | |
| 4351 | 361 | |
| 362 | (*** Case splitting ***) | |
| 3913 | 363 | |
| 5304 | 364 | structure SplitterData = | 
| 365 | struct | |
| 366 | structure Simplifier = Simplifier | |
| 5552 | 367 | val mk_eq = mk_eq | 
| 5304 | 368 | val meta_eq_to_iff = meta_eq_to_obj_eq | 
| 369 | val iffD = iffD2 | |
| 370 | val disjE = disjE | |
| 371 | val conjE = conjE | |
| 372 | val exE = exE | |
| 373 | val contrapos = contrapos | |
| 374 | val contrapos2 = contrapos2 | |
| 375 | val notnotD = notnotD | |
| 376 | end; | |
| 4681 | 377 | |
| 5304 | 378 | structure Splitter = SplitterFun(SplitterData); | 
| 2263 | 379 | |
| 5304 | 380 | val split_tac = Splitter.split_tac; | 
| 381 | val split_inside_tac = Splitter.split_inside_tac; | |
| 382 | val split_asm_tac = Splitter.split_asm_tac; | |
| 5307 | 383 | val op addsplits = Splitter.addsplits; | 
| 384 | val op delsplits = Splitter.delsplits; | |
| 5304 | 385 | val Addsplits = Splitter.Addsplits; | 
| 386 | val Delsplits = Splitter.Delsplits; | |
| 4718 
fc2ba9fb2135
new rewrite rules not1_or, not2_or, and if_eq_cancel
 oheimb parents: 
4681diff
changeset | 387 | |
| 2134 | 388 | (*In general it seems wrong to add distributive laws by default: they | 
| 389 | might cause exponential blow-up. But imp_disjL has been in for a while | |
| 390 | and cannot be removed without affecting existing proofs. Moreover, | |
| 391 | rewriting by "(P|Q --> R) = ((P-->R)&(Q-->R))" might be justified on the | |
| 392 | grounds that it allows simplification of R in the two cases.*) | |
| 393 | ||
| 5304 | 394 | fun gen_all th = forall_elim_vars (#maxidx(rep_thm th)+1) th; | 
| 395 | ||
| 2134 | 396 | val mksimps_pairs = | 
| 397 |   [("op -->", [mp]), ("op &", [conjunct1,conjunct2]),
 | |
| 398 |    ("All", [spec]), ("True", []), ("False", []),
 | |
| 4769 
bb60149fe21b
changed if_bool_eq to if_bool_eq_conj and added if_bool_eq_disj
 paulson parents: 
4744diff
changeset | 399 |    ("If", [if_bool_eq_conj RS iffD1])];
 | 
| 1758 | 400 | |
| 5552 | 401 | (* ###FIXME: move to Provers/simplifier.ML | 
| 5304 | 402 | val mk_atomize: (string * thm list) list -> thm -> thm list | 
| 403 | *) | |
| 5552 | 404 | (* ###FIXME: move to Provers/simplifier.ML *) | 
| 5304 | 405 | fun mk_atomize pairs = | 
| 406 | let fun atoms th = | |
| 407 | (case concl_of th of | |
| 408 |            Const("Trueprop",_) $ p =>
 | |
| 409 | (case head_of p of | |
| 410 | Const(a,_) => | |
| 411 | (case assoc(pairs,a) of | |
| 412 | Some(rls) => flat (map atoms ([th] RL rls)) | |
| 413 | | None => [th]) | |
| 414 | | _ => [th]) | |
| 415 | | _ => [th]) | |
| 416 | in atoms end; | |
| 417 | ||
| 5552 | 418 | fun mksimps pairs = (map mk_eq o mk_atomize pairs o gen_all); | 
| 5304 | 419 | |
| 7570 | 420 | fun unsafe_solver_tac prems = | 
| 421 | FIRST'[resolve_tac(reflexive_thm::TrueI::refl::prems), atac, etac FalseE]; | |
| 422 | val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac; | |
| 423 | ||
| 2636 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 oheimb parents: 
2595diff
changeset | 424 | (*No premature instantiation of variables during simplification*) | 
| 7570 | 425 | fun safe_solver_tac prems = | 
| 426 | FIRST'[match_tac(reflexive_thm::TrueI::refl::prems), | |
| 427 | eq_assume_tac, ematch_tac [FalseE]]; | |
| 428 | val safe_solver = mk_solver "HOL safe" safe_solver_tac; | |
| 2443 
a81d4c219c3c
factored out HOL_base_ss and val HOL_min_ss, added HOL_safe_min_ss
 oheimb parents: 
2263diff
changeset | 429 | |
| 2636 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 oheimb parents: 
2595diff
changeset | 430 | val HOL_basic_ss = empty_ss setsubgoaler asm_simp_tac | 
| 7570 | 431 | setSSolver safe_solver | 
| 2636 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 oheimb parents: 
2595diff
changeset | 432 | setSolver unsafe_solver | 
| 4677 | 433 | setmksimps (mksimps mksimps_pairs) | 
| 5552 | 434 | setmkeqTrue mk_eq_True; | 
| 2443 
a81d4c219c3c
factored out HOL_base_ss and val HOL_min_ss, added HOL_safe_min_ss
 oheimb parents: 
2263diff
changeset | 435 | |
| 3446 
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
 paulson parents: 
3282diff
changeset | 436 | val HOL_ss = | 
| 
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
 paulson parents: 
3282diff
changeset | 437 | HOL_basic_ss addsimps | 
| 
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
 paulson parents: 
3282diff
changeset | 438 | ([triv_forall_equality, (* prunes params *) | 
| 3654 | 439 | True_implies_equals, (* prune asms `True' *) | 
| 4718 
fc2ba9fb2135
new rewrite rules not1_or, not2_or, and if_eq_cancel
 oheimb parents: 
4681diff
changeset | 440 | if_True, if_False, if_cancel, if_eq_cancel, | 
| 5304 | 441 | imp_disjL, conj_assoc, disj_assoc, | 
| 3904 | 442 | de_Morgan_conj, de_Morgan_disj, imp_disj1, imp_disj2, not_imp, | 
| 5447 
df03d330aeab
Proved and added rewrite rule (@x. x=y) = y to simpset.
 nipkow parents: 
5307diff
changeset | 443 | disj_not1, not_all, not_ex, cases_simp, Eps_eq, Eps_sym_eq] | 
| 3446 
a14e5451f613
Addition of not_imp (which pushes negation into implication) as a default
 paulson parents: 
3282diff
changeset | 444 | @ ex_simps @ all_simps @ simp_thms) | 
| 4032 
4b1c69d8b767
For each datatype `t' there is now a theorem `split_t_case' of the form
 nipkow parents: 
3919diff
changeset | 445 | addsimprocs [defALL_regroup,defEX_regroup] | 
| 4744 
4469d498cd48
moved addsplits [expand_if] from HOL_basic_ss to HOL_ss;
 wenzelm parents: 
4743diff
changeset | 446 | addcongs [imp_cong] | 
| 4830 | 447 | addsplits [split_if]; | 
| 2082 | 448 | |
| 6293 | 449 | (*Simplifies x assuming c and y assuming ~c*) | 
| 450 | val prems = Goalw [if_def] | |
| 451 | "[| b=c; c ==> x=u; ~c ==> y=v |] ==> \ | |
| 452 | \ (if b then x else y) = (if c then u else v)"; | |
| 453 | by (asm_simp_tac (HOL_ss addsimps prems) 1); | |
| 454 | qed "if_cong"; | |
| 455 | ||
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7031diff
changeset | 456 | (*Prevents simplification of x and y: faster and allows the execution | 
| 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7031diff
changeset | 457 | of functional programs. NOW THE DEFAULT.*) | 
| 7031 | 458 | Goal "b=c ==> (if b then x else y) = (if c then x else y)"; | 
| 459 | by (etac arg_cong 1); | |
| 460 | qed "if_weak_cong"; | |
| 6293 | 461 | |
| 462 | (*Prevents simplification of t: much faster*) | |
| 7031 | 463 | Goal "a = b ==> (let x=a in t(x)) = (let x=b in t(x))"; | 
| 464 | by (etac arg_cong 1); | |
| 465 | qed "let_weak_cong"; | |
| 6293 | 466 | |
| 7031 | 467 | Goal "f(if c then x else y) = (if c then f x else f y)"; | 
| 468 | by (simp_tac (HOL_ss setloop (split_tac [split_if])) 1); | |
| 469 | qed "if_distrib"; | |
| 1655 | 470 | |
| 4327 | 471 | (*For expand_case_tac*) | 
| 7584 | 472 | val prems = Goal "[| P ==> Q(True); ~P ==> Q(False) |] ==> Q(P)"; | 
| 2948 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 paulson parents: 
2935diff
changeset | 473 | by (case_tac "P" 1); | 
| 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 paulson parents: 
2935diff
changeset | 474 | by (ALLGOALS (asm_simp_tac (HOL_ss addsimps prems))); | 
| 7584 | 475 | qed "expand_case"; | 
| 2948 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 paulson parents: 
2935diff
changeset | 476 | |
| 4327 | 477 | (*Used in Auth proofs. Typically P contains Vars that become instantiated | 
| 478 | during unification.*) | |
| 2948 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 paulson parents: 
2935diff
changeset | 479 | fun expand_case_tac P i = | 
| 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 paulson parents: 
2935diff
changeset | 480 |     res_inst_tac [("P",P)] expand_case i THEN
 | 
| 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 paulson parents: 
2935diff
changeset | 481 | Simp_tac (i+1) THEN | 
| 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 paulson parents: 
2935diff
changeset | 482 | Simp_tac i; | 
| 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 paulson parents: 
2935diff
changeset | 483 | |
| 7584 | 484 | (*This lemma restricts the effect of the rewrite rule u=v to the left-hand | 
| 485 |   side of an equality.  Used in {Integ,Real}/simproc.ML*)
 | |
| 486 | Goal "x=y ==> (x=z) = (y=z)"; | |
| 487 | by (asm_simp_tac HOL_ss 1); | |
| 488 | qed "restrict_to_left"; | |
| 2948 
f18035b1d531
Moved expand_case_tac from Auth/Message.ML to simpdata.ML
 paulson parents: 
2935diff
changeset | 489 | |
| 7357 | 490 | (* default simpset *) | 
| 7584 | 491 | val simpsetup = | 
| 492 | [fn thy => (simpset_ref_of thy := HOL_ss addcongs [if_weak_cong]; | |
| 493 | thy)]; | |
| 3615 
e5322197cfea
Moved some functions which used to be part of thy_data.ML
 berghofe parents: 
3577diff
changeset | 494 | |
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: 
4651diff
changeset | 495 | |
| 5219 | 496 | (*** integration of simplifier with classical reasoner ***) | 
| 2636 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 oheimb parents: 
2595diff
changeset | 497 | |
| 5219 | 498 | structure Clasimp = ClasimpFun | 
| 8473 | 499 | (structure Simplifier = Simplifier and Splitter = Splitter | 
| 500 | and Classical = Classical and Blast = Blast); | |
| 4652 
d24cca140eeb
factored out common code of HOL/simpdata.ML and FOL/simpdata.ML concerning
 oheimb parents: 
4651diff
changeset | 501 | open Clasimp; | 
| 2636 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 oheimb parents: 
2595diff
changeset | 502 | |
| 
4b30dbe4a020
added delcongs, Delcongs, unsafe_solver, safe_solver, HOL_basic_ss,
 oheimb parents: 
2595diff
changeset | 503 | val HOL_css = (HOL_cs, HOL_ss); | 
| 5975 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 504 | |
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 505 | |
| 8641 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 506 | (* "iff" attribute *) | 
| 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 507 | |
| 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 508 | val iff_add_global = Clasimp.change_global_css (op addIffs); | 
| 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 509 | val iff_add_local = Clasimp.change_local_css (op addIffs); | 
| 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 510 | |
| 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 511 | val iff_attrib_setup = | 
| 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 512 |   [Attrib.add_attributes [("iff", (Attrib.no_args iff_add_global, Attrib.no_args iff_add_local),
 | 
| 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 513 | "add rules to simpset and claset simultaneously")]]; | 
| 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 514 | |
| 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 515 | |
| 
978db2870862
change_global/local_css move to Provers/clasimp.ML;
 wenzelm parents: 
8473diff
changeset | 516 | |
| 5975 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 517 | (*** A general refutation procedure ***) | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 518 | |
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 519 | (* Parameters: | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 520 | |
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 521 | test: term -> bool | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 522 | tests if a term is at all relevant to the refutation proof; | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 523 | if not, then it can be discarded. Can improve performance, | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 524 | esp. if disjunctions can be discarded (no case distinction needed!). | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 525 | |
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 526 | prep_tac: int -> tactic | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 527 | A preparation tactic to be applied to the goal once all relevant premises | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 528 | have been moved to the conclusion. | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 529 | |
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 530 | ref_tac: int -> tactic | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 531 | the actual refutation tactic. Should be able to deal with goals | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 532 | [| A1; ...; An |] ==> False | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 533 | where the Ai are atomic, i.e. no top-level &, | or ? | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 534 | *) | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 535 | |
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 536 | fun refute_tac test prep_tac ref_tac = | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 537 | let val nnf_simps = | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 538 | [imp_conv_disj,iff_conv_conj_imp,de_Morgan_disj,de_Morgan_conj, | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 539 | not_all,not_ex,not_not]; | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 540 | val nnf_simpset = | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 541 | empty_ss setmkeqTrue mk_eq_True | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 542 | setmksimps (mksimps mksimps_pairs) | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 543 | addsimps nnf_simps; | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 544 | val prem_nnf_tac = full_simp_tac nnf_simpset; | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 545 | |
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 546 | val refute_prems_tac = | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 547 | REPEAT(eresolve_tac [conjE, exE] 1 ORELSE | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 548 | filter_prems_tac test 1 ORELSE | 
| 6301 | 549 | etac disjE 1) THEN | 
| 5975 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 550 | ref_tac 1; | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 551 | in EVERY'[TRY o filter_prems_tac test, | 
| 6128 | 552 | DETERM o REPEAT o etac rev_mp, prep_tac, rtac ccontr, prem_nnf_tac, | 
| 5975 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 553 | SELECT_GOAL (DEPTH_SOLVE refute_prems_tac)] | 
| 
cd19eaa90f45
Added a general refutation tactic which works by putting things into nnf first.
 nipkow parents: 
5552diff
changeset | 554 | end; |