134 (* [| H |- p->false; H |- p; q: pl |] ==> H |- q *) |
133 (* [| H |- p->false; H |- p; q: pl |] ==> H |- q *) |
135 val conseq_notE = standard (conseq_MP RS conseq_falseE); |
134 val conseq_notE = standard (conseq_MP RS conseq_falseE); |
136 |
135 |
137 (** The function eval **) |
136 (** The function eval **) |
138 |
137 |
139 val pl_ss = set_ss addsimps [pl_rec_var,pl_rec_false,pl_rec_imp]; |
138 val pl_ss = set_ss addsimps [pl_rec_var,pl_rec_false,pl_rec_imp] |
|
139 @ PL0.inject @ PL0.ineq; |
140 |
140 |
141 goalw PL.thy [eval_def] "tt[false] = False"; |
141 goalw PL.thy [eval_def] "tt[false] = False"; |
142 by (simp_tac pl_ss 1); |
142 by (simp_tac pl_ss 1); |
143 val eval_false = result(); |
143 val eval_false = result(); |
144 |
144 |
211 val imp_false = result(); |
211 val imp_false = result(); |
212 |
212 |
213 (*This formulation is required for strong induction hypotheses*) |
213 (*This formulation is required for strong induction hypotheses*) |
214 goal PL.thy "hyps(p,tt) |- if(tt[p], p, p->false)"; |
214 goal PL.thy "hyps(p,tt) |- if(tt[p], p, p->false)"; |
215 by (rtac (expand_if RS iffD2) 1); |
215 by (rtac (expand_if RS iffD2) 1); |
216 by(res_inst_tac[("x","p")]spec 1); |
216 by(PL0.induct_tac "p" 1); |
217 by (rtac pl_ind 1); |
|
218 by (ALLGOALS (simp_tac (pl_ss addsimps [conseq_I, conseq_H]))); |
217 by (ALLGOALS (simp_tac (pl_ss addsimps [conseq_I, conseq_H]))); |
219 by (fast_tac (set_cs addIs [weaken_left_Un1, weaken_left_Un2, |
218 by (fast_tac (set_cs addIs [weaken_left_Un1, weaken_left_Un2, |
220 weaken_right, imp_false] |
219 weaken_right, imp_false] |
221 addSEs [false_imp]) 1); |
220 addSEs [false_imp]) 1); |
222 val hyps_conseq_if = result(); |
221 val hyps_conseq_if = result(); |
248 |
247 |
249 (*** Completeness -- lemmas for reducing the set of assumptions ***) |
248 (*** Completeness -- lemmas for reducing the set of assumptions ***) |
250 |
249 |
251 (*For the case hyps(p,t)-insert(#v,Y) |- p; |
250 (*For the case hyps(p,t)-insert(#v,Y) |- p; |
252 we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *) |
251 we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *) |
253 goal PL.thy "!p.hyps(p, t-{v}) <= insert((#v)->false, hyps(p,t)-{#v})"; |
252 goal PL.thy "hyps(p, t-{v}) <= insert((#v)->false, hyps(p,t)-{#v})"; |
254 by (rtac pl_ind 1); |
253 by (PL0.induct_tac "p" 1); |
255 by (simp_tac pl_ss 1); |
254 by (simp_tac pl_ss 1); |
256 by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1); |
255 by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1); |
257 by (fast_tac (set_cs addSEs [sym RS var_neq_imp] addSDs [var_inject]) 1); |
256 by (simp_tac pl_ss 1); |
258 by (simp_tac pl_ss 1); |
257 by (fast_tac set_cs 1); |
259 by (fast_tac set_cs 1); |
258 val hyps_Diff = result(); |
260 val hyps_Diff = result() RS spec; |
|
261 |
259 |
262 (*For the case hyps(p,t)-insert(#v -> false,Y) |- p; |
260 (*For the case hyps(p,t)-insert(#v -> false,Y) |- p; |
263 we also have hyps(p,t)-{(#v)->false} <= hyps(p, insert(v,t)) *) |
261 we also have hyps(p,t)-{(#v)->false} <= hyps(p, insert(v,t)) *) |
264 goal PL.thy "!p.hyps(p, insert(v,t)) <= insert(#v, hyps(p,t)-{(#v)->false})"; |
262 goal PL.thy "hyps(p, insert(v,t)) <= insert(#v, hyps(p,t)-{(#v)->false})"; |
265 by (rtac pl_ind 1); |
263 by (PL0.induct_tac "p" 1); |
266 by (simp_tac pl_ss 1); |
264 by (simp_tac pl_ss 1); |
267 by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1); |
265 by (simp_tac (pl_ss addsimps [insert_subset] |
268 by (fast_tac (set_cs addSEs [var_neq_imp, imp_inject] addSDs [var_inject]) 1); |
266 setloop (split_tac [expand_if])) 1); |
269 by (simp_tac pl_ss 1); |
267 by (simp_tac pl_ss 1); |
270 by (fast_tac set_cs 1); |
268 by (fast_tac set_cs 1); |
271 val hyps_insert = result() RS spec; |
269 val hyps_insert = result(); |
272 |
270 |
273 (** Two lemmas for use with weaken_left **) |
271 (** Two lemmas for use with weaken_left **) |
274 |
272 |
275 goal Set.thy "B-C <= insert(a, B-insert(a,C))"; |
273 goal Set.thy "B-C <= insert(a, B-insert(a,C))"; |
276 by (fast_tac set_cs 1); |
274 by (fast_tac set_cs 1); |
280 by (fast_tac set_cs 1); |
278 by (fast_tac set_cs 1); |
281 val insert_Diff_subset2 = result(); |
279 val insert_Diff_subset2 = result(); |
282 |
280 |
283 (*The set hyps(p,t) is finite, and elements have the form #v or #v->false; |
281 (*The set hyps(p,t) is finite, and elements have the form #v or #v->false; |
284 could probably prove the stronger hyps(p,t) : Fin(hyps(p,{}) Un hyps(p,nat))*) |
282 could probably prove the stronger hyps(p,t) : Fin(hyps(p,{}) Un hyps(p,nat))*) |
285 goal PL.thy "!p.hyps(p,t) : Fin(UN v:{x.True}. {#v, (#v)->false})"; |
283 goal PL.thy "hyps(p,t) : Fin(UN v:{x.True}. {#v, (#v)->false})"; |
286 by (rtac pl_ind 1); |
284 by (PL0.induct_tac "p" 1); |
287 by (ALLGOALS (simp_tac (pl_ss setloop (split_tac [expand_if])) THEN' |
285 by (ALLGOALS (simp_tac (pl_ss setloop (split_tac [expand_if])) THEN' |
288 fast_tac (set_cs addSIs [Fin_0I, Fin_insertI, Fin_UnI]))); |
286 fast_tac (set_cs addSIs [Fin_0I, Fin_insertI, Fin_UnI]))); |
289 val hyps_finite = result() RS spec; |
287 val hyps_finite = result(); |
290 |
288 |
291 val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left; |
289 val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left; |
292 |
290 |
293 (*Induction on the finite set of assumptions hyps(p,t0). |
291 (*Induction on the finite set of assumptions hyps(p,t0). |
294 We may repeatedly subtract assumptions until none are left!*) |
292 We may repeatedly subtract assumptions until none are left!*) |