| author | paulson | 
| Tue, 24 Feb 1998 11:35:33 +0100 | |
| changeset 4648 | f04da668581c | 
| parent 4152 | 451104c223e2 | 
| child 5068 | fb28eaa07e01 | 
| 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 | ||
| 6 | For ex/prop-log.thy. Inductive definition of propositional logic. | |
| 7 | Soundness and completeness w.r.t. truth-tables. | |
| 8 | ||
| 9 | Prove: If H|=p then G|=p where G:Fin(H) | |
| 10 | *) | |
| 11 | ||
| 12 | open PropLog; | |
| 13 | ||
| 14 | (*** prop_rec -- by Vset recursion ***) | |
| 15 | ||
| 16 | (** conversion rules **) | |
| 17 | ||
| 18 | goal PropLog.thy "prop_rec(Fls,b,c,d) = b"; | |
| 19 | by (rtac (prop_rec_def RS def_Vrec RS trans) 1); | |
| 515 | 20 | by (rewrite_goals_tac prop.con_defs); | 
| 2469 | 21 | by (Simp_tac 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 22 | qed "prop_rec_Fls"; | 
| 0 | 23 | |
| 24 | goal PropLog.thy "prop_rec(#v,b,c,d) = c(v)"; | |
| 25 | by (rtac (prop_rec_def RS def_Vrec RS trans) 1); | |
| 515 | 26 | by (rewrite_goals_tac prop.con_defs); | 
| 2469 | 27 | by (Simp_tac 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 28 | qed "prop_rec_Var"; | 
| 0 | 29 | |
| 30 | goal PropLog.thy "prop_rec(p=>q,b,c,d) = \ | |
| 31 | \ d(p, q, prop_rec(p,b,c,d), prop_rec(q,b,c,d))"; | |
| 32 | by (rtac (prop_rec_def RS def_Vrec RS trans) 1); | |
| 515 | 33 | by (rewrite_goals_tac prop.con_defs); | 
| 7 
268f93ab3bc4
Installation of new simplifier for ZF/ex.  The hom_ss example in misc.ML is
 lcp parents: 
0diff
changeset | 34 | by (simp_tac rank_ss 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 35 | qed "prop_rec_Imp"; | 
| 0 | 36 | |
| 2469 | 37 | Addsimps [prop_rec_Fls, prop_rec_Var, prop_rec_Imp]; | 
| 0 | 38 | |
| 39 | (*** Semantics of propositional logic ***) | |
| 40 | ||
| 41 | (** The function is_true **) | |
| 42 | ||
| 43 | goalw PropLog.thy [is_true_def] "is_true(Fls,t) <-> False"; | |
| 4091 | 44 | by (simp_tac (simpset() addsimps [one_not_0 RS not_sym]) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 45 | qed "is_true_Fls"; | 
| 0 | 46 | |
| 47 | goalw PropLog.thy [is_true_def] "is_true(#v,t) <-> v:t"; | |
| 4091 | 48 | by (simp_tac (simpset() addsimps [one_not_0 RS not_sym] | 
| 1461 | 49 | setloop (split_tac [expand_if])) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 50 | qed "is_true_Var"; | 
| 0 | 51 | |
| 52 | goalw PropLog.thy [is_true_def] | |
| 53 | "is_true(p=>q,t) <-> (is_true(p,t)-->is_true(q,t))"; | |
| 4091 | 54 | by (simp_tac (simpset() setloop (split_tac [expand_if])) 1); | 
| 760 | 55 | qed "is_true_Imp"; | 
| 0 | 56 | |
| 57 | (** The function hyps **) | |
| 58 | ||
| 59 | goalw PropLog.thy [hyps_def] "hyps(Fls,t) = 0"; | |
| 2469 | 60 | by (Simp_tac 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 61 | qed "hyps_Fls"; | 
| 0 | 62 | |
| 63 | goalw PropLog.thy [hyps_def] "hyps(#v,t) = {if(v:t, #v, #v=>Fls)}";
 | |
| 2469 | 64 | by (Simp_tac 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 65 | qed "hyps_Var"; | 
| 0 | 66 | |
| 67 | goalw PropLog.thy [hyps_def] "hyps(p=>q,t) = hyps(p,t) Un hyps(q,t)"; | |
| 2469 | 68 | by (Simp_tac 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 69 | qed "hyps_Imp"; | 
| 0 | 70 | |
| 2469 | 71 | Addsimps prop.intrs; | 
| 72 | Addsimps [is_true_Fls, is_true_Var, is_true_Imp, | |
| 73 | hyps_Fls, hyps_Var, hyps_Imp]; | |
| 0 | 74 | |
| 75 | (*** Proof theory of propositional logic ***) | |
| 76 | ||
| 515 | 77 | goalw PropLog.thy thms.defs "!!G H. G<=H ==> thms(G) <= thms(H)"; | 
| 0 | 78 | by (rtac lfp_mono 1); | 
| 515 | 79 | by (REPEAT (rtac thms.bnd_mono 1)); | 
| 0 | 80 | by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 81 | qed "thms_mono"; | 
| 0 | 82 | |
| 515 | 83 | val thms_in_pl = thms.dom_subset RS subsetD; | 
| 0 | 84 | |
| 515 | 85 | val ImpE = prop.mk_cases prop.con_defs "p=>q : prop"; | 
| 0 | 86 | |
| 515 | 87 | (*Stronger Modus Ponens rule: no typechecking!*) | 
| 88 | goal PropLog.thy "!!p q H. [| H |- p=>q; H |- p |] ==> H |- q"; | |
| 89 | by (rtac thms.MP 1); | |
| 0 | 90 | by (REPEAT (eresolve_tac [asm_rl, thms_in_pl, thms_in_pl RS ImpE] 1)); | 
| 760 | 91 | qed "thms_MP"; | 
| 0 | 92 | |
| 93 | (*Rule is called I for Identity Combinator, not for Introduction*) | |
| 515 | 94 | goal PropLog.thy "!!p H. p:prop ==> H |- p=>p"; | 
| 95 | by (rtac (thms.S RS thms_MP RS thms_MP) 1); | |
| 96 | by (rtac thms.K 5); | |
| 97 | by (rtac thms.K 4); | |
| 98 | by (REPEAT (ares_tac prop.intrs 1)); | |
| 760 | 99 | qed "thms_I"; | 
| 0 | 100 | |
| 101 | (** Weakening, left and right **) | |
| 102 | ||
| 103 | (* [| 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 | 104 | bind_thm ("weaken_left", (thms_mono RS subsetD));
 | 
| 0 | 105 | |
| 106 | (* H |- p ==> cons(a,H) |- p *) | |
| 107 | val weaken_left_cons = subset_consI RS weaken_left; | |
| 108 | ||
| 109 | val weaken_left_Un1 = Un_upper1 RS weaken_left; | |
| 110 | val weaken_left_Un2 = Un_upper2 RS weaken_left; | |
| 111 | ||
| 515 | 112 | goal PropLog.thy "!!H p q. [| H |- q; p:prop |] ==> H |- p=>q"; | 
| 113 | by (rtac (thms.K RS thms_MP) 1); | |
| 0 | 114 | by (REPEAT (ares_tac [thms_in_pl] 1)); | 
| 760 | 115 | qed "weaken_right"; | 
| 0 | 116 | |
| 117 | (*The deduction theorem*) | |
| 515 | 118 | goal PropLog.thy "!!p q H. [| cons(p,H) |- q; p:prop |] ==> H |- p=>q"; | 
| 119 | by (etac thms.induct 1); | |
| 4091 | 120 | by (fast_tac (claset() addIs [thms_I, thms.H RS weaken_right]) 1); | 
| 121 | by (fast_tac (claset() addIs [thms.K RS weaken_right]) 1); | |
| 122 | by (fast_tac (claset() addIs [thms.S RS weaken_right]) 1); | |
| 123 | by (fast_tac (claset() addIs [thms.DN RS weaken_right]) 1); | |
| 124 | by (fast_tac (claset() addIs [thms.S RS thms_MP RS thms_MP]) 1); | |
| 760 | 125 | qed "deduction"; | 
| 0 | 126 | |
| 127 | ||
| 128 | (*The cut rule*) | |
| 515 | 129 | goal PropLog.thy "!!H p q. [| H|-p; cons(p,H) |- q |] ==> H |- q"; | 
| 0 | 130 | by (rtac (deduction RS thms_MP) 1); | 
| 131 | by (REPEAT (ares_tac [thms_in_pl] 1)); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 132 | qed "cut"; | 
| 0 | 133 | |
| 515 | 134 | goal PropLog.thy "!!H p. [| H |- Fls; p:prop |] ==> H |- p"; | 
| 135 | by (rtac (thms.DN RS thms_MP) 1); | |
| 0 | 136 | by (rtac weaken_right 2); | 
| 515 | 137 | by (REPEAT (ares_tac (prop.intrs@[consI1]) 1)); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 138 | qed "thms_FlsE"; | 
| 0 | 139 | |
| 140 | (* [| H |- p=>Fls; H |- p; q: prop |] ==> H |- q *) | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 141 | bind_thm ("thms_notE", (thms_MP RS thms_FlsE));
 | 
| 0 | 142 | |
| 143 | (*Soundness of the rules wrt truth-table semantics*) | |
| 515 | 144 | goalw PropLog.thy [logcon_def] "!!H. H |- p ==> H |= p"; | 
| 145 | by (etac thms.induct 1); | |
| 4091 | 146 | by (fast_tac (claset() addSDs [is_true_Imp RS iffD1 RS mp]) 5); | 
| 2469 | 147 | by (ALLGOALS Asm_simp_tac); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 148 | qed "soundness"; | 
| 0 | 149 | |
| 150 | (*** Towards the completeness proof ***) | |
| 151 | ||
| 515 | 152 | val [premf,premq] = goal PropLog.thy | 
| 0 | 153 | "[| H |- p=>Fls; q: prop |] ==> H |- p=>q"; | 
| 154 | by (rtac (premf RS thms_in_pl RS ImpE) 1); | |
| 155 | by (rtac deduction 1); | |
| 156 | by (rtac (premf RS weaken_left_cons RS thms_notE) 1); | |
| 515 | 157 | by (REPEAT (ares_tac [premq, consI1, thms.H] 1)); | 
| 760 | 158 | qed "Fls_Imp"; | 
| 0 | 159 | |
| 515 | 160 | val [premp,premq] = goal PropLog.thy | 
| 0 | 161 | "[| H |- p; H |- q=>Fls |] ==> H |- (p=>q)=>Fls"; | 
| 162 | by (cut_facts_tac ([premp,premq] RL [thms_in_pl]) 1); | |
| 163 | by (etac ImpE 1); | |
| 164 | by (rtac deduction 1); | |
| 165 | by (rtac (premq RS weaken_left_cons RS thms_MP) 1); | |
| 515 | 166 | by (rtac (consI1 RS thms.H RS thms_MP) 1); | 
| 0 | 167 | by (rtac (premp RS weaken_left_cons) 2); | 
| 515 | 168 | by (REPEAT (ares_tac prop.intrs 1)); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 169 | qed "Imp_Fls"; | 
| 0 | 170 | |
| 171 | (*Typical example of strengthening the induction formula*) | |
| 515 | 172 | val [major] = goal PropLog.thy | 
| 0 | 173 | "p: prop ==> hyps(p,t) |- if(is_true(p,t), p, p=>Fls)"; | 
| 174 | by (rtac (expand_if RS iffD2) 1); | |
| 515 | 175 | by (rtac (major RS prop.induct) 1); | 
| 4091 | 176 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [thms_I, thms.H]))); | 
| 177 | by (safe_tac (claset() addSEs [Fls_Imp RS weaken_left_Un1, | |
| 1461 | 178 | Fls_Imp RS weaken_left_Un2])); | 
| 4091 | 179 | by (ALLGOALS (fast_tac (claset() addIs [weaken_left_Un1, weaken_left_Un2, | 
| 1461 | 180 | weaken_right, Imp_Fls]))); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 181 | qed "hyps_thms_if"; | 
| 0 | 182 | |
| 183 | (*Key lemma for completeness; yields a set of assumptions satisfying p*) | |
| 515 | 184 | val [premp,sat] = goalw PropLog.thy [logcon_def] | 
| 0 | 185 | "[| p: prop; 0 |= p |] ==> hyps(p,t) |- p"; | 
| 186 | by (rtac (sat RS spec RS mp RS if_P RS subst) 1 THEN | |
| 187 | rtac (premp RS hyps_thms_if) 2); | |
| 2469 | 188 | by (Fast_tac 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 189 | qed "logcon_thms_p"; | 
| 0 | 190 | |
| 191 | (*For proving certain theorems in our new propositional logic*) | |
| 192 | val thms_cs = | |
| 515 | 193 | ZF_cs addSIs (prop.intrs @ [deduction]) | 
| 194 | addIs [thms_in_pl, thms.H, thms.H RS thms_MP]; | |
| 0 | 195 | |
| 196 | (*The excluded middle in the form of an elimination rule*) | |
| 515 | 197 | val prems = goal PropLog.thy | 
| 0 | 198 | "[| p: prop; q: prop |] ==> H |- (p=>q) => ((p=>Fls)=>q) => q"; | 
| 199 | by (rtac (deduction RS deduction) 1); | |
| 515 | 200 | by (rtac (thms.DN RS thms_MP) 1); | 
| 0 | 201 | by (ALLGOALS (best_tac (thms_cs addSIs prems))); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 202 | qed "thms_excluded_middle"; | 
| 0 | 203 | |
| 204 | (*Hard to prove directly because it requires cuts*) | |
| 515 | 205 | val prems = goal PropLog.thy | 
| 0 | 206 | "[| cons(p,H) |- q; cons(p=>Fls,H) |- q; p: prop |] ==> H |- q"; | 
| 207 | by (rtac (thms_excluded_middle RS thms_MP RS thms_MP) 1); | |
| 515 | 208 | by (REPEAT (resolve_tac (prems@prop.intrs@[deduction,thms_in_pl]) 1)); | 
| 760 | 209 | qed "thms_excluded_middle_rule"; | 
| 0 | 210 | |
| 211 | (*** Completeness -- lemmas for reducing the set of assumptions ***) | |
| 212 | ||
| 213 | (*For the case hyps(p,t)-cons(#v,Y) |- p; | |
| 214 |   we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *)
 | |
| 515 | 215 | val [major] = goal PropLog.thy | 
| 0 | 216 |     "p: prop ==> hyps(p, t-{v}) <= cons(#v=>Fls, hyps(p,t)-{#v})";
 | 
| 515 | 217 | by (rtac (major RS prop.induct) 1); | 
| 2469 | 218 | by (Simp_tac 1); | 
| 4091 | 219 | by (asm_simp_tac (simpset() setloop (split_tac [expand_if])) 1); | 
| 220 | by (fast_tac (claset() addSEs prop.free_SEs) 1); | |
| 2469 | 221 | by (Asm_simp_tac 1); | 
| 222 | by (Fast_tac 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 223 | qed "hyps_Diff"; | 
| 0 | 224 | |
| 225 | (*For the case hyps(p,t)-cons(#v => Fls,Y) |- p; | |
| 226 |   we also have hyps(p,t)-{#v=>Fls} <= hyps(p, cons(v,t)) *)
 | |
| 515 | 227 | val [major] = goal PropLog.thy | 
| 0 | 228 |     "p: prop ==> hyps(p, cons(v,t)) <= cons(#v, hyps(p,t)-{#v=>Fls})";
 | 
| 515 | 229 | by (rtac (major RS prop.induct) 1); | 
| 2469 | 230 | by (Simp_tac 1); | 
| 4091 | 231 | by (asm_simp_tac (simpset() setloop (split_tac [expand_if])) 1); | 
| 232 | by (fast_tac (claset() addSEs prop.free_SEs) 1); | |
| 2469 | 233 | by (Asm_simp_tac 1); | 
| 234 | by (Fast_tac 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 235 | qed "hyps_cons"; | 
| 0 | 236 | |
| 237 | (** Two lemmas for use with weaken_left **) | |
| 238 | ||
| 2873 | 239 | goal ZF.thy "B-C <= cons(a, B-cons(a,C))"; | 
| 2469 | 240 | by (Fast_tac 1); | 
| 760 | 241 | qed "cons_Diff_same"; | 
| 0 | 242 | |
| 2873 | 243 | goal ZF.thy "cons(a, B-{c}) - D <= cons(a, B-cons(c,D))";
 | 
| 2469 | 244 | by (Fast_tac 1); | 
| 760 | 245 | qed "cons_Diff_subset2"; | 
| 0 | 246 | |
| 247 | (*The set hyps(p,t) is finite, and elements have the form #v or #v=>Fls; | |
| 248 | could probably prove the stronger hyps(p,t) : Fin(hyps(p,0) Un hyps(p,nat))*) | |
| 515 | 249 | val [major] = goal PropLog.thy | 
| 0 | 250 |     "p: prop ==> hyps(p,t) : Fin(UN v:nat. {#v, #v=>Fls})";
 | 
| 515 | 251 | by (rtac (major RS prop.induct) 1); | 
| 4091 | 252 | by (asm_simp_tac (simpset() addsimps [UN_I] | 
| 1461 | 253 | setloop (split_tac [expand_if])) 2); | 
| 2469 | 254 | by (ALLGOALS Asm_simp_tac); | 
| 4091 | 255 | by (fast_tac (claset() addIs Fin.intrs) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 256 | qed "hyps_finite"; | 
| 0 | 257 | |
| 258 | val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left; | |
| 259 | ||
| 260 | (*Induction on the finite set of assumptions hyps(p,t0). | |
| 261 | We may repeatedly subtract assumptions until none are left!*) | |
| 515 | 262 | val [premp,sat] = goal PropLog.thy | 
| 0 | 263 | "[| p: prop; 0 |= p |] ==> ALL t. hyps(p,t) - hyps(p,t0) |- p"; | 
| 264 | by (rtac (premp RS hyps_finite RS Fin_induct) 1); | |
| 4091 | 265 | by (simp_tac (simpset() addsimps [premp, sat, logcon_thms_p, Diff_0]) 1); | 
| 4152 | 266 | by Safe_tac; | 
| 0 | 267 | (*Case hyps(p,t)-cons(#v,Y) |- p *) | 
| 268 | by (rtac thms_excluded_middle_rule 1); | |
| 515 | 269 | by (etac prop.Var_I 3); | 
| 0 | 270 | by (rtac (cons_Diff_same RS weaken_left) 1); | 
| 271 | by (etac spec 1); | |
| 272 | by (rtac (cons_Diff_subset2 RS weaken_left) 1); | |
| 273 | by (rtac (premp RS hyps_Diff RS Diff_weaken_left) 1); | |
| 274 | by (etac spec 1); | |
| 275 | (*Case hyps(p,t)-cons(#v => Fls,Y) |- p *) | |
| 276 | by (rtac thms_excluded_middle_rule 1); | |
| 515 | 277 | by (etac prop.Var_I 3); | 
| 0 | 278 | by (rtac (cons_Diff_same RS weaken_left) 2); | 
| 279 | by (etac spec 2); | |
| 280 | by (rtac (cons_Diff_subset2 RS weaken_left) 1); | |
| 281 | by (rtac (premp RS hyps_cons RS Diff_weaken_left) 1); | |
| 282 | by (etac spec 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 283 | qed "completeness_0_lemma"; | 
| 0 | 284 | |
| 285 | (*The base case for completeness*) | |
| 515 | 286 | val [premp,sat] = goal PropLog.thy "[| p: prop; 0 |= p |] ==> 0 |- p"; | 
| 0 | 287 | by (rtac (Diff_cancel RS subst) 1); | 
| 288 | 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 | 289 | qed "completeness_0"; | 
| 0 | 290 | |
| 291 | (*A semantic analogue of the Deduction Theorem*) | |
| 515 | 292 | goalw PropLog.thy [logcon_def] "!!H p q. [| cons(p,H) |= q |] ==> H |= p=>q"; | 
| 2469 | 293 | by (Simp_tac 1); | 
| 294 | by (Fast_tac 1); | |
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 295 | qed "logcon_Imp"; | 
| 0 | 296 | |
| 515 | 297 | goal PropLog.thy "!!H. H: Fin(prop) ==> ALL p:prop. H |= p --> H |- p"; | 
| 0 | 298 | by (etac Fin_induct 1); | 
| 4091 | 299 | by (safe_tac (claset() addSIs [completeness_0])); | 
| 0 | 300 | by (rtac (weaken_left_cons RS thms_MP) 1); | 
| 4091 | 301 | by (fast_tac (claset() addSIs (logcon_Imp::prop.intrs)) 1); | 
| 0 | 302 | by (fast_tac thms_cs 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 303 | qed "completeness_lemma"; | 
| 0 | 304 | |
| 305 | val completeness = completeness_lemma RS bspec RS mp; | |
| 306 | ||
| 515 | 307 | val [finite] = goal PropLog.thy "H: Fin(prop) ==> H |- p <-> H |= p & p:prop"; | 
| 4091 | 308 | by (fast_tac (claset() addSEs [soundness, finite RS completeness, | 
| 1461 | 309 | thms_in_pl]) 1); | 
| 782 
200a16083201
added bind_thm for theorems defined by "standard ..."
 clasohm parents: 
760diff
changeset | 310 | qed "thms_iff"; | 
| 0 | 311 | |
| 312 | writeln"Reached end of file."; |