diff --git a/thys/Jinja/Compiler/Correctness1.thy b/thys/Jinja/Compiler/Correctness1.thy --- a/thys/Jinja/Compiler/Correctness1.thy +++ b/thys/Jinja/Compiler/Correctness1.thy @@ -1,760 +1,761 @@ (* Title: Jinja/Compiler/Correctness1.thy Author: Tobias Nipkow Copyright TUM 2003 *) section \Correctness of Stage 1\ theory Correctness1 imports J1WellForm Compiler1 begin subsection\Correctness of program compilation\ primrec unmod :: "expr\<^sub>1 \ nat \ bool" and unmods :: "expr\<^sub>1 list \ nat \ bool" where "unmod (new C) i = True" | "unmod (Cast C e) i = unmod e i" | "unmod (Val v) i = True" | "unmod (e\<^sub>1 \bop\ e\<^sub>2) i = (unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (Var i) j = True" | "unmod (i:=e) j = (i \ j \ unmod e j)" | "unmod (e\F{D}) i = unmod e i" | "unmod (e\<^sub>1\F{D}:=e\<^sub>2) i = (unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (e\M(es)) i = (unmod e i \ unmods es i)" | "unmod {j:T; e} i = unmod e i" | "unmod (e\<^sub>1;;e\<^sub>2) i = (unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (if (e) e\<^sub>1 else e\<^sub>2) i = (unmod e i \ unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (while (e) c) i = (unmod e i \ unmod c i)" | "unmod (throw e) i = unmod e i" | "unmod (try e\<^sub>1 catch(C i) e\<^sub>2) j = (unmod e\<^sub>1 j \ (if i=j then False else unmod e\<^sub>2 j))" | "unmods ([]) i = True" | "unmods (e#es) i = (unmod e i \ unmods es i)" lemma hidden_unmod: "\Vs. hidden Vs i \ unmod (compE\<^sub>1 Vs e) i" and "\Vs. hidden Vs i \ unmods (compEs\<^sub>1 Vs es) i" (*<*) -apply(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) -apply (simp_all add:hidden_inacc) -apply(auto simp add:hidden_def) -done +proof(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) + case TryCatch + then show ?case by(simp add:hidden_inacc)(auto simp add:hidden_def) +qed (simp_all add:hidden_inacc) (*>*) lemma eval\<^sub>1_preserves_unmod: "\ P \\<^sub>1 \e,(h,ls)\ \ \e',(h',ls')\; unmod e i; i < size ls \ \ ls ! i = ls' ! i" and "\ P \\<^sub>1 \es,(h,ls)\ [\] \es',(h',ls')\; unmods es i; i < size ls \ \ ls ! i = ls' ! i" (*<*) -apply(induct rule:eval\<^sub>1_evals\<^sub>1_inducts) -apply(auto dest!:eval\<^sub>1_preserves_len split:if_split_asm) -done +proof(induct rule:eval\<^sub>1_evals\<^sub>1_inducts) +qed(auto dest!:eval\<^sub>1_preserves_len evals\<^sub>1_preserves_len split:if_split_asm) (*>*) lemma LAss_lem: "\x \ set xs; size xs \ size ys \ \ m\<^sub>1 \\<^sub>m m\<^sub>2(xs[\]ys) \ m\<^sub>1(x\y) \\<^sub>m m\<^sub>2(xs[\]ys[last_index xs x := y])" (*<*) by(simp add:map_le_def fun_upds_apply eq_sym_conv) (*>*) lemma Block_lem: fixes l :: "'a \ 'b" assumes 0: "l \\<^sub>m [Vs [\] ls]" and 1: "l' \\<^sub>m [Vs [\] ls', V\v]" and hidden: "V \ set Vs \ ls ! last_index Vs V = ls' ! last_index Vs V" and size: "size ls = size ls'" "size Vs < size ls'" shows "l'(V := l V) \\<^sub>m [Vs [\] ls']" (*<*) proof - have "l'(V := l V) \\<^sub>m [Vs [\] ls', V\v](V := l V)" using 1 by(rule map_le_upd) also have "\ = [Vs [\] ls'](V := l V)" by simp also have "\ \\<^sub>m [Vs [\] ls']" proof (cases "l V") case None thus ?thesis by simp next case (Some w) hence "[Vs [\] ls] V = Some w" using 0 by(force simp add: map_le_def split:if_splits) hence VinVs: "V \ set Vs" and w: "w = ls ! last_index Vs V" using size by(auto simp add:fun_upds_apply split:if_splits) hence "w = ls' ! last_index Vs V" using hidden[OF VinVs] by simp hence "[Vs [\] ls'](V := l V) = [Vs [\] ls']" using Some size VinVs by(simp add: map_upds_upd_conv_last_index) thus ?thesis by simp qed finally show ?thesis . qed (*>*) (*<*) declare fun_upd_apply[simp del] (*>*) text\\noindent The main theorem:\ theorem assumes wf: "wwf_J_prog P" shows eval\<^sub>1_eval: "P \ \e,(h,l)\ \ \e',(h',l')\ \ (\Vs ls. \ fv e \ set Vs; l \\<^sub>m [Vs[\]ls]; size Vs + max_vars e \ size ls \ \ \ls'. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs e,(h,ls)\ \ \fin\<^sub>1 e',(h',ls')\ \ l' \\<^sub>m [Vs[\]ls'])" (*<*) (is "_ \ (\Vs ls. PROP ?P e h l e' h' l' Vs ls)" is "_ \ (\Vs ls. \ _; _; _ \ \ \ls'. ?Post e h l e' h' l' Vs ls ls')") (*>*) and evals\<^sub>1_evals: "P \ \es,(h,l)\ [\] \es',(h',l')\ \ (\Vs ls. \ fvs es \ set Vs; l \\<^sub>m [Vs[\]ls]; size Vs + max_varss es \ size ls \ \ \ls'. compP\<^sub>1 P \\<^sub>1 \compEs\<^sub>1 Vs es,(h,ls)\ [\] \compEs\<^sub>1 Vs es',(h',ls')\ \ l' \\<^sub>m [Vs[\]ls'])" (*<*) (is "_ \ (\Vs ls. PROP ?Ps es h l es' h' l' Vs ls)" is "_ \ (\Vs ls. \ _; _; _\ \ \ls'. ?Posts es h l es' h' l' Vs ls ls')") proof (induct rule:eval_evals_inducts) case Nil thus ?case by(fastforce intro!:Nil\<^sub>1) next case (Cons e h l v h' l' es es' h\<^sub>2 l\<^sub>2) have "PROP ?P e h l (Val v) h' l' Vs ls" by fact with Cons.prems obtain ls' where 1: "?Post e h l (Val v) h' l' Vs ls ls'" "size ls = size ls'" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h' l' es' h\<^sub>2 l\<^sub>2 Vs ls'" by fact with 1 Cons.prems obtain ls\<^sub>2 where 2: "?Posts es h' l' es' h\<^sub>2 l\<^sub>2 Vs ls' ls\<^sub>2" by(auto) from 1 2 Cons show ?case by(auto intro!:Cons\<^sub>1) next case ConsThrow thus ?case by(fastforce intro!:ConsThrow\<^sub>1 dest: eval_final) next case (Block e h l V e' h' l' T) let ?Vs = "Vs @ [V]" have IH: "\fv e \ set ?Vs; l(V := None) \\<^sub>m [?Vs [\] ls]; size ?Vs + max_vars e \ size ls\ \ \ls'. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e,(h,ls)\ \ \fin\<^sub>1 e',(h', ls')\ \ l' \\<^sub>m [?Vs [\] ls']" and fv: "fv {V:T; e} \ set Vs" and rel: "l \\<^sub>m [Vs [\] ls]" and len: "length Vs + max_vars {V:T; e} \ length ls" by fact+ have len': "length Vs < length ls" using len by auto have "fv e \ set ?Vs" using fv by auto moreover have "l(V := None) \\<^sub>m [?Vs [\] ls]" using rel len' by simp moreover have "size ?Vs + max_vars e \ size ls" using len by simp ultimately obtain ls' where 1: "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e,(h,ls)\ \ \fin\<^sub>1 e',(h',ls')\" and rel': "l' \\<^sub>m [?Vs [\] ls']" using IH by blast have [simp]: "length ls = length ls'" by(rule eval\<^sub>1_preserves_len[OF 1]) show "\ls'. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs {V:T; e},(h,ls)\ \ \fin\<^sub>1 e',(h',ls')\ \ l'(V := l V) \\<^sub>m [Vs [\] ls']" (is "\ls'. ?R ls'") proof show "?R ls'" proof show "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs {V:T; e},(h,ls)\ \ \fin\<^sub>1 e',(h',ls')\" using 1 by(simp add:Block\<^sub>1) next show "l'(V := l V) \\<^sub>m [Vs [\] ls']" proof - have "l' \\<^sub>m [Vs [\] ls', V \ ls' ! length Vs]" using len' rel' by simp moreover { assume VinVs: "V \ set Vs" hence "hidden (Vs @ [V]) (last_index Vs V)" by(rule hidden_last_index) hence "unmod (compE\<^sub>1 (Vs @ [V]) e) (last_index Vs V)" by(rule hidden_unmod) moreover have "last_index Vs V < length ls" using len' VinVs by simp ultimately have "ls ! last_index Vs V = ls' ! last_index Vs V" by(rule eval\<^sub>1_preserves_unmod[OF 1]) } ultimately show ?thesis using Block_lem[OF rel] len' by auto qed qed qed next case (TryThrow e' h l a h' l' D fs C V e\<^sub>2) have "PROP ?P e' h l (Throw a) h' l' Vs ls" by fact with TryThrow.prems obtain ls' where 1: "?Post e' h l (Throw a) h' l' Vs ls ls'" by(auto) show ?case using 1 TryThrow.hyps by(auto intro!:eval\<^sub>1_evals\<^sub>1.TryThrow\<^sub>1) next case (TryCatch e\<^sub>1 h l a h\<^sub>1 l\<^sub>1 D fs C e\<^sub>2 V e' h\<^sub>2 l\<^sub>2) let ?e = "try e\<^sub>1 catch(C V) e\<^sub>2" have IH\<^sub>1: "\fv e\<^sub>1 \ set Vs; l \\<^sub>m [Vs [\] ls]; size Vs + max_vars e\<^sub>1 \ length ls\ \ \ls\<^sub>1. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs e\<^sub>1,(h,ls)\ \ \fin\<^sub>1 (Throw a),(h\<^sub>1,ls\<^sub>1)\ \ l\<^sub>1 \\<^sub>m [Vs [\] ls\<^sub>1]" and fv: "fv ?e \ set Vs" and rel: "l \\<^sub>m [Vs [\] ls]" and len: "length Vs + max_vars ?e \ length ls" by fact+ have "fv e\<^sub>1 \ set Vs" using fv by auto moreover have "length Vs + max_vars e\<^sub>1 \ length ls" using len by(auto) ultimately obtain ls\<^sub>1 where 1: "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs e\<^sub>1,(h,ls)\ \ \Throw a,(h\<^sub>1,ls\<^sub>1)\" and rel\<^sub>1: "l\<^sub>1 \\<^sub>m [Vs [\] ls\<^sub>1]" using IH\<^sub>1 rel by fastforce from 1 have [simp]: "size ls = size ls\<^sub>1" by(rule eval\<^sub>1_preserves_len) let ?Vs = "Vs @ [V]" let ?ls = "(ls\<^sub>1[size Vs:=Addr a])" have IH\<^sub>2: "\fv e\<^sub>2 \ set ?Vs; l\<^sub>1(V \ Addr a) \\<^sub>m [?Vs [\] ?ls]; length ?Vs + max_vars e\<^sub>2 \ length ?ls\ \ \ls\<^sub>2. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e\<^sub>2,(h\<^sub>1,?ls)\ \ \fin\<^sub>1 e',(h\<^sub>2, ls\<^sub>2)\ \ l\<^sub>2 \\<^sub>m [?Vs [\] ls\<^sub>2]" by fact have len\<^sub>1: "size Vs < size ls\<^sub>1" using len by(auto) have "fv e\<^sub>2 \ set ?Vs" using fv by auto moreover have "l\<^sub>1(V \ Addr a) \\<^sub>m [?Vs [\] ?ls]" using rel\<^sub>1 len\<^sub>1 by simp moreover have "length ?Vs + max_vars e\<^sub>2 \ length ?ls" using len by(auto) ultimately obtain ls\<^sub>2 where 2: "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e\<^sub>2,(h\<^sub>1,?ls)\ \ \fin\<^sub>1 e',(h\<^sub>2, ls\<^sub>2)\" and rel\<^sub>2: "l\<^sub>2 \\<^sub>m [?Vs [\] ls\<^sub>2]" using IH\<^sub>2 by blast from 2 have [simp]: "size ls\<^sub>1 = size ls\<^sub>2" by(fastforce dest: eval\<^sub>1_preserves_len) show "\ls\<^sub>2. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs ?e,(h,ls)\ \ \fin\<^sub>1 e',(h\<^sub>2,ls\<^sub>2)\ \ l\<^sub>2(V := l\<^sub>1 V) \\<^sub>m [Vs [\] ls\<^sub>2]" (is "\ls\<^sub>2. ?R ls\<^sub>2") proof show "?R ls\<^sub>2" proof have hp: "h\<^sub>1 a = Some (D, fs)" by fact have "P \ D \\<^sup>* C" by fact hence caught: "compP\<^sub>1 P \ D \\<^sup>* C" by simp from TryCatch\<^sub>1[OF 1 _ caught len\<^sub>1 2, OF hp] show "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs ?e,(h,ls)\ \ \fin\<^sub>1 e',(h\<^sub>2,ls\<^sub>2)\" by simp next show "l\<^sub>2(V := l\<^sub>1 V) \\<^sub>m [Vs [\] ls\<^sub>2]" proof - have "l\<^sub>2 \\<^sub>m [Vs [\] ls\<^sub>2, V \ ls\<^sub>2 ! length Vs]" using len\<^sub>1 rel\<^sub>2 by simp moreover { assume VinVs: "V \ set Vs" hence "hidden (Vs @ [V]) (last_index Vs V)" by(rule hidden_last_index) hence "unmod (compE\<^sub>1 (Vs @ [V]) e\<^sub>2) (last_index Vs V)" by(rule hidden_unmod) moreover have "last_index Vs V < length ?ls" using len\<^sub>1 VinVs by simp ultimately have "?ls ! last_index Vs V = ls\<^sub>2 ! last_index Vs V" by(rule eval\<^sub>1_preserves_unmod[OF 2]) moreover have "last_index Vs V < size Vs" using VinVs by simp ultimately have "ls\<^sub>1 ! last_index Vs V = ls\<^sub>2 ! last_index Vs V" using len\<^sub>1 by(simp del:size_last_index_conv) } ultimately show ?thesis using Block_lem[OF rel\<^sub>1] len\<^sub>1 by simp qed qed qed next case Try thus ?case by(fastforce intro!:Try\<^sub>1) next case Throw thus ?case by(fastforce intro!:Throw\<^sub>1) next case ThrowNull thus ?case by(fastforce intro!:ThrowNull\<^sub>1) next case ThrowThrow thus ?case by(fastforce intro!:ThrowThrow\<^sub>1) next case (CondT e h l h\<^sub>1 l\<^sub>1 e\<^sub>1 e' h\<^sub>2 l\<^sub>2 e\<^sub>2) have "PROP ?P e h l true h\<^sub>1 l\<^sub>1 Vs ls" by fact with CondT.prems obtain ls\<^sub>1 where 1: "?Post e h l true h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 e' h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 CondT.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 e' h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 show ?case by(auto intro!:CondT\<^sub>1) next case (CondF e h l h\<^sub>1 l\<^sub>1 e\<^sub>2 e' h\<^sub>2 l\<^sub>2 e\<^sub>1 Vs ls) have "PROP ?P e h l false h\<^sub>1 l\<^sub>1 Vs ls" by fact with CondF.prems obtain ls\<^sub>1 where 1: "?Post e h l false h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 e' h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 CondF.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 e' h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 show ?case by(auto intro!:CondF\<^sub>1) next case CondThrow thus ?case by(fastforce intro!:CondThrow\<^sub>1) next case (Seq e h l v h\<^sub>1 l\<^sub>1 e\<^sub>1 e' h\<^sub>2 l\<^sub>2) have "PROP ?P e h l (Val v) h\<^sub>1 l\<^sub>1 Vs ls" by fact with Seq.prems obtain ls\<^sub>1 where 1: "?Post e h l (Val v) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 e' h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 Seq.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 e' h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 Seq show ?case by(auto intro!:Seq\<^sub>1) next case SeqThrow thus ?case by(fastforce intro!:SeqThrow\<^sub>1) next case WhileF thus ?case by(fastforce intro!:eval\<^sub>1_evals\<^sub>1.intros) next case (WhileT e h l h\<^sub>1 l\<^sub>1 c v h\<^sub>2 l\<^sub>2 e' h\<^sub>3 l\<^sub>3) have "PROP ?P e h l true h\<^sub>1 l\<^sub>1 Vs ls" by fact with WhileT.prems obtain ls\<^sub>1 where 1: "?Post e h l true h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P c h\<^sub>1 l\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 WhileT.prems obtain ls\<^sub>2 where 2: "?Post c h\<^sub>1 l\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" "size ls\<^sub>1 = size ls\<^sub>2" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P (While (e) c) h\<^sub>2 l\<^sub>2 e' h\<^sub>3 l\<^sub>3 Vs ls\<^sub>2" by fact with 1 2 WhileT.prems obtain ls\<^sub>3 where 3: "?Post (While (e) c) h\<^sub>2 l\<^sub>2 e' h\<^sub>3 l\<^sub>3 Vs ls\<^sub>2 ls\<^sub>3" by(auto) from 1 2 3 show ?case by(auto intro!:WhileT\<^sub>1) next case (WhileBodyThrow e h l h\<^sub>1 l\<^sub>1 c e' h\<^sub>2 l\<^sub>2) have "PROP ?P e h l true h\<^sub>1 l\<^sub>1 Vs ls" by fact with WhileBodyThrow.prems obtain ls\<^sub>1 where 1: "?Post e h l true h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P c h\<^sub>1 l\<^sub>1 (throw e') h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 WhileBodyThrow.prems obtain ls\<^sub>2 where 2: "?Post c h\<^sub>1 l\<^sub>1 (throw e') h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by auto from 1 2 show ?case by(auto intro!:WhileBodyThrow\<^sub>1) next case WhileCondThrow thus ?case by(fastforce intro!:WhileCondThrow\<^sub>1) next case New thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case NewFail thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case Cast thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case CastNull thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case CastThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (CastFail e h l a h\<^sub>1 l\<^sub>1 D fs C) have "PROP ?P e h l (addr a) h\<^sub>1 l\<^sub>1 Vs ls" by fact with CastFail.prems obtain ls\<^sub>1 where 1: "?Post e h l (addr a) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" by auto show ?case using 1 CastFail.hyps by(auto intro!:CastFail\<^sub>1[where D=D]) next case Val thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (BinOp e h l v\<^sub>1 h\<^sub>1 l\<^sub>1 e\<^sub>1 v\<^sub>2 h\<^sub>2 l\<^sub>2 bop v) have "PROP ?P e h l (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 Vs ls" by fact with BinOp.prems obtain ls\<^sub>1 where 1: "?Post e h l (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 (Val v\<^sub>2) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 BinOp.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 (Val v\<^sub>2) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 BinOp show ?case by(auto intro!:BinOp\<^sub>1) next case (BinOpThrow2 e\<^sub>0 h l v\<^sub>1 h\<^sub>1 l\<^sub>1 e\<^sub>1 e h\<^sub>2 l\<^sub>2 bop) have "PROP ?P e\<^sub>0 h l (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 Vs ls" by fact with BinOpThrow2.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>0 h l (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 BinOpThrow2.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 BinOpThrow2 show ?case by(auto intro!:BinOpThrow\<^sub>2\<^sub>1) next case BinOpThrow1 thus ?case by(fastforce intro!:eval\<^sub>1_evals\<^sub>1.intros) next case Var thus ?case by(force intro!:Var\<^sub>1 simp add: map_le_def fun_upds_apply) next case LAss thus ?case by(fastforce simp add: LAss_lem intro:eval\<^sub>1_evals\<^sub>1.intros dest:eval\<^sub>1_preserves_len) next case LAssThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case FAcc thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case FAccNull thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case FAccThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (FAss e\<^sub>1 h l a h\<^sub>1 l\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 C fs fs' F D h\<^sub>2') have "PROP ?P e\<^sub>1 h l (addr a) h\<^sub>1 l\<^sub>1 Vs ls" by fact with FAss.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l (addr a) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAss.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAss show ?case by(auto intro!:FAss\<^sub>1) next case (FAssNull e\<^sub>1 h l h\<^sub>1 l\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 F D) have "PROP ?P e\<^sub>1 h l null h\<^sub>1 l\<^sub>1 Vs ls" by fact with FAssNull.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l null h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAssNull.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAssNull show ?case by(auto intro!:FAssNull\<^sub>1) next case FAssThrow1 thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (FAssThrow2 e\<^sub>1 h l v h\<^sub>1 l\<^sub>1 e\<^sub>2 e h\<^sub>2 l\<^sub>2 F D) have "PROP ?P e\<^sub>1 h l (Val v) h\<^sub>1 l\<^sub>1 Vs ls" by fact with FAssThrow2.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l (Val v) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAssThrow2.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAssThrow2 show ?case by(auto intro!:FAssThrow\<^sub>2\<^sub>1) next case (CallNull e h l h\<^sub>1 l\<^sub>1 es vs h\<^sub>2 l\<^sub>2 M) have "PROP ?P e h l null h\<^sub>1 l\<^sub>1 Vs ls" by fact with CallNull.prems obtain ls\<^sub>1 where 1: "?Post e h l null h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h\<^sub>1 l\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 CallNull.prems obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 CallNull show ?case by (auto simp add: comp_def elim!: CallNull\<^sub>1) next case CallObjThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (CallParamsThrow e h l v h\<^sub>1 l\<^sub>1 es vs ex es' h\<^sub>2 l\<^sub>2 M) have "PROP ?P e h l (Val v) h\<^sub>1 l\<^sub>1 Vs ls" by fact with CallParamsThrow.prems obtain ls\<^sub>1 where 1: "?Post e h l (Val v) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h\<^sub>1 l\<^sub>1 (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 CallParamsThrow.prems obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 CallParamsThrow show ?case by (auto simp add: comp_def elim!: CallParamsThrow\<^sub>1 dest!:evals_final) next case (Call e h l a h\<^sub>1 l\<^sub>1 es vs h\<^sub>2 l\<^sub>2 C fs M Ts T pns body D l\<^sub>2' b' h\<^sub>3 l\<^sub>3) have "PROP ?P e h l (addr a) h\<^sub>1 l\<^sub>1 Vs ls" by fact with Call.prems obtain ls\<^sub>1 where 1: "?Post e h l (addr a) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h\<^sub>1 l\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 Call.prems obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" "size ls\<^sub>1 = size ls\<^sub>2" by(auto intro!:evals\<^sub>1_preserves_len) let ?Vs = "this#pns" let ?ls = "Addr a # vs @ replicate (max_vars body) undefined" have mdecl: "P \ C sees M: Ts\T = (pns, body) in D" by fact have fv_body: "fv body \ set ?Vs" and wf_size: "size Ts = size pns" using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def) have mdecl\<^sub>1: "compP\<^sub>1 P \ C sees M: Ts\T = (compE\<^sub>1 ?Vs body) in D" using sees_method_compP[OF mdecl, of "\(pns,e). compE\<^sub>1 (this#pns) e"] by(simp) have [simp]: "l\<^sub>2' = [this \ Addr a, pns [\] vs]" by fact have Call_size: "size vs = size pns" by fact have "PROP ?P body h\<^sub>2 l\<^sub>2' b' h\<^sub>3 l\<^sub>3 ?Vs ?ls" by fact with 1 2 fv_body Call_size Call.prems obtain ls\<^sub>3 where 3: "?Post body h\<^sub>2 l\<^sub>2' b' h\<^sub>3 l\<^sub>3 ?Vs ?ls ls\<^sub>3" by(auto) have hp: "h\<^sub>2 a = Some (C, fs)" by fact from 1 2 3 hp mdecl\<^sub>1 wf_size Call_size show ?case by(fastforce simp add: comp_def intro!: Call\<^sub>1 dest!:evals_final) qed (*>*) subsection\Preservation of well-formedness\ text\The compiler preserves well-formedness. Is less trivial than it may appear. We start with two simple properties: preservation of well-typedness\ lemma compE\<^sub>1_pres_wt: "\Vs Ts U. \ P,[Vs[\]Ts] \ e :: U; size Ts = size Vs \ \ compP f P,Ts \\<^sub>1 compE\<^sub>1 Vs e :: U" and "\Vs Ts Us. \ P,[Vs[\]Ts] \ es [::] Us; size Ts = size Vs \ \ compP f P,Ts \\<^sub>1 compEs\<^sub>1 Vs es [::] Us" (*<*) -apply(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) -apply clarsimp -apply(fastforce) -apply clarsimp -apply(fastforce split:bop.splits) -apply (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) -apply (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) -apply (fastforce) -apply (fastforce) -apply (fastforce dest!: sees_method_compP[where f = f]) -apply (fastforce simp:nth_append) -apply (fastforce) -apply (fastforce) -apply (fastforce) -apply (fastforce) -apply (fastforce simp:nth_append) -apply simp -apply (fastforce) -done +proof(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) + case Var then show ?case + by (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) +next + case LAss then show ?case + by (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) +next + case Call then show ?case + by (fastforce dest!: sees_method_compP[where f = f]) +next + case Block then show ?case by (fastforce simp:nth_append) +next + case TryCatch then show ?case by (fastforce simp:nth_append) +qed fastforce+ (*>*) text\\noindent and the correct block numbering:\ lemma \: "\Vs n. size Vs = n \ \ (compE\<^sub>1 Vs e) n" and \s: "\Vs n. size Vs = n \ \s (compEs\<^sub>1 Vs es) n" -(*<*)apply (induction e and es rule: \.induct \s.induct) - apply (auto dest: sym) - apply (metis length_append_singleton) - apply (metis length_append_singleton) - done(*>*) +(*<*)by (induct e and es rule: \.induct \s.induct) + (force | simp,metis length_append_singleton)+(*>*) text\The main complication is preservation of definite assignment @{term"\"}.\ lemma image_last_index: "A \ set(xs@[x]) \ last_index (xs @ [x]) ` A = (if x \ A then insert (size xs) (last_index xs ` (A-{x})) else last_index xs ` A)" (*<*) by(auto simp:image_def) (*>*) lemma A_compE\<^sub>1_None[simp]: "\Vs. \ e = None \ \ (compE\<^sub>1 Vs e) = None" and "\Vs. \s es = None \ \s (compEs\<^sub>1 Vs es) = None" (*<*)by(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct)(auto simp:hyperset_defs)(*>*) lemma A_compE\<^sub>1: "\A Vs. \ \ e = \A\; fv e \ set Vs \ \ \ (compE\<^sub>1 Vs e) = \last_index Vs ` A\" and "\A Vs. \ \s es = \A\; fvs es \ set Vs \ \ \s (compEs\<^sub>1 Vs es) = \last_index Vs ` A\" (*<*) proof(induct e and es rule: fv.induct fvs.induct) case (Block V' T e) hence "fv e \ set (Vs@[V'])" by fastforce moreover obtain B where "\ e = \B\" using Block.prems by(simp add: hyperset_defs) moreover from calculation have "B \ set (Vs@[V'])" by(auto dest!:A_fv) ultimately show ?case using Block by(auto simp add: hyperset_defs image_last_index last_index_size_conv) next case (TryCatch e\<^sub>1 C V' e\<^sub>2) hence fve\<^sub>2: "fv e\<^sub>2 \ set (Vs@[V'])" by auto show ?case proof (cases "\ e\<^sub>1") assume A\<^sub>1: "\ e\<^sub>1 = None" then obtain A\<^sub>2 where A\<^sub>2: "\ e\<^sub>2 = \A\<^sub>2\" using TryCatch by(simp add:hyperset_defs) hence "A\<^sub>2 \ set (Vs@[V'])" using TryCatch.prems A_fv[OF A\<^sub>2] by simp blast thus ?thesis using TryCatch fve\<^sub>2 A\<^sub>1 A\<^sub>2 by(auto simp add:hyperset_defs image_last_index last_index_size_conv) next fix A\<^sub>1 assume A\<^sub>1: "\ e\<^sub>1 = \A\<^sub>1\" show ?thesis proof (cases "\ e\<^sub>2") assume A\<^sub>2: "\ e\<^sub>2 = None" then show ?case using TryCatch A\<^sub>1 by(simp add:hyperset_defs) next fix A\<^sub>2 assume A\<^sub>2: "\ e\<^sub>2 = \A\<^sub>2\" have "A\<^sub>1 \ set Vs" using TryCatch.prems A_fv[OF A\<^sub>1] by simp blast moreover have "A\<^sub>2 \ set (Vs@[V'])" using TryCatch.prems A_fv[OF A\<^sub>2] by simp blast ultimately show ?thesis using TryCatch A\<^sub>1 A\<^sub>2 by (auto simp add: Diff_subset_conv last_index_size_conv subsetD hyperset_defs dest!: sym [of _ A]) qed qed next case (Cond e e\<^sub>1 e\<^sub>2) { assume "\ e = None \ \ e\<^sub>1 = None \ \ e\<^sub>2 = None" hence ?case using Cond by (auto simp add: hyperset_defs) } moreover { fix A A\<^sub>1 A\<^sub>2 assume "\ e = \A\" and A\<^sub>1: "\ e\<^sub>1 = \A\<^sub>1\" and A\<^sub>2: "\ e\<^sub>2 = \A\<^sub>2\" moreover have "A\<^sub>1 \ set Vs" using Cond.prems A_fv[OF A\<^sub>1] by simp blast moreover have "A\<^sub>2 \ set Vs" using Cond.prems A_fv[OF A\<^sub>2] by simp blast ultimately have ?case using Cond by(auto simp add:hyperset_defs image_Un inj_on_image_Int[OF inj_on_last_index]) } ultimately show ?case by fastforce qed (auto simp add:hyperset_defs) (*>*) lemma D_None[iff]: "\ (e::'a exp) None" and [iff]: "\s (es::'a exp list) None" (*<*)by(induct e and es rule: \.induct \s.induct)(simp_all)(*>*) lemma D_last_index_compE\<^sub>1: "\A Vs. \ A \ set Vs; fv e \ set Vs \ \ \ e \A\ \ \ (compE\<^sub>1 Vs e) \last_index Vs ` A\" and "\A Vs. \ A \ set Vs; fvs es \ set Vs \ \ \s es \A\ \ \s (compEs\<^sub>1 Vs es) \last_index Vs ` A\" (*<*) proof(induct e and es rule: \.induct \s.induct) case (BinOp e\<^sub>1 bop e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using BinOp by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] BinOp.prems by auto have "A \ A\<^sub>1 \ set Vs" using BinOp.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using BinOp Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (FAss e\<^sub>1 F D e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using FAss by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] FAss.prems by auto have "A \ A\<^sub>1 \ set Vs" using FAss.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using FAss Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (Call e\<^sub>1 M es) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using Call by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] Call.prems by auto have "A \ A\<^sub>1 \ set Vs" using Call.prems A_fv[OF Some] by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` (A \ A\<^sub>1)\" using Call Some by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (TryCatch e\<^sub>1 C V e\<^sub>2) have "\ A\{V} \ set(Vs@[V]); fv e\<^sub>2 \ set(Vs@[V]); \ e\<^sub>2 \A\{V}\\ \ \ (compE\<^sub>1 (Vs@[V]) e\<^sub>2) \last_index (Vs@[V]) ` (A\{V})\" by fact hence "\ (compE\<^sub>1 (Vs@[V]) e\<^sub>2) \last_index (Vs@[V]) ` (A\{V})\" using TryCatch.prems by(simp add:Diff_subset_conv) moreover have "last_index (Vs@[V]) ` A \ last_index Vs ` A \ {size Vs}" using TryCatch.prems by(auto simp add: image_last_index split:if_split_asm) ultimately show ?case using TryCatch by(auto simp:hyperset_defs elim!:D_mono') next case (Seq e\<^sub>1 e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using Seq by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] Seq.prems by auto have "A \ A\<^sub>1 \ set Vs" using Seq.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using Seq Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (Cond e e\<^sub>1 e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e) \last_index Vs ` A\" by simp show ?case proof (cases "\ e") case None thus ?thesis using Cond by simp next case (Some B) have indexB: "\ (compE\<^sub>1 Vs e) = \last_index Vs ` B\" using A_compE\<^sub>1[OF Some] Cond.prems by auto have "A \ B \ set Vs" using Cond.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` (A \ B)\" and "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ B)\" using Cond Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A \ last_index Vs ` B\" and "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` B\" by(simp add: image_Un)+ thus ?thesis using IH\<^sub>1 indexB by auto qed next case (While e\<^sub>1 e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using While by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] While.prems by auto have "A \ A\<^sub>1 \ set Vs" using While.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using While Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (Block V T e) have "\ A-{V} \ set(Vs@[V]); fv e \ set(Vs@[V]); \ e \A-{V}\ \ \ \ (compE\<^sub>1 (Vs@[V]) e) \last_index (Vs@[V]) ` (A-{V})\" by fact hence "\ (compE\<^sub>1 (Vs@[V]) e) \last_index (Vs@[V]) ` (A-{V})\" using Block.prems by(simp add:Diff_subset_conv) moreover have "size Vs \ last_index Vs ` A" using Block.prems by(auto simp add:image_def size_last_index_conv) ultimately show ?case using Block by(auto simp add: image_last_index Diff_subset_conv hyperset_defs elim!: D_mono') next case (Cons_exp e\<^sub>1 es) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using Cons_exp by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] Cons_exp.prems by auto have "A \ A\<^sub>1 \ set Vs" using Cons_exp.prems A_fv[OF Some] by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` (A \ A\<^sub>1)\" using Cons_exp Some by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed qed (simp_all add:hyperset_defs) (*>*) lemma last_index_image_set: "distinct xs \ last_index xs ` set xs = {..*) lemma D_compE\<^sub>1: "\ \ e \set Vs\; fv e \ set Vs; distinct Vs \ \ \ (compE\<^sub>1 Vs e) \{.." (*<*)by(fastforce dest!: D_last_index_compE\<^sub>1[OF subset_refl] simp add:last_index_image_set)(*>*) lemma D_compE\<^sub>1': assumes "\ e \set(V#Vs)\" and "fv e \ set(V#Vs)" and "distinct(V#Vs)" shows "\ (compE\<^sub>1 (V#Vs) e) \{..length Vs}\" (*<*) proof - have "{..size Vs} = {..1) qed (*>*) lemma compP\<^sub>1_pres_wf: "wf_J_prog P \ wf_J\<^sub>1_prog (compP\<^sub>1 P)" (*<*) -apply simp -apply(rule wf_prog_compPI) - prefer 2 apply assumption -apply(case_tac m) -apply(simp add:wf_mdecl_def wf_J\<^sub>1_mdecl_def wf_J_mdecl) -apply(clarify) -apply(frule WT_fv) -apply(fastforce intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1' \) -done +proof - + assume wf: "wf_J_prog P" + let ?f = "(\(pns, body). compE\<^sub>1 (this # pns) body)" + let ?wf\<^sub>2 = "wf_J\<^sub>1_mdecl" + + { fix C M b Ts T m + assume cM: "P \ C sees M : Ts\T = m in C" + and wfm: "wf_mdecl wf_J_mdecl P C (M, Ts, T, m)" + obtain pns body where [simp]: "m = (pns, body)" by(cases m) simp + let ?E = "[pns [\] Ts, this \ Class C]" + obtain T' where WT: "P,?E \ body :: T'" and subT: "P \ T' \ T" + using wfm by(auto simp: wf_mdecl_def) + have fv: "fv body \ dom ?E" by(rule WT_fv[OF WT]) + have "wf_mdecl ?wf\<^sub>2 (compP ?f P) C (M, Ts, T, ?f m)" using cM wfm fv + by(clarsimp simp add:wf_mdecl_def wf_J\<^sub>1_mdecl_def) + (fastforce intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1' \) + } + then show ?thesis by simp (rule wf_prog_compPI[OF _ wf]) +qed (*>*) end diff --git a/thys/Jinja/Compiler/Correctness2.thy b/thys/Jinja/Compiler/Correctness2.thy --- a/thys/Jinja/Compiler/Correctness2.thy +++ b/thys/Jinja/Compiler/Correctness2.thy @@ -1,1090 +1,1145 @@ (* Title: Jinja/Compiler/Correctness2.thy Author: Tobias Nipkow Copyright TUM 2003 *) section \Correctness of Stage 2\ theory Correctness2 imports "HOL-Library.Sublist" Compiler2 begin (*<*)hide_const (open) Throw(*>*) subsection\Instruction sequences\ text\How to select individual instructions and subsequences of instructions from a program given the class, method and program counter.\ definition before :: "jvm_prog \ cname \ mname \ nat \ instr list \ bool" ("(_,_,_,_/ \ _)" [51,0,0,0,51] 50) where "P,C,M,pc \ is \ prefix is (drop pc (instrs_of P C M))" definition at :: "jvm_prog \ cname \ mname \ nat \ instr \ bool" ("(_,_,_,_/ \ _)" [51,0,0,0,51] 50) where "P,C,M,pc \ i \ (\is. drop pc (instrs_of P C M) = i#is)" lemma [simp]: "P,C,M,pc \ []" (*<*)by(simp add:before_def)(*>*) lemma [simp]: "P,C,M,pc \ (i#is) = (P,C,M,pc \ i \ P,C,M,pc + 1 \ is)" (*<*)by(fastforce simp add:before_def at_def prefix_def drop_Suc drop_tl)(*>*) (*<*) declare drop_drop[simp del] (*>*) lemma [simp]: "P,C,M,pc \ (is\<^sub>1 @ is\<^sub>2) = (P,C,M,pc \ is\<^sub>1 \ P,C,M,pc + size is\<^sub>1 \ is\<^sub>2)" (*<*) -apply(simp add:before_def prefix_def) -apply(subst add.commute) -apply(simp add: drop_drop[symmetric]) -apply fastforce -done +by(subst add.commute) + (fastforce simp add:before_def prefix_def drop_drop[symmetric]) (*>*) (*<*) declare drop_drop[simp] (*>*) lemma [simp]: "P,C,M,pc \ i \ instrs_of P C M ! pc = i" (*<*)by(clarsimp simp add:at_def strict_prefix_def nth_via_drop)(*>*) lemma beforeM: "P \ C sees M: Ts\T = body in D \ compP\<^sub>2 P,D,M,0 \ compE\<^sub>2 body @ [Return]" (*<*) -apply(drule sees_method_idemp) -apply(simp add:before_def compP\<^sub>2_def compMb\<^sub>2_def) -done +by(drule sees_method_idemp) + (simp add:before_def compP\<^sub>2_def compMb\<^sub>2_def) (*>*) text\This lemma executes a single instruction by rewriting:\ lemma [simp]: "P,C,M,pc \ instr \ (P \ (None, h, (vs,ls,C,M,pc) # frs) -jvm\ \') = ((None, h, (vs,ls,C,M,pc) # frs) = \' \ (\\. exec(P,(None, h, (vs,ls,C,M,pc) # frs)) = Some \ \ P \ \ -jvm\ \'))" (*<*) -apply(simp only: exec_all_def) -apply(blast intro: converse_rtranclE converse_rtrancl_into_rtrancl) -done +by(simp only: exec_all_def) + (blast intro: converse_rtranclE converse_rtrancl_into_rtrancl) (*>*) subsection\Exception tables\ definition pcs :: "ex_table \ nat set" where "pcs xt \ \(f,t,C,h,d) \ set xt. {f ..< t}" lemma pcs_subset: shows "\pc d. pcs(compxE\<^sub>2 e pc d) \ {pc..2 e)}" and "\pc d. pcs(compxEs\<^sub>2 es pc d) \ {pc..2 es)}" (*<*) -apply(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) -apply (simp_all add:pcs_def) -apply (fastforce split:bop.splits)+ -done +proof(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) + case Cast then show ?case by (fastforce simp:pcs_def) +next + case BinOp then show ?case by (fastforce simp:pcs_def split:bop.splits) +next + case LAss then show ?case by (fastforce simp: pcs_def) +next + case FAcc then show ?case by (fastforce simp: pcs_def) +next + case FAss then show ?case by (fastforce simp: pcs_def) +next + case Call then show ?case by (fastforce simp: pcs_def) +next + case Seq then show ?case by (fastforce simp: pcs_def) +next + case Cond then show ?case by (fastforce simp: pcs_def) +next + case While then show ?case by (fastforce simp: pcs_def) +next + case throw then show ?case by (fastforce simp: pcs_def) +next + case TryCatch then show ?case by (fastforce simp: pcs_def) +next + case Cons_exp then show ?case by (fastforce simp: pcs_def) +qed (simp_all add:pcs_def) (*>*) lemma [simp]: "pcs [] = {}" (*<*)by(simp add:pcs_def)(*>*) lemma [simp]: "pcs (x#xt) = {fst x ..< fst(snd x)} \ pcs xt" (*<*)by(auto simp add: pcs_def)(*>*) lemma [simp]: "pcs(xt\<^sub>1 @ xt\<^sub>2) = pcs xt\<^sub>1 \ pcs xt\<^sub>2" (*<*)by(simp add:pcs_def)(*>*) lemma [simp]: "pc < pc\<^sub>0 \ pc\<^sub>0+size(compE\<^sub>2 e) \ pc \ pc \ pcs(compxE\<^sub>2 e pc\<^sub>0 d)" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc < pc\<^sub>0 \ pc\<^sub>0+size(compEs\<^sub>2 es) \ pc \ pc \ pcs(compxEs\<^sub>2 es pc\<^sub>0 d)" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc\<^sub>1 + size(compE\<^sub>2 e\<^sub>1) \ pc\<^sub>2 \ pcs(compxE\<^sub>2 e\<^sub>1 pc\<^sub>1 d\<^sub>1) \ pcs(compxE\<^sub>2 e\<^sub>2 pc\<^sub>2 d\<^sub>2) = {}" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc\<^sub>1 + size(compE\<^sub>2 e) \ pc\<^sub>2 \ pcs(compxE\<^sub>2 e pc\<^sub>1 d\<^sub>1) \ pcs(compxEs\<^sub>2 es pc\<^sub>2 d\<^sub>2) = {}" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc \ pcs xt\<^sub>0 \ match_ex_table P C pc (xt\<^sub>0 @ xt\<^sub>1) = match_ex_table P C pc xt\<^sub>1" (*<*)by (induct xt\<^sub>0) (auto simp: matches_ex_entry_def)(*>*) lemma [simp]: "\ x \ set xt; pc \ pcs xt \ \ \ matches_ex_entry P D pc x" (*<*)by(auto simp:matches_ex_entry_def pcs_def)(*>*) lemma [simp]: assumes xe: "xe \ set(compxE\<^sub>2 e pc d)" and outside: "pc' < pc \ pc+size(compE\<^sub>2 e) \ pc'" shows "\ matches_ex_entry P C pc' xe" (*<*) proof assume "matches_ex_entry P C pc' xe" with xe have "pc' \ pcs(compxE\<^sub>2 e pc d)" by(force simp add:matches_ex_entry_def pcs_def) with outside show False by simp qed (*>*) lemma [simp]: assumes xe: "xe \ set(compxEs\<^sub>2 es pc d)" and outside: "pc' < pc \ pc+size(compEs\<^sub>2 es) \ pc'" shows "\ matches_ex_entry P C pc' xe" (*<*) proof assume "matches_ex_entry P C pc' xe" with xe have "pc' \ pcs(compxEs\<^sub>2 es pc d)" by(force simp add:matches_ex_entry_def pcs_def) with outside show False by simp qed (*>*) lemma match_ex_table_app[simp]: "\xte \ set xt\<^sub>1. \ matches_ex_entry P D pc xte \ match_ex_table P D pc (xt\<^sub>1 @ xt) = match_ex_table P D pc xt" (*<*)by(induct xt\<^sub>1) simp_all(*>*) lemma [simp]: "\x \ set xtab. \ matches_ex_entry P C pc x \ match_ex_table P C pc xtab = None" (*<*)using match_ex_table_app[where ?xt = "[]"] by fastforce(*>*) lemma match_ex_entry: "matches_ex_entry P C pc (start, end, catch_type, handler) = (start \ pc \ pc < end \ P \ C \\<^sup>* catch_type)" (*<*)by(simp add:matches_ex_entry_def)(*>*) definition caught :: "jvm_prog \ pc \ heap \ addr \ ex_table \ bool" where "caught P pc h a xt \ (\entry \ set xt. matches_ex_entry P (cname_of h a) pc entry)" definition beforex :: "jvm_prog \ cname \ mname \ ex_table \ nat set \ nat \ bool" ("(2_,/_,/_ \/ _ /'/ _,/_)" [51,0,0,0,0,51] 50) where "P,C,M \ xt / I,d \ (\xt\<^sub>0 xt\<^sub>1. ex_table_of P C M = xt\<^sub>0 @ xt @ xt\<^sub>1 \ pcs xt\<^sub>0 \ I = {} \ pcs xt \ I \ (\pc \ I. \C pc' d'. match_ex_table P C pc xt\<^sub>1 = \(pc',d')\ \ d' \ d))" definition dummyx :: "jvm_prog \ cname \ mname \ ex_table \ nat set \ nat \ bool" ("(2_,_,_ \/ _ '/_,_)" [51,0,0,0,0,51] 50) where "P,C,M \ xt/I,d \ P,C,M \ xt/I,d" +abbreviation +"beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1 + \ ex_table_of P C M = xt\<^sub>0 @ xt @ xt\<^sub>1 \ pcs xt\<^sub>0 \ I = {} + \ pcs xt \ I \ (\pc \ I. \C pc' d'. match_ex_table P C pc xt\<^sub>1 = \(pc',d')\ \ d' \ d)" + +lemma beforex_beforex\<^sub>0_eq: + "P,C,M \ xt / I,d \ \xt\<^sub>0 xt\<^sub>1. beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1" +using beforex_def by auto + + lemma beforexD1: "P,C,M \ xt / I,d \ pcs xt \ I" (*<*)by(auto simp add:beforex_def)(*>*) lemma beforex_mono: "\ P,C,M \ xt/I,d'; d' \ d \ \ P,C,M \ xt/I,d" (*<*)by(fastforce simp:beforex_def)(*>*) lemma [simp]: "P,C,M \ xt/I,d \ P,C,M \ xt/I,Suc d" (*<*)by(fastforce intro:beforex_mono)(*>*) lemma beforex_append[simp]: "pcs xt\<^sub>1 \ pcs xt\<^sub>2 = {} \ P,C,M \ xt\<^sub>1 @ xt\<^sub>2/I,d = (P,C,M \ xt\<^sub>1/I-pcs xt\<^sub>2,d \ P,C,M \ xt\<^sub>2/I-pcs xt\<^sub>1,d \ P,C,M \ xt\<^sub>1@xt\<^sub>2/I,d)" -(*<*) -apply(rule iffI) - prefer 2 - apply(simp add:dummyx_def) -apply(auto simp add: beforex_def dummyx_def) - apply(rule_tac x = xt\<^sub>0 in exI) - apply auto -apply(rule_tac x = "xt\<^sub>0@xt\<^sub>1" in exI) -apply auto -done +(*<*)(is "?Q \ ?P = (?P1 \ ?P2 \ ?P3)" is "?Q \ ?P = ?P123") +proof - + assume pcs: ?Q + show ?thesis proof(rule iffI) + assume "?P123" then show ?P by(simp add:dummyx_def) + next + assume hyp: ?P + let ?xt = "xt\<^sub>1 @ xt\<^sub>2" + let ?beforex = "beforex\<^sub>0 P C M d" + obtain xt\<^sub>0 xt\<^sub>1' where beforex: "?beforex I ?xt xt\<^sub>0 xt\<^sub>1'" + using hyp by(clarsimp simp: beforex_def) + have "\xt\<^sub>0 xt\<^sub>1'. ?beforex (I - pcs xt\<^sub>2) xt\<^sub>1 xt\<^sub>0 xt\<^sub>1'" \ \?P1\ + using pcs beforex by(rule_tac x=xt\<^sub>0 in exI) auto + moreover have "\xt\<^sub>0 xt\<^sub>1'. ?beforex (I - pcs xt\<^sub>1) xt\<^sub>2 xt\<^sub>0 xt\<^sub>1'" \ \?P2\ + using pcs beforex by(rule_tac x="xt\<^sub>0@xt\<^sub>1" in exI) auto + moreover have ?P3 using hyp by(simp add: dummyx_def) + ultimately show ?P123 by (simp add: beforex_def) + qed +qed (*>*) lemma beforex_appendD1: - "\ P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d; - pcs xt\<^sub>1 \ J; J \ I; J \ pcs xt\<^sub>2 = {} \ - \ P,C,M \ xt\<^sub>1 / J,d" +assumes bx: "P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d" + and pcs: "pcs xt\<^sub>1 \ J" and JI: "J \ I" and Jpcs: "J \ pcs xt\<^sub>2 = {}" +shows "P,C,M \ xt\<^sub>1 / J,d" (*<*) -apply(auto simp:beforex_def) -apply(rule exI,rule exI,rule conjI, rule refl) -apply(rule conjI, blast) -apply(auto) -apply(subgoal_tac "pc \ pcs xt\<^sub>2") - prefer 2 apply blast -apply (auto split:if_split_asm) -done +proof - + let ?beforex = "beforex\<^sub>0 P C M d" + obtain xt\<^sub>0 xt\<^sub>1' where bx': "?beforex I (xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)]) xt\<^sub>0 xt\<^sub>1'" + using bx by(clarsimp simp:beforex_def) + let ?xt0 = xt\<^sub>0 and ?xt1 = "xt\<^sub>2 @ (f, t, D, h, d) # xt\<^sub>1'" + have "pcs xt\<^sub>0 \ J = {}" using bx' JI by blast + moreover { + fix pc C pc' d' assume pcJ: "pc\J" + then have "pc \ pcs xt\<^sub>2" using bx' Jpcs by blast + then have "match_ex_table P C pc (xt\<^sub>2 @ (f, t, D, h, d) # xt\<^sub>1') + = \(pc', d')\ \ d' \ d" + using bx' JI pcJ by (auto split:if_split_asm) + } + ultimately have "?beforex J xt\<^sub>1 ?xt0 ?xt1" using bx' pcs by simp + then show ?thesis using beforex_def by blast +qed (*>*) lemma beforex_appendD2: - "\ P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d; - pcs xt\<^sub>2 \ J; J \ I; J \ pcs xt\<^sub>1 = {} \ - \ P,C,M \ xt\<^sub>2 / J,d" +assumes bx: "P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d" + and pcs: "pcs xt\<^sub>2 \ J" and JI: "J \ I" and Jpcs: "J \ pcs xt\<^sub>1 = {}" +shows "P,C,M \ xt\<^sub>2 / J,d" (*<*) -apply(auto simp:beforex_def) -apply(rule_tac x = "xt\<^sub>0 @ xt\<^sub>1" in exI) -apply fastforce -done +proof - + let ?beforex = "beforex\<^sub>0 P C M d" + obtain xt\<^sub>0 xt\<^sub>1' where bx': "?beforex I (xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)]) xt\<^sub>0 xt\<^sub>1'" + using bx by(clarsimp simp:beforex_def) + then have "\xt\<^sub>1''. beforex\<^sub>0 P C M d J xt\<^sub>2 (xt\<^sub>0 @ xt\<^sub>1) xt\<^sub>1''" + using assms by fastforce + then show ?thesis using beforex_def by blast +qed (*>*) lemma beforexM: "P \ C sees M: Ts\T = body in D \ compP\<^sub>2 P,D,M \ compxE\<^sub>2 body 0 0/{..2 body)},0" (*<*) -apply(drule sees_method_idemp) -apply(drule sees_method_compP[where f = compMb\<^sub>2]) -apply(simp add:beforex_def compP\<^sub>2_def compMb\<^sub>2_def) -apply(rule_tac x = "[]" in exI) -using pcs_subset apply fastforce -done +proof - + assume cM: "P \ C sees M: Ts\T = body in D" + let ?xt0 = "[]" + have "\xt1. beforex\<^sub>0 (compP\<^sub>2 P) D M 0 ({..2 body)}) (compxE\<^sub>2 body 0 0) ?xt0 xt1" + using sees_method_compP[where f = compMb\<^sub>2, OF sees_method_idemp[OF cM]] + pcs_subset by(fastforce simp add: compP\<^sub>2_def compMb\<^sub>2_def) + then show ?thesis using beforex_def by blast +qed (*>*) lemma match_ex_table_SomeD2: - "\ match_ex_table P D pc (ex_table_of P C M) = \(pc',d')\; - P,C,M \ xt/I,d; \x \ set xt. \ matches_ex_entry P D pc x; pc \ I \ - \ d' \ d" +assumes met: "match_ex_table P D pc (ex_table_of P C M) = \(pc',d')\" + and bx: "P,C,M \ xt/I,d" + and nmet: "\x \ set xt. \ matches_ex_entry P D pc x" and pcI: "pc \ I" +shows "d' \ d" (*<*) -apply(auto simp:beforex_def) -apply(subgoal_tac "pc \ pcs xt\<^sub>0") -apply auto -done +proof - + obtain xt\<^sub>0 xt\<^sub>1 where bx': "beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1" + using bx by(clarsimp simp:beforex_def) + then have "pc \ pcs xt\<^sub>0" using pcI by blast + then show ?thesis using bx' met nmet pcI by simp +qed (*>*) lemma match_ex_table_SomeD1: "\ match_ex_table P D pc (ex_table_of P C M) = \(pc',d')\; P,C,M \ xt / I,d; pc \ I; pc \ pcs xt \ \ d' \ d" (*<*)by(auto elim: match_ex_table_SomeD2)(*>*) subsection\The correctness proof\ (*<*) declare nat_add_distrib[simp] caught_def[simp] declare fun_upd_apply[simp del] (*>*) definition handle :: "jvm_prog \ cname \ mname \ addr \ heap \ val list \ val list \ nat \ frame list \ jvm_state" where "handle P C M a h vs ls pc frs = find_handler P a h ((vs,ls,C,M,pc) # frs)" lemma handle_Cons: "\ P,C,M \ xt/I,d; d \ size vs; pc \ I; \x \ set xt. \ matches_ex_entry P (cname_of h xa) pc x \ \ handle P C M xa h (v # vs) ls pc frs = handle P C M xa h vs ls pc frs" (*<*)by(auto simp:handle_def Suc_diff_le dest: match_ex_table_SomeD2)(*>*) lemma handle_append: - "\ P,C,M \ xt/I,d; d \ size vs; pc \ I; pc \ pcs xt \ \ - handle P C M xa h (ws @ vs) ls pc frs = handle P C M xa h vs ls pc frs" +assumes bx: "P,C,M \ xt/I,d" and d: "d \ size vs" + and pcI: "pc \ I" and pc_not: "pc \ pcs xt" +shows "handle P C M xa h (ws @ vs) ls pc frs = handle P C M xa h vs ls pc frs" (*<*) -apply(auto simp:handle_def) -apply(rename_tac pc' d') -apply(subgoal_tac "size ws \ length ws + length vs - d'") -apply(simp add:drop_all) -apply(fastforce dest:match_ex_table_SomeD2 split:nat_diff_split) -done +proof - + { fix pc' d' + assume "match_ex_table P (cname_of h xa) pc (ex_table_of P C M) = \(pc', d')\" + then have "length ws \ length ws + length vs - d'" + using assms by(fastforce dest:match_ex_table_SomeD2 split:nat_diff_split) + } + then show ?thesis by(simp add: handle_def) +qed (*>*) lemma aux_isin[simp]: "\ B \ A; a \ B \ \ a \ A" (*<*)by blast(*>*) lemma fixes P\<^sub>1 defines [simp]: "P \ compP\<^sub>2 P\<^sub>1" shows Jcc: "P\<^sub>1 \\<^sub>1 \e,(h\<^sub>0,ls\<^sub>0)\ \ \ef,(h\<^sub>1,ls\<^sub>1)\ \ (\C M pc v xa vs frs I. \ P,C,M,pc \ compE\<^sub>2 e; P,C,M \ compxE\<^sub>2 e pc (size vs)/I,size vs; {pc..2 e)} \ I \ \ (ef = Val v \ P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 e))#frs)) \ (ef = Throw xa \ (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compE\<^sub>2 e) \ \ caught P pc\<^sub>1 h\<^sub>1 xa (compxE\<^sub>2 e pc (size vs)) \ P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ handle P C M xa h\<^sub>1 vs ls\<^sub>1 pc\<^sub>1 frs)))" (*<*) (is "_ \ (\C M pc v xa vs frs I. PROP ?P e h\<^sub>0 ls\<^sub>0 ef h\<^sub>1 ls\<^sub>1 C M pc v xa vs frs I)") (*>*) and "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>0,ls\<^sub>0)\ [\] \fs,(h\<^sub>1,ls\<^sub>1)\ \ (\C M pc ws xa es' vs frs I. \ P,C,M,pc \ compEs\<^sub>2 es; P,C,M \ compxEs\<^sub>2 es pc (size vs)/I,size vs; {pc..2 es)} \ I \ \ (fs = map Val ws \ P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(rev ws @ vs,ls\<^sub>1,C,M,pc+size(compEs\<^sub>2 es))#frs)) \ (fs = map Val ws @ Throw xa # es' \ (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compEs\<^sub>2 es) \ \ caught P pc\<^sub>1 h\<^sub>1 xa (compxEs\<^sub>2 es pc (size vs)) \ P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ handle P C M xa h\<^sub>1 vs ls\<^sub>1 pc\<^sub>1 frs)))" (*<*) (is "_ \ (\C M pc ws xa es' vs frs I. PROP ?Ps es h\<^sub>0 ls\<^sub>0 fs h\<^sub>1 ls\<^sub>1 C M pc ws xa es' vs frs I)") proof (induct rule:eval\<^sub>1_evals\<^sub>1_inducts) case New\<^sub>1 thus ?case by (clarsimp simp add:blank_def fun_eq_iff) next case NewFail\<^sub>1 thus ?case by(auto simp: handle_def pcs_def) next case (Cast\<^sub>1 e h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1 D fs C') let ?pc = "pc + length(compE\<^sub>2 e)" have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc)#frs)" using Cast\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc+1)#frs)" using Cast\<^sub>1 by (auto simp add:cast_ok_def) finally show ?case by auto next case (CastNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 C') let ?pc = "pc + length(compE\<^sub>2 e)" have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc)#frs)" using CastNull\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc+1)#frs)" using CastNull\<^sub>1 by (auto simp add:cast_ok_def) finally show ?case by auto next case (CastFail\<^sub>1 e h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1 D fs C') let ?pc = "pc + length(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt ClassCast" have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc)#frs)" using CastFail\<^sub>1 by fastforce also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc frs" using CastFail\<^sub>1 by (auto simp add:handle_def cast_ok_def) also have "handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc frs = handle P C M ?xa h\<^sub>1 vs ls\<^sub>1 ?pc frs" using CastFail\<^sub>1.prems by(auto simp:handle_Cons) finally have exec: "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ \". show ?case (is "?N \ (?eq \ (\pc\<^sub>1. ?H pc\<^sub>1))") proof show ?N by simp next have "?eq \ ?H ?pc" using exec by auto thus "?eq \ (\pc\<^sub>1. ?H pc\<^sub>1)" by blast qed next case CastThrow\<^sub>1 thus ?case by fastforce next case Val\<^sub>1 thus ?case by simp next case Var\<^sub>1 thus ?case by auto next case (BinOp\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 e\<^sub>2 v\<^sub>2 h\<^sub>2 ls\<^sub>2 bop w) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 (Val v\<^sub>2) h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 v\<^sub>2 xa (v\<^sub>1#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using BinOp\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>2,(v\<^sub>2#v\<^sub>1#vs,ls\<^sub>2,C,M,?pc\<^sub>2)#frs)" using BinOp\<^sub>1.prems IH\<^sub>2 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>2,(w#vs,ls\<^sub>2,C,M,?pc\<^sub>2+1)#frs)" using BinOp\<^sub>1 by(cases bop) auto finally show ?case by (auto split: bop.splits simp:add.assoc) next case BinOpThrow\<^sub>1\<^sub>1 thus ?case by(fastforce) next case (BinOpThrow\<^sub>2\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 e\<^sub>2 e h\<^sub>2 ls\<^sub>2 bop) let ?pc = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?\\<^sub>1 = "(None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc)#frs)" have 1: "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ ?\\<^sub>1" using BinOpThrow\<^sub>2\<^sub>1 by fastforce show ?case (is "?N \ (?eq \ (\pc\<^sub>2. ?H pc\<^sub>2))") proof show ?N by simp next { assume ?eq moreover have "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 (throw e) h\<^sub>2 ls\<^sub>2 C M ?pc v xa (v\<^sub>1#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact ultimately obtain pc\<^sub>2 where pc\<^sub>2: "?pc \ pc\<^sub>2 \ pc\<^sub>2 < ?pc + size(compE\<^sub>2 e\<^sub>2) \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc (size vs + 1))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (v\<^sub>1#vs) ls\<^sub>2 pc\<^sub>2 frs" using BinOpThrow\<^sub>2\<^sub>1.prems by fastforce have 3: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 vs ls\<^sub>2 pc\<^sub>2 frs" using 2 BinOpThrow\<^sub>2\<^sub>1.prems pc\<^sub>2 by(auto simp:handle_Cons) have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF 1 3] by auto hence "\pc\<^sub>2. ?H pc\<^sub>2" by iprover } thus "?eq \ (\pc\<^sub>2. ?H pc\<^sub>2)" by iprover qed next case (FAcc\<^sub>1 e h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1 C fs F D w) let ?pc = "pc + length(compE\<^sub>2 e)" have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc)#frs)" using FAcc\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc+1)#frs)" using FAcc\<^sub>1 by auto finally show ?case by auto next case (FAccNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 F D) let ?pc = "pc + length(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt NullPointer" have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc)#frs)" using FAccNull\<^sub>1 by fastforce also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc frs" using FAccNull\<^sub>1.prems by(fastforce simp:split_beta handle_def simp del: split_paired_Ex) also have "handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc frs = handle P C M ?xa h\<^sub>1 vs ls\<^sub>1 ?pc frs" using FAccNull\<^sub>1.prems by(auto simp add:handle_Cons) finally show ?case by (auto intro: exI[where x = ?pc]) next case FAccThrow\<^sub>1 thus ?case by fastforce next case (LAss\<^sub>1 e h\<^sub>0 ls\<^sub>0 w h\<^sub>1 ls\<^sub>1 i ls\<^sub>2) let ?pc = "pc + length(compE\<^sub>2 e)" have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc)#frs)" using LAss\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>1,(Unit#vs,ls\<^sub>2,C,M,?pc+2)#frs)" using LAss\<^sub>1 by auto finally show ?case using LAss\<^sub>1 by auto next case LAssThrow\<^sub>1 thus ?case by fastforce next case (FAss\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 C fs fs' F D h\<^sub>2') let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 w xa (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using FAss\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>2,(w#Addr a#vs,ls\<^sub>2,C,M,?pc\<^sub>2)#frs)" using FAss\<^sub>1.prems IH\<^sub>2 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>2',(Unit#vs,ls\<^sub>2,C,M,?pc\<^sub>2+2)#frs)" using FAss\<^sub>1 by auto finally show ?case using FAss\<^sub>1 by (auto simp:add.assoc) next case (FAssNull\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 F D) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" let ?xa = "addr_of_sys_xcpt NullPointer" have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 w xa (Null#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using FAssNull\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>2,(w#Null#vs,ls\<^sub>2,C,M,?pc\<^sub>2)#frs)" using FAssNull\<^sub>1.prems IH\<^sub>2 by fastforce also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (w#Null#vs) ls\<^sub>2 ?pc\<^sub>2 frs" using FAssNull\<^sub>1.prems by(fastforce simp:split_beta handle_def simp del: split_paired_Ex) also have "handle P C M ?xa h\<^sub>2 (w#Null#vs) ls\<^sub>2 ?pc\<^sub>2 frs = handle P C M ?xa h\<^sub>2 vs ls\<^sub>2 ?pc\<^sub>2 frs" using FAssNull\<^sub>1.prems by(auto simp add:handle_Cons) finally show ?case by (auto intro: exI[where x = ?pc\<^sub>2]) next case (FAssThrow\<^sub>2\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 w h\<^sub>1 ls\<^sub>1 e\<^sub>2 e' h\<^sub>2 ls\<^sub>2 F D) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?\\<^sub>1 = "(None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" have 1: "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ ?\\<^sub>1" using FAssThrow\<^sub>2\<^sub>1 by fastforce show ?case (is "?N \ (?eq \ (\pc\<^sub>2. ?H pc\<^sub>2))") proof show ?N by simp next { assume ?eq moreover have "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 v xa (w#vs) frs (I - pcs (compxE\<^sub>2 e\<^sub>1 pc (length vs)))" by fact ultimately obtain pc\<^sub>2 where pc\<^sub>2: "?pc\<^sub>1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1 + size(compE\<^sub>2 e\<^sub>2) \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1 (size vs + 1))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (w#vs) ls\<^sub>2 pc\<^sub>2 frs" using FAssThrow\<^sub>2\<^sub>1.prems by fastforce have 3: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 vs ls\<^sub>2 pc\<^sub>2 frs" using 2 FAssThrow\<^sub>2\<^sub>1.prems pc\<^sub>2 by(auto simp:handle_Cons) have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF 1 3] by auto hence "\pc\<^sub>2. ?H pc\<^sub>2" by iprover } thus "?eq \ (\pc\<^sub>2. ?H pc\<^sub>2)" by iprover qed next case FAssThrow\<^sub>1\<^sub>1 thus ?case by fastforce next case (Call\<^sub>1 e h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 Ca fs M' Ts T body D ls\<^sub>2' f h\<^sub>3 ls\<^sub>3) have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) let ?\\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc)#frs)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>1 = "(None,h\<^sub>1,(Addr a # vs, ls\<^sub>1, C,M,?pc\<^sub>1)#frs)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" let ?frs\<^sub>2 = "(rev pvs @ Addr a # vs, ls\<^sub>2, C,M,?pc\<^sub>2)#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2)" let ?frs\<^sub>2' = "([], ls\<^sub>2', D,M',0) # ?frs\<^sub>2" let ?\\<^sub>2' = "(None, h\<^sub>2, ?frs\<^sub>2')" have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 pvs xa (map Val pvs) (Addr a # vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact have "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using Call\<^sub>1 by fastforce also have "P \ \ -jvm\ ?\\<^sub>2" using IH_es Call\<^sub>1.prems by fastforce also have "P \ \ -jvm\ ?\\<^sub>2'" using Call\<^sub>1 by(auto simp add: nth_append compMb\<^sub>2_def) finally have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>2'". have "P\<^sub>1 \ Ca sees M': Ts\T = body in D" by fact then have M'_in_D: "P\<^sub>1 \ D sees M': Ts\T = body in D" by(rule sees_method_idemp) hence M'_code: "compP\<^sub>2 P\<^sub>1,D,M',0 \ compE\<^sub>2 body @ [Return]" and M'_xtab: "compP\<^sub>2 P\<^sub>1,D,M' \ compxE\<^sub>2 body 0 0/{..2 body)},0" by(rule beforeM, rule beforexM) have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' f h\<^sub>3 ls\<^sub>3 D M' 0 v xa [] ?frs\<^sub>2 ({..2 body)})" by fact show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>2' -jvm\ (None,h\<^sub>3,([v],ls\<^sub>3,D,M',size(compE\<^sub>2 body))#?frs\<^sub>2)" using val IH_body Call\<^sub>1.prems M'_code M'_xtab by (fastforce simp del:split_paired_Ex) also have "P \ \ -jvm\ (None, h\<^sub>3, (v # vs, ls\<^sub>2, C,M,?pc\<^sub>2+1)#frs)" using Call\<^sub>1 M'_code M'_in_D by(auto simp: nth_append compMb\<^sub>2_def) finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof assume throw: ?throw with IH_body obtain pc\<^sub>2 where pc\<^sub>2: "0 \ pc\<^sub>2 \ pc\<^sub>2 < size(compE\<^sub>2 body) \ \ caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and 2: "P \ ?\\<^sub>2' -jvm\ handle P D M' xa h\<^sub>3 [] ls\<^sub>3 pc\<^sub>2 ?frs\<^sub>2" using Call\<^sub>1.prems M'_code M'_xtab by (fastforce simp del:split_paired_Ex) have "handle P D M' xa h\<^sub>3 [] ls\<^sub>3 pc\<^sub>2 ?frs\<^sub>2 = handle P C M xa h\<^sub>3 (rev pvs @ Addr a # vs) ls\<^sub>2 ?pc\<^sub>2 frs" using pc\<^sub>2 M'_in_D by(auto simp add:handle_def) also have "\ = handle P C M xa h\<^sub>3 vs ls\<^sub>2 ?pc\<^sub>2 frs" using Call\<^sub>1.prems by(auto simp add:handle_append handle_Cons) finally have "?H ?pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF 1 2] by auto thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case (CallParamsThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 w h\<^sub>1 ls\<^sub>1 es es' h\<^sub>2 ls\<^sub>2 pvs ex es'' M') let ?\\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc)#frs)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>1 = "(None,h\<^sub>1,(w # vs, ls\<^sub>1, C,M,?pc\<^sub>1)#frs)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using CallParamsThrow\<^sub>1 by fastforce show ?case (is "?N \ (?eq \ (\pc\<^sub>2. ?H pc\<^sub>2))") proof show ?N by simp next { assume ?eq moreover have "PROP ?Ps es h\<^sub>1 ls\<^sub>1 es' h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 pvs xa es'' (w#vs) frs (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact ultimately have "\pc\<^sub>2. (?pc\<^sub>1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1 + size(compEs\<^sub>2 es) \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es ?pc\<^sub>1 (size vs + 1))) \ P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (w#vs) ls\<^sub>2 pc\<^sub>2 frs" (is "\pc\<^sub>2. ?PC pc\<^sub>2 \ ?Exec pc\<^sub>2") using CallParamsThrow\<^sub>1 by force then obtain pc\<^sub>2 where pc\<^sub>2: "?PC pc\<^sub>2" and 2: "?Exec pc\<^sub>2" by iprover have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF 1 2] CallParamsThrow\<^sub>1 by(auto simp:handle_Cons) hence "\pc\<^sub>2. ?H pc\<^sub>2" by iprover } thus "?eq \ (\pc\<^sub>2. ?H pc\<^sub>2)" by iprover qed next case (CallNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 M') have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" let ?xa = "addr_of_sys_xcpt NullPointer" have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 pvs xa (map Val pvs) (Null#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using CallNull\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>2,(rev pvs@Null#vs,ls\<^sub>2,C,M,?pc\<^sub>2)#frs)" using CallNull\<^sub>1 IH_es by fastforce also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (rev pvs@Null#vs) ls\<^sub>2 ?pc\<^sub>2 frs" using CallNull\<^sub>1.prems by(auto simp:split_beta handle_def nth_append simp del: split_paired_Ex) also have "handle P C M ?xa h\<^sub>2 (rev pvs@Null#vs) ls\<^sub>2 ?pc\<^sub>2 frs = handle P C M ?xa h\<^sub>2 vs ls\<^sub>2 ?pc\<^sub>2 frs" using CallNull\<^sub>1.prems by(auto simp:handle_Cons handle_append) finally show ?case by (auto intro: exI[where x = ?pc\<^sub>2]) next case CallObjThrow\<^sub>1 thus ?case by fastforce next case Block\<^sub>1 thus ?case by auto next case (Seq\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 w h\<^sub>1 ls\<^sub>1 e\<^sub>2 e\<^sub>2' h\<^sub>2 ls\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1)#frs)" have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using Seq\<^sub>1 by fastforce also have "P \ \ -jvm\ ?\\<^sub>1" using Seq\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". let ?pc\<^sub>2 = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>2)" have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 e\<^sub>2' h\<^sub>2 ls\<^sub>2 C M (?pc\<^sub>1+1) v xa vs frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note eval\<^sub>1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2)#frs)" using val Seq\<^sub>1.prems IH\<^sub>2 by fastforce finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof assume throw: ?throw then obtain pc\<^sub>2 where pc\<^sub>2: "?pc\<^sub>1+1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2 \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 (?pc\<^sub>1+1) (size vs))" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 vs ls\<^sub>2 pc\<^sub>2 frs" using IH\<^sub>2 Seq\<^sub>1.prems by fastforce have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by auto thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case SeqThrow\<^sub>1 thus ?case by fastforce next case (CondT\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 e\<^sub>1 e' h\<^sub>2 ls\<^sub>2 e\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1)#frs)" have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool(True)#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using CondT\<^sub>1 by (fastforce simp add: Int_Un_distrib) also have "P \ \ -jvm\ ?\\<^sub>1" using CondT\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". let ?pc\<^sub>1' = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2' = "?pc\<^sub>1' + 1 + length(compE\<^sub>2 e\<^sub>2)" show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note eval\<^sub>1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>1')#frs)" using val CondT\<^sub>1 by(fastforce simp:Int_Un_distrib) also have "P \ \ -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2')#frs)" using CondT\<^sub>1 by(auto simp:add.assoc) finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof let ?d = "size vs" let ?I = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>2 (?pc\<^sub>1'+1) ?d)" assume throw: ?throw moreover have "PROP ?P e\<^sub>1 h\<^sub>1 ls\<^sub>1 e' h\<^sub>2 ls\<^sub>2 C M (?pc\<^sub>1+1) v xa vs frs ?I" by fact ultimately obtain pc\<^sub>2 where pc\<^sub>2: "?pc\<^sub>1+1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1' \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) (size vs))" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 vs ls\<^sub>2 pc\<^sub>2 frs" using CondT\<^sub>1.prems by (fastforce simp:Int_Un_distrib) have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by auto thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case (CondF\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 e\<^sub>2 e' h\<^sub>2 ls\<^sub>2 e\<^sub>1) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?pc\<^sub>2 = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>1)+ 1" let ?pc\<^sub>2' = "?pc\<^sub>2 + length(compE\<^sub>2 e\<^sub>2)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>2)#frs)" have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool(False)#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using CondF\<^sub>1 by (fastforce simp add: Int_Un_distrib) also have "P \ \ -jvm\ ?\\<^sub>1" using CondF\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note eval\<^sub>1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2')#frs)" using val CondF\<^sub>1 by(fastforce simp:Int_Un_distrib) finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof let ?d = "size vs" let ?I = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) ?d)" assume throw: ?throw moreover have "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 e' h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>2 v xa vs frs ?I" by fact ultimately obtain pc\<^sub>2 where pc\<^sub>2: "?pc\<^sub>2 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2' \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>2 ?d)" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 vs ls\<^sub>2 pc\<^sub>2 frs" using CondF\<^sub>1.prems by(fastforce simp:Int_Un_distrib) have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by auto thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case (CondThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 f h\<^sub>1 ls\<^sub>1 e\<^sub>1 e\<^sub>2) let ?d = "size vs" let ?xt\<^sub>1 = "compxE\<^sub>2 e\<^sub>1 (pc+size(compE\<^sub>2 e)+1) ?d" let ?xt\<^sub>2 = "compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e)+size(compE\<^sub>2 e\<^sub>1)+2) ?d" let ?I = "I - (pcs ?xt\<^sub>1 \ pcs ?xt\<^sub>2)" have "pcs(compxE\<^sub>2 e pc ?d) \ pcs(?xt\<^sub>1 @ ?xt\<^sub>2) = {}" using CondThrow\<^sub>1.prems by (simp add:Int_Un_distrib) moreover have "PROP ?P e h\<^sub>0 ls\<^sub>0 (throw f) h\<^sub>1 ls\<^sub>1 C M pc v xa vs frs ?I" by fact ultimately show ?case using CondThrow\<^sub>1.prems by fastforce next case (WhileF\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 c) let ?pc = "pc + length(compE\<^sub>2 e)" let ?pc' = "?pc + length(compE\<^sub>2 c) + 3" have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Bool False#vs,ls\<^sub>1,C,M,?pc)#frs)" using WhileF\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc')#frs)" using WhileF\<^sub>1 by (auto simp:add.assoc) also have "P \ \ -jvm\ (None,h\<^sub>1,(Unit#vs,ls\<^sub>1,C,M,?pc'+1)#frs)" using WhileF\<^sub>1.prems by (auto simp:eval_nat_numeral) finally show ?case by (simp add:add.assoc eval_nat_numeral) next case (WhileT\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 c v\<^sub>1 h\<^sub>2 ls\<^sub>2 e\<^sub>3 h\<^sub>3 ls\<^sub>3) let ?pc = "pc + length(compE\<^sub>2 e)" let ?pc' = "?pc + length(compE\<^sub>2 c) + 1" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?\\<^sub>2 = "(None,h\<^sub>2,(vs,ls\<^sub>2,C,M,pc)#frs)" have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool True#vs,ls\<^sub>1,C,M,?pc)#frs)" using WhileT\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc+1)#frs)" using WhileT\<^sub>1.prems by auto also have "P \ \ -jvm\ (None,h\<^sub>2,(v\<^sub>1#vs,ls\<^sub>2,C,M,?pc')#frs)" using WhileT\<^sub>1 by(fastforce) also have "P \ \ -jvm\ ?\\<^sub>2" using WhileT\<^sub>1.prems by auto finally have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>2". show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>2 -jvm\ (None,h\<^sub>3,(v#vs,ls\<^sub>3,C,M,?pc'+3)#frs)" using val WhileT\<^sub>1 by (auto simp add:add.assoc eval_nat_numeral) finally show ?trans by(simp add:add.assoc eval_nat_numeral) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof assume throw: ?throw moreover have "PROP ?P (while (e) c) h\<^sub>2 ls\<^sub>2 e\<^sub>3 h\<^sub>3 ls\<^sub>3 C M pc v xa vs frs I" by fact ultimately obtain pc\<^sub>2 where pc\<^sub>2: "pc \ pc\<^sub>2 \ pc\<^sub>2 < ?pc'+3 \ \ caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 (while (e) c) pc (size vs))" and 2: "P \ ?\\<^sub>2 -jvm\ handle P C M xa h\<^sub>3 vs ls\<^sub>3 pc\<^sub>2 frs" using WhileT\<^sub>1.prems by (auto simp:add.assoc eval_nat_numeral) have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF 1 2] by auto thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case WhileCondThrow\<^sub>1 thus ?case by fastforce next case (WhileBodyThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 c e' h\<^sub>2 ls\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1)#frs)" have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool(True)#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using WhileBodyThrow\<^sub>1 by (fastforce simp add: Int_Un_distrib) also have "P \ \ -jvm\ ?\\<^sub>1" using WhileBodyThrow\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". let ?pc\<^sub>1' = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 c)" show ?case (is "?Norm \ ?Err") proof show ?Norm by simp next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof assume throw: ?throw moreover have "PROP ?P c h\<^sub>1 ls\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 C M (?pc\<^sub>1+1) v xa vs frs (I - pcs (compxE\<^sub>2 e pc (size vs)))" by fact ultimately obtain pc\<^sub>2 where pc\<^sub>2: "?pc\<^sub>1+1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1' \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 c (?pc\<^sub>1+1) (size vs))" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 vs ls\<^sub>2 pc\<^sub>2 frs" using WhileBodyThrow\<^sub>1.prems by (fastforce simp:Int_Un_distrib) have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by auto thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case (Throw\<^sub>1 e h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1) let ?pc = "pc + size(compE\<^sub>2 e)" show ?case (is "?Norm \ ?Err") proof show ?Norm by simp next show ?Err (is "?throw \ (\pc\<^sub>1. ?H pc\<^sub>1)") proof assume ?throw hence "P \ (None, h\<^sub>0, (vs, ls\<^sub>0, C, M, pc) # frs) -jvm\ (None, h\<^sub>1, (Addr xa#vs, ls\<^sub>1, C, M, ?pc) # frs)" using Throw\<^sub>1 by fastforce also have "P \ \ -jvm\ handle P C M xa h\<^sub>1 (Addr xa#vs) ls\<^sub>1 ?pc frs" using Throw\<^sub>1.prems by(auto simp add:handle_def) also have "handle P C M xa h\<^sub>1 (Addr xa#vs) ls\<^sub>1 ?pc frs = handle P C M xa h\<^sub>1 vs ls\<^sub>1 ?pc frs" using Throw\<^sub>1.prems by(auto simp add:handle_Cons) finally have "?H ?pc" by simp thus "\pc\<^sub>1. ?H pc\<^sub>1" by iprover qed qed next case (ThrowNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1) let ?pc = "pc + size(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt NullPointer" show ?case (is "?Norm \ ?Err") proof show ?Norm by simp next show ?Err (is "?throw \ (\pc\<^sub>1. ?H pc\<^sub>1)") proof assume throw: ?throw have "P \ (None, h\<^sub>0, (vs, ls\<^sub>0, C, M, pc) # frs) -jvm\ (None, h\<^sub>1, (Null#vs, ls\<^sub>1, C, M, ?pc) # frs)" using ThrowNull\<^sub>1 by fastforce also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc frs" using ThrowNull\<^sub>1.prems by(auto simp add:handle_def) also have "handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc frs = handle P C M ?xa h\<^sub>1 vs ls\<^sub>1 ?pc frs" using ThrowNull\<^sub>1.prems by(auto simp add:handle_Cons) finally have "?H ?pc" using throw by simp thus "\pc\<^sub>1. ?H pc\<^sub>1" by iprover qed qed next case ThrowThrow\<^sub>1 thus ?case by fastforce next case (Try\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 Ci i e\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>1' = "?pc\<^sub>1 + 2 + length(compE\<^sub>2 e\<^sub>2)" have "P,C,M \ compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs) / I,size vs" by fact hence "P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs) / {pc..2 e\<^sub>1)},size vs" using Try\<^sub>1.prems by (fastforce simp:beforex_def split:if_split_asm) hence "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using Try\<^sub>1 by auto also have "P \ \ -jvm\ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1')#frs)" using Try\<^sub>1.prems by auto finally show ?case by (auto simp:add.assoc) next case (TryCatch\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1 D fs Ci i e\<^sub>2 e\<^sub>2' h\<^sub>2 ls\<^sub>2) let ?e = "try e\<^sub>1 catch(Ci i) e\<^sub>2" let ?xt = "compxE\<^sub>2 ?e pc (size vs)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?ls\<^sub>1 = "ls\<^sub>1[i := Addr a]" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>1' = "?pc\<^sub>1 + 2" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,?ls\<^sub>1,C,M, ?pc\<^sub>1') # frs)" have I: "{pc..2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \ I" and beforex: "P,C,M \ ?xt/I,size vs" by fact+ have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,((Addr a)#vs,ls\<^sub>1,C,M, ?pc\<^sub>1+1) # frs)" proof - have "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 (Throw a) h\<^sub>1 ls\<^sub>1 C M pc w a vs frs {pc..2 e\<^sub>1)}" by fact moreover have "P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs)/{pc..1},size vs" using beforex I pcs_subset by(force elim!: beforex_appendD1) ultimately have "\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < ?pc\<^sub>1 \ \ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs)) \ P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 vs ls\<^sub>1 pc\<^sub>1 frs" using TryCatch\<^sub>1.prems by auto then obtain pc\<^sub>1 where pc\<^sub>1_in_e\<^sub>1: "pc \ pc\<^sub>1" "pc\<^sub>1 < ?pc\<^sub>1" and pc\<^sub>1_not_caught: "\ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs))" and 0: "P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 vs ls\<^sub>1 pc\<^sub>1 frs" by iprover from beforex obtain xt\<^sub>0 xt\<^sub>1 where ex_tab: "ex_table_of P C M = xt\<^sub>0 @ ?xt @ xt\<^sub>1" and disj: "pcs xt\<^sub>0 \ I = {}" by(auto simp:beforex_def) have hp: "h\<^sub>1 a = Some (D, fs)" "P\<^sub>1 \ D \\<^sup>* Ci" by fact+ have "pc\<^sub>1 \ pcs xt\<^sub>0" using pc\<^sub>1_in_e\<^sub>1 I disj by auto with pc\<^sub>1_in_e\<^sub>1 pc\<^sub>1_not_caught hp show ?thesis using ex_tab 0 by(simp add:handle_def matches_ex_entry_def) qed also have "P \ \ -jvm\ ?\\<^sub>1" using TryCatch\<^sub>1 by auto finally have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" . let ?pc\<^sub>2 = "?pc\<^sub>1' + length(compE\<^sub>2 e\<^sub>2)" let ?I\<^sub>2 = "{?pc\<^sub>1' ..< ?pc\<^sub>2}" have "P,C,M \ compxE\<^sub>2 ?e pc (size vs) / I,size vs" by fact hence beforex\<^sub>2: "P,C,M \ compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1' (size vs) / ?I\<^sub>2, size vs" using I pcs_subset[of _ ?pc\<^sub>1'] by(auto elim!:beforex_appendD2) have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ?ls\<^sub>1 e\<^sub>2' h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1' v xa vs frs ?I\<^sub>2" by fact show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2)#frs)" using val beforex\<^sub>2 IH\<^sub>2 TryCatch\<^sub>1.prems by auto finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof assume throw: ?throw then obtain pc\<^sub>2 where pc\<^sub>2: "?pc\<^sub>1+2 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2 \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1' (size vs))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 vs ls\<^sub>2 pc\<^sub>2 frs" using IH\<^sub>2 beforex\<^sub>2 TryCatch\<^sub>1.prems by auto have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF 1 2] by (simp add:match_ex_entry) (fastforce) thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case (TryThrow\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1 D fs Ci i e\<^sub>2) let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?e = "try e\<^sub>1 catch(Ci i) e\<^sub>2" let ?xt = "compxE\<^sub>2 ?e pc (size vs)" have I: "{pc..2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \ I" and beforex: "P,C,M \ ?xt/I,size vs" by fact+ have "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 (Throw a) h\<^sub>1 ls\<^sub>1 C M pc w a vs frs {pc..2 e\<^sub>1)}" by fact moreover have "P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs)/{pc..1},size vs" using beforex I pcs_subset by(force elim!: beforex_appendD1) ultimately have "\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < ?pc\<^sub>1 \ \ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs)) \ P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 vs ls\<^sub>1 pc\<^sub>1 frs" using TryThrow\<^sub>1.prems by auto then obtain pc\<^sub>1 where pc\<^sub>1_in_e\<^sub>1: "pc \ pc\<^sub>1" "pc\<^sub>1 < ?pc\<^sub>1" and pc\<^sub>1_not_caught: "\ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs))" and 0: "P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 vs ls\<^sub>1 pc\<^sub>1 frs" by iprover show ?case (is "?N \ (?eq \ (\pc\<^sub>2. ?H pc\<^sub>2))") proof show ?N by simp next { assume ?eq with TryThrow\<^sub>1 pc\<^sub>1_in_e\<^sub>1 pc\<^sub>1_not_caught 0 have "?H pc\<^sub>1" by (simp add:match_ex_entry) auto hence "\pc\<^sub>2. ?H pc\<^sub>2" by iprover } thus "?eq \ (\pc\<^sub>2. ?H pc\<^sub>2)" by iprover qed next case Nil\<^sub>1 thus ?case by simp next case (Cons\<^sub>1 e h\<^sub>0 ls\<^sub>0 v h\<^sub>1 ls\<^sub>1 es fs h\<^sub>2 ls\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?\\<^sub>1 = "(None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using Cons\<^sub>1 by fastforce let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" have IHs: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 fs h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 (tl ws) xa es' (v#vs) frs (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(rev(ws) @ vs,ls\<^sub>2,C,M,?pc\<^sub>2)#frs)" using val IHs Cons\<^sub>1.prems by fastforce finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof assume throw: ?throw then obtain pc\<^sub>2 where pc\<^sub>2: "?pc\<^sub>1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2 \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es ?pc\<^sub>1 (size vs + 1))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (v#vs) ls\<^sub>2 pc\<^sub>2 frs" using IHs Cons\<^sub>1.prems by(fastforce simp:Cons_eq_append_conv neq_Nil_conv) have "?H pc\<^sub>2" using Cons\<^sub>1.prems pc\<^sub>2 jvm_trans[OF 1 2] by (auto simp add: handle_Cons) thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case ConsThrow\<^sub>1 thus ?case by (fastforce simp:Cons_eq_append_conv) qed (*>*) (*FIXME move! *) lemma atLeast0AtMost[simp]: "{0::nat..n} = {..n}" by auto lemma atLeast0LessThan[simp]: "{0::nat.. addr option" where "exception (Throw a) = Some a" | "exception e = None" lemma comp\<^sub>2_correct: assumes "method": "P\<^sub>1 \ C sees M:Ts\T = body in C" and eval: "P\<^sub>1 \\<^sub>1 \body,(h,ls)\ \ \e',(h',ls')\" shows "compP\<^sub>2 P\<^sub>1 \ (None,h,[([],ls,C,M,0)]) -jvm\ (exception e',h',[])" (*<*) (is "_ \ ?\\<^sub>0 -jvm\ ?\\<^sub>1") proof - let ?P = "compP\<^sub>2 P\<^sub>1" have code: "?P,C,M,0 \ compE\<^sub>2 body" using beforeM[OF "method"] by auto have xtab: "?P,C,M \ compxE\<^sub>2 body 0 (size[])/{..2 body)},size[]" using beforexM[OF "method"] by auto \ \Distinguish if e' is a value or an exception\ { fix v assume [simp]: "e' = Val v" have "?P \ ?\\<^sub>0 -jvm\ (None,h',[([v],ls',C,M,size(compE\<^sub>2 body))])" using Jcc[OF eval code xtab] by auto also have "?P \ \ -jvm\ ?\\<^sub>1" using beforeM[OF "method"] by auto finally have ?thesis . } moreover { fix a assume [simp]: "e' = Throw a" obtain pc where pc: "0 \ pc \ pc < size(compE\<^sub>2 body) \ \ caught ?P pc h' a (compxE\<^sub>2 body 0 0)" and 1: "?P \ ?\\<^sub>0 -jvm\ handle ?P C M a h' [] ls' pc []" using Jcc[OF eval code xtab] by fastforce from pc have "handle ?P C M a h' [] ls' pc [] = ?\\<^sub>1" using xtab "method" by(auto simp:handle_def compMb\<^sub>2_def) with 1 have ?thesis by simp } ultimately show ?thesis using eval\<^sub>1_final[OF eval] by(auto simp:final_def) qed (*>*) end diff --git a/thys/Jinja/Compiler/J1.thy b/thys/Jinja/Compiler/J1.thy --- a/thys/Jinja/Compiler/J1.thy +++ b/thys/Jinja/Compiler/J1.thy @@ -1,225 +1,221 @@ (* Title: Jinja/Compiler/J1.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) chapter \Compilation \label{cha:comp}\ section \An Intermediate Language\ theory J1 imports "../J/BigStep" begin type_synonym expr\<^sub>1 = "nat exp" type_synonym J\<^sub>1_prog = "expr\<^sub>1 prog" type_synonym state\<^sub>1 = "heap \ (val list)" primrec max_vars :: "'a exp \ nat" and max_varss :: "'a exp list \ nat" where "max_vars(new C) = 0" | "max_vars(Cast C e) = max_vars e" | "max_vars(Val v) = 0" | "max_vars(e\<^sub>1 \bop\ e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)" | "max_vars(Var V) = 0" | "max_vars(V:=e) = max_vars e" | "max_vars(e\F{D}) = max_vars e" | "max_vars(FAss e\<^sub>1 F D e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)" | "max_vars(e\M(es)) = max (max_vars e) (max_varss es)" | "max_vars({V:T; e}) = max_vars e + 1" | "max_vars(e\<^sub>1;;e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)" | "max_vars(if (e) e\<^sub>1 else e\<^sub>2) = max (max_vars e) (max (max_vars e\<^sub>1) (max_vars e\<^sub>2))" | "max_vars(while (b) e) = max (max_vars b) (max_vars e)" | "max_vars(throw e) = max_vars e" | "max_vars(try e\<^sub>1 catch(C V) e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2 + 1)" | "max_varss [] = 0" | "max_varss (e#es) = max (max_vars e) (max_varss es)" inductive eval\<^sub>1 :: "J\<^sub>1_prog \ expr\<^sub>1 \ state\<^sub>1 \ expr\<^sub>1 \ state\<^sub>1 \ bool" ("_ \\<^sub>1 ((1\_,/_\) \/ (1\_,/_\))" [51,0,0,0,0] 81) and evals\<^sub>1 :: "J\<^sub>1_prog \ expr\<^sub>1 list \ state\<^sub>1 \ expr\<^sub>1 list \ state\<^sub>1 \ bool" ("_ \\<^sub>1 ((1\_,/_\) [\]/ (1\_,/_\))" [51,0,0,0,0] 81) for P :: J\<^sub>1_prog where New\<^sub>1: "\ new_Addr h = Some a; P \ C has_fields FDTs; h' = h(a\(C,init_fields FDTs)) \ \ P \\<^sub>1 \new C,(h,l)\ \ \addr a,(h',l)\" | NewFail\<^sub>1: "new_Addr h = None \ P \\<^sub>1 \new C, (h,l)\ \ \THROW OutOfMemory,(h,l)\" | Cast\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,l)\; h a = Some(D,fs); P \ D \\<^sup>* C \ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \addr a,(h,l)\" | CastNull\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \null,s\<^sub>1\" | CastFail\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,l)\; h a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \THROW ClassCast,(h,l)\" | CastThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | Val\<^sub>1: "P \\<^sub>1 \Val v,s\ \ \Val v,s\" | BinOp\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v\<^sub>2,s\<^sub>2\; binop(bop,v\<^sub>1,v\<^sub>2) = Some v \ \ P \\<^sub>1 \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\ \ \Val v,s\<^sub>2\" | BinOpThrow\<^sub>1\<^sub>1: "P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \throw e,s\<^sub>1\ \ P \\<^sub>1 \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0\ \ \throw e,s\<^sub>1\" | BinOpThrow\<^sub>2\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \throw e,s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\ \ \throw e,s\<^sub>2\" | Var\<^sub>1: "\ ls!i = v; i < size ls \ \ P \\<^sub>1 \Var i,(h,ls)\ \ \Val v,(h,ls)\" | LAss\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \Val v,(h,ls)\; i < size ls; ls' = ls[i := v] \ \ P \\<^sub>1 \i:= e,s\<^sub>0\ \ \unit,(h,ls')\" | LAssThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \i:= e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAcc\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,ls)\; h a = Some(C,fs); fs(F,D) = Some v \ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \Val v,(h,ls)\" | FAccNull\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \THROW NullPointer,s\<^sub>1\" | FAccThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAss\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v,(h\<^sub>2,l\<^sub>2)\; h\<^sub>2 a = Some(C,fs); fs' = fs((F,D)\v); h\<^sub>2' = h\<^sub>2(a\(C,fs')) \ \ P \\<^sub>1 \e\<^sub>1\F{D}:= e\<^sub>2,s\<^sub>0\ \ \unit,(h\<^sub>2',l\<^sub>2)\" | FAssNull\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \null,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v,s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>1\F{D}:= e\<^sub>2,s\<^sub>0\ \ \THROW NullPointer,s\<^sub>2\" | FAssThrow\<^sub>1\<^sub>1: "P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e\<^sub>1\F{D}:= e\<^sub>2,s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAssThrow\<^sub>2\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \throw e',s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>1\F{D}:= e\<^sub>2,s\<^sub>0\ \ \throw e',s\<^sub>2\" | CallObjThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \throw e',s\<^sub>1\" | CallNull\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \map Val vs,s\<^sub>2\ \ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \THROW NullPointer,s\<^sub>2\" | Call\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \map Val vs,(h\<^sub>2,ls\<^sub>2)\; h\<^sub>2 a = Some(C,fs); P \ C sees M:Ts\T = body in D; size vs = size Ts; ls\<^sub>2' = (Addr a) # vs @ replicate (max_vars body) undefined; P \\<^sub>1 \body,(h\<^sub>2,ls\<^sub>2')\ \ \e',(h\<^sub>3,ls\<^sub>3)\ \ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \e',(h\<^sub>3,ls\<^sub>2)\" | CallParamsThrow\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \es',s\<^sub>2\; es' = map Val vs @ throw ex # es\<^sub>2 \ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \throw ex,s\<^sub>2\" | Block\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \e',s\<^sub>1\ \ P \\<^sub>1 \Block i T e,s\<^sub>0\ \ \e',s\<^sub>1\" | Seq\<^sub>1: "\ P \\<^sub>1 \e\<^sub>0,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \e\<^sub>1,s\<^sub>1\ \ \e\<^sub>2,s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>0;;e\<^sub>1,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" | SeqThrow\<^sub>1: "P \\<^sub>1 \e\<^sub>0,s\<^sub>0\ \ \throw e,s\<^sub>1\ \ P \\<^sub>1 \e\<^sub>0;;e\<^sub>1,s\<^sub>0\ \ \throw e,s\<^sub>1\" | CondT\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \\<^sub>1 \e\<^sub>1,s\<^sub>1\ \ \e',s\<^sub>2\ \ \ P \\<^sub>1 \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \ \e',s\<^sub>2\" | CondF\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \false,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \e',s\<^sub>2\ \ \ P \\<^sub>1 \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \ \e',s\<^sub>2\" | CondThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\ \ \throw e',s\<^sub>1\" | WhileF\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \false,s\<^sub>1\ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \unit,s\<^sub>1\" | WhileT\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \\<^sub>1 \c,s\<^sub>1\ \ \Val v\<^sub>1,s\<^sub>2\; P \\<^sub>1 \while (e) c,s\<^sub>2\ \ \e\<^sub>3,s\<^sub>3\ \ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \e\<^sub>3,s\<^sub>3\" | WhileCondThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \throw e',s\<^sub>1\" | WhileBodyThrow\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \\<^sub>1 \c,s\<^sub>1\ \ \throw e',s\<^sub>2\\ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \throw e',s\<^sub>2\" | Throw\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,s\<^sub>1\ \ P \\<^sub>1 \throw e,s\<^sub>0\ \ \Throw a,s\<^sub>1\" | ThrowNull\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \\<^sub>1 \throw e,s\<^sub>0\ \ \THROW NullPointer,s\<^sub>1\" | ThrowThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \throw e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | Try\<^sub>1: "P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\ \ P \\<^sub>1 \try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\" | TryCatch\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Throw a,(h\<^sub>1,ls\<^sub>1)\; h\<^sub>1 a = Some(D,fs); P \ D \\<^sup>* C; i < length ls\<^sub>1; P \\<^sub>1 \e\<^sub>2,(h\<^sub>1,ls\<^sub>1[i:=Addr a])\ \ \e\<^sub>2',(h\<^sub>2,ls\<^sub>2)\ \ \ P \\<^sub>1 \try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',(h\<^sub>2,ls\<^sub>2)\" | TryThrow\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Throw a,(h\<^sub>1,ls\<^sub>1)\; h\<^sub>1 a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \\<^sub>1 \try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\ \ \Throw a,(h\<^sub>1,ls\<^sub>1)\" | Nil\<^sub>1: "P \\<^sub>1 \[],s\ [\] \[],s\" | Cons\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \es',s\<^sub>2\ \ \ P \\<^sub>1 \e#es,s\<^sub>0\ [\] \Val v # es',s\<^sub>2\" | ConsThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e#es,s\<^sub>0\ [\] \throw e' # es, s\<^sub>1\" (*<*) lemmas eval\<^sub>1_evals\<^sub>1_induct = eval\<^sub>1_evals\<^sub>1.induct [split_format (complete)] and eval\<^sub>1_evals\<^sub>1_inducts = eval\<^sub>1_evals\<^sub>1.inducts [split_format (complete)] (*>*) lemma eval\<^sub>1_preserves_len: "P \\<^sub>1 \e\<^sub>0,(h\<^sub>0,ls\<^sub>0)\ \ \e\<^sub>1,(h\<^sub>1,ls\<^sub>1)\ \ length ls\<^sub>0 = length ls\<^sub>1" and evals\<^sub>1_preserves_len: "P \\<^sub>1 \es\<^sub>0,(h\<^sub>0,ls\<^sub>0)\ [\] \es\<^sub>1,(h\<^sub>1,ls\<^sub>1)\ \ length ls\<^sub>0 = length ls\<^sub>1" (*<*)by (induct rule:eval\<^sub>1_evals\<^sub>1_inducts, simp_all)(*>*) lemma evals\<^sub>1_preserves_elen: "\es' s s'. P \\<^sub>1 \es,s\ [\] \es',s'\ \ length es = length es'" -(*<*) -apply(induct es type:list) -apply (auto elim:evals\<^sub>1.cases) -done -(*>*) +(*<*)by(induct es type:list) (auto elim:evals\<^sub>1.cases)(*>*) lemma eval\<^sub>1_final: "P \\<^sub>1 \e,s\ \ \e',s'\ \ final e'" and evals\<^sub>1_final: "P \\<^sub>1 \es,s\ [\] \es',s'\ \ finals es'" (*<*)by(induct rule:eval\<^sub>1_evals\<^sub>1.inducts, simp_all)(*>*) end diff --git a/thys/Jinja/Compiler/J1WellForm.thy b/thys/Jinja/Compiler/J1WellForm.thy --- a/thys/Jinja/Compiler/J1WellForm.thy +++ b/thys/Jinja/Compiler/J1WellForm.thy @@ -1,223 +1,209 @@ (* Title: Jinja/Compiler/WellType1.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Well-Formedness of Intermediate Language\ theory J1WellForm imports "../J/JWellForm" J1 begin subsection "Well-Typedness" type_synonym env\<^sub>1 = "ty list" \ \type environment indexed by variable number\ inductive WT\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1, expr\<^sub>1 , ty ] \ bool" ("(_,_ \\<^sub>1/ _ :: _)" [51,51,51]50) and WTs\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1, expr\<^sub>1 list, ty list] \ bool" ("(_,_ \\<^sub>1/ _ [::] _)" [51,51,51]50) for P :: J\<^sub>1_prog where WTNew\<^sub>1: "is_class P C \ P,E \\<^sub>1 new C :: Class C" | WTCast\<^sub>1: "\ P,E \\<^sub>1 e :: Class D; is_class P C; P \ C \\<^sup>* D \ P \ D \\<^sup>* C \ \ P,E \\<^sub>1 Cast C e :: Class C" | WTVal\<^sub>1: "typeof v = Some T \ P,E \\<^sub>1 Val v :: T" | WTVar\<^sub>1: "\ E!i = T; i < size E \ \ P,E \\<^sub>1 Var i :: T" | WTBinOp\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1 :: T\<^sub>1; P,E \\<^sub>1 e\<^sub>2 :: T\<^sub>2; case bop of Eq \ (P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1) \ T = Boolean | Add \ T\<^sub>1 = Integer \ T\<^sub>2 = Integer \ T = Integer \ \ P,E \\<^sub>1 e\<^sub>1 \bop\ e\<^sub>2 :: T" | WTLAss\<^sub>1: "\ E!i = T; i < size E; P,E \\<^sub>1 e :: T'; P \ T' \ T \ \ P,E \\<^sub>1 i:=e :: Void" | WTFAcc\<^sub>1: "\ P,E \\<^sub>1 e :: Class C; P \ C sees F:T in D \ \ P,E \\<^sub>1 e\F{D} :: T" | WTFAss\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1 :: Class C; P \ C sees F:T in D; P,E \\<^sub>1 e\<^sub>2 :: T'; P \ T' \ T \ \ P,E \\<^sub>1 e\<^sub>1\F{D} := e\<^sub>2 :: Void" | WTCall\<^sub>1: "\ P,E \\<^sub>1 e :: Class C; P \ C sees M:Ts' \ T = m in D; P,E \\<^sub>1 es [::] Ts; P \ Ts [\] Ts' \ \ P,E \\<^sub>1 e\M(es) :: T" | WTBlock\<^sub>1: "\ is_type P T; P,E@[T] \\<^sub>1 e::T' \ \ P,E \\<^sub>1 {i:T; e} :: T'" | WTSeq\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1::T\<^sub>1; P,E \\<^sub>1 e\<^sub>2::T\<^sub>2 \ \ P,E \\<^sub>1 e\<^sub>1;;e\<^sub>2 :: T\<^sub>2" | WTCond\<^sub>1: "\ P,E \\<^sub>1 e :: Boolean; P,E \\<^sub>1 e\<^sub>1::T\<^sub>1; P,E \\<^sub>1 e\<^sub>2::T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1; P \ T\<^sub>1 \ T\<^sub>2 \ T = T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ T = T\<^sub>1 \ \ P,E \\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 :: T" | WTWhile\<^sub>1: "\ P,E \\<^sub>1 e :: Boolean; P,E \\<^sub>1 c::T \ \ P,E \\<^sub>1 while (e) c :: Void" | WTThrow\<^sub>1: "P,E \\<^sub>1 e :: Class C \ P,E \\<^sub>1 throw e :: Void" | WTTry\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1 :: T; P,E@[Class C] \\<^sub>1 e\<^sub>2 :: T; is_class P C \ \ P,E \\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 :: T" | WTNil\<^sub>1: "P,E \\<^sub>1 [] [::] []" | WTCons\<^sub>1: "\ P,E \\<^sub>1 e :: T; P,E \\<^sub>1 es [::] Ts \ \ P,E \\<^sub>1 e#es [::] T#Ts" (*<*) declare WT\<^sub>1_WTs\<^sub>1.intros[intro!] declare WTNil\<^sub>1[iff] lemmas WT\<^sub>1_WTs\<^sub>1_induct = WT\<^sub>1_WTs\<^sub>1.induct [split_format (complete)] and WT\<^sub>1_WTs\<^sub>1_inducts = WT\<^sub>1_WTs\<^sub>1.inducts [split_format (complete)] inductive_cases eee[elim!]: "P,E \\<^sub>1 Val v :: T" "P,E \\<^sub>1 Var i :: T" "P,E \\<^sub>1 Cast D e :: T" "P,E \\<^sub>1 i:=e :: T" "P,E \\<^sub>1 {i:U; e} :: T" "P,E \\<^sub>1 e\<^sub>1;;e\<^sub>2 :: T" "P,E \\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 :: T" "P,E \\<^sub>1 while (e) c :: T" "P,E \\<^sub>1 throw e :: T" "P,E \\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 :: T" "P,E \\<^sub>1 e\F{D} :: T" "P,E \\<^sub>1 e\<^sub>1\F{D}:=e\<^sub>2 :: T" "P,E \\<^sub>1 e\<^sub>1 \bop\ e\<^sub>2 :: T" "P,E \\<^sub>1 new C :: T" "P,E \\<^sub>1 e\M(es) :: T" "P,E \\<^sub>1 [] [::] Ts" "P,E \\<^sub>1 e#es [::] Ts" (*>*) lemma WTs\<^sub>1_same_size: "\Ts. P,E \\<^sub>1 es [::] Ts \ size es = size Ts" (*<*)by (induct es type:list) auto(*>*) lemma WT\<^sub>1_unique: "P,E \\<^sub>1 e :: T\<^sub>1 \ (\T\<^sub>2. P,E \\<^sub>1 e :: T\<^sub>2 \ T\<^sub>1 = T\<^sub>2)" and "P,E \\<^sub>1 es [::] Ts\<^sub>1 \ (\Ts\<^sub>2. P,E \\<^sub>1 es [::] Ts\<^sub>2 \ Ts\<^sub>1 = Ts\<^sub>2)" (*<*) -apply(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) -apply blast -apply blast -apply clarsimp -apply blast -apply clarsimp -apply(case_tac bop) -apply clarsimp -apply clarsimp -apply blast -apply (blast dest:sees_field_idemp sees_field_fun) -apply blast -apply (blast dest:sees_method_idemp sees_method_fun) -apply blast -apply blast -apply blast -apply blast -apply clarify -apply blast -apply blast -apply blast -done +proof(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) + case WTVal\<^sub>1 then show ?case by clarsimp +next + case (WTBinOp\<^sub>1 E e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 bop T) + then show ?case by(case_tac bop) force+ +next + case WTFAcc\<^sub>1 then show ?case + by (blast dest:sees_field_idemp sees_field_fun) +next + case WTCall\<^sub>1 then show ?case + by (blast dest:sees_method_idemp sees_method_fun) +qed blast+ (*>*) lemma assumes wf: "wf_prog p P" shows WT\<^sub>1_is_type: "P,E \\<^sub>1 e :: T \ set E \ types P \ is_type P T" and "P,E \\<^sub>1 es [::] Ts \ True" (*<*) -apply(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) -apply simp -apply simp -apply (simp add:typeof_lit_is_type) -apply (blast intro:nth_mem) -apply(simp split:bop.splits) -apply simp -apply (simp add:sees_field_is_type[OF _ wf]) -apply simp -apply(fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def) -apply simp -apply simp -apply blast -apply simp -apply simp -apply simp -apply simp -apply simp -done +proof(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) + case WTVal\<^sub>1 then show ?case by (simp add:typeof_lit_is_type) +next + case WTVar\<^sub>1 then show ?case by (blast intro:nth_mem) +next + case WTBinOp\<^sub>1 then show ?case by (simp split:bop.splits) +next + case WTFAcc\<^sub>1 then show ?case + by (simp add:sees_field_is_type[OF _ wf]) +next + case WTCall\<^sub>1 then show ?case + by (fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def) +next + case WTCond\<^sub>1 then show ?case by blast +qed simp+ (*>*) subsection\Well-formedness\ \ \Indices in blocks increase by 1\ primrec \ :: "expr\<^sub>1 \ nat \ bool" and \s :: "expr\<^sub>1 list \ nat \ bool" where "\ (new C) i = True" | "\ (Cast C e) i = \ e i" | "\ (Val v) i = True" | "\ (e\<^sub>1 \bop\ e\<^sub>2) i = (\ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (Var j) i = True" | "\ (e\F{D}) i = \ e i" | "\ (j:=e) i = \ e i" | "\ (e\<^sub>1\F{D} := e\<^sub>2) i = (\ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (e\M(es)) i = (\ e i \ \s es i)" | "\ ({j:T ; e}) i = (i = j \ \ e (i+1))" | "\ (e\<^sub>1;;e\<^sub>2) i = (\ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (if (e) e\<^sub>1 else e\<^sub>2) i = (\ e i \ \ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (throw e) i = \ e i" | "\ (while (e) c) i = (\ e i \ \ c i)" | "\ (try e\<^sub>1 catch(C j) e\<^sub>2) i = (\ e\<^sub>1 i \ i=j \ \ e\<^sub>2 (i+1))" | "\s [] i = True" | "\s (e#es) i = (\ e i \ \s es i)" definition wf_J\<^sub>1_mdecl :: "J\<^sub>1_prog \ cname \ expr\<^sub>1 mdecl \ bool" where "wf_J\<^sub>1_mdecl P C \ \(M,Ts,T,body). (\T'. P,Class C#Ts \\<^sub>1 body :: T' \ P \ T' \ T) \ \ body \{..size Ts}\ \ \ body (size Ts + 1)" lemma wf_J\<^sub>1_mdecl[simp]: "wf_J\<^sub>1_mdecl P C (M,Ts,T,body) \ ((\T'. P,Class C#Ts \\<^sub>1 body :: T' \ P \ T' \ T) \ \ body \{..size Ts}\ \ \ body (size Ts + 1))" (*<*)by (simp add:wf_J\<^sub>1_mdecl_def)(*>*) abbreviation "wf_J\<^sub>1_prog == wf_prog wf_J\<^sub>1_mdecl" end diff --git a/thys/Jinja/Compiler/PCompiler.thy b/thys/Jinja/Compiler/PCompiler.thy --- a/thys/Jinja/Compiler/PCompiler.thy +++ b/thys/Jinja/Compiler/PCompiler.thy @@ -1,278 +1,350 @@ (* Title: Jinja/Compiler/PCompiler.thy Author: Tobias Nipkow Copyright TUM 2003 *) section \Program Compilation\ theory PCompiler imports "../Common/WellForm" begin definition compM :: "('a \ 'b) \ 'a mdecl \ 'b mdecl" where "compM f \ \(M, Ts, T, m). (M, Ts, T, f m)" definition compC :: "('a \ 'b) \ 'a cdecl \ 'b cdecl" where "compC f \ \(C,D,Fdecls,Mdecls). (C,D,Fdecls, map (compM f) Mdecls)" definition compP :: "('a \ 'b) \ 'a prog \ 'b prog" where "compP f \ map (compC f)" text\Compilation preserves the program structure. Therfore lookup functions either commute with compilation (like method lookup) or are preserved by it (like the subclass relation).\ lemma map_of_map4: "map_of (map (\(x,a,b,c).(x,a,b,f c)) ts) = map_option (\(a,b,c).(a,b,f c)) \ (map_of ts)" (*<*) -apply(induct ts) - apply simp -apply(rule ext) -apply fastforce -done +proof(induct ts) + case Nil then show ?case by simp +qed fastforce (*>*) lemma class_compP: "class P C = Some (D, fs, ms) \ class (compP f P) C = Some (D, fs, map (compM f) ms)" (*<*)by(simp add:class_def compP_def compC_def map_of_map4)(*>*) lemma class_compPD: "class (compP f P) C = Some (D, fs, cms) \ \ms. class P C = Some(D,fs,ms) \ cms = map (compM f) ms" (*<*)by(clarsimp simp add:class_def compP_def compC_def map_of_map4)(*>*) lemma [simp]: "is_class (compP f P) C = is_class P C" (*<*)by(auto simp:is_class_def dest: class_compP class_compPD)(*>*) lemma [simp]: "class (compP f P) C = map_option (\c. snd(compC f (C,c))) (class P C)" (*<*) -apply(simp add:compP_def compC_def class_def map_of_map4) -apply(simp add:split_def) -done +by(simp add:compP_def compC_def class_def map_of_map4) + (simp add:split_def) (*>*) lemma sees_methods_compP: "P \ C sees_methods Mm \ compP f P \ C sees_methods (map_option (\((Ts,T,m),D). ((Ts,T,f m),D)) \ Mm)" -(*<*) -apply(erule Methods.induct) - apply(rule sees_methods_Object) - apply(erule class_compP) - apply(rule ext) - apply(simp add:compM_def map_of_map4 option.map_comp) - apply(case_tac "map_of ms x") - apply simp - apply fastforce -apply(rule sees_methods_rec) - apply(erule class_compP) - apply assumption - apply assumption -apply(rule ext) -apply(simp add:map_add_def compM_def map_of_map4 option.map_comp split:option.split) -done +(*<*)(is "?P \ compP f P \ C sees_methods (?map Mm)") +proof(induct rule: Methods.induct) + case Object: (sees_methods_Object D fs ms Mm) + let ?Mm1 = "\x. map_option ((\m. (m, Object)) \ (\(Ts, T, m). (Ts, T, f m))) (map_of ms x)" + let ?Mm2 = "\x. map_option (case_prod (\(Ts, T, m). + Pair (Ts, T, f m)) \ (\m. (m, Object))) (map_of ms x)" + have Mm_eq: "\x. ?Mm1 x = ?Mm2 x" + proof - + fix x show "?Mm1 x = ?Mm2 x" + proof(cases "map_of ms x") + case None then show ?thesis by simp + qed fastforce + qed + + have Mm: "Mm = map_option (\m. (m, Object)) \ map_of ms" by fact + let ?Mm = "map_option (\m. (m, Object)) \ map_of (map (compM f) ms)" + let ?Mm' = "?map Mm" + have "?Mm' = ?Mm" + by(rule ext) (simp add:Mm Mm_eq compM_def map_of_map4 option.map_comp) + then show ?case by(rule sees_methods_Object[OF class_compP[OF Object(1)]]) +next + case rec: (sees_methods_rec C D fs ms Mm Mm') + have Mm': "Mm' = Mm ++ (map_option (\m. (m, C)) \ map_of ms)" by fact + let ?Mm' = "?map Mm'" + let ?Mm'' = "(?map Mm) ++ (map_option (\m. (m, C)) \ map_of (map (compM f) ms))" + have "?Mm' = ?Mm''" + by(rule ext) (simp add:Mm' map_add_def compM_def map_of_map4) + moreover have "compP f P \ C sees_methods ?Mm''" + using sees_methods_rec[OF class_compP[OF rec(1)] rec(2,4)] by fast + ultimately show "compP f P \ C sees_methods ?Mm'" by simp +qed (*>*) lemma sees_method_compP: "P \ C sees M: Ts\T = m in D \ compP f P \ C sees M: Ts\T = (f m) in D" (*<*)by(fastforce elim:sees_methods_compP simp add:Method_def)(*>*) lemma [simp]: "P \ C sees M: Ts\T = m in D \ method (compP f P) C M = (D,Ts,T,f m)" (*<*) -apply(drule sees_method_compP) -apply(simp add:method_def) -apply(rule the_equality) - apply simp -apply(fastforce dest:sees_method_fun) -done +proof - + let ?P = "\(D, Ts, T, m). compP f P \ C sees M: Ts\T = m in D" + let ?a = "(D, Ts, T, f m)" + assume cM: "P \ C sees M: Ts\T = m in D" + have compP_cM: "?P ?a" using sees_method_compP[OF cM] by simp + moreover { + fix x assume "?P x" then have "x = ?a" + using compP_cM by(fastforce dest:sees_method_fun) + } + ultimately have "(THE x. ?P x) = ?a" by(rule the_equality) + then show ?thesis by(simp add:method_def) +qed (*>*) lemma sees_methods_compPD: "\ cP \ C sees_methods Mm'; cP = compP f P \ \ \Mm. P \ C sees_methods Mm \ Mm' = (map_option (\((Ts,T,m),D). ((Ts,T,f m),D)) \ Mm)" -(*<*) -apply(erule Methods.induct) - apply(clarsimp simp:compC_def) - apply(rule exI) - apply(rule conjI, erule sees_methods_Object) - apply(rule refl) - apply(rule ext) - apply(simp add:compM_def map_of_map4 option.map_comp) - apply(case_tac "map_of b x") - apply simp - apply fastforce -apply(clarsimp simp:compC_def) -apply(rule exI, rule conjI) -apply(erule (2) sees_methods_rec) - apply(rule refl) -apply(rule ext) -apply(simp add:map_add_def compM_def map_of_map4 option.map_comp split:option.split) -done +(*<*)(is "\ ?P; ?Q \ \ \Mm. P \ C sees_methods Mm \ Mm' = (?map Mm)") +proof(induct rule: Methods.induct) + case Object: (sees_methods_Object D fs ms Mm) + then obtain ms' where P_Obj: "class P Object = \(D, fs, ms')\" + and ms: "ms = map (compM f) ms'" by(clarsimp simp:compC_def) + + let ?Mm1 = "\x. map_option ((\m. (m, Object)) \ (\(Ts, T, m). (Ts, T, f m))) (map_of ms' x)" + let ?Mm2 = "\x. map_option (case_prod (\(Ts, T, m). Pair (Ts, T, f m)) \ (\m. (m, Object))) + (map_of ms' x)" + have Mm_eq: "\x. ?Mm1 x = ?Mm2 x" + proof - + fix x show "?Mm1 x = ?Mm2 x" + proof(cases "map_of ms' x") + case None then show ?thesis by simp + qed fastforce + qed + + let ?Mm = "map_option (\m. (m,Object)) \ map_of ms'" + let ?Mm' = "?map ?Mm" + have Mm: "Mm = map_option (\m. (m, Object)) \ map_of ms" by fact + have "P \ Object sees_methods ?Mm" + using sees_methods_Object[OF P_Obj] by simp + moreover have "Mm = ?Mm'" + by(rule ext) (simp add:Mm_eq Mm ms compM_def map_of_map4 option.map_comp) + ultimately show ?case by fast +next + case rec: (sees_methods_rec C D fs ms Mm Mm') + then obtain ms' Mm\<^sub>D where P_D: "class P C = \(D, fs, ms')\" + and ms: "ms = map (compM f) ms'" and C_nObj: "C \ Object" + and Mm\<^sub>D: "P \ D sees_methods Mm\<^sub>D" + and Mm: "Mm = (\a. map_option (case_prod (\(Ts, T, m). Pair (Ts, T, f m))) (Mm\<^sub>D a))" + by(clarsimp simp:compC_def) + + let ?Mm = "Mm\<^sub>D ++ (map_option (\m. (m, C)) \ map_of ms')" + let ?Mm1 = "Mm ++ (map_option (\m. (m, C)) \ map_of ms)" + let ?Mm2 = "Mm ++ (map_option (\m. (m, C)) \ map_of (map (compM f) ms'))" + let ?Mm3 = "?map ?Mm" + have "Mm' = ?Mm1" by fact + also have "\ = ?Mm2" using ms by simp + also have "\ = ?Mm3" + by(rule ext)(simp add:Mm map_add_def compM_def map_of_map4) + moreover have "P \ C sees_methods ?Mm" + using sees_methods_rec[OF P_D C_nObj Mm\<^sub>D] by simp + ultimately show ?case by fast +qed (*>*) lemma sees_method_compPD: "compP f P \ C sees M: Ts\T = fm in D \ \m. P \ C sees M: Ts\T = m in D \ f m = fm" (*<*) -apply(simp add:Method_def) -apply clarify -apply(drule sees_methods_compPD[OF _ refl]) -apply clarsimp -apply blast -done +proof - + assume "compP f P \ C sees M: Ts\T = fm in D" + then obtain Mm where Mm: "compP f P \ C sees_methods Mm" + and MmM: "Mm M = \((Ts, T, fm), D)\" + by(clarsimp simp:Method_def) + show ?thesis using sees_methods_compPD[OF Mm refl] MmM + by(fastforce simp: Method_def) +qed (*>*) lemma [simp]: "subcls1(compP f P) = subcls1 P" (*<*) by(fastforce simp add: is_class_def compC_def intro:subcls1I order_antisym dest:subcls1D) (*>*) lemma compP_widen[simp]: "(compP f P \ T \ T') = (P \ T \ T')" (*<*)by(cases T')(simp_all add:widen_Class)(*>*) lemma [simp]: "(compP f P \ Ts [\] Ts') = (P \ Ts [\] Ts')" (*<*) -apply(induct Ts) - apply simp -apply(cases Ts') -apply(auto simp:fun_of_def) -done +proof(induct Ts) + case (Cons a Ts) + then show ?case by(cases Ts')(auto simp:fun_of_def) +qed simp (*>*) lemma [simp]: "is_type (compP f P) T = is_type P T" (*<*)by(cases T) simp_all(*>*) lemma [simp]: "(compP (f::'a\'b) P \ C has_fields FDTs) = (P \ C has_fields FDTs)" (*<*) (is "?A = ?B") proof { fix cP::"'b prog" assume "cP \ C has_fields FDTs" hence "cP = compP f P \ P \ C has_fields FDTs" proof induct case has_fields_Object thus ?case by(fast intro:Fields.has_fields_Object dest:class_compPD) next case has_fields_rec thus ?case by(fast intro:Fields.has_fields_rec dest:class_compPD) qed } note lem = this assume ?A with lem show ?B by blast next assume ?B thus ?A proof induct case has_fields_Object thus ?case by(fast intro:Fields.has_fields_Object class_compP) next case has_fields_rec thus ?case by(fast intro:Fields.has_fields_rec class_compP) qed qed (*>*) lemma [simp]: "fields (compP f P) C = fields P C" (*<*)by(simp add:fields_def)(*>*) lemma [simp]: "(compP f P \ C sees F:T in D) = (P \ C sees F:T in D)" (*<*)by(simp add:sees_field_def)(*>*) lemma [simp]: "field (compP f P) F D = field P F D" (*<*)by(simp add:field_def)(*>*) subsection\Invariance of @{term wf_prog} under compilation\ lemma [iff]: "distinct_fst (compP f P) = distinct_fst P" (*<*) -apply(simp add:distinct_fst_def compP_def compC_def) -apply(induct P) -apply (auto simp:image_iff) -done +by (induct P) + (auto simp:distinct_fst_def compP_def compC_def image_iff) (*>*) lemma [iff]: "distinct_fst (map (compM f) ms) = distinct_fst ms" (*<*) -apply(simp add:distinct_fst_def compM_def) -apply(induct ms) -apply (auto simp:image_iff) -done +by (induct ms) + (auto simp:distinct_fst_def compM_def image_iff) (*>*) lemma [iff]: "wf_syscls (compP f P) = wf_syscls P" (*<*)by(simp add:wf_syscls_def compP_def compC_def image_def Bex_def)(*>*) lemma [iff]: "wf_fdecl (compP f P) = wf_fdecl P" (*<*)by(simp add:wf_fdecl_def)(*>*) lemma set_compP: "((C,D,fs,ms') \ set(compP f P)) = (\ms. (C,D,fs,ms) \ set P \ ms' = map (compM f) ms)" (*<*)by(fastforce simp add:compP_def compC_def image_iff Bex_def)(*>*) lemma wf_cdecl_compPI: "\ \C M Ts T m. \ wf_mdecl wf\<^sub>1 P C (M,Ts,T,m); P \ C sees M:Ts\T = m in C \ \ wf_mdecl wf\<^sub>2 (compP f P) C (M,Ts,T, f m); \x\set P. wf_cdecl wf\<^sub>1 P x; x \ set (compP f P); wf_prog p P \ \ wf_cdecl wf\<^sub>2 (compP f P) x" (*<*) -apply(clarsimp simp add:wf_cdecl_def Ball_def set_compP) -apply(rename_tac C D fs ms) -apply(rule conjI) - apply (clarsimp simp:compM_def) - apply (drule (2) mdecl_visible) - apply simp -apply(clarify) -apply(drule sees_method_compPD[where f = f]) -apply clarsimp -apply(fastforce simp:image_iff compM_def) -done +proof - + assume + wfm: "\C M Ts T m. \ wf_mdecl wf\<^sub>1 P C (M,Ts,T,m); P \ C sees M:Ts\T = m in C \ + \ wf_mdecl wf\<^sub>2 (compP f P) C (M,Ts,T, f m)" + and wfc: "\x\set P. wf_cdecl wf\<^sub>1 P x" + and compP: "x \ set (compP f P)" and wf: "wf_prog p P" + obtain C D fs ms where x: "x = (C, D, fs, map (compM f) ms)" + and x_set: "(C, D, fs, ms) \ set P" + using compP by(case_tac x) (clarsimp simp: set_compP) + have wfc': "wf_cdecl wf\<^sub>1 P (C, D, fs, ms)" using wfc x_set by fast + let ?P = "compP f P" and ?ms = "compM f ` set ms" + { fix M Ts T m + assume M: "(M,Ts,T,m) \ set ms" + then have "wf_mdecl wf\<^sub>1 P C (M, Ts, T, m)" using wfc' + by(simp add:wf_cdecl_def) + moreover have cM: "P \ C sees M : Ts\T = m in C" using M + by(rule mdecl_visible[OF wf x_set]) + ultimately have "wf_mdecl wf\<^sub>2 (compP f P) C (M, Ts, T, f m)" + by(rule wfm) + } + then have "\m \ ?ms. wf_mdecl wf\<^sub>2 ?P C m" + by (clarsimp simp:compM_def) + moreover have "C \ Object \ + (\(M,Ts,T,m)\?ms. + \D' Ts' T' m'. ?P \ D sees M:Ts' \ T' = m' in D' \ + P \ Ts' [\] Ts \ P \ T \ T')" + proof - + { fix M Ts T m D' Ts' T' m' + assume "C \ Object" and "(M,Ts,T,m)\?ms" + and dM: "?P \ D sees M:Ts' \ T' = m' in D'" + then have "P \ Ts' [\] Ts \ P \ T \ T'" + using wfc' sees_method_compPD[OF dM] + by(fastforce simp:wf_cdecl_def image_iff compM_def) + } + then show ?thesis by fast + qed + moreover have "(\f\set fs. wf_fdecl P f) \ distinct_fst fs \ distinct_fst ms + \ (C \ Object \ is_class P D \ \ P \ D \\<^sup>* C)" using wfc' + by(simp add: wf_cdecl_def) + ultimately show ?thesis using x by(simp add:wf_cdecl_def) +qed (*>*) lemma wf_prog_compPI: assumes lift: "\C M Ts T m. \ P \ C sees M:Ts\T = m in C; wf_mdecl wf\<^sub>1 P C (M,Ts,T,m) \ \ wf_mdecl wf\<^sub>2 (compP f P) C (M,Ts,T, f m)" and wf: "wf_prog wf\<^sub>1 P" shows "wf_prog wf\<^sub>2 (compP f P)" (*<*) using wf by (simp add:wf_prog_def) (blast intro:wf_cdecl_compPI lift wf) (*>*) end diff --git a/thys/JinjaDCI/Compiler/Correctness1.thy b/thys/JinjaDCI/Compiler/Correctness1.thy --- a/thys/JinjaDCI/Compiler/Correctness1.thy +++ b/thys/JinjaDCI/Compiler/Correctness1.thy @@ -1,1058 +1,1064 @@ (* Title: JinjaDCI/Compiler/Correctness1.thy Author: Tobias Nipkow, Susannah Mansky Copyright TUM 2003, UIUC 2019-20 Based on the Jinja theory Compiler/Correctness1.thy by Tobias Nipkow *) section \ Correctness of Stage 1 \ theory Correctness1 imports J1WellForm Compiler1 begin subsection\Correctness of program compilation \ primrec unmod :: "expr\<^sub>1 \ nat \ bool" and unmods :: "expr\<^sub>1 list \ nat \ bool" where "unmod (new C) i = True" | "unmod (Cast C e) i = unmod e i" | "unmod (Val v) i = True" | "unmod (e\<^sub>1 \bop\ e\<^sub>2) i = (unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (Var i) j = True" | "unmod (i:=e) j = (i \ j \ unmod e j)" | "unmod (e\F{D}) i = unmod e i" | "unmod (C\\<^sub>sF{D}) i = True" | "unmod (e\<^sub>1\F{D}:=e\<^sub>2) i = (unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (C\\<^sub>sF{D}:=e\<^sub>2) i = unmod e\<^sub>2 i" | "unmod (e\M(es)) i = (unmod e i \ unmods es i)" | "unmod (C\\<^sub>sM(es)) i = unmods es i" | "unmod {j:T; e} i = unmod e i" | "unmod (e\<^sub>1;;e\<^sub>2) i = (unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (if (e) e\<^sub>1 else e\<^sub>2) i = (unmod e i \ unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (while (e) c) i = (unmod e i \ unmod c i)" | "unmod (throw e) i = unmod e i" | "unmod (try e\<^sub>1 catch(C i) e\<^sub>2) j = (unmod e\<^sub>1 j \ (if i=j then False else unmod e\<^sub>2 j))" | "unmod (INIT C (Cs,b) \ e) i = unmod e i" | "unmod (RI(C,e);Cs \ e') i = (unmod e i \ unmod e' i)" | "unmods ([]) i = True" | "unmods (e#es) i = (unmod e i \ unmods es i)" lemma hidden_unmod: "\Vs. hidden Vs i \ unmod (compE\<^sub>1 Vs e) i" and "\Vs. hidden Vs i \ unmods (compEs\<^sub>1 Vs es) i" (*<*) -apply(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) -apply (simp_all add:hidden_inacc) -apply(auto simp add:hidden_def) -done +proof(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) + case TryCatch + then show ?case by(simp add:hidden_inacc)(auto simp add:hidden_def) +qed (simp_all add:hidden_inacc) (*>*) lemma eval\<^sub>1_preserves_unmod: "\ P \\<^sub>1 \e,(h,ls,sh)\ \ \e',(h',ls',sh')\; unmod e i; i < size ls \ \ ls ! i = ls' ! i" and "\ P \\<^sub>1 \es,(h,ls,sh)\ [\] \es',(h',ls',sh')\; unmods es i; i < size ls \ \ ls ! i = ls' ! i" (*<*) proof(induct rule:eval\<^sub>1_evals\<^sub>1_inducts) case (RInitInitFail\<^sub>1 e h l sh a h' l' sh' C sfs i sh'' D Cs e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1) have "final (throw a)" using eval\<^sub>1_final[OF RInitInitFail\<^sub>1.hyps(1)] by simp then show ?case using RInitInitFail\<^sub>1 by(auto simp: eval\<^sub>1_preserves_len) qed(auto dest!:eval\<^sub>1_preserves_len evals\<^sub>1_preserves_len split:if_split_asm) (*>*) lemma LAss_lem: "\x \ set xs; size xs \ size ys \ \ m\<^sub>1 \\<^sub>m m\<^sub>2(xs[\]ys) \ m\<^sub>1(x\y) \\<^sub>m m\<^sub>2(xs[\]ys[last_index xs x := y])" (*<*) by(simp add:map_le_def fun_upds_apply eq_sym_conv) (*>*) lemma Block_lem: fixes l :: "'a \ 'b" assumes 0: "l \\<^sub>m [Vs [\] ls]" and 1: "l' \\<^sub>m [Vs [\] ls', V\v]" and hidden: "V \ set Vs \ ls ! last_index Vs V = ls' ! last_index Vs V" and size: "size ls = size ls'" "size Vs < size ls'" shows "l'(V := l V) \\<^sub>m [Vs [\] ls']" (*<*) proof - have "l'(V := l V) \\<^sub>m [Vs [\] ls', V\v](V := l V)" using 1 by(rule map_le_upd) also have "\ = [Vs [\] ls'](V := l V)" by simp also have "\ \\<^sub>m [Vs [\] ls']" proof (cases "l V") case None thus ?thesis by simp next case (Some w) hence "[Vs [\] ls] V = Some w" using 0 by(force simp add: map_le_def split:if_splits) hence VinVs: "V \ set Vs" and w: "w = ls ! last_index Vs V" using size by(auto simp add:fun_upds_apply split:if_splits) hence "w = ls' ! last_index Vs V" using hidden[OF VinVs] by simp hence "[Vs [\] ls'](V := l V) = [Vs [\] ls']" using Some size VinVs by(simp add: map_upds_upd_conv_last_index) thus ?thesis by simp qed finally show ?thesis . qed (*>*) (*<*) declare fun_upd_apply[simp del] (*>*) text\\noindent The main theorem: \ theorem assumes wf: "wwf_J_prog P" shows eval\<^sub>1_eval: "P \ \e,(h,l,sh)\ \ \e',(h',l',sh')\ \ (\Vs ls. \ fv e \ set Vs; l \\<^sub>m [Vs[\]ls]; size Vs + max_vars e \ size ls \ \ \ls'. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs e,(h,ls,sh)\ \ \fin\<^sub>1 e',(h',ls',sh')\ \ l' \\<^sub>m [Vs[\]ls'])" (*<*) (is "_ \ (\Vs ls. PROP ?P e h l sh e' h' l' sh' Vs ls)" is "_ \ (\Vs ls. \ _; _; _ \ \ \ls'. ?Post e h l sh e' h' l' sh' Vs ls ls')") (*>*) and evals\<^sub>1_evals: "P \ \es,(h,l,sh)\ [\] \es',(h',l',sh')\ \ (\Vs ls. \ fvs es \ set Vs; l \\<^sub>m [Vs[\]ls]; size Vs + max_varss es \ size ls \ \ \ls'. compP\<^sub>1 P \\<^sub>1 \compEs\<^sub>1 Vs es,(h,ls,sh)\ [\] \compEs\<^sub>1 Vs es',(h',ls',sh')\ \ l' \\<^sub>m [Vs[\]ls'])" (*<*) (is "_ \ (\Vs ls. PROP ?Ps es h l sh es' h' l' sh' Vs ls)" is "_ \ (\Vs ls. \ _; _; _\ \ \ls'. ?Posts es h l sh es' h' l' sh' Vs ls ls')") proof (induct rule:eval_evals_inducts) case Nil thus ?case by(fastforce intro!:Nil\<^sub>1) next case (Cons e h l sh v h' l' sh' es es' h\<^sub>2 l\<^sub>2 sh\<^sub>2) have "PROP ?P e h l sh (Val v) h' l' sh' Vs ls" by fact with Cons.prems obtain ls' where 1: "?Post e h l sh (Val v) h' l' sh' Vs ls ls'" "size ls = size ls'" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h' l' sh' es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls'" by fact with 1 Cons.prems obtain ls\<^sub>2 where 2: "?Posts es h' l' sh' es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls' ls\<^sub>2" by(auto) from 1 2 Cons show ?case by(auto intro!:Cons\<^sub>1) next case ConsThrow thus ?case by(fastforce intro!:ConsThrow\<^sub>1 dest: eval_final) next case (Block e h l V sh e' h' l' sh' T) let ?Vs = "Vs @ [V]" have IH: "\fv e \ set ?Vs; l(V := None) \\<^sub>m [?Vs [\] ls]; size ?Vs + max_vars e \ size ls\ \ \ls'. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e,(h,ls,sh)\ \ \fin\<^sub>1 e',(h', ls',sh')\ \ l' \\<^sub>m [?Vs [\] ls']" and fv: "fv {V:T; e} \ set Vs" and rel: "l \\<^sub>m [Vs [\] ls]" and len: "length Vs + max_vars {V:T; e} \ length ls" by fact+ have len': "length Vs < length ls" using len by auto have "fv e \ set ?Vs" using fv by auto moreover have "l(V := None) \\<^sub>m [?Vs [\] ls]" using rel len' by simp moreover have "size ?Vs + max_vars e \ size ls" using len by simp ultimately obtain ls' where 1: "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e,(h,ls,sh)\ \ \fin\<^sub>1 e',(h',ls',sh')\" and rel': "l' \\<^sub>m [?Vs [\] ls']" using IH by blast have [simp]: "length ls = length ls'" by(rule eval\<^sub>1_preserves_len[OF 1]) show "\ls'. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs {V:T; e},(h,ls,sh)\ \ \fin\<^sub>1 e',(h',ls',sh')\ \ l'(V := l V) \\<^sub>m [Vs [\] ls']" (is "\ls'. ?R ls'") proof show "?R ls'" proof show "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs {V:T; e},(h,ls,sh)\ \ \fin\<^sub>1 e',(h',ls',sh')\" using 1 by(simp add:Block\<^sub>1) next show "l'(V := l V) \\<^sub>m [Vs [\] ls']" proof - have "l' \\<^sub>m [Vs [\] ls', V \ ls' ! length Vs]" using len' rel' by simp moreover { assume VinVs: "V \ set Vs" hence "hidden (Vs @ [V]) (last_index Vs V)" by(rule hidden_last_index) hence "unmod (compE\<^sub>1 (Vs @ [V]) e) (last_index Vs V)" by(rule hidden_unmod) moreover have "last_index Vs V < length ls" using len' VinVs by simp ultimately have "ls ! last_index Vs V = ls' ! last_index Vs V" by(rule eval\<^sub>1_preserves_unmod[OF 1]) } ultimately show ?thesis using Block_lem[OF rel] len' by auto qed qed qed next case (TryThrow e' h l sh a h' l' sh' D fs C V e\<^sub>2) have "PROP ?P e' h l sh (Throw a) h' l' sh' Vs ls" by fact with TryThrow.prems obtain ls' where 1: "?Post e' h l sh (Throw a) h' l' sh' Vs ls ls'" by(auto) show ?case using 1 TryThrow.hyps by(auto intro!:eval\<^sub>1_evals\<^sub>1.TryThrow\<^sub>1) next case (TryCatch e\<^sub>1 h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C e\<^sub>2 V e' h\<^sub>2 l\<^sub>2 sh\<^sub>2) let ?e = "try e\<^sub>1 catch(C V) e\<^sub>2" have IH\<^sub>1: "\fv e\<^sub>1 \ set Vs; l \\<^sub>m [Vs [\] ls]; size Vs + max_vars e\<^sub>1 \ length ls\ \ \ls\<^sub>1. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs e\<^sub>1,(h,ls,sh)\ \ \fin\<^sub>1 (Throw a),(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\ \ l\<^sub>1 \\<^sub>m [Vs [\] ls\<^sub>1]" and fv: "fv ?e \ set Vs" and rel: "l \\<^sub>m [Vs [\] ls]" and len: "length Vs + max_vars ?e \ length ls" by fact+ have "fv e\<^sub>1 \ set Vs" using fv by auto moreover have "length Vs + max_vars e\<^sub>1 \ length ls" using len by(auto) ultimately obtain ls\<^sub>1 where 1: "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs e\<^sub>1,(h,ls,sh)\ \ \Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\" and rel\<^sub>1: "l\<^sub>1 \\<^sub>m [Vs [\] ls\<^sub>1]" using IH\<^sub>1 rel by fastforce from 1 have [simp]: "size ls = size ls\<^sub>1" by(rule eval\<^sub>1_preserves_len) let ?Vs = "Vs @ [V]" let ?ls = "(ls\<^sub>1[size Vs:=Addr a])" have IH\<^sub>2: "\fv e\<^sub>2 \ set ?Vs; l\<^sub>1(V \ Addr a) \\<^sub>m [?Vs [\] ?ls]; length ?Vs + max_vars e\<^sub>2 \ length ?ls\ \ \ls\<^sub>2. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e\<^sub>2,(h\<^sub>1,?ls,sh\<^sub>1)\ \ \fin\<^sub>1 e',(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\ \ l\<^sub>2 \\<^sub>m [?Vs [\] ls\<^sub>2]" by fact have len\<^sub>1: "size Vs < size ls\<^sub>1" using len by(auto) have "fv e\<^sub>2 \ set ?Vs" using fv by auto moreover have "l\<^sub>1(V \ Addr a) \\<^sub>m [?Vs [\] ?ls]" using rel\<^sub>1 len\<^sub>1 by simp moreover have "length ?Vs + max_vars e\<^sub>2 \ length ?ls" using len by(auto) ultimately obtain ls\<^sub>2 where 2: "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e\<^sub>2,(h\<^sub>1,?ls,sh\<^sub>1)\ \ \fin\<^sub>1 e',(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" and rel\<^sub>2: "l\<^sub>2 \\<^sub>m [?Vs [\] ls\<^sub>2]" using IH\<^sub>2 by blast from 2 have [simp]: "size ls\<^sub>1 = size ls\<^sub>2" by(fastforce dest: eval\<^sub>1_preserves_len) show "\ls\<^sub>2. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs ?e,(h,ls,sh)\ \ \fin\<^sub>1 e',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\ \ l\<^sub>2(V := l\<^sub>1 V) \\<^sub>m [Vs [\] ls\<^sub>2]" (is "\ls\<^sub>2. ?R ls\<^sub>2") proof show "?R ls\<^sub>2" proof have hp: "h\<^sub>1 a = Some (D, fs)" by fact have "P \ D \\<^sup>* C" by fact hence caught: "compP\<^sub>1 P \ D \\<^sup>* C" by simp from TryCatch\<^sub>1[OF 1 _ caught len\<^sub>1 2, OF hp] show "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs ?e,(h,ls,sh)\ \ \fin\<^sub>1 e',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\" by simp next show "l\<^sub>2(V := l\<^sub>1 V) \\<^sub>m [Vs [\] ls\<^sub>2]" proof - have "l\<^sub>2 \\<^sub>m [Vs [\] ls\<^sub>2, V \ ls\<^sub>2 ! length Vs]" using len\<^sub>1 rel\<^sub>2 by simp moreover { assume VinVs: "V \ set Vs" hence "hidden (Vs @ [V]) (last_index Vs V)" by(rule hidden_last_index) hence "unmod (compE\<^sub>1 (Vs @ [V]) e\<^sub>2) (last_index Vs V)" by(rule hidden_unmod) moreover have "last_index Vs V < length ?ls" using len\<^sub>1 VinVs by simp ultimately have "?ls ! last_index Vs V = ls\<^sub>2 ! last_index Vs V" by(rule eval\<^sub>1_preserves_unmod[OF 2]) moreover have "last_index Vs V < size Vs" using VinVs by simp ultimately have "ls\<^sub>1 ! last_index Vs V = ls\<^sub>2 ! last_index Vs V" using len\<^sub>1 by(simp del:size_last_index_conv) } ultimately show ?thesis using Block_lem[OF rel\<^sub>1] len\<^sub>1 by simp qed qed qed next case Try thus ?case by(fastforce intro!:Try\<^sub>1) next case Throw thus ?case by(fastforce intro!:Throw\<^sub>1) next case ThrowNull thus ?case by(fastforce intro!:ThrowNull\<^sub>1) next case ThrowThrow thus ?case by(fastforce intro!:ThrowThrow\<^sub>1) next case (CondT e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 e\<^sub>2) have "PROP ?P e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with CondT.prems obtain ls\<^sub>1 where 1: "?Post e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 CondT.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 show ?case by(auto intro!:CondT\<^sub>1) next case (CondF e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 e\<^sub>1 Vs ls) have "PROP ?P e h l sh false h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with CondF.prems obtain ls\<^sub>1 where 1: "?Post e h l sh false h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 CondF.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 show ?case by(auto intro!:CondF\<^sub>1) next case CondThrow thus ?case by(fastforce intro!:CondThrow\<^sub>1) next case (Seq e h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2) have "PROP ?P e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with Seq.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 Seq.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 Seq show ?case by(auto intro!:Seq\<^sub>1) next case SeqThrow thus ?case by(fastforce intro!:SeqThrow\<^sub>1) next case WhileF thus ?case by(fastforce intro!:eval\<^sub>1_evals\<^sub>1.intros) next case (WhileT e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 c v h\<^sub>2 l\<^sub>2 sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) have "PROP ?P e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with WhileT.prems obtain ls\<^sub>1 where 1: "?Post e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 WhileT.prems obtain ls\<^sub>2 where 2: "?Post c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" "size ls\<^sub>1 = size ls\<^sub>2" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P (While (e) c) h\<^sub>2 l\<^sub>2 sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 Vs ls\<^sub>2" by fact with 1 2 WhileT.prems obtain ls\<^sub>3 where 3: "?Post (While (e) c) h\<^sub>2 l\<^sub>2 sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 Vs ls\<^sub>2 ls\<^sub>3" by(auto) from 1 2 3 show ?case by(auto intro!:WhileT\<^sub>1) next case (WhileBodyThrow e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 c e' h\<^sub>2 l\<^sub>2 sh\<^sub>2) have "PROP ?P e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with WhileBodyThrow.prems obtain ls\<^sub>1 where 1: "?Post e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 WhileBodyThrow.prems obtain ls\<^sub>2 where 2: "?Post c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by auto from 1 2 show ?case by(auto intro!:WhileBodyThrow\<^sub>1) next case WhileCondThrow thus ?case by(fastforce intro!:WhileCondThrow\<^sub>1) next case New thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case NewFail thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case NewInit then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case NewInitOOM then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case NewInitThrow then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case Cast thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case CastNull thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case CastThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (CastFail e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C) have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with CastFail.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" by auto show ?case using 1 CastFail.hyps by(auto intro!:CastFail\<^sub>1[where D=D]) next case Val thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (BinOp e h l sh v\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 v\<^sub>2 h\<^sub>2 l\<^sub>2 sh\<^sub>2 bop v) have "PROP ?P e h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with BinOp.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v\<^sub>2) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 BinOp.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v\<^sub>2) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 BinOp show ?case by(auto intro!:BinOp\<^sub>1) next case (BinOpThrow2 e\<^sub>0 h l sh v\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 e h\<^sub>2 l\<^sub>2 sh\<^sub>2 bop) have "PROP ?P e\<^sub>0 h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with BinOpThrow2.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>0 h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 BinOpThrow2.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 BinOpThrow2 show ?case by(auto intro!:BinOpThrow\<^sub>2\<^sub>1) next case BinOpThrow1 thus ?case by(fastforce intro!:eval\<^sub>1_evals\<^sub>1.intros) next case Var thus ?case by(force intro!:Var\<^sub>1 simp add: map_le_def fun_upds_apply) next case LAss thus ?case by(fastforce simp add: LAss_lem intro:eval\<^sub>1_evals\<^sub>1.intros dest:eval\<^sub>1_preserves_len) next case LAssThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case FAcc thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case FAccNull thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case FAccThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (FAccNone e h l sh a h' l' sh' C fs F D) have "PROP ?P e h l sh (addr a) h' l' sh' Vs ls" by fact with FAccNone.prems obtain ls\<^sub>2 where 2: "?Post e h l sh (addr a) h' l' sh' Vs ls ls\<^sub>2" by(auto) from 2 FAccNone show ?case by(rule_tac x = ls\<^sub>2 in exI, auto elim!: FAccNone\<^sub>1) next case (FAccStatic e h l sh a h' l' sh' C fs F t D) have "PROP ?P e h l sh (addr a) h' l' sh' Vs ls" by fact with FAccStatic.prems obtain ls\<^sub>2 where 2: "?Post e h l sh (addr a) h' l' sh' Vs ls ls\<^sub>2" by(auto) from 2 FAccStatic show ?case by(rule_tac x = ls\<^sub>2 in exI, auto elim!: FAccStatic\<^sub>1) next case SFAcc then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (SFAccInit C F t D sh h l v' h' l' sh' sfs i v) have "PROP ?P (INIT D ([D],False) \ unit) h l sh (Val v') h' l' sh' Vs ls" by fact with SFAccInit.prems obtain ls\<^sub>2 where 1: "?Post (INIT D ([D],False) \ unit) h l sh (Val v') h' l' sh' Vs ls ls\<^sub>2" by(auto) from 1 SFAccInit show ?case by(rule_tac x = ls\<^sub>2 in exI, auto intro: SFAccInit\<^sub>1) next case (SFAccInitThrow C F t D sh h l a h' l' sh') have "PROP ?P (INIT D ([D],False) \ unit) h l sh (throw a) h' l' sh' Vs ls" by fact with SFAccInitThrow.prems obtain ls\<^sub>2 where 1: "?Post (INIT D ([D],False) \ unit) h l sh (throw a) h' l' sh' Vs ls ls\<^sub>2" by(auto) from 1 SFAccInitThrow show ?case by(rule_tac x = ls\<^sub>2 in exI, auto intro: SFAccInitThrow\<^sub>1) next case SFAccNone then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case SFAccNonStatic then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (FAss e\<^sub>1 h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs fs' F D h\<^sub>2') have "PROP ?P e\<^sub>1 h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with FAss.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAss.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAss show ?case by(auto intro!:FAss\<^sub>1) next case (FAssNull e\<^sub>1 h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 F D) have "PROP ?P e\<^sub>1 h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with FAssNull.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAssNull.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAssNull show ?case by(auto intro!:FAssNull\<^sub>1) next case FAssThrow1 thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (FAssThrow2 e\<^sub>1 h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 e h\<^sub>2 l\<^sub>2 sh\<^sub>2 F D) have "PROP ?P e\<^sub>1 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with FAssThrow2.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAssThrow2.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAssThrow2 show ?case by(auto intro!:FAssThrow\<^sub>2\<^sub>1) next case (FAssNone e\<^sub>1 h l sh a h' l' sh' e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F D) have "PROP ?P e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls" by fact with FAssNone.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAssNone.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAssNone show ?case by(auto intro!:FAssNone\<^sub>1) next case (FAssStatic e\<^sub>1 h l sh a h' l' sh' e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F t D) have "PROP ?P e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls" by fact with FAssStatic.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAssStatic.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAssStatic show ?case by(auto intro!:FAssStatic\<^sub>1) next case SFAss then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (SFAssInit e\<^sub>2 h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D v' h' l' sh' sfs i sfs' sh'') have "PROP ?P e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with SFAssInit.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) then have Init_size: "length Vs \ length ls\<^sub>1" using SFAssInit.prems(3) by linarith have "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h' l' sh' Vs ls\<^sub>1" by fact with 1 Init_size SFAssInit.prems obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h' l' sh' Vs ls\<^sub>1 ls\<^sub>2" by auto from 1 2 SFAssInit show ?case by(auto simp add: comp_def intro!: SFAssInit\<^sub>1 dest!:evals_final) next case (SFAssInitThrow e\<^sub>2 h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D a h\<^sub>2 l\<^sub>2 sh\<^sub>2) have "PROP ?P e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with SFAssInitThrow.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) then have Init_size: "length Vs \ length ls\<^sub>1" using SFAssInitThrow.prems(3) by linarith have "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 Init_size SFAssInitThrow.prems obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by auto from 1 2 SFAssInitThrow show ?case by(auto simp add: comp_def intro!: SFAssInitThrow\<^sub>1 dest!:evals_final) next case SFAssThrow then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (SFAssNone e\<^sub>2 h l sh v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C F D) have "PROP ?P e\<^sub>2 h l sh (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls" by fact with SFAssNone.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h l sh (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls ls\<^sub>2" by(auto) from 2 SFAssNone show ?case by(rule_tac x = ls\<^sub>2 in exI, auto elim!: SFAssNone\<^sub>1) next case SFAssNonStatic then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (CallNull e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 es vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 M) have "PROP ?P e h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with CallNull.prems obtain ls\<^sub>1 where 1: "?Post e h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 CallNull.prems obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 CallNull show ?case by (auto simp add: comp_def elim!: CallNull\<^sub>1) next case CallObjThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (CallParamsThrow e h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 es vs ex es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 M) have "PROP ?P e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with CallParamsThrow.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 CallParamsThrow.prems obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 CallParamsThrow show ?case by (auto simp add: comp_def elim!: CallParamsThrow\<^sub>1 dest!:evals_final) next case (CallNone e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M) have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with CallNone.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 CallNone.prems obtain ls\<^sub>2 where 2: "?Posts ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 CallNone show ?case by (auto simp add: comp_def elim!: CallNone\<^sub>1 dest!:evals_final sees_method_compPD) next case (CallStatic e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body D) have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with CallStatic.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) let ?Vs = pns have mdecl: "P \ C sees M,Static: Ts\T = (pns, body) in D" by fact have mdecl\<^sub>1: "compP\<^sub>1 P \ C sees M,Static: Ts\T = (compE\<^sub>1 ?Vs body) in D" using sees_method_compP[OF mdecl, of "\b (pns,e). compE\<^sub>1 (case b of NonStatic \ this#pns | Static \ pns) e"] by(simp) have "PROP ?Ps ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 CallStatic.prems obtain ls\<^sub>2 where 2: "?Posts ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 mdecl\<^sub>1 CallStatic show ?case by (auto simp add: comp_def elim!: CallStatic\<^sub>1 dest!:evals_final) next case (Call e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 es vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body D l\<^sub>2' b' h\<^sub>3 l\<^sub>3 sh\<^sub>3) have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with Call.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 Call.prems obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" "size ls\<^sub>1 = size ls\<^sub>2" by(auto intro!:evals\<^sub>1_preserves_len) let ?Vs = "this#pns" let ?ls = "Addr a # vs @ replicate (max_vars body) undefined" have mdecl: "P \ C sees M,NonStatic: Ts\T = (pns, body) in D" by fact have fv_body: "fv body \ set ?Vs" and wf_size: "size Ts = size pns" using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def) have mdecl\<^sub>1: "compP\<^sub>1 P \ C sees M,NonStatic: Ts\T = (compE\<^sub>1 ?Vs body) in D" using sees_method_compP[OF mdecl, of "\b (pns,e). compE\<^sub>1 (case b of NonStatic \ this#pns | Static \ pns) e"] by(simp) have [simp]: "l\<^sub>2' = [this \ Addr a, pns [\] vs]" by fact have Call_size: "size vs = size pns" by fact have "PROP ?P body h\<^sub>2 l\<^sub>2' sh\<^sub>2 b' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls" by fact with 1 2 fv_body Call_size Call.prems obtain ls\<^sub>3 where 3: "?Post body h\<^sub>2 l\<^sub>2' sh\<^sub>2 b' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls ls\<^sub>3" by(auto) have hp: "h\<^sub>2 a = Some (C, fs)" by fact from 1 2 3 hp mdecl\<^sub>1 wf_size Call_size show ?case by(fastforce simp add: comp_def intro!: Call\<^sub>1 dest!:evals_final) next case (SCallParamsThrow es h l sh vs ex es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 C M) have "PROP ?Ps es h l sh (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls" by fact with SCallParamsThrow.prems obtain ls\<^sub>2 where 2: "?Posts es h l sh (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls ls\<^sub>2" by(auto) from 2 SCallParamsThrow show ?case by (fastforce simp add: comp_def elim!: SCallParamsThrow\<^sub>1 dest!:evals_final) next case (SCall ps h l sh vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C M Ts T pns body D sfs l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) have "PROP ?Ps ps h l sh (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls" by fact with SCall.prems obtain ls\<^sub>2 where 2: "?Posts ps h l sh (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls ls\<^sub>2" "size ls = size ls\<^sub>2" by(auto intro!:evals\<^sub>1_preserves_len) let ?Vs = "pns" let ?ls = "vs @ replicate (max_vars body) undefined" have mdecl: "P \ C sees M,Static: Ts\T = (pns, body) in D" by fact have fv_body: "fv body \ set ?Vs" and wf_size: "size Ts = size pns" using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def) have mdecl\<^sub>1: "compP\<^sub>1 P \ C sees M,Static: Ts\T = (compE\<^sub>1 ?Vs body) in D" using sees_method_compP[OF mdecl, of "\b (pns,e). compE\<^sub>1 (case b of NonStatic \ this#pns | Static \ pns) e"] by(simp) have [simp]: "l\<^sub>2' = [pns [\] vs]" by fact have SCall_size: "size vs = size pns" by fact have "PROP ?P body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls" by fact with 2 fv_body SCall_size SCall.prems obtain ls\<^sub>3 where 3: "?Post body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls ls\<^sub>3" by(auto) have shp: "sh\<^sub>2 D = \(sfs, Done)\ \ M = clinit \ sh\<^sub>2 D = \(sfs, Processing)\" by fact from 2 3 shp mdecl\<^sub>1 wf_size SCall_size show ?case by(fastforce simp add: comp_def intro!: SCall\<^sub>1 dest!:evals_final) next case (SCallNone ps h l sh vs h' l' sh' C M) have "PROP ?Ps ps h l sh (map Val vs) h' l' sh' Vs ls" by fact with SCallNone.prems obtain ls\<^sub>2 where 2: "?Posts ps h l sh (map Val vs) h' l' sh' Vs ls ls\<^sub>2" by(auto) from 2 SCallNone show ?case by(rule_tac x = ls\<^sub>2 in exI, auto simp add: comp_def elim!: SCallNone\<^sub>1 dest!:evals_final sees_method_compPD) next case (SCallNonStatic ps h l sh vs h' l' sh' C M Ts T pns body D) let ?Vs = "this#pns" have mdecl: "P \ C sees M,NonStatic: Ts\T = (pns, body) in D" by fact have mdecl\<^sub>1: "compP\<^sub>1 P \ C sees M,NonStatic: Ts\T = (compE\<^sub>1 ?Vs body) in D" using sees_method_compP[OF mdecl, of "\b (pns,e). compE\<^sub>1 (case b of NonStatic \ this#pns | Static \ pns) e"] by(simp) have "PROP ?Ps ps h l sh (map Val vs) h' l' sh' Vs ls" by fact with SCallNonStatic.prems obtain ls\<^sub>2 where 2: "?Posts ps h l sh (map Val vs) h' l' sh' Vs ls ls\<^sub>2" by(auto) from 2 mdecl\<^sub>1 SCallNonStatic show ?case by (auto simp add: comp_def elim!: SCallNonStatic\<^sub>1 dest!:evals_final) next case (SCallInitThrow ps h l sh vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D a h\<^sub>2 l\<^sub>2 sh\<^sub>2) have "PROP ?Ps ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with SCallInitThrow.prems obtain ls\<^sub>1 where 1: "?Posts ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1" by(auto intro!:evals\<^sub>1_preserves_len) then have Init_size: "length Vs \ length ls\<^sub>1" using SCallInitThrow.prems(3) by linarith have "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 Init_size SCallInitThrow.prems obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by auto let ?Vs = "pns" have mdecl: "P \ C sees M,Static: Ts\T = (pns, body) in D" by fact have mdecl\<^sub>1: "compP\<^sub>1 P \ C sees M,Static: Ts\T = (compE\<^sub>1 ?Vs body) in D" using sees_method_compP[OF mdecl, of "\b (pns,e). compE\<^sub>1 (case b of NonStatic \ this#pns | Static \ pns) e"] by(simp) from 1 2 mdecl\<^sub>1 SCallInitThrow show ?case by(auto simp add: comp_def intro!: SCallInitThrow\<^sub>1 dest!:evals_final) next case (SCallInit ps h l sh vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) have "PROP ?Ps ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with SCallInit.prems obtain ls\<^sub>1 where 1: "?Posts ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1" by(auto intro!:evals\<^sub>1_preserves_len) then have Init_size: "length Vs \ length ls\<^sub>1" using SCallInit.prems(3) by linarith have "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 Init_size SCallInit.prems obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by auto let ?Vs = "pns" let ?ls = "vs @ replicate (max_vars body) undefined" have mdecl: "P \ C sees M,Static: Ts\T = (pns, body) in D" by fact have fv_body: "fv body \ set ?Vs" and wf_size: "size Ts = size pns" using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def) have mdecl\<^sub>1: "compP\<^sub>1 P \ C sees M,Static: Ts\T = (compE\<^sub>1 ?Vs body) in D" using sees_method_compP[OF mdecl, of "\b (pns,e). compE\<^sub>1 (case b of NonStatic \ this#pns | Static \ pns) e"] by(simp) have [simp]: "l\<^sub>2' = [pns [\] vs]" by fact have SCall_size: "size vs = size pns" by fact have nclinit: "M \ clinit" by fact have "PROP ?P body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls" by fact with 2 fv_body SCall_size SCallInit.prems obtain ls\<^sub>3 where 3: "?Post body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls ls\<^sub>3" by(auto) have shp: "\sfs. sh\<^sub>1 D = \(sfs, Done)\" by fact from 1 2 3 shp mdecl\<^sub>1 wf_size SCall_size nclinit show ?case by(auto simp add: comp_def intro!: SCallInit\<^sub>1 dest!:evals_final) next \ \ init cases \ case InitFinal then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (InitNone sh C C' Cs e h l e' h' l' sh') let ?sh1 = "sh(C \ (sblank P C, Prepared))" have "PROP ?P (INIT C' (C # Cs,False) \ e) h l ?sh1 e' h' l' sh' Vs ls" by fact with InitNone.prems obtain ls\<^sub>2 where 2: "?Post (INIT C' (C # Cs,False) \ e) h l ?sh1 e' h' l' sh' Vs ls ls\<^sub>2" by(auto) from 2 InitNone show ?case by (auto elim!: InitNone\<^sub>1) next case InitDone then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case InitProcessing then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case InitError then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case InitObject then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (InitNonObject sh C sfs D fs ms sh' C' Cs e h l e' h1 l1 sh1) let ?f = "(\b (pns,body). compE\<^sub>1 (case b of NonStatic \ this#pns | Static \ pns) body)" have cls: "class (compP ?f P) C = \(D,fs,map (compM ?f) ms)\" by(rule class_compP[OF InitNonObject.hyps(3)]) have "PROP ?P (INIT C' (D # C # Cs,False) \ e) h l sh' e' h1 l1 sh1 Vs ls" by fact with InitNonObject.prems obtain ls\<^sub>2 where 2: "?Post (INIT C' (D # C # Cs,False) \ e) h l sh' e' h1 l1 sh1 Vs ls ls\<^sub>2" by(auto) from 2 cls InitNonObject show ?case by (auto elim!: InitNonObject\<^sub>1) next case InitRInit then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (RInit e h l sh v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1) have "PROP ?P e h l sh (Val v) h' l' sh' Vs ls" by fact with RInit.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v) h' l' sh' Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P (INIT C' (Cs,True) \ e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1" by fact with 1 RInit.prems obtain ls\<^sub>2 where 2: "?Post (INIT C' (Cs,True) \ e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 RInit show ?case by (auto elim!: RInit\<^sub>1) next case (RInitInitFail e h l sh a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1) have "PROP ?P e h l sh (throw a) h' l' sh' Vs ls" by fact with RInitInitFail.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (throw a) h' l' sh' Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have fv: "fv (RI (D,throw a) ; Cs \ e') \ set Vs" using RInitInitFail.hyps(1) eval_final RInitInitFail.prems(1) subset_eq by fastforce have l': "l' \\<^sub>m [Vs [\] ls\<^sub>1]" by (simp add: "1"(1)) have "PROP ?P (RI (D,throw a) ; Cs \ e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1" by fact with 1 eval_final[OF RInitInitFail.hyps(1)] RInitInitFail.prems obtain ls\<^sub>2 where 2: "?Post (RI (D,throw a) ; Cs \ e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1 ls\<^sub>2" by fastforce from 1 2 RInitInitFail show ?case by(fastforce simp add: comp_def intro!: RInitInitFail\<^sub>1 dest!:eval_final) next case RInitFailFinal then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) qed (*>*) subsection\Preservation of well-formedness\ text\ The compiler preserves well-formedness. Is less trivial than it may appear. We start with two simple properties: preservation of well-typedness \ lemma compE\<^sub>1_pres_wt: "\Vs Ts U. \ P,[Vs[\]Ts] \ e :: U; size Ts = size Vs \ \ compP f P,Ts \\<^sub>1 compE\<^sub>1 Vs e :: U" and "\Vs Ts Us. \ P,[Vs[\]Ts] \ es [::] Us; size Ts = size Vs \ \ compP f P,Ts \\<^sub>1 compEs\<^sub>1 Vs es [::] Us" (*<*) -apply(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) -apply clarsimp -apply(fastforce) -apply clarsimp -apply(fastforce split:bop.splits) -apply (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) -apply (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) -apply(fastforce) -apply(fastforce) -apply(fastforce) -apply(fastforce) -apply (fastforce dest!: sees_method_compP[where f = f]) -apply (fastforce dest!: sees_method_compP[where f = f]) -apply (fastforce simp:nth_append) -apply (fastforce) -apply (fastforce) -apply (fastforce) -apply (fastforce) -apply (fastforce simp:nth_append) -apply simp -apply simp -apply simp -apply (fastforce) -done +proof(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) + case Var then show ?case + by (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) +next + case LAss then show ?case + by (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) +next + case Call then show ?case + by (fastforce dest!: sees_method_compP[where f = f]) +next + case SCall then show ?case + by (fastforce dest!: sees_method_compP[where f = f]) +next + case Block then show ?case by (fastforce simp:nth_append) +next + case TryCatch then show ?case by (fastforce simp:nth_append) +qed fastforce+ (*>*) text\\noindent and the correct block numbering: \ lemma \: "\Vs n. size Vs = n \ \ (compE\<^sub>1 Vs e) n" and \s: "\Vs n. size Vs = n \ \s (compEs\<^sub>1 Vs es) n" (*<*)by (induct e and es rule: \.induct \s.induct) (force | simp,metis length_append_singleton)+(*>*) text\ The main complication is preservation of definite assignment @{term"\"}. \ lemma image_last_index: "A \ set(xs@[x]) \ last_index (xs @ [x]) ` A = (if x \ A then insert (size xs) (last_index xs ` (A-{x})) else last_index xs ` A)" (*<*) by(auto simp:image_def) (*>*) lemma A_compE\<^sub>1_None[simp]: "\Vs. \ e = None \ \ (compE\<^sub>1 Vs e) = None" and "\Vs. \s es = None \ \s (compEs\<^sub>1 Vs es) = None" (*<*)by(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct)(auto simp:hyperset_defs)(*>*) lemma A_compE\<^sub>1: "\A Vs. \ \ e = \A\; fv e \ set Vs \ \ \ (compE\<^sub>1 Vs e) = \last_index Vs ` A\" and "\A Vs. \ \s es = \A\; fvs es \ set Vs \ \ \s (compEs\<^sub>1 Vs es) = \last_index Vs ` A\" (*<*) proof(induct e and es rule: fv.induct fvs.induct) case (Block V' T e) hence "fv e \ set (Vs@[V'])" by fastforce moreover obtain B where "\ e = \B\" using Block.prems by(simp add: hyperset_defs) moreover from calculation have "B \ set (Vs@[V'])" by(auto dest!:A_fv) ultimately show ?case using Block by(auto simp add: hyperset_defs image_last_index last_index_size_conv) next case (TryCatch e\<^sub>1 C V' e\<^sub>2) hence fve\<^sub>2: "fv e\<^sub>2 \ set (Vs@[V'])" by auto show ?case proof (cases "\ e\<^sub>1") assume A\<^sub>1: "\ e\<^sub>1 = None" then obtain A\<^sub>2 where A\<^sub>2: "\ e\<^sub>2 = \A\<^sub>2\" using TryCatch by(simp add:hyperset_defs) hence "A\<^sub>2 \ set (Vs@[V'])" using TryCatch.prems A_fv[OF A\<^sub>2] by simp blast thus ?thesis using TryCatch fve\<^sub>2 A\<^sub>1 A\<^sub>2 by(auto simp add:hyperset_defs image_last_index last_index_size_conv) next fix A\<^sub>1 assume A\<^sub>1: "\ e\<^sub>1 = \A\<^sub>1\" show ?thesis proof (cases "\ e\<^sub>2") assume A\<^sub>2: "\ e\<^sub>2 = None" then show ?case using TryCatch A\<^sub>1 by(simp add:hyperset_defs) next fix A\<^sub>2 assume A\<^sub>2: "\ e\<^sub>2 = \A\<^sub>2\" have "A\<^sub>1 \ set Vs" using TryCatch.prems A_fv[OF A\<^sub>1] by simp blast moreover have "A\<^sub>2 \ set (Vs@[V'])" using TryCatch.prems A_fv[OF A\<^sub>2] by simp blast ultimately show ?thesis using TryCatch A\<^sub>1 A\<^sub>2 by (auto simp add: Diff_subset_conv last_index_size_conv subsetD hyperset_defs dest!: sym [of _ A]) qed qed next case (Cond e e\<^sub>1 e\<^sub>2) { assume "\ e = None \ \ e\<^sub>1 = None \ \ e\<^sub>2 = None" hence ?case using Cond by(auto simp add:hyperset_defs image_Un) } moreover { fix A A\<^sub>1 A\<^sub>2 assume "\ e = \A\" and A\<^sub>1: "\ e\<^sub>1 = \A\<^sub>1\" and A\<^sub>2: "\ e\<^sub>2 = \A\<^sub>2\" moreover have "A\<^sub>1 \ set Vs" using Cond.prems A_fv[OF A\<^sub>1] by simp blast moreover have "A\<^sub>2 \ set Vs" using Cond.prems A_fv[OF A\<^sub>2] by simp blast ultimately have ?case using Cond by(auto simp add:hyperset_defs image_Un inj_on_image_Int[OF inj_on_last_index]) } ultimately show ?case by fastforce qed (auto simp add:hyperset_defs) (*>*) lemma D_None[iff]: "\ (e::'a exp) None" and [iff]: "\s (es::'a exp list) None" (*<*)by(induct e and es rule: \.induct \s.induct)(simp_all)(*>*) lemma D_last_index_compE\<^sub>1: "\A Vs. \ A \ set Vs; fv e \ set Vs \ \ \ e \A\ \ \ (compE\<^sub>1 Vs e) \last_index Vs ` A\" and "\A Vs. \ A \ set Vs; fvs es \ set Vs \ \ \s es \A\ \ \s (compEs\<^sub>1 Vs es) \last_index Vs ` A\" (*<*) proof(induct e and es rule: \.induct \s.induct) case (BinOp e\<^sub>1 bop e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using BinOp by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] BinOp.prems by auto have "A \ A\<^sub>1 \ set Vs" using BinOp.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using BinOp Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (FAss e\<^sub>1 F D e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using FAss by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] FAss.prems by auto have "A \ A\<^sub>1 \ set Vs" using FAss.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using FAss Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (Call e\<^sub>1 M es) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using Call by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] Call.prems by auto have "A \ A\<^sub>1 \ set Vs" using Call.prems A_fv[OF Some] by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` (A \ A\<^sub>1)\" using Call Some by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (TryCatch e\<^sub>1 C V e\<^sub>2) have "\ A\{V} \ set(Vs@[V]); fv e\<^sub>2 \ set(Vs@[V]); \ e\<^sub>2 \A\{V}\\ \ \ (compE\<^sub>1 (Vs@[V]) e\<^sub>2) \last_index (Vs@[V]) ` (A\{V})\" by fact hence "\ (compE\<^sub>1 (Vs@[V]) e\<^sub>2) \last_index (Vs@[V]) ` (A\{V})\" using TryCatch.prems by(simp add:Diff_subset_conv) moreover have "last_index (Vs@[V]) ` A \ last_index Vs ` A \ {size Vs}" using TryCatch.prems by(auto simp add: image_last_index split:if_split_asm) ultimately show ?case using TryCatch by(auto simp:hyperset_defs elim!:D_mono') next case (Seq e\<^sub>1 e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using Seq by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] Seq.prems by auto have "A \ A\<^sub>1 \ set Vs" using Seq.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using Seq Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (Cond e e\<^sub>1 e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e) \last_index Vs ` A\" by simp show ?case proof (cases "\ e") case None thus ?thesis using Cond by simp next case (Some B) have indexB: "\ (compE\<^sub>1 Vs e) = \last_index Vs ` B\" using A_compE\<^sub>1[OF Some] Cond.prems by auto have "A \ B \ set Vs" using Cond.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` (A \ B)\" and "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ B)\" using Cond Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A \ last_index Vs ` B\" and "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` B\" by(simp add: image_Un)+ thus ?thesis using IH\<^sub>1 indexB by auto qed next case (While e\<^sub>1 e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using While by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] While.prems by auto have "A \ A\<^sub>1 \ set Vs" using While.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using While Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (Block V T e) have "\ A-{V} \ set(Vs@[V]); fv e \ set(Vs@[V]); \ e \A-{V}\ \ \ \ (compE\<^sub>1 (Vs@[V]) e) \last_index (Vs@[V]) ` (A-{V})\" by fact hence "\ (compE\<^sub>1 (Vs@[V]) e) \last_index (Vs@[V]) ` (A-{V})\" using Block.prems by(simp add:Diff_subset_conv) moreover have "size Vs \ last_index Vs ` A" using Block.prems by(auto simp add:image_def size_last_index_conv) ultimately show ?case using Block by(auto simp add: image_last_index Diff_subset_conv hyperset_defs elim!: D_mono') next case (Cons_exp e\<^sub>1 es) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using Cons_exp by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] Cons_exp.prems by auto have "A \ A\<^sub>1 \ set Vs" using Cons_exp.prems A_fv[OF Some] by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` (A \ A\<^sub>1)\" using Cons_exp Some by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed qed (simp_all add:hyperset_defs) (*>*) lemma last_index_image_set: "distinct xs \ last_index xs ` set xs = {..*) lemma D_compE\<^sub>1: "\ \ e \set Vs\; fv e \ set Vs; distinct Vs \ \ \ (compE\<^sub>1 Vs e) \{.." (*<*)by(fastforce dest!: D_last_index_compE\<^sub>1[OF subset_refl] simp add:last_index_image_set)(*>*) lemma D_compE\<^sub>1': assumes "\ e \set(V#Vs)\" and "fv e \ set(V#Vs)" and "distinct(V#Vs)" shows "\ (compE\<^sub>1 (V#Vs) e) \{..length Vs}\" (*<*) proof - have "{..size Vs} = {..1) qed (*>*) lemma compP\<^sub>1_pres_wf: "wf_J_prog P \ wf_J\<^sub>1_prog (compP\<^sub>1 P)" (*<*) -apply simp -apply(rule wf_prog_compPI) - prefer 2 apply assumption -apply(case_tac m) -apply(simp add:wf_mdecl_def wf_J\<^sub>1_mdecl_def) -apply(clarify) apply(rename_tac C M b Ts T x1 x2 pns body) -apply(case_tac b) - apply clarsimp - apply(frule WT_fv) - apply(auto intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1 \)[1] -apply clarsimp -apply(frule WT_fv) -apply(fastforce intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1' \) -done +proof - + assume wf: "wf_J_prog P" + let ?f = "(\b (pns, body). + compE\<^sub>1 (case b of Static \ pns | NonStatic \ this # pns) body)" + let ?wf\<^sub>2 = "wf_J\<^sub>1_mdecl" + + { fix C M b Ts T m + assume cM: "P \ C sees M, b : Ts\T = m in C" + and wfm: "wf_mdecl wf_J_mdecl P C (M, b, Ts, T, m)" + obtain pns body where [simp]: "m = (pns, body)" by(cases m) simp + let ?E = "\b. case b of Static \ [pns [\] Ts] | NonStatic \ [pns [\] Ts, this \ Class C]" + obtain T' where WT: "P,?E b \ body :: T'" and subT: "P \ T' \ T" + using wfm by(cases b) (auto simp: wf_mdecl_def) + have fv: "fv body \ dom (?E b)" by(rule WT_fv[OF WT]) + have "wf_mdecl ?wf\<^sub>2 (compP ?f P) C (M, b, Ts, T, ?f b m)" + proof(cases b) + case Static then show ?thesis using cM wfm fv + by(auto simp:wf_mdecl_def wf_J\<^sub>1_mdecl_def + intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1 \) + next + case NonStatic then show ?thesis using cM wfm fv + by(clarsimp simp add:wf_mdecl_def wf_J\<^sub>1_mdecl_def) + (fastforce intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1' \) + qed + } + then show ?thesis by simp (rule wf_prog_compPI[OF _ wf]) +qed (*>*) end diff --git a/thys/JinjaDCI/Compiler/Correctness2.thy b/thys/JinjaDCI/Compiler/Correctness2.thy --- a/thys/JinjaDCI/Compiler/Correctness2.thy +++ b/thys/JinjaDCI/Compiler/Correctness2.thy @@ -1,3202 +1,3314 @@ (* Title: JinjaDCI/Compiler/Correctness2.thy Author: Tobias Nipkow, Susannah Mansky Copyright TUM 2003, UIUC 2019-20 Based on the Jinja theory Compiler/Correctness2.thy by Tobias Nipkow *) section \ Correctness of Stage 2 \ theory Correctness2 imports "HOL-Library.Sublist" Compiler2 J1WellForm "../J/EConform" begin (*<*)hide_const (open) Throw(*>*) subsection\ Instruction sequences \ text\ How to select individual instructions and subsequences of instructions from a program given the class, method and program counter. \ definition before :: "jvm_prog \ cname \ mname \ nat \ instr list \ bool" ("(_,_,_,_/ \ _)" [51,0,0,0,51] 50) where "P,C,M,pc \ is \ prefix is (drop pc (instrs_of P C M))" definition at :: "jvm_prog \ cname \ mname \ nat \ instr \ bool" ("(_,_,_,_/ \ _)" [51,0,0,0,51] 50) where "P,C,M,pc \ i \ (\is. drop pc (instrs_of P C M) = i#is)" lemma [simp]: "P,C,M,pc \ []" (*<*)by(simp add:before_def)(*>*) lemma [simp]: "P,C,M,pc \ (i#is) = (P,C,M,pc \ i \ P,C,M,pc + 1 \ is)" (*<*)by(fastforce simp add:before_def at_def prefix_def drop_Suc drop_tl)(*>*) (*<*) declare drop_drop[simp del] (*>*) lemma [simp]: "P,C,M,pc \ (is\<^sub>1 @ is\<^sub>2) = (P,C,M,pc \ is\<^sub>1 \ P,C,M,pc + size is\<^sub>1 \ is\<^sub>2)" (*<*) -apply(simp add:before_def prefix_def) -apply(subst add.commute) -apply(simp add: drop_drop[symmetric]) -apply fastforce -done +by(subst add.commute) + (fastforce simp add:before_def prefix_def drop_drop[symmetric]) (*>*) (*<*) declare drop_drop[simp] (*>*) lemma [simp]: "P,C,M,pc \ i \ instrs_of P C M ! pc = i" (*<*)by(clarsimp simp add:at_def strict_prefix_def nth_via_drop)(*>*) lemma beforeM: "P \ C sees M,b: Ts\T = body in D \ compP\<^sub>2 P,D,M,0 \ compE\<^sub>2 body @ [Return]" -(*<*) -apply(drule sees_method_idemp) -apply(simp add:before_def compP\<^sub>2_def compMb\<^sub>2_def) -done -(*>*) +(*<*)by(drule sees_method_idemp) (simp add:before_def compMb\<^sub>2_def)(*>*) text\ This lemma executes a single instruction by rewriting: \ lemma [simp]: "P,C,M,pc \ instr \ (P \ (None, h, (vs,ls,C,M,pc,ics) # frs, sh) -jvm\ \') = ((None, h, (vs,ls,C,M,pc,ics) # frs, sh) = \' \ (\\. exec(P,(None, h, (vs,ls,C,M,pc,ics) # frs, sh)) = Some \ \ P \ \ -jvm\ \'))" (*<*) -apply(simp only: exec_all_def) -apply(blast intro: converse_rtranclE converse_rtrancl_into_rtrancl) -done +by(simp only: exec_all_def) + (blast intro: converse_rtranclE converse_rtrancl_into_rtrancl) (*>*) subsection\ Exception tables \ definition pcs :: "ex_table \ nat set" where "pcs xt \ \(f,t,C,h,d) \ set xt. {f ..< t}" lemma pcs_subset: shows "(\pc d. pcs(compxE\<^sub>2 e pc d) \ {pc..2 e)})" and "(\pc d. pcs(compxEs\<^sub>2 es pc d) \ {pc..2 es)})" (*<*) -apply(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) -apply (simp_all add:pcs_def) -apply (fastforce split:bop.splits)+ -done +proof(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) + case Cast then show ?case by (fastforce simp:pcs_def) +next + case BinOp then show ?case by (fastforce simp:pcs_def split:bop.splits) +next + case LAss then show ?case by (fastforce simp: pcs_def) +next + case FAcc then show ?case by (fastforce simp: pcs_def) +next + case FAss then show ?case by (fastforce simp: pcs_def) +next + case SFAss then show ?case by (fastforce simp: pcs_def) +next + case Call then show ?case by (fastforce simp: pcs_def) +next + case SCall then show ?case by (fastforce simp: pcs_def) +next + case Seq then show ?case by (fastforce simp: pcs_def) +next + case Cond then show ?case by (fastforce simp: pcs_def) +next + case While then show ?case by (fastforce simp: pcs_def) +next + case throw then show ?case by (fastforce simp: pcs_def) +next + case TryCatch then show ?case by (fastforce simp: pcs_def) +next + case Cons_exp then show ?case by (fastforce simp: pcs_def) +qed (simp_all add:pcs_def) (*>*) lemma [simp]: "pcs [] = {}" (*<*)by(simp add:pcs_def)(*>*) lemma [simp]: "pcs (x#xt) = {fst x ..< fst(snd x)} \ pcs xt" (*<*)by(auto simp add: pcs_def)(*>*) lemma [simp]: "pcs(xt\<^sub>1 @ xt\<^sub>2) = pcs xt\<^sub>1 \ pcs xt\<^sub>2" (*<*)by(simp add:pcs_def)(*>*) lemma [simp]: "pc < pc\<^sub>0 \ pc\<^sub>0+size(compE\<^sub>2 e) \ pc \ pc \ pcs(compxE\<^sub>2 e pc\<^sub>0 d)" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc < pc\<^sub>0 \ pc\<^sub>0+size(compEs\<^sub>2 es) \ pc \ pc \ pcs(compxEs\<^sub>2 es pc\<^sub>0 d)" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc\<^sub>1 + size(compE\<^sub>2 e\<^sub>1) \ pc\<^sub>2 \ pcs(compxE\<^sub>2 e\<^sub>1 pc\<^sub>1 d\<^sub>1) \ pcs(compxE\<^sub>2 e\<^sub>2 pc\<^sub>2 d\<^sub>2) = {}" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc\<^sub>1 + size(compE\<^sub>2 e) \ pc\<^sub>2 \ pcs(compxE\<^sub>2 e pc\<^sub>1 d\<^sub>1) \ pcs(compxEs\<^sub>2 es pc\<^sub>2 d\<^sub>2) = {}" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc \ pcs xt\<^sub>0 \ match_ex_table P C pc (xt\<^sub>0 @ xt\<^sub>1) = match_ex_table P C pc xt\<^sub>1" (*<*)by (induct xt\<^sub>0) (auto simp: matches_ex_entry_def)(*>*) lemma [simp]: "\ x \ set xt; pc \ pcs xt \ \ \ matches_ex_entry P D pc x" (*<*)by(auto simp:matches_ex_entry_def pcs_def)(*>*) lemma [simp]: assumes xe: "xe \ set(compxE\<^sub>2 e pc d)" and outside: "pc' < pc \ pc+size(compE\<^sub>2 e) \ pc'" shows "\ matches_ex_entry P C pc' xe" (*<*) proof assume "matches_ex_entry P C pc' xe" with xe have "pc' \ pcs(compxE\<^sub>2 e pc d)" by(force simp add:matches_ex_entry_def pcs_def) with outside show False by simp qed (*>*) lemma [simp]: assumes xe: "xe \ set(compxEs\<^sub>2 es pc d)" and outside: "pc' < pc \ pc+size(compEs\<^sub>2 es) \ pc'" shows "\ matches_ex_entry P C pc' xe" (*<*) proof assume "matches_ex_entry P C pc' xe" with xe have "pc' \ pcs(compxEs\<^sub>2 es pc d)" by(force simp add:matches_ex_entry_def pcs_def) with outside show False by simp qed (*>*) lemma match_ex_table_app[simp]: "\xte \ set xt\<^sub>1. \ matches_ex_entry P D pc xte \ match_ex_table P D pc (xt\<^sub>1 @ xt) = match_ex_table P D pc xt" (*<*)by(induct xt\<^sub>1) simp_all(*>*) lemma [simp]: "\x \ set xtab. \ matches_ex_entry P C pc x \ match_ex_table P C pc xtab = None" (*<*)using match_ex_table_app[where ?xt = "[]"] by fastforce(*>*) lemma match_ex_entry: "matches_ex_entry P C pc (start, end, catch_type, handler) = (start \ pc \ pc < end \ P \ C \\<^sup>* catch_type)" (*<*)by(simp add:matches_ex_entry_def)(*>*) definition caught :: "jvm_prog \ pc \ heap \ addr \ ex_table \ bool" where "caught P pc h a xt \ (\entry \ set xt. matches_ex_entry P (cname_of h a) pc entry)" definition beforex :: "jvm_prog \ cname \ mname \ ex_table \ nat set \ nat \ bool" ("(2_,/_,/_ \/ _ /'/ _,/_)" [51,0,0,0,0,51] 50) where "P,C,M \ xt / I,d \ (\xt\<^sub>0 xt\<^sub>1. ex_table_of P C M = xt\<^sub>0 @ xt @ xt\<^sub>1 \ pcs xt\<^sub>0 \ I = {} \ pcs xt \ I \ (\pc \ I. \C pc' d'. match_ex_table P C pc xt\<^sub>1 = \(pc',d')\ \ d' \ d))" definition dummyx :: "jvm_prog \ cname \ mname \ ex_table \ nat set \ nat \ bool" ("(2_,_,_ \/ _ '/_,_)" [51,0,0,0,0,51] 50) where "P,C,M \ xt/I,d \ P,C,M \ xt/I,d" +abbreviation +"beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1 + \ ex_table_of P C M = xt\<^sub>0 @ xt @ xt\<^sub>1 \ pcs xt\<^sub>0 \ I = {} + \ pcs xt \ I \ (\pc \ I. \C pc' d'. match_ex_table P C pc xt\<^sub>1 = \(pc',d')\ \ d' \ d)" + +lemma beforex_beforex\<^sub>0_eq: + "P,C,M \ xt / I,d \ \xt\<^sub>0 xt\<^sub>1. beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1" +using beforex_def by auto + lemma beforexD1: "P,C,M \ xt / I,d \ pcs xt \ I" (*<*)by(auto simp add:beforex_def)(*>*) lemma beforex_mono: "\ P,C,M \ xt/I,d'; d' \ d \ \ P,C,M \ xt/I,d" (*<*)by(fastforce simp:beforex_def)(*>*) lemma [simp]: "P,C,M \ xt/I,d \ P,C,M \ xt/I,Suc d" (*<*)by(fastforce intro:beforex_mono)(*>*) lemma beforex_append[simp]: "pcs xt\<^sub>1 \ pcs xt\<^sub>2 = {} \ P,C,M \ xt\<^sub>1 @ xt\<^sub>2/I,d = (P,C,M \ xt\<^sub>1/I-pcs xt\<^sub>2,d \ P,C,M \ xt\<^sub>2/I-pcs xt\<^sub>1,d \ P,C,M \ xt\<^sub>1@xt\<^sub>2/I,d)" -(*<*) -apply(rule iffI) - prefer 2 - apply(simp add:dummyx_def) -apply(auto simp add: beforex_def dummyx_def) - apply(rule_tac x = xt\<^sub>0 in exI) - apply auto -apply(rule_tac x = "xt\<^sub>0@xt\<^sub>1" in exI) -apply auto -done +(*<*)(is "?Q \ ?P = (?P1 \ ?P2 \ ?P3)" is "?Q \ ?P = ?P123") +proof - + assume pcs: ?Q + show ?thesis proof(rule iffI) + assume "?P123" then show ?P by(simp add:dummyx_def) + next + assume hyp: ?P + let ?xt = "xt\<^sub>1 @ xt\<^sub>2" + let ?beforex = "beforex\<^sub>0 P C M d" + obtain xt\<^sub>0 xt\<^sub>1' where beforex: "?beforex I ?xt xt\<^sub>0 xt\<^sub>1'" + using hyp by(clarsimp simp: beforex_def) + have "\xt\<^sub>0 xt\<^sub>1'. ?beforex (I - pcs xt\<^sub>2) xt\<^sub>1 xt\<^sub>0 xt\<^sub>1'" \ \?P1\ + using pcs beforex by(rule_tac x=xt\<^sub>0 in exI) auto + moreover have "\xt\<^sub>0 xt\<^sub>1'. ?beforex (I - pcs xt\<^sub>1) xt\<^sub>2 xt\<^sub>0 xt\<^sub>1'" \ \?P2\ + using pcs beforex by(rule_tac x="xt\<^sub>0@xt\<^sub>1" in exI) auto + moreover have ?P3 using hyp by(simp add: dummyx_def) + ultimately show ?P123 by (simp add: beforex_def) + qed +qed (*>*) lemma beforex_appendD1: - "\ P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d; - pcs xt\<^sub>1 \ J; J \ I; J \ pcs xt\<^sub>2 = {} \ - \ P,C,M \ xt\<^sub>1 / J,d" +assumes bx: "P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d" + and pcs: "pcs xt\<^sub>1 \ J" and JI: "J \ I" and Jpcs: "J \ pcs xt\<^sub>2 = {}" +shows "P,C,M \ xt\<^sub>1 / J,d" (*<*) -apply(auto simp:beforex_def) -apply(rule exI,rule exI,rule conjI, rule refl) -apply(rule conjI, blast) -apply(auto) -apply(subgoal_tac "pc \ pcs xt\<^sub>2") - prefer 2 apply blast -apply (auto split:if_split_asm) -done +proof - + let ?beforex = "beforex\<^sub>0 P C M d" + obtain xt\<^sub>0 xt\<^sub>1' where bx': "?beforex I (xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)]) xt\<^sub>0 xt\<^sub>1'" + using bx by(clarsimp simp:beforex_def) + let ?xt0 = xt\<^sub>0 and ?xt1 = "xt\<^sub>2 @ (f, t, D, h, d) # xt\<^sub>1'" + have "pcs xt\<^sub>0 \ J = {}" using bx' JI by blast + moreover { + fix pc C pc' d' assume pcJ: "pc\J" + then have "pc \ pcs xt\<^sub>2" using bx' Jpcs by blast + then have "match_ex_table P C pc (xt\<^sub>2 @ (f, t, D, h, d) # xt\<^sub>1') + = \(pc', d')\ \ d' \ d" + using bx' JI pcJ by (auto split:if_split_asm) + } + ultimately have "?beforex J xt\<^sub>1 ?xt0 ?xt1" using bx' pcs by simp + then show ?thesis using beforex_def by blast +qed (*>*) lemma beforex_appendD2: - "\ P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d; - pcs xt\<^sub>2 \ J; J \ I; J \ pcs xt\<^sub>1 = {} \ - \ P,C,M \ xt\<^sub>2 / J,d" +assumes bx: "P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d" + and pcs: "pcs xt\<^sub>2 \ J" and JI: "J \ I" and Jpcs: "J \ pcs xt\<^sub>1 = {}" +shows "P,C,M \ xt\<^sub>2 / J,d" (*<*) -apply(auto simp:beforex_def) -apply(rule_tac x = "xt\<^sub>0 @ xt\<^sub>1" in exI) -apply fastforce -done +proof - + let ?beforex = "beforex\<^sub>0 P C M d" + obtain xt\<^sub>0 xt\<^sub>1' where bx': "?beforex I (xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)]) xt\<^sub>0 xt\<^sub>1'" + using bx by(clarsimp simp:beforex_def) + then have "\xt\<^sub>1''. beforex\<^sub>0 P C M d J xt\<^sub>2 (xt\<^sub>0 @ xt\<^sub>1) xt\<^sub>1''" + using assms by fastforce + then show ?thesis using beforex_def by blast +qed (*>*) lemma beforexM: "P \ C sees M,b: Ts\T = body in D \ compP\<^sub>2 P,D,M \ compxE\<^sub>2 body 0 0/{..2 body)},0" (*<*) -apply(drule sees_method_idemp) -apply(drule sees_method_compP[where f = compMb\<^sub>2]) -apply(simp add:beforex_def compP\<^sub>2_def compMb\<^sub>2_def) -apply(rule_tac x = "[]" in exI) -using pcs_subset apply fastforce -done +proof - + assume cM: "P \ C sees M,b: Ts\T = body in D" + let ?xt0 = "[]" + have "\xt1. beforex\<^sub>0 (compP\<^sub>2 P) D M 0 ({..2 body)}) (compxE\<^sub>2 body 0 0) ?xt0 xt1" + using sees_method_compP[where f = compMb\<^sub>2, OF sees_method_idemp[OF cM]] + pcs_subset by(fastforce simp add: compP\<^sub>2_def compMb\<^sub>2_def) + then show ?thesis using beforex_def by blast +qed (*>*) lemma match_ex_table_SomeD2: - "\ match_ex_table P D pc (ex_table_of P C M) = \(pc',d')\; - P,C,M \ xt/I,d; \x \ set xt. \ matches_ex_entry P D pc x; pc \ I \ - \ d' \ d" +assumes met: "match_ex_table P D pc (ex_table_of P C M) = \(pc',d')\" + and bx: "P,C,M \ xt/I,d" + and nmet: "\x \ set xt. \ matches_ex_entry P D pc x" and pcI: "pc \ I" +shows "d' \ d" (*<*) -apply(auto simp:beforex_def) -apply(subgoal_tac "pc \ pcs xt\<^sub>0") -apply auto -done +proof - + obtain xt\<^sub>0 xt\<^sub>1 where bx': "beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1" + using bx by(clarsimp simp:beforex_def) + then have "pc \ pcs xt\<^sub>0" using pcI by blast + then show ?thesis using bx' met nmet pcI by simp +qed (*>*) lemma match_ex_table_SomeD1: "\ match_ex_table P D pc (ex_table_of P C M) = \(pc',d')\; P,C,M \ xt / I,d; pc \ I; pc \ pcs xt \ \ d' \ d" (*<*)by(auto elim: match_ex_table_SomeD2)(*>*) subsection\ The correctness proof \ (*<*) declare nat_add_distrib[simp] caught_def[simp] declare fun_upd_apply[simp del] (*>*) definition handle :: "jvm_prog \ cname \ mname \ addr \ heap \ val list \ val list \ nat \ init_call_status \ frame list \ sheap \ jvm_state" where "handle P C M a h vs ls pc ics frs sh = find_handler P a h ((vs,ls,C,M,pc,ics) # frs) sh" lemma aux_isin[simp]: "\ B \ A; a \ B \ \ a \ A" (*<*)by blast(*>*) lemma handle_frs_tl_neq: "ics_of f \ No_ics \ (xp, h, f#frs, sh) \ handle P C M xa h' vs l pc ics frs sh'" by(simp add: handle_def find_handler_frs_tl_neq del: find_handler.simps) subsubsection "Correctness proof inductive hypothesis" \ \ frame definitions for use by correctness proof inductive hypothesis \ fun calling_to_called :: "frame \ frame" where "calling_to_called (stk,loc,D,M,pc,ics) = (stk,loc,D,M,pc,case ics of Calling C Cs \ Called (C#Cs))" fun calling_to_scalled :: "frame \ frame" where "calling_to_scalled (stk,loc,D,M,pc,ics) = (stk,loc,D,M,pc,case ics of Calling C Cs \ Called Cs)" fun calling_to_calling :: "frame \ cname \ frame" where "calling_to_calling (stk,loc,D,M,pc,ics) C' = (stk,loc,D,M,pc,case ics of Calling C Cs \ Calling C' (C#Cs))" fun calling_to_throwing :: "frame \ addr \ frame" where "calling_to_throwing (stk,loc,D,M,pc,ics) a = (stk,loc,D,M,pc,case ics of Calling C Cs \ Throwing (C#Cs) a)" fun calling_to_sthrowing :: "frame \ addr \ frame" where "calling_to_sthrowing (stk,loc,D,M,pc,ics) a = (stk,loc,D,M,pc,case ics of Calling C Cs \ Throwing Cs a)" \ \ pieces of the correctness proof's inductive hypothesis, which depend on the expression being compiled) \ fun Jcc_cond :: "J\<^sub>1_prog \ ty list \ cname \ mname \ val list \ pc \ init_call_status \ nat set \ heap \ sheap \ expr\<^sub>1 \ bool" where "Jcc_cond P E C M vs pc ics I h sh (INIT C\<^sub>0 (Cs,b) \ e') = ((\T. P,E,h,sh \\<^sub>1 INIT C\<^sub>0 (Cs,b) \ e' : T) \ unit = e' \ ics = No_ics)" | "Jcc_cond P E C M vs pc ics I h sh (RI(C',e\<^sub>0);Cs \ e') = (((e\<^sub>0 = C'\\<^sub>sclinit([]) \ (\T. P,E,h,sh \\<^sub>1 RI(C',e\<^sub>0);Cs \ e':T)) \ ((\a. e\<^sub>0 = Throw a) \ (\C \ set(C'#Cs). is_class P C))) \ unit = e' \ ics = No_ics)" | "Jcc_cond P E C M vs pc ics I h sh (C'\\<^sub>sM'(es)) = (let e = (C'\\<^sub>sM'(es)) in if M' = clinit \ es = [] then (\T. P,E,h,sh \\<^sub>1 e:T) \ (\Cs. ics = Called Cs) else (compP\<^sub>2 P,C,M,pc \ compE\<^sub>2 e \ compP\<^sub>2 P,C,M \ compxE\<^sub>2 e pc (size vs)/I,size vs \ {pc..2 e)} \ I \ \sub_RI e \ ics = No_ics) )" | "Jcc_cond P E C M vs pc ics I h sh e = (compP\<^sub>2 P,C,M,pc \ compE\<^sub>2 e \ compP\<^sub>2 P,C,M \ compxE\<^sub>2 e pc (size vs)/I,size vs \ {pc..2 e)} \ I \ \sub_RI e \ ics = No_ics)" fun Jcc_frames :: "jvm_prog \ cname \ mname \ val list \ val list \ pc \ init_call_status \ frame list \ expr\<^sub>1 \ frame list" where "Jcc_frames P C M vs ls pc ics frs (INIT C\<^sub>0 (C'#Cs,b) \ e') = (case b of False \ (vs,ls,C,M,pc,Calling C' Cs) # frs | True \ (vs,ls,C,M,pc,Called (C'#Cs)) # frs )" | "Jcc_frames P C M vs ls pc ics frs (INIT C\<^sub>0 (Nil,b) \ e') = (vs,ls,C,M,pc,Called [])#frs" | "Jcc_frames P C M vs ls pc ics frs (RI(C',e\<^sub>0);Cs \ e') = (case e\<^sub>0 of Throw a \ (vs,ls,C,M,pc,Throwing (C'#Cs) a) # frs | _ \ (vs,ls,C,M,pc,Called (C'#Cs)) # frs )" | "Jcc_frames P C M vs ls pc ics frs (C'\\<^sub>sM'(es)) = (if M' = clinit \ es = [] then create_init_frame P C'#(vs,ls,C,M,pc,ics)#frs else (vs,ls,C,M,pc,ics)#frs )" | "Jcc_frames P C M vs ls pc ics frs e = (vs,ls,C,M,pc,ics)#frs" fun Jcc_rhs :: "J\<^sub>1_prog \ ty list \ cname \ mname \ val list \ val list \ pc \ init_call_status \ frame list \ heap \ val list \ sheap \ val \ expr\<^sub>1 \ jvm_state" where "Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v (INIT C\<^sub>0 (Cs,b) \ e') = (None,h',(vs,ls,C,M,pc,Called [])#frs,sh')" | "Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v (RI(C',e\<^sub>0);Cs \ e') = (None,h',(vs,ls,C,M,pc,Called [])#frs,sh')" | "Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v (C'\\<^sub>sM'(es)) = (let e = (C'\\<^sub>sM'(es)) in if M' = clinit \ es = [] then (None,h',(vs,ls,C,M,pc,ics)#frs,sh'(C'\(fst(the(sh' C')),Done))) else (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 e),ics)#frs,sh') )" | "Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v e = (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 e),ics)#frs,sh')" fun Jcc_err :: "jvm_prog \ cname \ mname \ heap \ val list \ val list \ pc \ init_call_status \ frame list \ sheap \ nat set \ heap \ val list \ sheap \ addr \ expr\<^sub>1 \ bool" where "Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa (INIT C\<^sub>0 (Cs,b) \ e') = (\vs'. P \ (None,h,Jcc_frames P C M vs ls pc ics frs (INIT C\<^sub>0 (Cs,b) \ e'),sh) -jvm\ handle P C M xa h' (vs'@vs) ls pc ics frs sh')" | "Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa (RI(C',e\<^sub>0);Cs \ e') = (\vs'. P \ (None,h,Jcc_frames P C M vs ls pc ics frs (RI(C',e\<^sub>0);Cs \ e'),sh) -jvm\ handle P C M xa h' (vs'@vs) ls pc ics frs sh')" | "Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa (C'\\<^sub>sM'(es)) = (let e = (C'\\<^sub>sM'(es)) in if M' = clinit \ es = [] then case ics of Called Cs \ P \ (None,h,Jcc_frames P C M vs ls pc ics frs e,sh) -jvm\ (None,h',(vs,ls,C,M,pc,Throwing Cs xa)#frs,(sh'(C' \ (fst(the(sh' C')),Error)))) else (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compE\<^sub>2 e) \ \ caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (size vs)) \ (\vs'. P \ (None,h,Jcc_frames P C M vs ls pc ics frs e,sh) -jvm\ handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh')) )" | "Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa e = (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compE\<^sub>2 e) \ \ caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (size vs)) \ (\vs'. P \ (None,h,Jcc_frames P C M vs ls pc ics frs e,sh) -jvm\ handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh'))" fun Jcc_pieces :: "J\<^sub>1_prog \ ty list \ cname \ mname \ heap \ val list \ val list \ pc \ init_call_status \ frame list \ sheap \ nat set \ heap \ val list \ sheap \ val \ addr \ expr\<^sub>1 \ bool \ frame list \ jvm_state \ bool" where "Jcc_pieces P E C M h vs ls pc ics frs sh I h' ls' sh' v xa e = (Jcc_cond P E C M vs pc ics I h sh e, Jcc_frames (compP\<^sub>2 P) C M vs ls pc ics frs e, Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v e, Jcc_err (compP\<^sub>2 P) C M h vs ls pc ics frs sh I h' ls' sh' xa e)" \ \ @{text Jcc_pieces} lemmas \ lemma nsub_RI_Jcc_pieces: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and nsub: "\sub_RI e" shows "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa e = (let cond = P,C,M,pc \ compE\<^sub>2 e \ P,C,M \ compxE\<^sub>2 e pc (size vs)/I,size vs \ {pc..2 e)} \ I \ ics = No_ics; frs' = (vs,ls,C,M,pc,ics)#frs; rhs = (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 e),ics)#frs,sh'); err = (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compE\<^sub>2 e) \ \ caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (size vs)) \ (\vs'. P \ (None,h,frs',sh) -jvm\ handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh')) in (cond, frs',rhs, err) )" proof - have NC: "\C'. e \ C'\\<^sub>sclinit([])" using assms(2) proof(cases e) qed(simp_all) then show ?thesis using assms proof(cases e) case (SCall C M es) then have "M \ clinit" using nsub by simp then show ?thesis using SCall nsub proof(cases es) qed(simp_all) qed(simp_all) qed lemma Jcc_pieces_Cast: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e) = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e = (True, frs\<^sub>0, (xp',h',(v'#vs',ls',C\<^sub>0,M',pc' - 1,ics')#frs',sh'), (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compE\<^sub>2 e) \ \ caught P pc\<^sub>1 h\<^sub>1 xa (compxE\<^sub>2 e pc (size vs)) \ (\vs'. P \ (None,h\<^sub>0,frs\<^sub>0,sh\<^sub>0) -jvm\ handle P C M xa h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)))" proof - have pc: "{pc..2 e)} \ I" using assms by clarsimp show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] pc by clarsimp qed lemma Jcc_pieces_BinOp1: assumes "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e \bop\ e') = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "\err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 (I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs')))) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e = (True, frs\<^sub>0, (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C\<^sub>0,M',pc' - size (compE\<^sub>2 e') - 1,ics')#frs',sh\<^sub>1), err)" proof - have bef: "compP compMb\<^sub>2 P,C\<^sub>0,M' \ compxE\<^sub>2 e pc (length vs) / I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs'))),length vs" using assms by clarsimp have vs: "vs = vs'" using assms by simp show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] bef vs by clarsimp qed lemma Jcc_pieces_BinOp2: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e \bop\ e') = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "\err. Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (v\<^sub>1#vs) ls\<^sub>1 (pc + size (compE\<^sub>2 e)) ics frs sh\<^sub>1 (I - pcs (compxE\<^sub>2 e pc (length vs'))) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v' xa e' = (True, (v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs, (xp',h',(v'#v\<^sub>1#vs',ls',C\<^sub>0,M',pc' - 1,ics')#frs',sh'), (\pc\<^sub>1. pc + size (compE\<^sub>2 e) \ pc\<^sub>1 \ pc\<^sub>1 < pc + size (compE\<^sub>2 e) + length (compE\<^sub>2 e') \ \ caught P pc\<^sub>1 h\<^sub>2 xa (compxE\<^sub>2 e' (pc + size (compE\<^sub>2 e)) (Suc (length vs))) \ (\vs'. P \ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs,sh\<^sub>1) -jvm\ handle P C M xa h\<^sub>2 (vs'@v\<^sub>1#vs) ls\<^sub>2 pc\<^sub>1 ics frs sh\<^sub>2)))" proof - have bef: "compP compMb\<^sub>2 P\<^sub>1,C\<^sub>0,M' \ compxE\<^sub>2 e pc (length vs) / I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs'))),length vs" using assms by clarsimp have vs: "vs = vs'" using assms by simp show ?thesis using assms nsub_RI_Jcc_pieces[where e=e'] bef vs by clarsimp qed lemma Jcc_pieces_FAcc: assumes "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\F{D}) = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "\err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e = (True, frs\<^sub>0, (xp',h',(v'#vs',ls',C\<^sub>0,M',pc' - 1,ics')#frs',sh'), err)" proof - have pc: "{pc..2 e)} \ I" using assms by clarsimp then show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by clarsimp qed lemma Jcc_pieces_LAss: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (i:=e) = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e = (True, frs\<^sub>0, (xp',h',(v'#vs',ls',C\<^sub>0,M',pc' - 2,ics')#frs',sh'), (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compE\<^sub>2 e) \ \ caught P pc\<^sub>1 h\<^sub>1 xa (compxE\<^sub>2 e pc (size vs)) \ (\vs'. P \ (None,h\<^sub>0,frs\<^sub>0,sh\<^sub>0) -jvm\ handle P C M xa h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)))" proof - have pc: "{pc..2 e)} \ I" using assms by clarsimp show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] pc by clarsimp qed lemma Jcc_pieces_FAss1: assumes "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\F{D}:=e') = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "\err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 (I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs')))) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e = (True, frs\<^sub>0, (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C\<^sub>0,M',pc' - size (compE\<^sub>2 e') - 2,ics')#frs',sh\<^sub>1), err)" proof - show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by clarsimp qed lemma Jcc_pieces_FAss2: assumes "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\F{D}:=e') = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "Jcc_pieces P E C M h\<^sub>1 (v\<^sub>1#vs) ls\<^sub>1 (pc + size (compE\<^sub>2 e)) ics frs sh\<^sub>1 (I - pcs (compxE\<^sub>2 e pc (length vs'))) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v' xa e' = (True, (v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs, (xp',h',(v'#v\<^sub>1#vs',ls',C\<^sub>0,M',pc' - 2,ics')#frs',sh'), (\pc\<^sub>1. (pc + size (compE\<^sub>2 e)) \ pc\<^sub>1 \ pc\<^sub>1 < pc + size (compE\<^sub>2 e) + size(compE\<^sub>2 e') \ \ caught (compP\<^sub>2 P) pc\<^sub>1 h\<^sub>2 xa (compxE\<^sub>2 e' (pc + size (compE\<^sub>2 e)) (size (v\<^sub>1#vs))) \ (\vs'. (compP\<^sub>2 P) \ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs,sh\<^sub>1) -jvm\ handle (compP\<^sub>2 P) C M xa h\<^sub>2 (vs'@v\<^sub>1#vs) ls\<^sub>2 pc\<^sub>1 ics frs sh\<^sub>2)))" proof - show ?thesis using assms nsub_RI_Jcc_pieces[where e=e'] by clarsimp qed lemma Jcc_pieces_SFAss: assumes "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h' ls' sh' v xa (C'\\<^sub>sF{D}:=e) = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "\err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e = (True, frs\<^sub>0, (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C\<^sub>0,M',pc' - 2,ics')#frs',sh\<^sub>1), err)" proof - have pc: "{pc..2 e)} \ I" using assms by clarsimp show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] pc by clarsimp qed lemma Jcc_pieces_Call1: assumes "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>3 ls\<^sub>3 sh\<^sub>3 v xa (e\M\<^sub>0(es)) = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C',M',pc',ics')#frs',sh'), err)" shows "\err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs')))) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e = (True, frs\<^sub>0, (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C',M',pc' - size (compEs\<^sub>2 es) - 1,ics')#frs',sh\<^sub>1), err)" proof - show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by clarsimp qed lemma Jcc_pieces_clinit: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and cond: "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (C1\\<^sub>sclinit([]))" shows "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C1\\<^sub>sclinit([])) = (True, create_init_frame P C1 # (vs,ls,C,M,pc,ics)#frs, (None, h', (vs,ls,C,M,pc,ics)#frs, sh'(C1\(fst(the(sh' C1)),Done))), P \ (None,h,create_init_frame P C1 # (vs,ls,C,M,pc,ics)#frs,sh) -jvm\ (case ics of Called Cs \ (None,h',(vs,ls,C,M,pc,Throwing Cs xa)#frs,(sh'(C1 \ (fst(the(sh' C1)),Error))))))" using assms by(auto split: init_call_status.splits list.splits bool.splits) lemma Jcc_pieces_SCall_clinit_body: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa (C1\\<^sub>sclinit([])) = (True, frs', rhs', err')" and method: "P\<^sub>1 \ C1 sees clinit,Static: []\Void = body in D" shows "Jcc_pieces P\<^sub>1 [] D clinit h\<^sub>2 [] (replicate (max_vars body) undefined) 0 No_ics (tl frs') sh\<^sub>2 {..2 body)} h\<^sub>3 ls\<^sub>3 sh\<^sub>3 v xa body = (True, frs', (None,h\<^sub>3,([v],ls\<^sub>3,D,clinit,size(compE\<^sub>2 body), No_ics)#tl frs',sh\<^sub>3), \pc\<^sub>1. 0 \ pc\<^sub>1 \ pc\<^sub>1 < size(compE\<^sub>2 body) \ \ caught P pc\<^sub>1 h\<^sub>3 xa (compxE\<^sub>2 body 0 0) \ (\vs'. P \ (None,h\<^sub>2,frs',sh\<^sub>2) -jvm\ handle P D clinit xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>1 No_ics (tl frs') sh\<^sub>3))" proof - have M_in_D: "P\<^sub>1 \ D sees clinit,Static: []\Void = body in D" using method by(rule sees_method_idemp) hence M_code: "compP\<^sub>2 P\<^sub>1,D,clinit,0 \ compE\<^sub>2 body @ [Return]" and M_xtab: "compP\<^sub>2 P\<^sub>1,D,clinit \ compxE\<^sub>2 body 0 0/{..2 body)},0" by(rule beforeM, rule beforexM) have nsub: "\sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf method]) then show ?thesis using assms nsub_RI_Jcc_pieces M_code M_xtab by clarsimp qed lemma Jcc_pieces_Cons: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "P,C,M,pc \ compEs\<^sub>2 (e#es)" and "P,C,M \ compxEs\<^sub>2 (e#es) pc (size vs)/I,size vs" and "{pc..2 (e#es))} \ I" and "ics = No_ics" and "\sub_RIs (e#es)" shows "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs)))) h' ls' sh' v xa e = (True, (vs, ls, C, M, pc, ics) # frs, (None, h', (v#vs, ls', C, M, pc + length (compE\<^sub>2 e), ics) # frs, sh'), \pc\<^sub>1\pc. pc\<^sub>1 < pc + length (compE\<^sub>2 e) \ \ caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (length vs)) \ (\vs'. P \ (None, h, (vs, ls, C, M, pc, ics) # frs, sh) -jvm\ handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh'))" proof - show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by auto qed lemma Jcc_pieces_InitNone: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" shows "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs (sh(C\<^sub>0 \ (sblank P C\<^sub>0, Prepared))) I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), \vs'. P \ (None,h,frs',(sh(C\<^sub>0 \ (sblank P\<^sub>1 C\<^sub>0, Prepared)))) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh')" proof - have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \ e)" using assms by simp then obtain T where "P\<^sub>1,E,h,sh \\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \ unit : T" by fastforce then have "P\<^sub>1,E,h,sh(C\<^sub>0 \ (sblank P\<^sub>1 C\<^sub>0, Prepared)) \\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \ unit : T" by(auto simp: fun_upd_apply) then have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h (sh(C\<^sub>0 \ (sblank P\<^sub>1 C\<^sub>0, Prepared))) (INIT C' (C\<^sub>0 # Cs,False) \ unit))" by(simp only: exI) then show ?thesis using assms by clarsimp qed lemma Jcc_pieces_InitDP: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" shows "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (Cs,True) \ e) = (True, (calling_to_scalled (hd frs'))#(tl frs'), (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), \vs'. P \ (None,h,calling_to_scalled (hd frs')#(tl frs'),sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh')" proof - have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \ e)" using assms by simp then obtain T where "P\<^sub>1,E,h,sh \\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \ unit : T" by fastforce then have "P\<^sub>1,E,h,sh \\<^sub>1 INIT C' (Cs,True) \ unit : T" by (auto; metis list.sel(2) list.set_sel(2)) then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' (Cs,True) \ unit))" by(simp only: exI) show ?thesis using assms wtrt proof(cases Cs) case (Cons C1 Cs1) then show ?thesis using assms wtrt by(case_tac "method P C1 clinit") clarsimp qed(clarsimp) qed lemma Jcc_pieces_InitError: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" and err: "sh C\<^sub>0 = Some(sfs,Error)" shows "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (RI (C\<^sub>0, THROW NoClassDefFoundError);Cs \ e) = (True, (calling_to_throwing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs'), (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), \vs'. P \ (None,h, (calling_to_throwing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs'),sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh')" proof - show ?thesis using assms proof(cases Cs) case (Cons C1 Cs1) then show ?thesis using assms by(case_tac "method P C1 clinit", case_tac "method P C\<^sub>0 clinit") clarsimp qed(clarsimp) qed lemma Jcc_pieces_InitObj: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' (sh(C\<^sub>0 \ (sfs,Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" shows "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs (sh(C\<^sub>0 \ (sfs,Processing))) I h' l' sh'' v xa (INIT C' (C\<^sub>0 # Cs,True) \ e) = (True, calling_to_called (hd frs')#(tl frs'), (None, h', (vs, l, C, M, pc, Called []) # frs, sh''), \vs'. P \ (None,h,calling_to_called (hd frs')#(tl frs'),sh') -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'')" proof - have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \ e)" using assms by simp then obtain T where "P\<^sub>1,E,h,sh \\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \ unit : T" by fastforce then have "P\<^sub>1,E,h,sh(C\<^sub>0 \ (sfs,Processing)) \\<^sub>1 INIT C' (C\<^sub>0 # Cs,True) \ unit : T" using assms by clarsimp (auto simp: fun_upd_apply) then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h (sh(C\<^sub>0 \ (sfs,Processing))) (INIT C' (C\<^sub>0 # Cs,True) \ unit))" by(simp only: exI) show ?thesis using assms wtrt by clarsimp qed lemma Jcc_pieces_InitNonObj: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "is_class P\<^sub>1 D" and "D \ set (C\<^sub>0#Cs)" and "\C \ set (C\<^sub>0#Cs). P\<^sub>1 \ C \\<^sup>* D" and pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' (sh(C\<^sub>0 \ (sfs,Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" shows "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs (sh(C\<^sub>0 \ (sfs,Processing))) I h' l' sh'' v xa (INIT C' (D # C\<^sub>0 # Cs,False) \ e) = (True, calling_to_calling (hd frs') D#(tl frs'), (None, h', (vs, l, C, M, pc, Called []) # frs, sh''), \vs'. P \ (None,h,calling_to_calling (hd frs') D#(tl frs'),sh') -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'')" proof - have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \ e)" using assms by simp then obtain T where "P\<^sub>1,E,h,sh \\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \ unit : T" by fastforce then have "P\<^sub>1,E,h,sh(C\<^sub>0 \ (sfs,Processing)) \\<^sub>1 INIT C' (D # C\<^sub>0 # Cs,False) \ unit : T" using assms by clarsimp (auto simp: fun_upd_apply) then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h (sh(C\<^sub>0 \ (sfs,Processing))) (INIT C' (D # C\<^sub>0 # Cs,False) \ unit))" by(simp only: exI) show ?thesis using assms wtrt by clarsimp qed lemma Jcc_pieces_InitRInit: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,True) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" shows "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; Cs \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), \vs'. P \ (None,h,frs',sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh')" proof - have cond: "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,True) \ e)" using assms by simp then have clinit: "\T. P\<^sub>1,E,h,sh \\<^sub>1 C\<^sub>0\\<^sub>sclinit([]) : T" using wf by clarsimp (auto simp: is_class_def intro: wf\<^sub>1_types_clinit) then obtain T where cT: "P\<^sub>1,E,h,sh \\<^sub>1 C\<^sub>0\\<^sub>sclinit([]) : T" by blast obtain T where "P\<^sub>1,E,h,sh \\<^sub>1 INIT C' (C\<^sub>0 # Cs,True) \ unit : T" using cond by fastforce then have "P\<^sub>1,E,h,sh \\<^sub>1 RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; Cs \ unit : T" using assms by (auto intro: cT) then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; Cs \ unit))" by(simp only: exI) then show ?thesis using assms by simp qed lemma Jcc_pieces_RInit_clinit: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([]));Cs \ e) = (True, frs', (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)" shows "Jcc_pieces P\<^sub>1 E C M h vs l pc (Called Cs) (tl frs') sh I h' l' sh' v xa (C\<^sub>0\\<^sub>sclinit([])) = (True, create_init_frame P C\<^sub>0#(vs,l,C,M,pc,Called Cs)#tl frs', (None, h', (vs,l,C,M,pc,Called Cs)#tl frs', sh'(C\<^sub>0\(fst(the(sh' C\<^sub>0)),Done))), P \ (None,h,create_init_frame P C\<^sub>0#(vs,l,C,M,pc,Called Cs)#tl frs',sh) -jvm\ (None,h',(vs, l, C, M, pc, Throwing Cs xa) # tl frs',sh'(C\<^sub>0 \ (fst(the(sh' C\<^sub>0)),Error))))" proof - have cond: "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([]));Cs \ e)" using assms by simp then have wtrt: "\T. P\<^sub>1,E,h,sh \\<^sub>1 C\<^sub>0\\<^sub>sclinit([]) : T" using wf by clarsimp (auto simp: is_class_def intro: wf\<^sub>1_types_clinit) then show ?thesis using assms by clarsimp qed lemma Jcc_pieces_RInit_Init: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1" and proc: "\C' \ set Cs. \sfs. sh'' C' = \(sfs,Processing)\" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([]));Cs \ e) = (True, frs', (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)" shows "Jcc_pieces P\<^sub>1 E C M h' vs l pc ics frs sh'' I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (INIT (last (C\<^sub>0#Cs)) (Cs,True) \ e) = (True, (vs, l, C, M, pc, Called Cs) # frs, (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), \vs'. P \ (None,h',(vs, l, C, M, pc, Called Cs) # frs,sh'') -jvm\ handle P C M xa h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1)" proof - have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([]));Cs \ e)" using assms by simp then have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; Cs \ unit))" by simp then obtain T where riwt: "P\<^sub>1,E,h,sh \\<^sub>1 RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([]));Cs \ unit : T" by meson then have "P\<^sub>1,E,h',sh'' \\<^sub>1 INIT (last (C\<^sub>0#Cs)) (Cs,True) \ unit : T" using proc proof(cases Cs) qed(auto) then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h' sh'' (INIT (last (C\<^sub>0#Cs)) (Cs,True) \ unit))" by(simp only: exI) show ?thesis using assms wtrt proof(cases Cs) case (Cons C1 Cs1) then show ?thesis using assms wtrt by(case_tac "method P C1 clinit") clarsimp qed(clarsimp) qed lemma Jcc_pieces_RInit_RInit: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (C\<^sub>0,e);D#Cs \ e') = (True, frs', rhs, err)" and hd: "hd frs' = (vs1,l1,C1,M1,pc1,ics1)" shows "Jcc_pieces P\<^sub>1 E C M h' vs l pc ics frs sh'' I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (D,Throw xa) ; Cs \ e') = (True, (vs1, l1, C1, M1, pc1, Throwing (D#Cs) xa) # tl frs', (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), \vs'. P \ (None,h',(vs1, l1, C1, M1, pc1, Throwing (D#Cs) xa) # tl frs',sh'') -jvm\ handle P C M xa h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1)" using assms by(case_tac "method P D clinit", cases "e = C\<^sub>0\\<^sub>sclinit([])") clarsimp+ subsubsection "JVM stepping lemmas" lemma jvm_Invoke: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "P,C,M,pc \ Invoke M' (length Ts)" and ha: "h\<^sub>2 a = \(Ca, fs)\" and method: "P\<^sub>1 \ Ca sees M', NonStatic : Ts\T = body in D" and len: "length pvs = length Ts" and "ls\<^sub>2' = Addr a # pvs @ replicate (max_vars body) undefined" shows "P \ (None, h\<^sub>2, (rev pvs @ Addr a # vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2) -jvm\ (None, h\<^sub>2, ([], ls\<^sub>2', D, M', 0, No_ics) # (rev pvs @ Addr a # vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2)" proof - have cname: "cname_of h\<^sub>2 (the_Addr ((rev pvs @ Addr a # vs) ! length Ts)) = Ca" using ha method len by(auto simp: nth_append) have r: "(rev pvs @ Addr a # vs) ! (length Ts) = Addr a" using len by(auto simp: nth_append) have exm: "\Ts T m D b. P \ Ca sees M',b:Ts \ T = m in D" using sees_method_compP[OF method] by fastforce show ?thesis using assms cname r exm by simp qed lemma jvm_Invokestatic: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "P,C,M,pc \ Invokestatic C' M' (length Ts)" and sh: "sh\<^sub>2 D = Some(sfs,Done)" and method: "P\<^sub>1 \ C' sees M', Static : Ts\T = body in D" and len: "length pvs = length Ts" and "ls\<^sub>2' = pvs @ replicate (max_vars body) undefined" shows "P \ (None, h\<^sub>2, (rev pvs @ vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2) -jvm\ (None, h\<^sub>2, ([], ls\<^sub>2', D, M', 0, No_ics) # (rev pvs @ vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2)" proof - have exm: "\Ts T m D b. P \ C' sees M',b:Ts \ T = m in D" using sees_method_compP[OF method] by fastforce show ?thesis using assms exm by simp qed lemma jvm_Invokestatic_Called: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "P,C,M,pc \ Invokestatic C' M' (length Ts)" and sh: "sh\<^sub>2 D = Some(sfs,i)" and method: "P\<^sub>1 \ C' sees M', Static : Ts\T = body in D" and len: "length pvs = length Ts" and "ls\<^sub>2' = pvs @ replicate (max_vars body) undefined" shows "P \ (None, h\<^sub>2, (rev pvs @ vs, ls\<^sub>2, C, M, pc, Called []) # frs, sh\<^sub>2) -jvm\ (None, h\<^sub>2, ([], ls\<^sub>2', D, M', 0, No_ics) # (rev pvs @ vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2)" proof - have exm: "\Ts T m D b. P \ C' sees M',b:Ts \ T = m in D" using sees_method_compP[OF method] by fastforce show ?thesis using assms exm by simp qed lemma jvm_Return_Init: "P,D,clinit,0 \ compE\<^sub>2 body @ [Return] \ P \ (None, h, (vs, ls, D, clinit, size(compE\<^sub>2 body), No_ics) # frs, sh) -jvm\ (None, h, frs, sh(D\(fst(the(sh D)),Done)))" -apply(simp add: exec_all_def1, rule r_into_rtrancl, rule exec_1I) -apply(cases frs, auto) -done +(is "?P \ P \ ?s1 -jvm\ ?s2") +proof - + assume ?P + then have "exec (P, ?s1) = \?s2\" by(cases frs) auto + then have "(?s1, ?s2) \ (exec_1 P)\<^sup>*" + by(rule exec_1I[THEN r_into_rtrancl]) + then show ?thesis by(simp add: exec_all_def1) +qed lemma jvm_InitNone: "\ ics_of f = Calling C Cs; sh C = None \ \ P \ (None,h,f#frs,sh) -jvm\ (None,h,f#frs,sh(C \ (sblank P C, Prepared)))" -apply(simp add: exec_all_def1, rule r_into_rtrancl, rule exec_1I) -apply(cases f) apply(rename_tac ics, case_tac ics, simp_all) -done +(is "\ ?P; ?Q \ \ P \ ?s1 -jvm\ ?s2") +proof - + assume assms: ?P ?Q + then obtain stk1 loc1 C1 M1 pc1 ics1 where "f = (stk1,loc1,C1,M1,pc1,ics1)" + by(cases f) simp + then have "exec (P, ?s1) = \?s2\" using assms + by(case_tac ics1) simp_all + then have "(?s1, ?s2) \ (exec_1 P)\<^sup>*" + by(rule exec_1I[THEN r_into_rtrancl]) + then show ?thesis by(simp add: exec_all_def1) +qed lemma jvm_InitDP: "\ ics_of f = Calling C Cs; sh C = \(sfs,i)\; i = Done \ i = Processing \ \ P \ (None,h,f#frs,sh) -jvm\ (None,h,(calling_to_scalled f)#frs,sh)" -apply(simp add: exec_all_def1, rule r_into_rtrancl, rule exec_1I) -apply(cases f) -apply(erule_tac P = "i = Done" in disjE) - apply simp_all -done +(is "\ ?P; ?Q; ?R \ \ P \ ?s1 -jvm\ ?s2") +proof - + assume assms: ?P ?Q ?R + then obtain stk1 loc1 C1 M1 pc1 ics1 where "f = (stk1,loc1,C1,M1,pc1,ics1)" + by(cases f) simp + then have "exec (P, ?s1) = \?s2\" using assms + by(case_tac i) simp_all + then have "(?s1, ?s2) \ (exec_1 P)\<^sup>*" + by(rule exec_1I[THEN r_into_rtrancl]) + then show ?thesis by(simp add: exec_all_def1) +qed lemma jvm_InitError: "sh C = \(sfs,Error)\ \ P \ (None,h,(vs,ls,C\<^sub>0,M,pc,Calling C Cs)#frs,sh) -jvm\ (None,h,(vs,ls,C\<^sub>0,M,pc,Throwing Cs (addr_of_sys_xcpt NoClassDefFoundError))#frs,sh)" by(clarsimp simp: exec_all_def1 intro!: r_into_rtrancl exec_1I) lemma exec_ErrorThrowing: "sh C = \(sfs,Error)\ \ exec (P, (None,h,calling_to_throwing (stk,loc,D,M,pc,Calling C Cs) a#frs,sh)) = Some (None,h,calling_to_sthrowing (stk,loc,D,M,pc,Calling C Cs) a #frs,sh)" by(clarsimp simp: exec_all_def1 fun_upd_idem_iff intro!: r_into_rtrancl exec_1I) lemma jvm_InitObj: "\ sh C = Some(sfs,Prepared); C = Object; sh' = sh(C \ (sfs,Processing)) \ \ P \ (None, h, (vs,ls,C\<^sub>0,M,pc,Calling C Cs)#frs, sh) -jvm\ (None, h, (vs,ls,C\<^sub>0,M,pc,Called (C#Cs))#frs,sh')" -apply(simp add: exec_all_def1, rule r_into_rtrancl, rule exec_1I) -apply(case_tac "method P C clinit", simp) -done +(is "\ ?P; ?Q; ?R \ \ P \ ?s1 -jvm\ ?s2") +proof - + assume assms: ?P ?Q ?R + then have "exec (P, ?s1) = \?s2\" + by(case_tac "method P C clinit") simp + then have "(?s1, ?s2) \ (exec_1 P)\<^sup>*" + by(rule exec_1I[THEN r_into_rtrancl]) + then show ?thesis by(simp add: exec_all_def1) +qed lemma jvm_InitNonObj: "\ sh C = Some(sfs,Prepared); C \ Object; class P C = Some (D,r); sh' = sh(C \ (sfs,Processing)) \ \ P \ (None, h, (vs,ls,C\<^sub>0,M,pc,Calling C Cs)#frs, sh) -jvm\ (None, h, (vs,ls,C\<^sub>0,M,pc,Calling D (C#Cs))#frs, sh')" -apply(simp add: exec_all_def1, rule r_into_rtrancl, rule exec_1I) -apply(case_tac "method P C clinit", simp) -done +(is "\ ?P; ?Q; ?R; ?S \ \ P \ ?s1 -jvm\ ?s2") +proof - + assume assms: ?P ?Q ?R ?S + then have "exec (P, ?s1) = \?s2\" + by(case_tac "method P C clinit") simp + then have "(?s1, ?s2) \ (exec_1 P)\<^sup>*" + by(rule exec_1I[THEN r_into_rtrancl]) + then show ?thesis by(simp add: exec_all_def1) +qed lemma jvm_RInit_throw: "P \ (None,h,(vs,l,C,M,pc,Throwing [] xa) # frs,sh) -jvm\ handle P C M xa h vs l pc No_ics frs sh" -apply(simp add: exec_all_def1, rule r_into_rtrancl, rule exec_1I) -apply(simp add: handle_def split: bool.splits) -done +(is "P \ ?s1 -jvm\ ?s2") +proof - + have "exec (P, ?s1) = \?s2\" + by(simp add: handle_def split: bool.splits) + then have "(?s1, ?s2) \ (exec_1 P)\<^sup>*" + by(rule exec_1I[THEN r_into_rtrancl]) + then show ?thesis by(simp add: exec_all_def1) +qed lemma jvm_RInit_throw': "P \ (None,h,(vs,l,C,M,pc,Throwing [C'] xa) # frs,sh) -jvm\ handle P C M xa h vs l pc No_ics frs (sh(C':=Some(fst(the(sh C')), Error)))" -apply(simp add: exec_all_def1) -apply(rule_tac y = "(None,h,(vs,l,C,M,pc,Throwing [] xa) # frs,sh(C':=Some(fst(the(sh C')), Error)))" in rtrancl_trans) - apply(rule r_into_rtrancl, rule exec_1I) - apply(simp add: handle_def) -apply(cut_tac jvm_RInit_throw) -apply(simp add: exec_all_def1) -done +(is "P \ ?s1 -jvm\ ?s2") +proof - + let ?sy = "(None,h,(vs,l,C,M,pc,Throwing [] xa) # frs,sh(C':=Some(fst(the(sh C')), Error)))" + have "exec (P, ?s1) = \?sy\" by simp + then have "(?s1, ?sy) \ (exec_1 P)\<^sup>*" + by(rule exec_1I[THEN r_into_rtrancl]) + also have "(?sy, ?s2) \ (exec_1 P)\<^sup>*" + using jvm_RInit_throw by(simp add: exec_all_def1) + ultimately show ?thesis by(simp add: exec_all_def1) +qed lemma jvm_Called: "P \ (None, h, (vs, l, C, M, pc, Called (C\<^sub>0 # Cs)) # frs, sh) -jvm\ (None, h, create_init_frame P C\<^sub>0 # (vs, l, C, M, pc, Called Cs) # frs, sh)" by(simp add: exec_all_def1 r_into_rtrancl exec_1I) lemma jvm_Throwing: "P \ (None, h, (vs, l, C, M, pc, Throwing (C\<^sub>0#Cs) xa') # frs, sh) -jvm\ (None, h, (vs, l, C, M, pc, Throwing Cs xa') # frs, sh(C\<^sub>0 \ (fst (the (sh C\<^sub>0)), Error)))" by(simp add: exec_all_def1 r_into_rtrancl exec_1I) subsubsection "Other lemmas for correctness proof" lemma assumes wf:"wf_prog wf_md P" and ex: "class P C = Some a" shows create_init_frame_wf_eq: "create_init_frame (compP\<^sub>2 P) C = (stk,loc,D,M,pc,ics) \ D=C" using wf_sees_clinit[OF wf ex] by(cases "method P C clinit", auto) lemma beforex_try: - "\ {pc..2(try e\<^sub>1 catch(Ci i) e\<^sub>2))} \ I; - P,C,M \ compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs) / I,size vs \ - \ P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs) / {pc..2 e\<^sub>1)},size vs" -apply(clarsimp simp:beforex_def split:if_split_asm) -apply(rename_tac xt\<^sub>0 xt\<^sub>1) apply(rule_tac x=xt\<^sub>0 in exI) -apply(auto simp: pcs_subset(1)) -using atLeastLessThan_iff by blast +assumes pcI: "{pc..2(try e\<^sub>1 catch(Ci i) e\<^sub>2))} \ I" + and bx: "P,C,M \ compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs) / I,size vs" +shows "P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs) / {pc..2 e\<^sub>1)},size vs" +proof - + obtain xt\<^sub>0 xt\<^sub>1 where + "beforex\<^sub>0 P C M (size vs) I (compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs)) xt\<^sub>0 xt\<^sub>1" + using bx by(clarsimp simp:beforex_def) + then have "\xt1. beforex\<^sub>0 P C M (size vs) {pc..2 e\<^sub>1)} + (compxE\<^sub>2 e\<^sub>1 pc (size vs)) xt\<^sub>0 xt1" + using pcI pcs_subset(1) atLeastLessThan_iff by simp blast + then show ?thesis using beforex_def by blast +qed \ \ Evaluation of initialization expressions \ (* --needs J1 and EConform; version for eval found in Equivalence *) lemma shows eval\<^sub>1_init_return: "P \\<^sub>1 \e,s\ \ \e',s'\ \ iconf (shp\<^sub>1 s) e \ (\Cs b. e = INIT C' (Cs,b) \ unit) \ (\C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs@[C'] \ unit) \ (\e\<^sub>0. e = RI(C',e\<^sub>0);Nil \ unit) \ (val_of e' = Some v \ (\sfs i. shp\<^sub>1 s' C' = \(sfs,i)\ \ (i = Done \ i = Processing))) \ (throw_of e' = Some a \ (\sfs i. shp\<^sub>1 s' C' = \(sfs,Error)\))" and "P \\<^sub>1 \es,s\ [\] \es',s'\ \ True" proof(induct rule: eval\<^sub>1_evals\<^sub>1.inducts) case (InitFinal\<^sub>1 e s e' s' C b) then show ?case by(auto simp: initPD_def dest: eval\<^sub>1_final_same) next case (InitDone\<^sub>1 sh C sfs C' Cs e h l e' s') then have "final e'" using eval\<^sub>1_final by simp then show ?case proof(rule finalE) fix v assume e': "e' = Val v" then show ?thesis using InitDone\<^sub>1 initPD_def proof(cases Cs) qed(auto) next fix a assume e': "e' = throw a" then show ?thesis using InitDone\<^sub>1 initPD_def proof(cases Cs) qed(auto) qed next case (InitProcessing\<^sub>1 sh C sfs C' Cs e h l e' s') then have "final e'" using eval\<^sub>1_final by simp then show ?case proof(rule finalE) fix v assume e': "e' = Val v" then show ?thesis using InitProcessing\<^sub>1 initPD_def proof(cases Cs) qed(auto) next fix a assume e': "e' = throw a" then show ?thesis using InitProcessing\<^sub>1 initPD_def proof(cases Cs) qed(auto) qed next case (InitError\<^sub>1 sh C sfs Cs e h l e' s' C') show ?case proof(cases Cs) case Nil then show ?thesis using InitError\<^sub>1 by simp next case (Cons C2 list) then have "final e'" using InitError\<^sub>1 eval\<^sub>1_final by simp then show ?thesis proof(rule finalE) fix v assume e': "e' = Val v" show ?thesis using InitError\<^sub>1.hyps(2) e' rinit\<^sub>1_throwE by blast next fix a assume e': "e' = throw a" then show ?thesis using Cons InitError\<^sub>1 cons_to_append[of list] by clarsimp qed qed next case (InitRInit\<^sub>1 C Cs h l sh e' s' C') show ?case proof(cases Cs) case Nil then show ?thesis using InitRInit\<^sub>1 by simp next case (Cons C' list) then show ?thesis using InitRInit\<^sub>1 Cons cons_to_append[of list] by clarsimp qed next case (RInit\<^sub>1 e s v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 s\<^sub>1) then have final: "final e\<^sub>1" using eval\<^sub>1_final by simp then show ?case proof(cases Cs) case Nil show ?thesis using final proof(rule finalE) fix v assume e': "e\<^sub>1 = Val v" show ?thesis using RInit\<^sub>1 Nil by(clarsimp, meson fun_upd_same initPD_def) next fix a assume e': "e\<^sub>1 = throw a" show ?thesis using RInit\<^sub>1 Nil by(clarsimp, meson fun_upd_same initPD_def) qed next case (Cons a list) show ?thesis using final proof(rule finalE) fix v assume e': "e\<^sub>1 = Val v" then show ?thesis using RInit\<^sub>1 Cons by(clarsimp, metis last.simps last_appendR list.distinct(1)) next fix a assume e': "e\<^sub>1 = throw a" then show ?thesis using RInit\<^sub>1 Cons by(clarsimp, metis last.simps last_appendR list.distinct(1)) qed qed next case (RInitInitFail\<^sub>1 e s a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 s\<^sub>1) then have final: "final e\<^sub>1" using eval\<^sub>1_final by simp then show ?case proof(rule finalE) fix v assume e': "e\<^sub>1 = Val v" then show ?thesis using RInitInitFail\<^sub>1 by(clarsimp, meson exp.distinct(101) rinit\<^sub>1_throwE) next fix a' assume e': "e\<^sub>1 = Throw a'" then have "iconf (sh'(C \ (sfs, Error))) a" using RInitInitFail\<^sub>1.hyps(1) eval\<^sub>1_final by fastforce then show ?thesis using RInitInitFail\<^sub>1 e' by(clarsimp, meson Cons_eq_append_conv list.inject) qed qed(auto simp: fun_upd_same) lemma init\<^sub>1_Val_PD: "P \\<^sub>1 \INIT C' (Cs,b) \ unit,s\ \ \Val v,s'\ \ iconf (shp\<^sub>1 s) (INIT C' (Cs,b) \ unit) \ \sfs i. shp\<^sub>1 s' C' = \(sfs,i)\ \ (i = Done \ i = Processing)" by(drule_tac v = v in eval\<^sub>1_init_return, simp+) lemma init\<^sub>1_throw_PD: "P \\<^sub>1 \INIT C' (Cs,b) \ unit,s\ \ \throw a,s'\ \ iconf (shp\<^sub>1 s) (INIT C' (Cs,b) \ unit) \ \sfs i. shp\<^sub>1 s' C' = \(sfs,Error)\" by(drule_tac a = a in eval\<^sub>1_init_return, simp+) -lemma rinit\<^sub>1_Val_PD: "P \\<^sub>1 \RI(C,e\<^sub>0);Cs \ unit,s\ \ \Val v,s'\ - \ iconf (shp\<^sub>1 s) (RI(C,e\<^sub>0);Cs \ unit) \ last(C#Cs) = C' - \ \sfs i. shp\<^sub>1 s' C' = \(sfs,i)\ \ (i = Done \ i = Processing)" -apply(drule_tac C' = C' and v = v in eval\<^sub>1_init_return, simp_all) -apply (metis append_butlast_last_id) -done +lemma rinit\<^sub>1_Val_PD: +assumes eval: "P \\<^sub>1 \RI(C,e\<^sub>0);Cs \ unit,s\ \ \Val v,s'\" + and iconf: "iconf (shp\<^sub>1 s) (RI(C,e\<^sub>0);Cs \ unit)" and last: "last(C#Cs) = C'" +shows "\sfs i. shp\<^sub>1 s' C' = \(sfs,i)\ \ (i = Done \ i = Processing)" +proof(cases Cs) + case Nil + then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] last by simp +next + case (Cons a list) + then have nNil: "Cs \ []" by simp + then have "\Cs'. Cs = Cs' @ [C']" using last append_butlast_last_id[OF nNil] + by(rule_tac x="butlast Cs" in exI) simp + then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] by simp +qed -lemma rinit\<^sub>1_throw_PD: "P \\<^sub>1 \RI(C,e\<^sub>0);Cs \ unit,s\ \ \throw a,s'\ - \ iconf (shp\<^sub>1 s) (RI(C,e\<^sub>0);Cs \ unit) \ last(C#Cs) = C' - \ \sfs i. shp\<^sub>1 s' C' = \(sfs,Error)\" -apply(drule_tac C' = C' and a = a in eval\<^sub>1_init_return, simp_all) -apply (metis append_butlast_last_id) -done +lemma rinit\<^sub>1_throw_PD: +assumes eval: "P \\<^sub>1 \RI(C,e\<^sub>0);Cs \ unit,s\ \ \throw a,s'\" + and iconf: "iconf (shp\<^sub>1 s) (RI(C,e\<^sub>0);Cs \ unit)" and last: "last(C#Cs) = C'" +shows "\sfs. shp\<^sub>1 s' C' = \(sfs,Error)\" +proof(cases Cs) + case Nil + then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] last by simp +next + case (Cons a list) + then have nNil: "Cs \ []" by simp + then have "\Cs'. Cs = Cs' @ [C']" using last append_butlast_last_id[OF nNil] + by(rule_tac x="butlast Cs" in exI) simp + then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] by simp +qed subsubsection "The proof" lemma fixes P\<^sub>1 defines [simp]: "P \ compP\<^sub>2 P\<^sub>1" assumes wf: "wf_J\<^sub>1_prog P\<^sub>1" shows Jcc: "P\<^sub>1 \\<^sub>1 \e,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\ \ \ef,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\ \ (\E C M pc ics v xa vs frs I. \ Jcc_cond P\<^sub>1 E C M vs pc ics I h\<^sub>0 sh\<^sub>0 e \ \ (ef = Val v \ P \ (None,h\<^sub>0,Jcc_frames P C M vs ls\<^sub>0 pc ics frs e,sh\<^sub>0) -jvm\ Jcc_rhs P\<^sub>1 E C M vs ls\<^sub>0 pc ics frs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v e) \ (ef = Throw xa \ Jcc_err P C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 xa e) )" (*<*) (is "_ \ (\E C M pc ics v xa vs frs I. PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 ef h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I)") (*>*) and "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\ [\] \fs,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\ \ (\C M pc ics ws xa es' vs frs I. \ P,C,M,pc \ compEs\<^sub>2 es; P,C,M \ compxEs\<^sub>2 es pc (size vs)/I,size vs; {pc..2 es)} \ I; ics = No_ics; \sub_RIs es \ \ (fs = map Val ws \ P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(rev ws @ vs,ls\<^sub>1,C,M,pc+size(compEs\<^sub>2 es),ics)#frs,sh\<^sub>1)) \ (fs = map Val ws @ Throw xa # es' \ (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compEs\<^sub>2 es) \ \ caught P pc\<^sub>1 h\<^sub>1 xa (compxEs\<^sub>2 es pc (size vs)) \ (\vs'. P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ handle P C M xa h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1))))" (*<*) (is "_ \ (\C M pc ics ws xa es' vs frs I. PROP ?Ps es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 fs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C M pc ics ws xa es' vs frs I)") proof (induct rule:eval\<^sub>1_evals\<^sub>1_inducts) case New\<^sub>1 thus ?case by auto next case (NewFail\<^sub>1 sh C' sfs h ls) let ?xa = "addr_of_sys_xcpt OutOfMemory" have "P \ (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\ handle P C M ?xa h vs ls pc ics frs sh" using NewFail\<^sub>1 by(clarsimp simp: handle_def) then show ?case by(auto intro!: exI[where x="[]"]) next case (NewInit\<^sub>1 sh C' h ls v' h' ls' sh' a FDTs h'') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (new C') = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (new C')),ics)#frs,sh'), err)" using NewInit\<^sub>1.prems(1) by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' ([C'],False) \ unit))" using has_fields_is_class[OF NewInit\<^sub>1.hyps(5)] by auto then obtain err' where pcs': "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v' xa (INIT C' ([C'],False) \ unit) = (True, (vs,ls,C,M,pc,Calling C' []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), err')" using NewInit\<^sub>1.prems(1) by auto have IH: "PROP ?P (INIT C' ([C'],False) \ unit) h ls sh (Val v') h' ls' sh' E C M pc ics v' xa vs frs I" by fact have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF NewInit\<^sub>1.hyps(2)]) obtain sfs i where sh': "sh' C' = Some(sfs,i)" using init\<^sub>1_Val_PD[OF NewInit\<^sub>1.hyps(2)] by clarsimp have "P \ (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\ (None,h,(vs,ls,C,M,pc,Calling C' [])#frs,sh)" proof(cases "sh C'") case None then show ?thesis using NewInit\<^sub>1.prems by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using NewInit\<^sub>1.hyps(1) NewInit\<^sub>1.prems Some by(cases ics; case_tac i) auto qed also have "P \ \ -jvm\ (None, h', (vs, ls, C, M, pc, Called []) # frs, sh')" using IH pcs' by auto also have "P \ \ -jvm\ (None, h'', (Addr a#vs, ls, C, M, Suc pc, ics) # frs, sh')" using NewInit\<^sub>1.hyps(1,2,4-6) NewInit\<^sub>1.prems sh' by(cases ics) auto finally show ?case using pcs ls by clarsimp next case (NewInitOOM\<^sub>1 sh C' h ls v' h' ls' sh') let ?xa = "addr_of_sys_xcpt OutOfMemory" obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (new C') = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (new C')),ics)#frs,sh'), err)" using NewInitOOM\<^sub>1.prems(1) by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' ([C'],False) \ unit))" using NewInitOOM\<^sub>1.hyps(5) by auto then obtain err' where pcs': "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v' xa (INIT C' ([C'],False) \ unit) = (True, (vs,ls,C,M,pc,Calling C' []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), err')" using NewInitOOM\<^sub>1.prems(1) by auto have IH: "PROP ?P (INIT C' ([C'],False) \ unit) h ls sh (Val v') h' ls' sh' E C M pc ics v' xa vs frs I" by fact have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF NewInitOOM\<^sub>1.hyps(2)]) have "iconf (shp\<^sub>1 (h, ls, sh)) (INIT C' ([C'],False) \ unit)" by simp then obtain sfs i where sh': "sh' C' = Some(sfs,i)" using init\<^sub>1_Val_PD[OF NewInitOOM\<^sub>1.hyps(2)] by clarsimp have "P \ (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\ (None,h,(vs,ls,C,M,pc,Calling C' [])#frs,sh)" proof(cases "sh C'") case None then show ?thesis using NewInitOOM\<^sub>1.prems by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using NewInitOOM\<^sub>1.hyps(1) NewInitOOM\<^sub>1.prems Some by(cases ics; case_tac i) auto qed also have "P \ \ -jvm\ (None, h', (vs, ls, C, M, pc, Called []) # frs, sh')" using IH pcs' by auto also have "P \ \ -jvm\ handle P C M ?xa h' vs ls pc ics frs sh'" using NewInitOOM\<^sub>1.hyps(1,2,4,5) NewInitOOM\<^sub>1.prems sh' by(auto simp: handle_def) finally show ?case using pcs ls by(simp, metis (no_types) append_Nil le_refl lessI) next case (NewInitThrow\<^sub>1 sh C' h ls a h' ls' sh') obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (new C') = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (new C')),ics)#frs,sh'), err)" using NewInitThrow\<^sub>1.prems(1) by clarsimp obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF NewInitThrow\<^sub>1.hyps(2)] by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' ([C'],False) \ unit))" using NewInitThrow\<^sub>1.hyps(4) by auto then obtain vs' where pcs': "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v a' (INIT C' ([C'],False) \ unit) = (True, (vs,ls,C,M,pc,Calling C' []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), P \ (None,h,(vs,ls,C,M,pc,Calling C' []) # frs,sh) -jvm\ handle P C M a' h' (vs'@vs) ls pc ics frs sh')" using NewInitThrow\<^sub>1.prems(1) by simp blast have IH: "PROP ?P (INIT C' ([C'],False) \ unit) h ls sh (throw a) h' ls' sh' E C M pc ics v a' vs frs I" by fact have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF NewInitThrow\<^sub>1.hyps(2)]) then have "P \ (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\ (None,h,(vs,ls,C,M,pc,Calling C' []) # frs,sh)" proof(cases "sh C'") case None then show ?thesis using NewInitThrow\<^sub>1.prems by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using NewInitThrow\<^sub>1.hyps(1) NewInitThrow\<^sub>1.prems Some by(cases ics; case_tac i) auto qed also have "P \ \ -jvm\ handle P C M a' h' (vs'@vs) ls pc ics frs sh'" using IH pcs' throw by auto finally show ?case using throw ls by auto next case (Cast\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs C') let ?pc = "pc + length(compE\<^sub>2 e)" obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)" using Cast\<^sub>1.prems(1) by auto have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact then have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using Jcc_pieces_Cast[OF assms(1) pcs, of "Addr a"] Cast\<^sub>1.prems pcs by auto also have "P \ \ -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)" using Cast\<^sub>1 by (auto simp add:cast_ok_def) finally show ?case by auto next case (CastNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C') let ?pc = "pc + length(compE\<^sub>2 e)" obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)" using CastNull\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs I" by fact then have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using Jcc_pieces_Cast[OF assms(1) pcs, of Null] CastNull\<^sub>1.prems pcs by auto also have "P \ \ -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)" using CastNull\<^sub>1 by (auto simp add:cast_ok_def) finally show ?case by auto next case (CastFail\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs C') let ?pc = "pc + length(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt ClassCast" obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)" using CastFail\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact then have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using Jcc_pieces_Cast[OF assms(1) pcs, of "Addr a"] CastFail\<^sub>1.prems pcs by auto also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using CastFail\<^sub>1 by (auto simp add:handle_def cast_ok_def) finally have exec: "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ \". show ?case (is "?N \ (?eq \ ?err)") proof show ?N by simp next { assume ?eq then have ?err using exec by (auto intro!: exI[where x="?pc"] exI[where x="[Addr a]"]) } thus "?eq \ ?err" by simp qed next case (CastThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C') obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)" using CastThrow\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact show ?case using IH Jcc_pieces_Cast[OF assms(1) pcs, of v] CastThrow\<^sub>1.prems pcs less_SucI by(simp, blast) next case Val\<^sub>1 thus ?case by auto next case Var\<^sub>1 thus ?case by auto next case (BinOp\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 v\<^sub>2 h\<^sub>2 ls\<^sub>2 sh\<^sub>2 bop w) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1 \bop\ e\<^sub>2) = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1 \bop\ e\<^sub>2)),ics)#frs,sh\<^sub>2), err)" using BinOp\<^sub>1.prems(1) by clarsimp let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val v\<^sub>1) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v\<^sub>1 xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v\<^sub>2) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v\<^sub>2 xa (v\<^sub>1#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using IH\<^sub>1 Jcc_pieces_BinOp1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v\<^sub>1] by simp also have "P \ \ -jvm\ (None,h\<^sub>2,(v\<^sub>2#v\<^sub>1#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using IH\<^sub>2 Jcc_pieces_BinOp2[OF assms(1) pcs, of h\<^sub>1 v\<^sub>1 ls\<^sub>1 sh\<^sub>1 v\<^sub>2] by (simp add: add.assoc) also have "P \ \ -jvm\ (None,h\<^sub>2,(w#vs,ls\<^sub>2,C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2)" using BinOp\<^sub>1 by(cases bop) auto finally show ?case using pcs by (auto split: bop.splits simp:add.assoc) next case (BinOpThrow\<^sub>1\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e h\<^sub>1 ls\<^sub>1 sh\<^sub>1 bop e\<^sub>2) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<^sub>1 \bop\ e\<^sub>2) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<^sub>1 \bop\ e\<^sub>2)),ics)#frs,sh\<^sub>1), err)" using BinOpThrow\<^sub>1\<^sub>1.prems(1) by clarsimp have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact show ?case using IH\<^sub>1 Jcc_pieces_BinOp1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v] BinOpThrow\<^sub>1\<^sub>1.prems nsub_RI_Jcc_pieces by auto next case (BinOpThrow\<^sub>2\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e h\<^sub>2 ls\<^sub>2 sh\<^sub>2 bop) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1 \bop\ e\<^sub>2) = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1 \bop\ e\<^sub>2)),ics)#frs,sh\<^sub>2), err)" using BinOpThrow\<^sub>2\<^sub>1.prems(1) by clarsimp let ?pc = "pc + length(compE\<^sub>2 e\<^sub>1)" have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val v\<^sub>1) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v\<^sub>1 xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc ics v xa (v\<^sub>1#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact let ?\\<^sub>1 = "(None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" have 1: "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ ?\\<^sub>1" using IH\<^sub>1 Jcc_pieces_BinOp1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v\<^sub>1] by simp have "(throw e = Val v \ P \ (None, h\<^sub>0, Jcc_frames P C M vs ls\<^sub>0 pc ics frs (e\<^sub>1 \bop\ e\<^sub>2), sh\<^sub>0) -jvm\ Jcc_rhs P\<^sub>1 E C M vs ls\<^sub>0 pc ics frs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v (e\<^sub>1 \bop\ e\<^sub>2)) \ (throw e = Throw xa \ (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compE\<^sub>2 (e\<^sub>1 \bop\ e\<^sub>2)) \ \ caught P pc\<^sub>1 h\<^sub>2 xa (compxE\<^sub>2 (e\<^sub>1 \bop\ e\<^sub>2) pc (size vs)) \ (\vs'. P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>1 ics frs sh\<^sub>2)))" (is "?N \ (?eq \ (\pc\<^sub>2. ?H pc\<^sub>2))") proof show ?N by simp next { assume ?eq then obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc \ pc\<^sub>2 \ pc\<^sub>2 < ?pc + size(compE\<^sub>2 e\<^sub>2) \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc (size vs + 1))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@v\<^sub>1#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using IH\<^sub>2 Jcc_pieces_BinOp2[OF assms(1) pcs, of h\<^sub>1 v\<^sub>1 ls\<^sub>1 sh\<^sub>1 v] BinOpThrow\<^sub>2\<^sub>1.prems by clarsimp then have "?H pc\<^sub>2" using jvm_trans[OF 1 2] by(auto intro!: exI[where x="vs'@[v\<^sub>1]"]) hence "\pc\<^sub>2. ?H pc\<^sub>2" by iprover } thus "?eq \ (\pc\<^sub>2. ?H pc\<^sub>2)" by iprover qed then show ?case using pcs by simp blast next case (FAcc\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' fs F T D w) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\F{D}) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\F{D})),ics)#frs,sh\<^sub>1), err)" using FAcc\<^sub>1.prems(1) by clarsimp have "P\<^sub>1 \ D sees F,NonStatic:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAcc\<^sub>1.hyps(4)]]) then have field: "field P D F = (D,NonStatic,T)" by simp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact let ?pc = "pc + length(compE\<^sub>2 e)" have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_FAcc[OF pcs, of "Addr a"] pcs by simp also have "P \ \ -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)" using FAcc\<^sub>1 field by auto finally have "P \ (None, h\<^sub>0, frs', sh\<^sub>0) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)" by auto then show ?case using pcs by auto next case (FAccNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 F D) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\F{D}) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\F{D})),ics)#frs,sh\<^sub>1), err)" using FAccNull\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs I" by fact let ?pc = "pc + length(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt NullPointer" have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_FAcc[OF pcs, of Null] by simp also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using FAccNull\<^sub>1.prems by(fastforce simp:split_beta handle_def simp del: split_paired_Ex) finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[Null]"]) next case (FAccThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 F D) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\F{D}) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\F{D})),ics)#frs,sh\<^sub>1), err)" using FAccThrow\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact show ?case using IH Jcc_pieces_FAcc[OF pcs, of v] FAccThrow\<^sub>1.prems nsub_RI_Jcc_pieces less_Suc_eq by auto next case (FAccNone\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C fs F D) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\F{D}) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\F{D})),ics)#frs,sh\<^sub>1), err)" using FAccNone\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact let ?pc = "pc + length(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt NoSuchFieldError" have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_FAcc[OF pcs, of "Addr a"] by simp also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using FAccNone\<^sub>1 by(cases ics; clarsimp simp:split_beta handle_def simp del: split_paired_Ex) finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[Addr a]"]) next case (FAccStatic\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' fs F T D) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\F{D}) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\F{D})),ics)#frs,sh\<^sub>1), err)" using FAccStatic\<^sub>1.prems(1) by clarsimp have "P\<^sub>1 \ D sees F,Static:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAccStatic\<^sub>1.hyps(4)]]) then have field: "field P D F = (D,Static,T)" by simp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact let ?pc = "pc + length(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError" have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_FAcc[OF pcs, of "Addr a"] by simp also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using FAccStatic\<^sub>1 field by(fastforce simp:split_beta handle_def simp del: split_paired_Ex) finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[Addr a]"]) next case (SFAcc\<^sub>1 C' F t D sh sfs v' h ls) have has: "P\<^sub>1 \ D has F,Static:t in D" by(rule has_field_idemp[OF SFAcc\<^sub>1.hyps(1)]) have "P\<^sub>1 \ D sees F,Static:t in D" by(rule has_field_sees[OF has]) then have field: "field P D F = (D,Static,t)" by simp then have "P \ (None,h,Jcc_frames P C M vs ls pc ics frs (C'\\<^sub>sF{D}),sh) -jvm\ (None,h,(v'#vs,ls,C,M,Suc pc,ics)#frs,sh)" using SFAcc\<^sub>1 has by(cases ics) auto then show ?case by clarsimp next case (SFAccInit\<^sub>1 C' F t D sh h ls v' h' ls' sh' sfs i v'') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C'\\<^sub>sF{D}) = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D})),ics)#frs,sh'), err)" using SFAccInit\<^sub>1.prems(1) by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT D ([D],False) \ unit))" using has_field_is_class'[OF SFAccInit\<^sub>1.hyps(1)] by auto then obtain err' where pcs': "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v' xa (INIT D ([D],False) \ unit) = (True, (vs,ls,C,M,pc,Calling D []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), err')" using SFAccInit\<^sub>1.prems(1) by auto have IH: "PROP ?P (INIT D ([D],False) \ unit) h ls sh (Val v') h' ls' sh' E C M pc ics v' xa vs frs I" by fact have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF SFAccInit\<^sub>1.hyps(3)]) have has: "P\<^sub>1 \ D has F,Static:t in D" by(rule has_field_idemp[OF SFAccInit\<^sub>1.hyps(1)]) have "P\<^sub>1 \ D sees F,Static:t in D" by(rule has_field_sees[OF has]) then have field: "field P D F = (D,Static,t)" by simp have "P \ (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\ (None,h,(vs,ls,C,M,pc,Calling D [])#frs,sh)" proof(cases "sh D") case None then show ?thesis using SFAccInit\<^sub>1.hyps(1,2,5,6) SFAccInit\<^sub>1.prems field by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using SFAccInit\<^sub>1.hyps(1,2,5,6) SFAccInit\<^sub>1.prems field Some by(cases ics; case_tac i) auto qed also have "P \ ... -jvm\ (None, h', (vs, ls, C, M, pc, Called []) # frs, sh')" using IH pcs' by auto also have "P \ ... -jvm\ (None, h', (v''#vs, ls, C, M, Suc pc, ics) # frs, sh')" using SFAccInit\<^sub>1.hyps(1,2,5,6) SFAccInit\<^sub>1.prems has field by(cases ics) auto finally show ?case using pcs ls by clarsimp next case (SFAccInitThrow\<^sub>1 C' F t D sh h ls a h' ls' sh') obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C'\\<^sub>sF{D}) = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D})),ics)#frs,sh'), err)" using SFAccInitThrow\<^sub>1.prems(1) by clarsimp obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF SFAccInitThrow\<^sub>1.hyps(3)] by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT D ([D],False) \ unit))" using has_field_is_class'[OF SFAccInitThrow\<^sub>1.hyps(1)] by auto then obtain vs' where pcs': "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v a' (INIT D ([D],False) \ unit) = (True, (vs,ls,C,M,pc,Calling D []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), P \ (None,h,(vs,ls,C,M,pc,Calling D []) # frs,sh) -jvm\ handle P C M a' h' (vs'@vs) ls pc ics frs sh')" using SFAccInitThrow\<^sub>1.prems(1) by simp blast have IH: "PROP ?P (INIT D ([D],False) \ unit) h ls sh (throw a) h' ls' sh' E C M pc ics v a' vs frs I" by fact have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF SFAccInitThrow\<^sub>1.hyps(3)]) have has: "P\<^sub>1 \ D has F,Static:t in D" by(rule has_field_idemp[OF SFAccInitThrow\<^sub>1.hyps(1)]) have "P\<^sub>1 \ D sees F,Static:t in D" by(rule has_field_sees[OF has]) then have field: "field P D F = (D,Static,t)" by simp then have "P \ (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\ (None,h,(vs,ls,C,M,pc,Calling D []) # frs,sh)" proof(cases "sh D") case None then show ?thesis using SFAccInitThrow\<^sub>1.hyps(1,2) SFAccInitThrow\<^sub>1.prems field by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using SFAccInitThrow\<^sub>1.hyps(1,2) SFAccInitThrow\<^sub>1.prems field Some by(cases ics; case_tac i) auto qed also have "P \ \ -jvm\ handle P C M a' h' (vs'@vs) ls pc ics frs sh'" using IH pcs' throw by auto finally show ?case using throw ls by auto next case (SFAccNone\<^sub>1 C' F D h\<^sub>1 ls\<^sub>1 sh\<^sub>1) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\\<^sub>sF{D}) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D})),ics)#frs,sh\<^sub>1), err)" by clarsimp let ?xa = "addr_of_sys_xcpt NoSuchFieldError" have "P \ (None,h\<^sub>1,frs',sh\<^sub>1) -jvm\ handle P C M ?xa h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1" using SFAccNone\<^sub>1 pcs by(cases ics; clarsimp simp:split_beta handle_def simp del: split_paired_Ex) then show ?case using pcs by(auto intro!: exI[where x = pc] exI[where x="[]"]) next case (SFAccNonStatic\<^sub>1 C' F t D h\<^sub>1 ls\<^sub>1 sh\<^sub>1) let ?frs' = "(vs, ls\<^sub>1, C, M, pc, ics) # frs" let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError" have "P\<^sub>1 \ D sees F,NonStatic:t in D" by(rule has_field_sees[OF has_field_idemp[OF SFAccNonStatic\<^sub>1.hyps(1)]]) then have field: "field P D F = (D,NonStatic,t)" by simp have "P \ (None,h\<^sub>1,?frs',sh\<^sub>1) -jvm\ handle P C M ?xa h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1" using SFAccNonStatic\<^sub>1 proof(cases ics) case No_ics then show ?thesis using SFAccNonStatic\<^sub>1 field by (auto simp:split_beta handle_def simp del: split_paired_Ex) qed(simp_all) then show ?case by (auto intro!: exI[where x = pc] exI[where x="[]"]) next case (LAss\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 i ls\<^sub>2) let ?pc = "pc + length(compE\<^sub>2 e)" obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (i:=e) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (i:=e)),ics)#frs,sh\<^sub>1), err)" using LAss\<^sub>1.prems(1) by auto have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact then have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using Jcc_pieces_LAss[OF assms(1) pcs, of w] LAss\<^sub>1.prems pcs by auto also have "P \ \ -jvm\ (None,h\<^sub>1,(Unit#vs,ls\<^sub>2,C,M,?pc+2,ics)#frs,sh\<^sub>1)" using LAss\<^sub>1 by (auto simp add:cast_ok_def) finally show ?case by auto next case (LAssThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 i) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (i:=e) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (i:=e)),ics)#frs,sh\<^sub>1), err)" using LAssThrow\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact show ?case using IH Jcc_pieces_LAss[OF assms(1) pcs, of v] LAssThrow\<^sub>1.prems pcs less_SucI by(simp, blast) next case (FAss\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs F T D fs' h\<^sub>2') obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\F{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)" using FAss\<^sub>1.prems(1) by clarsimp have "P\<^sub>1 \ D sees F,NonStatic:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAss\<^sub>1.hyps(6)]]) then have field: "field P D F = (D,NonStatic,T)" by simp let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 "Addr a"] by simp also have "P \ \ -jvm\ (None,h\<^sub>2,(w#Addr a#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 "Addr a" ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc) also have "P \ \ -jvm\ (None,h\<^sub>2',(Unit#vs,ls\<^sub>2,C,M,?pc\<^sub>2+2,ics)#frs,sh\<^sub>2)" using FAss\<^sub>1 field by auto finally show ?case using pcs by (auto simp:add.assoc) next case (FAssNull\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 F D) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\F{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)" using FAssNull\<^sub>1.prems(1) by clarsimp let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" let ?xa = "addr_of_sys_xcpt NullPointer" have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Null#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 Null] by simp also have "P \ \ -jvm\ (None,h\<^sub>2,(w#Null#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 Null ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc) also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (w#Null#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using FAssNull\<^sub>1 by(fastforce simp:split_beta handle_def simp del: split_paired_Ex) finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="w#[Null]"]) next case (FAssThrow\<^sub>2\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 F D) let ?frs' = "(vs, ls\<^sub>0, C, M, pc, ics) # frs" obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\F{D} := e\<^sub>2) = (True, ?frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)" using FAssThrow\<^sub>2\<^sub>1.prems(1) by clarsimp let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?\\<^sub>1 = "(None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v xa (w#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have 1: "P \ (None,h\<^sub>0,?frs',sh\<^sub>0) -jvm\ ?\\<^sub>1" using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 w] by simp show ?case (is "?N \ (?eq \ ?err)") proof show ?N by simp next { assume ?eq moreover have "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v xa (w#vs) frs (I - pcs (compxE\<^sub>2 e\<^sub>1 pc (length vs)))" by fact ultimately obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc\<^sub>1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1 + size(compE\<^sub>2 e\<^sub>2) \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1 (size vs + 1))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@w#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using FAssThrow\<^sub>2\<^sub>1.prems Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 w ls\<^sub>1 sh\<^sub>1] by auto have ?err using Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 w ls\<^sub>1 sh\<^sub>1] pc\<^sub>2 jvm_trans[OF 1 2] by(auto intro!: exI[where x=pc\<^sub>2] exI[where x="vs'@[w]"]) } thus "?eq \ ?err" by simp qed next case (FAssThrow\<^sub>1\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 F D e\<^sub>2) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<^sub>1\F{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)" using FAssThrow\<^sub>1\<^sub>1.prems(1) by clarsimp have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact show ?case using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v] FAssThrow\<^sub>1\<^sub>1.prems nsub_RI_Jcc_pieces by auto next case (FAssNone\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs F D) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\F{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)" using FAssNone\<^sub>1.prems(1) by clarsimp let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" let ?xa = "addr_of_sys_xcpt NoSuchFieldError" have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 "Addr a"] by simp also have "P \ \ -jvm\ (None,h\<^sub>2,(w#Addr a#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 "Addr a" ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc) also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (w#Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using FAssNone\<^sub>1 by(fastforce simp:split_beta handle_def simp del: split_paired_Ex) finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="w#[Addr a]"]) next case (FAssStatic\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs F T D) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\F{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)" using FAssStatic\<^sub>1.prems(1) by clarsimp have "P\<^sub>1 \ D sees F,Static:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAssStatic\<^sub>1.hyps(6)]]) then have field: "field P D F = (D,Static,T)" by simp let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError" have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 "Addr a"] by simp also have "P \ \ -jvm\ (None,h\<^sub>2,(w#Addr a#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 "Addr a" ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc) also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (w#Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using FAssStatic\<^sub>1 field by(fastforce simp:split_beta handle_def simp del: split_paired_Ex) finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="w#[Addr a]"]) next case (SFAss\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F T D sfs sfs' sh\<^sub>1') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\\<^sub>sF{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)" using SFAss\<^sub>1.prems(1) by clarsimp have "P\<^sub>1 \ D sees F,Static:T in D" by(rule has_field_sees[OF has_field_idemp[OF SFAss\<^sub>1.hyps(3)]]) then have field: "field P D F = (D,Static,T)" by simp have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)" have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp also have "P \ \ -jvm\ (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1')" using SFAss\<^sub>1.hyps(3-6) SFAss\<^sub>1.prems(1) field by auto also have "P \ ... -jvm\ (None,h\<^sub>1,(Unit#vs,ls\<^sub>1,C,M,?pc+2,ics)#frs,sh\<^sub>1')" using SFAss\<^sub>1 by auto finally show ?case using pcs by auto next case (SFAssInit\<^sub>1 e\<^sub>2 h ls sh w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F t D v' h' ls' sh' sfs i sfs' sh'') let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)" obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh'' v xa (C'\\<^sub>sF{D}:=e\<^sub>2) = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D}:=e\<^sub>2)),ics)#frs,sh''), err)" using SFAssInit\<^sub>1.prems(1) by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \ unit))" using has_field_is_class'[OF SFAssInit\<^sub>1.hyps(3)] by auto then obtain err' where pcs': "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1 I h' ls' sh' v' xa (INIT D ([D],False) \ unit) = (True, (w#vs,ls\<^sub>1,C,M,?pc,Calling D []) # frs, (None,h',(w#vs,ls\<^sub>1,C,M,?pc,Called [])#frs,sh'), err')" using SFAssInit\<^sub>1.prems(1) by simp have ls: "ls\<^sub>1 = ls'" by(rule init\<^sub>1_same_loc[OF SFAssInit\<^sub>1.hyps(5)]) have has: "P\<^sub>1 \ D has F,Static:t in D" by(rule has_field_idemp[OF SFAssInit\<^sub>1.hyps(3)]) have "P\<^sub>1 \ D sees F,Static:t in D" by(rule has_field_sees[OF has]) then have field: "field P D F = (D,Static,t)" by simp have IH: "PROP ?P e\<^sub>2 h ls sh (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact have IHI: "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v') h' ls' sh' E C M ?pc ics v' xa (w#vs) frs I" by fact have "P \ (None,h,frs',sh) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_SFAss[OF pcs, where v'=w] by simp also have "P \ \ -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,Calling D [])#frs,sh\<^sub>1)" proof(cases "sh\<^sub>1 D") case None then show ?thesis using None SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems field by(cases ics, auto) next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems field Some by(cases ics; case_tac i) auto qed also have "P \ ... -jvm\ (None, h', (w#vs, ls\<^sub>1, C, M, ?pc, Called []) # frs, sh')" using IHI pcs' by clarsimp also have "P \ ... -jvm\ (None, h', (vs, ls\<^sub>1, C, M, ?pc + 1, ics) # frs, sh'')" using SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems has field by(cases ics) auto also have "P \ ... -jvm\ (None, h', (Unit#vs, ls\<^sub>1, C, M, ?pc + 2, ics) # frs, sh'')" using SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems has field by(cases ics) auto finally show ?case using pcs ls by simp blast next case (SFAssInitThrow\<^sub>1 e\<^sub>2 h ls sh w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F t D a h' ls' sh') let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)" obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C'\\<^sub>sF{D}:=e\<^sub>2) = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D}:=e\<^sub>2)),ics)#frs,sh'), err)" using SFAssInitThrow\<^sub>1.prems(1) by clarsimp obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF SFAssInitThrow\<^sub>1.hyps(5)] by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \ unit))" using has_field_is_class'[OF SFAssInitThrow\<^sub>1.hyps(3)] by auto then obtain vs' where pcs': "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1 I h' ls' sh' v a' (INIT D ([D],False) \ unit) = (True, (w#vs,ls\<^sub>1,C,M,?pc,Calling D []) # frs, (None,h',(w#vs,ls\<^sub>1,C,M,?pc,Called [])#frs,sh'), P \ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,Calling D []) # frs,sh\<^sub>1) -jvm\ handle P C M a' h' (vs'@w#vs) ls\<^sub>1 ?pc ics frs sh')" using SFAssInitThrow\<^sub>1.prems(1) by simp blast have ls: "ls\<^sub>1 = ls'" by(rule init\<^sub>1_same_loc[OF SFAssInitThrow\<^sub>1.hyps(5)]) have has: "P\<^sub>1 \ D has F,Static:t in D" by(rule has_field_idemp[OF SFAssInitThrow\<^sub>1.hyps(3)]) have "P\<^sub>1 \ D sees F,Static:t in D" by(rule has_field_sees[OF has]) then have field: "field P D F = (D,Static,t)" by simp have IH: "PROP ?P e\<^sub>2 h ls sh (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact have IHI: "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw a) h' ls' sh' E C M ?pc ics v a' (w#vs) frs I" by fact have "P \ (None,h,(vs, ls, C, M, pc, ics) # frs,sh) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp blast also have "P \ \ -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,Calling D [])#frs,sh\<^sub>1)" proof(cases "sh\<^sub>1 D") case None then show ?thesis using SFAssInitThrow\<^sub>1.hyps(1,3,4,5) SFAssInitThrow\<^sub>1.prems field by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using SFAssInitThrow\<^sub>1.hyps(1,3,4,5) SFAssInitThrow\<^sub>1.prems field Some by(cases ics; case_tac i) auto qed also have "P \ ... -jvm\ handle P C M a' h' (vs'@w#vs) ls\<^sub>1 ?pc ics frs sh'" using IHI pcs' throw by auto finally show ?case using throw ls by(auto intro!: exI[where x = ?pc] exI[where x="vs'@[w]"]) next case (SFAssThrow\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F D) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\\<^sub>sF{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)" using SFAssThrow\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact show ?case using IH Jcc_pieces_SFAss[OF pcs, where v'=v] SFAssThrow\<^sub>1.prems nsub_RI_Jcc_pieces less_Suc_eq by auto next case (SFAssNone\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F D) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\\<^sub>sF{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)" using SFAssNone\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)" let ?xa = "addr_of_sys_xcpt NoSuchFieldError" have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using SFAssNone\<^sub>1 by(cases ics; clarsimp simp add: handle_def) finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[w]"]) next case (SFAssNonStatic\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F T D) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\\<^sub>sF{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)" using SFAssNonStatic\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)" let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError" have "P\<^sub>1 \ D sees F,NonStatic:T in D" by(rule has_field_sees[OF has_field_idemp[OF SFAssNonStatic\<^sub>1.hyps(3)]]) then have field: "field P D F = (D,NonStatic,T)" by simp have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using SFAssNonStatic\<^sub>1 proof(cases ics) case No_ics then show ?thesis using SFAssNonStatic\<^sub>1 field by (auto simp:split_beta handle_def simp del: split_paired_Ex) qed(simp_all) finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[w]"]) next case (Call\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 Ca fs M' Ts T body D ls\<^sub>2' f h\<^sub>3 ls\<^sub>3 sh\<^sub>3) let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs" let ?\\<^sub>0 = "(None,h\<^sub>0,?frs\<^sub>0,sh\<^sub>0)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>1 = "(None,h\<^sub>1,(Addr a#vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" let ?frs\<^sub>2 = "(rev pvs @ Addr a # vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" let ?frs\<^sub>2' = "([], ls\<^sub>2', D,M',0,No_ics) # ?frs\<^sub>2" let ?\\<^sub>2' = "(None, h\<^sub>2, ?frs\<^sub>2', sh\<^sub>2)" have nclinit: "M' \ clinit" using wf_sees_clinit1[OF wf] visible_method_exists[OF Call\<^sub>1.hyps(6)] sees_method_idemp[OF Call\<^sub>1.hyps(6)] by fastforce have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) have invoke: "P,C,M,?pc\<^sub>2 \ Invoke M' (length Ts)" using Call\<^sub>1.hyps(7) Call\<^sub>1.prems(1) by clarsimp have nsub: "\ sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf Call\<^sub>1.hyps(6)]) obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa (e\M'(es)) = (True, ?frs\<^sub>0, (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>3), err)" using Call\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa (map Val pvs) (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact have "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp also have "P \ \ -jvm\ ?\\<^sub>2" using IH_es Call\<^sub>1.prems by fastforce also have "P \ \ -jvm\ ?\\<^sub>2'" using jvm_Invoke[OF assms(1) invoke _ Call\<^sub>1.hyps(6-8)] Call\<^sub>1.hyps(5) Call\<^sub>1.prems(1) by simp finally have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>2'". have "P\<^sub>1 \ Ca sees M',NonStatic: Ts\T = body in D" by fact then have M'_in_D: "P\<^sub>1 \ D sees M',NonStatic: Ts\T = body in D" by(rule sees_method_idemp) have M'_code: "compP\<^sub>2 P\<^sub>1,D,M',0 \ compE\<^sub>2 body @ [Return]" using beforeM M'_in_D by simp have M'_xtab: "compP\<^sub>2 P\<^sub>1,D,M' \ compxE\<^sub>2 body 0 0/{..2 body)},0" using M'_in_D by(rule beforexM) have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 f h\<^sub>3 ls\<^sub>3 sh\<^sub>3 (Class D # Ts) D M' 0 No_ics v xa [] ?frs\<^sub>2 ({..2 body)})" by fact have cond: "Jcc_cond P\<^sub>1 (Class D # Ts) D M' [] 0 No_ics {..2 body)} h\<^sub>2 sh\<^sub>2 body" using nsub_RI_Jcc_pieces[OF assms(1) nsub] M'_code M'_xtab by clarsimp show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>2' -jvm\ (None,h\<^sub>3,([v],ls\<^sub>3,D,M',size(compE\<^sub>2 body),No_ics)#?frs\<^sub>2,sh\<^sub>3)" using val IH_body Call\<^sub>1.prems M'_code cond nsub_RI_Jcc_pieces nsub by auto also have "P \ \ -jvm\ (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>3)" using Call\<^sub>1.hyps(7) M'_code M'_in_D nclinit by(cases T, auto) finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw with IH_body obtain pc\<^sub>2 vs' where pc\<^sub>2: "0 \ pc\<^sub>2 \ pc\<^sub>2 < size(compE\<^sub>2 body) \ \ caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and 2: "P \ ?\\<^sub>2' -jvm\ handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3" using Call\<^sub>1.prems M'_code M'_xtab cond nsub_RI_Jcc_pieces nsub by (auto simp del:split_paired_Ex) have "handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3 = handle P C M xa h\<^sub>3 (rev pvs @ Addr a # vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>3" using pc\<^sub>2 M'_in_D nclinit by(auto simp add:handle_def) then show "?err" using pc\<^sub>2 jvm_trans[OF 1 2] by(auto intro!:exI[where x="?pc\<^sub>2"] exI[where x="rev pvs@[Addr a]"]) qed qed next case (CallParamsThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 pvs ex es'' M') let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc,ics)#frs,sh\<^sub>0)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>1 = "(None,h\<^sub>1,(w # vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\M'(es)) = (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)" using CallParamsThrow\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp have Isubs: "{?pc\<^sub>1..2} \ I - pcs (compxE\<^sub>2 e pc (length vs))" using CallParamsThrow\<^sub>1.prems by clarsimp show ?case (is "?N \ (?eq \ ?err)") proof show ?N by simp next { assume ?eq moreover have "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa es'' (w#vs) frs (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact ultimately obtain vs' where "\pc\<^sub>2. (?pc\<^sub>1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1 + size(compEs\<^sub>2 es) \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es ?pc\<^sub>1 (size vs + 1))) \ P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@w#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" (is "\pc\<^sub>2. ?PC pc\<^sub>2 \ ?Exec pc\<^sub>2") using CallParamsThrow\<^sub>1 Isubs by auto then obtain pc\<^sub>2 where pc\<^sub>2: "?PC pc\<^sub>2" and 2: "?Exec pc\<^sub>2" by iprover then have "?err" using pc\<^sub>2 jvm_trans[OF 1 2] by(auto intro!: exI[where x="pc\<^sub>2"] exI[where x="vs'@[w]"]) } thus "?eq \ ?err" by simp qed next case (CallNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 M') have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" let ?xa = "addr_of_sys_xcpt NullPointer" obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\M'(es)) = (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)" using CallNull\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa (map Val pvs) (Null#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact have Isubs: "{pc + length (compE\<^sub>2 e)..2 e) + length (compEs\<^sub>2 es)} \ I - pcs (compxE\<^sub>2 e pc (length vs))" using CallNull\<^sub>1.prems by clarsimp have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using Jcc_pieces_Call1[OF pcs] IH by clarsimp also have "P \ \ -jvm\ (None,h\<^sub>2,(rev pvs@Null#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using CallNull\<^sub>1 IH_es Isubs by auto also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (rev pvs@Null#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using CallNull\<^sub>1.prems by(auto simp:split_beta handle_def nth_append simp del: split_paired_Ex) finally show ?case by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="rev pvs@[Null]"]) next case (CallObjThrow\<^sub>1 e h ls sh e' h' ls' sh' M' es) obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (e\M'(es)) = (True, (vs, ls, C,M,pc,ics)#frs, (None, h', (v#vs, ls', C,M,pc+size(compE\<^sub>2 (e\M'(es))),ics)#frs,sh'), err)" using CallObjThrow\<^sub>1.prems(1) by clarsimp obtain a' where throw: "throw e' = Throw a'" using eval\<^sub>1_final[OF CallObjThrow\<^sub>1.hyps(1)] by clarsimp have IH: "PROP ?P e h ls sh (throw e') h' ls' sh' E C M pc ics v a' vs frs (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact show ?case using IH Jcc_pieces_Call1[OF pcs] throw CallObjThrow\<^sub>1.prems nsub_RI_Jcc_pieces by auto next case (CallNone\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs M') let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs" let ?\\<^sub>0 = "(None,h\<^sub>0,?frs\<^sub>0,sh\<^sub>0)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>1 = "(None,h\<^sub>1,(Addr a#vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" let ?frs\<^sub>2 = "(rev pvs @ Addr a # vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" let ?xa = "addr_of_sys_xcpt NoSuchMethodError" have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) have aux: "(rev pvs @ Addr a # vs) ! length pvs = Addr a" by (metis length_rev nth_append_length) have nmeth: "\(\b Ts T body D. P \ C' sees M', b : Ts\T = body in D)" using sees_method_compPD CallNone\<^sub>1.hyps(6) by fastforce obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\M'(es)) = (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)" using CallNone\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa (map Val pvs) (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact have "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp also have "P \ \ -jvm\ ?\\<^sub>2" using IH_es CallNone\<^sub>1.prems by fastforce also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (rev pvs@Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using CallNone\<^sub>1.hyps(5) CallNone\<^sub>1.prems aux nmeth by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2", auto simp: handle_def) finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="rev pvs@[Addr a]"]) next case (CallStatic\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs M' Ts T body D) let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs" let ?\\<^sub>0 = "(None,h\<^sub>0,?frs\<^sub>0,sh\<^sub>0)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>1 = "(None,h\<^sub>1,(Addr a#vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" let ?frs\<^sub>2 = "(rev pvs @ Addr a # vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError" have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) have aux: "(rev pvs @ Addr a # vs) ! length pvs = Addr a" by (metis length_rev nth_append_length) obtain body' where method: "P \ C' sees M', Static : Ts\T = body' in D" by (metis CallStatic\<^sub>1.hyps(6) P_def compP\<^sub>2_def sees_method_compP) obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\M'(es)) = (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)" using CallStatic\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa (map Val pvs) (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact have "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp also have "P \ \ -jvm\ ?\\<^sub>2" using IH_es CallStatic\<^sub>1.prems by fastforce also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (rev pvs@Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using CallStatic\<^sub>1.hyps(5) CallStatic\<^sub>1.prems aux method by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2") (auto simp: handle_def; meson frames_of.cases) finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="rev pvs@[Addr a]"]) next case (SCallParamsThrow\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 pvs ex es'' C' M') show ?case proof(cases "M' = clinit \ es = []") case clinit: True then show ?thesis using SCallParamsThrow\<^sub>1.hyps(1,3) evals\<^sub>1_cases(1) by fastforce next case nclinit: False let ?\\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)" have Isubs: "{pc..2 es)} \ I" using SCallParamsThrow\<^sub>1.prems nclinit by clarsimp show ?thesis (is "?N \ (?eq \ ?err)") proof show ?N by simp next { assume ?eq moreover have "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa es'' vs frs I" by fact ultimately have "\pc\<^sub>2. (pc \ pc\<^sub>2 \ pc\<^sub>2 < pc + size(compEs\<^sub>2 es) \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es pc (size vs))) \ (\vs'. P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2)" (is "\pc\<^sub>2. ?PC pc\<^sub>2 \ ?Exec pc\<^sub>2") using SCallParamsThrow\<^sub>1 Isubs nclinit by auto then obtain pc\<^sub>2 where pc\<^sub>2: "?PC pc\<^sub>2" and 2: "?Exec pc\<^sub>2" by iprover then have "?err" using pc\<^sub>2 2 by(auto intro: exI[where x="pc\<^sub>2"]) } thus "?eq \ ?err" by iprover qed qed next case (SCallNone\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' M') show ?case proof(cases "M' = clinit \ es = []") case clinit: True then show ?thesis using SCallNone\<^sub>1.hyps(3) SCallNone\<^sub>1.prems by auto next case nclinit: False let ?\\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)" let ?frs\<^sub>2 = "(rev pvs @ vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" let ?xa = "addr_of_sys_xcpt NoSuchMethodError" have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) have nmeth: "\(\b Ts T body D. P \ C' sees M', b : Ts\T = body in D)" using sees_method_compPD SCallNone\<^sub>1.hyps(3) by fastforce have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa (map Val pvs) vs frs I" by fact have "P \ ?\\<^sub>1 -jvm\ ?\\<^sub>2" using IH_es SCallNone\<^sub>1.prems nclinit by auto fastforce+ also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (rev pvs@vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using SCallNone\<^sub>1.prems nmeth nclinit by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2", auto simp: handle_def) finally show ?thesis using nclinit by (auto intro: exI[where x = ?pc\<^sub>2]) qed next case (SCallNonStatic\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' M' Ts T body D) show ?case proof(cases "M' = clinit \ es = []") case clinit: True then show ?thesis using SCallNonStatic\<^sub>1.hyps(3) SCallNonStatic\<^sub>1.prems sees_method_fun by fastforce next case nclinit: False let ?\\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)" let ?frs\<^sub>2 = "(rev pvs @ vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError" have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) obtain body' where method: "P \ C' sees M', NonStatic : Ts\T = body' in D" by (metis SCallNonStatic\<^sub>1.hyps(3) P_def compP\<^sub>2_def sees_method_compP) have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa (map Val pvs) vs frs I" by fact have "P \ ?\\<^sub>1 -jvm\ ?\\<^sub>2" using IH_es SCallNonStatic\<^sub>1.prems nclinit by auto fastforce+ also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (rev pvs@vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using SCallNonStatic\<^sub>1.prems method nclinit by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2") (auto simp: handle_def; meson frames_of.cases) finally show ?thesis using nclinit by (auto intro: exI[where x = ?pc\<^sub>2]) qed next case (SCallInitThrow\<^sub>1 es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 pvs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' M' Ts T body D a h\<^sub>2 ls\<^sub>2 sh\<^sub>2) show ?case proof(cases "M' = clinit \ es = []") case clinit: True then show ?thesis using SCallInitThrow\<^sub>1 by simp next case nclinit: False let ?\\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc,ics)#frs,sh\<^sub>0)" let ?pc\<^sub>1 = "pc + length(compEs\<^sub>2 es)" let ?frs\<^sub>1 = "(rev pvs @ vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs" let ?\\<^sub>1 = "(None,h\<^sub>1,?frs\<^sub>1,sh\<^sub>1)" let ?frs\<^sub>1' = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Calling D [])#frs" let ?\\<^sub>1' = "(None,h\<^sub>1,?frs\<^sub>1',sh\<^sub>1)" let ?frs\<^sub>2 = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Called [])#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" have ls: "ls\<^sub>1 = ls\<^sub>2" by(rule init\<^sub>1_same_loc[OF SCallInitThrow\<^sub>1.hyps(6)]) have method: "\m'. P \ C' sees M',Static:Ts\T = m' in D" using SCallInitThrow\<^sub>1.hyps(3) by (metis P_def compP\<^sub>2_def sees_method_compP) obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF SCallInitThrow\<^sub>1.hyps(6)] by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \ unit))" using sees_method_is_class'[OF SCallInitThrow\<^sub>1.hyps(3)] by auto then obtain err' where pcs': "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (rev pvs@vs) ls\<^sub>1 ?pc\<^sub>1 ics frs sh\<^sub>1 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (INIT D ([D],False) \ unit) = (True, ?frs\<^sub>1', (None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2), err')" using SCallInitThrow\<^sub>1.prems(1) nclinit by auto have IHI: "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v a' (rev pvs@vs) frs I" by fact have IH_es: "PROP ?Ps es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (map Val pvs) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C M pc ics pvs xa (map Val pvs) vs frs I" by fact have "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using IH_es SCallInitThrow\<^sub>1.prems nclinit by auto fastforce+ also have "P \ \ -jvm\ ?\\<^sub>1'" proof(cases "sh\<^sub>1 D") case None then show ?thesis using SCallInitThrow\<^sub>1.hyps(1,3-6) SCallInitThrow\<^sub>1.prems method by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using SCallInitThrow\<^sub>1.hyps(1,3-6) SCallInitThrow\<^sub>1.prems method Some by(cases ics; case_tac i, auto) qed also obtain vs' where "P \ \ -jvm\ handle P C M a' h\<^sub>2 (vs'@rev pvs@vs) ls\<^sub>1 ?pc\<^sub>1 ics frs sh\<^sub>2" using IHI pcs' throw by auto finally show ?thesis using nclinit throw ls by(auto intro!: exI[where x="?pc\<^sub>1"] exI[where x="vs'@rev pvs"]) qed next case (SCallInit\<^sub>1 es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 pvs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' M' Ts T body D v' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 ls\<^sub>2' e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3) show ?case proof(cases "M' = clinit \ es = []") case clinit: True then show ?thesis using SCallInit\<^sub>1 by simp next case nclinit: False let ?\\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc,ics)#frs,sh\<^sub>0)" let ?pc\<^sub>1 = "pc + length(compEs\<^sub>2 es)" let ?frs\<^sub>1 = "(rev pvs @ vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs" let ?\\<^sub>1 = "(None,h\<^sub>1,?frs\<^sub>1,sh\<^sub>1)" let ?frs\<^sub>1' = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Calling D [])#frs" let ?\\<^sub>1' = "(None,h\<^sub>1,?frs\<^sub>1',sh\<^sub>1)" let ?frs\<^sub>2 = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Called [])#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" let ?frs\<^sub>2' = "([], ls\<^sub>2', D,M',0,No_ics) # ?frs\<^sub>1" let ?\\<^sub>2' = "(None, h\<^sub>2, ?frs\<^sub>2', sh\<^sub>2)" have nclinit': "M' \ clinit" by fact have ics: "ics = No_ics" using SCallInit\<^sub>1.hyps(5) SCallInit\<^sub>1.prems by simp have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>0, ls\<^sub>0, sh\<^sub>0)\ [\] \map Val pvs,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) have invoke: "P,C,M,?pc\<^sub>1 \ Invokestatic C' M' (length Ts)" using SCallInit\<^sub>1.hyps(8) SCallInit\<^sub>1.prems nclinit by(auto simp: add.assoc) have nsub: "\ sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf SCallInit\<^sub>1.hyps(3)]) have ls: "ls\<^sub>1 = ls\<^sub>2" by(rule init\<^sub>1_same_loc[OF SCallInit\<^sub>1.hyps(6)]) obtain sfs i where sh\<^sub>2: "sh\<^sub>2 D = Some(sfs,i)" using init\<^sub>1_Val_PD[OF SCallInit\<^sub>1.hyps(6)] by clarsimp have method: "\m'. P \ C' sees M',Static:Ts\T = m' in D" using SCallInit\<^sub>1.hyps(3) by (metis P_def compP\<^sub>2_def sees_method_compP) have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \ unit))" using sees_method_is_class'[OF SCallInit\<^sub>1.hyps(3)] by auto then obtain err' where pcs': "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (rev pvs@vs) ls\<^sub>1 ?pc\<^sub>1 ics frs sh\<^sub>1 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v' xa (INIT D ([D],False) \ unit) = (True, ?frs\<^sub>1', (None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2), err')" using SCallInit\<^sub>1.prems(1) nclinit by auto have IHI: "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v') h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v' xa (rev pvs@vs) frs I" by fact have IH_es: "PROP ?Ps es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (map Val pvs) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C M pc ics pvs xa (map Val pvs) vs frs I" by fact have "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using IH_es SCallInit\<^sub>1.prems nclinit by auto fastforce+ also have "P \ \ -jvm\ ?\\<^sub>1'" proof(cases "sh\<^sub>1 D") case None then show ?thesis using SCallInit\<^sub>1.hyps(1,3-6,8-10) SCallInit\<^sub>1.prems method by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using SCallInit\<^sub>1.hyps(1,3-6,8-10) SCallInit\<^sub>1.prems method Some by(cases ics; case_tac i, auto) qed also have "P \ \ -jvm\ ?\\<^sub>2" using IHI pcs' by auto also have "P \ \ -jvm\ ?\\<^sub>2'" using jvm_Invokestatic_Called[OF assms(1) invoke _ SCallInit\<^sub>1.hyps(3,8,9)] sh\<^sub>2 ics by auto finally have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>2'". have "P\<^sub>1 \ C' sees M',Static: Ts\T = body in D" by fact then have M'_in_D: "P\<^sub>1 \ D sees M',Static: Ts\T = body in D" by(rule sees_method_idemp) have M'_code: "compP\<^sub>2 P\<^sub>1,D,M',0 \ compE\<^sub>2 body @ [Return]" using beforeM M'_in_D by simp have M'_xtab: "compP\<^sub>2 P\<^sub>1,D,M' \ compxE\<^sub>2 body 0 0/{..2 body)},0" using M'_in_D by(rule beforexM) have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3 (Class D # Ts) D M' 0 No_ics v xa [] ?frs\<^sub>1 ({..2 body)})" by fact have cond: "Jcc_cond P\<^sub>1 (Class D # Ts) D M' [] 0 No_ics {..2 body)} h\<^sub>2 sh\<^sub>2 body" using nsub_RI_Jcc_pieces[OF assms(1) nsub] M'_code M'_xtab by clarsimp show ?thesis (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>2' -jvm\ (None,h\<^sub>3,([v],ls\<^sub>3,D,M',size(compE\<^sub>2 body),No_ics)#?frs\<^sub>1,sh\<^sub>3)" using val IH_body SCallInit\<^sub>1.prems M'_code cond nsub_RI_Jcc_pieces nsub by auto also have "P \ \ -jvm\ (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>3)" using SCallInit\<^sub>1.hyps(8) M'_code M'_in_D ls nclinit' by(cases T, auto) finally show ?trans using nclinit by(auto simp:add.assoc) qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw with IH_body obtain pc\<^sub>2 vs' where pc\<^sub>2: "0 \ pc\<^sub>2 \ pc\<^sub>2 < size(compE\<^sub>2 body) \ \ caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and 2: "P \ ?\\<^sub>2' -jvm\ handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>1 sh\<^sub>3" using SCallInit\<^sub>1.prems M'_code M'_xtab cond nsub_RI_Jcc_pieces nsub by (auto simp del:split_paired_Ex) have "handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>1 sh\<^sub>3 = handle P C M xa h\<^sub>3 (rev pvs @ vs) ls\<^sub>2 ?pc\<^sub>1 ics frs sh\<^sub>3" using pc\<^sub>2 M'_in_D ls nclinit' by(auto simp add:handle_def) then show "?err" using pc\<^sub>2 jvm_trans[OF 1 2] nclinit by(auto intro!:exI[where x="?pc\<^sub>1"] exI[where x="rev pvs"]) qed qed qed next case (SCall\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' M' Ts T body D sfs ls\<^sub>2' e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3) show ?case proof(cases "M' = clinit \ es = []") case clinit: True then have s1: "pvs = []" "h\<^sub>1 = h\<^sub>2" "ls\<^sub>1 = ls\<^sub>2" "sh\<^sub>1 = sh\<^sub>2" using SCall\<^sub>1.hyps(1) evals\<^sub>1_cases(1) by blast+ then have ls\<^sub>2': "ls\<^sub>2' = replicate (max_vars body) undefined" using SCall\<^sub>1.hyps(6) clinit by simp let ?frs = "create_init_frame P C' # (vs, ls\<^sub>1, C,M,pc,ics)#frs" let ?\\<^sub>1 = "(None,h\<^sub>1,?frs,sh\<^sub>1)" have method: "P\<^sub>1 \ C' sees clinit,Static: []\Void = body in C'" using SCall\<^sub>1.hyps(3) clinit s1(1) wf_sees_clinit[OF wf] by (metis is_class_def option.collapse sees_method_fun sees_method_is_class) then have M_code: "compP\<^sub>2 P\<^sub>1,C',clinit,0 \ compE\<^sub>2 body @ [Return]" by(rule beforeM) have pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1 I h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa (C'\\<^sub>sclinit([])) = (True, ?frs, (None, h\<^sub>3, tl ?frs, sh\<^sub>3(C'\(fst(the(sh\<^sub>3 C')),Done))), P \ (None, h\<^sub>1, ?frs, sh\<^sub>1) -jvm\ (case ics of Called Cs \ (None, h\<^sub>3, (vs, ls\<^sub>1, C, M, pc, Throwing Cs xa) # frs, sh\<^sub>3(C' \ (fst (the (sh\<^sub>3 C')), Error)))))" using Jcc_pieces_clinit[OF assms(1),of E C M vs pc ics I h\<^sub>1 sh\<^sub>1 C' ls\<^sub>1 frs h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa] SCall\<^sub>1.prems(1) clinit s1(1) by clarsimp have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3 [] C' clinit 0 No_ics v xa [] (tl ?frs) ({..2 body)})" by fact show ?thesis (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val then have "P \ ?\\<^sub>1 -jvm\ (None, h\<^sub>3, ([v], ls\<^sub>3, C', clinit, size(compE\<^sub>2 body), No_ics) # tl ?frs,sh\<^sub>3)" using IH_body Jcc_pieces_SCall_clinit_body[OF assms(1) wf pcs method] s1 ls\<^sub>2' by clarsimp also have "P \ \ -jvm\ (None, h\<^sub>3, tl ?frs, sh\<^sub>3(C'\(fst(the(sh\<^sub>3 C')),Done)))" using jvm_Return_Init[OF M_code] by simp finally show ?trans using pcs s1 clinit by simp qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw with IH_body obtain pc\<^sub>2 vs2 where pc\<^sub>2: "0 \ pc\<^sub>2 \ pc\<^sub>2 < size(compE\<^sub>2 body) \ \ caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C' clinit xa h\<^sub>3 vs2 ls\<^sub>3 pc\<^sub>2 No_ics (tl ?frs) sh\<^sub>3" using SCall\<^sub>1.prems Jcc_pieces_SCall_clinit_body[OF assms(1) wf pcs method] s1 ls\<^sub>2' by clarsimp show ?err using SCall\<^sub>1.prems(1) clinit proof(cases ics) case (Called Cs) note 2 also have "handle P C' clinit xa h\<^sub>3 vs2 ls\<^sub>3 pc\<^sub>2 No_ics (tl ?frs) sh\<^sub>3 = (None, h\<^sub>3, (vs, ls\<^sub>1, C, M, pc, Throwing (C'#Cs) xa) # frs, sh\<^sub>3)" using Called pc\<^sub>2 method by(simp add: handle_def) also have "P \ \ -jvm\ (None, h\<^sub>3, (vs, ls\<^sub>1, C, M, pc, Throwing Cs xa) # frs, sh\<^sub>3(C' \ (fst (the (sh\<^sub>3 C')), Error)))" using Called jvm_Throwing by simp finally show ?thesis using pcs clinit Called by(clarsimp intro!: exI[where x="[]"]) qed(auto) qed qed next case nclinit: False let ?\\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)" let ?frs\<^sub>2 = "(rev pvs @ vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" let ?frs\<^sub>2' = "([], ls\<^sub>2', D,M',0,No_ics) # ?frs\<^sub>2" let ?\\<^sub>2' = "(None, h\<^sub>2, ?frs\<^sub>2', sh\<^sub>2)" have nclinit': "M' \ clinit" using wf_sees_clinit1[OF wf] visible_method_exists[OF SCall\<^sub>1.hyps(3)] sees_method_idemp[OF SCall\<^sub>1.hyps(3)] nclinit SCall\<^sub>1.hyps(5) evals\<^sub>1_preserves_elen[OF SCall\<^sub>1.hyps(1)] by fastforce have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) have invoke: "P,C,M,?pc\<^sub>2 \ Invokestatic C' M' (length Ts)" using SCall\<^sub>1.hyps(5) SCall\<^sub>1.prems nclinit by(auto simp: add.assoc) have nsub: "\ sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf SCall\<^sub>1.hyps(3)]) have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa (map Val pvs) vs frs I" by fact have "P \ ?\\<^sub>1 -jvm\ ?\\<^sub>2" using IH_es SCall\<^sub>1.prems nclinit by auto fastforce+ also have "P \ \ -jvm\ ?\\<^sub>2'" using jvm_Invokestatic[OF assms(1) invoke _ SCall\<^sub>1.hyps(3,5,6)] SCall\<^sub>1.hyps(4) SCall\<^sub>1.prems nclinit by auto finally have 1: "P \ ?\\<^sub>1 -jvm\ ?\\<^sub>2'". have "P\<^sub>1 \ C' sees M',Static: Ts\T = body in D" by fact then have M'_in_D: "P\<^sub>1 \ D sees M',Static: Ts\T = body in D" by(rule sees_method_idemp) have M'_code: "compP\<^sub>2 P\<^sub>1,D,M',0 \ compE\<^sub>2 body @ [Return]" using beforeM M'_in_D by simp have M'_xtab: "compP\<^sub>2 P\<^sub>1,D,M' \ compxE\<^sub>2 body 0 0/{..2 body)},0" using M'_in_D by(rule beforexM) have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3 (Class D # Ts) D M' 0 No_ics v xa [] ?frs\<^sub>2 ({..2 body)})" by fact have cond: "Jcc_cond P\<^sub>1 (Class D # Ts) D M' [] 0 No_ics {..2 body)} h\<^sub>2 sh\<^sub>2 body" using nsub_RI_Jcc_pieces[OF assms(1) nsub] M'_code M'_xtab by clarsimp show ?thesis (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>2' -jvm\ (None,h\<^sub>3,([v],ls\<^sub>3,D,M',size(compE\<^sub>2 body),No_ics)#?frs\<^sub>2,sh\<^sub>3)" using val IH_body SCall\<^sub>1.prems M'_code cond nsub_RI_Jcc_pieces nsub by auto also have "P \ \ -jvm\ (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>3)" using SCall\<^sub>1.hyps(5) M'_code M'_in_D nclinit' by(cases T, auto) finally show ?trans using nclinit by(auto simp:add.assoc) qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw with IH_body obtain pc\<^sub>2 vs' where pc\<^sub>2: "0 \ pc\<^sub>2 \ pc\<^sub>2 < size(compE\<^sub>2 body) \ \ caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and 2: "P \ ?\\<^sub>2' -jvm\ handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3" using SCall\<^sub>1.prems M'_code M'_xtab cond nsub_RI_Jcc_pieces nsub by (auto simp del:split_paired_Ex) have "handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3 = handle P C M xa h\<^sub>3 (rev pvs @ vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>3" using pc\<^sub>2 M'_in_D nclinit' by(auto simp add:handle_def) then show "?err" using pc\<^sub>2 jvm_trans[OF 1 2] nclinit by(auto intro:exI[where x="?pc\<^sub>2"]) qed qed qed next case Block\<^sub>1 then show ?case using nsub_RI_Jcc_pieces by auto next case (Seq\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>1)" let ?I = "I - pcs (compxE\<^sub>2 e\<^sub>2 (Suc ?pc\<^sub>1) (length vs))" have Isub: "{pc..2 e\<^sub>1)} \ ?I" using Seq\<^sub>1.prems by clarsimp have IH: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs ?I" by fact have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using Seq\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by auto also have "P \ \ -jvm\ ?\\<^sub>1" using Seq\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". let ?pc\<^sub>2 = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>2)" let ?I' = "I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs))" have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (?pc\<^sub>1+1) ics v xa vs frs ?I'" by fact have Isub2: "{Suc (pc + length (compE\<^sub>2 e\<^sub>1))..2 e\<^sub>1) + length (compE\<^sub>2 e\<^sub>2))} \ ?I'" using Seq\<^sub>1.prems by clarsimp show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note eval\<^sub>1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using val Seq\<^sub>1.prems nsub_RI_Jcc_pieces IH\<^sub>2 Isub2 by auto finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw then obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc\<^sub>1+1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2 \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 (?pc\<^sub>1+1) (size vs))" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using IH\<^sub>2 Seq\<^sub>1.prems nsub_RI_Jcc_pieces Isub2 by auto show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2]) qed qed next case (SeqThrow\<^sub>1 e\<^sub>0 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>1) let ?I = "I - pcs (compxE\<^sub>2 e\<^sub>1 (Suc (pc + length (compE\<^sub>2 e\<^sub>0))) (length vs))" obtain a' where throw: "throw e = Throw a'" using eval\<^sub>1_final[OF SeqThrow\<^sub>1.hyps(1)] by clarsimp have Isub: "{pc..2 e\<^sub>0)} \ ?I" using SeqThrow\<^sub>1.prems by clarsimp have "PROP ?P e\<^sub>0 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v a' vs frs ?I" by fact then show ?case using SeqThrow\<^sub>1.prems throw nsub_RI_Jcc_pieces Isub by auto next case (CondT\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>1 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>1)" let ?d = "size vs" let ?xt\<^sub>1 = "compxE\<^sub>2 e\<^sub>1 (pc+size(compE\<^sub>2 e)+1) ?d" let ?xt\<^sub>2 = "compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e)+size(compE\<^sub>2 e\<^sub>1)+2) ?d" let ?I = "I - (pcs ?xt\<^sub>1 \ pcs ?xt\<^sub>2)" have Isub: "{pc..2 e)} \ ?I" using CondT\<^sub>1.prems by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 true h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool True) xa vs frs ?I" by fact have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool(True)#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using CondT\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by(auto simp: Int_Un_distrib) also have "P \ \ -jvm\ ?\\<^sub>1" using CondT\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". let ?pc\<^sub>1' = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2' = "?pc\<^sub>1' + 1 + length(compE\<^sub>2 e\<^sub>2)" let ?I' = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>2 (?pc\<^sub>1'+1) ?d)" have IH2: "PROP ?P e\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (?pc\<^sub>1+1) ics v xa vs frs ?I'" by fact show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note eval\<^sub>1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>1',ics)#frs,sh\<^sub>2)" using val CondT\<^sub>1.prems nsub_RI_Jcc_pieces IH2 by(fastforce simp:Int_Un_distrib) also have "P \ \ -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2',ics)#frs,sh\<^sub>2)" using CondT\<^sub>1 nsub_RI_Jcc_pieces by(auto simp:add.assoc) finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw moreover note IH2 ultimately obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc\<^sub>1+1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1' \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) (size vs))" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using CondT\<^sub>1.prems nsub_RI_Jcc_pieces by (fastforce simp:Int_Un_distrib) show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2]) qed qed next case (CondF\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>1) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?pc\<^sub>2 = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>1)+ 1" let ?pc\<^sub>2' = "?pc\<^sub>2 + length(compE\<^sub>2 e\<^sub>2)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>1)" let ?d = "size vs" let ?xt\<^sub>1 = "compxE\<^sub>2 e\<^sub>1 (pc+size(compE\<^sub>2 e)+1) ?d" let ?xt\<^sub>2 = "compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e)+size(compE\<^sub>2 e\<^sub>1)+2) ?d" let ?I = "I - (pcs ?xt\<^sub>1 \ pcs ?xt\<^sub>2)" let ?I' = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) ?d)" have pcs: "pcs(compxE\<^sub>2 e pc ?d) \ pcs(?xt\<^sub>1 @ ?xt\<^sub>2) = {}" using CondF\<^sub>1.prems by (simp add:Int_Un_distrib) have Isub: "{pc..2 e)} \ ?I" using CondF\<^sub>1.prems by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 false h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool False) xa vs frs ?I" by fact have IH2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>2 ics v xa vs frs ?I'" by fact have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool(False)#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using CondF\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub pcs by auto also have "P \ \ -jvm\ ?\\<^sub>1" using CondF\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note eval\<^sub>1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2',ics)#frs,sh\<^sub>2)" using val CondF\<^sub>1.prems nsub_RI_Jcc_pieces IH2 by(fastforce simp:Int_Un_distrib) finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ ?err") proof let ?I' = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) ?d)" assume throw: ?throw then obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc\<^sub>2 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2' \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>2 ?d)" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using CondF\<^sub>1.prems nsub_RI_Jcc_pieces IH2 by(fastforce simp:Int_Un_distrib) show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2]) qed qed next case (CondThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 f h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>1 e\<^sub>2) let ?d = "size vs" let ?xt\<^sub>1 = "compxE\<^sub>2 e\<^sub>1 (pc+size(compE\<^sub>2 e)+1) ?d" let ?xt\<^sub>2 = "compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e)+size(compE\<^sub>2 e\<^sub>1)+2) ?d" let ?I = "I - (pcs ?xt\<^sub>1 \ pcs ?xt\<^sub>2)" have Isub: "{pc..2 e)} \ ?I" using CondThrow\<^sub>1.prems by clarsimp have "pcs(compxE\<^sub>2 e pc ?d) \ pcs(?xt\<^sub>1 @ ?xt\<^sub>2) = {}" using CondThrow\<^sub>1.prems by (simp add:Int_Un_distrib) moreover have "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw f) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs ?I" by fact ultimately show ?case using CondThrow\<^sub>1.prems nsub_RI_Jcc_pieces Isub by auto next case (WhileF\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 c) let ?pc = "pc + length(compE\<^sub>2 e)" let ?pc' = "?pc + length(compE\<^sub>2 c) + 3" have Isub: "{pc..2 e)} \ I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))" using WhileF\<^sub>1.prems by clarsimp have Isub2: "{Suc (pc + length (compE\<^sub>2 e))..2 e) + length (compE\<^sub>2 c))} \ I - pcs (compxE\<^sub>2 e pc (length vs))" using WhileF\<^sub>1.prems by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 false h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool False) xa vs frs (I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs)))" by fact have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Bool False#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using WhileF\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by auto also have "P \ \ -jvm\ (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc',ics)#frs,sh\<^sub>1)" using WhileF\<^sub>1 by (auto simp:add.assoc) also have "P \ \ -jvm\ (None,h\<^sub>1,(Unit#vs,ls\<^sub>1,C,M,?pc'+1,ics)#frs,sh\<^sub>1)" using WhileF\<^sub>1.prems by (auto simp:eval_nat_numeral) finally show ?case by (simp add:add.assoc eval_nat_numeral) next case (WhileT\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 c v\<^sub>1 h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>3 h\<^sub>3 ls\<^sub>3 sh\<^sub>3) let ?pc = "pc + length(compE\<^sub>2 e)" let ?pc' = "?pc + length(compE\<^sub>2 c) + 1" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?\\<^sub>2 = "(None,h\<^sub>2,(vs,ls\<^sub>2,C,M,pc,ics)#frs,sh\<^sub>2)" have Isub: "{pc..2 e)} \ I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))" using WhileT\<^sub>1.prems by clarsimp have Isub2: "{Suc (pc + length (compE\<^sub>2 e))..2 e) + length (compE\<^sub>2 c))} \ I - pcs (compxE\<^sub>2 e pc (length vs))" using WhileT\<^sub>1.prems by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 true h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool True) xa vs frs (I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs)))" by fact have IH2: "PROP ?P c h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v\<^sub>1) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (Suc ?pc) ics v\<^sub>1 xa vs frs (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool True#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using WhileT\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by auto also have "P \ \ -jvm\ (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)" using WhileT\<^sub>1.prems by auto also have "P \ \ -jvm\ (None,h\<^sub>2,(v\<^sub>1#vs,ls\<^sub>2,C,M,?pc',ics)#frs,sh\<^sub>2)" using WhileT\<^sub>1.prems nsub_RI_Jcc_pieces IH2 Isub2 by auto also have "P \ \ -jvm\ ?\\<^sub>2" using WhileT\<^sub>1.prems by auto finally have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>2". show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>2 -jvm\ (None,h\<^sub>3,(v#vs,ls\<^sub>3,C,M,?pc'+3,ics)#frs,sh\<^sub>3)" using val WhileT\<^sub>1 by (auto simp add:add.assoc eval_nat_numeral) finally show ?trans by(simp add:add.assoc eval_nat_numeral) qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw moreover have "PROP ?P (while (e) c) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>3 h\<^sub>3 ls\<^sub>3 sh\<^sub>3 E C M pc ics v xa vs frs I" by fact ultimately obtain pc\<^sub>2 vs' where pc\<^sub>2: "pc \ pc\<^sub>2 \ pc\<^sub>2 < ?pc'+3 \ \ caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 (while (e) c) pc (size vs))" and 2: "P \ ?\\<^sub>2 -jvm\ handle P C M xa h\<^sub>3 (vs'@vs) ls\<^sub>3 pc\<^sub>2 ics frs sh\<^sub>3" using WhileT\<^sub>1.prems by (auto simp:add.assoc eval_nat_numeral) show "?err" using pc\<^sub>2 jvm_trans[OF 1 2] by(auto intro: exI[where x=pc\<^sub>2]) qed qed next case (WhileCondThrow\<^sub>1 e h ls sh e' h' ls' sh' c) let ?I = "I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))" obtain a' where throw: "throw e' = Throw a'" using eval\<^sub>1_final[OF WhileCondThrow\<^sub>1.hyps(1)] by clarsimp have Isub: "{pc..2 e)} \ ?I" using WhileCondThrow\<^sub>1.prems by clarsimp have "PROP ?P e h ls sh (throw e') h' ls' sh' E C M pc ics v a' vs frs ?I" by fact then show ?case using WhileCondThrow\<^sub>1.prems throw nsub_RI_Jcc_pieces Isub by auto next case (WhileBodyThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 c e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>1)" let ?I = "I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))" have Isub: "{pc..2 e)} \ ?I" using WhileBodyThrow\<^sub>1.prems by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 true h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool True) xa vs frs ?I" by fact then have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool(True)#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using WhileBodyThrow\<^sub>1.prems nsub_RI_Jcc_pieces Isub by auto also have "P \ \ -jvm\ ?\\<^sub>1" using WhileBodyThrow\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". let ?pc\<^sub>1' = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 c)" show ?case (is "?Norm \ ?Err") proof show ?Norm by simp next show ?Err (is "?throw \ ?err") proof assume throw: ?throw moreover have "PROP ?P c h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (?pc\<^sub>1+1) ics v xa vs frs (I - pcs (compxE\<^sub>2 e pc (size vs)))" by fact ultimately obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc\<^sub>1+1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1' \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 c (?pc\<^sub>1+1) (size vs))" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using WhileBodyThrow\<^sub>1.prems nsub_RI_Jcc_pieces by (fastforce simp:Int_Un_distrib) show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2]) qed qed next case (Throw\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1) let ?pc = "pc + size(compE\<^sub>2 e)" have Isub: "{pc..2 e)} \ I" using Throw\<^sub>1.prems by clarsimp show ?case (is "?Norm \ ?Err") proof show ?Norm by simp next show ?Err (is "?throw \ ?err") proof assume throw:?throw have "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) a vs frs I" by fact then have "P \ (None, h\<^sub>0, (vs, ls\<^sub>0, C, M, pc, ics) # frs, sh\<^sub>0) -jvm\ (None, h\<^sub>1, (Addr xa#vs, ls\<^sub>1, C, M, ?pc, ics) # frs, sh\<^sub>1)" using Throw\<^sub>1 nsub_RI_Jcc_pieces Isub throw by auto also have "P \ \ -jvm\ handle P C M xa h\<^sub>1 (Addr xa#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using Throw\<^sub>1.prems by(auto simp add:handle_def) finally show "?err" by(auto intro!: exI[where x="?pc"] exI[where x="[Addr xa]"]) qed qed next case (ThrowNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1) let ?pc = "pc + size(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt NullPointer" have Isub: "{pc..2 e)} \ I" using ThrowNull\<^sub>1.prems by clarsimp show ?case (is "?Norm \ ?Err") proof show ?Norm by simp next show ?Err (is "?throw \ ?err") proof assume throw: ?throw have "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs I" by fact then have "P \ (None, h\<^sub>0, (vs, ls\<^sub>0, C, M, pc, ics) # frs, sh\<^sub>0) -jvm\ (None, h\<^sub>1, (Null#vs, ls\<^sub>1, C, M, ?pc, ics) # frs, sh\<^sub>1)" using ThrowNull\<^sub>1.prems nsub_RI_Jcc_pieces Isub by auto also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using ThrowNull\<^sub>1.prems by(auto simp add:handle_def) finally show "?err" using throw by(auto intro!: exI[where x="?pc"] exI[where x="[Null]"]) qed qed next case (ThrowThrow\<^sub>1 e h ls sh e' h' ls' sh') obtain a' where throw: "throw e' = Throw a'" using eval\<^sub>1_final[OF ThrowThrow\<^sub>1.hyps(1)] by clarsimp have Isub: "{pc..2 e)} \ I" using ThrowThrow\<^sub>1.prems by clarsimp have "PROP ?P e h ls sh (throw e') h' ls' sh' E C M pc ics v a' vs frs I" by fact then show ?case using ThrowThrow\<^sub>1.prems throw nsub_RI_Jcc_pieces Isub by auto next case (Try\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 Ci i e\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>1' = "?pc\<^sub>1 + 2 + length(compE\<^sub>2 e\<^sub>2)" have "{pc..2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \ I" using Try\<^sub>1.prems by simp also have "P,C,M \ compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs) / I,size vs" using Try\<^sub>1.prems by simp ultimately have "P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs) / {pc..2 e\<^sub>1)},size vs" by(rule beforex_try) hence "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using Try\<^sub>1 nsub_RI_Jcc_pieces by auto blast also have "P \ \ -jvm\ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1',ics)#frs,sh\<^sub>1)" using Try\<^sub>1.prems by auto finally show ?case by (auto simp:add.assoc) next case (TryCatch\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs Ci i e\<^sub>2 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2) let ?e = "try e\<^sub>1 catch(Ci i) e\<^sub>2" let ?xt = "compxE\<^sub>2 ?e pc (size vs)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?ls\<^sub>1 = "ls\<^sub>1[i := Addr a]" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>1' = "?pc\<^sub>1 + 2" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,?ls\<^sub>1,C,M, ?pc\<^sub>1',ics) # frs,sh\<^sub>1)" have I: "{pc..2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \ I" and beforex: "P,C,M \ ?xt/I,size vs" using TryCatch\<^sub>1.prems by simp+ have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,((Addr a)#vs,ls\<^sub>1,C,M, ?pc\<^sub>1+1,ics) # frs,sh\<^sub>1)" proof - have ics: "ics = No_ics" using TryCatch\<^sub>1.prems by auto have "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Throw a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v a vs frs {pc..2 e\<^sub>1)}" by fact moreover have "P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs)/{pc..1},size vs" using beforex I pcs_subset by(force elim!: beforex_appendD1) ultimately have "\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < ?pc\<^sub>1 \ \ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs)) \ (\vs'. P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)" using TryCatch\<^sub>1.prems nsub_RI_Jcc_pieces by auto then obtain pc\<^sub>1 vs' where pc\<^sub>1_in_e\<^sub>1: "pc \ pc\<^sub>1" "pc\<^sub>1 < ?pc\<^sub>1" and pc\<^sub>1_not_caught: "\ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs))" and 0: "P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1" by iprover from beforex obtain xt\<^sub>0 xt\<^sub>1 where ex_tab: "ex_table_of P C M = xt\<^sub>0 @ ?xt @ xt\<^sub>1" and disj: "pcs xt\<^sub>0 \ I = {}" by(auto simp:beforex_def) have hp: "h\<^sub>1 a = Some (D, fs)" "P\<^sub>1 \ D \\<^sup>* Ci" by fact+ have "pc\<^sub>1 \ pcs xt\<^sub>0" using pc\<^sub>1_in_e\<^sub>1 I disj by auto with pc\<^sub>1_in_e\<^sub>1 pc\<^sub>1_not_caught hp show ?thesis using ex_tab 0 ics by(simp add:handle_def matches_ex_entry_def) qed also have "P \ \ -jvm\ ?\\<^sub>1" using TryCatch\<^sub>1 by auto finally have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" . let ?pc\<^sub>2 = "?pc\<^sub>1' + length(compE\<^sub>2 e\<^sub>2)" let ?I\<^sub>2 = "{?pc\<^sub>1' ..< ?pc\<^sub>2}" have "P,C,M \ compxE\<^sub>2 ?e pc (size vs) / I,size vs" by fact hence beforex\<^sub>2: "P,C,M \ compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1' (size vs) / ?I\<^sub>2, size vs" using I pcs_subset[of _ ?pc\<^sub>1'] by(auto elim!:beforex_appendD2) have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ?ls\<^sub>1 sh\<^sub>1 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1' ics v xa vs frs ?I\<^sub>2" by fact show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using val beforex\<^sub>2 IH\<^sub>2 TryCatch\<^sub>1.prems nsub_RI_Jcc_pieces by auto finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw then obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc\<^sub>1+2 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2 \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1' (size vs))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using IH\<^sub>2 beforex\<^sub>2 TryCatch\<^sub>1.prems nsub_RI_Jcc_pieces by auto show ?err using pc\<^sub>2 jvm_trans[OF 1 2] by (simp add:match_ex_entry) (auto intro: exI[where x=pc\<^sub>2]) qed qed next case (TryThrow\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs Ci i e\<^sub>2) let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?e = "try e\<^sub>1 catch(Ci i) e\<^sub>2" let ?xt = "compxE\<^sub>2 ?e pc (size vs)" have I: "{pc..2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \ I" and beforex: "P,C,M \ ?xt/I,size vs" using TryThrow\<^sub>1.prems by simp+ have "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Throw a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v a vs frs {pc..2 e\<^sub>1)}" by fact moreover have "P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs)/{pc..1},size vs" using beforex I pcs_subset by(force elim!: beforex_appendD1) ultimately have "\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < ?pc\<^sub>1 \ \ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs)) \ (\vs'. P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)" using TryThrow\<^sub>1.prems nsub_RI_Jcc_pieces by auto then obtain pc\<^sub>1 vs' where pc\<^sub>1_in_e\<^sub>1: "pc \ pc\<^sub>1" "pc\<^sub>1 < ?pc\<^sub>1" and pc\<^sub>1_not_caught: "\ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs))" and 0: "P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1" by iprover show ?case (is "?N \ (?eq \ ?err)") proof show ?N by simp next { assume ?eq with TryThrow\<^sub>1 pc\<^sub>1_in_e\<^sub>1 pc\<^sub>1_not_caught 0 have "?err" by (simp add:match_ex_entry) auto } thus "?eq \ ?err" by iprover qed next case Nil\<^sub>1 thus ?case by simp next case (Cons\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es fs h\<^sub>2 ls\<^sub>2 sh\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?\\<^sub>1 = "(None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val v) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 [] C M pc ics v xa vs frs (I - pcs (compxEs\<^sub>2 es ?pc\<^sub>1 (Suc (length vs))))" by fact then have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using Jcc_pieces_Cons[OF _ Cons\<^sub>1.prems(1-5)] by auto let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" have IHs: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 fs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics (tl ws) xa es' (v#vs) frs (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(rev(ws) @ vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using val IHs Cons\<^sub>1.prems by fastforce finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof assume throw: ?throw then obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc\<^sub>1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2 \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es ?pc\<^sub>1 (size vs + 1))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@v#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using IHs Cons\<^sub>1.prems by(fastforce simp:Cons_eq_append_conv neq_Nil_conv) have "?H pc\<^sub>2" using Cons\<^sub>1.prems pc\<^sub>2 jvm_trans[OF 1 2] by(auto intro!: exI[where x="vs'@[v]"]) thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case (ConsThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es) then show ?case using Jcc_pieces_Cons[OF _ ConsThrow\<^sub>1.prems(1-5)] by (fastforce simp:Cons_eq_append_conv) next case InitFinal\<^sub>1 then show ?case using eval\<^sub>1_final_same[OF InitFinal\<^sub>1.hyps(1)] by clarsimp next case (InitNone\<^sub>1 sh C\<^sub>0 C' Cs e h l e' h' l' sh') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" using InitNone\<^sub>1.prems(1) by clarsimp let ?sh = "(sh(C\<^sub>0 \ (sblank P\<^sub>1 C\<^sub>0, Prepared)))" obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs" and frs\<^sub>1: "frs' \ Nil" using pcs by clarsimp then have 1: "P \ (None,h,frs',sh) -jvm\ (None,h,frs',?sh)" using InitNone\<^sub>1 jvm_InitNone[where P = P] by(cases frs', simp+) show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 note 1 also have "P \ (None,h,frs',?sh) -jvm\ (None,h',(vs,l,C,M,pc,Called [])#frs,sh')" using InitNone\<^sub>1.hyps(3)[of E] Jcc_pieces_InitNone[OF assms(1) pcs] InitNone\<^sub>1.prems val by clarsimp finally have ?jvm1 using pcs by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 note 1 also obtain vs' where "P \ (None,h,frs',?sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'" using InitNone\<^sub>1.hyps(3)[of E] Jcc_pieces_InitNone[OF assms(1) pcs] throw by clarsimp presburger finally have ?err using pcs by auto } thus "?e2 \ ?err" by simp qed next case (InitDone\<^sub>1 sh C\<^sub>0 sfs C' Cs e h l e' h' l' sh') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" using InitDone\<^sub>1.prems(1) by clarsimp let ?frs' = "(calling_to_scalled (hd frs'))#(tl frs')" have IH: "PROP ?P (INIT C' (Cs,True) \ e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I" by fact obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs" and frs\<^sub>1: "frs' \ Nil" using pcs by clarsimp then have 1: "P \ (None,h,frs',sh) -jvm\ (None,h,?frs',sh)" using InitDone\<^sub>1 jvm_InitDP[where P = P] by(cases frs', simp+) show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 note 1 also have "P \ (None,h,?frs',sh) -jvm\ (None,h',(vs,l,C,M,pc,Called [])#frs,sh')" using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitDone\<^sub>1.prems val by clarsimp finally have ?jvm1 using pcs by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 note 1 also obtain vs' where "P \ (None,h,?frs',sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'" using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitDone\<^sub>1.prems throw by clarsimp finally have ?err using pcs by auto } thus "?e2 \ ?err" by simp qed next case (InitProcessing\<^sub>1 sh C\<^sub>0 sfs C' Cs e h l e' h' l' sh') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" using InitProcessing\<^sub>1.prems(1) by clarsimp let ?frs' = "(calling_to_scalled (hd frs'))#(tl frs')" have IH: "PROP ?P (INIT C' (Cs,True) \ e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I" by fact obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs" and frs\<^sub>1: "frs' \ Nil" using pcs by clarsimp then have 1: "P \ (None,h,frs',sh) -jvm\ (None,h,?frs',sh)" using InitProcessing\<^sub>1 jvm_InitDP[where P = P] by(cases frs', simp+) show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 note 1 also have "P \ (None,h,?frs',sh) -jvm\ (None,h',(vs,l,C,M,pc,Called [])#frs,sh')" using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitProcessing\<^sub>1.prems val by clarsimp finally have ?jvm1 using pcs by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 note 1 also obtain vs' where "P \ (None,h,?frs',sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'" using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitProcessing\<^sub>1.prems throw by clarsimp finally have ?err using pcs by auto } thus "?e2 \ ?err" by simp qed next case (InitError\<^sub>1 sh C\<^sub>0 sfs Cs e h l e' h' l' sh' C') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" using InitError\<^sub>1.prems(1) by clarsimp let ?e\<^sub>0 = "THROW NoClassDefFoundError" let ?frs' = "(calling_to_sthrowing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs')" have IH: "PROP ?P (RI (C\<^sub>0,?e\<^sub>0) ; Cs \ e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I" by fact obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs" and frs\<^sub>1: "frs' \ Nil" and tl: "tl frs' = frs" using pcs by clarsimp then have 1: "P \ (None,h,frs',sh) -jvm\ (None,h,?frs',sh)" proof(cases frs') case (Cons a list) obtain vs' l' C' M' pc' ics' where a: "a = (vs',l',C',M',pc',ics')" by(cases a) then have "ics' = Calling C\<^sub>0 Cs" using Cons ics by simp then show ?thesis using Cons a IH InitError\<^sub>1.prems jvm_InitError[where P = P] InitError\<^sub>1.hyps(1) by simp qed(simp) show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 then have False using val rinit\<^sub>1_throw[OF InitError\<^sub>1.hyps(2)] by blast then have ?jvm1 using pcs by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 let ?frs = "(calling_to_throwing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs')" have exec: "exec (P, (None,h,?frs,sh)) = Some (None,h,?frs',sh)" using exec_ErrorThrowing[where sh=sh, OF InitError\<^sub>1.hyps(1)] ics by(cases "hd frs'", simp) obtain vs' where 2: "P \ (None,h,?frs,sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'" using IH Jcc_pieces_InitError[OF assms(1) pcs InitError\<^sub>1.hyps(1)] throw by clarsimp have neq: "(None, h, ?frs, sh) \ handle P C M xa h' (vs' @ vs) l pc ics frs sh'" using tl ics by(cases "hd frs'", simp add: handle_frs_tl_neq) note 1 also have "P \ (None,h,?frs',sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'" using exec_1_exec_all_conf[OF exec 2] neq by simp finally have ?err using pcs by auto } thus "?e2 \ ?err" by simp qed next case (InitObject\<^sub>1 sh C\<^sub>0 sfs sh' C' Cs e h l e' h' l' sh'') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' (sh(C\<^sub>0 \ (sfs, Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" using InitObject\<^sub>1.prems(1) by clarsimp let ?frs' = "(calling_to_called (hd frs'))#(tl frs')" have IH: "PROP ?P (INIT C' (C\<^sub>0#Cs,True) \ e) h l sh' e' h' l' sh'' E C M pc ics v xa vs frs I" by fact obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs" and frs\<^sub>1: "frs' \ Nil" using pcs by clarsimp then have 1: "P \ (None,h,frs',sh) -jvm\ (None,h,?frs',sh')" proof(cases frs') case (Cons a list) obtain vs' l' C' M' pc' ics' where a: "a = (vs',l',C',M',pc',ics')" by(cases a) then have "ics' = Calling C\<^sub>0 Cs" using Cons ics by simp then show ?thesis using Cons Nil a IH InitObject\<^sub>1 jvm_InitObj[where P = P] by simp qed(simp) show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 note 1 also have "P \ (None,h,?frs',sh') -jvm\ (None,h',(vs,l,C,M,pc,Called [])#frs,sh'')" using IH Jcc_pieces_InitObj[OF assms(1) pcs] InitObject\<^sub>1 val by simp finally have ?jvm1 using pcs by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 note 1 also obtain vs' where "P \ (None,h,?frs',sh') -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh''" using IH Jcc_pieces_InitObj[OF assms(1) pcs] InitObject\<^sub>1 throw by clarsimp finally have ?err using pcs by auto } thus "?e2 \ ?err" by simp qed next case (InitNonObject\<^sub>1 sh C\<^sub>0 sfs D a b sh' C' Cs e h l e' h' l' sh'') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' (sh(C\<^sub>0 \ (sfs,Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" using InitNonObject\<^sub>1.prems(1) by clarsimp let ?frs' = "(calling_to_calling (hd frs') D)#(tl frs')" have cls1: "is_class P\<^sub>1 D" using InitNonObject\<^sub>1.hyps(2,3) class_wf wf wf_cdecl_supD by blast have cls_aux: "distinct (C\<^sub>0#Cs) \ supercls_lst P\<^sub>1 (C\<^sub>0#Cs)" using InitNonObject\<^sub>1.prems(1) by auto then have cls2: "D \ set (C\<^sub>0 # Cs)" proof - have "distinct (D # C\<^sub>0 # Cs)" using InitNonObject\<^sub>1.hyps(2,3) cls_aux wf wf_supercls_distinct_app by blast then show "D \ set (C\<^sub>0 # Cs)" by (metis distinct.simps(2)) qed have cls3: "\C\set (C\<^sub>0 # Cs). P\<^sub>1 \ C \\<^sup>* D" using InitNonObject\<^sub>1.hyps(2,3) cls_aux by (metis r_into_rtrancl rtrancl_into_rtrancl set_ConsD subcls1.subcls1I supercls_lst.simps(1)) have IH: "PROP ?P (INIT C' (D # C\<^sub>0 # Cs,False) \ e) h l sh' e' h' l' sh'' E C M pc ics v xa vs frs I" by fact obtain r where cls: "class P C\<^sub>0 = \(D, r)\" using InitNonObject\<^sub>1.hyps(3) by (metis assms class_compP compP\<^sub>2_def) obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs" and frs\<^sub>1: "frs' \ Nil" using pcs by clarsimp then have 1: "P \ (None,h,frs',sh) -jvm\ (None,h,?frs',sh')" proof(cases frs') case (Cons a list) obtain vs' l' C' M' pc' ics' where a: "a = (vs',l',C',M',pc',ics')" by(cases a) then have "ics' = Calling C\<^sub>0 Cs" using Cons ics by simp then show ?thesis using Cons a IH InitNonObject\<^sub>1 jvm_InitNonObj[OF _ _ cls] by simp qed(simp) show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 note 1 also have "P \ (None,h,?frs',sh') -jvm\ (None,h',(vs,l,C,M,pc,Called [])#frs,sh'')" using IH Jcc_pieces_InitNonObj[OF assms(1) cls1 cls2 cls3 pcs] InitNonObject\<^sub>1 val by simp finally have ?jvm1 using pcs by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 note 1 also obtain vs' where "P \ (None,h,?frs',sh') -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh''" using IH Jcc_pieces_InitNonObj[OF assms(1) cls1 cls2 cls3 pcs] InitNonObject\<^sub>1 throw by clarsimp finally have ?err using pcs by auto } thus "?e2 \ ?err" by simp qed next case (InitRInit\<^sub>1 C\<^sub>0 Cs e h l sh e' h' l' sh' C') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,True) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" using InitRInit\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; Cs \ e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I" by fact show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 have "P \ (None,h,frs',sh) -jvm\ (None,h',(vs,l,C,M,pc,Called [])#frs,sh')" using IH Jcc_pieces_InitRInit[OF assms(1,2) pcs] InitRInit\<^sub>1.prems val by simp then have ?jvm1 using pcs by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 obtain vs' where "P \ (None,h,frs',sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'" using IH Jcc_pieces_InitRInit[OF assms(1,2) pcs] InitRInit\<^sub>1 throw by clarsimp then have ?err using pcs by auto } thus "?e2 \ ?err" by simp qed next case (RInit\<^sub>1 e h l sh v1 h' l' sh' C\<^sub>0 sfs i sh'' C' Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1) let ?frs = "(vs,l,C,M,pc,Called (C\<^sub>0#Cs)) # frs" let ?frs' = "(vs,l,C,M,pc,Called Cs) # frs" have clinit: "e = C\<^sub>0\\<^sub>sclinit([])" using RInit\<^sub>1 by (metis Jcc_cond.simps(2) eval\<^sub>1_final_same exp.distinct(101) final_def) then obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; Cs \ e') = (True, ?frs, (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)" using RInit\<^sub>1.prems(1) by simp have shC: "\C'\set Cs. \sfs. sh C' = \(sfs, Processing)\" using RInit\<^sub>1.prems(1) clinit by clarsimp then have shC'': "\C'\set Cs. \sfs. sh'' C' = \(sfs, Processing)\" using clinit\<^sub>1_proc_pres[OF wf] RInit\<^sub>1.hyps(1) clinit RInit\<^sub>1.hyps(4) RInit\<^sub>1.prems(1) by (auto simp: fun_upd_apply) have loc: "l = l'" using clinit\<^sub>1_loc_pres RInit\<^sub>1.hyps(1) clinit by simp have IH: "PROP ?P e h l sh (Val v1) h' l' sh' E C M pc (Called Cs) v1 xa vs (tl ?frs') I" by fact then have IH': "PROP ?P (C\<^sub>0\\<^sub>sclinit([])) h l sh (Val v1) h' l' sh' E C M pc (Called Cs) v1 xa vs (tl ?frs') I" using clinit by simp have IH2: "PROP ?P (INIT C' (Cs,True) \ e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact have "P \ (None,h,?frs,sh) -jvm\ (None,h,create_init_frame P C\<^sub>0 # ?frs',sh)" by(rule jvm_Called) also have "P \ \ -jvm\ (None,h',?frs',sh'')" using IH' Jcc_pieces_RInit_clinit[OF assms(1-2) pcs,of h' l' sh'] RInit\<^sub>1.hyps(3,4) by simp finally have jvm1: "P \ (None,h,?frs,sh) -jvm\ (None,h',?frs',sh'')" . show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 note jvm1 also have "P \ (None,h',?frs',sh'') -jvm\ (None,h\<^sub>1,(vs,l,C,M,pc,Called [])#frs,sh\<^sub>1)" using IH2 Jcc_pieces_RInit_Init[OF assms(1-2) shC'' pcs,of h'] RInit\<^sub>1.hyps(5) loc val by auto finally have ?jvm1 using pcs clinit by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 note jvm1 also obtain vs' where "P \ (None,h',?frs',sh'') -jvm\ handle P C M xa h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1" using IH2 Jcc_pieces_RInit_Init[OF assms(1-2) shC'' pcs,of h'] RInit\<^sub>1.hyps(5) loc throw by auto finally have ?err using pcs clinit by auto } thus "?e2 \ ?err" by simp qed next case (RInitInitFail\<^sub>1 e h l sh a h' l' sh' C\<^sub>0 sfs i sh'' D Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1) let ?frs = "(vs,l,C,M,pc,Called (C\<^sub>0#D#Cs)) # frs" let ?frs' = "(vs,l,C,M,pc,Called (D#Cs)) # frs" let "?frsT" = "\xa1. (vs,l,C,M,pc,Throwing (C\<^sub>0#D#Cs) xa1) # frs" let "?frsT'" = "\xa1. (vs,l,C,M,pc,Throwing (D#Cs) xa1) # frs" obtain xa' where xa': "throw a = Throw xa'" by (metis RInitInitFail\<^sub>1.hyps(1) eval\<^sub>1_final exp.distinct(101) final_def) have e\<^sub>1: "e\<^sub>1 = Throw xa'" using xa' rinit\<^sub>1_throw RInitInitFail\<^sub>1.hyps(5) by simp show ?case proof(cases "e = C\<^sub>0\\<^sub>sclinit([])") case clinit: True then obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa' (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; D # Cs \ e') = (True, ?frs, (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)" using RInitInitFail\<^sub>1.prems(1) by simp have loc: "l = l'" using clinit\<^sub>1_loc_pres RInitInitFail\<^sub>1.hyps(1) clinit by simp have IH: "PROP ?P e h l sh (throw a) h' l' sh' E C M pc (Called (D#Cs)) v xa' vs frs I" by fact then have IH': "PROP ?P (C\<^sub>0\\<^sub>sclinit([])) h l sh (Throw xa') h' l' sh' E C M pc (Called (D#Cs)) v xa' vs frs I" using clinit xa' by simp have IH2: "PROP ?P (RI (D,throw a) ; Cs \ e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 E C M pc ics v xa' vs frs I" by fact have "P \ (None,h,?frs,sh) -jvm\ (None,h,create_init_frame P C\<^sub>0 # ?frs',sh)" by(rule jvm_Called) also have "P \ \ -jvm\ (None,h',(vs, l, C, M, pc, Throwing (D#Cs) xa') # frs,sh'')" using IH' Jcc_pieces_RInit_clinit[OF assms(1-2) pcs,of h' l' sh'] RInitInitFail\<^sub>1.hyps(3,4) by simp also obtain vs'' where "P \ \ -jvm\ handle P C M xa' h\<^sub>1 (vs''@vs) l pc ics frs sh\<^sub>1" using IH2 pcs Jcc_pieces_RInit_RInit[OF assms(1) pcs] RInitInitFail\<^sub>1.hyps(3,4) xa' loc e\<^sub>1 xa' by clarsimp finally show ?thesis using pcs e\<^sub>1 clinit by auto next case throw: False then have eT: "e = Throw xa'" "h = h'" "l = l'" "sh = sh'" using xa' RInitInitFail\<^sub>1.prems(1) eval\<^sub>1_final_same[OF RInitInitFail\<^sub>1.hyps(1)] by clarsimp+ obtain a' where "class P\<^sub>1 C\<^sub>0 = \a'\" using RInitInitFail\<^sub>1.prems by(auto simp: is_class_def) then obtain stk' loc' M' pc' ics' where "create_init_frame P C\<^sub>0 = (stk',loc',C\<^sub>0,M',pc',ics')" using create_init_frame_wf_eq[OF wf] by(cases "create_init_frame P C\<^sub>0", simp) then obtain rhs err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh'' v xa' (RI (C\<^sub>0,e) ; D#Cs \ e') = (True, ?frsT xa', rhs, err)" using RInitInitFail\<^sub>1.prems(1) eT by clarsimp have IH2: "PROP ?P (RI (D,throw a) ; Cs \ e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 E C M pc ics v xa' vs frs I" by fact have "P \ (None,h,?frsT xa',sh') -jvm\ (None,h,?frsT' xa',sh'(C\<^sub>0 \ (fst (the (sh' C\<^sub>0)), Error)))" by(rule jvm_Throwing) also obtain vs' where "P \ ... -jvm\ handle P C M xa' h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1" using IH2 Jcc_pieces_RInit_RInit[OF assms(1) pcs] RInitInitFail\<^sub>1.hyps(3,4) eT e\<^sub>1 xa' by clarsimp finally show ?thesis using pcs e\<^sub>1 throw eT by auto qed next case (RInitFailFinal\<^sub>1 e h l sh a h' l' sh' C\<^sub>0 sfs i sh'' e'') let ?frs = "(vs,l,C,M,pc,Called [C\<^sub>0]) # frs" let ?frs' = "(vs,l,C,M,pc,Called []) # frs" let "?frsT" = "\xa1. (vs,l,C,M,pc,Throwing [C\<^sub>0] xa1) # frs" let "?frsT'" = "\xa1. (vs,l,C,M,pc,Throwing [] xa1) # frs" obtain xa' where xa': "throw a = Throw xa'" by (metis RInitFailFinal\<^sub>1.hyps(1) eval\<^sub>1_final exp.distinct(101) final_def) show ?case proof(cases "e = C\<^sub>0\\<^sub>sclinit([])") case clinit: True then obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh'' v xa' (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; [] \ unit) = (True, ?frs, (None, h', ?frs', sh''), err)" using RInitFailFinal\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h l sh (throw a) h' l' sh' E C M pc (Called []) v xa' vs frs I" by fact then have IH': "PROP ?P (C\<^sub>0\\<^sub>sclinit([])) h l sh (throw a) h' l' sh' E C M pc (Called []) v xa' vs frs I" using clinit by simp have "P \ (None,h,?frs,sh) -jvm\ (None,h,create_init_frame P C\<^sub>0 # ?frs',sh)" by(rule jvm_Called) also have "P \ \ -jvm\ (None,h',?frsT' xa',sh'')" using IH' Jcc_pieces_RInit_clinit[OF assms(1-2) pcs,of h' l' sh'] xa' RInitFailFinal\<^sub>1.hyps(3,4) by simp also have "P \ \ -jvm\ handle (compP compMb\<^sub>2 P\<^sub>1) C M xa' h' vs l pc No_ics frs sh''" using RInitFailFinal\<^sub>1.hyps(3,4) jvm_RInit_throw[where h=h' and sh=sh''] by simp finally show ?thesis using xa' pcs clinit by(clarsimp intro!: exI[where x="[]"]) next case throw: False then have eT: "e = Throw xa'" "h = h'" "sh = sh'" using xa' RInitFailFinal\<^sub>1.prems(1) eval\<^sub>1_final_same[OF RInitFailFinal\<^sub>1.hyps(1)] by clarsimp+ obtain a where "class P\<^sub>1 C\<^sub>0 = \a\" using RInitFailFinal\<^sub>1.prems by(auto simp: is_class_def) then obtain stk' loc' M' pc' ics' where "create_init_frame P C\<^sub>0 = (stk',loc',C\<^sub>0,M',pc',ics')" using create_init_frame_wf_eq[OF wf] by(cases "create_init_frame P C\<^sub>0", simp) then obtain rhs err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh'' v xa' (RI (C\<^sub>0,e) ; [] \ unit) = (True, ?frsT xa', rhs, err)" using RInitFailFinal\<^sub>1.prems(1) eT by clarsimp have "P \ (None,h,?frsT xa',sh) -jvm\ (None,h,?frsT' xa',sh(C\<^sub>0 \ (fst (the (sh C\<^sub>0)), Error)))" by(rule jvm_Throwing) also have "P \ \ -jvm\ handle P C M xa' h' vs l pc No_ics frs sh''" using RInitFailFinal\<^sub>1.hyps(3,4) jvm_RInit_throw[where h=h and sh=sh''] eT by simp finally show ?thesis using pcs xa' by(clarsimp intro!: exI[where x="[]"]) qed qed (*>*) (*FIXME move! *) lemma atLeast0AtMost[simp]: "{0::nat..n} = {..n}" by auto lemma atLeast0LessThan[simp]: "{0::nat.. addr option" where "exception (Throw a) = Some a" | "exception e = None" lemma comp\<^sub>2_correct: assumes wf: "wf_J\<^sub>1_prog P\<^sub>1" and "method": "P\<^sub>1 \ C sees M,b:Ts\T = body in C" and eval: "P\<^sub>1 \\<^sub>1 \body,(h,ls,sh)\ \ \e',(h',ls',sh')\" and nclinit: "M \ clinit" shows "compP\<^sub>2 P\<^sub>1 \ (None,h,[([],ls,C,M,0,No_ics)],sh) -jvm\ (exception e',h',[],sh')" (*<*) (is "_ \ ?\\<^sub>0 -jvm\ ?\\<^sub>1") proof - let ?P = "compP\<^sub>2 P\<^sub>1" let ?E = "case b of Static \ Ts | NonStatic \ Class C#Ts" have nsub: "\sub_RI body" using sees_wf\<^sub>1_nsub_RI[OF wf method] by simp have code: "?P,C,M,0 \ compE\<^sub>2 body" using beforeM[OF "method"] by auto have xtab: "?P,C,M \ compxE\<^sub>2 body 0 (size[])/{..2 body)},size[]" using beforexM[OF "method"] by auto have cond: "Jcc_cond P\<^sub>1 ?E C M [] 0 No_ics {..2 body)} h sh body" using nsub_RI_Jcc_pieces nsub code xtab by auto \ \Distinguish if e' is a value or an exception\ { fix v assume [simp]: "e' = Val v" have "?P \ ?\\<^sub>0 -jvm\ (None,h',[([v],ls',C,M,size(compE\<^sub>2 body),No_ics)],sh')" using Jcc[OF wf eval cond] nsub_RI_Jcc_pieces[OF _ nsub] by auto also have "?P \ \ -jvm\ ?\\<^sub>1" using beforeM[OF "method"] nclinit by auto finally have ?thesis . } moreover { fix a assume [simp]: "e' = Throw a" obtain pc vs' where pc: "0 \ pc \ pc < size(compE\<^sub>2 body) \ \ caught ?P pc h' a (compxE\<^sub>2 body 0 0)" and 1: "?P \ ?\\<^sub>0 -jvm\ handle ?P C M a h' vs' ls' pc No_ics [] sh'" using Jcc[OF wf eval cond] nsub_RI_Jcc_pieces[OF _ nsub] by auto meson from pc have "handle ?P C M a h' vs' ls' pc No_ics [] sh' = ?\\<^sub>1" using xtab "method" nclinit by(auto simp:handle_def compMb\<^sub>2_def) with 1 have ?thesis by simp } ultimately show ?thesis using eval\<^sub>1_final[OF eval] by(auto simp:final_def) qed (*>*) end diff --git a/thys/JinjaDCI/Compiler/J1.thy b/thys/JinjaDCI/Compiler/J1.thy --- a/thys/JinjaDCI/Compiler/J1.thy +++ b/thys/JinjaDCI/Compiler/J1.thy @@ -1,535 +1,548 @@ (* Title: JinjaDCI/Compiler/J1.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory Compiler/J1.thy by Tobias Nipkow *) chapter \ Compilation \label{cha:comp} \ section \ An Intermediate Language \ theory J1 imports "../J/BigStep" begin type_synonym expr\<^sub>1 = "nat exp" type_synonym J\<^sub>1_prog = "expr\<^sub>1 prog" type_synonym state\<^sub>1 = "heap \ (val list) \ sheap" definition hp\<^sub>1 :: "state\<^sub>1 \ heap" where "hp\<^sub>1 \ fst" definition lcl\<^sub>1 :: "state\<^sub>1 \ val list" where "lcl\<^sub>1 \ fst \ snd" definition shp\<^sub>1 :: "state\<^sub>1 \ sheap" where "shp\<^sub>1 \ snd \ snd" (*<*) declare hp\<^sub>1_def[simp] lcl\<^sub>1_def[simp] shp\<^sub>1_def[simp] (*>*) primrec max_vars :: "'a exp \ nat" and max_varss :: "'a exp list \ nat" where "max_vars(new C) = 0" | "max_vars(Cast C e) = max_vars e" | "max_vars(Val v) = 0" | "max_vars(e\<^sub>1 \bop\ e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)" | "max_vars(Var V) = 0" | "max_vars(V:=e) = max_vars e" | "max_vars(e\F{D}) = max_vars e" | "max_vars(C\\<^sub>sF{D}) = 0" | "max_vars(FAss e\<^sub>1 F D e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)" | "max_vars(SFAss C F D e\<^sub>2) = max_vars e\<^sub>2" | "max_vars(e\M(es)) = max (max_vars e) (max_varss es)" | "max_vars(C\\<^sub>sM(es)) = max_varss es" | "max_vars({V:T; e}) = max_vars e + 1" | "max_vars(e\<^sub>1;;e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)" | "max_vars(if (e) e\<^sub>1 else e\<^sub>2) = max (max_vars e) (max (max_vars e\<^sub>1) (max_vars e\<^sub>2))" | "max_vars(while (b) e) = max (max_vars b) (max_vars e)" | "max_vars(throw e) = max_vars e" | "max_vars(try e\<^sub>1 catch(C V) e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2 + 1)" | "max_vars(INIT C (Cs,b) \ e) = max_vars e" | "max_vars(RI(C,e);Cs \ e') = max (max_vars e) (max_vars e')" | "max_varss [] = 0" | "max_varss (e#es) = max (max_vars e) (max_varss es)" inductive eval\<^sub>1 :: "J\<^sub>1_prog \ expr\<^sub>1 \ state\<^sub>1 \ expr\<^sub>1 \ state\<^sub>1 \ bool" ("_ \\<^sub>1 ((1\_,/_\) \/ (1\_,/_\))" [51,0,0,0,0] 81) and evals\<^sub>1 :: "J\<^sub>1_prog \ expr\<^sub>1 list \ state\<^sub>1 \ expr\<^sub>1 list \ state\<^sub>1 \ bool" ("_ \\<^sub>1 ((1\_,/_\) [\]/ (1\_,/_\))" [51,0,0,0,0] 81) for P :: J\<^sub>1_prog where New\<^sub>1: "\ sh C = Some (sfs, Done); new_Addr h = Some a; P \ C has_fields FDTs; h' = h(a\blank P C) \ \ P \\<^sub>1 \new C,(h,l,sh)\ \ \addr a,(h',l,sh)\" | NewFail\<^sub>1: "\ sh C = Some (sfs, Done); new_Addr h = None \ \ P \\<^sub>1 \new C, (h,l,sh)\ \ \THROW OutOfMemory,(h,l,sh)\" | NewInit\<^sub>1: "\ \sfs. sh C = Some (sfs, Done); P \\<^sub>1 \INIT C ([C],False) \ unit,(h,l,sh)\ \ \Val v',(h',l',sh')\; new_Addr h' = Some a; P \ C has_fields FDTs; h'' = h'(a\blank P C) \ \ P \\<^sub>1 \new C,(h,l,sh)\ \ \addr a,(h'',l',sh')\" | NewInitOOM\<^sub>1: "\ \sfs. sh C = Some (sfs, Done); P \\<^sub>1 \INIT C ([C],False) \ unit,(h,l,sh)\ \ \Val v',(h',l',sh')\; new_Addr h' = None; is_class P C \ \ P \\<^sub>1 \new C,(h,l,sh)\ \ \THROW OutOfMemory,(h',l',sh')\" | NewInitThrow\<^sub>1: "\ \sfs. sh C = Some (sfs, Done); P \\<^sub>1 \INIT C ([C],False) \ unit,(h,l,sh)\ \ \throw a,s'\; is_class P C \ \ P \\<^sub>1 \new C,(h,l,sh)\ \ \throw a,s'\" | Cast\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,l,sh)\; h a = Some(D,fs); P \ D \\<^sup>* C \ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \addr a,(h,l,sh)\" | CastNull\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \null,s\<^sub>1\" | CastFail\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,l,sh)\; h a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \THROW ClassCast,(h,l,sh)\" | CastThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | Val\<^sub>1: "P \\<^sub>1 \Val v,s\ \ \Val v,s\" | BinOp\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v\<^sub>2,s\<^sub>2\; binop(bop,v\<^sub>1,v\<^sub>2) = Some v \ \ P \\<^sub>1 \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\ \ \Val v,s\<^sub>2\" | BinOpThrow\<^sub>1\<^sub>1: "P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \throw e,s\<^sub>1\ \ P \\<^sub>1 \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0\ \ \throw e,s\<^sub>1\" | BinOpThrow\<^sub>2\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \throw e,s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\ \ \throw e,s\<^sub>2\" | Var\<^sub>1: "\ ls!i = v; i < size ls \ \ P \\<^sub>1 \Var i,(h,ls,sh)\ \ \Val v,(h,ls,sh)\" | LAss\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \Val v,(h,ls,sh)\; i < size ls; ls' = ls[i := v] \ \ P \\<^sub>1 \i:= e,s\<^sub>0\ \ \unit,(h,ls',sh)\" | LAssThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \i:= e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAcc\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,ls,sh)\; h a = Some(C,fs); P \ C has F,NonStatic:t in D; fs(F,D) = Some v \ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \Val v,(h,ls,sh)\" | FAccNull\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \THROW NullPointer,s\<^sub>1\" | FAccThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAccNone\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,ls,sh)\; h a = Some(C,fs); \(\b t. P \ C has F,b:t in D) \ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \THROW NoSuchFieldError,(h,ls,sh)\" | FAccStatic\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,ls,sh)\; h a = Some(C,fs); P \ C has F,Static:t in D \ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \THROW IncompatibleClassChangeError,(h,ls,sh)\" | SFAcc\<^sub>1: "\ P \ C has F,Static:t in D; sh D = Some (sfs,Done); sfs F = Some v \ \ P \\<^sub>1 \C\\<^sub>sF{D},(h,ls,sh)\ \ \Val v,(h,ls,sh)\" | SFAccInit\<^sub>1: "\ P \ C has F,Static:t in D; \sfs. sh D = Some (sfs,Done); P \\<^sub>1 \INIT D ([D],False) \ unit,(h,ls,sh)\ \ \Val v',(h',ls',sh')\; sh' D = Some (sfs,i); sfs F = Some v \ \ P \\<^sub>1 \C\\<^sub>sF{D},(h,ls,sh)\ \ \Val v,(h',ls',sh')\" | SFAccInitThrow\<^sub>1: "\ P \ C has F,Static:t in D; \sfs. sh D = Some (sfs,Done); P \\<^sub>1 \INIT D ([D],False) \ unit,(h,ls,sh)\ \ \throw a,s'\ \ \ P \\<^sub>1 \C\\<^sub>sF{D},(h,ls,sh)\ \ \throw a,s'\" | SFAccNone\<^sub>1: "\ \(\b t. P \ C has F,b:t in D) \ \ P \\<^sub>1 \C\\<^sub>sF{D},s\ \ \THROW NoSuchFieldError,s\" | SFAccNonStatic\<^sub>1: "\ P \ C has F,NonStatic:t in D \ \ P \\<^sub>1 \C\\<^sub>sF{D},s\ \ \THROW IncompatibleClassChangeError,s\" | FAss\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\; h\<^sub>2 a = Some(C,fs); P \ C has F,NonStatic:t in D; fs' = fs((F,D)\v); h\<^sub>2' = h\<^sub>2(a\(C,fs')) \ \ P \\<^sub>1 \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \unit,(h\<^sub>2',l\<^sub>2,sh\<^sub>2)\" | FAssNull\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \null,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v,s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \THROW NullPointer,s\<^sub>2\" | FAssThrow\<^sub>1\<^sub>1: "P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAssThrow\<^sub>2\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \throw e',s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \throw e',s\<^sub>2\" | FAssNone\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\; h\<^sub>2 a = Some(C,fs); \(\b t. P \ C has F,b:t in D) \ \ P \\<^sub>1 \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \THROW NoSuchFieldError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\" | FAssStatic\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\; h\<^sub>2 a = Some(C,fs); P \ C has F,Static:t in D \ \ P \\<^sub>1 \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\" | SFAss\<^sub>1: "\ P \\<^sub>1 \e\<^sub>2,s\<^sub>0\ \ \Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\; P \ C has F,Static:t in D; sh\<^sub>1 D = Some(sfs,Done); sfs' = sfs(F\v); sh\<^sub>1' = sh\<^sub>1(D\(sfs',Done)) \ \ P \\<^sub>1 \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\ \ \unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1')\" | SFAssInit\<^sub>1: "\ P \\<^sub>1 \e\<^sub>2,s\<^sub>0\ \ \Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\; P \ C has F,Static:t in D; \sfs. sh\<^sub>1 D = Some(sfs,Done); P \\<^sub>1 \INIT D ([D],False) \ unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\ \ \Val v',(h',l',sh')\; sh' D = Some(sfs,i); sfs' = sfs(F\v); sh'' = sh'(D\(sfs',i)) \ \ P \\<^sub>1 \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\ \ \unit,(h',l',sh'')\" | SFAssInitThrow\<^sub>1: "\ P \\<^sub>1 \e\<^sub>2,s\<^sub>0\ \ \Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\; P \ C has F,Static:t in D; \sfs. sh\<^sub>1 D = Some(sfs,Done); P \\<^sub>1 \INIT D ([D],False) \ unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\ \ \throw a,s'\ \ \ P \\<^sub>1 \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\ \ \throw a,s'\" | SFAssThrow\<^sub>1: "P \\<^sub>1 \e\<^sub>2,s\<^sub>0\ \ \throw e',s\<^sub>2\ \ P \\<^sub>1 \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\ \ \throw e',s\<^sub>2\" | SFAssNone\<^sub>1: "\ P \\<^sub>1 \e\<^sub>2,s\<^sub>0\ \ \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\; \(\b t. P \ C has F,b:t in D) \ \ P \\<^sub>1 \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\ \ \THROW NoSuchFieldError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\" | SFAssNonStatic\<^sub>1: "\ P \\<^sub>1 \e\<^sub>2,s\<^sub>0\ \ \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\; P \ C has F,NonStatic:t in D \ \ P \\<^sub>1 \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\ \ \THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\" | CallObjThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \throw e',s\<^sub>1\" | CallNull\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \map Val vs,s\<^sub>2\ \ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \THROW NullPointer,s\<^sub>2\" | Call\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\; h\<^sub>2 a = Some(C,fs); P \ C sees M,NonStatic:Ts\T = body in D; size vs = size Ts; ls\<^sub>2' = (Addr a) # vs @ replicate (max_vars body) undefined; P \\<^sub>1 \body,(h\<^sub>2,ls\<^sub>2',sh\<^sub>2)\ \ \e',(h\<^sub>3,ls\<^sub>3,sh\<^sub>3)\ \ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \e',(h\<^sub>3,ls\<^sub>2,sh\<^sub>3)\" | CallParamsThrow\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \es',s\<^sub>2\; es' = map Val vs @ throw ex # es\<^sub>2 \ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \throw ex,s\<^sub>2\" | CallNone\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \ps,s\<^sub>1\ [\] \map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\; h\<^sub>2 a = Some(C,fs); \(\b Ts T body D. P \ C sees M,b:Ts\T = body in D) \ \ P \\<^sub>1 \e\M(ps),s\<^sub>0\ \ \THROW NoSuchMethodError,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\" | CallStatic\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \ps,s\<^sub>1\ [\] \map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\; h\<^sub>2 a = Some(C,fs); P \ C sees M,Static:Ts\T = body in D \ \ P \\<^sub>1 \e\M(ps),s\<^sub>0\ \ \THROW IncompatibleClassChangeError,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\" | SCallParamsThrow\<^sub>1: "\ P \\<^sub>1 \es,s\<^sub>0\ [\] \es',s\<^sub>2\; es' = map Val vs @ throw ex # es\<^sub>2 \ \ P \\<^sub>1 \C\\<^sub>sM(es),s\<^sub>0\ \ \throw ex,s\<^sub>2\" | SCallNone\<^sub>1: "\ P \\<^sub>1 \ps,s\<^sub>0\ [\] \map Val vs,s\<^sub>2\; \(\b Ts T body D. P \ C sees M,b:Ts\T = body in D) \ \ P \\<^sub>1 \C\\<^sub>sM(ps),s\<^sub>0\ \ \THROW NoSuchMethodError,s\<^sub>2\" | SCallNonStatic\<^sub>1: "\ P \\<^sub>1 \ps,s\<^sub>0\ [\] \map Val vs,s\<^sub>2\; P \ C sees M,NonStatic:Ts\T = body in D \ \ P \\<^sub>1 \C\\<^sub>sM(ps),s\<^sub>0\ \ \THROW IncompatibleClassChangeError,s\<^sub>2\" | SCallInitThrow\<^sub>1: "\ P \\<^sub>1 \ps,s\<^sub>0\ [\] \map Val vs,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\; P \ C sees M,Static:Ts\T = body in D; \sfs. sh\<^sub>1 D = Some(sfs,Done); M \ clinit; P \\<^sub>1 \INIT D ([D],False) \ unit,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\ \ \throw a,s'\ \ \ P \\<^sub>1 \C\\<^sub>sM(ps),s\<^sub>0\ \ \throw a,s'\" | SCallInit\<^sub>1: "\ P \\<^sub>1 \ps,s\<^sub>0\ [\] \map Val vs,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\; P \ C sees M,Static:Ts\T = body in D; \sfs. sh\<^sub>1 D = Some(sfs,Done); M \ clinit; P \\<^sub>1 \INIT D ([D],False) \ unit,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\ \ \Val v',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\; size vs = size Ts; ls\<^sub>2' = vs @ replicate (max_vars body) undefined; P \\<^sub>1 \body,(h\<^sub>2,ls\<^sub>2',sh\<^sub>2)\ \ \e',(h\<^sub>3,ls\<^sub>3,sh\<^sub>3)\ \ \ P \\<^sub>1 \C\\<^sub>sM(ps),s\<^sub>0\ \ \e',(h\<^sub>3,ls\<^sub>2,sh\<^sub>3)\" | SCall\<^sub>1: "\ P \\<^sub>1 \ps,s\<^sub>0\ [\] \map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\; P \ C sees M,Static:Ts\T = body in D; sh\<^sub>2 D = Some(sfs,Done) \ (M = clinit \ sh\<^sub>2 D = \(sfs, Processing)\); size vs = size Ts; ls\<^sub>2' = vs @ replicate (max_vars body) undefined; P \\<^sub>1 \body,(h\<^sub>2,ls\<^sub>2',sh\<^sub>2)\ \ \e',(h\<^sub>3,ls\<^sub>3,sh\<^sub>3)\ \ \ P \\<^sub>1 \C\\<^sub>sM(ps),s\<^sub>0\ \ \e',(h\<^sub>3,ls\<^sub>2,sh\<^sub>3)\" | Block\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \e',s\<^sub>1\ \ P \\<^sub>1 \Block i T e,s\<^sub>0\ \ \e',s\<^sub>1\" | Seq\<^sub>1: "\ P \\<^sub>1 \e\<^sub>0,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \e\<^sub>1,s\<^sub>1\ \ \e\<^sub>2,s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>0;;e\<^sub>1,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" | SeqThrow\<^sub>1: "P \\<^sub>1 \e\<^sub>0,s\<^sub>0\ \ \throw e,s\<^sub>1\ \ P \\<^sub>1 \e\<^sub>0;;e\<^sub>1,s\<^sub>0\ \ \throw e,s\<^sub>1\" | CondT\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \\<^sub>1 \e\<^sub>1,s\<^sub>1\ \ \e',s\<^sub>2\ \ \ P \\<^sub>1 \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \ \e',s\<^sub>2\" | CondF\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \false,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \e',s\<^sub>2\ \ \ P \\<^sub>1 \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \ \e',s\<^sub>2\" | CondThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\ \ \throw e',s\<^sub>1\" | WhileF\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \false,s\<^sub>1\ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \unit,s\<^sub>1\" | WhileT\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \\<^sub>1 \c,s\<^sub>1\ \ \Val v\<^sub>1,s\<^sub>2\; P \\<^sub>1 \while (e) c,s\<^sub>2\ \ \e\<^sub>3,s\<^sub>3\ \ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \e\<^sub>3,s\<^sub>3\" | WhileCondThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \throw e',s\<^sub>1\" | WhileBodyThrow\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \\<^sub>1 \c,s\<^sub>1\ \ \throw e',s\<^sub>2\\ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \throw e',s\<^sub>2\" | Throw\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,s\<^sub>1\ \ P \\<^sub>1 \throw e,s\<^sub>0\ \ \Throw a,s\<^sub>1\" | ThrowNull\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \\<^sub>1 \throw e,s\<^sub>0\ \ \THROW NullPointer,s\<^sub>1\" | ThrowThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \throw e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | Try\<^sub>1: "P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\ \ P \\<^sub>1 \try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\" | TryCatch\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\; h\<^sub>1 a = Some(D,fs); P \ D \\<^sup>* C; i < length ls\<^sub>1; P \\<^sub>1 \e\<^sub>2,(h\<^sub>1,ls\<^sub>1[i:=Addr a],sh\<^sub>1)\ \ \e\<^sub>2',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\ \ \ P \\<^sub>1 \try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\" | TryThrow\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\; h\<^sub>1 a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \\<^sub>1 \try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\ \ \Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\" | Nil\<^sub>1: "P \\<^sub>1 \[],s\ [\] \[],s\" | Cons\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \es',s\<^sub>2\ \ \ P \\<^sub>1 \e#es,s\<^sub>0\ [\] \Val v # es',s\<^sub>2\" | ConsThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e#es,s\<^sub>0\ [\] \throw e' # es, s\<^sub>1\" \ \ init rules \ | InitFinal\<^sub>1: "P \\<^sub>1 \e,s\ \ \e',s'\ \ P \\<^sub>1 \INIT C (Nil,b) \ e,s\ \ \e',s'\" | InitNone\<^sub>1: "\ sh C = None; P \\<^sub>1 \INIT C' (C#Cs,False) \ e,(h,l,sh(C \ (sblank P C, Prepared)))\ \ \e',s'\ \ \ P \\<^sub>1 \INIT C' (C#Cs,False) \ e,(h,l,sh)\ \ \e',s'\" | InitDone\<^sub>1: "\ sh C = Some(sfs,Done); P \\<^sub>1 \INIT C' (Cs,True) \ e,(h,l,sh)\ \ \e',s'\ \ \ P \\<^sub>1 \INIT C' (C#Cs,False) \ e,(h,l,sh)\ \ \e',s'\" | InitProcessing\<^sub>1: "\ sh C = Some(sfs,Processing); P \\<^sub>1 \INIT C' (Cs,True) \ e,(h,l,sh)\ \ \e',s'\ \ \ P \\<^sub>1 \INIT C' (C#Cs,False) \ e,(h,l,sh)\ \ \e',s'\" | InitError\<^sub>1: "\ sh C = Some(sfs,Error); P \\<^sub>1 \RI (C, THROW NoClassDefFoundError);Cs \ e,(h,l,sh)\ \ \e',s'\ \ \ P \\<^sub>1 \INIT C' (C#Cs,False) \ e,(h,l,sh)\ \ \e',s'\" | InitObject\<^sub>1: "\ sh C = Some(sfs,Prepared); C = Object; sh' = sh(C \ (sfs,Processing)); P \\<^sub>1 \INIT C' (C#Cs,True) \ e,(h,l,sh')\ \ \e',s'\ \ \ P \\<^sub>1 \INIT C' (C#Cs,False) \ e,(h,l,sh)\ \ \e',s'\" | InitNonObject\<^sub>1: "\ sh C = Some(sfs,Prepared); C \ Object; class P C = Some (D,r); sh' = sh(C \ (sfs,Processing)); P \\<^sub>1 \INIT C' (D#C#Cs,False) \ e,(h,l,sh')\ \ \e',s'\ \ \ P \\<^sub>1 \INIT C' (C#Cs,False) \ e,(h,l,sh)\ \ \e',s'\" | InitRInit\<^sub>1: "P \\<^sub>1 \RI (C,C\\<^sub>sclinit([]));Cs \ e,(h,l,sh)\ \ \e',s'\ \ P \\<^sub>1 \INIT C' (C#Cs,True) \ e,(h,l,sh)\ \ \e',s'\" | RInit\<^sub>1: "\ P \\<^sub>1 \e,s\ \ \Val v, (h',l',sh')\; sh' C = Some(sfs, i); sh'' = sh'(C \ (sfs, Done)); C' = last(C#Cs); P \\<^sub>1 \INIT C' (Cs,True) \ e', (h',l',sh'')\ \ \e\<^sub>1,s\<^sub>1\ \ \ P \\<^sub>1 \RI (C,e);Cs \ e',s\ \ \e\<^sub>1,s\<^sub>1\" | RInitInitFail\<^sub>1: "\ P \\<^sub>1 \e,s\ \ \throw a, (h',l',sh')\; sh' C = Some(sfs, i); sh'' = sh'(C \ (sfs, Error)); P \\<^sub>1 \RI (D,throw a);Cs \ e', (h',l',sh'')\ \ \e\<^sub>1,s\<^sub>1\ \ \ P \\<^sub>1 \RI (C,e);D#Cs \ e',s\ \ \e\<^sub>1,s\<^sub>1\" | RInitFailFinal\<^sub>1: "\ P \\<^sub>1 \e,s\ \ \throw a, (h',l',sh')\; sh' C = Some(sfs, i); sh'' = sh'(C \ (sfs, Error)) \ \ P \\<^sub>1 \RI (C,e);Nil \ e',s\ \ \throw a, (h',l',sh'')\" (*<*) lemmas eval\<^sub>1_evals\<^sub>1_induct = eval\<^sub>1_evals\<^sub>1.induct [split_format (complete)] and eval\<^sub>1_evals\<^sub>1_inducts = eval\<^sub>1_evals\<^sub>1.inducts [split_format (complete)] (*>*) inductive_cases eval\<^sub>1_cases [cases set]: "P \\<^sub>1 \new C,s\ \ \e',s'\" "P \\<^sub>1 \Cast C e,s\ \ \e',s'\" "P \\<^sub>1 \Val v,s\ \ \e',s'\" "P \\<^sub>1 \e\<^sub>1 \bop\ e\<^sub>2,s\ \ \e',s'\" "P \\<^sub>1 \Var v,s\ \ \e',s'\" "P \\<^sub>1 \V:=e,s\ \ \e',s'\" "P \\<^sub>1 \e\F{D},s\ \ \e',s'\" "P \\<^sub>1 \C\\<^sub>sF{D},s\ \ \e',s'\" "P \\<^sub>1 \e\<^sub>1\F{D}:=e\<^sub>2,s\ \ \e',s'\" "P \\<^sub>1 \C\\<^sub>sF{D}:=e\<^sub>2,s\ \ \e',s'\" "P \\<^sub>1 \e\M(es),s\ \ \e',s'\" "P \\<^sub>1 \C\\<^sub>sM(es),s\ \ \e',s'\" "P \\<^sub>1 \{V:T;e\<^sub>1},s\ \ \e',s'\" "P \\<^sub>1 \e\<^sub>1;;e\<^sub>2,s\ \ \e',s'\" "P \\<^sub>1 \if (e) e\<^sub>1 else e\<^sub>2,s\ \ \e',s'\" "P \\<^sub>1 \while (b) c,s\ \ \e',s'\" "P \\<^sub>1 \throw e,s\ \ \e',s'\" "P \\<^sub>1 \try e\<^sub>1 catch(C V) e\<^sub>2,s\ \ \e',s'\" "P \\<^sub>1 \INIT C (Cs,b) \ e,s\ \ \e',s'\" "P \\<^sub>1 \RI (C,e);Cs \ e\<^sub>0,s\ \ \e',s'\" inductive_cases evals\<^sub>1_cases [cases set]: "P \\<^sub>1 \[],s\ [\] \e',s'\" "P \\<^sub>1 \e#es,s\ [\] \e',s'\" (*>*) lemma eval\<^sub>1_final: "P \\<^sub>1 \e,s\ \ \e',s'\ \ final e'" and evals\<^sub>1_final: "P \\<^sub>1 \es,s\ [\] \es',s'\ \ finals es'" (*<*)by(induct rule:eval\<^sub>1_evals\<^sub>1.inducts, simp_all)(*>*) -lemma eval\<^sub>1_final_same: "\ P \\<^sub>1 \e,s\ \ \e',s'\; final e \ \ e = e' \ s = s'" +lemma eval\<^sub>1_final_same: +assumes eval: "P \\<^sub>1 \e,s\ \ \e',s'\" shows "final e \ e = e' \ s = s'" (*<*) -apply(erule finalE) - using eval\<^sub>1_cases(3) apply blast -by (metis eval\<^sub>1_cases(3,17) exp.distinct(101) exp.inject(3) val.distinct(13)) +proof(erule finalE) + fix v assume Val: "e = Val v" + then show ?thesis using eval eval\<^sub>1_cases(3) by blast +next + fix a assume "e = Throw a" + then show ?thesis using eval + by (metis eval\<^sub>1_cases(3,17) exp.distinct(101) exp.inject(3) val.distinct(13)) +qed (*>*) subsection "Property preservation" lemma eval\<^sub>1_preserves_len: "P \\<^sub>1 \e\<^sub>0,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\ \ \e\<^sub>1,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\ \ length ls\<^sub>0 = length ls\<^sub>1" and evals\<^sub>1_preserves_len: "P \\<^sub>1 \es\<^sub>0,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\ [\] \es\<^sub>1,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\ \ length ls\<^sub>0 = length ls\<^sub>1" (*<*)by (induct rule:eval\<^sub>1_evals\<^sub>1_inducts, simp_all)(*>*) lemma evals\<^sub>1_preserves_elen: "\es' s s'. P \\<^sub>1 \es,s\ [\] \es',s'\ \ length es = length es'" -(*<*) -apply(induct es type:list) -apply (auto elim:evals\<^sub>1.cases) -done -(*>*) +(*<*)by(induct es type:list) (auto elim:evals\<^sub>1.cases)(*>*) lemma clinit\<^sub>1_loc_pres: "P \\<^sub>1 \C\<^sub>0\\<^sub>sclinit([]),(h,l,sh)\ \ \e',(h',l',sh')\ \ l = l'" by(auto elim!: eval\<^sub>1_cases(12) evals\<^sub>1_cases(1)) lemma shows init\<^sub>1_ri\<^sub>1_same_loc: "P \\<^sub>1 \e,(h,l,sh)\ \ \e',(h',l',sh')\ \ (\C Cs b M a. e = INIT C (Cs,b) \ unit \ e = C\\<^sub>sM([]) \ e = RI (C,Throw a) ; Cs \ unit \ e = RI (C,C\\<^sub>sclinit([])) ; Cs \ unit \ l = l')" and "P \\<^sub>1 \es,(h,l,sh)\ [\] \es',(h',l',sh')\ \ True" proof(induct rule: eval\<^sub>1_evals\<^sub>1_inducts) case (RInitInitFail\<^sub>1 e h l sh a') then show ?case using eval\<^sub>1_final[of _ _ _ "throw a'"] by(fastforce dest: eval\<^sub>1_final_same[of _ "Throw a"]) next case RInitFailFinal\<^sub>1 then show ?case by(auto dest: eval\<^sub>1_final_same) qed(auto dest: evals\<^sub>1_cases eval\<^sub>1_cases(17) eval\<^sub>1_final_same) lemma init\<^sub>1_same_loc: "P \\<^sub>1 \INIT C (Cs,b) \ unit,(h,l,sh)\ \ \e',(h',l',sh')\ \ l = l'" by(simp add: init\<^sub>1_ri\<^sub>1_same_loc) theorem eval\<^sub>1_hext: "P \\<^sub>1 \e,(h,l,sh)\ \ \e',(h',l',sh')\ \ h \ h'" and evals\<^sub>1_hext: "P \\<^sub>1 \es,(h,l,sh)\ [\] \es',(h',l',sh')\ \ h \ h'" (*<*) proof (induct rule: eval\<^sub>1_evals\<^sub>1_inducts) case New\<^sub>1 thus ?case by(fastforce intro!: hext_new intro:LeastI simp:new_Addr_def split:if_split_asm simp del:fun_upd_apply) next case NewInit\<^sub>1 thus ?case by (meson hext_new hext_trans new_Addr_SomeD) next case FAss\<^sub>1 thus ?case by(auto simp:sym[THEN hext_upd_obj] simp del:fun_upd_apply elim!: hext_trans) qed (auto elim!: hext_trans) (*>*) subsection "Initialization" lemma rinit\<^sub>1_throw: "P\<^sub>1 \\<^sub>1 \RI (D,Throw xa) ; Cs \ e,(h, l, sh)\ \ \e',(h', l', sh')\ \ e' = Throw xa" -apply(induct Cs arbitrary: D h l sh h' l' sh') - apply(drule eval\<^sub>1_cases(20), auto elim: eval\<^sub>1_cases) -apply(drule eval\<^sub>1_cases(20), auto elim: eval\<^sub>1_cases dest: eval\<^sub>1_final_same simp: final_def) -done +proof(induct Cs arbitrary: D h l sh h' l' sh') + case Nil then show ?case + proof(rule eval\<^sub>1_cases(20)) qed(auto elim: eval\<^sub>1_cases) +next + case (Cons C Cs) show ?case using Cons.prems + proof(induct rule: eval\<^sub>1_cases(20)) + case RInit\<^sub>1: 1 + then show ?case using Cons.hyps by(auto elim: eval\<^sub>1_cases) + next + case RInitInitFail\<^sub>1: 2 + then show ?case using Cons.hyps eval\<^sub>1_final_same final_def by blast + next + case RInitFailFinal\<^sub>1: 3 then show ?case by simp + qed +qed lemma rinit\<^sub>1_throwE: "P \\<^sub>1 \RI (C,throw e) ; Cs \ e\<^sub>0,s\ \ \e',s'\ \ \a s\<^sub>t. e' = throw a \ P \\<^sub>1 \throw e,s\ \ \throw a,s\<^sub>t\" proof(induct Cs arbitrary: C e s) case Nil then show ?case proof(rule eval\<^sub>1_cases(20)) \ \ RI \ fix v h' l' sh' assume "P \\<^sub>1 \throw e,s\ \ \Val v,(h', l', sh')\" then show ?case using eval\<^sub>1_cases(17) by blast qed(auto) next case (Cons C' Cs') show ?case using Cons.prems(1) proof(rule eval\<^sub>1_cases(20)) \ \ RI \ fix v h' l' sh' assume "P \\<^sub>1 \throw e,s\ \ \Val v,(h', l', sh')\" then show ?case using eval\<^sub>1_cases(17) by blast next fix a h' l' sh' sfs i D Cs'' assume e''step: "P \\<^sub>1 \throw e,s\ \ \throw a,(h', l', sh')\" and shC: "sh' C = \(sfs, i)\" and riD: "P \\<^sub>1 \RI (D,throw a) ; Cs'' \ e\<^sub>0,(h', l', sh'(C \ (sfs, Error)))\ \ \e',s'\" and "C' # Cs' = D # Cs''" then show ?thesis using Cons.hyps eval\<^sub>1_final eval\<^sub>1_final_same by blast qed(simp) qed end diff --git a/thys/JinjaDCI/Compiler/J1WellForm.thy b/thys/JinjaDCI/Compiler/J1WellForm.thy --- a/thys/JinjaDCI/Compiler/J1WellForm.thy +++ b/thys/JinjaDCI/Compiler/J1WellForm.thy @@ -1,531 +1,519 @@ (* Title: JinjaDCI/Compiler/J1WellForm.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory Compiler/J1WellForm.thy by Tobias Nipkow *) section \ Well-Formedness of Intermediate Language \ theory J1WellForm imports "../J/JWellForm" J1 begin subsection "Well-Typedness" type_synonym env\<^sub>1 = "ty list" \ \type environment indexed by variable number\ inductive WT\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1, expr\<^sub>1 , ty ] \ bool" ("(_,_ \\<^sub>1/ _ :: _)" [51,51,51]50) and WTs\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1, expr\<^sub>1 list, ty list] \ bool" ("(_,_ \\<^sub>1/ _ [::] _)" [51,51,51]50) for P :: J\<^sub>1_prog where WTNew\<^sub>1: "is_class P C \ P,E \\<^sub>1 new C :: Class C" | WTCast\<^sub>1: "\ P,E \\<^sub>1 e :: Class D; is_class P C; P \ C \\<^sup>* D \ P \ D \\<^sup>* C \ \ P,E \\<^sub>1 Cast C e :: Class C" | WTVal\<^sub>1: "typeof v = Some T \ P,E \\<^sub>1 Val v :: T" | WTVar\<^sub>1: "\ E!i = T; i < size E \ \ P,E \\<^sub>1 Var i :: T" | WTBinOp\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1 :: T\<^sub>1; P,E \\<^sub>1 e\<^sub>2 :: T\<^sub>2; case bop of Eq \ (P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1) \ T = Boolean | Add \ T\<^sub>1 = Integer \ T\<^sub>2 = Integer \ T = Integer \ \ P,E \\<^sub>1 e\<^sub>1 \bop\ e\<^sub>2 :: T" | WTLAss\<^sub>1: "\ E!i = T; i < size E; P,E \\<^sub>1 e :: T'; P \ T' \ T \ \ P,E \\<^sub>1 i:=e :: Void" | WTFAcc\<^sub>1: "\ P,E \\<^sub>1 e :: Class C; P \ C sees F,NonStatic:T in D \ \ P,E \\<^sub>1 e\F{D} :: T" | WTSFAcc\<^sub>1: "\ P \ C sees F,Static:T in D \ \ P,E \\<^sub>1 C\\<^sub>sF{D} :: T" | WTFAss\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1 :: Class C; P \ C sees F,NonStatic:T in D; P,E \\<^sub>1 e\<^sub>2 :: T'; P \ T' \ T \ \ P,E \\<^sub>1 e\<^sub>1\F{D} := e\<^sub>2 :: Void" | WTSFAss\<^sub>1: "\ P \ C sees F,Static:T in D; P,E \\<^sub>1 e\<^sub>2 :: T'; P \ T' \ T \ \ P,E \\<^sub>1 C\\<^sub>sF{D}:=e\<^sub>2 :: Void" | WTCall\<^sub>1: "\ P,E \\<^sub>1 e :: Class C; P \ C sees M,NonStatic:Ts' \ T = m in D; P,E \\<^sub>1 es [::] Ts; P \ Ts [\] Ts' \ \ P,E \\<^sub>1 e\M(es) :: T" | WTSCall\<^sub>1: "\ P \ C sees M,Static:Ts \ T = m in D; P,E \\<^sub>1 es [::] Ts'; P \ Ts' [\] Ts; M \ clinit \ \ P,E \\<^sub>1 C\\<^sub>sM(es) :: T" | WTBlock\<^sub>1: "\ is_type P T; P,E@[T] \\<^sub>1 e::T' \ \ P,E \\<^sub>1 {i:T; e} :: T'" | WTSeq\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1::T\<^sub>1; P,E \\<^sub>1 e\<^sub>2::T\<^sub>2 \ \ P,E \\<^sub>1 e\<^sub>1;;e\<^sub>2 :: T\<^sub>2" | WTCond\<^sub>1: "\ P,E \\<^sub>1 e :: Boolean; P,E \\<^sub>1 e\<^sub>1::T\<^sub>1; P,E \\<^sub>1 e\<^sub>2::T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1; P \ T\<^sub>1 \ T\<^sub>2 \ T = T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ T = T\<^sub>1 \ \ P,E \\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 :: T" | WTWhile\<^sub>1: "\ P,E \\<^sub>1 e :: Boolean; P,E \\<^sub>1 c::T \ \ P,E \\<^sub>1 while (e) c :: Void" | WTThrow\<^sub>1: "P,E \\<^sub>1 e :: Class C \ P,E \\<^sub>1 throw e :: Void" | WTTry\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1 :: T; P,E@[Class C] \\<^sub>1 e\<^sub>2 :: T; is_class P C \ \ P,E \\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 :: T" | WTNil\<^sub>1: "P,E \\<^sub>1 [] [::] []" | WTCons\<^sub>1: "\ P,E \\<^sub>1 e :: T; P,E \\<^sub>1 es [::] Ts \ \ P,E \\<^sub>1 e#es [::] T#Ts" (*<*) declare WT\<^sub>1_WTs\<^sub>1.intros[intro!] declare WTNil\<^sub>1[iff] lemmas WT\<^sub>1_WTs\<^sub>1_induct = WT\<^sub>1_WTs\<^sub>1.induct [split_format (complete)] and WT\<^sub>1_WTs\<^sub>1_inducts = WT\<^sub>1_WTs\<^sub>1.inducts [split_format (complete)] inductive_cases eee[elim!]: "P,E \\<^sub>1 Val v :: T" "P,E \\<^sub>1 Var i :: T" "P,E \\<^sub>1 Cast D e :: T" "P,E \\<^sub>1 i:=e :: T" "P,E \\<^sub>1 {i:U; e} :: T" "P,E \\<^sub>1 e\<^sub>1;;e\<^sub>2 :: T" "P,E \\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 :: T" "P,E \\<^sub>1 while (e) c :: T" "P,E \\<^sub>1 throw e :: T" "P,E \\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 :: T" "P,E \\<^sub>1 e\F{D} :: T" "P,E \\<^sub>1 C\\<^sub>sF{D} :: T" "P,E \\<^sub>1 e\<^sub>1\F{D}:=e\<^sub>2 :: T" "P,E \\<^sub>1 C\\<^sub>sF{D}:=e\<^sub>2 :: T" "P,E \\<^sub>1 e\<^sub>1 \bop\ e\<^sub>2 :: T" "P,E \\<^sub>1 new C :: T" "P,E \\<^sub>1 e\M(es) :: T" "P,E \\<^sub>1 C\\<^sub>sM(es) :: T" "P,E \\<^sub>1 [] [::] Ts" "P,E \\<^sub>1 e#es [::] Ts" (*>*) lemma init_nWT\<^sub>1 [simp]:"\P,E \\<^sub>1 INIT C (Cs,b) \ e :: T" by(auto elim:WT\<^sub>1.cases) lemma rinit_nWT\<^sub>1 [simp]:"\P,E \\<^sub>1 RI(C,e);Cs \ e' :: T" by(auto elim:WT\<^sub>1.cases) lemma WTs\<^sub>1_same_size: "\Ts. P,E \\<^sub>1 es [::] Ts \ size es = size Ts" (*<*)by (induct es type:list) auto(*>*) lemma WT\<^sub>1_unique: "P,E \\<^sub>1 e :: T\<^sub>1 \ (\T\<^sub>2. P,E \\<^sub>1 e :: T\<^sub>2 \ T\<^sub>1 = T\<^sub>2)" and WTs\<^sub>1_unique: "P,E \\<^sub>1 es [::] Ts\<^sub>1 \ (\Ts\<^sub>2. P,E \\<^sub>1 es [::] Ts\<^sub>2 \ Ts\<^sub>1 = Ts\<^sub>2)" (*<*) -apply(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) -apply blast -apply blast -apply clarsimp -apply blast -apply clarsimp -apply(case_tac bop) -apply clarsimp -apply clarsimp -apply blast -apply (blast dest:sees_field_idemp sees_field_fun) -apply (blast dest:sees_field_fun) -apply blast -apply (blast dest:sees_field_fun) -apply (blast dest:sees_method_idemp sees_method_fun) -apply (blast dest:sees_method_fun) -apply blast -apply blast -apply blast -apply blast -apply clarify -apply blast -apply blast -apply blast -done +proof(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) + case WTVal\<^sub>1 then show ?case by clarsimp +next + case (WTBinOp\<^sub>1 E e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 bop T) + then show ?case by(case_tac bop) force+ +next + case WTFAcc\<^sub>1 then show ?case + by (blast dest:sees_field_idemp sees_field_fun) +next + case WTSFAcc\<^sub>1 then show ?case by (blast dest:sees_field_fun) +next + case WTSFAss\<^sub>1 then show ?case by (blast dest:sees_field_fun) +next + case WTCall\<^sub>1 then show ?case + by (blast dest:sees_method_idemp sees_method_fun) +next + case WTSCall\<^sub>1 then show ?case by (blast dest:sees_method_fun) +qed blast+ (*>*) lemma assumes wf: "wf_prog p P" shows WT\<^sub>1_is_type: "P,E \\<^sub>1 e :: T \ set E \ types P \ is_type P T" and "P,E \\<^sub>1 es [::] Ts \ True" (*<*) -apply(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) -apply simp -apply simp -apply (simp add:typeof_lit_is_type) -apply (blast intro:nth_mem) -apply(simp split:bop.splits) -apply simp -apply (simp add:sees_field_is_type[OF _ wf]) -apply (simp add:sees_field_is_type[OF _ wf]) -apply simp -apply simp -apply(fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def) -apply(fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def) -apply simp -apply simp -apply blast -apply simp -apply simp -apply simp -apply simp -apply simp -done +proof(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) + case WTVal\<^sub>1 then show ?case by (simp add:typeof_lit_is_type) +next + case WTVar\<^sub>1 then show ?case by (blast intro:nth_mem) +next + case WTBinOp\<^sub>1 then show ?case by (simp split:bop.splits) +next + case WTFAcc\<^sub>1 then show ?case + by (simp add:sees_field_is_type[OF _ wf]) +next + case WTSFAcc\<^sub>1 then show ?case + by (simp add:sees_field_is_type[OF _ wf]) +next + case WTCall\<^sub>1 then show ?case + by (fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def) +next + case WTSCall\<^sub>1 then show ?case + by (fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def) +next + case WTCond\<^sub>1 then show ?case by blast +qed simp+ (*>*) lemma WT\<^sub>1_nsub_RI: "P,E \\<^sub>1 e :: T \ \sub_RI e" and WTs\<^sub>1_nsub_RIs: "P,E \\<^sub>1 es [::] Ts \ \sub_RIs es" proof(induct rule: WT\<^sub>1_WTs\<^sub>1.inducts) qed(simp_all) subsection\ Runtime Well-Typedness \ inductive WTrt\<^sub>1 :: "J\<^sub>1_prog \ heap \ sheap \ env\<^sub>1 \ expr\<^sub>1 \ ty \ bool" and WTrts\<^sub>1 :: "J\<^sub>1_prog \ heap \ sheap \ env\<^sub>1 \ expr\<^sub>1 list \ ty list \ bool" and WTrt2\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1,heap,sheap,expr\<^sub>1,ty] \ bool" ("_,_,_,_ \\<^sub>1 _ : _" [51,51,51,51]50) and WTrts2\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1,heap,sheap,expr\<^sub>1 list, ty list] \ bool" ("_,_,_,_ \\<^sub>1 _ [:] _" [51,51,51,51]50) for P :: J\<^sub>1_prog and h :: heap and sh :: sheap where "P,E,h,sh \\<^sub>1 e : T \ WTrt\<^sub>1 P h sh E e T" | "P,E,h,sh \\<^sub>1 es[:]Ts \ WTrts\<^sub>1 P h sh E es Ts" | WTrtNew\<^sub>1: "is_class P C \ P,E,h,sh \\<^sub>1 new C : Class C" | WTrtCast\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : T; is_refT T; is_class P C \ \ P,E,h,sh \\<^sub>1 Cast C e : Class C" | WTrtVal\<^sub>1: "typeof\<^bsub>h\<^esub> v = Some T \ P,E,h,sh \\<^sub>1 Val v : T" | WTrtVar\<^sub>1: "\ E!i = T; i < size E \ \ P,E,h,sh \\<^sub>1 Var i : T" | WTrtBinOpEq\<^sub>1: "\ P,E,h,sh \\<^sub>1 e\<^sub>1 : T\<^sub>1; P,E,h,sh \\<^sub>1 e\<^sub>2 : T\<^sub>2 \ \ P,E,h,sh \\<^sub>1 e\<^sub>1 \Eq\ e\<^sub>2 : Boolean" | WTrtBinOpAdd\<^sub>1: "\ P,E,h,sh \\<^sub>1 e\<^sub>1 : Integer; P,E,h,sh \\<^sub>1 e\<^sub>2 : Integer \ \ P,E,h,sh \\<^sub>1 e\<^sub>1 \Add\ e\<^sub>2 : Integer" | WTrtLAss\<^sub>1: "\ E!i = T; i < size E; P,E,h,sh \\<^sub>1 e : T'; P \ T' \ T \ \ P,E,h,sh \\<^sub>1 i:=e : Void" | WTrtFAcc\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : Class C; P \ C has F,NonStatic:T in D \ \ P,E,h,sh \\<^sub>1 e\F{D} : T" | WTrtFAccNT\<^sub>1: "P,E,h,sh \\<^sub>1 e : NT \ P,E,h,sh \\<^sub>1 e\F{D} : T" | WTrtSFAcc\<^sub>1: "\ P \ C has F,Static:T in D \ \ P,E,h,sh \\<^sub>1 C\\<^sub>sF{D} : T" | WTrtFAss\<^sub>1: "\ P,E,h,sh \\<^sub>1 e\<^sub>1 : Class C; P \ C has F,NonStatic:T in D; P,E,h,sh \\<^sub>1 e\<^sub>2 : T\<^sub>2; P \ T\<^sub>2 \ T \ \ P,E,h,sh \\<^sub>1 e\<^sub>1\F{D}:=e\<^sub>2 : Void" | WTrtFAssNT\<^sub>1: "\ P,E,h,sh \\<^sub>1 e\<^sub>1:NT; P,E,h,sh \\<^sub>1 e\<^sub>2 : T\<^sub>2 \ \ P,E,h,sh \\<^sub>1 e\<^sub>1\F{D}:=e\<^sub>2 : Void" | WTrtSFAss\<^sub>1: "\ P,E,h,sh \\<^sub>1 e\<^sub>2 : T\<^sub>2; P \ C has F,Static:T in D; P \ T\<^sub>2 \ T \ \ P,E,h,sh \\<^sub>1 C\\<^sub>sF{D}:=e\<^sub>2 : Void" | WTrtCall\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : Class C; P \ C sees M,NonStatic:Ts \ T = m in D; P,E,h,sh \\<^sub>1 es [:] Ts'; P \ Ts' [\] Ts \ \ P,E,h,sh \\<^sub>1 e\M(es) : T" | WTrtCallNT\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : NT; P,E,h,sh \\<^sub>1 es [:] Ts \ \ P,E,h,sh \\<^sub>1 e\M(es) : T" | WTrtSCall\<^sub>1: "\ P \ C sees M,Static:Ts \ T = m in D; P,E,h,sh \\<^sub>1 es [:] Ts'; P \ Ts' [\] Ts; M = clinit \ sh D = \(sfs,Processing)\ \ es = map Val vs \ \ P,E,h,sh \\<^sub>1 C\\<^sub>sM(es) : T" | WTrtBlock\<^sub>1: "P,E@[T],h,sh \\<^sub>1 e : T' \ P,E,h,sh \\<^sub>1 {i:T; e} : T'" | WTrtSeq\<^sub>1: "\ P,E,h,sh \\<^sub>1 e\<^sub>1:T\<^sub>1; P,E,h,sh \\<^sub>1 e\<^sub>2:T\<^sub>2 \ \ P,E,h,sh \\<^sub>1 e\<^sub>1;;e\<^sub>2 : T\<^sub>2" | WTrtCond\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : Boolean; P,E,h,sh \\<^sub>1 e\<^sub>1:T\<^sub>1; P,E,h,sh \\<^sub>1 e\<^sub>2:T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1; P \ T\<^sub>1 \ T\<^sub>2 \ T = T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ T = T\<^sub>1 \ \ P,E,h,sh \\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 : T" | WTrtWhile\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : Boolean; P,E,h,sh \\<^sub>1 c:T \ \ P,E,h,sh \\<^sub>1 while(e) c : Void" | WTrtThrow\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : T\<^sub>r; is_refT T\<^sub>r \ \ P,E,h,sh \\<^sub>1 throw e : T" | WTrtTry\<^sub>1: "\ P,E,h,sh \\<^sub>1 e\<^sub>1 : T\<^sub>1; P,E@[Class C],h,sh \\<^sub>1 e\<^sub>2 : T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ \ P,E,h,sh \\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 : T\<^sub>2" | WTrtInit\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : T; \C' \ set (C#Cs). is_class P C'; \sub_RI e; \C' \ set (tl Cs). \sfs. sh C' = \(sfs,Processing)\; b \ (\C' \ set Cs. \sfs. sh C' = \(sfs,Processing)\); distinct Cs; supercls_lst P Cs \ \ P,E,h,sh \\<^sub>1 INIT C (Cs, b) \ e : T" | WTrtRI\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : T; P,E,h,sh \\<^sub>1 e' : T'; \C' \ set (C#Cs). is_class P C'; \sub_RI e'; \C' \ set (C#Cs). not_init C' e; \C' \ set Cs. \sfs. sh C' = \(sfs,Processing)\; \sfs. sh C = \(sfs, Processing)\ \ (sh C = \(sfs, Error)\ \ e = THROW NoClassDefFoundError); distinct (C#Cs); supercls_lst P (C#Cs) \ \ P,E,h,sh \\<^sub>1 RI(C, e);Cs \ e' : T'" \ \well-typed expression lists\ | WTrtNil\<^sub>1: "P,E,h,sh \\<^sub>1 [] [:] []" | WTrtCons\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : T; P,E,h,sh \\<^sub>1 es [:] Ts \ \ P,E,h,sh \\<^sub>1 e#es [:] T#Ts" (*<*) declare WTrt\<^sub>1_WTrts\<^sub>1.intros[intro!] WTrtNil\<^sub>1[iff] declare WTrtFAcc\<^sub>1[rule del] WTrtFAccNT\<^sub>1[rule del] WTrtSFAcc\<^sub>1[rule del] WTrtFAss\<^sub>1[rule del] WTrtFAssNT\<^sub>1[rule del] WTrtSFAss\<^sub>1[rule del] WTrtCall\<^sub>1[rule del] WTrtCallNT\<^sub>1[rule del] WTrtSCall\<^sub>1[rule del] lemmas WTrt\<^sub>1_induct = WTrt\<^sub>1_WTrts\<^sub>1.induct [split_format (complete)] and WTrt\<^sub>1_inducts = WTrt\<^sub>1_WTrts\<^sub>1.inducts [split_format (complete)] (*>*) (*<*) inductive_cases WTrt\<^sub>1_elim_cases[elim!]: "P,E,h,sh \\<^sub>1 Val v : T" "P,E,h,sh \\<^sub>1 Var i : T" "P,E,h,sh \\<^sub>1 v :=e : T" "P,E,h,sh \\<^sub>1 {i:U; e} : T" "P,E,h,sh \\<^sub>1 e\<^sub>1;;e\<^sub>2 : T\<^sub>2" "P,E,h,sh \\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 : T" "P,E,h,sh \\<^sub>1 while(e) c : T" "P,E,h,sh \\<^sub>1 throw e : T" "P,E,h,sh \\<^sub>1 try e\<^sub>1 catch(C V) e\<^sub>2 : T" "P,E,h,sh \\<^sub>1 Cast D e : T" "P,E,h,sh \\<^sub>1 e\F{D} : T" "P,E,h,sh \\<^sub>1 C\\<^sub>sF{D} : T" "P,E,h,sh \\<^sub>1 e\F{D} := v : T" "P,E,h,sh \\<^sub>1 C\\<^sub>sF{D} := v : T" "P,E,h,sh \\<^sub>1 e\<^sub>1 \bop\ e\<^sub>2 : T" "P,E,h,sh \\<^sub>1 new C : T" "P,E,h,sh \\<^sub>1 e\M{D}(es) : T" "P,E,h,sh \\<^sub>1 C\\<^sub>sM{D}(es) : T" "P,E,h,sh \\<^sub>1 INIT C (Cs,b) \ e : T" "P,E,h,sh \\<^sub>1 RI(C,e);Cs \ e' : T" "P,E,h,sh \\<^sub>1 [] [:] Ts" "P,E,h,sh \\<^sub>1 e#es [:] Ts" (*>*) lemma WT\<^sub>1_implies_WTrt\<^sub>1: "P,E \\<^sub>1 e :: T \ P,E,h,sh \\<^sub>1 e : T" and WTs\<^sub>1_implies_WTrts\<^sub>1: "P,E \\<^sub>1 es [::] Ts \ P,E,h,sh \\<^sub>1 es [:] Ts" (*<*) -apply(induct rule: WT\<^sub>1_WTs\<^sub>1_inducts) -apply fast -apply (fast) -apply(fastforce dest:typeof_lit_typeof) -apply(fast) -apply(rename_tac E e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 bop T) apply(case_tac bop) - apply(fastforce) - apply(fastforce) -apply(fastforce) -apply(meson WTrtFAcc\<^sub>1 has_visible_field) -apply(meson WTrtSFAcc\<^sub>1 has_visible_field) -apply(meson WTrtFAss\<^sub>1 has_visible_field) -apply(meson WTrtSFAss\<^sub>1 has_visible_field) -apply(fastforce simp: WTrtCall\<^sub>1) -apply(fastforce simp: WTrtSCall\<^sub>1) -apply(fastforce) -apply(fastforce) -apply(fastforce simp: WTrtCond\<^sub>1) -apply(fastforce) -apply(fastforce) -apply(fastforce) -apply(simp) -apply(fast) -done +proof(induct rule: WT\<^sub>1_WTs\<^sub>1_inducts) + case WTVal\<^sub>1 then show ?case by (fastforce dest:typeof_lit_typeof) +next + case (WTBinOp\<^sub>1 E e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 bop T) + then show ?case by (case_tac bop) fastforce+ +next + case WTFAcc\<^sub>1 then show ?case + by (fastforce simp: WTrtFAcc\<^sub>1 has_visible_field) +next + case WTSFAcc\<^sub>1 then show ?case + by (fastforce simp: WTrtSFAcc\<^sub>1 has_visible_field) +next + case WTFAss\<^sub>1 then show ?case by (meson WTrtFAss\<^sub>1 has_visible_field) +next + case WTSFAss\<^sub>1 then show ?case by (meson WTrtSFAss\<^sub>1 has_visible_field) +next + case WTCall\<^sub>1 then show ?case by (fastforce simp: WTrtCall\<^sub>1) +next + case WTSCall\<^sub>1 then show ?case by (fastforce simp: WTrtSCall\<^sub>1) +qed fastforce+ (*>*) subsection\ Well-formedness\ \ \Indices in blocks increase by 1\ primrec \ :: "expr\<^sub>1 \ nat \ bool" and \s :: "expr\<^sub>1 list \ nat \ bool" where "\ (new C) i = True" | "\ (Cast C e) i = \ e i" | "\ (Val v) i = True" | "\ (e\<^sub>1 \bop\ e\<^sub>2) i = (\ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (Var j) i = True" | "\ (e\F{D}) i = \ e i" | "\ (C\\<^sub>sF{D}) i = True" | "\ (j:=e) i = \ e i" | "\ (e\<^sub>1\F{D} := e\<^sub>2) i = (\ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (C\\<^sub>sF{D} := e\<^sub>2) i = \ e\<^sub>2 i" | "\ (e\M(es)) i = (\ e i \ \s es i)" | "\ (C\\<^sub>sM(es)) i = \s es i" | "\ ({j:T ; e}) i = (i = j \ \ e (i+1))" | "\ (e\<^sub>1;;e\<^sub>2) i = (\ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (if (e) e\<^sub>1 else e\<^sub>2) i = (\ e i \ \ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (throw e) i = \ e i" | "\ (while (e) c) i = (\ e i \ \ c i)" | "\ (try e\<^sub>1 catch(C j) e\<^sub>2) i = (\ e\<^sub>1 i \ i=j \ \ e\<^sub>2 (i+1))" | "\ (INIT C (Cs,b) \ e) i = \ e i" | "\ (RI(C,e);Cs \ e') i = (\ e i \ \ e' i)" | "\s [] i = True" | "\s (e#es) i = (\ e i \ \s es i)" definition wf_J\<^sub>1_mdecl :: "J\<^sub>1_prog \ cname \ expr\<^sub>1 mdecl \ bool" where "wf_J\<^sub>1_mdecl P C \ \(M,b,Ts,T,body). \sub_RI body \ (case b of NonStatic \ (\T'. P,Class C#Ts \\<^sub>1 body :: T' \ P \ T' \ T) \ \ body \{..size Ts}\ \ \ body (size Ts + 1) | Static \ (\T'. P,Ts \\<^sub>1 body :: T' \ P \ T' \ T) \ \ body \{.. \ \ body (size Ts))" lemma wf_J\<^sub>1_mdecl_NonStatic[simp]: "wf_J\<^sub>1_mdecl P C (M,NonStatic,Ts,T,body) \ (\sub_RI body \ (\T'. P,Class C#Ts \\<^sub>1 body :: T' \ P \ T' \ T) \ \ body \{..size Ts}\ \ \ body (size Ts + 1))" (*<*)by (simp add:wf_J\<^sub>1_mdecl_def)(*>*) lemma wf_J\<^sub>1_mdecl_Static[simp]: "wf_J\<^sub>1_mdecl P C (M,Static,Ts,T,body) \ (\sub_RI body \ (\T'. P,Ts \\<^sub>1 body :: T' \ P \ T' \ T) \ \ body \{.. \ \ body (size Ts))" (*<*)by (simp add:wf_J\<^sub>1_mdecl_def)(*>*) abbreviation "wf_J\<^sub>1_prog == wf_prog wf_J\<^sub>1_mdecl" lemma sees_wf\<^sub>1_nsub_RI: - "\ wf_J\<^sub>1_prog P; P \ C sees M,b : Ts\T = body in D \ \ \sub_RI body" -apply(drule sees_wf_mdecl, simp) -apply(unfold wf_J\<^sub>1_mdecl_def wf_mdecl_def, simp) -done +assumes wf: "wf_J\<^sub>1_prog P" and cM: "P \ C sees M,b : Ts\T = body in D" +shows "\sub_RI body" +using sees_wf_mdecl[OF wf cM] by(simp add: wf_J\<^sub>1_mdecl_def wf_mdecl_def) lemma wf\<^sub>1_types_clinit: assumes wf:"wf_prog wf_md P" and ex: "class P C = Some a" and proc: "sh C = \(sfs, Processing)\" shows "P,E,h,sh \\<^sub>1 C\\<^sub>sclinit([]) : Void" proof - from ex obtain D fs ms where "a = (D,fs,ms)" by(cases a) then have sP: "(C, D, fs, ms) \ set P" using ex map_of_SomeD[of P C a] by(simp add: class_def) then have "wf_clinit ms" using assms by(unfold wf_prog_def wf_cdecl_def, auto) then obtain m where sm: "(clinit, Static, [], Void, m) \ set ms" by(unfold wf_clinit_def) auto then have "P \ C sees clinit,Static:[] \ Void = m in C" using mdecl_visible[OF wf sP sm] by simp then show ?thesis using WTrtSCall\<^sub>1 proc by blast qed lemma assumes wf: "wf_J\<^sub>1_prog P" shows eval\<^sub>1_proc_pres: "P \\<^sub>1 \e,(h,l,sh)\ \ \e',(h',l',sh')\ \ not_init C e \ \sfs. sh C = \(sfs, Processing)\ \ \sfs'. sh' C = \(sfs', Processing)\" and evals\<^sub>1_proc_pres: "P \\<^sub>1 \es,(h,l,sh)\ [\] \es',(h',l',sh')\ \ not_inits C es \ \sfs. sh C = \(sfs, Processing)\ \ \sfs'. sh' C = \(sfs', Processing)\" (*<*) proof(induct rule:eval\<^sub>1_evals\<^sub>1_inducts) case Call\<^sub>1 then show ?case using sees_wf\<^sub>1_nsub_RI[OF wf Call\<^sub>1.hyps(6)] nsub_RI_not_init by auto next case (SCallInit\<^sub>1 ps h l sh vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C' M Ts T body D v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) then show ?case using SCallInit\<^sub>1 sees_wf\<^sub>1_nsub_RI[OF wf SCallInit\<^sub>1.hyps(3)] nsub_RI_not_init[of body] by auto next case SCall\<^sub>1 then show ?case using sees_wf\<^sub>1_nsub_RI[OF wf SCall\<^sub>1.hyps(3)] nsub_RI_not_init by auto next case (InitNone\<^sub>1 sh C1 C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto next case (InitDone\<^sub>1 sh C sfs C' Cs h l e' a a b) then show ?case by(cases Cs, auto) next case (InitProcessing\<^sub>1 sh C sfs C' Cs h l e' a a b) then show ?case by(cases Cs, auto) next case (InitError\<^sub>1 sh C1 sfs Cs h l e' a a b C') then show ?case by(cases "C = C1") auto next case (InitObject\<^sub>1 sh C1 sfs sh' C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto next case (InitNonObject\<^sub>1 sh C1 sfs D a b sh' C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto next case (RInit\<^sub>1 e a a b v h' l' sh' C sfs i sh'' C' Cs e\<^sub>1 a a b) then show ?case by(cases Cs, auto) next case (RInitInitFail\<^sub>1 e h l sh a h' l' sh' C1 sfs i sh'' D Cs e\<^sub>1 h1 l1 sh1) then show ?case using eval\<^sub>1_final by fastforce qed(auto) +(*>*) lemma clinit\<^sub>1_proc_pres: "\ wf_J\<^sub>1_prog P; P \\<^sub>1 \C\<^sub>0\\<^sub>sclinit([]),(h,l,sh)\ \ \e',(h',l',sh')\; sh C' = \(sfs,Processing)\ \ \ \sfs. sh' C' = \(sfs,Processing)\" by(auto dest: eval\<^sub>1_proc_pres) end diff --git a/thys/JinjaDCI/Compiler/PCompiler.thy b/thys/JinjaDCI/Compiler/PCompiler.thy --- a/thys/JinjaDCI/Compiler/PCompiler.thy +++ b/thys/JinjaDCI/Compiler/PCompiler.thy @@ -1,314 +1,393 @@ (* Title: JinjaDCI/Compiler/PCompiler.thy Author: Tobias Nipkow, Susannah Mansky Copyright TUM 2003, UIUC 2019-20 Based on the Jinja theory Common/PCompiler.thy by Tobias Nipkow *) section \ Program Compilation \ theory PCompiler imports "../Common/WellForm" begin definition compM :: "(staticb \ 'a \ 'b) \ 'a mdecl \ 'b mdecl" where "compM f \ \(M, b, Ts, T, m). (M, b, Ts, T, f b m)" definition compC :: "(staticb \ 'a \ 'b) \ 'a cdecl \ 'b cdecl" where "compC f \ \(C,D,Fdecls,Mdecls). (C,D,Fdecls, map (compM f) Mdecls)" definition compP :: "(staticb \ 'a \ 'b) \ 'a prog \ 'b prog" where "compP f \ map (compC f)" text\ Compilation preserves the program structure. Therefore lookup functions either commute with compilation (like method lookup) or are preserved by it (like the subclass relation). \ lemma map_of_map4: "map_of (map (\(x,a,b,c).(x,a,b,f c)) ts) = map_option (\(a,b,c).(a,b,f c)) \ (map_of ts)" (*<*) -apply(induct ts) - apply simp -apply(rule ext) -apply fastforce -done +proof(induct ts) + case Nil then show ?case by simp +qed fastforce (*>*) lemma map_of_map245: "map_of (map (\(x,a,b,c,d).(x,a,b,c,f a c d)) ts) = map_option (\(a,b,c,d).(a,b,c,f a c d)) \ (map_of ts)" (*<*) -apply(induct ts) - apply simp -apply(rule ext) -apply fastforce -done +proof(induct ts) + case Nil then show ?case by simp +qed fastforce (*>*) lemma class_compP: "class P C = Some (D, fs, ms) \ class (compP f P) C = Some (D, fs, map (compM f) ms)" (*<*)by(simp add:class_def compP_def compC_def map_of_map4)(*>*) lemma class_compPD: "class (compP f P) C = Some (D, fs, cms) \ \ms. class P C = Some(D,fs,ms) \ cms = map (compM f) ms" (*<*)by(clarsimp simp add:class_def compP_def compC_def map_of_map4)(*>*) lemma [simp]: "is_class (compP f P) C = is_class P C" (*<*)by(auto simp:is_class_def dest: class_compP class_compPD)(*>*) lemma [simp]: "class (compP f P) C = map_option (\c. snd(compC f (C,c))) (class P C)" (*<*) -apply(simp add:compP_def compC_def class_def map_of_map4) -apply(simp add:split_def) -done +by(simp add:compP_def compC_def class_def map_of_map4) + (simp add:split_def) (*>*) lemma sees_methods_compP: "P \ C sees_methods Mm \ compP f P \ C sees_methods (map_option (\((b,Ts,T,m),D). ((b,Ts,T,f b m),D)) \ Mm)" -(*<*) -apply(erule Methods.induct) - apply(rule sees_methods_Object) - apply(erule class_compP) - apply(rule ext) - apply(simp add:compM_def map_of_map245 option.map_comp) - apply(case_tac "map_of ms x") - apply simp - apply fastforce -apply(rule sees_methods_rec) - apply(erule class_compP) - apply assumption - apply assumption -apply(rule ext) -apply(simp add:map_add_def compM_def map_of_map245 option.map_comp split:option.split) -done +(*<*)(is "?P \ compP f P \ C sees_methods (?map Mm)") +proof(induct rule: Methods.induct) + case Object: (sees_methods_Object D fs ms Mm) + let ?Mm1 = "\x. map_option ((\m. (m, Object)) \ (\(b, Ts, T, m). (b, Ts, T, f b m))) (map_of ms x)" + let ?Mm2 = "\x. map_option (case_prod (\(b, Ts, T, m). + Pair (b, Ts, T, f b m)) \ (\m. (m, Object))) (map_of ms x)" + have Mm_eq: "\x. ?Mm1 x = ?Mm2 x" + proof - + fix x show "?Mm1 x = ?Mm2 x" + proof(cases "map_of ms x") + case None then show ?thesis by simp + qed fastforce + qed + + have Mm: "Mm = map_option (\m. (m, Object)) \ map_of ms" by fact + let ?Mm = "map_option (\m. (m, Object)) \ map_of (map (compM f) ms)" + let ?Mm' = "?map Mm" + have "?Mm' = ?Mm" + by(rule ext) (simp add:Mm Mm_eq compM_def map_of_map245 option.map_comp) + then show ?case by(rule sees_methods_Object[OF class_compP[OF Object(1)]]) +next + case rec: (sees_methods_rec C D fs ms Mm Mm') + have Mm': "Mm' = Mm ++ (map_option (\m. (m, C)) \ map_of ms)" by fact + let ?Mm' = "?map Mm'" + let ?Mm'' = "(?map Mm) ++ (map_option (\m. (m, C)) \ map_of (map (compM f) ms))" + have "?Mm' = ?Mm''" + by(rule ext) (simp add:Mm' map_add_def compM_def map_of_map245) + moreover have "compP f P \ C sees_methods ?Mm''" + using sees_methods_rec[OF class_compP[OF rec(1)] rec(2,4)] by fast + ultimately show "compP f P \ C sees_methods ?Mm'" by simp +qed (*>*) lemma sees_method_compP: "P \ C sees M,b: Ts\T = m in D \ compP f P \ C sees M,b: Ts\T = (f b m) in D" (*<*)by(fastforce elim:sees_methods_compP simp add:Method_def)(*>*) lemma [simp]: "P \ C sees M,b: Ts\T = m in D \ method (compP f P) C M = (D,b,Ts,T,f b m)" (*<*) -apply(drule sees_method_compP) -apply(simp add:method_def) -apply(rule the_equality) - apply simp -apply(fastforce dest:sees_method_fun) -done +proof - + let ?P = "\(D, b, Ts, T, m). compP f P \ C sees M, b : Ts\T = m in D" + let ?a = "(D, b, Ts, T, f b m)" + assume cM: "P \ C sees M,b: Ts\T = m in D" + have compP_cM: "?P ?a" using sees_method_compP[OF cM] by simp + moreover { + fix x assume "?P x" then have "x = ?a" + using compP_cM by(fastforce dest:sees_method_fun) + } + ultimately have "(THE x. ?P x) = ?a" by(rule the_equality) + then show ?thesis by(simp add:method_def) +qed (*>*) lemma sees_methods_compPD: "\ cP \ C sees_methods Mm'; cP = compP f P \ \ \Mm. P \ C sees_methods Mm \ Mm' = (map_option (\((b,Ts,T,m),D). ((b,Ts,T,f b m),D)) \ Mm)" -(*<*) -apply(erule Methods.induct) - apply(clarsimp simp:compC_def) - apply(rule exI) - apply(rule conjI, erule sees_methods_Object) - apply(rule refl) - apply(rule ext) - apply(simp add:compM_def map_of_map245 option.map_comp) - apply(case_tac "map_of b x") - apply simp - apply fastforce -apply(clarsimp simp:compC_def) -apply(rule exI, rule conjI) - apply(erule (2) sees_methods_rec) - apply(rule refl) -apply(rule ext) -apply(simp add:map_add_def compM_def map_of_map245 option.map_comp split:option.split) -done +(*<*)(is "\ ?P; ?Q \ \ \Mm. P \ C sees_methods Mm \ Mm' = (?map Mm)") +proof(induct rule: Methods.induct) + case Object: (sees_methods_Object D fs ms Mm) + then obtain ms' where P_Obj: "class P Object = \(D, fs, ms')\" + and ms: "ms = map (compM f) ms'" by(clarsimp simp:compC_def) + + let ?Mm1 = "\x. map_option ((\m. (m, Object)) \ (\(b, Ts, T, m). (b, Ts, T, f b m))) (map_of ms' x)" + let ?Mm2 = "\x. map_option (case_prod (\(b, Ts, T, m). Pair (b, Ts, T, f b m)) \ (\m. (m, Object))) + (map_of ms' x)" + have Mm_eq: "\x. ?Mm1 x = ?Mm2 x" + proof - + fix x show "?Mm1 x = ?Mm2 x" + proof(cases "map_of ms' x") + case None then show ?thesis by simp + qed fastforce + qed + + let ?Mm = "map_option (\m. (m,Object)) \ map_of ms'" + let ?Mm' = "?map ?Mm" + have Mm: "Mm = map_option (\m. (m, Object)) \ map_of ms" by fact + have "P \ Object sees_methods ?Mm" + using sees_methods_Object[OF P_Obj] by simp + moreover have "Mm = ?Mm'" + by(rule ext) (simp add:Mm_eq Mm ms compM_def map_of_map245 option.map_comp) + ultimately show ?case by fast +next + case rec: (sees_methods_rec C D fs ms Mm Mm') + then obtain ms' Mm\<^sub>D where P_D: "class P C = \(D, fs, ms')\" + and ms: "ms = map (compM f) ms'" and C_nObj: "C \ Object" + and Mm\<^sub>D: "P \ D sees_methods Mm\<^sub>D" + and Mm: "Mm = (\a. map_option (case_prod (\(b, Ts, T, m). Pair (b, Ts, T, f b m))) (Mm\<^sub>D a))" + by(clarsimp simp:compC_def) + + let ?Mm = "Mm\<^sub>D ++ (map_option (\m. (m, C)) \ map_of ms')" + let ?Mm1 = "Mm ++ (map_option (\m. (m, C)) \ map_of ms)" + let ?Mm2 = "Mm ++ (map_option (\m. (m, C)) \ map_of (map (compM f) ms'))" + let ?Mm3 = "?map ?Mm" + have "Mm' = ?Mm1" by fact + also have "\ = ?Mm2" using ms by simp + also have "\ = ?Mm3" + by(rule ext)(simp add:Mm map_add_def compM_def map_of_map245) + moreover have "P \ C sees_methods ?Mm" + using sees_methods_rec[OF P_D C_nObj Mm\<^sub>D] by simp + ultimately show ?case by fast +qed (*>*) lemma sees_method_compPD: "compP f P \ C sees M,b: Ts\T = fm in D \ \m. P \ C sees M,b: Ts\T = m in D \ f b m = fm" (*<*) -apply(simp add:Method_def) -apply clarify -apply(drule sees_methods_compPD[OF _ refl]) -apply clarsimp -apply blast -done +proof - + assume "compP f P \ C sees M,b: Ts\T = fm in D" + then obtain Mm where Mm: "compP f P \ C sees_methods Mm" + and MmM: "Mm M = \((b, Ts, T, fm), D)\" + by(clarsimp simp:Method_def) + show ?thesis using sees_methods_compPD[OF Mm refl] MmM + by(fastforce simp: Method_def) +qed (*>*) lemma [simp]: "subcls1(compP f P) = subcls1 P" (*<*) by(fastforce simp add: is_class_def compC_def intro:subcls1I order_antisym dest:subcls1D) (*>*) lemma compP_widen[simp]: "(compP f P \ T \ T') = (P \ T \ T')" (*<*)by(cases T')(simp_all add:widen_Class)(*>*) lemma [simp]: "(compP f P \ Ts [\] Ts') = (P \ Ts [\] Ts')" (*<*) -apply(induct Ts) - apply simp -apply(cases Ts') - apply(auto simp:fun_of_def) -done +proof(induct Ts) + case (Cons a Ts) + then show ?case by(cases Ts')(auto simp:fun_of_def) +qed simp (*>*) lemma [simp]: "is_type (compP f P) T = is_type P T" (*<*)by(cases T) simp_all(*>*) lemma [simp]: "(compP (f::staticb\'a\'b) P \ C has_fields FDTs) = (P \ C has_fields FDTs)" (*<*) (is "?A = ?B") proof { fix cP::"'b prog" assume "cP \ C has_fields FDTs" hence "cP = compP f P \ P \ C has_fields FDTs" proof induct case has_fields_Object thus ?case by(fast intro:Fields.has_fields_Object dest:class_compPD) next case has_fields_rec thus ?case by(fast intro:Fields.has_fields_rec dest:class_compPD) qed } note lem = this assume ?A with lem show ?B by blast next assume ?B thus ?A proof induct case has_fields_Object thus ?case by(fast intro:Fields.has_fields_Object class_compP) next case has_fields_rec thus ?case by(fast intro:Fields.has_fields_rec class_compP) qed qed (*>*) lemma fields_compP [simp]: "fields (compP f P) C = fields P C" (*<*)by(simp add:fields_def)(*>*) lemma ifields_compP [simp]: "ifields (compP f P) C = ifields P C" (*<*)by(simp add:ifields_def)(*>*) lemma blank_compP [simp]: "blank (compP f P) C = blank P C" (*<*)by(simp add:blank_def)(*>*) lemma isfields_compP [simp]: "isfields (compP f P) C = isfields P C" (*<*)by(simp add:isfields_def)(*>*) lemma sblank_compP [simp]: "sblank (compP f P) C = sblank P C" (*<*)by(simp add:sblank_def)(*>*) lemma sees_fields_compP [simp]: "(compP f P \ C sees F,b:T in D) = (P \ C sees F,b:T in D)" (*<*)by(simp add:sees_field_def)(*>*) lemma has_field_compP [simp]: "(compP f P \ C has F,b:T in D) = (P \ C has F,b:T in D)" (*<*)by(simp add:has_field_def)(*>*) lemma field_compP [simp]: "field (compP f P) F D = field P F D" (*<*)by(simp add:field_def)(*>*) subsection\Invariance of @{term wf_prog} under compilation \ lemma [iff]: "distinct_fst (compP f P) = distinct_fst P" (*<*) -apply(simp add:distinct_fst_def compP_def compC_def) -apply(induct P) -apply (auto simp:image_iff) -done +by (induct P) + (auto simp:distinct_fst_def compP_def compC_def image_iff) (*>*) lemma [iff]: "distinct_fst (map (compM f) ms) = distinct_fst ms" (*<*) -apply(simp add:distinct_fst_def compM_def) -apply(induct ms) -apply (auto simp:image_iff) -done +by (induct ms) + (auto simp:distinct_fst_def compM_def image_iff) (*>*) lemma [iff]: "wf_syscls (compP f P) = wf_syscls P" (*<*)by(simp add:wf_syscls_def compP_def compC_def image_def Bex_def)(*>*) lemma [iff]: "wf_fdecl (compP f P) = wf_fdecl P" (*<*)by(simp add:wf_fdecl_def)(*>*) lemma wf_clinit_compM [iff]: "wf_clinit (map (compM f) ms) = wf_clinit ms" (*<*) -apply(simp add: wf_clinit_def compM_def) -apply(rule iffI) - apply clarsimp apply(rename_tac m) - apply(rule_tac x = m in exI, simp+) -apply clarsimp apply(rename_tac m) -apply(rule_tac x = "f Static m" in exI, simp add: rev_image_eqI) -done +proof(rule iffI) + assume "wf_clinit (map (compM f) ms)" + then obtain m where "(clinit, Static, [], Void, m) \ set ms" + by(clarsimp simp: wf_clinit_def compM_def) + then show "wf_clinit ms" by(fastforce simp: wf_clinit_def) +next + assume "wf_clinit ms" + then obtain m where "(clinit, Static, [], Void, m) \ set ms" + by(clarsimp simp: wf_clinit_def compM_def) + then have "\m. (clinit, Static, [], Void, m) + \ (\x. case x of (M, b, Ts, T, m) \ (M, b, Ts, T, f b m)) ` set ms" + by(rule_tac x = "f Static m" in exI) (simp add: rev_image_eqI) + then show "wf_clinit (map (compM f) ms)" + by(simp add: wf_clinit_def compM_def) +qed (*>*) lemma set_compP: "((C,D,fs,ms') \ set(compP f P)) = (\ms. (C,D,fs,ms) \ set P \ ms' = map (compM f) ms)" (*<*)by(fastforce simp add:compP_def compC_def image_iff Bex_def)(*>*) lemma wf_cdecl_compPI: "\ \C M b Ts T m. \ wf_mdecl wf\<^sub>1 P C (M,b,Ts,T,m); P \ C sees M,b:Ts\T = m in C \ \ wf_mdecl wf\<^sub>2 (compP f P) C (M,b,Ts,T, f b m); \x\set P. wf_cdecl wf\<^sub>1 P x; x \ set (compP f P); wf_prog p P \ \ wf_cdecl wf\<^sub>2 (compP f P) x" (*<*) -apply(clarsimp simp add:wf_cdecl_def Ball_def set_compP) -apply(rename_tac C D fs ms) -apply(rule conjI) - apply (clarsimp simp:compM_def) - apply (drule (2) mdecl_visible) - apply simp -apply(clarify) -apply(drule sees_method_compPD[where f = f]) -apply clarsimp -apply(fastforce simp:image_iff compM_def) -done +proof - + assume + wfm: "\C M b Ts T m. \ wf_mdecl wf\<^sub>1 P C (M,b,Ts,T,m); P \ C sees M,b:Ts\T = m in C \ + \ wf_mdecl wf\<^sub>2 (compP f P) C (M,b,Ts,T, f b m)" + and wfc: "\x\set P. wf_cdecl wf\<^sub>1 P x" + and compP: "x \ set (compP f P)" and wf: "wf_prog p P" + obtain C D fs ms where x: "x = (C, D, fs, map (compM f) ms)" + and x_set: "(C, D, fs, ms) \ set P" + using compP by(case_tac x) (clarsimp simp: set_compP) + have wfc': "wf_cdecl wf\<^sub>1 P (C, D, fs, ms)" using wfc x_set by fast + let ?P = "compP f P" and ?ms = "compM f ` set ms" + { fix M b Ts T m + assume M: "(M,b,Ts,T,m) \ set ms" + then have "wf_mdecl wf\<^sub>1 P C (M, b, Ts, T, m)" using wfc' + by(simp add:wf_cdecl_def) + moreover have cM: "P \ C sees M, b : Ts\T = m in C" using M + by(rule mdecl_visible[OF wf x_set]) + ultimately have "wf_mdecl wf\<^sub>2 (compP f P) C (M, b, Ts, T, f b m)" + by(rule wfm) + } + then have "\m \ ?ms. wf_mdecl wf\<^sub>2 ?P C m" + by (clarsimp simp:compM_def) + moreover have "C \ Object \ + (\(M,b,Ts,T,m)\?ms. + \D' b' Ts' T' m'. ?P \ D sees M,b':Ts' \ T' = m' in D' \ + b = b' \ P \ Ts' [\] Ts \ P \ T \ T')" + proof - + { fix M b Ts T m D' b' Ts' T' m' + assume "C \ Object" and "(M,b,Ts,T,m)\?ms" + and dM: "?P \ D sees M,b':Ts' \ T' = m' in D'" + then have "b = b' \ P \ Ts' [\] Ts \ P \ T \ T'" + using wfc' sees_method_compPD[OF dM] + by(fastforce simp:wf_cdecl_def image_iff compM_def) + } + then show ?thesis by fast + qed + moreover have "(\f\set fs. wf_fdecl P f) \ distinct_fst fs + \ distinct_fst ms \ wf_clinit ms + \ (C \ Object \ is_class P D \ \ P \ D \\<^sup>* C)" using wfc' + by(simp add: wf_cdecl_def) + ultimately show ?thesis using x by(simp add:wf_cdecl_def) +qed (*>*) lemma wf_prog_compPI: assumes lift: "\C M b Ts T m. \ P \ C sees M,b:Ts\T = m in C; wf_mdecl wf\<^sub>1 P C (M,b,Ts,T,m) \ \ wf_mdecl wf\<^sub>2 (compP f P) C (M,b,Ts,T, f b m)" and wf: "wf_prog wf\<^sub>1 P" shows "wf_prog wf\<^sub>2 (compP f P)" (*<*) using wf by (simp add:wf_prog_def) (blast intro:wf_cdecl_compPI lift wf) (*>*) end