| author | wenzelm | 
| Thu, 14 Dec 2000 19:37:43 +0100 | |
| changeset 10674 | 2cc6415c1801 | 
| parent 9713 | 2c5b42311eb0 | 
| child 12720 | f8a134b9a57f | 
| permissions | -rw-r--r-- | 
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 1 | (* Title: Sequents/simpdata.ML | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 4 | Copyright 1999 University of Cambridge | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 5 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 6 | Instantiation of the generic simplifier for LK | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 7 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 8 | Borrows from the DC simplifier of Soren Heilmann. | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 9 | *) | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 10 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 11 | (*** Rewrite rules ***) | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 12 | |
| 9713 | 13 | fun prove_fun s = | 
| 14 | (writeln s; | |
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 15 | prove_goal LK.thy s | 
| 9713 | 16 | (fn prems => [ (cut_facts_tac prems 1), | 
| 7123 | 17 | (fast_tac (pack() add_safes [subst]) 1) ])); | 
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 18 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 19 | val conj_simps = map prove_fun | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 20 | ["|- P & True <-> P", "|- True & P <-> P", | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 21 | "|- P & False <-> False", "|- False & P <-> False", | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 22 | "|- P & P <-> P", " |- P & P & Q <-> P & Q", | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 23 | "|- P & ~P <-> False", "|- ~P & P <-> False", | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 24 | "|- (P & Q) & R <-> P & (Q & R)"]; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 25 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 26 | val disj_simps = map prove_fun | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 27 | ["|- P | True <-> True", "|- True | P <-> True", | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 28 | "|- P | False <-> P", "|- False | P <-> P", | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 29 | "|- P | P <-> P", "|- P | P | Q <-> P | Q", | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 30 | "|- (P | Q) | R <-> P | (Q | R)"]; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 31 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 32 | val not_simps = map prove_fun | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 33 | ["|- ~ False <-> True", "|- ~ True <-> False"]; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 34 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 35 | val imp_simps = map prove_fun | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 36 | ["|- (P --> False) <-> ~P", "|- (P --> True) <-> True", | 
| 9713 | 37 | "|- (False --> P) <-> True", "|- (True --> P) <-> P", | 
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 38 | "|- (P --> P) <-> True", "|- (P --> ~P) <-> ~P"]; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 39 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 40 | val iff_simps = map prove_fun | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 41 | ["|- (True <-> P) <-> P", "|- (P <-> True) <-> P", | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 42 | "|- (P <-> P) <-> True", | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 43 | "|- (False <-> P) <-> ~P", "|- (P <-> False) <-> ~P"]; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 44 | |
| 7123 | 45 | |
| 46 | val quant_simps = map prove_fun | |
| 9713 | 47 | ["|- (ALL x. P) <-> P", | 
| 7123 | 48 | "|- (ALL x. x=t --> P(x)) <-> P(t)", | 
| 49 | "|- (ALL x. t=x --> P(x)) <-> P(t)", | |
| 50 | "|- (EX x. P) <-> P", | |
| 9713 | 51 | "|- (EX x. x=t & P(x)) <-> P(t)", | 
| 7123 | 52 | "|- (EX x. t=x & P(x)) <-> P(t)"]; | 
| 53 | ||
| 54 | (*** Miniscoping: pushing quantifiers in | |
| 55 | We do NOT distribute of ALL over &, or dually that of EX over | | |
| 9713 | 56 | Baaz and Leitsch, On Skolemization and Proof Complexity (1994) | 
| 7123 | 57 | show that this step can increase proof length! | 
| 58 | ***) | |
| 59 | ||
| 60 | (*existential miniscoping*) | |
| 9713 | 61 | val ex_simps = map prove_fun | 
| 7123 | 62 | ["|- (EX x. P(x) & Q) <-> (EX x. P(x)) & Q", | 
| 9713 | 63 | "|- (EX x. P & Q(x)) <-> P & (EX x. Q(x))", | 
| 64 | "|- (EX x. P(x) | Q) <-> (EX x. P(x)) | Q", | |
| 65 | "|- (EX x. P | Q(x)) <-> P | (EX x. Q(x))", | |
| 66 | "|- (EX x. P(x) --> Q) <-> (ALL x. P(x)) --> Q", | |
| 67 | "|- (EX x. P --> Q(x)) <-> P --> (EX x. Q(x))"]; | |
| 7123 | 68 | |
| 69 | (*universal miniscoping*) | |
| 70 | val all_simps = map prove_fun | |
| 71 | ["|- (ALL x. P(x) & Q) <-> (ALL x. P(x)) & Q", | |
| 9713 | 72 | "|- (ALL x. P & Q(x)) <-> P & (ALL x. Q(x))", | 
| 73 | "|- (ALL x. P(x) --> Q) <-> (EX x. P(x)) --> Q", | |
| 74 | "|- (ALL x. P --> Q(x)) <-> P --> (ALL x. Q(x))", | |
| 75 | "|- (ALL x. P(x) | Q) <-> (ALL x. P(x)) | Q", | |
| 76 | "|- (ALL x. P | Q(x)) <-> P | (ALL x. Q(x))"]; | |
| 7123 | 77 | |
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 78 | (*These are NOT supplied by default!*) | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 79 | val distrib_simps = map prove_fun | 
| 9713 | 80 | ["|- P & (Q | R) <-> P&Q | P&R", | 
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 81 | "|- (Q | R) & P <-> Q&P | R&P", | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 82 | "|- (P | Q --> R) <-> (P --> R) & (Q --> R)"]; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 83 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 84 | (** Conversion into rewrite rules **) | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 85 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 86 | fun gen_all th = forall_elim_vars (#maxidx(rep_thm th)+1) th; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 87 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 88 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 89 | (*Make atomic rewrite rules*) | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 90 | fun atomize r = | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 91 | case concl_of r of | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 92 |    Const("Trueprop",_) $ Abs(_,_,a) $ Abs(_,_,c) =>
 | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 93 | (case (forms_of_seq a, forms_of_seq c) of | 
| 9713 | 94 | ([], [p]) => | 
| 95 | (case p of | |
| 96 |                Const("op -->",_)$_$_ => atomize(r RS mp_R)
 | |
| 97 |              | Const("op &",_)$_$_   => atomize(r RS conjunct1) @
 | |
| 98 | atomize(r RS conjunct2) | |
| 99 |              | Const("All",_)$_      => atomize(r RS spec)
 | |
| 100 |              | Const("True",_)       => []    (*True is DELETED*)
 | |
| 101 |              | Const("False",_)      => []    (*should False do something?*)
 | |
| 102 | | _ => [r]) | |
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 103 | | _ => []) (*ignore theorem unless it has precisely one conclusion*) | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 104 | | _ => [r]; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 105 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 106 | |
| 9259 | 107 | Goal "|- ~P ==> |- (P <-> False)"; | 
| 108 | by (etac (thinR RS cut) 1); | |
| 9713 | 109 | by (Fast_tac 1); | 
| 9259 | 110 | qed "P_iff_F"; | 
| 111 | ||
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 112 | val iff_reflection_F = P_iff_F RS iff_reflection; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 113 | |
| 9259 | 114 | Goal "|- P ==> |- (P <-> True)"; | 
| 115 | by (etac (thinR RS cut) 1); | |
| 9713 | 116 | by (Fast_tac 1); | 
| 9259 | 117 | qed "P_iff_T"; | 
| 118 | ||
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 119 | val iff_reflection_T = P_iff_T RS iff_reflection; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 120 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 121 | (*Make meta-equalities.*) | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 122 | fun mk_meta_eq th = case concl_of th of | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 123 |     Const("==",_)$_$_           => th
 | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 124 |   | Const("Trueprop",_) $ Abs(_,_,a) $ Abs(_,_,c) =>
 | 
| 9713 | 125 | (case (forms_of_seq a, forms_of_seq c) of | 
| 126 | ([], [p]) => | |
| 127 | (case p of | |
| 128 |                       (Const("op =",_)$_$_)   => th RS eq_reflection
 | |
| 129 |                     | (Const("op <->",_)$_$_) => th RS iff_reflection
 | |
| 130 |                     | (Const("Not",_)$_)      => th RS iff_reflection_F
 | |
| 131 | | _ => th RS iff_reflection_T) | |
| 132 |            | _ => error ("addsimps: unable to use theorem\n" ^
 | |
| 133 | string_of_thm th)); | |
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 134 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 135 | |
| 7123 | 136 | (*Replace premises x=y, X<->Y by X==Y*) | 
| 9713 | 137 | val mk_meta_prems = | 
| 138 | rule_by_tactic | |
| 7123 | 139 | (REPEAT_FIRST (resolve_tac [meta_eq_to_obj_eq, def_imp_iff])); | 
| 140 | ||
| 9713 | 141 | (*Congruence rules for = or <-> (instead of ==)*) | 
| 7123 | 142 | fun mk_meta_cong rl = | 
| 143 | standard(mk_meta_eq (mk_meta_prems rl)) | |
| 144 | handle THM _ => | |
| 145 |   error("Premises and conclusion of congruence rules must use =-equality or <->");
 | |
| 146 | ||
| 147 | ||
| 148 | (*** Named rewrite rules ***) | |
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 149 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 150 | fun prove nm thm = qed_goal nm LK.thy thm | 
| 9713 | 151 | (fn prems => [ (cut_facts_tac prems 1), | 
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 152 | (fast_tac LK_pack 1) ]); | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 153 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 154 | prove "conj_commute" "|- P&Q <-> Q&P"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 155 | prove "conj_left_commute" "|- P&(Q&R) <-> Q&(P&R)"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 156 | val conj_comms = [conj_commute, conj_left_commute]; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 157 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 158 | prove "disj_commute" "|- P|Q <-> Q|P"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 159 | prove "disj_left_commute" "|- P|(Q|R) <-> Q|(P|R)"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 160 | val disj_comms = [disj_commute, disj_left_commute]; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 161 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 162 | prove "conj_disj_distribL" "|- P&(Q|R) <-> (P&Q | P&R)"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 163 | prove "conj_disj_distribR" "|- (P|Q)&R <-> (P&R | Q&R)"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 164 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 165 | prove "disj_conj_distribL" "|- P|(Q&R) <-> (P|Q) & (P|R)"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 166 | prove "disj_conj_distribR" "|- (P&Q)|R <-> (P|R) & (Q|R)"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 167 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 168 | prove "imp_conj_distrib" "|- (P --> (Q&R)) <-> (P-->Q) & (P-->R)"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 169 | prove "imp_conj" "|- ((P&Q)-->R) <-> (P --> (Q --> R))"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 170 | prove "imp_disj" "|- (P|Q --> R) <-> (P-->R) & (Q-->R)"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 171 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 172 | prove "imp_disj1" "|- (P-->Q) | R <-> (P-->Q | R)"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 173 | prove "imp_disj2" "|- Q | (P-->R) <-> (P-->Q | R)"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 174 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 175 | prove "de_Morgan_disj" "|- (~(P | Q)) <-> (~P & ~Q)"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 176 | prove "de_Morgan_conj" "|- (~(P & Q)) <-> (~P | ~Q)"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 177 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 178 | prove "not_iff" "|- ~(P <-> Q) <-> (P <-> ~Q)"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 179 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 180 | |
| 9713 | 181 | val [p1,p2] = Goal | 
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 182 | "[| |- P <-> P'; |- P' ==> |- Q <-> Q' |] ==> |- (P-->Q) <-> (P'-->Q')"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 183 | by (lemma_tac p1 1); | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 184 | by (Safe_tac 1); | 
| 9713 | 185 | by (REPEAT (rtac cut 1 | 
| 186 | THEN | |
| 187 | DEPTH_SOLVE_1 (resolve_tac [thinL, thinR, p2 COMP monotonic] 1) | |
| 188 | THEN | |
| 189 | Safe_tac 1)); | |
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 190 | qed "imp_cong"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 191 | |
| 9713 | 192 | val [p1,p2] = Goal | 
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 193 | "[| |- P <-> P'; |- P' ==> |- Q <-> Q' |] ==> |- (P&Q) <-> (P'&Q')"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 194 | by (lemma_tac p1 1); | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 195 | by (Safe_tac 1); | 
| 9713 | 196 | by (REPEAT (rtac cut 1 | 
| 197 | THEN | |
| 198 | DEPTH_SOLVE_1 (resolve_tac [thinL, thinR, p2 COMP monotonic] 1) | |
| 199 | THEN | |
| 200 | Safe_tac 1)); | |
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 201 | qed "conj_cong"; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 202 | |
| 7123 | 203 | Goal "|- (x=y) <-> (y=x)"; | 
| 204 | by (fast_tac (pack() add_safes [subst]) 1); | |
| 205 | qed "eq_sym_conv"; | |
| 206 | ||
| 207 | ||
| 208 | (** if-then-else rules **) | |
| 209 | ||
| 210 | Goalw [If_def] "|- (if True then x else y) = x"; | |
| 211 | by (Fast_tac 1); | |
| 212 | qed "if_True"; | |
| 213 | ||
| 214 | Goalw [If_def] "|- (if False then x else y) = y"; | |
| 215 | by (Fast_tac 1); | |
| 216 | qed "if_False"; | |
| 217 | ||
| 218 | Goalw [If_def] "|- P ==> |- (if P then x else y) = x"; | |
| 219 | by (etac (thinR RS cut) 1); | |
| 220 | by (Fast_tac 1); | |
| 221 | qed "if_P"; | |
| 222 | ||
| 223 | Goalw [If_def] "|- ~P ==> |- (if P then x else y) = y"; | |
| 224 | by (etac (thinR RS cut) 1); | |
| 225 | by (Fast_tac 1); | |
| 226 | qed "if_not_P"; | |
| 227 | ||
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 228 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 229 | open Simplifier; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 230 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 231 | (*** Standard simpsets ***) | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 232 | |
| 7123 | 233 | val triv_rls = [FalseL, TrueR, basic, refl, iff_refl, reflexive_thm]; | 
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 234 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 235 | fun unsafe_solver prems = FIRST'[resolve_tac (triv_rls@prems), | 
| 9713 | 236 | assume_tac]; | 
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 237 | (*No premature instantiation of variables during simplification*) | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 238 | fun safe_solver prems = FIRST'[fn i => DETERM (match_tac (triv_rls@prems) i), | 
| 9713 | 239 | eq_assume_tac]; | 
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 240 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 241 | (*No simprules, but basic infrastructure for simplification*) | 
| 9713 | 242 | val LK_basic_ss = | 
| 243 | empty_ss setsubgoaler asm_simp_tac | |
| 244 | setSSolver (mk_solver "safe" safe_solver) | |
| 245 | setSolver (mk_solver "unsafe" unsafe_solver) | |
| 246 | setmksimps (map mk_meta_eq o atomize o gen_all) | |
| 247 | setmkcong mk_meta_cong; | |
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 248 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 249 | val LK_simps = | 
| 7123 | 250 | [triv_forall_equality, (* prunes params *) | 
| 9713 | 251 | refl RS P_iff_T] @ | 
| 252 | conj_simps @ disj_simps @ not_simps @ | |
| 7123 | 253 | imp_simps @ iff_simps @quant_simps @ all_simps @ ex_simps @ | 
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 254 | [de_Morgan_conj, de_Morgan_disj, imp_disj1, imp_disj2] @ | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 255 | map prove_fun | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 256 | ["|- P | ~P", "|- ~P | P", | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 257 | "|- ~ ~ P <-> P", "|- (~P --> P) <-> P", | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 258 | "|- (~P <-> ~Q) <-> (P<->Q)"]; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 259 | |
| 9713 | 260 | val LK_ss = | 
| 261 | LK_basic_ss addsimps LK_simps | |
| 262 | addeqcongs [left_cong] | |
| 263 | addcongs [imp_cong]; | |
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 264 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 265 | simpset_ref() := LK_ss; | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 266 | |
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 267 | |
| 7123 | 268 | (* To create substition rules *) | 
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 269 | |
| 7123 | 270 | qed_goal "eq_imp_subst" LK.thy "|- a=b ==> $H, A(a), $G |- $E, A(b), $F" | 
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 271 | (fn prems => | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 272 | [cut_facts_tac prems 1, | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 273 | asm_simp_tac LK_basic_ss 1]); | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 274 | |
| 7123 | 275 | Goal "|- P(if Q then x else y) <-> ((Q --> P(x)) & (~Q --> P(y)))"; | 
| 276 | by (res_inst_tac [ ("P","Q") ] cut 1);
 | |
| 277 | by (simp_tac (simpset() addsimps [if_P]) 2); | |
| 278 | by (res_inst_tac [ ("P","~Q") ] cut 1);
 | |
| 279 | by (simp_tac (simpset() addsimps [if_not_P]) 2); | |
| 280 | by (Fast_tac 1); | |
| 281 | qed "split_if"; | |
| 7098 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: diff
changeset | 282 | |
| 7123 | 283 | Goal "|- (if P then x else x) = x"; | 
| 284 | by (lemma_tac split_if 1); | |
| 285 | by (Fast_tac 1); | |
| 286 | qed "if_cancel"; | |
| 287 | ||
| 288 | Goal "|- (if x=y then y else x) = x"; | |
| 289 | by (lemma_tac split_if 1); | |
| 290 | by (Safe_tac 1); | |
| 291 | by (rtac symL 1); | |
| 292 | by (rtac basic 1); | |
| 293 | qed "if_eq_cancel"; | |
| 294 | ||
| 295 | (*Putting in automatic case splits seems to require a lot of work.*) |