author | wenzelm |
Tue, 28 Oct 1997 17:58:35 +0100 | |
changeset 4029 | 22f2d1b17f97 |
parent 16 | 0b033d50ca1c |
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); |
|
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); |
|
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); |
|
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 |
|
7
268f93ab3bc4
Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents:
0
diff
changeset
|
73 |
addsimps Prop.intrs |
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 |
||
79 |
structure PropThms = Inductive_Fun |
|
80 |
(val thy = PropLog.thy; |
|
81 |
val rec_doms = [("thms","prop")]; |
|
82 |
val sintrs = |
|
83 |
["[| p:H; p:prop |] ==> H |- p", |
|
84 |
"[| p:prop; q:prop |] ==> H |- p=>q=>p", |
|
85 |
"[| p:prop; q:prop; r:prop |] ==> H |- (p=>q=>r) => (p=>q) => p=>r", |
|
86 |
"p:prop ==> H |- ((p=>Fls) => Fls) => p", |
|
87 |
"[| H |- p=>q; H |- p; p:prop; q:prop |] ==> H |- q"]; |
|
88 |
val monos = []; |
|
89 |
val con_defs = []; |
|
90 |
val type_intrs = Prop.intrs; |
|
91 |
val type_elims = []); |
|
92 |
||
93 |
goalw PropThms.thy PropThms.defs "!!G H. G<=H ==> thms(G) <= thms(H)"; |
|
94 |
by (rtac lfp_mono 1); |
|
95 |
by (REPEAT (rtac PropThms.bnd_mono 1)); |
|
96 |
by (REPEAT (ares_tac (univ_mono::basic_monos) 1)); |
|
97 |
val thms_mono = result(); |
|
98 |
||
99 |
val thms_in_pl = PropThms.dom_subset RS subsetD; |
|
100 |
||
101 |
val [thms_H, thms_K, thms_S, thms_DN, weak_thms_MP] = PropThms.intrs; |
|
102 |
||
103 |
(*Modus Ponens rule -- this stronger version avoids typecheck*) |
|
104 |
goal PropThms.thy "!!p q H. [| H |- p=>q; H |- p |] ==> H |- q"; |
|
105 |
by (rtac weak_thms_MP 1); |
|
106 |
by (REPEAT (eresolve_tac [asm_rl, thms_in_pl, thms_in_pl RS ImpE] 1)); |
|
107 |
val thms_MP = result(); |
|
108 |
||
109 |
(*Rule is called I for Identity Combinator, not for Introduction*) |
|
110 |
goal PropThms.thy "!!p H. p:prop ==> H |- p=>p"; |
|
111 |
by (rtac (thms_S RS thms_MP RS thms_MP) 1); |
|
112 |
by (rtac thms_K 5); |
|
113 |
by (rtac thms_K 4); |
|
114 |
by (REPEAT (ares_tac [ImpI] 1)); |
|
115 |
val thms_I = result(); |
|
116 |
||
117 |
(** Weakening, left and right **) |
|
118 |
||
119 |
(* [| G<=H; G|-p |] ==> H|-p Order of premises is convenient with RS*) |
|
120 |
val weaken_left = standard (thms_mono RS subsetD); |
|
121 |
||
122 |
(* H |- p ==> cons(a,H) |- p *) |
|
123 |
val weaken_left_cons = subset_consI RS weaken_left; |
|
124 |
||
125 |
val weaken_left_Un1 = Un_upper1 RS weaken_left; |
|
126 |
val weaken_left_Un2 = Un_upper2 RS weaken_left; |
|
127 |
||
128 |
goal PropThms.thy "!!H p q. [| H |- q; p:prop |] ==> H |- p=>q"; |
|
129 |
by (rtac (thms_K RS thms_MP) 1); |
|
130 |
by (REPEAT (ares_tac [thms_in_pl] 1)); |
|
131 |
val weaken_right = result(); |
|
132 |
||
133 |
(*The deduction theorem*) |
|
134 |
goal PropThms.thy "!!p q H. [| cons(p,H) |- q; p:prop |] ==> H |- p=>q"; |
|
135 |
by (etac PropThms.induct 1); |
|
136 |
by (fast_tac (ZF_cs addIs [thms_I, thms_H RS weaken_right]) 1); |
|
137 |
by (fast_tac (ZF_cs addIs [thms_K RS weaken_right]) 1); |
|
138 |
by (fast_tac (ZF_cs addIs [thms_S RS weaken_right]) 1); |
|
139 |
by (fast_tac (ZF_cs addIs [thms_DN RS weaken_right]) 1); |
|
140 |
by (fast_tac (ZF_cs addIs [thms_S RS thms_MP RS thms_MP]) 1); |
|
141 |
val deduction = result(); |
|
142 |
||
143 |
||
144 |
(*The cut rule*) |
|
145 |
goal PropThms.thy "!!H p q. [| H|-p; cons(p,H) |- q |] ==> H |- q"; |
|
146 |
by (rtac (deduction RS thms_MP) 1); |
|
147 |
by (REPEAT (ares_tac [thms_in_pl] 1)); |
|
148 |
val cut = result(); |
|
149 |
||
150 |
goal PropThms.thy "!!H p. [| H |- Fls; p:prop |] ==> H |- p"; |
|
151 |
by (rtac (thms_DN RS thms_MP) 1); |
|
152 |
by (rtac weaken_right 2); |
|
153 |
by (REPEAT (ares_tac (Prop.intrs@[consI1]) 1)); |
|
154 |
val thms_FlsE = result(); |
|
155 |
||
156 |
(* [| H |- p=>Fls; H |- p; q: prop |] ==> H |- q *) |
|
157 |
val thms_notE = standard (thms_MP RS thms_FlsE); |
|
158 |
||
159 |
(*Soundness of the rules wrt truth-table semantics*) |
|
7
268f93ab3bc4
Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents:
0
diff
changeset
|
160 |
goalw PropThms.thy [sat_def] "!!H. H |- p ==> H |= p"; |
268f93ab3bc4
Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents:
0
diff
changeset
|
161 |
by (etac PropThms.induct 1); |
0 | 162 |
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
|
163 |
by (ALLGOALS (asm_simp_tac prop_ss)); |
0 | 164 |
val soundness = result(); |
165 |
||
166 |
(*** Towards the completeness proof ***) |
|
167 |
||
168 |
val [premf,premq] = goal PropThms.thy |
|
169 |
"[| H |- p=>Fls; q: prop |] ==> H |- p=>q"; |
|
170 |
by (rtac (premf RS thms_in_pl RS ImpE) 1); |
|
171 |
by (rtac deduction 1); |
|
172 |
by (rtac (premf RS weaken_left_cons RS thms_notE) 1); |
|
173 |
by (REPEAT (ares_tac [premq, consI1, thms_H] 1)); |
|
174 |
val Fls_Imp = result(); |
|
175 |
||
176 |
val [premp,premq] = goal PropThms.thy |
|
177 |
"[| H |- p; H |- q=>Fls |] ==> H |- (p=>q)=>Fls"; |
|
178 |
by (cut_facts_tac ([premp,premq] RL [thms_in_pl]) 1); |
|
179 |
by (etac ImpE 1); |
|
180 |
by (rtac deduction 1); |
|
181 |
by (rtac (premq RS weaken_left_cons RS thms_MP) 1); |
|
182 |
by (rtac (consI1 RS thms_H RS thms_MP) 1); |
|
183 |
by (rtac (premp RS weaken_left_cons) 2); |
|
184 |
by (REPEAT (ares_tac Prop.intrs 1)); |
|
185 |
val Imp_Fls = result(); |
|
186 |
||
187 |
(*Typical example of strengthening the induction formula*) |
|
188 |
val [major] = goal PropThms.thy |
|
189 |
"p: prop ==> hyps(p,t) |- if(is_true(p,t), p, p=>Fls)"; |
|
190 |
by (rtac (expand_if RS iffD2) 1); |
|
191 |
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
|
192 |
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
|
193 |
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
|
194 |
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
|
195 |
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
|
196 |
weaken_right, Imp_Fls]))); |
0 | 197 |
val hyps_thms_if = result(); |
198 |
||
199 |
(*Key lemma for completeness; yields a set of assumptions satisfying p*) |
|
200 |
val [premp,sat] = goalw PropThms.thy [sat_def] |
|
201 |
"[| p: prop; 0 |= p |] ==> hyps(p,t) |- p"; |
|
202 |
by (rtac (sat RS spec RS mp RS if_P RS subst) 1 THEN |
|
203 |
rtac (premp RS hyps_thms_if) 2); |
|
204 |
by (fast_tac ZF_cs 1); |
|
205 |
val sat_thms_p = result(); |
|
206 |
||
207 |
(*For proving certain theorems in our new propositional logic*) |
|
208 |
val thms_cs = |
|
209 |
ZF_cs addSIs [FlsI, VarI, ImpI, deduction] |
|
210 |
addIs [thms_in_pl, thms_H, thms_H RS thms_MP]; |
|
211 |
||
212 |
(*The excluded middle in the form of an elimination rule*) |
|
213 |
val prems = goal PropThms.thy |
|
214 |
"[| p: prop; q: prop |] ==> H |- (p=>q) => ((p=>Fls)=>q) => q"; |
|
215 |
by (rtac (deduction RS deduction) 1); |
|
216 |
by (rtac (thms_DN RS thms_MP) 1); |
|
217 |
by (ALLGOALS (best_tac (thms_cs addSIs prems))); |
|
218 |
val thms_excluded_middle = result(); |
|
219 |
||
220 |
(*Hard to prove directly because it requires cuts*) |
|
221 |
val prems = goal PropThms.thy |
|
222 |
"[| cons(p,H) |- q; cons(p=>Fls,H) |- q; p: prop |] ==> H |- q"; |
|
223 |
by (rtac (thms_excluded_middle RS thms_MP RS thms_MP) 1); |
|
224 |
by (REPEAT (resolve_tac (prems@Prop.intrs@[deduction,thms_in_pl]) 1)); |
|
225 |
val thms_excluded_middle_rule = result(); |
|
226 |
||
227 |
(*** Completeness -- lemmas for reducing the set of assumptions ***) |
|
228 |
||
229 |
(*For the case hyps(p,t)-cons(#v,Y) |- p; |
|
230 |
we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *) |
|
231 |
val [major] = goal PropThms.thy |
|
232 |
"p: prop ==> hyps(p, t-{v}) <= cons(#v=>Fls, hyps(p,t)-{#v})"; |
|
233 |
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
|
234 |
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
|
235 |
by (asm_simp_tac (prop_ss setloop (split_tac [expand_if])) 1); |
0 | 236 |
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
|
237 |
by (asm_simp_tac prop_ss 1); |
0 | 238 |
by (fast_tac ZF_cs 1); |
239 |
val hyps_Diff = result(); |
|
240 |
||
241 |
(*For the case hyps(p,t)-cons(#v => Fls,Y) |- p; |
|
242 |
we also have hyps(p,t)-{#v=>Fls} <= hyps(p, cons(v,t)) *) |
|
243 |
val [major] = goal PropThms.thy |
|
244 |
"p: prop ==> hyps(p, cons(v,t)) <= cons(#v, hyps(p,t)-{#v=>Fls})"; |
|
245 |
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
|
246 |
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
|
247 |
by (asm_simp_tac (prop_ss setloop (split_tac [expand_if])) 1); |
0 | 248 |
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
|
249 |
by (asm_simp_tac prop_ss 1); |
0 | 250 |
by (fast_tac ZF_cs 1); |
251 |
val hyps_cons = result(); |
|
252 |
||
253 |
(** Two lemmas for use with weaken_left **) |
|
254 |
||
255 |
goal ZF.thy "B-C <= cons(a, B-cons(a,C))"; |
|
256 |
by (fast_tac ZF_cs 1); |
|
257 |
val cons_Diff_same = result(); |
|
258 |
||
259 |
goal ZF.thy "cons(a, B-{c}) - D <= cons(a, B-cons(c,D))"; |
|
260 |
by (fast_tac ZF_cs 1); |
|
261 |
val cons_Diff_subset2 = result(); |
|
262 |
||
263 |
(*The set hyps(p,t) is finite, and elements have the form #v or #v=>Fls; |
|
264 |
could probably prove the stronger hyps(p,t) : Fin(hyps(p,0) Un hyps(p,nat))*) |
|
265 |
val [major] = goal PropThms.thy |
|
266 |
"p: prop ==> hyps(p,t) : Fin(UN v:nat. {#v, #v=>Fls})"; |
|
267 |
by (rtac (major RS Prop.induct) 1); |
|
16
0b033d50ca1c
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
7
diff
changeset
|
268 |
by (asm_simp_tac (prop_ss addsimps [Fin_0I, Fin_consI, 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
|
269 |
setloop (split_tac [expand_if])) 2); |
268f93ab3bc4
Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents:
0
diff
changeset
|
270 |
by (ALLGOALS (asm_simp_tac (prop_ss addsimps [Un_0, Fin_0I, Fin_UnI]))); |
0 | 271 |
val hyps_finite = result(); |
272 |
||
273 |
val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left; |
|
274 |
||
275 |
(*Induction on the finite set of assumptions hyps(p,t0). |
|
276 |
We may repeatedly subtract assumptions until none are left!*) |
|
277 |
val [premp,sat] = goal PropThms.thy |
|
278 |
"[| p: prop; 0 |= p |] ==> ALL t. hyps(p,t) - hyps(p,t0) |- p"; |
|
279 |
by (rtac (premp RS hyps_finite RS Fin_induct) 1); |
|
7
268f93ab3bc4
Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents:
0
diff
changeset
|
280 |
by (simp_tac (prop_ss addsimps [premp, sat, sat_thms_p, Diff_0]) 1); |
0 | 281 |
by (safe_tac ZF_cs); |
282 |
(*Case hyps(p,t)-cons(#v,Y) |- p *) |
|
283 |
by (rtac thms_excluded_middle_rule 1); |
|
284 |
by (etac VarI 3); |
|
285 |
by (rtac (cons_Diff_same RS weaken_left) 1); |
|
286 |
by (etac spec 1); |
|
287 |
by (rtac (cons_Diff_subset2 RS weaken_left) 1); |
|
288 |
by (rtac (premp RS hyps_Diff RS Diff_weaken_left) 1); |
|
289 |
by (etac spec 1); |
|
290 |
(*Case hyps(p,t)-cons(#v => Fls,Y) |- p *) |
|
291 |
by (rtac thms_excluded_middle_rule 1); |
|
292 |
by (etac VarI 3); |
|
293 |
by (rtac (cons_Diff_same RS weaken_left) 2); |
|
294 |
by (etac spec 2); |
|
295 |
by (rtac (cons_Diff_subset2 RS weaken_left) 1); |
|
296 |
by (rtac (premp RS hyps_cons RS Diff_weaken_left) 1); |
|
297 |
by (etac spec 1); |
|
298 |
val completeness_0_lemma = result(); |
|
299 |
||
300 |
(*The base case for completeness*) |
|
301 |
val [premp,sat] = goal PropThms.thy "[| p: prop; 0 |= p |] ==> 0 |- p"; |
|
302 |
by (rtac (Diff_cancel RS subst) 1); |
|
303 |
by (rtac (sat RS (premp RS completeness_0_lemma RS spec)) 1); |
|
304 |
val completeness_0 = result(); |
|
305 |
||
306 |
(*A semantic analogue of the Deduction Theorem*) |
|
307 |
goalw PropThms.thy [sat_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
|
308 |
by (simp_tac prop_ss 1); |
0 | 309 |
by (fast_tac ZF_cs 1); |
310 |
val sat_Imp = result(); |
|
311 |
||
312 |
goal PropThms.thy "!!H. H: Fin(prop) ==> ALL p:prop. H |= p --> H |- p"; |
|
313 |
by (etac Fin_induct 1); |
|
314 |
by (safe_tac (ZF_cs addSIs [completeness_0])); |
|
315 |
by (rtac (weaken_left_cons RS thms_MP) 1); |
|
316 |
by (fast_tac (ZF_cs addSIs [sat_Imp,ImpI]) 1); |
|
317 |
by (fast_tac thms_cs 1); |
|
318 |
val completeness_lemma = result(); |
|
319 |
||
320 |
val completeness = completeness_lemma RS bspec RS mp; |
|
321 |
||
322 |
val [finite] = goal PropThms.thy "H: Fin(prop) ==> H |- p <-> H |= p & p:prop"; |
|
323 |
by (fast_tac (ZF_cs addSEs [soundness, finite RS completeness, |
|
324 |
thms_in_pl]) 1); |
|
325 |
val thms_iff = result(); |
|
326 |
||
327 |
writeln"Reached end of file."; |