| author | paulson | 
| Wed, 19 Jul 2000 10:59:59 +0200 | |
| changeset 9389 | 17c707841ad3 | 
| parent 6154 | 6a00a5baef2b | 
| child 11233 | 34c81a796ee3 | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/ex/prop-log.ML | 
| 0 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Tobias Nipkow & Lawrence C Paulson | 
| 0 | 4 | Copyright 1992 University of Cambridge | 
| 5 | ||
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5137diff
changeset | 6 | Inductive definition of propositional logic. | 
| 0 | 7 | Soundness and completeness w.r.t. truth-tables. | 
| 8 | ||
| 9 | Prove: If H|=p then G|=p where G:Fin(H) | |
| 10 | *) | |
| 11 | ||
| 6046 | 12 | Addsimps prop.intrs; | 
| 0 | 13 | |
| 14 | (*** Semantics of propositional logic ***) | |
| 15 | ||
| 16 | (** The function is_true **) | |
| 17 | ||
| 5068 | 18 | Goalw [is_true_def] "is_true(Fls,t) <-> False"; | 
| 5618 | 19 | by (Simp_tac 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 20 | qed "is_true_Fls"; | 
| 0 | 21 | |
| 5068 | 22 | Goalw [is_true_def] "is_true(#v,t) <-> v:t"; | 
| 5618 | 23 | by (Simp_tac 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 24 | qed "is_true_Var"; | 
| 0 | 25 | |
| 5618 | 26 | Goalw [is_true_def] "is_true(p=>q,t) <-> (is_true(p,t)-->is_true(q,t))"; | 
| 27 | by (Simp_tac 1); | |
| 760 | 28 | qed "is_true_Imp"; | 
| 0 | 29 | |
| 6046 | 30 | Addsimps [is_true_Fls, is_true_Var, is_true_Imp]; | 
| 0 | 31 | |
| 32 | ||
| 33 | (*** Proof theory of propositional logic ***) | |
| 34 | ||
| 5137 | 35 | Goalw thms.defs "G<=H ==> thms(G) <= thms(H)"; | 
| 0 | 36 | by (rtac lfp_mono 1); | 
| 515 | 37 | by (REPEAT (rtac thms.bnd_mono 1)); | 
| 0 | 38 | by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 39 | qed "thms_mono"; | 
| 0 | 40 | |
| 515 | 41 | val thms_in_pl = thms.dom_subset RS subsetD; | 
| 0 | 42 | |
| 6141 | 43 | val ImpE = prop.mk_cases "p=>q : prop"; | 
| 0 | 44 | |
| 515 | 45 | (*Stronger Modus Ponens rule: no typechecking!*) | 
| 5137 | 46 | Goal "[| H |- p=>q; H |- p |] ==> H |- q"; | 
| 515 | 47 | by (rtac thms.MP 1); | 
| 0 | 48 | by (REPEAT (eresolve_tac [asm_rl, thms_in_pl, thms_in_pl RS ImpE] 1)); | 
| 760 | 49 | qed "thms_MP"; | 
| 0 | 50 | |
| 51 | (*Rule is called I for Identity Combinator, not for Introduction*) | |
| 5137 | 52 | Goal "p:prop ==> H |- p=>p"; | 
| 515 | 53 | by (rtac (thms.S RS thms_MP RS thms_MP) 1); | 
| 54 | by (rtac thms.K 5); | |
| 55 | by (rtac thms.K 4); | |
| 56 | by (REPEAT (ares_tac prop.intrs 1)); | |
| 760 | 57 | qed "thms_I"; | 
| 0 | 58 | |
| 59 | (** Weakening, left and right **) | |
| 60 | ||
| 61 | (* [| G<=H; G|-p |] ==> H|-p Order of premises is convenient with RS*) | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 62 | bind_thm ("weaken_left", (thms_mono RS subsetD));
 | 
| 0 | 63 | |
| 64 | (* H |- p ==> cons(a,H) |- p *) | |
| 65 | val weaken_left_cons = subset_consI RS weaken_left; | |
| 66 | ||
| 67 | val weaken_left_Un1 = Un_upper1 RS weaken_left; | |
| 68 | val weaken_left_Un2 = Un_upper2 RS weaken_left; | |
| 69 | ||
| 5137 | 70 | Goal "[| H |- q; p:prop |] ==> H |- p=>q"; | 
| 515 | 71 | by (rtac (thms.K RS thms_MP) 1); | 
| 0 | 72 | by (REPEAT (ares_tac [thms_in_pl] 1)); | 
| 760 | 73 | qed "weaken_right"; | 
| 0 | 74 | |
| 75 | (*The deduction theorem*) | |
| 5137 | 76 | Goal "[| cons(p,H) |- q; p:prop |] ==> H |- p=>q"; | 
| 515 | 77 | by (etac thms.induct 1); | 
| 6154 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6141diff
changeset | 78 | by (blast_tac (claset() addIs [thms_I, thms.H RS weaken_right]) 1); | 
| 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6141diff
changeset | 79 | by (blast_tac (claset() addIs [thms.K RS weaken_right]) 1); | 
| 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6141diff
changeset | 80 | by (blast_tac (claset() addIs [thms.S RS weaken_right]) 1); | 
| 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6141diff
changeset | 81 | by (blast_tac (claset() addIs [thms.DN RS weaken_right]) 1); | 
| 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6141diff
changeset | 82 | by (blast_tac (claset() addIs [thms.S RS thms_MP RS thms_MP]) 1); | 
| 760 | 83 | qed "deduction"; | 
| 0 | 84 | |
| 85 | ||
| 86 | (*The cut rule*) | |
| 5137 | 87 | Goal "[| H|-p; cons(p,H) |- q |] ==> H |- q"; | 
| 0 | 88 | by (rtac (deduction RS thms_MP) 1); | 
| 89 | by (REPEAT (ares_tac [thms_in_pl] 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 90 | qed "cut"; | 
| 0 | 91 | |
| 5137 | 92 | Goal "[| H |- Fls; p:prop |] ==> H |- p"; | 
| 515 | 93 | by (rtac (thms.DN RS thms_MP) 1); | 
| 0 | 94 | by (rtac weaken_right 2); | 
| 515 | 95 | by (REPEAT (ares_tac (prop.intrs@[consI1]) 1)); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 96 | qed "thms_FlsE"; | 
| 0 | 97 | |
| 98 | (* [| H |- p=>Fls; H |- p; q: prop |] ==> H |- q *) | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 99 | bind_thm ("thms_notE", (thms_MP RS thms_FlsE));
 | 
| 0 | 100 | |
| 101 | (*Soundness of the rules wrt truth-table semantics*) | |
| 5137 | 102 | Goalw [logcon_def] "H |- p ==> H |= p"; | 
| 515 | 103 | by (etac thms.induct 1); | 
| 6046 | 104 | by Auto_tac; | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 105 | qed "soundness"; | 
| 0 | 106 | |
| 107 | (*** Towards the completeness proof ***) | |
| 108 | ||
| 515 | 109 | val [premf,premq] = goal PropLog.thy | 
| 0 | 110 | "[| H |- p=>Fls; q: prop |] ==> H |- p=>q"; | 
| 111 | by (rtac (premf RS thms_in_pl RS ImpE) 1); | |
| 112 | by (rtac deduction 1); | |
| 113 | by (rtac (premf RS weaken_left_cons RS thms_notE) 1); | |
| 515 | 114 | by (REPEAT (ares_tac [premq, consI1, thms.H] 1)); | 
| 760 | 115 | qed "Fls_Imp"; | 
| 0 | 116 | |
| 515 | 117 | val [premp,premq] = goal PropLog.thy | 
| 0 | 118 | "[| H |- p; H |- q=>Fls |] ==> H |- (p=>q)=>Fls"; | 
| 119 | by (cut_facts_tac ([premp,premq] RL [thms_in_pl]) 1); | |
| 120 | by (etac ImpE 1); | |
| 121 | by (rtac deduction 1); | |
| 122 | by (rtac (premq RS weaken_left_cons RS thms_MP) 1); | |
| 515 | 123 | by (rtac (consI1 RS thms.H RS thms_MP) 1); | 
| 0 | 124 | by (rtac (premp RS weaken_left_cons) 2); | 
| 515 | 125 | by (REPEAT (ares_tac prop.intrs 1)); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 126 | qed "Imp_Fls"; | 
| 0 | 127 | |
| 128 | (*Typical example of strengthening the induction formula*) | |
| 6068 | 129 | Goal "p: prop ==> hyps(p,t) |- (if is_true(p,t) then p else p=>Fls)"; | 
| 6046 | 130 | by (Simp_tac 1); | 
| 6154 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6141diff
changeset | 131 | by (induct_tac "p" 1); | 
| 4091 | 132 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [thms_I, thms.H]))); | 
| 133 | by (safe_tac (claset() addSEs [Fls_Imp RS weaken_left_Un1, | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5137diff
changeset | 134 | Fls_Imp RS weaken_left_Un2])); | 
| 6154 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6141diff
changeset | 135 | by (ALLGOALS (blast_tac (claset() addIs [weaken_left_Un1, weaken_left_Un2, | 
| 5618 | 136 | weaken_right, Imp_Fls]))); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 137 | qed "hyps_thms_if"; | 
| 0 | 138 | |
| 139 | (*Key lemma for completeness; yields a set of assumptions satisfying p*) | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5137diff
changeset | 140 | Goalw [logcon_def] "[| p: prop; 0 |= p |] ==> hyps(p,t) |- p"; | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5137diff
changeset | 141 | by (dtac hyps_thms_if 1); | 
| 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5137diff
changeset | 142 | by (Asm_full_simp_tac 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 143 | qed "logcon_thms_p"; | 
| 0 | 144 | |
| 145 | (*For proving certain theorems in our new propositional logic*) | |
| 146 | val thms_cs = | |
| 515 | 147 | ZF_cs addSIs (prop.intrs @ [deduction]) | 
| 148 | addIs [thms_in_pl, thms.H, thms.H RS thms_MP]; | |
| 0 | 149 | |
| 150 | (*The excluded middle in the form of an elimination rule*) | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5137diff
changeset | 151 | Goal "[| p: prop; q: prop |] ==> H |- (p=>q) => ((p=>Fls)=>q) => q"; | 
| 0 | 152 | by (rtac (deduction RS deduction) 1); | 
| 515 | 153 | by (rtac (thms.DN RS thms_MP) 1); | 
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5137diff
changeset | 154 | by (ALLGOALS (blast_tac thms_cs)); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 155 | qed "thms_excluded_middle"; | 
| 0 | 156 | |
| 157 | (*Hard to prove directly because it requires cuts*) | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5137diff
changeset | 158 | Goal "[| cons(p,H) |- q; cons(p=>Fls,H) |- q; p: prop |] ==> H |- q"; | 
| 0 | 159 | by (rtac (thms_excluded_middle RS thms_MP RS thms_MP) 1); | 
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5137diff
changeset | 160 | by (REPEAT (ares_tac (prop.intrs@[deduction,thms_in_pl]) 1)); | 
| 760 | 161 | qed "thms_excluded_middle_rule"; | 
| 0 | 162 | |
| 163 | (*** Completeness -- lemmas for reducing the set of assumptions ***) | |
| 164 | ||
| 165 | (*For the case hyps(p,t)-cons(#v,Y) |- p; | |
| 166 |   we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *)
 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5137diff
changeset | 167 | Goal "p: prop ==> hyps(p, t-{v}) <= cons(#v=>Fls, hyps(p,t)-{#v})";
 | 
| 6154 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6141diff
changeset | 168 | by (induct_tac "p" 1); | 
| 6141 | 169 | by Auto_tac; | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 170 | qed "hyps_Diff"; | 
| 0 | 171 | |
| 172 | (*For the case hyps(p,t)-cons(#v => Fls,Y) |- p; | |
| 173 |   we also have hyps(p,t)-{#v=>Fls} <= hyps(p, cons(v,t)) *)
 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5137diff
changeset | 174 | Goal "p: prop ==> hyps(p, cons(v,t)) <= cons(#v, hyps(p,t)-{#v=>Fls})";
 | 
| 6154 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6141diff
changeset | 175 | by (induct_tac "p" 1); | 
| 6141 | 176 | by Auto_tac; | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 177 | qed "hyps_cons"; | 
| 0 | 178 | |
| 179 | (** Two lemmas for use with weaken_left **) | |
| 180 | ||
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5137diff
changeset | 181 | Goal "B-C <= cons(a, B-cons(a,C))"; | 
| 2469 | 182 | by (Fast_tac 1); | 
| 760 | 183 | qed "cons_Diff_same"; | 
| 0 | 184 | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5137diff
changeset | 185 | Goal "cons(a, B-{c}) - D <= cons(a, B-cons(c,D))";
 | 
| 2469 | 186 | by (Fast_tac 1); | 
| 760 | 187 | qed "cons_Diff_subset2"; | 
| 0 | 188 | |
| 189 | (*The set hyps(p,t) is finite, and elements have the form #v or #v=>Fls; | |
| 190 | could probably prove the stronger hyps(p,t) : Fin(hyps(p,0) Un hyps(p,nat))*) | |
| 5325 
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
 paulson parents: 
5137diff
changeset | 191 | Goal "p: prop ==> hyps(p,t) : Fin(UN v:nat. {#v, #v=>Fls})";
 | 
| 6154 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6141diff
changeset | 192 | by (induct_tac "p" 1); | 
| 5618 | 193 | by (asm_simp_tac (simpset() addsimps [UN_I]) 2); | 
| 2469 | 194 | by (ALLGOALS Asm_simp_tac); | 
| 6154 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6141diff
changeset | 195 | by (blast_tac (claset() addIs Fin.intrs) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 196 | qed "hyps_finite"; | 
| 0 | 197 | |
| 198 | val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left; | |
| 199 | ||
| 200 | (*Induction on the finite set of assumptions hyps(p,t0). | |
| 201 | We may repeatedly subtract assumptions until none are left!*) | |
| 515 | 202 | val [premp,sat] = goal PropLog.thy | 
| 0 | 203 | "[| p: prop; 0 |= p |] ==> ALL t. hyps(p,t) - hyps(p,t0) |- p"; | 
| 204 | by (rtac (premp RS hyps_finite RS Fin_induct) 1); | |
| 4091 | 205 | by (simp_tac (simpset() addsimps [premp, sat, logcon_thms_p, Diff_0]) 1); | 
| 4152 | 206 | by Safe_tac; | 
| 0 | 207 | (*Case hyps(p,t)-cons(#v,Y) |- p *) | 
| 208 | by (rtac thms_excluded_middle_rule 1); | |
| 515 | 209 | by (etac prop.Var_I 3); | 
| 0 | 210 | by (rtac (cons_Diff_same RS weaken_left) 1); | 
| 211 | by (etac spec 1); | |
| 212 | by (rtac (cons_Diff_subset2 RS weaken_left) 1); | |
| 213 | by (rtac (premp RS hyps_Diff RS Diff_weaken_left) 1); | |
| 214 | by (etac spec 1); | |
| 215 | (*Case hyps(p,t)-cons(#v => Fls,Y) |- p *) | |
| 216 | by (rtac thms_excluded_middle_rule 1); | |
| 515 | 217 | by (etac prop.Var_I 3); | 
| 0 | 218 | by (rtac (cons_Diff_same RS weaken_left) 2); | 
| 219 | by (etac spec 2); | |
| 220 | by (rtac (cons_Diff_subset2 RS weaken_left) 1); | |
| 221 | by (rtac (premp RS hyps_cons RS Diff_weaken_left) 1); | |
| 222 | by (etac spec 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 223 | qed "completeness_0_lemma"; | 
| 0 | 224 | |
| 225 | (*The base case for completeness*) | |
| 515 | 226 | val [premp,sat] = goal PropLog.thy "[| p: prop; 0 |= p |] ==> 0 |- p"; | 
| 0 | 227 | by (rtac (Diff_cancel RS subst) 1); | 
| 228 | by (rtac (sat RS (premp RS completeness_0_lemma RS spec)) 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 229 | qed "completeness_0"; | 
| 0 | 230 | |
| 231 | (*A semantic analogue of the Deduction Theorem*) | |
| 5137 | 232 | Goalw [logcon_def] "[| cons(p,H) |= q |] ==> H |= p=>q"; | 
| 6154 
6a00a5baef2b
automatic insertion of datatype intr rules into claset
 paulson parents: 
6141diff
changeset | 233 | by Auto_tac; | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 234 | qed "logcon_Imp"; | 
| 0 | 235 | |
| 5137 | 236 | Goal "H: Fin(prop) ==> ALL p:prop. H |= p --> H |- p"; | 
| 0 | 237 | by (etac Fin_induct 1); | 
| 4091 | 238 | by (safe_tac (claset() addSIs [completeness_0])); | 
| 0 | 239 | by (rtac (weaken_left_cons RS thms_MP) 1); | 
| 5137 | 240 | by (blast_tac (claset() addSIs (logcon_Imp::prop.intrs)) 1); | 
| 241 | by (blast_tac thms_cs 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 242 | qed "completeness_lemma"; | 
| 0 | 243 | |
| 244 | val completeness = completeness_lemma RS bspec RS mp; | |
| 245 | ||
| 515 | 246 | val [finite] = goal PropLog.thy "H: Fin(prop) ==> H |- p <-> H |= p & p:prop"; | 
| 4091 | 247 | by (fast_tac (claset() addSEs [soundness, finite RS completeness, | 
| 5618 | 248 | thms_in_pl]) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 249 | qed "thms_iff"; | 
| 0 | 250 | |
| 251 | writeln"Reached end of file."; |