ex/PL.ML
changeset 93 8c9be2e9236d
parent 81 fded09018308
child 102 18d44ab74672
equal deleted inserted replaced
92:bcd0ee8d71aa 93:8c9be2e9236d
   133 (* [| H |- p->false;  H |- p;  q: pl |] ==> H |- q *)
   133 (* [| H |- p->false;  H |- p;  q: pl |] ==> H |- q *)
   134 val conseq_notE = standard (conseq_MP RS conseq_falseE);
   134 val conseq_notE = standard (conseq_MP RS conseq_falseE);
   135 
   135 
   136 (** The function eval **)
   136 (** The function eval **)
   137 
   137 
   138 val pl_ss = set_ss addsimps [pl_rec_var,pl_rec_false,pl_rec_imp] @ PL0.pl.simps;
   138 val pl_ss = set_ss addsimps [pl_rec_var,pl_rec_false,pl_rec_imp] @ PL.pl.simps;
   139 
   139 
   140 goalw PL.thy [eval_def] "tt[false] = False";
   140 goalw PL.thy [eval_def] "tt[false] = False";
   141 by (simp_tac pl_ss 1);
   141 by (simp_tac pl_ss 1);
   142 val eval_false = result();
   142 val eval_false = result();
   143 
   143 
   155 
   155 
   156 goalw PL.thy [hyps_def] "hyps(false,tt) = {}";
   156 goalw PL.thy [hyps_def] "hyps(false,tt) = {}";
   157 by (simp_tac pl_ss 1);
   157 by (simp_tac pl_ss 1);
   158 val hyps_false = result();
   158 val hyps_false = result();
   159 
   159 
   160 goalw PL.thy [hyps_def] "hyps(#v,tt) = {if(v:tt, #v, (#v)->false)}";
   160 goalw PL.thy [hyps_def] "hyps(#v,tt) = {if(v:tt, #v, #v->false)}";
   161 by (simp_tac pl_ss 1);
   161 by (simp_tac pl_ss 1);
   162 val hyps_var = result();
   162 val hyps_var = result();
   163 
   163 
   164 goalw PL.thy [hyps_def] "hyps(p->q,tt) = hyps(p,tt) Un hyps(q,tt)";
   164 goalw PL.thy [hyps_def] "hyps(p->q,tt) = hyps(p,tt) Un hyps(q,tt)";
   165 by (simp_tac pl_ss 1);
   165 by (simp_tac pl_ss 1);
   210 val imp_false = result();
   210 val imp_false = result();
   211 
   211 
   212 (*This formulation is required for strong induction hypotheses*)
   212 (*This formulation is required for strong induction hypotheses*)
   213 goal PL.thy "hyps(p,tt) |- if(tt[p], p, p->false)";
   213 goal PL.thy "hyps(p,tt) |- if(tt[p], p, p->false)";
   214 by (rtac (expand_if RS iffD2) 1);
   214 by (rtac (expand_if RS iffD2) 1);
   215 by(PL0.pl.induct_tac "p" 1);
   215 by(PL.pl.induct_tac "p" 1);
   216 by (ALLGOALS (simp_tac (pl_ss addsimps [conseq_I, conseq_H])));
   216 by (ALLGOALS (simp_tac (pl_ss addsimps [conseq_I, conseq_H])));
   217 by (fast_tac (set_cs addIs [weaken_left_Un1, weaken_left_Un2, 
   217 by (fast_tac (set_cs addIs [weaken_left_Un1, weaken_left_Un2, 
   218 			   weaken_right, imp_false]
   218 			   weaken_right, imp_false]
   219                     addSEs [false_imp]) 1);
   219                     addSEs [false_imp]) 1);
   220 val hyps_conseq_if = result();
   220 val hyps_conseq_if = result();
   246 
   246 
   247 (*** Completeness -- lemmas for reducing the set of assumptions ***)
   247 (*** Completeness -- lemmas for reducing the set of assumptions ***)
   248 
   248 
   249 (*For the case hyps(p,t)-insert(#v,Y) |- p;
   249 (*For the case hyps(p,t)-insert(#v,Y) |- p;
   250   we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *)
   250   we also have hyps(p,t)-{#v} <= hyps(p, t-{v}) *)
   251 goal PL.thy "hyps(p, t-{v}) <= insert((#v)->false, hyps(p,t)-{#v})";
   251 goal PL.thy "hyps(p, t-{v}) <= insert(#v->false, hyps(p,t)-{#v})";
   252 by (PL0.pl.induct_tac "p" 1);
   252 by (PL.pl.induct_tac "p" 1);
   253 by (simp_tac pl_ss 1);
   253 by (simp_tac pl_ss 1);
   254 by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1);
   254 by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1);
   255 by (simp_tac pl_ss 1);
   255 by (simp_tac pl_ss 1);
   256 by (fast_tac set_cs 1);
   256 by (fast_tac set_cs 1);
   257 val hyps_Diff = result();
   257 val hyps_Diff = result();
   258 
   258 
   259 (*For the case hyps(p,t)-insert(#v -> false,Y) |- p;
   259 (*For the case hyps(p,t)-insert(#v -> false,Y) |- p;
   260   we also have hyps(p,t)-{(#v)->false} <= hyps(p, insert(v,t)) *)
   260   we also have hyps(p,t)-{#v->false} <= hyps(p, insert(v,t)) *)
   261 goal PL.thy "hyps(p, insert(v,t)) <= insert(#v, hyps(p,t)-{(#v)->false})";
   261 goal PL.thy "hyps(p, insert(v,t)) <= insert(#v, hyps(p,t)-{#v->false})";
   262 by (PL0.pl.induct_tac "p" 1);
   262 by (PL.pl.induct_tac "p" 1);
   263 by (simp_tac pl_ss 1);
   263 by (simp_tac pl_ss 1);
   264 by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1);
   264 by (simp_tac (pl_ss setloop (split_tac [expand_if])) 1);
   265 by (simp_tac pl_ss 1);
   265 by (simp_tac pl_ss 1);
   266 by (fast_tac set_cs 1);
   266 by (fast_tac set_cs 1);
   267 val hyps_insert = result();
   267 val hyps_insert = result();
   276 by (fast_tac set_cs 1);
   276 by (fast_tac set_cs 1);
   277 val insert_Diff_subset2 = result();
   277 val insert_Diff_subset2 = result();
   278 
   278 
   279 (*The set hyps(p,t) is finite, and elements have the form #v or #v->false;
   279 (*The set hyps(p,t) is finite, and elements have the form #v or #v->false;
   280  could probably prove the stronger hyps(p,t) : Fin(hyps(p,{}) Un hyps(p,nat))*)
   280  could probably prove the stronger hyps(p,t) : Fin(hyps(p,{}) Un hyps(p,nat))*)
   281 goal PL.thy "hyps(p,t) : Fin(UN v:{x.True}. {#v, (#v)->false})";
   281 goal PL.thy "hyps(p,t) : Fin(UN v:{x.True}. {#v, #v->false})";
   282 by (PL0.pl.induct_tac "p" 1);
   282 by (PL.pl.induct_tac "p" 1);
   283 by (ALLGOALS (simp_tac (pl_ss setloop (split_tac [expand_if])) THEN'
   283 by (ALLGOALS (simp_tac (pl_ss setloop (split_tac [expand_if])) THEN'
   284               fast_tac (set_cs addSIs [Fin_0I, Fin_insertI, Fin_UnI])));
   284               fast_tac (set_cs addSIs [Fin_0I, Fin_insertI, Fin_UnI])));
   285 val hyps_finite = result();
   285 val hyps_finite = result();
   286 
   286 
   287 val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left;
   287 val Diff_weaken_left = subset_refl RSN (2, Diff_mono) RS weaken_left;