156 |
156 |
157 goalw PL.thy [hyps_def] "hyps(false,tt) = {}"; |
157 goalw PL.thy [hyps_def] "hyps(false,tt) = {}"; |
158 by (simp_tac pl_ss 1); |
158 by (simp_tac pl_ss 1); |
159 val hyps_false = result(); |
159 val hyps_false = result(); |
160 |
160 |
161 goalw PL.thy [hyps_def] "hyps(#v,tt) = {if(v:tt, #v, #v->false)}"; |
161 goalw PL.thy [hyps_def] "hyps(#v,tt) = {if(v:tt, #v, (#v)->false)}"; |
162 by (simp_tac pl_ss 1); |
162 by (simp_tac pl_ss 1); |
163 val hyps_var = result(); |
163 val hyps_var = result(); |
164 |
164 |
165 goalw PL.thy [hyps_def] "hyps(p->q,tt) = hyps(p,tt) Un hyps(q,tt)"; |
165 goalw PL.thy [hyps_def] "hyps(p->q,tt) = hyps(p,tt) Un hyps(q,tt)"; |
166 by (simp_tac pl_ss 1); |
166 by (simp_tac pl_ss 1); |
248 |
248 |
249 (*** Completeness -- lemmas for reducing the set of assumptions ***) |
249 (*** Completeness -- lemmas for reducing the set of assumptions ***) |
250 |
250 |
251 (*For the case hyps(p,t)-insert(#v,Y) |- p; |
251 (*For the case hyps(p,t)-insert(#v,Y) |- p; |
252 we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *) |
252 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})"; |
253 goal PL.thy "!p.hyps(p, t-{v}) <= insert((#v)->false, hyps(p,t)-{#v})"; |
254 by (rtac pl_ind 1); |
254 by (rtac pl_ind 1); |
255 by (simp_tac pl_ss 1); |
255 by (simp_tac pl_ss 1); |
256 by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1); |
256 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); |
257 by (fast_tac (set_cs addSEs [sym RS var_neq_imp] addSDs [var_inject]) 1); |
258 by (simp_tac pl_ss 1); |
258 by (simp_tac pl_ss 1); |
259 by (fast_tac set_cs 1); |
259 by (fast_tac set_cs 1); |
260 val hyps_Diff = result() RS spec; |
260 val hyps_Diff = result() RS spec; |
261 |
261 |
262 (*For the case hyps(p,t)-insert(#v -> false,Y) |- p; |
262 (*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)) *) |
263 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})"; |
264 goal PL.thy "!p.hyps(p, insert(v,t)) <= insert(#v, hyps(p,t)-{(#v)->false})"; |
265 by (rtac pl_ind 1); |
265 by (rtac pl_ind 1); |
266 by (simp_tac pl_ss 1); |
266 by (simp_tac pl_ss 1); |
267 by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1); |
267 by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1); |
268 by (fast_tac (set_cs addSEs [var_neq_imp, imp_inject] addSDs [var_inject]) 1); |
268 by (fast_tac (set_cs addSEs [var_neq_imp, imp_inject] addSDs [var_inject]) 1); |
269 by (simp_tac pl_ss 1); |
269 by (simp_tac pl_ss 1); |
280 by (fast_tac set_cs 1); |
280 by (fast_tac set_cs 1); |
281 val insert_Diff_subset2 = result(); |
281 val insert_Diff_subset2 = result(); |
282 |
282 |
283 (*The set hyps(p,t) is finite, and elements have the form #v or #v->false; |
283 (*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))*) |
284 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})"; |
285 goal PL.thy "!p.hyps(p,t) : Fin(UN v:{x.True}. {#v, (#v)->false})"; |
286 by (rtac pl_ind 1); |
286 by (rtac pl_ind 1); |
287 by (ALLGOALS (simp_tac (pl_ss setloop (split_tac [expand_if])) THEN' |
287 by (ALLGOALS (simp_tac (pl_ss setloop (split_tac [expand_if])) THEN' |
288 fast_tac (set_cs addSIs [Fin_0I, Fin_insertI, Fin_UnI]))); |
288 fast_tac (set_cs addSIs [Fin_0I, Fin_insertI, Fin_UnI]))); |
289 val hyps_finite = result() RS spec; |
289 val hyps_finite = result() RS spec; |
290 |
290 |