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