diff --git a/thys/Jinja/J/TypeSafe.thy b/thys/Jinja/J/TypeSafe.thy --- a/thys/Jinja/J/TypeSafe.thy +++ b/thys/Jinja/J/TypeSafe.thy @@ -1,670 +1,719 @@ (* Title: Jinja/J/SmallTypeSafe.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Type Safety Proof\ theory TypeSafe imports Progress JWellForm begin subsection\Basic preservation lemmas\ text\First two easy preservation lemmas.\ theorem red_preserves_hconf: "P \ \e,(h,l)\ \ \e',(h',l')\ \ (\T E. \ P,E,h \ e : T; P \ h \ \ \ P \ h' \)" and reds_preserves_hconf: "P \ \es,(h,l)\ [\] \es',(h',l')\ \ (\Ts E. \ P,E,h \ es [:] Ts; P \ h \ \ \ P \ h' \)" (*<*) proof (induct rule:red_reds_inducts) case (RedNew h a C FDTs h' l) have new: "new_Addr h = Some a" and fields: "P \ C has_fields FDTs" and h': "h' = h(a\(C, init_fields FDTs))" and hconf: "P \ h \" by fact+ from new have None: "h a = None" by(rule new_Addr_SomeD) moreover have "P,h \ (C,init_fields FDTs) \" using fields by(rule oconf_init_fields) ultimately show "P \ h' \" using h' by(fast intro: hconf_new[OF hconf]) next case (RedFAss h a C fs F D v l) let ?fs' = "fs((F,D)\v)" have hconf: "P \ h \" and ha: "h a = Some(C,fs)" and wt: "P,E,h \ addr a\F{D}:=Val v : T" by fact+ from wt ha obtain TF Tv where typeofv: "typeof\<^bsub>h\<^esub> v = Some Tv" and has: "P \ C has F:TF in D" and sub: "P \ Tv \ TF" by auto have "P,h \ (C, ?fs') \" proof (rule oconf_fupd[OF has]) show "P,h \ (C, fs) \" using hconf ha by(simp add:hconf_def) show "P,h \ v :\ TF" using sub typeofv by(simp add:conf_def) qed with hconf ha show "P \ h(a\(C, ?fs')) \" by (rule hconf_upd_obj) qed auto (*>*) theorem red_preserves_lconf: "P \ \e,(h,l)\ \ \e',(h',l')\ \ (\T E. \ P,E,h \ e:T; P,h \ l (:\) E \ \ P,h' \ l' (:\) E)" and reds_preserves_lconf: "P \ \es,(h,l)\ [\] \es',(h',l')\ \ (\Ts E. \ P,E,h \ es[:]Ts; P,h \ l (:\) E \ \ P,h' \ l' (:\) E)" (*<*) proof(induct rule:red_reds_inducts) case RedNew thus ?case by(fast intro:lconf_hext red_hext_incr[OF red_reds.RedNew]) next case RedLAss thus ?case by(fastforce elim: lconf_upd simp:conf_def) next case RedFAss thus ?case by(fast intro:lconf_hext red_hext_incr[OF red_reds.RedFAss]) next case (InitBlockRed e h l V v e' h' l' v' T T') have red: "P \ \e, (h, l(V\v))\ \ \e',(h', l')\" and IH: "\T E . \ P,E,h \ e:T; P,h \ l(V\v) (:\) E \ \ P,h' \ l' (:\) E" and l'V: "l' V = Some v'" and lconf: "P,h \ l (:\) E" and wt: "P,E,h \ {V:T := Val v; e} : T'" by fact+ from lconf_hext[OF lconf red_hext_incr[OF red]] have "P,h' \ l (:\) E" . moreover from IH lconf wt have "P,h' \ l' (:\) E(V\T)" by(auto simp del: fun_upd_apply simp: fun_upd_same lconf_upd2 conf_def) ultimately show "P,h' \ l'(V := l V) (:\) E" by (fastforce simp:lconf_def split:if_split_asm) next case (BlockRedNone e h l V e' h' l' T T') have red: "P \ \e,(h, l(V := None))\ \ \e',(h', l')\" and IH: "\E T. \ P,E,h \ e : T; P,h \ l(V:=None) (:\) E \ \ P,h' \ l' (:\) E" and lconf: "P,h \ l (:\) E" and wt: "P,E,h \ {V:T; e} : T'" by fact+ from lconf_hext[OF lconf red_hext_incr[OF red]] have "P,h' \ l (:\) E" . moreover have "P,h' \ l' (:\) E(V\T)" by(rule IH, insert lconf wt, auto simp:lconf_def) ultimately show "P,h' \ l'(V := l V) (:\) E" by (fastforce simp:lconf_def split:if_split_asm) next case (BlockRedSome e h l V e' h' l' v T T') have red: "P \ \e,(h, l(V := None))\ \ \e',(h', l')\" and IH: "\E T. \P,E,h \ e : T; P,h \ l(V:=None) (:\) E\ \ P,h' \ l' (:\) E" and lconf: "P,h \ l (:\) E" and wt: "P,E,h \ {V:T; e} : T'" by fact+ from lconf_hext[OF lconf red_hext_incr[OF red]] have "P,h' \ l (:\) E" . moreover have "P,h' \ l' (:\) E(V\T)" by(rule IH, insert lconf wt, auto simp:lconf_def) ultimately show "P,h' \ l'(V := l V) (:\) E" by (fastforce simp:lconf_def split:if_split_asm) qed auto (*>*) text\Preservation of definite assignment more complex and requires a few lemmas first.\ lemma [iff]: "\A. \ length Vs = length Ts; length vs = length Ts\ \ \ (blocks (Vs,Ts,vs,e)) A = \ e (A \ \set Vs\)" (*<*) -apply(induct Vs Ts vs e rule:blocks_induct) -apply(simp_all add:hyperset_defs) -done +by (induct Vs Ts vs e rule:blocks_induct) + (simp_all add:hyperset_defs) (*>*) lemma red_lA_incr: "P \ \e,(h,l)\ \ \e',(h',l')\ \ \dom l\ \ \ e \ \dom l'\ \ \ e'" and reds_lA_incr: "P \ \es,(h,l)\ [\] \es',(h',l')\ \ \dom l\ \ \s es \ \dom l'\ \ \s es'" (*<*) -apply(induct rule:red_reds_inducts) -apply(simp_all del:fun_upd_apply add:hyperset_defs) -apply auto -apply(blast dest:red_lcl_incr)+ -done +proof(induct rule:red_reds_inducts) + case TryRed then show ?case + by(simp del:fun_upd_apply add:hyperset_defs) + (blast dest:red_lcl_incr)+ +next + case BinOpRed1 then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case BinOpRed2 then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case LAssRed then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case FAssRed1 then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case FAssRed2 then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case CallObj then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case CallParams then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case BlockRedNone then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case BlockRedSome then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case InitBlockRed then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case SeqRed then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case CondRed then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case RedCondT then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case RedCondF then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case ListRed1 then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case ListRed2 then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +qed (simp_all del:fun_upd_apply add:hyperset_defs) (*>*) text\Now preservation of definite assignment.\ lemma assumes wf: "wf_J_prog P" shows red_preserves_defass: "P \ \e,(h,l)\ \ \e',(h',l')\ \ \ e \dom l\ \ \ e' \dom l'\" and "P \ \es,(h,l)\ [\] \es',(h',l')\ \ \s es \dom l\ \ \s es' \dom l'\" (*<*) proof (induct rule:red_reds_inducts) case BinOpRed1 thus ?case by (auto elim!: D_mono[OF red_lA_incr]) next case FAssRed1 thus ?case by (auto elim!: D_mono[OF red_lA_incr]) next case CallObj thus ?case by (auto elim!: Ds_mono[OF red_lA_incr]) next case RedCall thus ?case - apply (auto dest!:sees_wf_mdecl[OF wf] simp:wf_mdecl_def elim!:D_mono') - by(auto simp:hyperset_defs) + by (auto dest!:sees_wf_mdecl[OF wf] simp:wf_mdecl_def hyperset_defs elim!:D_mono') next case InitBlockRed thus ?case by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply) next case BlockRedNone thus ?case by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply) next case BlockRedSome thus ?case by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply) next case SeqRed thus ?case by (auto elim!: D_mono[OF red_lA_incr]) next case CondRed thus ?case by (auto elim!: D_mono[OF red_lA_incr]) next case TryRed thus ?case by (fastforce dest:red_lcl_incr intro:D_mono' simp:hyperset_defs) next case ListRed1 thus ?case by (auto elim!: Ds_mono[OF red_lA_incr]) next case RedWhile thus ?case by(auto simp:hyperset_defs elim!:D_mono') qed (auto simp:hyperset_defs) (*>*) text\Combining conformance of heap and local variables:\ definition sconf :: "J_prog \ env \ state \ bool" ("_,_ \ _ \" [51,51,51]50) where "P,E \ s \ \ let (h,l) = s in P \ h \ \ P,h \ l (:\) E" lemma red_preserves_sconf: "\ P \ \e,s\ \ \e',s'\; P,E,hp s \ e : T; P,E \ s \ \ \ P,E \ s' \" (*<*) by(fastforce intro:red_preserves_hconf red_preserves_lconf simp add:sconf_def) (*>*) lemma reds_preserves_sconf: "\ P \ \es,s\ [\] \es',s'\; P,E,hp s \ es [:] Ts; P,E \ s \ \ \ P,E \ s' \" (*<*) by(fastforce intro:reds_preserves_hconf reds_preserves_lconf simp add:sconf_def) (*>*) subsection "Subject reduction" lemma wt_blocks: "\E. \ length Vs = length Ts; length vs = length Ts \ \ (P,E,h \ blocks(Vs,Ts,vs,e) : T) = (P,E(Vs[\]Ts),h \ e:T \ (\Ts'. map (typeof\<^bsub>h\<^esub>) vs = map Some Ts' \ P \ Ts' [\] Ts))" (*<*) -apply(induct Vs Ts vs e rule:blocks_induct) -apply (force simp add:rel_list_all2_Cons2) -apply simp_all -done +proof(induct Vs Ts vs e rule:blocks_induct) + case (1 V Vs T Ts v vs e) + then show ?case by(force simp add:rel_list_all2_Cons2) +qed simp_all (*>*) theorem assumes wf: "wf_J_prog P" shows subject_reduction2: "P \ \e,(h,l)\ \ \e',(h',l')\ \ (\E T. \ P,E \ (h,l) \; P,E,h \ e:T \ \ \T'. P,E,h' \ e':T' \ P \ T' \ T)" and subjects_reduction2: "P \ \es,(h,l)\ [\] \es',(h',l')\ \ (\E Ts. \ P,E \ (h,l) \; P,E,h \ es [:] Ts \ \ \Ts'. P,E,h' \ es' [:] Ts' \ P \ Ts' [\] Ts)" (*<*) proof (induct rule:red_reds_inducts) case (RedCall h l a C fs M Ts T pns body D vs E T') have hp: "hp(h,l) a = Some(C,fs)" and "method": "P \ C sees M: Ts\T = (pns,body) in D" and wt: "P,E,h \ addr a\M(map Val vs) : T'" by fact+ obtain Ts' where wtes: "P,E,h \ map Val vs [:] Ts'" and subs: "P \ Ts' [\] Ts" and T'isT: "T' = T" using wt "method" hp by (auto dest:sees_method_fun) from wtes subs have length_vs: "length vs = length Ts" by(fastforce simp:list_all2_iff dest!:WTrts_same_length) from sees_wf_mdecl[OF wf "method"] obtain T'' where wtabody: "P,[this#pns [\] Class D#Ts] \ body :: T''" and T''subT: "P \ T'' \ T" and length_pns: "length pns = length Ts" by(fastforce simp:wf_mdecl_def simp del:map_upds_twist) from wtabody have "P,Map.empty(this#pns [\] Class D#Ts),h \ body : T''" by(rule WT_implies_WTrt) hence "P,E(this#pns [\] Class D#Ts),h \ body : T''" by(rule WTrt_env_mono) simp hence "P,E,h \ blocks(this#pns, Class D#Ts, Addr a#vs, body) : T''" using wtes subs hp sees_method_decl_above[OF "method"] length_vs length_pns by(fastforce simp add:wt_blocks rel_list_all2_Cons2) with T''subT T'isT show ?case by blast next case RedNewFail thus ?case by (unfold sconf_def hconf_def) (fastforce elim!:typeof_OutOfMemory) next case CastRed thus ?case by(clarsimp simp:is_refT_def) (blast intro: widens_trans dest!:widen_Class[THEN iffD1]) next case RedCastFail thus ?case by (unfold sconf_def hconf_def) (fastforce elim!:typeof_ClassCast) next case (BinOpRed1 e\<^sub>1 h l e\<^sub>1' h' l' bop e\<^sub>2) have red: "P \ \e\<^sub>1,(h,l)\ \ \e\<^sub>1',(h',l')\" and IH: "\E T. \P,E \ (h,l) \; P,E,h \ e\<^sub>1:T\ \ \U. P,E,h' \ e\<^sub>1' : U \ P \ U \ T" and conf: "P,E \ (h,l) \" and wt: "P,E,h \ e\<^sub>1 \bop\ e\<^sub>2 : T" by fact+ have "P,E,h' \ e\<^sub>1' \bop\ e\<^sub>2 : T" proof (cases bop) assume [simp]: "bop = Eq" from wt obtain T\<^sub>1 T\<^sub>2 where [simp]: "T = Boolean" and wt\<^sub>1: "P,E,h \ e\<^sub>1 : T\<^sub>1" and wt\<^sub>2: "P,E,h \ e\<^sub>2 : T\<^sub>2" by auto show ?thesis using WTrt_hext_mono[OF wt\<^sub>2 red_hext_incr[OF red]] IH[OF conf wt\<^sub>1] by auto next assume [simp]: "bop = Add" from wt have [simp]: "T = Integer" and wt\<^sub>1: "P,E,h \ e\<^sub>1 : Integer" and wt\<^sub>2: "P,E,h \ e\<^sub>2 : Integer" by auto show ?thesis using IH[OF conf wt\<^sub>1] WTrt_hext_mono[OF wt\<^sub>2 red_hext_incr[OF red]] by auto qed thus ?case by auto next case (BinOpRed2 e\<^sub>2 h l e\<^sub>2' h' l' v\<^sub>1 bop) have red: "P \ \e\<^sub>2,(h,l)\ \ \e\<^sub>2',(h',l')\" and IH: "\E T. \P,E \ (h,l) \; P,E,h \ e\<^sub>2:T\ \ \U. P,E,h' \ e\<^sub>2' : U \ P \ U \ T" and conf: "P,E \ (h,l) \" and wt: "P,E,h \ (Val v\<^sub>1) \bop\ e\<^sub>2 : T" by fact+ have "P,E,h' \ (Val v\<^sub>1) \bop\ e\<^sub>2' : T" proof (cases bop) assume [simp]: "bop = Eq" from wt obtain T\<^sub>1 T\<^sub>2 where [simp]: "T = Boolean" and wt\<^sub>1: "P,E,h \ Val v\<^sub>1 : T\<^sub>1" and wt\<^sub>2: "P,E,h \ e\<^sub>2:T\<^sub>2" by auto show ?thesis using IH[OF conf wt\<^sub>2] WTrt_hext_mono[OF wt\<^sub>1 red_hext_incr[OF red]] by auto next assume [simp]: "bop = Add" from wt have [simp]: "T = Integer" and wt\<^sub>1: "P,E,h \ Val v\<^sub>1 : Integer" and wt\<^sub>2: "P,E,h \ e\<^sub>2 : Integer" by auto show ?thesis using IH[OF conf wt\<^sub>2] WTrt_hext_mono[OF wt\<^sub>1 red_hext_incr[OF red]] by auto qed thus ?case by auto next case (RedBinOp bop) thus ?case proof (cases bop) case Eq thus ?thesis using RedBinOp by auto next case Add thus ?thesis using RedBinOp by auto qed next case RedVar thus ?case by (fastforce simp:sconf_def lconf_def conf_def) next case LAssRed thus ?case by(blast intro:widen_trans) next case (FAccRed e h l e' h' l' F D) have IH: "\E T. \P,E \ (h,l) \; P,E,h \ e : T\ \ \U. P,E,h' \ e' : U \ P \ U \ T" and conf: "P,E \ (h,l) \" and wt: "P,E,h \ e\F{D} : T" by fact+ \ \The goal: ?case = @{prop ?case}\ \ \Now distinguish the two cases how wt can have arisen.\ { fix C assume wte: "P,E,h \ e : Class C" and has: "P \ C has F:T in D" from IH[OF conf wte] obtain U where wte': "P,E,h' \ e' : U" and UsubC: "P \ U \ Class C" by auto \ \Now distinguish what @{term U} can be.\ { assume "U = NT" hence ?case using wte' by(blast intro:WTrtFAccNT widen_refl) } moreover { fix C' assume U: "U = Class C'" and C'subC: "P \ C' \\<^sup>* C" from has_field_mono[OF has C'subC] wte' U have ?case by(blast intro:WTrtFAcc) } ultimately have ?case using UsubC by(simp add: widen_Class) blast } moreover { assume "P,E,h \ e : NT" hence "P,E,h' \ e' : NT" using IH[OF conf] by fastforce hence ?case by(fastforce intro:WTrtFAccNT widen_refl) } ultimately show ?case using wt by blast next case RedFAcc thus ?case by(fastforce simp:sconf_def hconf_def oconf_def conf_def has_field_def dest:has_fields_fun) next case RedFAccNull thus ?case by(fastforce intro: widen_refl WTThrow[OF WTVal] elim!: typeof_NullPointer simp: sconf_def hconf_def) next case (FAssRed1 e h l e' h' l' F D e\<^sub>2) have red: "P \ \e,(h,l)\ \ \e',(h',l')\" and IH: "\E T. \P,E \ (h,l) \; P,E,h \ e : T\ \ \U. P,E,h' \ e' : U \ P \ U \ T" and conf: "P,E \ (h,l) \" and wt: "P,E,h \ e\F{D}:=e\<^sub>2 : T" by fact+ from wt have void: "T = Void" by blast \ \We distinguish if @{term e} has type @{term NT} or a Class type\ \ \Remember ?case = @{term ?case}\ { assume "P,E,h \ e : NT" hence "P,E,h' \ e' : NT" using IH[OF conf] by fastforce moreover obtain T\<^sub>2 where "P,E,h \ e\<^sub>2 : T\<^sub>2" using wt by auto from this red_hext_incr[OF red] have "P,E,h' \ e\<^sub>2 : T\<^sub>2" by(rule WTrt_hext_mono) ultimately have ?case using void by(blast intro!:WTrtFAssNT) } moreover { fix C TF T\<^sub>2 assume wt\<^sub>1: "P,E,h \ e : Class C" and wt\<^sub>2: "P,E,h \ e\<^sub>2 : T\<^sub>2" and has: "P \ C has F:TF in D" and sub: "P \ T\<^sub>2 \ TF" obtain U where wt\<^sub>1': "P,E,h' \ e' : U" and UsubC: "P \ U \ Class C" using IH[OF conf wt\<^sub>1] by blast have wt\<^sub>2': "P,E,h' \ e\<^sub>2 : T\<^sub>2" by(rule WTrt_hext_mono[OF wt\<^sub>2 red_hext_incr[OF red]]) \ \Is @{term U} the null type or a class type?\ { assume "U = NT" with wt\<^sub>1' wt\<^sub>2' void have ?case by(blast intro!:WTrtFAssNT) } moreover { fix C' assume UClass: "U = Class C'" and "subclass": "P \ C' \\<^sup>* C" have "P,E,h' \ e' : Class C'" using wt\<^sub>1' UClass by auto moreover have "P \ C' has F:TF in D" by(rule has_field_mono[OF has "subclass"]) ultimately have ?case using wt\<^sub>2' sub void by(blast intro:WTrtFAss) } ultimately have ?case using UsubC by(auto simp add:widen_Class) } ultimately show ?case using wt by blast next case (FAssRed2 e\<^sub>2 h l e\<^sub>2' h' l' v F D) have red: "P \ \e\<^sub>2,(h,l)\ \ \e\<^sub>2',(h',l')\" and IH: "\E T. \P,E \ (h,l) \; P,E,h \ e\<^sub>2 : T\ \ \U. P,E,h' \ e\<^sub>2' : U \ P \ U \ T" and conf: "P,E \ (h,l) \" and wt: "P,E,h \ Val v\F{D}:=e\<^sub>2 : T" by fact+ from wt have [simp]: "T = Void" by auto from wt show ?case proof (rule WTrt_elim_cases) fix C TF T\<^sub>2 assume wt\<^sub>1: "P,E,h \ Val v : Class C" and has: "P \ C has F:TF in D" and wt\<^sub>2: "P,E,h \ e\<^sub>2 : T\<^sub>2" and TsubTF: "P \ T\<^sub>2 \ TF" have wt\<^sub>1': "P,E,h' \ Val v : Class C" by(rule WTrt_hext_mono[OF wt\<^sub>1 red_hext_incr[OF red]]) obtain T\<^sub>2' where wt\<^sub>2': "P,E,h' \ e\<^sub>2' : T\<^sub>2'" and T'subT: "P \ T\<^sub>2' \ T\<^sub>2" using IH[OF conf wt\<^sub>2] by blast have "P,E,h' \ Val v\F{D}:=e\<^sub>2' : Void" by(rule WTrtFAss[OF wt\<^sub>1' has wt\<^sub>2' widen_trans[OF T'subT TsubTF]]) thus ?case by auto next fix T\<^sub>2 assume null: "P,E,h \ Val v : NT" and wt\<^sub>2: "P,E,h \ e\<^sub>2 : T\<^sub>2" from null have "v = Null" by simp moreover obtain T\<^sub>2' where "P,E,h' \ e\<^sub>2' : T\<^sub>2' \ P \ T\<^sub>2' \ T\<^sub>2" using IH[OF conf wt\<^sub>2] by blast ultimately show ?thesis by(fastforce intro:WTrtFAssNT) qed next case RedFAss thus ?case by(auto simp del:fun_upd_apply) next case RedFAssNull thus ?case by(fastforce intro: WTThrow[OF WTVal] elim!:typeof_NullPointer simp:sconf_def hconf_def) next case (CallObj e h l e' h' l' M es) have red: "P \ \e,(h,l)\ \ \e',(h',l')\" and IH: "\E T. \P,E \ (h,l) \; P,E,h \ e : T\ \ \U. P,E,h' \ e' : U \ P \ U \ T" and conf: "P,E \ (h,l) \" and wt: "P,E,h \ e\M(es) : T" by fact+ \ \We distinguish if @{term e} has type @{term NT} or a Class type\ \ \Remember ?case = @{term ?case}\ { assume "P,E,h \ e:NT" hence "P,E,h' \ e' : NT" using IH[OF conf] by fastforce moreover fix Ts assume wtes: "P,E,h \ es [:] Ts" have "P,E,h' \ es [:] Ts" by(rule WTrts_hext_mono[OF wtes red_hext_incr[OF red]]) ultimately have ?case by(blast intro!:WTrtCallNT) } moreover { fix C D Ts Us pns body assume wte: "P,E,h \ e : Class C" and "method": "P \ C sees M:Ts\T = (pns,body) in D" and wtes: "P,E,h \ es [:] Us" and subs: "P \ Us [\] Ts" obtain U where wte': "P,E,h' \ e' : U" and UsubC: "P \ U \ Class C" using IH[OF conf wte] by blast \ \Is @{term U} the null type or a class type?\ { assume "U = NT" moreover have "P,E,h' \ es [:] Us" by(rule WTrts_hext_mono[OF wtes red_hext_incr[OF red]]) ultimately have ?case using wte' by(blast intro!:WTrtCallNT) } moreover { fix C' assume UClass: "U = Class C'" and "subclass": "P \ C' \\<^sup>* C" have "P,E,h' \ e' : Class C'" using wte' UClass by auto moreover obtain Ts' T' pns' body' D' where method': "P \ C' sees M:Ts'\T' = (pns',body') in D'" and subs': "P \ Ts [\] Ts'" and sub': "P \ T' \ T" using Call_lemma[OF "method" "subclass" wf] by fast moreover have "P,E,h' \ es [:] Us" by(rule WTrts_hext_mono[OF wtes red_hext_incr[OF red]]) ultimately have ?case using subs by(blast intro:WTrtCall rtrancl_trans widens_trans) } ultimately have ?case using UsubC by(auto simp add:widen_Class) } ultimately show ?case using wt by auto next case (CallParams es h l es' h' l' v M) have reds: "P \ \es,(h,l)\ [\] \es',(h',l')\" and IH: "\E Ts. \P,E \ (h,l) \; P,E,h \ es [:] Ts\ \ \Us. P,E,h' \ es' [:] Us \ P \ Us [\] Ts" and conf: "P,E \ (h,l) \" and wt: "P,E,h \ Val v\M(es) : T" by fact+ from wt show ?case proof (rule WTrt_elim_cases) fix C D Ts Us pns body assume wte: "P,E,h \ Val v : Class C" and "P \ C sees M:Ts\T = (pns,body) in D" and wtes: "P,E,h \ es [:] Us" and "P \ Us [\] Ts" moreover have "P,E,h' \ Val v : Class C" by(rule WTrt_hext_mono[OF wte reds_hext_incr[OF reds]]) moreover obtain Us' where "P,E,h' \ es' [:] Us' \ P \ Us' [\] Us" using IH[OF conf wtes] by blast ultimately show ?thesis by(blast intro:WTrtCall widens_trans) next fix Us assume null: "P,E,h \ Val v : NT" and wtes: "P,E,h \ es [:] Us" from null have "v = Null" by simp moreover obtain Us' where "P,E,h' \ es' [:] Us' \ P \ Us' [\] Us" using IH[OF conf wtes] by blast ultimately show ?thesis by(fastforce intro:WTrtCallNT) qed next case RedCallNull thus ?case by(fastforce intro: WTThrow[OF WTVal] elim!:typeof_NullPointer simp: sconf_def hconf_def) next case (InitBlockRed e h l V v e' h' l' v' T E T') have red: "P \ \e, (h,l(V\v))\ \ \e',(h',l')\" and IH: "\E T. \P,E \ (h,l(V\v)) \; P,E,h \ e : T\ \ \U. P,E,h' \ e' : U \ P \ U \ T" and v': "l' V = Some v'" and conf: "P,E \ (h,l) \" and wt: "P,E,h \ {V:T := Val v; e} : T'" by fact+ from wt obtain T\<^sub>1 where wt\<^sub>1: "typeof\<^bsub>h\<^esub> v = Some T\<^sub>1" and T1subT: "P \ T\<^sub>1 \ T" and wt\<^sub>2: "P,E(V\T),h \ e : T'" by auto have lconf\<^sub>2: "P,h \ l(V\v) (:\) E(V\T)" using conf wt\<^sub>1 T1subT by(simp add:sconf_def lconf_upd2 conf_def) have "\T\<^sub>1'. typeof\<^bsub>h'\<^esub> v' = Some T\<^sub>1' \ P \ T\<^sub>1' \ T" using v' red_preserves_lconf[OF red wt\<^sub>2 lconf\<^sub>2] by(fastforce simp:lconf_def conf_def) with IH conf lconf\<^sub>2 wt\<^sub>2 show ?case by (fastforce simp add:sconf_def) next case BlockRedNone thus ?case by(auto simp del:fun_upd_apply)(fastforce simp:sconf_def lconf_def) next case (BlockRedSome e h l V e' h' l' v T E Te) have red: "P \ \e,(h,l(V:=None))\ \ \e',(h',l')\" and IH: "\E T. \P,E \ (h,l(V:=None)) \; P,E,h \ e : T\ \ \T'. P,E,h' \ e' : T' \ P \ T' \ T" and Some: "l' V = Some v" and conf: "P,E \ (h,l) \" and wt: "P,E,h \ {V:T; e} : Te" by fact+ obtain Te' where IH': "P,E(V\T),h' \ e' : Te' \ P \ Te' \ Te" using IH conf wt by(fastforce simp:sconf_def lconf_def) have "P,h' \ l' (:\) E(V\T)" using conf wt by(fastforce intro:red_preserves_lconf[OF red] simp:sconf_def lconf_def) hence "P,h' \ v :\ T" using Some by(fastforce simp:lconf_def) with IH' show ?case by(fastforce simp:sconf_def conf_def fun_upd_same simp del:fun_upd_apply) next case SeqRed thus ?case by auto (blast dest:WTrt_hext_mono[OF _ red_hext_incr]) next case CondRed thus ?case by auto (blast intro:WTrt_hext_mono[OF _ red_hext_incr])+ next case ThrowRed thus ?case by(auto simp:is_refT_def)(blast dest:widen_Class[THEN iffD1])+ next case RedThrowNull thus ?case by(fastforce intro: WTThrow[OF WTVal] elim!:typeof_NullPointer simp:sconf_def hconf_def) next case TryRed thus ?case by auto (blast intro:widen_trans WTrt_hext_mono[OF _ red_hext_incr]) next case RedTryFail thus ?case by(fastforce intro: WTrtThrow[OF WTrtVal] simp:sconf_def hconf_def) next case ListRed1 thus ?case by(fastforce dest: WTrts_hext_mono[OF _ red_hext_incr]) next case ListRed2 thus ?case by(fastforce dest: hext_typeof_mono[OF reds_hext_incr]) qed fastforce+ (* esp all Throw propagation rules are dealt with here *) (*>*) corollary subject_reduction: "\ wf_J_prog P; P \ \e,s\ \ \e',s'\; P,E \ s \; P,E,hp s \ e:T \ \ \T'. P,E,hp s' \ e':T' \ P \ T' \ T" (*<*)by(cases s, cases s', fastforce dest:subject_reduction2)(*>*) corollary subjects_reduction: "\ wf_J_prog P; P \ \es,s\ [\] \es',s'\; P,E \ s \; P,E,hp s \ es[:]Ts \ \ \Ts'. P,E,hp s' \ es'[:]Ts' \ P \ Ts' [\] Ts" (*<*)by(cases s, cases s', fastforce dest:subjects_reduction2)(*>*) subsection \Lifting to \\*\\ text\Now all these preservation lemmas are first lifted to the transitive closure \dots\ lemma Red_preserves_sconf: assumes wf: "wf_J_prog P" and Red: "P \ \e,s\ \* \e',s'\" shows "\T. \ P,E,hp s \ e : T; P,E \ s \ \ \ P,E \ s' \" (*<*) using Red proof (induct rule:converse_rtrancl_induct2) case refl show ?case by fact next case step thus ?case by(blast intro:red_preserves_sconf dest: subject_reduction[OF wf]) qed (*>*) lemma Red_preserves_defass: assumes wf: "wf_J_prog P" and reds: "P \ \e,s\ \* \e',s'\" shows "\ e \dom(lcl s)\ \ \ e' \dom(lcl s')\" using reds proof (induct rule:converse_rtrancl_induct2) case refl thus ?case . next case (step e s e' s') thus ?case by(cases s,cases s')(auto dest:red_preserves_defass[OF wf]) qed lemma Red_preserves_type: assumes wf: "wf_J_prog P" and Red: "P \ \e,s\ \* \e',s'\" shows "!!T. \ P,E \ s\; P,E,hp s \ e:T \ \ \T'. P \ T' \ T \ P,E,hp s' \ e':T'" (*<*) using Red proof (induct rule:converse_rtrancl_induct2) case refl thus ?case by blast next case step thus ?case by(blast intro:widen_trans red_preserves_sconf dest:subject_reduction[OF wf]) qed (*>*) subsection \Lifting to \\\\ text\\dots and now to the big step semantics, just for fun.\ lemma eval_preserves_sconf: "\ wf_J_prog P; P \ \e,s\ \ \e',s'\; P,E \ e::T; P,E \ s\ \ \ P,E \ s'\" (*<*) by(blast intro:Red_preserves_sconf big_by_small WT_implies_WTrt wf_prog_wwf_prog) (*>*) lemma eval_preserves_type: assumes wf: "wf_J_prog P" shows "\ P \ \e,s\ \ \e',s'\; P,E \ s\; P,E \ e::T \ \ \T'. P \ T' \ T \ P,E,hp s' \ e':T'" (*<*) by(blast dest:big_by_small[OF wf_prog_wwf_prog[OF wf]] WT_implies_WTrt Red_preserves_type[OF wf]) (*>*) subsection "The final polish" text\The above preservation lemmas are now combined and packed nicely.\ definition wf_config :: "J_prog \ env \ state \ expr \ ty \ bool" ("_,_,_ \ _ : _ \" [51,0,0,0,0]50) where "P,E,s \ e:T \ \ P,E \ s \ \ P,E,hp s \ e:T" theorem Subject_reduction: assumes wf: "wf_J_prog P" shows "P \ \e,s\ \ \e',s'\ \ P,E,s \ e : T \ \ \T'. P,E,s' \ e' : T' \ \ P \ T' \ T" (*<*) by(force simp add: wf_config_def elim:red_preserves_sconf dest:subject_reduction[OF wf]) (*>*) theorem Subject_reductions: assumes wf: "wf_J_prog P" and reds: "P \ \e,s\ \* \e',s'\" shows "\T. P,E,s \ e:T \ \ \T'. P,E,s' \ e':T' \ \ P \ T' \ T" (*<*) using reds proof (induct rule:converse_rtrancl_induct2) case refl thus ?case by blast next case step thus ?case by(blast dest:Subject_reduction[OF wf] intro:widen_trans) qed (*>*) corollary Progress: assumes wf: "wf_J_prog P" shows "\ P,E,s \ e : T \; \ e \dom(lcl s)\; \ final e \ \ \e' s'. P \ \e,s\ \ \e',s'\" (*<*) using progress[OF wf_prog_wwf_prog[OF wf]] by(auto simp:wf_config_def sconf_def) (*>*) corollary TypeSafety: - "\ wf_J_prog P; P,E \ s \; P,E \ e::T; \ e \dom(lcl s)\; - P \ \e,s\ \* \e',s'\; \(\e'' s''. P \ \e',s'\ \ \e'',s''\) \ - \ (\v. e' = Val v \ P,hp s' \ v :\ T) \ +fixes s::state +assumes wf: "wf_J_prog P" and sconf: "P,E \ s \" and wt: "P,E \ e::T" + and \: "\ e \dom(lcl s)\" + and steps: "P \ \e,s\ \* \e',s'\" + and nstep: "\(\e'' s''. P \ \e',s'\ \ \e'',s''\)" +shows "(\v. e' = Val v \ P,hp s' \ v :\ T) \ (\a. e' = Throw a \ a \ dom(hp s'))" (*<*) -apply(subgoal_tac " P,E,s \ e:T \") - prefer 2 - apply(fastforce simp:wf_config_def dest:WT_implies_WTrt) -apply(frule (2) Subject_reductions) -apply(erule exE conjE)+ -apply(frule (2) Red_preserves_defass) -apply(subgoal_tac "final e'") - prefer 2 - apply(blast dest: Progress) -apply (fastforce simp:wf_config_def final_def conf_def dest: Progress) -done +proof - + have wfc: "P,E,s \ e:T \" using WT_implies_WTrt[OF wt] sconf + by(simp add:wf_config_def) + obtain T' where wfc': "P,E,s' \ e' : T' \" and T': "P \ T' \ T" + using Subject_reductions[OF wf steps wfc] by clarsimp + have \': "\ e' \dom (lcl s')\" + by(rule Red_preserves_defass[OF wf steps \]) + have fin': "final e'" using Progress[OF wf wfc' \'] nstep by blast + then show ?thesis using wfc wfc' T' + by(fastforce simp:wf_config_def final_def conf_def) +qed (*>*) end diff --git a/thys/JinjaDCI/J/EConform.thy b/thys/JinjaDCI/J/EConform.thy --- a/thys/JinjaDCI/J/EConform.thy +++ b/thys/JinjaDCI/J/EConform.thy @@ -1,337 +1,337 @@ (* Title: JinjaDCI/J/EConform.thy Author: Susannah Mansky 2019-20 UIUC *) section \ Expression conformance properties \ theory EConform imports SmallStep BigStep begin lemma cons_to_append: "list \ [] \ (\ls. a # list = ls @ [last list])" by (metis append_butlast_last_id last_ConsR list.simps(3)) subsection "Initialization conformance" \ \ returns class that can be initialized (if any) by top-level expression \ fun init_class :: "'m prog \ 'a exp \ cname option" where "init_class P (new C) = Some C" | "init_class P (C\\<^sub>sF{D}) = Some D" | "init_class P (C\\<^sub>sF{D}:=e\<^sub>2) = Some D" | "init_class P (C\\<^sub>sM(es)) = seeing_class P C M" | "init_class _ _ = None" lemma icheck_init_class: "icheck P C e \ init_class P e = \C\" proof(induct e) case (SFAss x1 x2 x3 e') then show ?case by(case_tac e') auto qed auto \ \ exp to take next small step (in particular, subexp that may contain initialization) \ fun ss_exp :: "'a exp \ 'a exp" and ss_exps :: "'a exp list \ 'a exp option" where "ss_exp (new C) = new C" | "ss_exp (Cast C e) = (case val_of e of Some v \ Cast C e | _ \ ss_exp e)" | "ss_exp (Val v) = Val v" | "ss_exp (e\<^sub>1 \bop\ e\<^sub>2) = (case val_of e\<^sub>1 of Some v \ (case val_of e\<^sub>2 of Some v \ e\<^sub>1 \bop\ e\<^sub>2 | _ \ ss_exp e\<^sub>2) | _ \ ss_exp e\<^sub>1)" | "ss_exp (Var V) = Var V" | "ss_exp (LAss V e) = (case val_of e of Some v \ LAss V e | _ \ ss_exp e)" | "ss_exp (e\F{D}) = (case val_of e of Some v \ e\F{D} | _ \ ss_exp e)" | "ss_exp (C\\<^sub>sF{D}) = C\\<^sub>sF{D}" | "ss_exp (e\<^sub>1\F{D}:=e\<^sub>2) = (case val_of e\<^sub>1 of Some v \ (case val_of e\<^sub>2 of Some v \ e\<^sub>1\F{D}:=e\<^sub>2 | _ \ ss_exp e\<^sub>2) | _ \ ss_exp e\<^sub>1)" | "ss_exp (C\\<^sub>sF{D}:=e\<^sub>2) = (case val_of e\<^sub>2 of Some v \ C\\<^sub>sF{D}:=e\<^sub>2 | _ \ ss_exp e\<^sub>2)" | "ss_exp (e\M(es)) = (case val_of e of Some v \ (case map_vals_of es of Some t \ e\M(es) | _ \ the(ss_exps es)) | _ \ ss_exp e)" | "ss_exp (C\\<^sub>sM(es)) = (case map_vals_of es of Some t \ C\\<^sub>sM(es) | _ \ the(ss_exps es))" | "ss_exp ({V:T; e}) = ss_exp e" | "ss_exp (e\<^sub>1;;e\<^sub>2) = (case val_of e\<^sub>1 of Some v \ ss_exp e\<^sub>2 | None \ (case lass_val_of e\<^sub>1 of Some p \ ss_exp e\<^sub>2 | None \ ss_exp e\<^sub>1))" | "ss_exp (if (b) e\<^sub>1 else e\<^sub>2) = (case bool_of b of Some True \ if (b) e\<^sub>1 else e\<^sub>2 | Some False \ if (b) e\<^sub>1 else e\<^sub>2 | _ \ ss_exp b)" | "ss_exp (while (b) e) = while (b) e" | "ss_exp (throw e) = (case val_of e of Some v \ throw e | _ \ ss_exp e)" | "ss_exp (try e\<^sub>1 catch(C V) e\<^sub>2) = (case val_of e\<^sub>1 of Some v \ try e\<^sub>1 catch(C V) e\<^sub>2 | _ \ ss_exp e\<^sub>1)" | "ss_exp (INIT C (Cs,b) \ e) = INIT C (Cs,b) \ e" | "ss_exp (RI (C,e);Cs \ e') = (case val_of e of Some v \ RI (C,e);Cs \ e | _ \ ss_exp e)" | "ss_exps([]) = None" | "ss_exps(e#es) = (case val_of e of Some v \ ss_exps es | _ \ Some (ss_exp e))" (*<*) lemmas ss_exp_ss_exps_induct = ss_exp_ss_exps.induct [ case_names New Cast Val BinOp Var LAss FAcc SFAcc FAss SFAss Call SCall Block Seq Cond While Throw Try Init RI Nil Cons ] (*>*) lemma icheck_ss_exp: assumes "icheck P C e" shows "ss_exp e = e" using assms proof(cases e) case (SFAss C F D e) then show ?thesis using assms proof(cases e)qed(auto) qed(auto) lemma ss_exps_Vals_None[simp]: "ss_exps (map Val vs) = None" by(induct vs) (auto) lemma ss_exps_Vals_NoneI: "ss_exps es = None \ \vs. es = map Val vs" using val_of_spec by(induct es) (auto) lemma ss_exps_throw_nVal: "\ val_of e = None; ss_exps (map Val vs @ throw e # es') = \e'\ \ \ e' = ss_exp e" by(induct vs) (auto) lemma ss_exps_throw_Val: "\ val_of e = \a\; ss_exps (map Val vs @ throw e # es') = \e'\ \ \ e' = throw e" by(induct vs) (auto) abbreviation curr_init :: "'m prog \ 'a exp \ cname option" where "curr_init P e \ init_class P (ss_exp e)" abbreviation curr_inits :: "'m prog \ 'a exp list \ cname option" where "curr_inits P es \ case ss_exps es of Some e \ init_class P e | _ \ None" lemma icheck_curr_init': "\e'. ss_exp e = e' \ icheck P C e' \ curr_init P e = \C\" and icheck_curr_inits': "\e. ss_exps es = \e\ \ icheck P C e \ curr_inits P es = \C\" proof(induct rule: ss_exp_ss_exps_induct) qed(simp_all add: icheck_init_class) lemma icheck_curr_init: "icheck P C e' \ ss_exp e = e' \ curr_init P e = \C\" by(rule icheck_curr_init') lemma icheck_curr_inits: "icheck P C e \ ss_exps es = \e\ \ curr_inits P es = \C\" by(rule icheck_curr_inits') definition initPD :: "sheap \ cname \ bool" where "initPD sh C \ \sfs i. sh C = Some (sfs, i) \ (i = Done \ i = Processing)" \ \ checks that @{text INIT} and @{text RI} conform and are only in the main computation \ fun iconf :: "sheap \ 'a exp \ bool" and iconfs :: " sheap \ 'a exp list \ bool" where "iconf sh (new C) = True" | "iconf sh (Cast C e) = iconf sh e" | "iconf sh (Val v) = True" | "iconf sh (e\<^sub>1 \bop\ e\<^sub>2) = (case val_of e\<^sub>1 of Some v \ iconf sh e\<^sub>2 | _ \ iconf sh e\<^sub>1 \ \sub_RI e\<^sub>2)" | "iconf sh (Var V) = True" | "iconf sh (LAss V e) = iconf sh e" | "iconf sh (e\F{D}) = iconf sh e" | "iconf sh (C\\<^sub>sF{D}) = True" | "iconf sh (e\<^sub>1\F{D}:=e\<^sub>2) = (case val_of e\<^sub>1 of Some v \ iconf sh e\<^sub>2 | _ \ iconf sh e\<^sub>1 \ \sub_RI e\<^sub>2)" | "iconf sh (C\\<^sub>sF{D}:=e\<^sub>2) = iconf sh e\<^sub>2" | "iconf sh (e\M(es)) = (case val_of e of Some v \ iconfs sh es | _ \ iconf sh e \ \sub_RIs es)" | "iconf sh (C\\<^sub>sM(es)) = iconfs sh es" | "iconf sh ({V:T; e}) = iconf sh e" | "iconf sh (e\<^sub>1;;e\<^sub>2) = (case val_of e\<^sub>1 of Some v \ iconf sh e\<^sub>2 | None \ (case lass_val_of e\<^sub>1 of Some p \ iconf sh e\<^sub>2 | None \ iconf sh e\<^sub>1 \ \sub_RI e\<^sub>2))" | "iconf sh (if (b) e\<^sub>1 else e\<^sub>2) = (iconf sh b \ \sub_RI e\<^sub>1 \ \sub_RI e\<^sub>2)" | "iconf sh (while (b) e) = (\sub_RI b \ \sub_RI e)" | "iconf sh (throw e) = iconf sh e" | "iconf sh (try e\<^sub>1 catch(C V) e\<^sub>2) = (iconf sh e\<^sub>1 \ \sub_RI e\<^sub>2)" | "iconf sh (INIT C (Cs,b) \ e) = ((case Cs of Nil \ initPD sh C | C'#Cs' \ last Cs = C) \ \sub_RI e)" | "iconf sh (RI (C,e);Cs \ e') = (iconf sh e \ \sub_RI e')" | "iconfs sh ([]) = True" | "iconfs sh (e#es) = (case val_of e of Some v \ iconfs sh es | _ \ iconf sh e \ \sub_RIs es)" lemma iconfs_map_throw: "iconfs sh (map Val vs @ throw e # es') \ iconf sh e" by(induct vs,auto) lemma nsub_RI_iconf_aux: "(\sub_RI (e::'a exp) \ (\e'. e' \ subexp e \ \sub_RI e' \ iconf sh e') \ iconf sh e) \ (\sub_RIs (es::'a exp list) \ (\e'. e' \ subexps es \ \sub_RI e' \ iconf sh e') \ iconfs sh es)" proof(induct rule: sub_RI_sub_RIs.induct) qed(auto) lemma nsub_RI_iconf_aux': "(\e'. subexp_of e' e \ \sub_RI e' \ iconf sh e') \ (\sub_RI e \ iconf sh e)" by(simp add: nsub_RI_iconf_aux) lemma nsub_RI_iconf: "\sub_RI e \ iconf sh e" and nsub_RIs_iconfs: "\sub_RIs es \ iconfs sh es" proof - let ?R = "\e. \sub_RI e \ iconf sh e" let ?Rs = "\es. \sub_RIs es \ iconfs sh es" have "(\e'. subexp_of e' e \ ?R e') \ ?R e" - and "(\e'. e' \ subexps es \ ?R e') \ ?Rs es" by(rule subexp_induct[where ?Rs = ?Rs]; clarsimp simp: nsub_RI_iconf_aux) - (rule subexps_induct; clarsimp simp: nsub_RI_iconf_aux) - then show "\sub_RI e \ iconf sh e" - and "\sub_RIs es \ iconfs sh es" by simp+ + moreover have "(\e'. e' \ subexps es \ ?R e') \ ?Rs es" + by(rule subexps_induct; clarsimp simp: nsub_RI_iconf_aux) + ultimately show "\sub_RI e \ iconf sh e" + and "\sub_RIs es \ iconfs sh es" by simp+ qed lemma lass_val_of_iconf: "lass_val_of e = \a\ \ iconf sh e" by(drule lass_val_of_nsub_RI, erule nsub_RI_iconf) lemma icheck_iconf: assumes "icheck P C e" shows "iconf sh e" using assms proof(cases e) case (SFAss C F D e) then show ?thesis using assms proof(cases e)qed(auto) next case (SCall C M es) then show ?thesis using assms by (auto simp: nsub_RIs_iconfs) next qed(auto) subsection "Indicator boolean conformance" \ \ checks that the given expression, indicator boolean pair is allowed in small-step (i.e., if @{term b} is True, then @{term e} is an initialization-calling expression to a class that is marked either @{term Processing} or @{term Done}) \ definition bconf :: "'m prog \ sheap \ 'a exp \ bool \ bool" ("_,_ \\<^sub>b '(_,_') \" [51,51,0,0] 50) where "P,sh \\<^sub>b (e,b) \ \ b \ (\C. icheck P C (ss_exp e) \ initPD sh C)" definition bconfs :: "'m prog \ sheap \ 'a exp list \ bool \ bool" ("_,_ \\<^sub>b '(_,_') \" [51,51,0,0] 50) where "P,sh \\<^sub>b (es,b) \ \ b \ (\C. (icheck P C (the(ss_exps es)) \ (curr_inits P es = Some C) \ initPD sh C))" \ \ bconf helper lemmas \ lemma bconf_nonVal[simp]: "P,sh \\<^sub>b (e,True) \ \ val_of e = None" by(cases e) (auto simp: bconf_def) lemma bconfs_nonVals[simp]: "P,sh \\<^sub>b (es,True) \ \ map_vals_of es = None" by(induct es) (auto simp: bconfs_def) lemma bconf_Cast[iff]: "P,sh \\<^sub>b (Cast C e,b) \ \ P,sh \\<^sub>b (e,b) \" by(cases b) (auto simp: bconf_def dest: val_of_spec) lemma bconf_BinOp[iff]: "P,sh \\<^sub>b (e1 \bop\ e2,b) \ \ (case val_of e1 of Some v \ P,sh \\<^sub>b (e2,b) \ | _ \ P,sh \\<^sub>b (e1,b) \)" by(cases b) (auto simp: bconf_def dest: val_of_spec) lemma bconf_LAss[iff]: "P,sh \\<^sub>b (LAss V e,b) \ \ P,sh \\<^sub>b (e,b) \" by(cases b) (auto simp: bconf_def dest: val_of_spec) lemma bconf_FAcc[iff]: "P,sh \\<^sub>b (e\F{D},b) \ \ P,sh \\<^sub>b (e,b) \" by(cases b) (auto simp: bconf_def dest: val_of_spec) lemma bconf_FAss[iff]: "P,sh \\<^sub>b (FAss e1 F D e2,b) \ \ (case val_of e1 of Some v \ P,sh \\<^sub>b (e2,b) \ | _ \ P,sh \\<^sub>b (e1,b) \)" by(cases b) (auto simp: bconf_def dest: val_of_spec) lemma bconf_SFAss[iff]: "val_of e2 = None \ P,sh \\<^sub>b (SFAss C F D e2,b) \ \ P,sh \\<^sub>b (e2,b) \" by(cases b) (auto simp: bconf_def) lemma bconfs_Vals[iff]: "P,sh \\<^sub>b (map Val vs, b) \ \ \ b" by(unfold bconfs_def) simp lemma bconf_Call[iff]: "P,sh \\<^sub>b (e\M(es),b) \ \ (case val_of e of Some v \ P,sh \\<^sub>b (es,b) \ | _ \ P,sh \\<^sub>b (e,b) \)" proof(cases b) case True then show ?thesis proof(cases "ss_exps es") case None then obtain vs where "es = map Val vs" using ss_exps_Vals_NoneI by auto then have mv: "map_vals_of es = \vs\" by simp then show ?thesis by(auto simp: bconf_def) (simp add: bconfs_def) next case (Some a) then show ?thesis by(auto simp: bconf_def) (auto simp: bconfs_def icheck_init_class) qed qed(simp add: bconf_def bconfs_def) lemma bconf_SCall[iff]: assumes mvn: "map_vals_of es = None" shows "P,sh \\<^sub>b (C\\<^sub>sM(es),b) \ \ P,sh \\<^sub>b (es,b) \" proof(cases b) case True then show ?thesis proof(cases "ss_exps es") case None then have "\vs. es = map Val vs" using ss_exps_Vals_NoneI by auto then show ?thesis using mvn finals_def by clarsimp next case (Some a) then show ?thesis by(auto simp: bconf_def) (auto simp: bconfs_def icheck_init_class) qed qed(simp add: bconf_def bconfs_def) lemma bconf_Cons[iff]: "P,sh \\<^sub>b (e#es,b) \ \ (case val_of e of Some v \ P,sh \\<^sub>b (es,b) \ | _ \ P,sh \\<^sub>b (e,b) \)" proof(cases b) case True then show ?thesis proof(cases "ss_exps es") case None then have "\vs. es = map Val vs" using ss_exps_Vals_NoneI by auto then show ?thesis using None by(auto simp: bconf_def bconfs_def icheck_init_class) next case (Some a) then show ?thesis by(auto simp: bconf_def bconfs_def icheck_init_class) qed qed(simp add: bconf_def bconfs_def) lemma bconf_InitBlock[iff]: "P,sh \\<^sub>b ({V:T; V:=Val v;; e\<^sub>2},b) \ \ P,sh \\<^sub>b (e\<^sub>2,b) \" by(cases b) (auto simp: bconf_def assigned_def) lemma bconf_Block[iff]: "P,sh \\<^sub>b ({V:T; e},b) \ \ P,sh \\<^sub>b (e,b) \" by(cases b) (auto simp: bconf_def) lemma bconf_Seq[iff]: "P,sh \\<^sub>b (e1;;e2,b) \ \ (case val_of e1 of Some v \ P,sh \\<^sub>b (e2,b) \ | _ \ (case lass_val_of e1 of Some p \ P,sh \\<^sub>b (e2,b) \ | None \ P,sh \\<^sub>b (e1,b) \))" by(cases b) (auto simp: bconf_def dest: val_of_spec lass_val_of_spec) lemma bconf_Cond[iff]: "P,sh \\<^sub>b (if (b) e\<^sub>1 else e\<^sub>2,b') \ \ P,sh \\<^sub>b (b,b') \" proof(cases "bool_of b") case None then show ?thesis by(auto simp: bconf_def) next case (Some a) then show ?thesis by(case_tac a) (auto simp: bconf_def dest: bool_of_specT bool_of_specF) qed lemma bconf_While[iff]: "P,sh \\<^sub>b (while (b) e,b') \ \ \b'" by(cases b) (auto simp: bconf_def) lemma bconf_Throw[iff]: "P,sh \\<^sub>b (throw e,b) \ \ P,sh \\<^sub>b (e,b) \" by(cases b) (auto simp: bconf_def dest: val_of_spec) lemma bconf_Try[iff]: "P,sh \\<^sub>b (try e\<^sub>1 catch(C V) e\<^sub>2,b) \ \ P,sh \\<^sub>b (e\<^sub>1,b) \" by(cases b) (auto simp: bconf_def dest: val_of_spec) lemma bconf_INIT[iff]: "P,sh \\<^sub>b (INIT C (Cs,b') \ e,b) \ \ \b" by(cases b) (auto simp: bconf_def) lemma bconf_RI[iff]: "P,sh \\<^sub>b (RI(C,e);Cs \ e',b) \ \ P,sh \\<^sub>b (e,b) \" by(cases b) (auto simp: bconf_def dest: val_of_spec) lemma bconfs_map_throw[iff]: "P,sh \\<^sub>b (map Val vs @ throw e # es',b) \ \ P,sh \\<^sub>b (e,b) \" by(induct vs) auto end diff --git a/thys/JinjaDCI/J/Equivalence.thy b/thys/JinjaDCI/J/Equivalence.thy --- a/thys/JinjaDCI/J/Equivalence.thy +++ b/thys/JinjaDCI/J/Equivalence.thy @@ -1,4292 +1,4312 @@ (* Title: JinjaDCI/J/Equivalence.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory J/Equivalence.thy by Tobias Nipkow *) section \ Equivalence of Big Step and Small Step Semantics \ theory Equivalence imports TypeSafe WWellForm begin subsection\Small steps simulate big step\ subsubsection "Init" text "The reduction of initialization expressions doesn't touch or use their on-hold expressions (the subexpression to the right of @{text \}) until the initialization operation completes. This function is used to prove this and related properties. It is then used for general reduction of initialization expressions separately from their on-hold expressions by using the on-hold expression @{term unit}, then putting the real on-hold expression back in at the end." fun init_switch :: "'a exp \ 'a exp \ 'a exp" where "init_switch (INIT C (Cs,b) \ e\<^sub>i) e = (INIT C (Cs,b) \ e)" | "init_switch (RI(C,e');Cs \ e\<^sub>i) e = (RI(C,e');Cs \ e)" | "init_switch e' e = e'" fun INIT_class :: "'a exp \ cname option" where "INIT_class (INIT C (Cs,b) \ e) = (if C = last (C#Cs) then Some C else None)" | "INIT_class (RI(C,e\<^sub>0);Cs \ e) = Some (last (C#Cs))" | "INIT_class _ = None" lemma init_red_init: "\ init_exp_of e\<^sub>0 = \e\; P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \ \e\<^sub>1,s\<^sub>1,b\<^sub>1\ \ \ (init_exp_of e\<^sub>1 = \e\ \ (INIT_class e\<^sub>0 = \C\ \ INIT_class e\<^sub>1 = \C\)) \ (e\<^sub>1 = e \ b\<^sub>1 = icheck P (the(INIT_class e\<^sub>0)) e) \ (\a. e\<^sub>1 = throw a)" by(erule red.cases, auto) lemma init_exp_switch[simp]: "init_exp_of e\<^sub>0 = \e\ \ init_exp_of (init_switch e\<^sub>0 e') = \e'\" by(cases e\<^sub>0, simp_all) lemma init_red_sync: "\ P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \ \e\<^sub>1,s\<^sub>1,b\<^sub>1\; init_exp_of e\<^sub>0 = \e\; e\<^sub>1 \ e \ \ (\e'. P \ \init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\ \ \init_switch e\<^sub>1 e',s\<^sub>1,b\<^sub>1\)" proof(induct rule: red.cases) qed(simp_all add: red_reds.intros) lemma init_red_sync_end: "\ P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \ \e\<^sub>1,s\<^sub>1,b\<^sub>1\; init_exp_of e\<^sub>0 = \e\; e\<^sub>1 = e; throw_of e = None \ \ (\e'. \sub_RI e' \ P \ \init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\ \ \e',s\<^sub>1,icheck P (the(INIT_class e\<^sub>0)) e'\)" proof(induct rule: red.cases) qed(simp_all add: red_reds.intros) lemma init_reds_sync_unit': "\ P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \* \Val v',s\<^sub>1,b\<^sub>1\; init_exp_of e\<^sub>0 = \unit\; INIT_class e\<^sub>0 = \C\ \ \ (\e'. \sub_RI e' \ P \ \init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\ \* \e',s\<^sub>1,icheck P (the(INIT_class e\<^sub>0)) e'\)" proof(induct rule:converse_rtrancl_induct3) case refl then show ?case by simp next case (step e0 s0 b0 e1 s1 b1) have "(init_exp_of e1 = \unit\ \ (INIT_class e0 = \C\ \ INIT_class e1 = \C\)) \ (e1 = unit \ b1 = icheck P (the(INIT_class e0)) unit) \ (\a. e1 = throw a)" using init_red_init[OF step.prems(1) step.hyps(1)] by simp then show ?case proof(rule disjE) assume assm: "init_exp_of e1 = \unit\ \ (INIT_class e0 = \C\ \ INIT_class e1 = \C\)" then have red: "P \ \init_switch e0 e',s0,b0\ \ \init_switch e1 e',s1,b1\" using init_red_sync[OF step.hyps(1) step.prems(1)] by simp have reds: "P \ \init_switch e1 e',s1,b1\ \* \e',s\<^sub>1,icheck P (the (INIT_class e0)) e'\" using step.hyps(3)[OF _ _ step.prems(3)] assm step.prems(2) by simp show ?thesis by(rule converse_rtrancl_into_rtrancl[OF red reds]) next assume "(e1 = unit \ b1 = icheck P (the(INIT_class e0)) unit) \ (\a. e1 = throw a)" then show ?thesis proof(rule disjE) assume assm: "e1 = unit \ b1 = icheck P (the(INIT_class e0)) unit" then have e1: "e1 = unit" by simp have sb: "s1 = s\<^sub>1" "b1 = b\<^sub>1" using reds_final_same[OF step.hyps(2)] assm by(simp_all add: final_def) then have step': "P \ \init_switch e0 e',s0,b0\ \ \e',s\<^sub>1,icheck P (the (INIT_class e0)) e'\" using init_red_sync_end[OF step.hyps(1) step.prems(1) e1 _ step.prems(3)] by auto then have "P \ \init_switch e0 e',s0,b0\ \* \e',s\<^sub>1,icheck P (the (INIT_class e0)) e'\" using r_into_rtrancl by auto then show ?thesis using step assm sb by simp next assume "\a. e1 = throw a" then obtain a where "e1 = throw a" by clarsimp then have tof: "throw_of e1 = \a\" by simp then show ?thesis using reds_throw[OF step.hyps(2) tof] by simp qed qed qed lemma init_reds_sync_unit_throw': "\ P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\; init_exp_of e\<^sub>0 = \unit\ \ \ (\e'. P \ \init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\)" proof(induct rule:converse_rtrancl_induct3) case refl then show ?case by simp next case (step e0 s0 b0 e1 s1 b1) have "init_exp_of e1 = \unit\ \ (\C. INIT_class e0 = \C\ \ INIT_class e1 = \C\) \ e1 = unit \ b1 = icheck P (the (INIT_class e0)) unit \ (\a. e1 = throw a)" using init_red_init[OF step.prems(1) step.hyps(1)] by auto then show ?case proof(rule disjE) assume assm: "init_exp_of e1 = \unit\ \ (\C. INIT_class e0 = \C\ \ INIT_class e1 = \C\)" then have "P \ \init_switch e0 e',s0,b0\ \ \init_switch e1 e',s1,b1\" using step init_red_sync[OF step.hyps(1) step.prems] by simp then show ?thesis using step assm by (meson converse_rtrancl_into_rtrancl) next assume "e1 = unit \ b1 = icheck P (the (INIT_class e0)) unit \ (\a. e1 = throw a)" then show ?thesis proof(rule disjE) assume "e1 = unit \ b1 = icheck P (the (INIT_class e0)) unit" then show ?thesis using step using final_def reds_final_same by blast next assume assm: "\a. e1 = throw a" then have "P \ \init_switch e0 e',s0,b0\ \ \e1,s1,b1\" using init_red_sync[OF step.hyps(1) step.prems] by clarsimp then show ?thesis using step by simp qed qed qed lemma init_reds_sync_unit: assumes "P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \* \Val v',s\<^sub>1,b\<^sub>1\" and "init_exp_of e\<^sub>0 = \unit\" and "INIT_class e\<^sub>0 = \C\" and "\sub_RI e'" shows "P \ \init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\ \* \e',s\<^sub>1,icheck P (the(INIT_class e\<^sub>0)) e'\" using init_reds_sync_unit'[OF assms] by clarsimp lemma init_reds_sync_unit_throw: assumes "P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\" and "init_exp_of e\<^sub>0 = \unit\" shows "P \ \init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\" using init_reds_sync_unit_throw'[OF assms] by clarsimp \ \ init reds lemmas \ lemma InitSeqReds: assumes "P \ \INIT C ([C],b) \ unit,s\<^sub>0,b\<^sub>0\ \* \Val v',s\<^sub>1,b\<^sub>1\" and "P \ \e,s\<^sub>1,icheck P C e\ \* \e\<^sub>2,s\<^sub>2,b\<^sub>2\" and "\sub_RI e" shows "P \ \INIT C ([C],b) \ e,s\<^sub>0,b\<^sub>0\ \* \e\<^sub>2,s\<^sub>2,b\<^sub>2\" using assms proof - have "P \ \init_switch (INIT C ([C],b) \ unit) e,s\<^sub>0,b\<^sub>0\ \* \e,s\<^sub>1,icheck P C e\" using init_reds_sync_unit[OF assms(1) _ _ assms(3)] by simp then show ?thesis using assms(2) by simp qed lemma InitSeqThrowReds: assumes "P \ \INIT C ([C],b) \ unit,s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\" shows "P \ \INIT C ([C],b) \ e,s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\" using assms proof - have "P \ \init_switch (INIT C ([C],b) \ unit) e,s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\" using init_reds_sync_unit_throw[OF assms(1)] by simp then show ?thesis by simp qed lemma InitNoneReds: "\ sh C = None; P \ \INIT C' (C # Cs,False) \ e,(h, l, sh(C \ (sblank P C, Prepared))),b\ \* \e',s',b'\ \ \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \* \e',s',b'\" (*<*) -apply(rule converse_rtrancl_into_rtrancl) - apply(erule InitNoneRed) -apply assumption -done +by(auto intro: converse_rtrancl_into_rtrancl + elim: InitNoneRed) (*>*) lemma InitDoneReds: "\ sh C = Some(sfs,Done); P \ \INIT C' (Cs,True) \ e,(h,l,sh),b\ \* \e',s',b'\ \ \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \* \e',s',b'\" (*<*) -apply(rule converse_rtrancl_into_rtrancl) - apply(erule RedInitDone) -apply assumption -done +by(auto intro: converse_rtrancl_into_rtrancl + elim: RedInitDone) (*>*) lemma InitProcessingReds: "\ sh C = Some(sfs,Processing); P \ \INIT C' (Cs,True) \ e,(h,l,sh),b\ \* \e',s',b'\ \ \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \* \e',s',b'\" (*<*) -apply(rule converse_rtrancl_into_rtrancl) - apply(erule RedInitProcessing) -apply assumption -done +by(auto intro: converse_rtrancl_into_rtrancl + elim: RedInitProcessing) (*>*) lemma InitErrorReds: "\ sh C = Some(sfs,Error); P \ \RI (C,THROW NoClassDefFoundError);Cs \ e,(h,l,sh),b\ \* \e',s',b'\ \ \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \* \e',s',b'\" (*<*) -apply(rule converse_rtrancl_into_rtrancl) - apply(erule RedInitError) -apply assumption -done +by(auto intro: converse_rtrancl_into_rtrancl + elim: RedInitError) (*>*) lemma InitObjectReds: "\ sh C = Some(sfs,Prepared); C = Object; sh' = sh(C \ (sfs,Processing)); P \ \INIT C' (C#Cs,True) \ e,(h,l,sh'),b\ \* \e',s',b'\ \ \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \* \e',s',b'\" (*<*) -apply(rule converse_rtrancl_into_rtrancl) - apply(erule (2) InitObjectRed) -apply assumption -done +by(auto intro: converse_rtrancl_into_rtrancl + elim: InitObjectRed) (*>*) lemma InitNonObjectReds: "\ sh C = Some(sfs,Prepared); C \ Object; class P C = Some (D,r); sh' = sh(C \ (sfs,Processing)); P \ \INIT C' (D#C#Cs,False) \ e,(h,l,sh'),b\ \* \e',s',b'\ \ \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \* \e',s',b'\" (*<*) -apply(rule converse_rtrancl_into_rtrancl) - apply(erule (3) InitNonObjectSuperRed) -apply assumption -done +by(auto intro: converse_rtrancl_into_rtrancl + elim: InitNonObjectSuperRed) (*>*) lemma RedsInitRInit: "P \ \RI (C,C\\<^sub>sclinit([]));Cs \ e,(h,l,sh),b\ \* \e',s',b'\ \ P \ \INIT C' (C#Cs,True) \ e,(h,l,sh),b\ \* \e',s',b'\" (*<*) -apply(rule converse_rtrancl_into_rtrancl) - apply(rule RedInitRInit) -apply assumption -done +by(auto intro: converse_rtrancl_into_rtrancl RedInitRInit) (*>*) lemmas rtrancl_induct3 = rtrancl_induct[of "(ax,ay,az)" "(bx,by,bz)", split_format (complete), consumes 1, case_names refl step] lemma RInitReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \RI (C,e);Cs \ e\<^sub>0, s, b\ \* \RI (C,e');Cs \ e\<^sub>0, s', b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule RInitRed) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] RInitRed[OF step(2)] + by simp +qed (*>*) lemma RedsRInit: - "\ P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b\<^sub>1\; - sh\<^sub>1 C = Some (sfs, i); sh\<^sub>2 = sh\<^sub>1(C \ (sfs,Done)); C' = last(C#Cs); - P \ \INIT C' (Cs,True) \ e,(h\<^sub>1,l\<^sub>1,sh\<^sub>2),b\<^sub>1\ \* \e',s',b'\ \ -\ P \ \RI (C, e\<^sub>0);Cs \ e,s\<^sub>0,b\<^sub>0\ \* \e',s',b'\" +assumes e\<^sub>0_steps: "P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b\<^sub>1\" + and "sh\<^sub>1 C = Some (sfs, i)" and "sh\<^sub>2 = sh\<^sub>1(C \ (sfs,Done))" + and "C' = last(C#Cs)" + and "P \ \INIT C' (Cs,True) \ e,(h\<^sub>1,l\<^sub>1,sh\<^sub>2),b\<^sub>1\ \* \e',s',b'\" +shows "P \ \RI (C, e\<^sub>0);Cs \ e,s\<^sub>0,b\<^sub>0\ \* \e',s',b'\" (*<*) -apply(rule rtrancl_trans) - apply(erule RInitReds) -apply(rule converse_rtrancl_into_rtrancl) - apply(erule (2) RedRInit) -apply assumption -done +proof - + let ?y = "(RI (C,Val v) ; Cs \ e,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b\<^sub>1)" + have "((RI (C,e\<^sub>0) ; Cs \ e, s\<^sub>0, b\<^sub>0), ?y) \ (red P)\<^sup>*" + by(rule RInitReds[OF e\<^sub>0_steps]) + also have "(?y, e', s', b') \ (red P)\<^sup>*" using assms(2-5) + by(auto intro: converse_rtrancl_into_rtrancl + elim: RedRInit) + ultimately show ?thesis by simp +qed (*>*) + lemma RInitInitThrowReds: - "\ P \ \e,s,b\ \* \Throw a, (h',l',sh'),b'\; - sh' C = Some(sfs, i); sh'' = sh'(C \ (sfs, Error)); - P \ \RI (D,Throw a);Cs \ e\<^sub>0, (h',l',sh''),b'\ \* \e\<^sub>1,s\<^sub>1,b\<^sub>1\ \ - \ P \ \RI (C,e);D#Cs \ e\<^sub>0,s,b\ \* \e\<^sub>1,s\<^sub>1,b\<^sub>1\" +assumes e_steps: "P \ \e,s,b\ \* \Throw a,(h',l',sh'),b'\" + and "sh' C = Some (sfs, i)" and "sh'' = sh'(C \ (sfs,Error))" + and "P \ \RI (D,Throw a);Cs \ e\<^sub>0, (h',l',sh''),b'\ \* \e\<^sub>1,s\<^sub>1,b\<^sub>1\" +shows "P \ \RI (C, e);D#Cs \ e\<^sub>0,s,b\ \* \e\<^sub>1,s\<^sub>1,b\<^sub>1\" (*<*) -apply(rule rtrancl_trans) - apply(erule RInitReds) -apply(rule converse_rtrancl_into_rtrancl) - apply(erule (1) RInitInitThrow) -apply assumption -done +proof - + let ?y = "(RI (C,Throw a) ; D # Cs \ e\<^sub>0,(h',l',sh'),b')" + have "((RI (C,e) ; D#Cs \ e\<^sub>0, s, b), ?y) \ (red P)\<^sup>*" + by(rule RInitReds[OF e_steps]) + also have "(?y, e\<^sub>1, s\<^sub>1, b\<^sub>1) \ (red P)\<^sup>*" using assms(2-4) + by(auto intro: converse_rtrancl_into_rtrancl + elim: RInitInitThrow) + ultimately show ?thesis by simp +qed (*>*) lemma RInitThrowReds: "\ P \ \e,s,b\ \* \Throw a, (h',l',sh'),b'\; sh' C = Some(sfs, i); sh'' = sh'(C \ (sfs, Error)) \ \ P \ \RI (C,e);Nil \ e\<^sub>0,s,b\ \* \Throw a, (h',l',sh''),b'\" (*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule RInitReds) -apply(erule RInitThrow) -apply assumption -done +by(auto intro: rtrancl_into_rtrancl + elim: RInitReds RInitThrow) (*>*) subsubsection "New" lemma NewInitDoneReds: "\ sh C = Some (sfs, Done); new_Addr h = Some a; P \ C has_fields FDTs; h' = h(a\blank P C) \ \ P \ \new C,(h,l,sh),False\ \* \addr a,(h',l,sh),False\" (*<*) -apply(rule converse_rtrancl_into_rtrancl) - apply(erule NewInitDoneRed) -apply(rule r_into_rtrancl) -apply(erule (2) RedNew) -done +by(auto intro: converse_rtrancl_into_rtrancl + elim: NewInitDoneRed RedNew[THEN r_into_rtrancl]) (*>*) lemma NewInitDoneReds2: "\ sh C = Some (sfs, Done); new_Addr h = None; is_class P C \ \ P \ \new C,(h,l,sh),False\ \* \THROW OutOfMemory, (h,l,sh), False\" (*<*) -apply(rule_tac converse_rtrancl_into_rtrancl) - apply(erule NewInitDoneRed) -apply(rule r_into_rtrancl) -apply(erule (1) RedNewFail) -done +by(auto intro: converse_rtrancl_into_rtrancl + elim: NewInitDoneRed RedNewFail[THEN r_into_rtrancl]) (*>*) lemma NewInitReds: "\ \sfs. shp s C = Some (sfs, Done); P \ \INIT C ([C],False) \ unit,s,False\ \* \Val v',(h',l',sh'),b'\; new_Addr h' = Some a; P \ C has_fields FDTs; h'' = h'(a\blank P C); is_class P C \ \ P \ \new C,s,False\ \* \addr a,(h'',l',sh'),False\" (*<*) apply(rule_tac b = "(INIT C ([C],False) \ new C,s,False)" in converse_rtrancl_into_rtrancl) apply(cases s, simp) apply (simp add: NewInitRed) apply(erule InitSeqReds, simp_all) apply(rule r_into_rtrancl, rule RedNew) apply simp+ done (*>*) lemma NewInitOOMReds: "\ \sfs. shp s C = Some (sfs, Done); P \ \INIT C ([C],False) \ unit,s,False\ \* \Val v',(h',l',sh'),b'\; new_Addr h' = None; is_class P C \ \ P \ \new C,s,False\ \* \THROW OutOfMemory,(h',l',sh'),False\" (*<*) apply(rule_tac b = "(INIT C ([C],False) \ new C,s,False)" in converse_rtrancl_into_rtrancl) apply(cases s, simp) apply (simp add: NewInitRed) apply(erule InitSeqReds, simp_all) apply(rule r_into_rtrancl, rule RedNewFail) apply simp+ done (*>*) lemma NewInitThrowReds: "\ \sfs. shp s C = Some (sfs, Done); is_class P C; P \ \INIT C ([C],False) \ unit,s,False\ \* \throw a,s',b'\ \ \ P \ \new C,s,False\ \* \throw a,s',b'\" (*<*) apply(rule_tac b = "(INIT C ([C],False) \ new C,s,False)" in converse_rtrancl_into_rtrancl) apply(cases s, simp) apply (simp add: NewInitRed) apply(erule InitSeqThrowReds) done (*>*) subsubsection "Cast" lemma CastReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \Cast C e,s,b\ \* \Cast C e',s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule CastRed) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] CastRed[OF step(2)] + by simp +qed (*>*) lemma CastRedsNull: "P \ \e,s,b\ \* \null,s',b'\ \ P \ \Cast C e,s,b\ \* \null,s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule CastReds) apply(rule RedCastNull) done (*>*) lemma CastRedsAddr: "\ P \ \e,s,b\ \* \addr a,s',b'\; hp s' a = Some(D,fs); P \ D \\<^sup>* C \ \ P \ \Cast C e,s,b\ \* \addr a,s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule CastReds) apply(cases s',simp) apply(erule (1) RedCast) done (*>*) lemma CastRedsFail: "\ P \ \e,s,b\ \* \addr a,s',b'\; hp s' a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \ \Cast C e,s,b\ \* \THROW ClassCast,s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule CastReds) apply(cases s',simp) apply(erule (1) RedCastFail) done (*>*) lemma CastRedsThrow: "\ P \ \e,s,b\ \* \throw a,s',b'\ \ \ P \ \Cast C e,s,b\ \* \throw a,s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule CastReds) apply(rule red_reds.CastThrow) done (*>*) subsubsection "LAss" lemma LAssReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \ V:=e,s,b\ \* \ V:=e',s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule LAssRed) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] LAssRed[OF step(2)] + by simp +qed (*>*) lemma LAssRedsVal: "\ P \ \e,s,b\ \* \Val v,(h',l',sh'),b'\ \ \ P \ \ V:=e,s,b\ \* \unit,(h',l'(V\v),sh'),b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule LAssReds) apply(rule RedLAss) done (*>*) lemma LAssRedsThrow: "\ P \ \e,s,b\ \* \throw a,s',b'\ \ \ P \ \ V:=e,s,b\ \* \throw a,s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule LAssReds) apply(rule red_reds.LAssThrow) done (*>*) subsubsection "BinOp" lemma BinOp1Reds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \ e \bop\ e\<^sub>2, s,b\ \* \e' \bop\ e\<^sub>2, s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule BinOpRed1) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] BinOpRed1[OF step(2)] + by simp +qed (*>*) lemma BinOp2Reds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \(Val v) \bop\ e, s,b\ \* \(Val v) \bop\ e', s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule BinOpRed2) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] BinOpRed2[OF step(2)] + by simp +qed (*>*) lemma BinOpRedsVal: "\ P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1,b\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v\<^sub>2,s\<^sub>2,b\<^sub>2\; binop(bop,v\<^sub>1,v\<^sub>2) = Some v \ \ P \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \Val v,s\<^sub>2,b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule BinOp1Reds) apply(rule rtrancl_into_rtrancl) apply(erule BinOp2Reds) apply(rule RedBinOp) apply simp done (*>*) lemma BinOpRedsThrow1: "P \ \e,s,b\ \* \throw e',s',b'\ \ P \ \e \bop\ e\<^sub>2, s,b\ \* \throw e', s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule BinOp1Reds) apply(rule red_reds.BinOpThrow1) done (*>*) lemma BinOpRedsThrow2: "\ P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1,b\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \throw e,s\<^sub>2,b\<^sub>2\\ \ P \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0, b\<^sub>0\ \* \throw e,s\<^sub>2,b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule BinOp1Reds) apply(rule rtrancl_into_rtrancl) apply(erule BinOp2Reds) apply(rule red_reds.BinOpThrow2) done (*>*) subsubsection "FAcc" lemma FAccReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \e\F{D}, s,b\ \* \e'\F{D}, s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule FAccRed) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] FAccRed[OF step(2)] + by simp +qed (*>*) lemma FAccRedsVal: "\ P \ \e,s,b\ \* \addr a,s',b'\; hp s' a = Some(C,fs); fs(F,D) = Some v; P \ C has F,NonStatic:t in D \ \ P \ \e\F{D},s,b\ \* \Val v,s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule FAccReds) apply(cases s',simp) apply(erule (2) RedFAcc) done (*>*) lemma FAccRedsNull: "P \ \e,s,b\ \* \null,s',b'\ \ P \ \e\F{D},s,b\ \* \THROW NullPointer,s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule FAccReds) apply(rule RedFAccNull) done (*>*) lemma FAccRedsNone: "\ P \ \e,s,b\ \* \addr a,s',b'\; hp s' a = Some(C,fs); \(\b t. P \ C has F,b:t in D) \ \ P \ \e\F{D},s,b\ \* \THROW NoSuchFieldError,s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule FAccReds) apply(cases s',simp) apply(erule RedFAccNone, simp) done (*>*) lemma FAccRedsStatic: "\ P \ \e,s,b\ \* \addr a,s',b'\; hp s' a = Some(C,fs); P \ C has F,Static:t in D \ \ P \ \e\F{D},s,b\ \* \THROW IncompatibleClassChangeError,s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule FAccReds) apply(cases s',simp) apply(erule (1) RedFAccStatic) done (*>*) lemma FAccRedsThrow: "P \ \e,s,b\ \* \throw a,s',b'\ \ P \ \e\F{D},s,b\ \* \throw a,s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule FAccReds) apply(rule red_reds.FAccThrow) done (*>*) subsubsection "SFAcc" lemma SFAccReds: "\ P \ C has F,Static:t in D; shp s D = Some(sfs,Done); sfs F = Some v \ \ P \ \C\\<^sub>sF{D},s,True\ \* \Val v,s,False\" (*<*) apply(rule r_into_rtrancl) apply(cases s,simp) apply(erule (2) RedSFAcc) done (*>*) lemma SFAccRedsNone: "\(\b t. P \ C has F,b:t in D) \ P \ \C\\<^sub>sF{D},s,b\ \* \THROW NoSuchFieldError,s,False\" (*<*) apply(rule r_into_rtrancl) apply(cases s,simp) apply(rule RedSFAccNone, simp) done (*>*) lemma SFAccRedsNonStatic: "P \ C has F,NonStatic:t in D \ P \ \C\\<^sub>sF{D},s,b\ \* \THROW IncompatibleClassChangeError,s,False\" (*<*) apply(rule r_into_rtrancl) apply(cases s,simp) apply(erule RedSFAccNonStatic) done (*>*) lemma SFAccInitDoneReds: "\ P \ C has F,Static:t in D; shp s D = Some (sfs,Done); sfs F = Some v \ \ P \ \C\\<^sub>sF{D}, s,b\ \* \Val v, s,False\" (*<*) apply(cases b) \ \ case True \ apply(rule r_into_rtrancl) apply(cases s, simp) apply(erule (2) RedSFAcc) \ \ case False \ apply(rule_tac b = "(C\\<^sub>sF{D},s,True)" in converse_rtrancl_into_rtrancl) apply(cases s, simp) apply(drule (2) SFAccInitDoneRed) apply(erule SFAccReds, simp+) done (*>*) lemma SFAccInitReds: "\ P \ C has F,Static:t in D; \sfs. shp s D = Some (sfs,Done); P \ \INIT D ([D],False) \ unit,s,False\ \* \Val v',s',b'\; shp s' D = Some (sfs,i); sfs F = Some v \ \ P \ \C\\<^sub>sF{D},s,False\ \* \Val v,s',False\" (*<*) apply(rule_tac b = "(INIT D ([D],False) \ C\\<^sub>sF{D},s,False)" in converse_rtrancl_into_rtrancl) apply(cases s, simp) apply(simp add: SFAccInitRed) apply(rule InitSeqReds, simp_all) apply(subgoal_tac "\T. P \ C has F,Static:T in D") prefer 2 apply fast apply(rule r_into_rtrancl) apply(cases s', simp) apply(erule (2) RedSFAcc) done (*>*) lemma SFAccInitThrowReds: "\ P \ C has F,Static:t in D; \sfs. shp s D = Some (sfs,Done); P \ \INIT D ([D],False) \ unit,s,False\ \* \throw a,s',b'\ \ \ P \ \C\\<^sub>sF{D},s,False\ \* \throw a,s',b'\" (*<*) apply(rule_tac b = "(INIT D ([D],False) \ C\\<^sub>sF{D},s,False)" in converse_rtrancl_into_rtrancl) apply(cases s, simp) apply (simp add: SFAccInitRed) apply(erule InitSeqThrowReds) done (*>*) subsubsection "FAss" lemma FAssReds1: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \e\F{D}:=e\<^sub>2, s,b\ \* \e'\F{D}:=e\<^sub>2, s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule FAssRed1) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] FAssRed1[OF step(2)] + by simp +qed (*>*) lemma FAssReds2: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \Val v\F{D}:=e, s,b\ \* \Val v\F{D}:=e', s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule FAssRed2) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] FAssRed2[OF step(2)] + by simp +qed (*>*) lemma FAssRedsVal: "\ P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \addr a,s\<^sub>1,b\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\; P \ C has F,NonStatic:t in D; Some(C,fs) = h\<^sub>2 a \ \ P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \unit, (h\<^sub>2(a\(C,fs((F,D)\v))),l\<^sub>2,sh\<^sub>2),b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule FAssReds1) apply(rule rtrancl_into_rtrancl) apply(erule FAssReds2) apply(rule RedFAss) apply simp+ done (*>*) lemma FAssRedsNull: "\ P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \null,s\<^sub>1,b\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v,s\<^sub>2,b\<^sub>2\ \ \ P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \THROW NullPointer, s\<^sub>2,b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule FAssReds1) apply(rule rtrancl_into_rtrancl) apply(erule FAssReds2) apply(rule RedFAssNull) done (*>*) lemma FAssRedsThrow1: "P \ \e,s,b\ \* \throw e',s',b'\ \ P \ \e\F{D}:=e\<^sub>2, s,b\ \* \throw e', s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule FAssReds1) apply(rule red_reds.FAssThrow1) done (*>*) lemma FAssRedsThrow2: "\ P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Val v,s\<^sub>1,b\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \throw e,s\<^sub>2,b\<^sub>2\ \ \ P \ \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \throw e,s\<^sub>2,b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule FAssReds1) apply(rule rtrancl_into_rtrancl) apply(erule FAssReds2) apply(rule red_reds.FAssThrow2) done (*>*) lemma FAssRedsNone: "\ P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \addr a,s\<^sub>1,b\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\; h\<^sub>2 a = Some(C,fs); \(\b t. P \ C has F,b:t in D) \ \ P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \THROW NoSuchFieldError, (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule FAssReds1) apply(rule rtrancl_into_rtrancl) apply(erule FAssReds2) apply(rule RedFAssNone) apply simp+ done (*>*) lemma FAssRedsStatic: "\ P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \addr a,s\<^sub>1,b\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\; h\<^sub>2 a = Some(C,fs); P \ C has F,Static:t in D \ \ P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \THROW IncompatibleClassChangeError, (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule FAssReds1) apply(rule rtrancl_into_rtrancl) apply(erule FAssReds2) apply(rule RedFAssStatic) apply simp+ done (*>*) subsubsection "SFAss" lemma SFAssReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \C\\<^sub>sF{D}:=e,s,b\ \* \C\\<^sub>sF{D}:=e',s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule SFAssRed) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] SFAssRed[OF step(2)] + by simp +qed (*>*) lemma SFAssRedsVal: "\ P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\; P \ C has F,Static:t in D; sh\<^sub>2 D = \(sfs,Done)\ \ \ P \ \C\\<^sub>sF{D}:=e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \unit, (h\<^sub>2,l\<^sub>2,sh\<^sub>2(D\(sfs(F\v), Done))),False\" (*<*) apply(rule rtrancl_trans) apply(erule SFAssReds) apply(cases b\<^sub>2) \ \ case True \ apply(rule r_into_rtrancl) apply(drule_tac l = l\<^sub>2 in RedSFAss, simp_all) \ \ case False \ apply(rule converse_rtrancl_into_rtrancl) apply(drule_tac sh = sh\<^sub>2 in SFAssInitDoneRed, simp_all) apply(rule r_into_rtrancl) apply(drule_tac l = l\<^sub>2 in RedSFAss, simp_all) done (*>*) lemma SFAssRedsThrow: "\ P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \throw e,s\<^sub>2,b\<^sub>2\ \ \ P \ \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \throw e,s\<^sub>2,b\<^sub>2\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule SFAssReds) apply(rule red_reds.SFAssThrow) done (*>*) lemma SFAssRedsNone: "\ P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\; \(\b t. P \ C has F,b:t in D) \ \ P \ \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \THROW NoSuchFieldError, (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule SFAssReds) apply(rule RedSFAssNone) apply simp done (*>*) lemma SFAssRedsNonStatic: "\ P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\; P \ C has F,NonStatic:t in D \ \ P \ \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule SFAssReds) apply(rule RedSFAssNonStatic) apply simp done (*>*) lemma SFAssInitReds: "\ P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\; P \ C has F,Static:t in D; \sfs. sh\<^sub>2 D = Some (sfs, Done); P \ \INIT D ([D],False) \ unit,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\ \* \Val v',(h',l',sh'),b'\; sh' D = Some(sfs,i); sfs' = sfs(F\v); sh'' = sh'(D\(sfs',i)) \ \ P \ \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \unit,(h',l',sh''),False\" (*<*) apply(rule rtrancl_trans) apply(erule SFAssReds) apply(rule_tac converse_rtrancl_into_rtrancl) apply(erule (1) SFAssInitRed) apply(erule InitSeqReds, simp_all) apply(subgoal_tac "\T. P \ C has F,Static:T in D") prefer 2 apply fast apply(simp,rule r_into_rtrancl) apply(erule (2) RedSFAss) apply simp done (*>*) lemma SFAssInitThrowReds: "\ P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\; P \ C has F,Static:t in D; \sfs. sh\<^sub>2 D = Some (sfs, Done); P \ \INIT D ([D],False) \ unit,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\ \* \throw a,s',b'\ \ \ P \ \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \throw a,s',b'\" (*<*) apply(rule rtrancl_trans) apply(erule SFAssReds) apply(rule converse_rtrancl_into_rtrancl) apply(erule (1) SFAssInitRed) apply(erule InitSeqThrowReds) done (*>*) subsubsection";;" lemma SeqReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \e;;e\<^sub>2, s,b\ \* \e';;e\<^sub>2, s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule SeqRed) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] SeqRed[OF step(2)] + by simp +qed (*>*) lemma SeqRedsThrow: "P \ \e,s,b\ \* \throw e',s',b'\ \ P \ \e;;e\<^sub>2, s,b\ \* \throw e', s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule SeqReds) apply(rule red_reds.SeqThrow) done (*>*) lemma SeqReds2: "\ P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1,b\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \e\<^sub>2',s\<^sub>2,b\<^sub>2\ \ \ P \ \e\<^sub>1;;e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \e\<^sub>2',s\<^sub>2,b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule SeqReds) apply(rule converse_rtrancl_into_rtrancl) apply(rule RedSeq) apply assumption done (*>*) subsubsection"If" lemma CondReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \if (e) e\<^sub>1 else e\<^sub>2,s,b\ \* \if (e') e\<^sub>1 else e\<^sub>2,s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule CondRed) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] CondRed[OF step(2)] + by simp +qed (*>*) lemma CondRedsThrow: "P \ \e,s,b\ \* \throw a,s',b'\ \ P \ \if (e) e\<^sub>1 else e\<^sub>2, s,b\ \* \throw a,s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule CondReds) apply(rule red_reds.CondThrow) done (*>*) lemma CondReds2T: "\P \ \e,s\<^sub>0,b\<^sub>0\ \* \true,s\<^sub>1,b\<^sub>1\; P \ \e\<^sub>1, s\<^sub>1,b\<^sub>1\ \* \e',s\<^sub>2,b\<^sub>2\ \ \ P \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \e',s\<^sub>2,b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule CondReds) apply(rule converse_rtrancl_into_rtrancl) apply(rule RedCondT) apply assumption done (*>*) lemma CondReds2F: "\P \ \e,s\<^sub>0,b\<^sub>0\ \* \false,s\<^sub>1,b\<^sub>1\; P \ \e\<^sub>2, s\<^sub>1,b\<^sub>1\ \* \e',s\<^sub>2,b\<^sub>2\ \ \ P \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \e',s\<^sub>2,b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule CondReds) apply(rule converse_rtrancl_into_rtrancl) apply(rule RedCondF) apply assumption done (*>*) subsubsection "While" lemma WhileFReds: "P \ \b,s,b\<^sub>0\ \* \false,s',b'\ \ P \ \while (b) c,s,b\<^sub>0\ \* \unit,s',b'\" (*<*) apply(rule converse_rtrancl_into_rtrancl) apply(rule RedWhile) apply(rule rtrancl_into_rtrancl) apply(erule CondReds) apply(rule RedCondF) done (*>*) lemma WhileRedsThrow: "P \ \b,s,b\<^sub>0\ \* \throw e,s',b'\ \ P \ \while (b) c,s,b\<^sub>0\ \* \throw e,s',b'\" (*<*) apply(rule converse_rtrancl_into_rtrancl) apply(rule RedWhile) apply(rule rtrancl_into_rtrancl) apply(erule CondReds) apply(rule red_reds.CondThrow) done (*>*) lemma WhileTReds: "\ P \ \b,s\<^sub>0,b\<^sub>0\ \* \true,s\<^sub>1,b\<^sub>1\; P \ \c,s\<^sub>1,b\<^sub>1\ \* \Val v\<^sub>1,s\<^sub>2,b\<^sub>2\; P \ \while (b) c,s\<^sub>2,b\<^sub>2\ \* \e,s\<^sub>3,b\<^sub>3\ \ \ P \ \while (b) c,s\<^sub>0,b\<^sub>0\ \* \e,s\<^sub>3,b\<^sub>3\" (*<*) apply(rule converse_rtrancl_into_rtrancl) apply(rule RedWhile) apply(rule rtrancl_trans) apply(erule CondReds) apply(rule converse_rtrancl_into_rtrancl) apply(rule RedCondT) apply(rule rtrancl_trans) apply(erule SeqReds) apply(rule converse_rtrancl_into_rtrancl) apply(rule RedSeq) apply assumption done (*>*) lemma WhileTRedsThrow: "\ P \ \b,s\<^sub>0,b\<^sub>0\ \* \true,s\<^sub>1,b\<^sub>1\; P \ \c,s\<^sub>1,b\<^sub>1\ \* \throw e,s\<^sub>2,b\<^sub>2\ \ \ P \ \while (b) c,s\<^sub>0,b\<^sub>0\ \* \throw e,s\<^sub>2,b\<^sub>2\" (*<*) apply(rule converse_rtrancl_into_rtrancl) apply(rule RedWhile) apply(rule rtrancl_trans) apply(erule CondReds) apply(rule converse_rtrancl_into_rtrancl) apply(rule RedCondT) apply(rule rtrancl_into_rtrancl) apply(erule SeqReds) apply(rule red_reds.SeqThrow) done (*>*) subsubsection"Throw" lemma ThrowReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \throw e,s,b\ \* \throw e',s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule ThrowRed) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] ThrowRed[OF step(2)] + by simp +qed (*>*) lemma ThrowRedsNull: "P \ \e,s,b\ \* \null,s',b'\ \ P \ \throw e,s,b\ \* \THROW NullPointer,s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule ThrowReds) apply(rule RedThrowNull) done (*>*) lemma ThrowRedsThrow: "P \ \e,s,b\ \* \throw a,s',b'\ \ P \ \throw e,s,b\ \* \throw a,s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule ThrowReds) apply(rule red_reds.ThrowThrow) done (*>*) subsubsection "InitBlock" lemma InitBlockReds_aux: "P \ \e,s,b\ \* \e',s',b'\ \ \h l sh h' l' sh' v. s = (h,l(V\v),sh) \ s' = (h',l',sh') \ P \ \{V:T := Val v; e},(h,l,sh),b\ \* \{V:T := Val(the(l' V)); e'},(h',l'(V:=(l V)),sh'),b'\" (*<*) apply(erule converse_rtrancl_induct3) apply(fastforce simp: fun_upd_same simp del:fun_upd_apply) apply clarify apply(rename_tac e0 X Y x3 b0 e1 h1 l1 sh1 b1 h0 l0 sh0 h2 l2 sh2 v0) apply(subgoal_tac "V \ dom l1") prefer 2 apply(drule red_lcl_incr) apply simp apply clarsimp apply(rename_tac v1) apply(rule converse_rtrancl_into_rtrancl) apply(rule InitBlockRed) apply assumption apply simp apply(erule_tac x = "l1(V := l0 V)" in allE) apply(erule_tac x = v1 in allE) apply(erule impE) apply(rule ext) apply(simp add:fun_upd_def) apply(simp add:fun_upd_def) done (*>*) lemma InitBlockReds: "P \ \e, (h,l(V\v),sh),b\ \* \e', (h',l',sh'),b'\ \ P \ \{V:T := Val v; e}, (h,l,sh),b\ \* \{V:T := Val(the(l' V)); e'}, (h',l'(V:=(l V)),sh'),b'\" (*<*)by(blast dest:InitBlockReds_aux)(*>*) lemma InitBlockRedsFinal: "\ P \ \e,(h,l(V\v),sh),b\ \* \e',(h',l',sh'),b'\; final e' \ \ P \ \{V:T := Val v; e},(h,l,sh),b\ \* \e',(h', l'(V := l V),sh'),b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule InitBlockReds) apply(fast elim!:finalE intro:RedInitBlock InitBlockThrow) done (*>*) subsubsection "Block" lemmas converse_rtranclE3 = converse_rtranclE [of "(xa,xb,xc)" "(za,zb,zc)", split_rule] lemma BlockRedsFinal: assumes reds: "P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" and fin: "final e\<^sub>2" shows "\h\<^sub>0 l\<^sub>0 sh\<^sub>0. s\<^sub>0 = (h\<^sub>0,l\<^sub>0(V:=None),sh\<^sub>0) \ P \ \{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2(V:=l\<^sub>0 V),sh\<^sub>2),b\<^sub>2\" (*<*) using reds proof (induct rule:converse_rtrancl_induct3) case refl thus ?case by(fastforce intro:finalE[OF fin] RedBlock BlockThrow simp del:fun_upd_apply) next case (step e\<^sub>0 s\<^sub>0 b\<^sub>0 e\<^sub>1 s\<^sub>1 b\<^sub>1) have red: "P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \ \e\<^sub>1,s\<^sub>1,b\<^sub>1\" and reds: "P \ \e\<^sub>1,s\<^sub>1,b\<^sub>1\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" and IH: "\h l sh. s\<^sub>1 = (h,l(V := None),sh) \ P \ \{V:T; e\<^sub>1},(h,l,sh),b\<^sub>1\ \* \e\<^sub>2,(h\<^sub>2, l\<^sub>2(V := l V),sh\<^sub>2),b\<^sub>2\" and s\<^sub>0: "s\<^sub>0 = (h\<^sub>0, l\<^sub>0(V := None),sh\<^sub>0)" by fact+ obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" using prod_cases3 by blast show ?case proof cases assume "assigned V e\<^sub>0" then obtain v e where e\<^sub>0: "e\<^sub>0 = V := Val v;; e" by (unfold assigned_def)blast from red e\<^sub>0 s\<^sub>0 have e\<^sub>1: "e\<^sub>1 = unit;;e" and s\<^sub>1: "s\<^sub>1 = (h\<^sub>0, l\<^sub>0(V \ v),sh\<^sub>0)" and b\<^sub>1: "b\<^sub>1 = b\<^sub>0" by auto from e\<^sub>1 fin have "e\<^sub>1 \ e\<^sub>2" by (auto simp:final_def) then obtain e' s' b' where red1: "P \ \e\<^sub>1,s\<^sub>1,b\<^sub>1\ \ \e',s',b'\" and reds': "P \ \e',s',b'\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" using converse_rtranclE3[OF reds] by blast from red1 e\<^sub>1 b\<^sub>1 have es': "e' = e" "s' = s\<^sub>1" "b' = b\<^sub>0" by auto show ?case using e\<^sub>0 s\<^sub>1 es' reds' by(auto intro!: InitBlockRedsFinal[OF _ fin] simp del:fun_upd_apply) next assume unass: "\ assigned V e\<^sub>0" show ?thesis proof (cases "l\<^sub>1 V") assume None: "l\<^sub>1 V = None" hence "P \ \{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\ \ \{V:T; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\" using s\<^sub>0 s\<^sub>1 red by(simp add: BlockRedNone[OF _ _ unass]) moreover have "P \ \{V:T; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\ \* \e\<^sub>2,(h\<^sub>2, l\<^sub>2(V := l\<^sub>0 V),sh\<^sub>2),b\<^sub>2\" using IH[of _ "l\<^sub>1(V := l\<^sub>0 V)"] s\<^sub>1 None by(simp add:fun_upd_idem) ultimately show ?case by(rule converse_rtrancl_into_rtrancl) next fix v assume Some: "l\<^sub>1 V = Some v" hence "P \ \{V:T;e\<^sub>0},(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\ \ \{V:T := Val v; e\<^sub>1},(h\<^sub>1,l\<^sub>1(V := l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\" using s\<^sub>0 s\<^sub>1 red by(simp add: BlockRedSome[OF _ _ unass]) moreover have "P \ \{V:T := Val v; e\<^sub>1},(h\<^sub>1,l\<^sub>1(V:= l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2(V:=l\<^sub>0 V),sh\<^sub>2),b\<^sub>2\" using InitBlockRedsFinal[OF _ fin,of _ _ "l\<^sub>1(V:=l\<^sub>0 V)" V] Some reds s\<^sub>1 by(simp add:fun_upd_idem) ultimately show ?case by(rule converse_rtrancl_into_rtrancl) qed qed qed (*>*) subsubsection "try-catch" lemma TryReds: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \try e catch(C V) e\<^sub>2,s,b\ \* \try e' catch(C V) e\<^sub>2,s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule TryRed) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] TryRed[OF step(2)] + by simp +qed (*>*) lemma TryRedsVal: "P \ \e,s,b\ \* \Val v,s',b'\ \ P \ \try e catch(C V) e\<^sub>2,s,b\ \* \Val v,s',b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule TryReds) apply(rule RedTry) done (*>*) lemma TryCatchRedsFinal: "\ P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Throw a,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b\<^sub>1\; h\<^sub>1 a = Some(D,fs); P \ D \\<^sup>* C; P \ \e\<^sub>2, (h\<^sub>1, l\<^sub>1(V \ Addr a),sh\<^sub>1),b\<^sub>1\ \* \e\<^sub>2', (h\<^sub>2,l\<^sub>2,sh\<^sub>2), b\<^sub>2\; final e\<^sub>2' \ \ P \ \try e\<^sub>1 catch(C V) e\<^sub>2, s\<^sub>0, b\<^sub>0\ \* \e\<^sub>2', (h\<^sub>2, l\<^sub>2(V := l\<^sub>1 V),sh\<^sub>2),b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule TryReds) apply(rule converse_rtrancl_into_rtrancl) apply(rule RedTryCatch) apply fastforce apply assumption apply(rule InitBlockRedsFinal) apply assumption apply(simp) done (*>*) lemma TryRedsFail: "\ P \ \e\<^sub>1,s,b\ \* \Throw a,(h,l,sh),b'\; h a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \ \try e\<^sub>1 catch(C V) e\<^sub>2,s,b\ \* \Throw a,(h,l,sh),b'\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule TryReds) apply(fastforce intro!: RedTryFail) done (*>*) subsubsection "List" lemma ListReds1: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \e#es,s,b\ [\]* \e' # es,s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule ListRed1) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] ListRed1[OF step(2)] + by simp +qed (*>*) lemma ListReds2: "P \ \es,s,b\ [\]* \es',s',b'\ \ P \ \Val v # es,s,b\ [\]* \Val v # es',s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule ListRed2) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] ListRed2[OF step(2)] + by simp +qed (*>*) lemma ListRedsVal: "\ P \ \e,s\<^sub>0,b\<^sub>0\ \* \Val v,s\<^sub>1,b\<^sub>1\; P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \es',s\<^sub>2,b\<^sub>2\ \ \ P \ \e#es,s\<^sub>0,b\<^sub>0\ [\]* \Val v # es',s\<^sub>2,b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule ListReds1) apply(erule ListReds2) done (*>*) subsubsection"Call" text\ First a few lemmas on what happens to free variables during redction. \ lemma assumes wf: "wwf_J_prog P" shows Red_fv: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ fv e' \ fv e" and "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ fvs es' \ fvs es" (*<*) proof (induct rule:red_reds_inducts) case (RedCall h a C fs M Ts T pns body D vs l sh b) hence "fv body \ {this} \ set pns" using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def) with RedCall.hyps show ?case by fastforce next case (RedSCall C M Ts T pns body D vs) hence "fv body \ set pns" using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def) with RedSCall.hyps show ?case by fastforce qed auto (*>*) lemma Red_dom_lcl: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ dom l' \ dom l \ fv e" and "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ dom l' \ dom l \ fvs es" (*<*) proof (induct rule:red_reds_inducts) case RedLAss thus ?case by(force split:if_splits) next case CallParams thus ?case by(force split:if_splits) next case BlockRedNone thus ?case by clarsimp (fastforce split:if_splits) next case BlockRedSome thus ?case by clarsimp (fastforce split:if_splits) next case InitBlockRed thus ?case by clarsimp (fastforce split:if_splits) qed auto (*>*) lemma Reds_dom_lcl: "\ wwf_J_prog P; P \ \e,(h,l,sh),b\ \* \e',(h',l',sh'),b'\ \ \ dom l' \ dom l \ fv e" (*<*) apply(erule converse_rtrancl_induct_red) apply blast apply(blast dest: Red_fv Red_dom_lcl) done (*>*) text\ Now a few lemmas on the behaviour of blocks during reduction. \ lemma override_on_upd_lemma: "(override_on f (g(a\b)) A)(a := g a) = override_on f g (insert a A)" (*<*) apply(rule ext) apply(simp add:override_on_def) done declare fun_upd_apply[simp del] map_upds_twist[simp del] (*>*) lemma blocksReds: "\l. \ length Vs = length Ts; length vs = length Ts; distinct Vs; P \ \e, (h,l(Vs [\] vs),sh),b\ \* \e', (h',l',sh'),b'\ \ \ P \ \blocks(Vs,Ts,vs,e), (h,l,sh),b\ \* \blocks(Vs,Ts,map (the \ l') Vs,e'), (h',override_on l' l (set Vs),sh'),b'\" (*<*) proof(induct Vs Ts vs e rule:blocks_induct) case (1 V Vs T Ts v vs e) show ?case using InitBlockReds[OF "1.hyps"[of "l(V\v)"]] "1.prems" by(auto simp:override_on_upd_lemma) qed auto (*>*) lemma blocksFinal: "\l. \ length Vs = length Ts; length vs = length Ts; final e \ \ P \ \blocks(Vs,Ts,vs,e), (h,l,sh),b\ \* \e, (h,l,sh),b\" (*<*) proof(induct Vs Ts vs e rule:blocks_induct) case 1 show ?case using "1.prems" InitBlockReds[OF "1.hyps"] by(fastforce elim!:finalE elim: rtrancl_into_rtrancl[OF _ RedInitBlock] rtrancl_into_rtrancl[OF _ InitBlockThrow]) qed auto (*>*) lemma blocksRedsFinal: assumes wf: "length Vs = length Ts" "length vs = length Ts" "distinct Vs" and reds: "P \ \e, (h,l(Vs [\] vs),sh),b\ \* \e', (h',l',sh'),b'\" and fin: "final e'" and l'': "l'' = override_on l' l (set Vs)" shows "P \ \blocks(Vs,Ts,vs,e), (h,l,sh),b\ \* \e', (h',l'',sh'),b'\" (*<*) proof - let ?bv = "blocks(Vs,Ts,map (the o l') Vs,e')" have "P \ \blocks(Vs,Ts,vs,e), (h,l,sh),b\ \* \?bv, (h',l'',sh'),b'\" using l'' by simp (rule blocksReds[OF wf reds]) also have "P \ \?bv, (h',l'',sh'),b'\ \* \e', (h',l'',sh'),b'\" using wf by(fastforce intro:blocksFinal fin) finally show ?thesis . qed (*>*) text\ An now the actual method call reduction lemmas. \ lemma CallRedsObj: "P \ \e,s,b\ \* \e',s',b'\ \ P \ \e\M(es),s,b\ \* \e'\M(es),s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule CallObj) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] CallObj[OF step(2)] + by simp +qed (*>*) lemma CallRedsParams: "P \ \es,s,b\ [\]* \es',s',b'\ \ P \ \(Val v)\M(es),s,b\ \* \(Val v)\M(es'),s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule CallParams) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] CallParams[OF step(2)] + by simp +qed (*>*) lemma CallRedsFinal: assumes wwf: "wwf_J_prog P" and "P \ \e,s\<^sub>0,b\<^sub>0\ \* \addr a,s\<^sub>1,b\<^sub>1\" "P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" "h\<^sub>2 a = Some(C,fs)" "P \ C sees M,NonStatic:Ts\T = (pns,body) in D" "size vs = size pns" and l\<^sub>2': "l\<^sub>2' = [this \ Addr a, pns[\]vs]" and body: "P \ \body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),b\<^sub>2\ \* \ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\" and "final ef" shows "P \ \e\M(es), s\<^sub>0,b\<^sub>0\ \* \ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\" (*<*) proof - have wf: "size Ts = size pns \ distinct pns \ this \ set pns" and wt: "fv body \ {this} \ set pns" using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ from body[THEN Red_lcl_add, of l\<^sub>2] have body': "P \ \body,(h\<^sub>2,l\<^sub>2(this\ Addr a, pns[\]vs),sh\<^sub>2),b\<^sub>2\ \* \ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\" by (simp add:l\<^sub>2') have "dom l\<^sub>3 \ {this} \ set pns" using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns) = l\<^sub>2" by(fastforce simp add:map_add_def override_on_def fun_eq_iff) have "P \ \e\M(es),s\<^sub>0,b\<^sub>0\ \* \(addr a)\M(es),s\<^sub>1,b\<^sub>1\" by(rule CallRedsObj)(rule assms(2)) also have "P \ \(addr a)\M(es),s\<^sub>1,b\<^sub>1\ \* \(addr a)\M(map Val vs),(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" by(rule CallRedsParams)(rule assms(3)) also have "P \ \(addr a)\M(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\ \ \blocks(this#pns, Class D#Ts, Addr a#vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" by(rule RedCall)(auto simp: assms wf, rule assms(5)) also (rtrancl_into_rtrancl) have "P \ \blocks(this#pns, Class D#Ts, Addr a#vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns),sh\<^sub>3),b\<^sub>3\" by(rule blocksRedsFinal, insert assms wf body', simp_all) finally show ?thesis using eql\<^sub>2 by simp qed (*>*) lemma CallRedsThrowParams: "\ P \ \e,s\<^sub>0,b\<^sub>0\ \* \Val v,s\<^sub>1,b\<^sub>1\; P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \map Val vs\<^sub>1 @ throw a # es\<^sub>2,s\<^sub>2,b\<^sub>2\ \ \ P \ \e\M(es),s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>2,b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule CallRedsObj) apply(rule rtrancl_into_rtrancl) apply(erule CallRedsParams) apply(rule CallThrowParams) apply simp done (*>*) lemma CallRedsThrowObj: "P \ \e,s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\ \ P \ \e\M(es),s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>1,b\<^sub>1\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule CallRedsObj) apply(rule CallThrowObj) done (*>*) lemma CallRedsNull: "\ P \ \e,s\<^sub>0,b\<^sub>0\ \* \null,s\<^sub>1,b\<^sub>1\; P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \map Val vs,s\<^sub>2,b\<^sub>2\ \ \ P \ \e\M(es),s\<^sub>0,b\<^sub>0\ \* \THROW NullPointer,s\<^sub>2,b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule CallRedsObj) apply(rule rtrancl_into_rtrancl) apply(erule CallRedsParams) apply(rule RedCallNull) done (*>*) lemma CallRedsNone: "\ P \ \e,s,b\ \* \addr a,s\<^sub>1,b\<^sub>1\; P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \map Val vs,s\<^sub>2,b\<^sub>2\; hp s\<^sub>2 a = Some(C,fs); \(\b Ts T m D. P \ C sees M,b:Ts\T = m in D) \ \ P \ \e\M(es),s,b\ \* \THROW NoSuchMethodError,s\<^sub>2,b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule CallRedsObj) apply(rule rtrancl_into_rtrancl) apply(erule CallRedsParams) apply(cases s\<^sub>2,simp) apply(erule RedCallNone, simp) done (*>*) lemma CallRedsStatic: "\ P \ \e,s,b\ \* \addr a,s\<^sub>1,b\<^sub>1\; P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \map Val vs,s\<^sub>2,b\<^sub>2\; hp s\<^sub>2 a = Some(C,fs); P \ C sees M,Static:Ts\T = m in D \ \ P \ \e\M(es),s,b\ \* \THROW IncompatibleClassChangeError,s\<^sub>2,b\<^sub>2\" (*<*) apply(rule rtrancl_trans) apply(erule CallRedsObj) apply(rule rtrancl_into_rtrancl) apply(erule CallRedsParams) apply(cases s\<^sub>2,simp) apply(erule RedCallStatic, simp) done (*>*) subsection\SCall\ lemma SCallRedsParams: "P \ \es,s,b\ [\]* \es',s',b'\ \ P \ \C\\<^sub>sM(es),s,b\ \* \C\\<^sub>sM(es'),s',b'\" (*<*) -apply(erule rtrancl_induct3) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule SCallParams) -done +proof(induct rule: rtrancl_induct3) + case refl show ?case by blast +next + case step show ?case + using rtrancl_into_rtrancl[OF step(3)] SCallParams[OF step(2)] + by simp +qed (*>*) lemma SCallRedsFinal: assumes wwf: "wwf_J_prog P" and "P \ \es,s\<^sub>0,b\<^sub>0\ [\]* \map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" "P \ C sees M,Static:Ts\T = (pns,body) in D" "sh\<^sub>2 D = Some(sfs,Done) \ (M = clinit \ sh\<^sub>2 D = \(sfs, Processing)\)" "size vs = size pns" and l\<^sub>2': "l\<^sub>2' = [pns[\]vs]" and body: "P \ \body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),False\ \* \ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\" and "final ef" shows "P \ \C\\<^sub>sM(es), s\<^sub>0,b\<^sub>0\ \* \ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\" (*<*) proof - have wf: "size Ts = size pns \ distinct pns" and wt: "fv body \ set pns" using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ from body[THEN Red_lcl_add, of l\<^sub>2] have body': "P \ \body,(h\<^sub>2,l\<^sub>2(pns[\]vs),sh\<^sub>2),False\ \* \ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\" by (simp add:l\<^sub>2') have "dom l\<^sub>3 \ set pns" using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns) = l\<^sub>2" by(fastforce simp add:map_add_def override_on_def fun_eq_iff) have b2T: "P \ \C\\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\ \* \C\\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),True\" proof(cases b\<^sub>2) case True then show ?thesis by simp next case False then show ?thesis using assms(3,4) by(auto elim: SCallInitDoneRed) qed have "P \ \C\\<^sub>sM(es),s\<^sub>0,b\<^sub>0\ \* \C\\<^sub>sM(map Val vs),(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" by(rule SCallRedsParams)(rule assms(2)) also have "P \ \C\\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\ \* \blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\" by(auto intro!: rtrancl_into_rtrancl[OF b2T] RedSCall assms(3) simp: assms wf) also (rtrancl_trans) have "P \ \blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\" by(rule blocksRedsFinal, insert assms wf body', simp_all) finally show ?thesis using eql\<^sub>2 by simp qed (*>*) lemma SCallRedsThrowParams: "\ P \ \es,s\<^sub>0,b\<^sub>0\ [\]* \map Val vs\<^sub>1 @ throw a # es\<^sub>2,s\<^sub>2,b\<^sub>2\ \ \ P \ \C\\<^sub>sM(es),s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>2,b\<^sub>2\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule SCallRedsParams) apply(rule SCallThrowParams) apply simp done (*>*) lemma SCallRedsNone: "\ P \ \es,s,b\ [\]* \map Val vs,s\<^sub>2,False\; \(\b Ts T m D. P \ C sees M,b:Ts\T = m in D) \ \ P \ \C\\<^sub>sM(es),s,b\ \* \THROW NoSuchMethodError,s\<^sub>2,False\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule SCallRedsParams) apply(cases s\<^sub>2,simp) apply(rule RedSCallNone, simp) done (*>*) lemma SCallRedsNonStatic: "\ P \ \es,s,b\ [\]* \map Val vs,s\<^sub>2,False\; P \ C sees M,NonStatic:Ts\T = m in D \ \ P \ \C\\<^sub>sM(es),s,b\ \* \THROW IncompatibleClassChangeError,s\<^sub>2,False\" (*<*) apply(rule rtrancl_into_rtrancl) apply(erule SCallRedsParams) apply(cases s\<^sub>2,simp) apply(rule RedSCallNonStatic, simp) done (*>*) lemma SCallInitThrowReds: assumes wwf: "wwf_J_prog P" and "P \ \es,s\<^sub>0,b\<^sub>0\ [\]* \map Val vs,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\" "P \ C sees M,Static:Ts\T = (pns,body) in D" "\sfs. sh\<^sub>1 D = Some(sfs,Done)" "M \ clinit" and "P \ \INIT D ([D],False) \ unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\ \* \throw a,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" shows "P \ \C\\<^sub>sM(es), s\<^sub>0,b\<^sub>0\ \* \throw a,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" (*<*) proof - have "P \ \C\\<^sub>sM(es),s\<^sub>0,b\<^sub>0\ \* \C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\" by(rule SCallRedsParams)(rule assms(2)) also have "P \ \C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\ \ \INIT D ([D],False) \ C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\" using SCallInitRed[OF assms(3)] assms(4-5) by auto also (rtrancl_into_rtrancl) have "P \ \INIT D ([D],False) \ C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\ \* \throw a,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" by(rule InitSeqThrowReds[OF assms(6)]) finally show ?thesis by simp qed (*>*) lemma SCallInitReds: assumes wwf: "wwf_J_prog P" and "P \ \es,s\<^sub>0,b\<^sub>0\ [\]* \map Val vs,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\" "P \ C sees M,Static:Ts\T = (pns,body) in D" "\sfs. sh\<^sub>1 D = Some(sfs,Done)" "M \ clinit" and "P \ \INIT D ([D],False) \ unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\ \* \Val v',(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" and "size vs = size pns" and l\<^sub>2': "l\<^sub>2' = [pns[\]vs]" and body: "P \ \body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),False\ \* \ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\" and "final ef" shows "P \ \C\\<^sub>sM(es),s\<^sub>0,b\<^sub>0\ \* \ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\" (*<*) proof - have wf: "size Ts = size pns \ distinct pns" and wt: "fv body \ set pns" using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ from body[THEN Red_lcl_add, of l\<^sub>2] have body': "P \ \body,(h\<^sub>2,l\<^sub>2(pns[\]vs),sh\<^sub>2),False\ \* \ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\" by (simp add:l\<^sub>2') have "dom l\<^sub>3 \ set pns" using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns) = l\<^sub>2" by(fastforce simp add:map_add_def override_on_def fun_eq_iff) have "icheck P D (C\\<^sub>sM(map Val vs)::'a exp)" using assms(3) by auto then have "P \ \C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),icheck P D (C\\<^sub>sM(map Val vs))\ \ \blocks(pns, Ts, vs, body), (h\<^sub>2, l\<^sub>2, sh\<^sub>2), False\" by (metis (full_types) assms(3) assms(7) local.wf red_reds.RedSCall) also have "P \ \blocks(pns, Ts, vs, body), (h\<^sub>2, l\<^sub>2, sh\<^sub>2), False\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\" by(rule blocksRedsFinal, insert assms wf body', simp_all) finally have trueReds: "P \ \C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),icheck P D (C\\<^sub>sM(map Val vs))\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\" by simp have "P \ \C\\<^sub>sM(es),s\<^sub>0,b\<^sub>0\ \* \C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\" by(rule SCallRedsParams)(rule assms(2)) also have "P \ \C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\ \ \INIT D ([D],False) \ C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\" using SCallInitRed[OF assms(3)] assms(4-5) by auto also (rtrancl_into_rtrancl) have "P \ \INIT D ([D],False) \ C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\" using InitSeqReds[OF assms(6) trueReds] assms(5) by simp finally show ?thesis using eql\<^sub>2 by simp qed (*>*) lemma SCallInitProcessingReds: assumes wwf: "wwf_J_prog P" and "P \ \es,s\<^sub>0,b\<^sub>0\ [\]* \map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" "P \ C sees M,Static:Ts\T = (pns,body) in D" "sh\<^sub>2 D = Some(sfs,Processing)" and "size vs = size pns" and l\<^sub>2': "l\<^sub>2' = [pns[\]vs]" and body: "P \ \body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),False\ \* \ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\" and "final ef" shows "P \ \C\\<^sub>sM(es),s\<^sub>0,b\<^sub>0\ \* \ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\" (*<*) proof - have wf: "size Ts = size pns \ distinct pns" and wt: "fv body \ set pns" using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ from body[THEN Red_lcl_add, of l\<^sub>2] have body': "P \ \body,(h\<^sub>2,l\<^sub>2(pns[\]vs),sh\<^sub>2),False\ \* \ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\" by (simp add:l\<^sub>2') have "dom l\<^sub>3 \ set pns" using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns) = l\<^sub>2" by(fastforce simp add:map_add_def override_on_def fun_eq_iff) have b2T: "P \ \C\\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\ \* \C\\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),True\" proof(cases b\<^sub>2) case True then show ?thesis by simp next case False show ?thesis proof(cases "M = clinit") case True then show ?thesis using False assms(3) red_reds.SCallInitDoneRed assms(4) by (simp add: r_into_rtrancl) next case nclinit: False have icheck: "icheck P D (C\\<^sub>sM(map Val vs))" using assms(3) by auto have "P \ \C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b\<^sub>2\ \ \INIT D ([D],False) \ C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using SCallInitRed[OF assms(3)] assms(4) False nclinit by simp also have "P \ \INIT D ([D],False) \ C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\ \ \INIT D (Nil,True) \ C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using RedInitProcessing assms(4) by simp also have "P \ \INIT D (Nil,True) \ C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\ \ \C\\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),True\" using RedInit[of "C\\<^sub>sM(map Val vs)" D _ _ _ P] icheck nclinit by (metis (full_types) nsub_RI_Vals sub_RI.simps(12)) finally show ?thesis by simp qed qed have "P \ \C\\<^sub>sM(es),s\<^sub>0,b\<^sub>0\ \* \C\\<^sub>sM(map Val vs),(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" by(rule SCallRedsParams)(rule assms(2)) also have "P \ \C\\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\ \* \blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\" by(auto intro!: rtrancl_into_rtrancl[OF b2T] RedSCall assms(3) simp: assms wf) also (rtrancl_trans) have "P \ \blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\" by(rule blocksRedsFinal, insert assms wf body', simp_all) finally show ?thesis using eql\<^sub>2 by simp qed (*>*) (***********************************) subsubsection "The main Theorem" lemma assumes wwf: "wwf_J_prog P" shows big_by_small: "P \ \e,s\ \ \e',s'\ \ (\b. iconf (shp s) e \ P,shp s \\<^sub>b (e,b) \ \ P \ \e,s,b\ \* \e',s',False\)" and bigs_by_smalls: "P \ \es,s\ [\] \es',s'\ \ (\b. iconfs (shp s) es \ P,shp s \\<^sub>b (es,b) \ \ P \ \es,s,b\ [\]* \es',s',False\)" (*<*) proof (induct rule: eval_evals.inducts) case New show ?case proof(cases b) case True then show ?thesis using RedNew[OF New.hyps(2-4)] by auto next case False then show ?thesis using New.hyps(1) NewInitDoneReds[OF _ New.hyps(2-4)] by auto qed next case NewFail show ?case proof(cases b) case True then show ?thesis using RedNewFail[OF NewFail.hyps(2)] NewFail.hyps(3) by fastforce next case False then show ?thesis using NewInitDoneReds2[OF _ NewFail.hyps(2)] NewFail by fastforce qed next case (NewInit sh C h l v' h' l' sh' a FDTs h'') show ?case proof(cases b) case True then obtain sfs where shC: "sh C = \(sfs, Processing)\" using NewInit.hyps(1) NewInit.prems by(clarsimp simp: bconf_def initPD_def) then have s': "(h',l',sh') = (h,l,sh)" using NewInit.hyps(2) init_ProcessingE by clarsimp then show ?thesis using RedNew[OF NewInit.hyps(4-6)] True by auto next case False then have init: "P \ \INIT C ([C],False) \ unit,(h, l, sh),False\ \* \Val v',(h', l', sh'),False\" using NewInit.hyps(3) by(auto simp: bconf_def) then show ?thesis using NewInit NewInitReds[OF _ init NewInit.hyps(4-6)] False has_fields_is_class[OF NewInit.hyps(5)] by auto qed next case (NewInitOOM sh C h l v' h' l' sh') show ?case proof(cases b) case True then obtain sfs where shC: "sh C = \(sfs, Processing)\" using NewInitOOM.hyps(1) NewInitOOM.prems by(clarsimp simp: bconf_def initPD_def) then have s': "(h',l',sh') = (h,l,sh)" using NewInitOOM.hyps(2) init_ProcessingE by clarsimp then show ?thesis using RedNewFail[OF NewInitOOM.hyps(4)] True r_into_rtrancl NewInitOOM.hyps(5) by auto next case False then have init: "P \ \INIT C ([C],False) \ unit,(h, l, sh),False\ \* \Val v',(h', l', sh'),False\" using NewInitOOM.hyps(3) by(auto simp: bconf_def) then show ?thesis using NewInitOOM.hyps(1) NewInitOOMReds[OF _ init NewInitOOM.hyps(4)] False NewInitOOM.hyps(5) by auto qed next case (NewInitThrow sh C h l a s') show ?case proof(cases b) case True then obtain sfs where shC: "sh C = \(sfs, Processing)\" using NewInitThrow.hyps(1) NewInitThrow.prems by(clarsimp simp: bconf_def initPD_def) then show ?thesis using NewInitThrow.hyps(2) init_ProcessingE by blast next case False then have init: "P \ \INIT C ([C],False) \ unit,(h, l, sh),b\ \* \throw a,s',False\" using NewInitThrow.hyps(3) by(auto simp: bconf_def) then show ?thesis using NewInitThrow NewInitThrowReds[of "(h,l,sh)" C P a s'] False by auto qed next case Cast then show ?case by(fastforce intro:CastRedsAddr) next case CastNull then show ?case by(fastforce intro: CastRedsNull) next case CastFail thus ?case by(fastforce intro!:CastRedsFail) next case CastThrow thus ?case by(fastforce dest!:eval_final intro:CastRedsThrow) next case Val then show ?case using exI[of _ b] by(simp add: bconf_def) next case (BinOp e\<^sub>1 s\<^sub>0 v\<^sub>1 s\<^sub>1 e\<^sub>2 v\<^sub>2 s\<^sub>2 bop v) show ?case proof(cases "val_of e\<^sub>1") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None BinOp.prems by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using None BinOp.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \Val v\<^sub>1,s\<^sub>1,False\" using iconf BinOp.hyps(2) by auto have binop: "P \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0,b\ \* \Val v\<^sub>1 \bop\ e\<^sub>2,s\<^sub>1,False\" by(rule BinOp1Reds[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] None BinOp by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \Val v\<^sub>2,s\<^sub>2,False\" using BinOp.hyps(4)[OF iconf2'] by auto then show ?thesis using BinOpRedsVal[OF b1 b2 BinOp.hyps(5)] by fast next case (Some a) then obtain b1 where b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \Val v\<^sub>1,s\<^sub>1,b1\" by (metis (no_types, lifting) BinOp.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have binop: "P \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0,b\ \* \Val v\<^sub>1 \bop\ e\<^sub>2,s\<^sub>1,b1\" by(rule BinOp1Reds[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] BinOp by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using BinOp.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,b1) \" using Red_preserves_bconf[OF wwf binop BinOp.prems] by simp then have b2: "P \ \e\<^sub>2,s\<^sub>1,b1\ \* \Val v\<^sub>2,s\<^sub>2,False\" using BinOp.hyps(4)[OF iconf2'] by auto then show ?thesis using BinOpRedsVal[OF b1 b2 BinOp.hyps(5)] by fast qed next case (BinOpThrow1 e\<^sub>1 s\<^sub>0 e s\<^sub>1 bop e\<^sub>2) show ?case proof(cases "val_of e\<^sub>1") case None then have "iconf (shp s\<^sub>0) e\<^sub>1" and "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using BinOpThrow1.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \throw e,s\<^sub>1,False\" using BinOpThrow1.hyps(2) by auto then have "P \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0,b\ \* \throw e,s\<^sub>1,False\" using BinOpThrow1 None by(auto dest!:eval_final simp: BinOpRedsThrow1[OF b1]) then show ?thesis by fast next case (Some a) then show ?thesis using eval_final_same[OF BinOpThrow1.hyps(1)] val_of_spec[OF Some] by auto qed next case (BinOpThrow2 e\<^sub>1 s\<^sub>0 v\<^sub>1 s\<^sub>1 e\<^sub>2 e s\<^sub>2 bop) show ?case proof(cases "val_of e\<^sub>1") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None BinOpThrow2.prems by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using None BinOpThrow2.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \Val v\<^sub>1,s\<^sub>1,False\" using iconf BinOpThrow2.hyps(2) by auto have binop: "P \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0,b\ \* \Val v\<^sub>1 \bop\ e\<^sub>2,s\<^sub>1,False\" by(rule BinOp1Reds[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] None BinOpThrow2 by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \throw e,s\<^sub>2,False\" using BinOpThrow2.hyps(4)[OF iconf2'] by auto then show ?thesis using BinOpRedsThrow2[OF b1 b2] by fast next case (Some a) then obtain b1 where b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \Val v\<^sub>1,s\<^sub>1,b1\" by (metis (no_types, lifting) BinOpThrow2.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have binop: "P \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0,b\ \* \Val v\<^sub>1 \bop\ e\<^sub>2,s\<^sub>1,b1\" by(rule BinOp1Reds[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] BinOpThrow2 by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using BinOpThrow2.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,b1) \" using Red_preserves_bconf[OF wwf binop BinOpThrow2.prems] by simp then have b2: "P \ \e\<^sub>2,s\<^sub>1,b1\ \* \throw e,s\<^sub>2,False\" using BinOpThrow2.hyps(4)[OF iconf2'] by auto then show ?thesis using BinOpRedsThrow2[OF b1 b2] by fast qed next case Var thus ?case by(auto dest:RedVar simp: bconf_def) next case LAss thus ?case by(auto dest: LAssRedsVal) next case LAssThrow thus ?case by(auto dest!:eval_final dest: LAssRedsThrow) next case FAcc thus ?case by(fastforce intro:FAccRedsVal) next case FAccNull thus ?case by(auto dest:FAccRedsNull) next case FAccThrow thus ?case by(auto dest!:eval_final dest:FAccRedsThrow) next case FAccNone then show ?case by(fastforce intro: FAccRedsNone) next case FAccStatic then show ?case by(fastforce intro: FAccRedsStatic) next case SFAcc show ?case proof(cases b) case True then show ?thesis using RedSFAcc SFAcc.hyps by auto next case False then show ?thesis using SFAcc.hyps SFAccInitDoneReds[OF SFAcc.hyps(1)] by auto qed next case (SFAccInit C F t D sh h l v' h' l' sh' sfs i v) show ?case proof(cases b) case True then obtain sfs where shC: "sh D = \(sfs, Processing)\" using SFAccInit.hyps(2) SFAccInit.prems by(clarsimp simp: bconf_def initPD_def) then have s': "(h',l',sh') = (h,l,sh)" using SFAccInit.hyps(3) init_ProcessingE by clarsimp then show ?thesis using RedSFAcc SFAccInit.hyps True by auto next case False then have init: "P \ \INIT D ([D],False) \ unit,(h, l, sh),False\ \* \Val v',(h', l', sh'),False\" using SFAccInit.hyps(4) by(auto simp: bconf_def) then show ?thesis using SFAccInit SFAccInitReds[OF _ _ init] False by auto qed next case (SFAccInitThrow C F t D sh h l a s') show ?case proof(cases b) case True then obtain sfs where shC: "sh D = \(sfs, Processing)\" using SFAccInitThrow.hyps(2) SFAccInitThrow.prems(2) by(clarsimp simp: bconf_def initPD_def) then show ?thesis using SFAccInitThrow.hyps(3) init_ProcessingE by blast next case False then have init: "P \ \INIT D ([D],False) \ unit,(h, l, sh),b\ \* \throw a,s',False\" using SFAccInitThrow.hyps(4) by(auto simp: bconf_def) then show ?thesis using SFAccInitThrow SFAccInitThrowReds False by auto qed next case SFAccNone then show ?case by(fastforce intro: SFAccRedsNone) next case SFAccNonStatic then show ?case by(fastforce intro: SFAccRedsNonStatic) next case (FAss e\<^sub>1 s\<^sub>0 a s\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F t D fs' h\<^sub>2') show ?case proof(cases "val_of e\<^sub>1") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAss.prems by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using None FAss.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \addr a,s\<^sub>1,False\" using iconf FAss.hyps(2) by auto have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \addr a\F{D} := e\<^sub>2,s\<^sub>1,False\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAss by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using FAss.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsVal[OF b1 b2 FAss.hyps(6) FAss.hyps(5)[THEN sym]] FAss.hyps(7,8) by fast next case (Some a') then obtain b1 where b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \addr a,s\<^sub>1,b1\" by (metis (no_types, lifting) FAss.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \addr a\F{D} := e\<^sub>2,s\<^sub>1,b1\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAss by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using FAss.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,b1) \" using Red_preserves_bconf[OF wwf fass FAss.prems] by simp then have b2: "P \ \e\<^sub>2,s\<^sub>1,b1\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using FAss.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsVal[OF b1 b2] FAss.hyps(5)[THEN sym] FAss.hyps(6-8) by fast qed next case (FAssNull e\<^sub>1 s\<^sub>0 s\<^sub>1 e\<^sub>2 v s\<^sub>2 F D) show ?case proof(cases "val_of e\<^sub>1") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using FAssNull.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using FAssNull.prems(2) None by simp then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \null,s\<^sub>1,False\" using FAssNull.hyps(2)[OF iconf] by auto have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \null\F{D} := e\<^sub>2,s\<^sub>1,False\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssNull by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \Val v,s\<^sub>2,False\" using FAssNull.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsNull[OF b1 b2] by fast next case (Some a') then obtain b1 where b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \null,s\<^sub>1,b1\" by (metis (no_types, lifting) FAssNull.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \null\F{D} := e\<^sub>2,s\<^sub>1,b1\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssNull by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using FAssNull.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,b1) \" using Red_preserves_bconf[OF wwf fass FAssNull.prems] by simp then have b2: "P \ \e\<^sub>2,s\<^sub>1,b1\ \* \Val v,s\<^sub>2,False\" using FAssNull.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsNull[OF b1 b2] by fast qed next case (FAssThrow1 e\<^sub>1 s\<^sub>0 e' s\<^sub>1 F D e\<^sub>2) show ?case proof(cases "val_of e\<^sub>1") case None then have "iconf (shp s\<^sub>0) e\<^sub>1" and "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using FAssThrow1.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \throw e',s\<^sub>1,False\" using FAssThrow1.hyps(2) by auto then have "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \throw e',s\<^sub>1,False\" using FAssThrow1 None by(auto dest!:eval_final simp: FAssRedsThrow1[OF b1]) then show ?thesis by fast next case (Some a) then show ?thesis using eval_final_same[OF FAssThrow1.hyps(1)] val_of_spec[OF Some] by auto qed next case (FAssThrow2 e\<^sub>1 s\<^sub>0 v s\<^sub>1 e\<^sub>2 e' s\<^sub>2 F D) show ?case proof(cases "val_of e\<^sub>1") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAssThrow2.prems by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using None FAssThrow2.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \Val v,s\<^sub>1,False\" using iconf FAssThrow2.hyps(2) by auto have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \Val v\F{D} := e\<^sub>2,s\<^sub>1,False\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssThrow2 by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \throw e',s\<^sub>2,False\" using FAssThrow2.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsThrow2[OF b1 b2] by fast next case (Some a') then obtain b1 where b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \Val v,s\<^sub>1,b1\" by (metis (no_types, lifting) FAssThrow2.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \Val v\F{D} := e\<^sub>2,s\<^sub>1,b1\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssThrow2 by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using FAssThrow2.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,b1) \" using Red_preserves_bconf[OF wwf fass FAssThrow2.prems] by simp then have b2: "P \ \e\<^sub>2,s\<^sub>1,b1\ \* \throw e',s\<^sub>2,False\" using FAssThrow2.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsThrow2[OF b1 b2] by fast qed next case (FAssNone e\<^sub>1 s\<^sub>0 a s\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F D) show ?case proof(cases "val_of e\<^sub>1") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAssNone.prems by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using None FAssNone.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \addr a,s\<^sub>1,False\" using iconf FAssNone.hyps(2) by auto have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \addr a\F{D} := e\<^sub>2,s\<^sub>1,False\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssNone by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using FAssNone.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsNone[OF b1 b2 FAssNone.hyps(5,6)] by fast next case (Some a') then obtain b1 where b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \addr a,s\<^sub>1,b1\" by (metis (no_types, lifting) FAssNone.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \addr a\F{D} := e\<^sub>2,s\<^sub>1,b1\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssNone by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using FAssNone.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,b1) \" using Red_preserves_bconf[OF wwf fass FAssNone.prems] by simp then have b2: "P \ \e\<^sub>2,s\<^sub>1,b1\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using FAssNone.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsNone[OF b1 b2 FAssNone.hyps(5,6)] by fast qed next case (FAssStatic e\<^sub>1 s\<^sub>0 a s\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F t D) show ?case proof(cases "val_of e\<^sub>1") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAssStatic.prems by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>1,b) \" using None FAssStatic.prems by auto then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \addr a,s\<^sub>1,False\" using iconf FAssStatic.hyps(2) by auto have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \addr a\F{D} := e\<^sub>2,s\<^sub>1,False\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssStatic by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using FAssStatic.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsStatic[OF b1 b2 FAssStatic.hyps(5,6)] by fast next case (Some a') then obtain b1 where b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \addr a,s\<^sub>1,b1\" by (metis (no_types, lifting) FAssStatic.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\<^sub>1\F{D} := e\<^sub>2,s\<^sub>0,b\ \* \addr a\F{D} := e\<^sub>2,s\<^sub>1,b1\" by(rule FAssReds1[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssStatic by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using FAssStatic.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,b1) \" using Red_preserves_bconf[OF wwf fass FAssStatic.prems] by simp then have b2: "P \ \e\<^sub>2,s\<^sub>1,b1\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using FAssStatic.hyps(4)[OF iconf2'] by auto then show ?thesis using FAssRedsStatic[OF b1 b2 FAssStatic.hyps(5,6)] by fast qed next case (SFAss e\<^sub>2 s\<^sub>0 v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D sfs sfs' sh\<^sub>1') show ?case proof(cases "val_of e\<^sub>2") case None then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using SFAss.prems(2) by simp then have b1: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using SFAss by auto thus ?thesis using SFAssRedsVal[OF b1 SFAss.hyps(3,4)] SFAss.hyps(5,6) by fast next case (Some a) then obtain b1 where b1: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b1\" by (metis (no_types, lifting) SFAss.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) thus ?thesis using SFAssRedsVal[OF b1 SFAss.hyps(3,4)] SFAss.hyps(5,6) by fast qed next case (SFAssInit e\<^sub>2 s\<^sub>0 v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D v' h' l' sh' sfs i sfs' sh'') then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" by simp show ?case proof(cases "val_of e\<^sub>2") case None then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using SFAssInit.prems(2) by simp then have reds: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using SFAssInit.hyps(2)[OF iconf bconf] by auto then have init: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \Val v',(h', l', sh'),False\" using SFAssInit.hyps(6) by(auto simp: bconf_def) then show ?thesis using SFAssInit SFAssInitReds[OF reds SFAssInit.hyps(3) _ init] by auto next case (Some v2) show ?thesis proof(cases b) case False then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" by(simp add: bconf_def) then have reds: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using SFAssInit.hyps(2)[OF iconf bconf] by auto then have init: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \Val v',(h', l', sh'),False\" using SFAssInit.hyps(6) by(auto simp: bconf_def) then show ?thesis using SFAssInit SFAssInitReds[OF reds SFAssInit.hyps(3) _ init] by auto next case True have e\<^sub>2: "e\<^sub>2 = Val v2" using val_of_spec[OF Some] by simp then have vs: "v2 = v \ s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF SFAssInit.hyps(1)] by simp then obtain sfs where shC: "sh\<^sub>1 D = \(sfs, Processing)\" using SFAssInit.hyps(3,4) SFAssInit.prems(2) Some True by(cases e\<^sub>2, auto simp: bconf_def initPD_def dest: sees_method_fun) then have s': "(h',l',sh') = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using SFAssInit.hyps(5) init_ProcessingE by clarsimp then show ?thesis using SFAssInit.hyps(3,7-9) True e\<^sub>2 red_reds.RedSFAss vs by auto qed qed next case (SFAssInitThrow e\<^sub>2 s\<^sub>0 v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D a s') then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" by simp show ?case proof(cases "val_of e\<^sub>2") case None then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using SFAssInitThrow.prems(2) by simp then have reds: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using SFAssInitThrow.hyps(2)[OF iconf bconf] by auto then have init: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \throw a,s',False\" using SFAssInitThrow.hyps(6) by(auto simp: bconf_def) then show ?thesis using SFAssInitThrow SFAssInitThrowReds[OF reds _ _ init] by auto next case (Some v2) show ?thesis proof(cases b) case False then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" by(simp add: bconf_def) then have reds: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using SFAssInitThrow.hyps(2)[OF iconf bconf] by auto then have init: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \throw a,s',False\" using SFAssInitThrow.hyps(6) by(auto simp: bconf_def) then show ?thesis using SFAssInitThrow SFAssInitThrowReds[OF reds _ _ init] by auto next case True obtain v2 where e\<^sub>2: "e\<^sub>2 = Val v2" using val_of_spec[OF Some] by simp then have vs: "v2 = v \ s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF SFAssInitThrow.hyps(1)] by simp then obtain sfs where shC: "sh\<^sub>1 D = \(sfs, Processing)\" using SFAssInitThrow.hyps(4) SFAssInitThrow.prems(2) Some True by(cases e\<^sub>2, auto simp: bconf_def initPD_def dest: sees_method_fun) then show ?thesis using SFAssInitThrow.hyps(5) init_ProcessingE by blast qed qed next case (SFAssThrow e\<^sub>2 s\<^sub>0 e' s\<^sub>2 C F D) show ?case proof(cases "val_of e\<^sub>2") case None then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using SFAssThrow.prems(2) None by simp then have b1: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \throw e',s\<^sub>2,False\" using SFAssThrow by auto thus ?thesis using SFAssRedsThrow[OF b1] by fast next case (Some a) then show ?thesis using eval_final_same[OF SFAssThrow.hyps(1)] val_of_spec[OF Some] by auto qed next case (SFAssNone e\<^sub>2 s\<^sub>0 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C F D) show ?case proof(cases "val_of e\<^sub>2") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" using SFAssNone by simp then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using SFAssNone.prems(2) None by simp then have b1: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using SFAssNone.hyps(2) iconf by auto thus ?thesis using SFAssRedsNone[OF b1 SFAssNone.hyps(3)] by fast next case (Some a) then obtain b1 where b1: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b1\" by (metis (no_types, lifting) SFAssNone.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) thus ?thesis using SFAssRedsNone[OF b1 SFAssNone.hyps(3)] by fast qed next case (SFAssNonStatic e\<^sub>2 s\<^sub>0 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C F t D) show ?case proof(cases "val_of e\<^sub>2") case None then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" using SFAssNonStatic by simp then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>2,b) \" using SFAssNonStatic.prems(2) None by simp then have b1: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using SFAssNonStatic.hyps(2) iconf by auto thus ?thesis using SFAssRedsNonStatic[OF b1 SFAssNonStatic.hyps(3)] by fast next case (Some a) then obtain b' where b1: "P \ \e\<^sub>2,s\<^sub>0,b\ \* \Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b'\" by (metis (no_types, lifting) SFAssNonStatic.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) thus ?thesis using SFAssRedsNonStatic[OF b1 SFAssNonStatic.hyps(3)] by fast qed next case (CallObjThrow e s\<^sub>0 e' s\<^sub>1 M ps) show ?case proof(cases "val_of e") case None then have "iconf (shp s\<^sub>0) e" and "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using CallObjThrow.prems by auto then have b1: "P \ \e,s\<^sub>0,b\ \* \throw e',s\<^sub>1,False\" using CallObjThrow.hyps(2) by auto then have "P \ \e\M(ps),s\<^sub>0,b\ \* \throw e',s\<^sub>1,False\" using CallObjThrow None by(auto dest!:eval_final simp: CallRedsThrowObj[OF b1]) then show ?thesis by fast next case (Some a) then show ?thesis using eval_final_same[OF CallObjThrow.hyps(1)] val_of_spec[OF Some] by auto qed next case (CallNull e s\<^sub>0 s\<^sub>1 ps vs s\<^sub>2 M) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s\<^sub>0) e" using CallNull.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using CallNull.prems(2) None by simp then have b1: "P \ \e,s\<^sub>0,b\ \* \null,s\<^sub>1,False\" using CallNull.hyps(2)[OF iconf] by auto have call: "P \ \e\M(ps),s\<^sub>0,b\ \* \null\M(ps),s\<^sub>1,False\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None CallNull by auto have "P,shp s\<^sub>1 \\<^sub>b (ps,False) \" by(simp add: bconfs_def) then have b2: "P \ \ps,s\<^sub>1,False\ [\]* \map Val vs,s\<^sub>2,False\" using CallNull.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsNull[OF b1 b2] by fast next case (Some a') then obtain b1 where b1: "P \ \e,s\<^sub>0,b\ \* \null,s\<^sub>1,b1\" by (metis (no_types, lifting) CallNull.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\M(ps),s\<^sub>0,b\ \* \null\M(ps),s\<^sub>1,b1\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf fass] CallNull by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using CallNull.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (ps,b1) \" using Red_preserves_bconf[OF wwf fass CallNull.prems] by simp then have b2: "P \ \ps,s\<^sub>1,b1\ [\]* \map Val vs,s\<^sub>2,False\" using CallNull.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsNull[OF b1 b2] by fast qed next case (CallParamsThrow e s\<^sub>0 v s\<^sub>1 es vs ex es' s\<^sub>2 M) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s\<^sub>0) e" using CallParamsThrow.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using CallParamsThrow.prems(2) None by simp then have b1: "P \ \e,s\<^sub>0,b\ \* \Val v,s\<^sub>1,False\" using CallParamsThrow.hyps(2)[OF iconf] by auto have call: "P \ \e\M(es),s\<^sub>0,b\ \* \Val v\M(es),s\<^sub>1,False\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) es" using Red_preserves_iconf[OF wwf call] None CallParamsThrow by auto have "P,shp s\<^sub>1 \\<^sub>b (es,False) \" by(simp add: bconfs_def) then have b2: "P \ \es,s\<^sub>1,False\ [\]* \map Val vs @ throw ex # es',s\<^sub>2,False\" using CallParamsThrow.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsThrowParams[OF b1 b2] by fast next case (Some a') then obtain b1 where b1: "P \ \e,s\<^sub>0,b\ \* \Val v,s\<^sub>1,b1\" by (metis (no_types, lifting) CallParamsThrow.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\M(es),s\<^sub>0,b\ \* \Val v\M(es),s\<^sub>1,b1\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) es" using Red_preserves_iconf[OF wwf fass] CallParamsThrow by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (es,b) \" using CallParamsThrow.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (es,b1) \" using Red_preserves_bconf[OF wwf fass CallParamsThrow.prems] by simp then have b2: "P \ \es,s\<^sub>1,b1\ [\]* \map Val vs @ throw ex # es',s\<^sub>2,False\" using CallParamsThrow.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsThrowParams[OF b1 b2] by fast qed next case (CallNone e s\<^sub>0 a s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s\<^sub>0) e" using CallNone.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using CallNone.prems(2) None by simp then have b1: "P \ \e,s\<^sub>0,b\ \* \addr a,s\<^sub>1,False\" using CallNone.hyps(2)[OF iconf] by auto have call: "P \ \e\M(ps),s\<^sub>0,b\ \* \addr a\M(ps),s\<^sub>1,False\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None CallNone by auto have "P,shp s\<^sub>1 \\<^sub>b (ps,False) \" by(simp add: bconfs_def) then have b2: "P \ \ps,s\<^sub>1,False\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using CallNone.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsNone[OF b1 b2 _ CallNone.hyps(6)] CallNone.hyps(5) by fastforce next case (Some a') then obtain b1 where b1: "P \ \e,s\<^sub>0,b\ \* \addr a,s\<^sub>1,b1\" by (metis (no_types, lifting) CallNone.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fass: "P \ \e\M(ps),s\<^sub>0,b\ \* \addr a\M(ps),s\<^sub>1,b1\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf fass] CallNone by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using CallNone.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (ps,b1) \" using Red_preserves_bconf[OF wwf fass CallNone.prems] by simp then have b2: "P \ \ps,s\<^sub>1,b1\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using CallNone.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsNone[OF b1 b2 _ CallNone.hyps(6)] CallNone.hyps(5) by fastforce qed next case (CallStatic e s\<^sub>0 a s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T m D) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s\<^sub>0) e" using CallStatic.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using CallStatic.prems(2) None by simp then have b1: "P \ \e,s\<^sub>0,b\ \* \addr a,s\<^sub>1,False\" using CallStatic.hyps(2)[OF iconf] by auto have call: "P \ \e\M(ps),s\<^sub>0,b\ \* \addr a\M(ps),s\<^sub>1,False\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None CallStatic by auto have "P,shp s\<^sub>1 \\<^sub>b (ps,False) \" by(simp add: bconfs_def) then have b2: "P \ \ps,s\<^sub>1,False\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using CallStatic.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsStatic[OF b1 b2 _ CallStatic.hyps(6)] CallStatic.hyps(5) by fastforce next case (Some a') then obtain b1 where b1: "P \ \e,s\<^sub>0,b\ \* \addr a,s\<^sub>1,b1\" by (metis (no_types, lifting) CallStatic.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have call: "P \ \e\M(ps),s\<^sub>0,b\ \* \addr a\M(ps),s\<^sub>1,b1\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] CallStatic by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using CallStatic.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (ps,b1) \" using Red_preserves_bconf[OF wwf call CallStatic.prems] by simp then have b2: "P \ \ps,s\<^sub>1,b1\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using CallStatic.hyps(4)[OF iconf2'] by auto then show ?thesis using CallRedsStatic[OF b1 b2 _ CallStatic.hyps(6)] CallStatic.hyps(5) by fastforce qed next case (Call e s\<^sub>0 a s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body D l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s\<^sub>0) e" using Call.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using Call.prems(2) None by simp then have b1: "P \ \e,s\<^sub>0,b\ \* \addr a,s\<^sub>1,False\" using Call.hyps(2)[OF iconf] by auto have call: "P \ \e\M(ps),s\<^sub>0,b\ \* \addr a\M(ps),s\<^sub>1,False\" by(rule CallRedsObj[OF b1]) then have iconf2: "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None Call by auto have "P,shp s\<^sub>1 \\<^sub>b (ps,False) \" by(simp add: bconfs_def) then have b2: "P \ \ps,s\<^sub>1,False\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using Call.hyps(4)[OF iconf2] by simp have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body" by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf Call.hyps(6)]]) have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \\<^sub>b (body,False) \" by(simp add: bconf_def) then have b3: "P \ \body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using Call.hyps(10)[OF iconf3] by simp show ?thesis by(rule CallRedsFinal[OF wwf b1 b2 Call.hyps(5-8) b3 eval_final[OF Call.hyps(9)]]) next case (Some a') then obtain b1 where b1: "P \ \e,s\<^sub>0,b\ \* \addr a,s\<^sub>1,b1\" by (metis (no_types, lifting) Call.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have call: "P \ \e\M(ps),s\<^sub>0,b\ \* \addr a\M(ps),s\<^sub>1,b1\" by(rule CallRedsObj[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] Call by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using Call.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (ps,b1) \" using Red_preserves_bconf[OF wwf call Call.prems] by simp then have b2: "P \ \ps,s\<^sub>1,b1\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using Call.hyps(4)[OF iconf2'] by auto have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body" by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf Call.hyps(6)]]) have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \\<^sub>b (body,False) \" by(simp add: bconf_def) then have b3: "P \ \body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using Call.hyps(10)[OF iconf3] by simp show ?thesis by(rule CallRedsFinal[OF wwf b1 b2 Call.hyps(5-8) b3 eval_final[OF Call.hyps(9)]]) qed next case (SCallParamsThrow es s\<^sub>0 vs ex es' s\<^sub>2 C M) show ?case proof(cases "map_vals_of es") case None then have iconf: "iconfs (shp s\<^sub>0) es" using SCallParamsThrow.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (es,b) \" using SCallParamsThrow.prems(2) None by simp then have b1: "P \ \es,s\<^sub>0,b\ [\]* \map Val vs @ throw ex # es',s\<^sub>2,False\" using SCallParamsThrow.hyps(2)[OF iconf] by simp show ?thesis using SCallRedsThrowParams[OF b1] by simp next case (Some vs') then have "es = map Val vs'" by(rule map_vals_of_spec) then show ?thesis using evals_finals_same[OF _ SCallParamsThrow.hyps(1)] map_Val_nthrow_neq by auto qed next case (SCallNone ps s\<^sub>0 vs s\<^sub>2 C M) show ?case proof(cases "map_vals_of ps") case None then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallNone.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using SCallNone.prems(2) None by simp then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,s\<^sub>2,False\" using SCallNone.hyps(2)[OF iconf] by auto then show ?thesis using SCallRedsNone[OF b1 SCallNone.hyps(3)] SCallNone.hyps(1) by simp next case (Some vs') then have ps: "ps = map Val vs'" by(rule map_vals_of_spec) then have s\<^sub>0: "s\<^sub>0 = s\<^sub>2" using SCallNone.hyps(1) evals_finals_same by blast then show ?thesis using RedSCallNone[OF SCallNone.hyps(3)] ps by(cases s\<^sub>2, auto) qed next case (SCallNonStatic ps s\<^sub>0 vs s\<^sub>2 C M Ts T m D) show ?case proof(cases "map_vals_of ps") case None then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallNonStatic.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using SCallNonStatic.prems(2) None by simp then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,s\<^sub>2,False\" using SCallNonStatic.hyps(2)[OF iconf] by auto then show ?thesis using SCallRedsNonStatic[OF b1 SCallNonStatic.hyps(3)] SCallNonStatic.hyps(1) by simp next case (Some vs') then have ps: "ps = map Val vs'" by(rule map_vals_of_spec) then have s\<^sub>0: "s\<^sub>0 = s\<^sub>2" using SCallNonStatic.hyps(1) evals_finals_same by blast then show ?thesis using RedSCallNonStatic[OF SCallNonStatic.hyps(3)] ps by(cases s\<^sub>2, auto) qed next case (SCallInitThrow ps s\<^sub>0 vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D a s') show ?case proof(cases "map_vals_of ps") case None then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallInitThrow.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using SCallInitThrow.prems(2) None by simp then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using SCallInitThrow.hyps(2)[OF iconf] by auto have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \\<^sub>b (INIT D ([D],False) \ unit,False) \" by(simp add: bconf_def) then have b2: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \throw a,s',False\" using SCallInitThrow.hyps(7) by auto then show ?thesis using SCallInitThrowReds[OF wwf b1 SCallInitThrow.hyps(3-5)] by(cases s', auto) next case (Some vs') have ps: "ps = map Val vs'" by(rule map_vals_of_spec[OF Some]) then have vs: "vs = vs' \ s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using evals_finals_same[OF _ SCallInitThrow.hyps(1)] map_Val_eq by auto show ?thesis proof(cases b) case True obtain sfs where shC: "sh\<^sub>1 D = \(sfs, Processing)\" using SCallInitThrow.hyps(3,4) SCallInitThrow.prems(2) True Some vs by(auto simp: bconf_def initPD_def dest: sees_method_fun) then show ?thesis using init_ProcessingE[OF _ SCallInitThrow.hyps(6)] by blast next case False then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using ps vs by simp have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \\<^sub>b (INIT D ([D],False) \ unit,False) \" by(simp add: bconf_def) then have b2: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \throw a,s',False\" using SCallInitThrow.hyps(7) by auto then show ?thesis using SCallInitThrowReds[OF wwf b1 SCallInitThrow.hyps(3-5)] by(cases s', auto) qed qed next case (SCallInit ps s\<^sub>0 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) show ?case proof(cases "map_vals_of ps") case None then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallInit.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using SCallInit.prems(2) None by simp then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using SCallInit.hyps(2)[OF iconf] by auto have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \\<^sub>b (INIT D ([D],False) \ unit,False) \" by(simp add: bconf_def) then have b2: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \Val v',(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using SCallInit.hyps(7) by auto have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body" by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCallInit.hyps(3)]]) have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \\<^sub>b (body,False) \" by(simp add: bconf_def) then have b3: "P \ \body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using SCallInit.hyps(11)[OF iconf3] by simp show ?thesis by(rule SCallInitReds[OF wwf b1 SCallInit.hyps(3-5) b2 SCallInit.hyps(8-9) b3 eval_final[OF SCallInit.hyps(10)]]) next case (Some vs') then have ps: "ps = map Val vs'" by(rule map_vals_of_spec) then have vs: "vs = vs' \ s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using evals_finals_same[OF _ SCallInit.hyps(1)] map_Val_eq by auto show ?thesis proof(cases b) case True then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b\" using ps vs by simp obtain sfs where shC: "sh\<^sub>1 D = \(sfs, Processing)\" using SCallInit.hyps(3,4) SCallInit.prems(2) True Some vs by(auto simp: bconf_def initPD_def dest: sees_method_fun) then have s': "(h\<^sub>1, l\<^sub>1, sh\<^sub>1) = (h\<^sub>2, l\<^sub>2, sh\<^sub>2)" using init_ProcessingE[OF _ SCallInit.hyps(6)] by simp have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body" by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCallInit.hyps(3)]]) have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \\<^sub>b (body,False) \" by(simp add: bconf_def) then have b3: "P \ \body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using SCallInit.hyps(11)[OF iconf3] by simp then have b3': "P \ \body,(h\<^sub>1, l\<^sub>2', sh\<^sub>1),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using s' by simp then show ?thesis using SCallInitProcessingReds[OF wwf b1 SCallInit.hyps(3) shC SCallInit.hyps(8-9) b3' eval_final[OF SCallInit.hyps(10)]] s' by simp next case False then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using ps vs by simp have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \\<^sub>b (INIT D ([D],False) \ unit,False) \" by(simp add: bconf_def) then have b2: "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\ \* \Val v',(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using SCallInit.hyps(7) by auto have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body" by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCallInit.hyps(3)]]) have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \\<^sub>b (body,False) \" by(simp add: bconf_def) then have b3: "P \ \body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using SCallInit.hyps(11)[OF iconf3] by simp show ?thesis by(rule SCallInitReds[OF wwf b1 SCallInit.hyps(3-5) b2 SCallInit.hyps(8-9) b3 eval_final[OF SCallInit.hyps(10)]]) qed qed next case (SCall ps s\<^sub>0 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) show ?case proof(cases "map_vals_of ps") case None then have iconf: "iconfs (shp s\<^sub>0) ps" using SCall.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (ps,b) \" using SCall.prems(2) None by simp then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using SCall.hyps(2)[OF iconf] by auto have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body" by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCall.hyps(3)]]) have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \\<^sub>b (body,False) \" by(simp add: bconf_def) then have b2: "P \ \body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using SCall.hyps(8)[OF iconf3] by simp show ?thesis by(rule SCallRedsFinal[OF wwf b1 SCall.hyps(3-6) b2 eval_final[OF SCall.hyps(7)]]) next case (Some vs') then have ps: "ps = map Val vs'" by(rule map_vals_of_spec) then have vs: "vs = vs' \ s\<^sub>0 = (h\<^sub>2, l\<^sub>2, sh\<^sub>2)" using evals_finals_same[OF _ SCall.hyps(1)] map_Val_eq by auto then have b1: "P \ \ps,s\<^sub>0,b\ [\]* \map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b\" using ps by simp have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body" by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCall.hyps(3)]]) have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \\<^sub>b (body,False) \" by(simp add: bconf_def) then have b2: "P \ \body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\ \* \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\" using SCall.hyps(8)[OF iconf3] by simp show ?thesis by(rule SCallRedsFinal[OF wwf b1 SCall.hyps(3-6) b2 eval_final[OF SCall.hyps(7)]]) qed next case (Block e\<^sub>0 h\<^sub>0 l\<^sub>0 V sh\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 T) have iconf: "iconf (shp (h\<^sub>0, l\<^sub>0(V := None), sh\<^sub>0)) e\<^sub>0" using Block.prems(1) by (auto simp: assigned_def) have bconf: "P,shp (h\<^sub>0, l\<^sub>0(V := None), sh\<^sub>0) \\<^sub>b (e\<^sub>0,b) \" using Block.prems(2) by(auto simp: bconf_def) then have b': "P \ \e\<^sub>0,(h\<^sub>0, l\<^sub>0(V := None), sh\<^sub>0),b\ \* \e\<^sub>1,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" using Block.hyps(2)[OF iconf] by auto have fin: "final e\<^sub>1" using Block by(auto dest: eval_final) thus ?case using BlockRedsFinal[OF b' fin] by simp next case (Seq e\<^sub>0 s\<^sub>0 v s\<^sub>1 e\<^sub>1 e\<^sub>2 s\<^sub>2) then have iconf: "iconf (shp s\<^sub>0) e\<^sub>0" using Seq.prems(1) by(auto dest: val_of_spec lass_val_of_spec) have b1: "\b1. P \ \e\<^sub>0,s\<^sub>0,b\ \* \Val v,s\<^sub>1,b1\" proof(cases "val_of e\<^sub>0") case None show ?thesis proof(cases "lass_val_of e\<^sub>0") case lNone:None then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e\<^sub>0,b) \" using Seq.prems(2) None by simp then have "P \ \e\<^sub>0,s\<^sub>0,b\ \* \Val v,s\<^sub>1,False\" using iconf Seq.hyps(2) by auto then show ?thesis by fast next case (Some p) obtain V' v' where p: "p = (V',v')" and e\<^sub>0: "e\<^sub>0 = V':=Val v'" using lass_val_of_spec[OF Some] by(cases p, auto) obtain h l sh h' l' sh' where s\<^sub>0: "s\<^sub>0 = (h,l,sh)" and s\<^sub>1: "s\<^sub>1 = (h',l',sh')" by(cases s\<^sub>0, cases s\<^sub>1) then have eval: "P \ \e\<^sub>0,(h,l,sh)\ \ \Val v,(h',l',sh')\" using Seq.hyps(1) by simp then have s\<^sub>1': "Val v = unit \ h' = h \ l' = l(V' \ v') \ sh' = sh" using lass_val_of_eval[OF Some eval] p e\<^sub>0 by simp then have "P \ \e\<^sub>0,s\<^sub>0,b\ \ \Val v,s\<^sub>1,b\" using e\<^sub>0 s\<^sub>0 s\<^sub>1 by(auto intro: RedLAss) then show ?thesis by auto qed next case (Some a) then have "e\<^sub>0 = Val v" and "s\<^sub>0 = s\<^sub>1" using Seq.hyps(1) eval_cases(3) val_of_spec by blast+ then show ?thesis using Seq by auto qed then obtain b1 where b1': "P \ \e\<^sub>0,s\<^sub>0,b\ \* \Val v,s\<^sub>1,b1\" by clarsimp have seq: "P \ \e\<^sub>0;;e\<^sub>1,s\<^sub>0,b\ \* \Val v;;e\<^sub>1,s\<^sub>1,b1\" by(rule SeqReds[OF b1']) then have iconf2: "iconf (shp s\<^sub>1) e\<^sub>1" using Red_preserves_iconf[OF wwf seq] Seq nsub_RI_iconf by auto have "P,shp s\<^sub>1 \\<^sub>b (Val v;; e\<^sub>1,b1) \" by(rule Red_preserves_bconf[OF wwf seq Seq.prems]) then have bconf2: "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>1,b1) \" by simp have b2: "P \ \e\<^sub>1,s\<^sub>1,b1\ \* \e\<^sub>2,s\<^sub>2,False\" by(rule Seq.hyps(4)[OF iconf2 bconf2]) then show ?case using SeqReds2[OF b1' b2] by fast next case (SeqThrow e\<^sub>0 s\<^sub>0 a s\<^sub>1 e\<^sub>1 b) have notVal: "val_of e\<^sub>0 = None" "lass_val_of e\<^sub>0 = None" using SeqThrow.hyps(1) eval_throw_nonVal eval_throw_nonLAss by auto thus ?case using SeqThrow notVal by(auto dest!:eval_final dest: SeqRedsThrow) next case (CondT e s\<^sub>0 s\<^sub>1 e\<^sub>1 e' s\<^sub>2 e\<^sub>2) then have iconf: "iconf (shp s\<^sub>0) e" using CondT.prems(1) by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using CondT.prems(2) by auto then have b1: "P \ \e,s\<^sub>0,b\ \* \true,s\<^sub>1,False\" using iconf CondT.hyps(2) by auto have cond: "P \ \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0,b\ \* \if (true) e\<^sub>1 else e\<^sub>2,s\<^sub>1,False\" by(rule CondReds[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>1" using Red_preserves_iconf[OF wwf cond] CondT nsub_RI_iconf by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>1,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>1,s\<^sub>1,False\ \* \e',s\<^sub>2,False\" using CondT.hyps(4)[OF iconf2'] by auto then show ?case using CondReds2T[OF b1 b2] by fast next case (CondF e s\<^sub>0 s\<^sub>1 e\<^sub>2 e' s\<^sub>2 e\<^sub>1) then have iconf: "iconf (shp s\<^sub>0) e" using CondF.prems(1) by auto have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using CondF.prems(2) by auto then have b1: "P \ \e,s\<^sub>0,b\ \* \false,s\<^sub>1,False\" using iconf CondF.hyps(2) by auto have cond: "P \ \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0,b\ \* \if (false) e\<^sub>1 else e\<^sub>2,s\<^sub>1,False\" by(rule CondReds[OF b1]) then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf cond] CondF nsub_RI_iconf by auto have "P,shp s\<^sub>1 \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,s\<^sub>1,False\ \* \e',s\<^sub>2,False\" using CondF.hyps(4)[OF iconf2'] by auto then show ?case using CondReds2F[OF b1 b2] by fast next case CondThrow thus ?case by(auto dest!:eval_final dest:CondRedsThrow) next case (WhileF e s\<^sub>0 s\<^sub>1 c) then have iconf: "iconf (shp s\<^sub>0) e" using nsub_RI_iconf by auto then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using WhileF.prems(2) by(simp add: bconf_def) then have b': "P \ \e,s\<^sub>0,b\ \* \false,s\<^sub>1,False\" using WhileF.hyps(2) iconf by auto thus ?case using WhileFReds[OF b'] by fast next case (WhileT e s\<^sub>0 s\<^sub>1 c v\<^sub>1 s\<^sub>2 e\<^sub>3 s\<^sub>3) then have iconf: "iconf (shp s\<^sub>0) e" using nsub_RI_iconf by auto then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using WhileT.prems(2) by(simp add: bconf_def) then have b1: "P \ \e,s\<^sub>0,b\ \* \true,s\<^sub>1,False\" using WhileT.hyps(2) iconf by auto have iconf2: "iconf (shp s\<^sub>1) c" using WhileT.prems(1) nsub_RI_iconf by auto have bconf2: "P,shp s\<^sub>1 \\<^sub>b (c,False) \" by(simp add: bconf_def) then have b2: "P \ \c,s\<^sub>1,False\ \* \Val v\<^sub>1,s\<^sub>2,False\" using WhileT.hyps(4) iconf2 by auto have iconf3: "iconf (shp s\<^sub>2) (while (e) c)" using WhileT.prems(1) by auto have "P,shp s\<^sub>2 \\<^sub>b (while (e) c,False) \" by(simp add: bconf_def) then have b3: "P \ \while (e) c,s\<^sub>2,False\ \* \e\<^sub>3,s\<^sub>3,False\" using WhileT.hyps(6) iconf3 by auto show ?case using WhileTReds[OF b1 b2 b3] by fast next case WhileCondThrow thus ?case by (metis (no_types, lifting) WhileRedsThrow iconf.simps(16) bconf_While bconf_def nsub_RI_iconf) next case (WhileBodyThrow e s\<^sub>0 s\<^sub>1 c e' s\<^sub>2) then have iconf: "iconf (shp s\<^sub>0) e" using nsub_RI_iconf by auto then have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using WhileBodyThrow.prems(2) by(simp add: bconf_def) then have b1: "P \ \e,s\<^sub>0,b\ \* \true,s\<^sub>1,False\" using WhileBodyThrow.hyps(2) iconf by auto have iconf2: "iconf (shp s\<^sub>1) c" using WhileBodyThrow.prems(1) nsub_RI_iconf by auto have bconf2: "P,shp s\<^sub>1 \\<^sub>b (c,False) \" by(simp add: bconf_def) then have b2: "P \ \c,s\<^sub>1,False\ \* \throw e',s\<^sub>2,False\" using WhileBodyThrow.hyps(4) iconf2 by auto show ?case using WhileTRedsThrow[OF b1 b2] by fast next case Throw thus ?case by (meson ThrowReds iconf.simps(17) bconf_Throw) next case ThrowNull thus ?case by (meson ThrowRedsNull iconf.simps(17) bconf_Throw) next case ThrowThrow thus ?case by (meson ThrowRedsThrow iconf.simps(17) bconf_Throw) next case Try thus ?case by (meson TryRedsVal iconf.simps(18) bconf_Try) next case (TryCatch e\<^sub>1 s\<^sub>0 a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C e\<^sub>2 V e\<^sub>2' h\<^sub>2 l\<^sub>2 sh\<^sub>2) then have b1: "P \ \e\<^sub>1,s\<^sub>0,b\ \* \Throw a,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" by auto have Try: "P \ \try e\<^sub>1 catch(C V) e\<^sub>2,s\<^sub>0,b\ \* \try (Throw a) catch(C V) e\<^sub>2,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\" by(rule TryReds[OF b1]) have iconf: "iconf sh\<^sub>1 e\<^sub>2" using Red_preserves_iconf[OF wwf Try] TryCatch nsub_RI_iconf by auto have bconf: "P,shp (h\<^sub>1, l\<^sub>1(V \ Addr a), sh\<^sub>1) \\<^sub>b (e\<^sub>2,False) \" by(simp add: bconf_def) then have b2: "P \ \e\<^sub>2,(h\<^sub>1, l\<^sub>1(V \ Addr a), sh\<^sub>1),False\ \* \e\<^sub>2',(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\" using TryCatch.hyps(6) iconf by auto thus ?case using TryCatchRedsFinal[OF b1] TryCatch.hyps(3-5) by (meson eval_final) next case TryThrow thus ?case by (meson TryRedsFail iconf.simps(18) bconf_Try) next case Nil thus ?case by(auto simp: bconfs_def) next case (Cons e s\<^sub>0 v s\<^sub>1 es es' s\<^sub>2) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s\<^sub>0) e" using Cons.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using Cons.prems(2) None by simp then have b1: "P \ \e,s\<^sub>0,b\ \* \Val v,s\<^sub>1,False\" using Cons.hyps(2) iconf by auto have cons: "P \ \e # es,s\<^sub>0,b\ [\]* \Val v # es,s\<^sub>1,False\" by(rule ListReds1[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) es" using Reds_preserves_iconf[OF wwf cons] None Cons by auto have "P,shp s\<^sub>1 \\<^sub>b (es,False) \" by(simp add: bconfs_def) then have b2: "P \ \es,s\<^sub>1,False\ [\]* \es',s\<^sub>2,False\" using Cons.hyps(4)[OF iconf2'] by auto show ?thesis using ListRedsVal[OF b1 b2] by auto next case (Some a) then obtain b1 where b1: "P \ \e,s\<^sub>0,b\ \* \Val v,s\<^sub>1,b1\" by (metis (no_types, lifting) Cons.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have cons: "P \ \e # es,s\<^sub>0,b\ [\]* \Val v # es,s\<^sub>1,b1\" by(rule ListReds1[OF b1]) then have iconf2': "iconfs (shp s\<^sub>1) es" using Reds_preserves_iconf[OF wwf cons] Cons by auto have bconf2: "P,shp s\<^sub>0 \\<^sub>b (es,b) \" using Cons.prems Some by simp then have "P,shp s\<^sub>1 \\<^sub>b (es,b1) \" using Reds_preserves_bconf[OF wwf cons Cons.prems] by simp then have b2: "P \ \es,s\<^sub>1,b1\ [\]* \es',s\<^sub>2,False\" using Cons.hyps(4)[OF iconf2'] by auto show ?thesis using ListRedsVal[OF b1 b2] by auto qed next case (ConsThrow e s\<^sub>0 e' s\<^sub>1 es) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s\<^sub>0) e" using ConsThrow.prems(1) by simp have bconf: "P,shp s\<^sub>0 \\<^sub>b (e,b) \" using ConsThrow.prems(2) None by simp then have b1: "P \ \e,s\<^sub>0,b\ \* \throw e',s\<^sub>1,False\" using ConsThrow.hyps(2) iconf by auto have cons: "P \ \e # es,s\<^sub>0,b\ [\]* \throw e' # es,s\<^sub>1,False\" by(rule ListReds1[OF b1]) then show ?thesis by fast next case (Some a) then show ?thesis using eval_final_same[OF ConsThrow.hyps(1)] val_of_spec[OF Some] by auto qed next case (InitFinal e s e' s' C b') then have "\sub_RI e" by simp then show ?case using InitFinal RedInit[of e C b' s b P] by (meson converse_rtrancl_into_rtrancl nsub_RI_iconf red_preserves_bconf RedInit) next case (InitNone sh C C' Cs e h l e' s') then have init: "P \ \INIT C' (C # Cs,False) \ e,(h, l, sh(C \ (sblank P C, Prepared))),b\ \* \e',s',False\" by(simp add: bconf_def) show ?case by(rule InitNoneReds[OF InitNone.hyps(1) init]) next case (InitDone sh C sfs C' Cs e h l e' s') then have "iconf (shp (h, l, sh)) (INIT C' (Cs,True) \ e)" using InitDone.hyps(1) proof(cases Cs) case Nil then have "C = C'" "\sub_RI e" using InitDone.prems(1) by simp+ then show ?thesis using Nil InitDone.hyps(1) by(simp add: initPD_def) qed(auto) then have init: "P \ \INIT C' (Cs,True) \ e,(h, l, sh),b\ \* \e',s',False\" using InitDone by(simp add: bconf_def) show ?case by(rule InitDoneReds[OF InitDone.hyps(1) init]) next case (InitProcessing sh C sfs C' Cs e h l e' s') then have "iconf (shp (h, l, sh)) (INIT C' (Cs,True) \ e)" using InitProcessing.hyps(1) proof(cases Cs) case Nil then have "C = C'" "\sub_RI e" using InitProcessing.prems(1) by simp+ then show ?thesis using Nil InitProcessing.hyps(1) by(simp add: initPD_def) qed(auto) then have init: "P \ \INIT C' (Cs,True) \ e,(h, l, sh),b\ \* \e',s',False\" using InitProcessing by(simp add: bconf_def) show ?case by(rule InitProcessingReds[OF InitProcessing.hyps(1) init]) next case InitError thus ?case by(fastforce intro: InitErrorReds simp: bconf_def) next case InitObject thus ?case by(fastforce intro: InitObjectReds simp: bconf_def) next case InitNonObject thus ?case by(fastforce intro: InitNonObjectReds simp: bconf_def) next case InitRInit thus ?case by(fastforce intro: RedsInitRInit simp: bconf_def) next case (RInit e s v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 s\<^sub>1) then have iconf2: "iconf (shp (h', l', sh'')) (INIT C' (Cs,True) \ e')" by(auto simp: initPD_def fun_upd_same list_nonempty_induct) show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s) e" using RInit.prems(1) by simp have bconf: "P,shp s \\<^sub>b (e,b) \" using RInit.prems(2) None by simp then have b1: "P \ \e,s,b\ \* \Val v,(h',l',sh'),False\" using RInit.hyps(2)[OF iconf] by auto have "P,shp (h', l', sh'') \\<^sub>b (INIT C' (Cs,True) \ e',False) \" by(simp add: bconf_def) then have b2: "P \ \INIT C' (Cs,True) \ e',(h',l',sh''),False\ \* \e\<^sub>1,s\<^sub>1,False\" using RInit.hyps(7)[OF iconf2] by auto then show ?thesis using RedsRInit[OF b1 RInit.hyps(3-5) b2] by fast next case (Some a') then obtain b1 where b1: "P \ \e,s,b\ \* \Val v,(h',l',sh'),b1\" by (metis (no_types, lifting) RInit.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fin: "final e" by(simp add: val_of_spec[OF Some]) have "\b" using RInit.prems(2) Some by(simp add: bconf_def) then have nb1: "\b1" using reds_final_same[OF b1 fin] by simp have "P,shp (h', l', sh'') \\<^sub>b (INIT C' (Cs,True) \ e',b1) \" using nb1 by(simp add: bconf_def) then have b2: "P \ \INIT C' (Cs,True) \ e',(h', l', sh''),b1\ \* \e\<^sub>1,s\<^sub>1,False\" using RInit.hyps(7)[OF iconf2] by auto then show ?thesis using RedsRInit[OF b1 RInit.hyps(3-5) b2] by fast qed next case (RInitInitFail e s a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 s\<^sub>1) have fin: "final (throw a)" using eval_final[OF RInitInitFail.hyps(1)] by simp then obtain a' where a': "throw a = Throw a'" by auto have iconf2: "iconf (shp (h', l', sh'')) (RI (D,Throw a') ; Cs \ e')" using RInitInitFail.prems(1) by auto show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s) e" using RInitInitFail.prems(1) by simp have bconf: "P,shp s \\<^sub>b (e,b) \" using RInitInitFail.prems(2) None by simp then have b1: "P \ \e,s,b\ \* \Throw a',(h',l',sh'),False\" using RInitInitFail.hyps(2)[OF iconf] a' by auto have "P,shp (h', l', sh'') \\<^sub>b (RI (D,Throw a') ; Cs \ e',False) \" by(simp add: bconf_def) then have b2: "P \ \RI (D,Throw a') ; Cs \ e',(h',l',sh''),False\ \* \e\<^sub>1,s\<^sub>1,False\" using RInitInitFail.hyps(6) iconf2 a' by auto show ?thesis using RInitInitThrowReds[OF b1 RInitInitFail.hyps(3-4) b2] by fast next case (Some a1) then obtain b1 where b1: "P \ \e,s,b\ \* \Throw a',(h',l',sh'),b1\" using a' by (metis (no_types, lifting) RInitInitFail.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec) have fin: "final e" by(simp add: val_of_spec[OF Some]) have "\b" using RInitInitFail.prems(2) Some by(simp add: bconf_def) then have nb1: "\b1" using reds_final_same[OF b1 fin] by simp have "P,shp (h', l', sh'') \\<^sub>b (RI (D,Throw a') ; Cs \ e',b1) \" using nb1 by(simp add: bconf_def) then have b2: "P \ \RI (D,Throw a') ; Cs \ e',(h', l', sh''),b1\ \* \e\<^sub>1,s\<^sub>1,False\" using RInitInitFail.hyps(6) iconf2 a' by auto show ?thesis using RInitInitThrowReds[OF b1 RInitInitFail.hyps(3-4) b2] by fast qed next case (RInitFailFinal e s a h' l' sh' C sfs i sh'' e') have fin: "final (throw a)" using eval_final[OF RInitFailFinal.hyps(1)] by simp then obtain a' where a': "throw a = Throw a'" by auto show ?case proof(cases "val_of e") case None then have iconf: "iconf (shp s) e" using RInitFailFinal.prems(1) by simp have bconf: "P,shp s \\<^sub>b (e,b) \" using RInitFailFinal.prems(2) None by simp then have b1: "P \ \e,s,b\ \* \Throw a',(h',l',sh'),False\" using RInitFailFinal.hyps(2)[OF iconf] a' by auto show ?thesis using RInitThrowReds[OF b1 RInitFailFinal.hyps(3-4)] a' by fast next case (Some a1) then show ?thesis using eval_final_same[OF RInitFailFinal.hyps(1)] val_of_spec[OF Some] by auto qed qed (*>*) subsection\Big steps simulates small step\ text\ This direction was carried out by Norbert Schirmer and Daniel Wasserrab (and modified to include statics and DCI by Susannah Mansky). \ text \ The big step equivalent of @{text RedWhile}: \ lemma unfold_while: "P \ \while(b) c,s\ \ \e',s'\ = P \ \if(b) (c;;while(b) c) else (unit),s\ \ \e',s'\" (*<*) proof assume "P \ \while (b) c,s\ \ \e',s'\" thus "P \ \if (b) (c;; while (b) c) else unit,s\ \ \e',s'\" by cases (fastforce intro: eval_evals.intros)+ next assume "P \ \if (b) (c;; while (b) c) else unit,s\ \ \e',s'\" thus "P \ \while (b) c,s\ \ \e',s'\" proof (cases) fix a assume e': "e' = throw a" assume "P \ \b,s\ \ \throw a,s'\" hence "P \ \while(b) c,s\ \ \throw a,s'\" by (rule WhileCondThrow) with e' show ?thesis by simp next fix s\<^sub>1 assume eval_false: "P \ \b,s\ \ \false,s\<^sub>1\" and eval_unit: "P \ \unit,s\<^sub>1\ \ \e',s'\" with eval_unit have "s' = s\<^sub>1" "e' = unit" by (auto elim: eval_cases) moreover from eval_false have "P \ \while (b) c,s\ \ \unit,s\<^sub>1\" by - (rule WhileF, simp) ultimately show ?thesis by simp next fix s\<^sub>1 assume eval_true: "P \ \b,s\ \ \true,s\<^sub>1\" and eval_rest: "P \ \c;; while (b) c,s\<^sub>1\\\e',s'\" from eval_rest show ?thesis proof (cases) fix s\<^sub>2 v\<^sub>1 assume "P \ \c,s\<^sub>1\ \ \Val v\<^sub>1,s\<^sub>2\" "P \ \while (b) c,s\<^sub>2\ \ \e',s'\" with eval_true show "P \ \while(b) c,s\ \ \e',s'\" by (rule WhileT) next fix a assume "P \ \c,s\<^sub>1\ \ \throw a,s'\" "e' = throw a" with eval_true show "P \ \while(b) c,s\ \ \e',s'\" by (iprover intro: WhileBodyThrow) qed qed qed (*>*) lemma blocksEval: "\Ts vs l l'. \size ps = size Ts; size ps = size vs; P \ \blocks(ps,Ts,vs,e),(h,l,sh)\ \ \e',(h',l',sh')\ \ \ \ l''. P \ \e,(h,l(ps[\]vs),sh)\ \ \e',(h',l'',sh')\" (*<*) proof (induct ps) case Nil then show ?case by fastforce next case (Cons p ps') have length_eqs: "length (p # ps') = length Ts" "length (p # ps') = length vs" by fact+ then obtain T Ts' where Ts: "Ts = T#Ts'" by (cases "Ts") simp obtain v vs' where vs: "vs = v#vs'" using length_eqs by (cases "vs") simp have "P \ \blocks (p # ps', Ts, vs, e),(h,l,sh)\ \ \e',(h', l',sh')\" by fact with Ts vs have "P \ \{p:T := Val v; blocks (ps', Ts', vs', e)},(h,l,sh)\ \ \e',(h', l',sh')\" by simp then obtain l''' where eval_ps': "P \ \blocks (ps', Ts', vs', e),(h, l(p\v), sh)\ \ \e',(h', l''', sh')\" and l''': "l'=l'''(p:=l p)" by (auto elim!: eval_cases) then obtain l'' where hyp: "P \ \e,(h, l(p\v)(ps'[\]vs'), sh)\ \ \e',(h', l'', sh')\" using length_eqs Ts vs Cons.hyps [OF _ _ eval_ps'] by auto from hyp show "\l''. P \ \e,(h, l(p # ps'[\]vs), sh)\ \ \e',(h', l'', sh')\" using Ts vs by auto qed (*>*) lemma assumes wf: "wwf_J_prog P" shows eval_restrict_lcl: "P \ \e,(h,l,sh)\ \ \e',(h',l',sh')\ \ (\W. fv e \ W \ P \ \e,(h,l|`W,sh)\ \ \e',(h',l'|`W,sh')\)" and "P \ \es,(h,l,sh)\ [\] \es',(h',l',sh')\ \ (\W. fvs es \ W \ P \ \es,(h,l|`W,sh)\ [\] \es',(h',l'|`W,sh')\)" (*<*) proof(induct rule:eval_evals_inducts) case (Block e\<^sub>0 h\<^sub>0 l\<^sub>0 V sh\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 T) have IH: "\W. fv e\<^sub>0 \ W \ P \ \e\<^sub>0,(h\<^sub>0,l\<^sub>0(V:=None)|`W,sh\<^sub>0)\ \ \e\<^sub>1,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\" by fact have "fv({V:T; e\<^sub>0}) \ W" by fact+ hence "fv e\<^sub>0 - {V} \ W" by simp_all hence "fv e\<^sub>0 \ insert V W" by fast from IH[OF this] have "P \ \e\<^sub>0,(h\<^sub>0, (l\<^sub>0|`W)(V := None), sh\<^sub>0)\ \ \e\<^sub>1,(h\<^sub>1, l\<^sub>1|`insert V W, sh\<^sub>1)\" by fastforce from eval_evals.Block[OF this] show ?case by fastforce next case Seq thus ?case by simp (blast intro:eval_evals.Seq) next case New thus ?case by(simp add:eval_evals.intros) next case NewFail thus ?case by(simp add:eval_evals.intros) next case (NewInit sh C h l v' h' l' sh' a h'') have "fv(INIT C ([C],False) \ unit) \ W" by simp then have "P \ \INIT C ([C],False) \ unit,(h, l |` W, sh)\ \ \Val v',(h', l' |` W, sh')\" by (simp add: NewInit.hyps(3)) thus ?case using NewInit.hyps(1,4-6) eval_evals.NewInit by blast next case (NewInitOOM sh C h l v' h' l' sh') have "fv(INIT C ([C],False) \ unit) \ W" by simp then have "P \ \INIT C ([C],False) \ unit,(h, l |` W, sh)\ \ \Val v',(h', l' |` W, sh')\" by (simp add: NewInitOOM.hyps(3)) thus ?case using NewInitOOM.hyps(1,4,5) eval_evals.NewInitOOM by auto next case NewInitThrow thus ?case by(simp add:eval_evals.intros) next case Cast thus ?case by simp (blast intro:eval_evals.Cast) next case CastNull thus ?case by simp (blast intro:eval_evals.CastNull) next case CastFail thus ?case by simp (blast intro:eval_evals.CastFail) next case CastThrow thus ?case by(simp add:eval_evals.intros) next case Val thus ?case by(simp add:eval_evals.intros) next case BinOp thus ?case by simp (blast intro:eval_evals.BinOp) next case BinOpThrow1 thus ?case by simp (blast intro:eval_evals.BinOpThrow1) next case BinOpThrow2 thus ?case by simp (blast intro:eval_evals.BinOpThrow2) next case Var thus ?case by(simp add:eval_evals.intros) next case (LAss e h\<^sub>0 l\<^sub>0 sh\<^sub>0 v h l sh l' V) have IH: "\W. fv e \ W \ P \ \e,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\ \ \Val v,(h,l|`W,sh)\" and [simp]: "l' = l(V \ v)" by fact+ have "fv (V:=e) \ W" by fact hence fv: "fv e \ W" and VinW: "V \ W" by auto from eval_evals.LAss[OF IH[OF fv] refl, of V] VinW show ?case by simp next case LAssThrow thus ?case by(fastforce intro: eval_evals.LAssThrow) next case FAcc thus ?case by simp (blast intro: eval_evals.FAcc) next case FAccNull thus ?case by(fastforce intro: eval_evals.FAccNull) next case FAccThrow thus ?case by(fastforce intro: eval_evals.FAccThrow) next case FAccNone thus ?case by(metis eval_evals.FAccNone fv.simps(7)) next case FAccStatic thus ?case by(metis eval_evals.FAccStatic fv.simps(7)) next case SFAcc thus ?case by simp (blast intro: eval_evals.SFAcc) next case SFAccInit thus ?case by simp (blast intro: eval_evals.SFAccInit) next case SFAccInitThrow thus ?case by simp (blast intro: eval_evals.SFAccInitThrow) next case SFAccNone thus ?case by simp (blast intro: eval_evals.SFAccNone) next case SFAccNonStatic thus ?case by simp (blast intro: eval_evals.SFAccNonStatic) next case FAss thus ?case by simp (blast intro: eval_evals.FAss) next case FAssNull thus ?case by simp (blast intro: eval_evals.FAssNull) next case FAssThrow1 thus ?case by simp (blast intro: eval_evals.FAssThrow1) next case FAssThrow2 thus ?case by simp (blast intro: eval_evals.FAssThrow2) next case FAssNone thus ?case by simp (blast intro: eval_evals.FAssNone) next case FAssStatic thus ?case by simp (blast intro: eval_evals.FAssStatic) next case SFAss thus ?case by simp (blast intro: eval_evals.SFAss) next case SFAssInit thus ?case by simp (blast intro: eval_evals.SFAssInit) next case SFAssInitThrow thus ?case by simp (blast intro: eval_evals.SFAssInitThrow) next case SFAssThrow thus ?case by simp (blast intro: eval_evals.SFAssThrow) next case SFAssNone thus ?case by simp (blast intro: eval_evals.SFAssNone) next case SFAssNonStatic thus ?case by simp (blast intro: eval_evals.SFAssNonStatic) next case CallObjThrow thus ?case by simp (blast intro: eval_evals.intros) next case CallNull thus ?case by simp (blast intro: eval_evals.CallNull) next case (CallNone e h l sh a h' l' sh' ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M) have f1: "P \ \e,(h, l |` W, sh)\ \ \addr a,(h', l' |` W, sh')\" by (metis (no_types) fv.simps(11) le_sup_iff local.CallNone(2) local.CallNone(7)) have "P \ \ps,(h', l' |` W, sh')\ [\] \map Val vs, (h\<^sub>2, l\<^sub>2 |` W, sh\<^sub>2)\" using local.CallNone(4) local.CallNone(7) by auto then show ?case using f1 eval_evals.CallNone local.CallNone(5) local.CallNone(6) by auto next case CallStatic thus ?case by (metis (no_types, lifting) eval_evals.CallStatic fv.simps(11) le_sup_iff) next case CallParamsThrow thus ?case by simp (blast intro: eval_evals.CallParamsThrow) next case (Call e h\<^sub>0 l\<^sub>0 sh\<^sub>0 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 l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) have IHe: "\W. fv e \ W \ P \ \e,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\ \ \addr a,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\" and IHps: "\W. fvs ps \ W \ P \ \ps,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2|`W,sh\<^sub>2)\" and IHbd: "\W. fv body \ W \ P \ \body,(h\<^sub>2,l\<^sub>2'|`W,sh\<^sub>2)\ \ \e',(h\<^sub>3,l\<^sub>3|`W,sh\<^sub>3)\" and h\<^sub>2a: "h\<^sub>2 a = Some (C, fs)" and "method": "P \ C sees M,NonStatic: Ts\T = (pns, body) in D" and same_len: "size vs = size pns" and l\<^sub>2': "l\<^sub>2' = [this \ Addr a, pns [\] vs]" by fact+ have "fv (e\M(ps)) \ W" by fact hence fve: "fv e \ W" and fvps: "fvs(ps) \ W" by auto have wfmethod: "size Ts = size pns \ this \ set pns" and fvbd: "fv body \ {this} \ set pns" using "method" wf by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ show ?case using IHbd[OF fvbd] l\<^sub>2' same_len wfmethod h\<^sub>2a eval_evals.Call[OF IHe[OF fve] IHps[OF fvps] _ "method" same_len l\<^sub>2'] by (simp add:subset_insertI) next case (SCallNone ps h l sh vs h' l' sh' C M) have "P \ \ps,(h, l |` W, sh)\ [\] \map Val vs,(h', l' |` W, sh')\" using SCallNone.hyps(2) SCallNone.prems by auto then show ?case using SCallNone.hyps(3) eval_evals.SCallNone by auto next case SCallNonStatic thus ?case by (metis eval_evals.SCallNonStatic fv.simps(12)) next case SCallParamsThrow thus ?case by simp (blast intro: eval_evals.SCallParamsThrow) next case SCallInitThrow thus ?case by simp (blast intro: eval_evals.SCallInitThrow) next case SCallInit thus ?case by simp (blast intro: eval_evals.SCallInit) next case (SCall ps h\<^sub>0 l\<^sub>0 sh\<^sub>0 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 IHps: "\W. fvs ps \ W \ P \ \ps,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2|`W,sh\<^sub>2)\" and IHbd: "\W. fv body \ W \ P \ \body,(h\<^sub>2,l\<^sub>2'|`W,sh\<^sub>2)\ \ \e',(h\<^sub>3,l\<^sub>3|`W,sh\<^sub>3)\" and sh\<^sub>2D: "sh\<^sub>2 D = Some (sfs, Done) \ M = clinit \ sh\<^sub>2 D = \(sfs, Processing)\" and "method": "P \ C sees M,Static: Ts\T = (pns, body) in D" and same_len: "size vs = size pns" and l\<^sub>2': "l\<^sub>2' = [pns [\] vs]" by fact+ have "fv (C\\<^sub>sM(ps)) \ W" by fact hence fvps: "fvs(ps) \ W" by auto have wfmethod: "size Ts = size pns" and fvbd: "fv body \ set pns" using "method" wf by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ show ?case using IHbd[OF fvbd] l\<^sub>2' same_len wfmethod sh\<^sub>2D eval_evals.SCall[OF IHps[OF fvps] "method" _ same_len l\<^sub>2'] by (simp add:subset_insertI) next case SeqThrow thus ?case by simp (blast intro: eval_evals.SeqThrow) next case CondT thus ?case by simp (blast intro: eval_evals.CondT) next case CondF thus ?case by simp (blast intro: eval_evals.CondF) next case CondThrow thus ?case by simp (blast intro: eval_evals.CondThrow) next case WhileF thus ?case by simp (blast intro: eval_evals.WhileF) next case WhileT thus ?case by simp (blast intro: eval_evals.WhileT) next case WhileCondThrow thus ?case by simp (blast intro: eval_evals.WhileCondThrow) next case WhileBodyThrow thus ?case by simp (blast intro: eval_evals.WhileBodyThrow) next case Throw thus ?case by simp (blast intro: eval_evals.Throw) next case ThrowNull thus ?case by simp (blast intro: eval_evals.ThrowNull) next case ThrowThrow thus ?case by simp (blast intro: eval_evals.ThrowThrow) next case Try thus ?case by simp (blast intro: eval_evals.Try) next case (TryCatch e\<^sub>1 h\<^sub>0 l\<^sub>0 sh\<^sub>0 a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C e\<^sub>2 V e\<^sub>2' h\<^sub>2 l\<^sub>2 sh\<^sub>2) have IH\<^sub>1: "\W. fv e\<^sub>1 \ W \ P \ \e\<^sub>1,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\ \ \Throw a,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\" and IH\<^sub>2: "\W. fv e\<^sub>2 \ W \ P \ \e\<^sub>2,(h\<^sub>1,l\<^sub>1(V\Addr a)|`W,sh\<^sub>1)\ \ \e\<^sub>2',(h\<^sub>2,l\<^sub>2|`W,sh\<^sub>2)\" and lookup: "h\<^sub>1 a = Some(D, fs)" and subtype: "P \ D \\<^sup>* C" by fact+ have "fv (try e\<^sub>1 catch(C V) e\<^sub>2) \ W" by fact hence fv\<^sub>1: "fv e\<^sub>1 \ W" and fv\<^sub>2: "fv e\<^sub>2 \ insert V W" by auto have IH\<^sub>2': "P \ \e\<^sub>2,(h\<^sub>1,(l\<^sub>1|`W)(V \ Addr a),sh\<^sub>1)\ \ \e\<^sub>2',(h\<^sub>2,l\<^sub>2|`insert V W,sh\<^sub>2)\" using IH\<^sub>2[OF fv\<^sub>2] fun_upd_restrict[of l\<^sub>1 W] (*FIXME just l|W instead of l|(W-V) in simp rule??*) by simp with eval_evals.TryCatch[OF IH\<^sub>1[OF fv\<^sub>1] _ subtype IH\<^sub>2'] lookup show ?case by fastforce next case TryThrow thus ?case by simp (blast intro: eval_evals.TryThrow) next case Nil thus ?case by (simp add: eval_evals.Nil) next case Cons thus ?case by simp (blast intro: eval_evals.Cons) next case ConsThrow thus ?case by simp (blast intro: eval_evals.ConsThrow) next case InitFinal thus ?case by (simp add: eval_evals.InitFinal) next case InitNone thus ?case by(blast intro: eval_evals.InitNone) next case InitDone thus ?case by (simp add: InitDone.hyps(2) InitDone.prems eval_evals.InitDone) next case InitProcessing thus ?case by (simp add: eval_evals.InitProcessing) next case InitError thus ?case using eval_evals.InitError by auto next case InitObject thus ?case by(simp add: eval_evals.InitObject) next case InitNonObject thus ?case by(simp add: eval_evals.InitNonObject) next case InitRInit thus ?case by(simp add: eval_evals.InitRInit) 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 f1: "fv e \ W \ fv (INIT C' (Cs,True) \ e') \ W" using RInit.prems by auto then have f2: "P \ \e,(h, l |` W, sh)\ \ \Val v,(h', l' |` W, sh')\" using RInit.hyps(2) by blast have "P \ \INIT C' (Cs,True) \ e', (h', l' |` W, sh'')\ \ \e\<^sub>1,(h\<^sub>1, l\<^sub>1 |` W, sh\<^sub>1)\" using f1 by (meson RInit.hyps(7)) then show ?case using f2 RInit.hyps(3) RInit.hyps(4) RInit.hyps(5) eval_evals.RInit by blast 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 f1: "fv e \ W" "fv e' \ W" using RInitInitFail.prems by auto then have f2: "P \ \e,(h, l |` W, sh)\ \ \throw a,(h', l' |` W, sh')\" using RInitInitFail.hyps(2) by blast then have f2': "fv (throw a) \ W" using eval_final[OF f2] by auto then have f1': "fv (RI (D,throw a);Cs \ e') \ W" using f1 by auto have "P \ \RI (D,throw a);Cs \ e', (h', l' |` W, sh'')\ \ \e\<^sub>1,(h\<^sub>1, l\<^sub>1 |` W, sh\<^sub>1)\" using f1' by (meson RInitInitFail.hyps(6)) then show ?case using f2 by (simp add: RInitInitFail.hyps(3,4) eval_evals.RInitInitFail) next case (RInitFailFinal e h l sh a h' l' sh' sh'' C) have f1: "fv e \ W" using RInitFailFinal.prems by auto then have f2: "P \ \e,(h, l |` W, sh)\ \ \throw a,(h', l' |` W, sh')\" using RInitFailFinal.hyps(2) by blast then have f2': "fv (throw a) \ W" using eval_final[OF f2] by auto then show ?case using f2 RInitFailFinal.hyps(3,4) eval_evals.RInitFailFinal by blast qed (*>*) lemma eval_notfree_unchanged: "P \ \e,(h,l,sh)\ \ \e',(h',l',sh')\ \ (\V. V \ fv e \ l' V = l V)" and "P \ \es,(h,l,sh)\ [\] \es',(h',l',sh')\ \ (\V. V \ fvs es \ l' V = l V)" (*<*) proof(induct rule:eval_evals_inducts) case LAss thus ?case by(simp add:fun_upd_apply) next case Block thus ?case by (simp only:fun_upd_apply split:if_splits) fastforce next case TryCatch thus ?case by (simp only:fun_upd_apply split:if_splits) fastforce next case (RInitInitFail 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 "fv (throw a) = {}" using RInitInitFail.hyps(1) eval_final final_fv by blast then have "fv e \ fv (RI (D,throw a) ; Cs \ unit) \ fv (RI (C,e) ; D#Cs \ unit)" by auto then show ?case using RInitInitFail.hyps(2,6) RInitInitFail.prems by fastforce qed simp_all (*>*) lemma eval_closed_lcl_unchanged: "\ P \ \e,(h,l,sh)\ \ \e',(h',l',sh')\; fv e = {} \ \ l' = l" (*<*)by(fastforce dest:eval_notfree_unchanged simp add:fun_eq_iff [where 'b="val option"])(*>*) lemma list_eval_Throw: assumes eval_e: "P \ \throw x,s\ \ \e',s'\" shows "P \ \map Val vs @ throw x # es',s\ [\] \map Val vs @ e' # es',s'\" (*<*) proof - from eval_e obtain a where e': "e' = Throw a" by (cases) (auto dest!: eval_final) { fix es have "\vs. es = map Val vs @ throw x # es' \ P \ \es,s\[\]\map Val vs @ e' # es',s'\" proof (induct es type: list) case Nil thus ?case by simp next case (Cons e es vs) have e_es: "e # es = map Val vs @ throw x # es'" by fact show "P \ \e # es,s\ [\] \map Val vs @ e' # es',s'\" proof (cases vs) case Nil with e_es obtain "e=throw x" "es=es'" by simp moreover from eval_e e' have "P \ \throw x # es,s\ [\] \Throw a # es,s'\" by (iprover intro: ConsThrow) ultimately show ?thesis using Nil e' by simp next case (Cons v vs') have vs: "vs = v # vs'" by fact with e_es obtain e: "e=Val v" and es:"es= map Val vs' @ throw x # es'" by simp from e have "P \ \e,s\ \ \Val v,s\" by (iprover intro: eval_evals.Val) moreover from es have "P \ \es,s\ [\] \map Val vs' @ e' # es',s'\" by (rule Cons.hyps) ultimately show "P \ \e#es,s\ [\] \map Val vs @ e' # es',s'\" using vs by (auto intro: eval_evals.Cons) qed qed } thus ?thesis by simp qed (*>*) \ \ separate evaluation of first subexp of a sequence \ lemma seq_ext: assumes IH: "\e' s'. P \ \e'',s''\ \ \e',s'\ \ P \ \e,s\ \ \e',s'\" and seq: "P \ \e'' ;; e\<^sub>0,s''\ \ \e',s'\" shows "P \ \e ;; e\<^sub>0,s\ \ \e',s'\" proof(rule eval_cases(14)[OF seq]) \ \ Seq \ fix v' s\<^sub>1 assume e'': "P \ \e'',s''\ \ \Val v',s\<^sub>1\" and estep: "P \ \e\<^sub>0,s\<^sub>1\ \ \e',s'\" have "P \ \e,s\ \ \Val v',s\<^sub>1\" using e'' IH by simp then show ?thesis using estep Seq by simp next fix e\<^sub>t assume e'': "P \ \e'',s''\ \ \throw e\<^sub>t,s'\" and e': "e' = throw e\<^sub>t" have "P \ \e,s\ \ \throw e\<^sub>t,s'\" using e'' IH by simp then show ?thesis using eval_evals.SeqThrow e' by simp qed \ \ separate evaluation of @{text RI} subexp, val case \ lemma rinit_Val_ext: assumes ri: "P \ \RI (C,e'') ; Cs \ e\<^sub>0,s''\ \ \Val v',s\<^sub>1\" and IH: "\e' s'. P \ \e'',s''\ \ \e',s'\ \ P \ \e,s\ \ \e',s'\" shows "P \ \RI (C,e) ; Cs \ e\<^sub>0,s\ \ \Val v',s\<^sub>1\" proof(rule eval_cases(20)[OF ri]) \ \ RI \ fix v'' h' l' sh' sfs i assume e''step: "P \ \e'',s''\ \ \Val v'',(h', l', sh')\" and shC: "sh' C = \(sfs, i)\" and init: "P \ \INIT (if Cs = [] then C else last Cs) (Cs,True) \ e\<^sub>0,(h', l', sh'(C \ (sfs, Done)))\ \ \Val v',s\<^sub>1\" have "P \ \e,s\ \ \Val v'',(h', l', sh')\" using IH[OF e''step] by simp then show ?thesis using RInit init shC by auto next fix a h' l' sh' sfs i D Cs' assume e''step: "P \ \e'',s''\ \ \throw a,(h', l', sh')\" and riD: "P \ \RI (D,throw a) ; Cs' \ e\<^sub>0,(h', l', sh'(C \ (sfs, Error)))\ \ \Val v',s\<^sub>1\" have "P \ \e,s\ \ \throw a,(h', l', sh')\" using IH[OF e''step] by simp then show ?thesis using riD rinit_throwE by blast qed(simp) \ \ separate evaluation of @{text RI} subexp, throw case \ lemma rinit_throw_ext: assumes ri: "P \ \RI (C,e'') ; Cs \ e\<^sub>0,s''\ \ \throw e\<^sub>t,s'\" and IH: "\e' s'. P \ \e'',s''\ \ \e',s'\ \ P \ \e,s\ \ \e',s'\" shows "P \ \RI (C,e) ; Cs \ e\<^sub>0,s\ \ \throw e\<^sub>t,s'\" proof(rule eval_cases(20)[OF ri]) \ \ RI \ fix v h' l' sh' sfs i assume e''step: "P \ \e'',s''\ \ \Val v,(h', l', sh')\" and shC: "sh' C = \(sfs, i)\" and init: "P \ \INIT (if Cs = [] then C else last Cs) (Cs,True) \ e\<^sub>0,(h', l', sh'(C \ (sfs, Done)))\ \ \throw e\<^sub>t,s'\" have "P \ \e,s\ \ \Val v,(h', l', sh')\" using IH[OF e''step] by simp then show ?thesis using RInit init shC by auto next fix a h' l' sh' sfs i D Cs' assume e''step: "P \ \e'',s''\ \ \throw a,(h', l', sh')\" and shC: "sh' C = \(sfs, i)\" and riD: "P \ \RI (D,throw a) ; Cs' \ e\<^sub>0,(h', l', sh'(C \ (sfs, Error)))\ \ \throw e\<^sub>t,s'\" and cons: "Cs = D # Cs'" have estep': "P \ \e,s\ \ \throw a,(h', l', sh')\" using IH[OF e''step] by simp then show ?thesis using RInitInitFail cons riD shC by simp next fix a h' l' sh' sfs i assume "throw e\<^sub>t = throw a" and "s' = (h', l', sh'(C \ (sfs, Error)))" and "P \ \e'',s''\ \ \throw a,(h', l', sh')\" and "sh' C = \(sfs, i)\" and "Cs = []" then show ?thesis using RInitFailFinal IH by auto qed \ \ separate evaluation of @{text RI} subexp \ lemma rinit_ext: assumes IH: "\e' s'. P \ \e'',s''\ \ \e',s'\ \ P \ \e,s\ \ \e',s'\" shows "\e' s'. P \ \RI (C,e'') ; Cs \ e\<^sub>0,s''\ \ \e',s'\ \ P \ \RI (C,e) ; Cs \ e\<^sub>0,s\ \ \e',s'\" proof - fix e' s' assume ri'': "P \ \RI (C,e'') ; Cs \ e\<^sub>0,s''\ \ \e',s'\" then have "final e'" using eval_final by simp then show "P \ \RI (C,e) ; Cs \ e\<^sub>0,s\ \ \e',s'\" proof(rule finalE) fix v assume "e' = Val v" then show ?thesis using rinit_Val_ext[OF _ IH] ri'' by simp next fix a assume "e' = throw a" then show ?thesis using rinit_throw_ext[OF _ IH] ri'' by simp qed qed \ \ @{text INIT} and @{text RI} return either @{text Val} with @{text Done} or @{text Processing} flag or @{text Throw} with @{text Error} flag \ lemma shows eval_init_return: "P \ \e,s\ \ \e',s'\ \ iconf (shp 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 s' C' = \(sfs,i)\ \ (i = Done \ i = Processing))) \ (throw_of e' = Some a \ (\sfs i. shp s' C' = \(sfs,Error)\))" and "P \ \es,s\ [\] \es',s'\ \ True" proof(induct rule: eval_evals.inducts) case (InitFinal e s e' s' C b) then show ?case by(fastforce simp: initPD_def dest: eval_final_same) next case (InitDone sh C sfs C' Cs e h l e' s') then have "final e'" using eval_final by simp then show ?case proof(rule finalE) fix v assume e': "e' = Val v" then show ?thesis using InitDone initPD_def proof(cases Cs) qed(auto) next fix a assume e': "e' = throw a" then show ?thesis using InitDone initPD_def proof(cases Cs) qed(auto) qed next case (InitProcessing sh C sfs C' Cs e h l e' s') then have "final e'" using eval_final by simp then show ?case proof(rule finalE) fix v assume e': "e' = Val v" then show ?thesis using InitProcessing initPD_def proof(cases Cs) qed(auto) next fix a assume e': "e' = throw a" then show ?thesis using InitProcessing initPD_def proof(cases Cs) qed(auto) qed next case (InitError sh C sfs Cs e h l e' s' C') show ?case proof(cases Cs) case Nil then show ?thesis using InitError by simp next case (Cons C2 list) then have "final e'" using InitError eval_final by simp then show ?thesis proof(rule finalE) fix v assume e': "e' = Val v" then show ?thesis using InitError proof - obtain ccss :: "char list list" and bb :: bool where "INIT C' (C # Cs,False) \ e = INIT C' (ccss,bb) \ unit" using InitError.prems(2) by blast then show ?thesis using InitError.hyps(2) e' by(auto dest!: rinit_throwE) qed next fix a assume e': "e' = throw a" then show ?thesis using Cons InitError cons_to_append[of list] by clarsimp qed qed next case (InitRInit C Cs h l sh e' s' C') show ?case proof(cases Cs) case Nil then show ?thesis using InitRInit by simp next case (Cons C' list) then show ?thesis using InitRInit Cons cons_to_append[of list] by clarsimp qed next case (RInit 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_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 Nil by(auto simp: fun_upd_same initPD_def) next fix a assume e': "e\<^sub>1 = throw a" show ?thesis using RInit Nil by(auto simp: fun_upd_same initPD_def) qed next case (Cons a list) show ?thesis proof(rule finalE[OF final]) fix v assume e': "e\<^sub>1 = Val v" then show ?thesis using RInit 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 Cons by clarsimp (metis last.simps last_appendR list.distinct(1)) qed qed next case (RInitInitFail 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_final by simp then show ?case proof(rule finalE) fix v assume e': "e\<^sub>1 = Val v" then show ?thesis using RInitInitFail by clarsimp (meson exp.distinct(101) rinit_throwE) next fix a' assume e': "e\<^sub>1 = Throw a'" then have "iconf (sh'(C \ (sfs, Error))) a" using RInitInitFail.hyps(1) eval_final by fastforce then show ?thesis using RInitInitFail e' by clarsimp (meson Cons_eq_append_conv list.inject) qed qed(auto simp: fun_upd_same) lemma init_Val_PD: "P \ \INIT C' (Cs,b) \ unit,s\ \ \Val v,s'\ \ iconf (shp s) (INIT C' (Cs,b) \ unit) \ \sfs i. shp s' C' = \(sfs,i)\ \ (i = Done \ i = Processing)" by(drule_tac v = v in eval_init_return, simp+) lemma init_throw_PD: "P \ \INIT C' (Cs,b) \ unit,s\ \ \throw a,s'\ \ iconf (shp s) (INIT C' (Cs,b) \ unit) \ \sfs i. shp s' C' = \(sfs,Error)\" by(drule_tac a = a in eval_init_return, simp+) lemma rinit_Val_PD: "P \ \RI(C,e\<^sub>0);Cs \ unit,s\ \ \Val v,s'\ \ iconf (shp s) (RI(C,e\<^sub>0);Cs \ unit) \ last(C#Cs) = C' \ \sfs i. shp s' C' = \(sfs,i)\ \ (i = Done \ i = Processing)" apply(drule_tac C' = C' and v = v in eval_init_return, simp_all) apply (metis append_butlast_last_id) done lemma rinit_throw_PD: "P \ \RI(C,e\<^sub>0);Cs \ unit,s\ \ \throw a,s'\ \ iconf (shp s) (RI(C,e\<^sub>0);Cs \ unit) \ last(C#Cs) = C' \ \sfs i. shp s' C' = \(sfs,Error)\" apply(drule_tac C' = C' and a = a in eval_init_return, simp_all) apply (metis append_butlast_last_id) done \ \ combining the above to get evaluation of @{text INIT} in a sequence \ (* Hiermit kann man die ganze pair-Splitterei in den automatischen Taktiken abschalten. Wieder anschalten siehe nach dem Beweis. *) (*<*) declare split_paired_All [simp del] split_paired_Ex [simp del] (*>*) lemma eval_init_seq': "P \ \e,s\ \ \e',s'\ \ (\C Cs b e\<^sub>i. e = INIT C (Cs,b) \ e\<^sub>i) \ (\C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs \ e\<^sub>i) \ (\C Cs b e\<^sub>i. e = INIT C (Cs,b) \ e\<^sub>i \ P \ \(INIT C (Cs,b) \ unit);; e\<^sub>i,s\ \ \e',s'\) \ (\C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs \ e\<^sub>i \ P \ \(RI(C,e\<^sub>0);Cs \ unit);; e\<^sub>i,s\ \ \e',s'\)" and "P \ \es,s\ [\] \es',s'\ \ True" proof(induct rule: eval_evals.inducts) case InitFinal then show ?case by(auto simp: Seq[OF eval_evals.InitFinal[OF Val[where v=Unit]]]) next case (InitNone sh) then show ?case using seq_ext[OF eval_evals.InitNone[where sh=sh and e=unit, OF InitNone.hyps(1)]] by fastforce next case (InitDone sh) then show ?case using seq_ext[OF eval_evals.InitDone[where sh=sh and e=unit, OF InitDone.hyps(1)]] by fastforce next case (InitProcessing sh) then show ?case using seq_ext[OF eval_evals.InitProcessing[where sh=sh and e=unit, OF InitProcessing.hyps(1)]] by fastforce next case (InitError sh) then show ?case using seq_ext[OF eval_evals.InitError[where sh=sh and e=unit, OF InitError.hyps(1)]] by fastforce next case (InitObject sh) then show ?case using seq_ext[OF eval_evals.InitObject[where sh=sh and e=unit, OF InitObject.hyps(1)]] by fastforce next case (InitNonObject sh) then show ?case using seq_ext[OF eval_evals.InitNonObject[where sh=sh and e=unit, OF InitNonObject.hyps(1)]] by fastforce next case (InitRInit C Cs e h l sh e' s' C') then show ?case using seq_ext[OF eval_evals.InitRInit[where e=unit]] by fastforce next case RInit then show ?case using seq_ext[OF eval_evals.RInit[where e=unit, OF RInit.hyps(1)]] by fastforce next case RInitInitFail then show ?case using seq_ext[OF eval_evals.RInitInitFail[where e=unit, OF RInitInitFail.hyps(1)]] by fastforce next case RInitFailFinal then show ?case using eval_evals.RInitFailFinal eval_evals.SeqThrow by auto qed(auto) lemma eval_init_seq: "P \ \INIT C (Cs,b) \ e,(h, l, sh)\ \ \e',s'\ \ P \ \(INIT C (Cs,b) \ unit);; e,(h, l, sh)\ \ \e',s'\" by(auto dest: eval_init_seq') text \ The key lemma: \ lemma assumes wf: "wwf_J_prog P" shows extend_1_eval: "P \ \e,s,b\ \ \e'',s'',b''\ \ P,shp s \\<^sub>b (e,b) \ \ (\s' e'. P \ \e'',s''\ \ \e',s'\ \ P \ \e,s\ \ \e',s'\)" and extend_1_evals: "P \ \es,s,b\ [\] \es'',s'',b''\ \ P,shp s \\<^sub>b (es,b) \ \ (\s' es'. P \ \es'',s''\ [\] \es',s'\ \ P \ \es,s\ [\] \es',s'\)" proof (induct rule: red_reds.inducts) case (RedNew h a C FDTs h' l sh) then have e':"e' = addr a" and s':"s' = (h(a \ blank P C), l, sh)" using eval_cases(3) by fastforce+ obtain sfs i where shC: "sh C = \(sfs, i)\" and "i = Done \ i = Processing" using RedNew by (clarsimp simp: bconf_def initPD_def) then show ?case proof(cases i) case Done then show ?thesis using RedNew shC e' s' New by simp next case Processing then have shC': "\sfs. sh C = Some(sfs,Done)" and shP: "sh C = Some(sfs,Processing)" using shC by simp+ then have init: "P \ \INIT C ([C],False) \ unit,(h,l,sh)\ \ \unit,(h,l,sh)\" by(simp add: InitFinal InitProcessing Val) have "P \ \new C,(h, l, sh)\ \ \addr a,(h(a \ blank P C),l,sh)\" using RedNew shC' by(auto intro: NewInit[OF _ init]) then show ?thesis using e' s' by simp qed(auto) next case (RedNewFail h C l sh) then have e':"e' = THROW OutOfMemory" and s':"s' = (h, l, sh)" using eval_final_same final_def by fastforce+ obtain sfs i where shC: "sh C = \(sfs, i)\" and "i = Done \ i = Processing" using RedNewFail by (clarsimp simp: bconf_def initPD_def) then show ?case proof(cases i) case Done then show ?thesis using RedNewFail shC e' s' NewFail by simp next case Processing then have shC': "\sfs. sh C = Some(sfs,Done)" and shP: "sh C = Some(sfs,Processing)" using shC by simp+ then have init: "P \ \INIT C ([C],False) \ unit,(h,l,sh)\ \ \unit,(h,l,sh)\" by(simp add: InitFinal InitProcessing Val) have "P \ \new C,(h, l, sh)\ \ \THROW OutOfMemory,(h,l,sh)\" using RedNewFail shC' by(auto intro: NewInitOOM[OF _ init]) then show ?thesis using e' s' by simp qed(auto) next case (NewInitRed sh C h l) then have seq: "P \ \(INIT C ([C],False) \ unit);; new C,(h, l, sh)\ \ \e',s'\" using eval_init_seq by simp then show ?case proof(rule eval_cases(14)) \ \ Seq \ fix v s\<^sub>1 assume init: "P \ \INIT C ([C],False) \ unit,(h, l, sh)\ \ \Val v,s\<^sub>1\" and new: "P \ \new C,s\<^sub>1\ \ \e',s'\" obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1) then obtain sfs i where shC: "sh\<^sub>1 C = \(sfs, i)\" and iDP: "i = Done \ i = Processing" using init_Val_PD[OF init] by auto show ?thesis proof(rule eval_cases(1)[OF new]) \ \ New \ fix sha ha a FDTs la assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = addr a" and s': "s' = (ha(a \ blank P C), la, sha)" and addr: "new_Addr ha = \a\" and fields: "P \ C has_fields FDTs" then show ?thesis using NewInit[OF _ _ addr fields] NewInitRed.hyps init by simp next fix sha ha la assume "s\<^sub>1 = (ha, la, sha)" and "e' = THROW OutOfMemory" and "s' = (ha, la, sha)" and "new_Addr ha = None" then show ?thesis using NewInitOOM NewInitRed.hyps init by simp next fix sha ha la v' h' l' sh' a FDTs assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = addr a" and s': "s' = (h'(a \ blank P C), l', sh')" and shaC: "\sfs. sha C \ \(sfs, Done)\" and init': "P \ \INIT C ([C],False) \ unit,(ha, la, sha)\ \ \Val v',(h', l', sh')\" and addr: "new_Addr h' = \a\" and fields: "P \ C has_fields FDTs" then have i: "i = Processing" using iDP shC s\<^sub>1 by simp then have "(h', l', sh') = (ha, la, sha)" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shC by blast then show ?thesis using NewInit NewInitRed.hyps s\<^sub>1a addr fields init shaC e' s' by auto next fix sha ha la v' h' l' sh' assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = THROW OutOfMemory" and s': "s' = (h', l', sh')" and "\sfs. sha C \ \(sfs, Done)\" and init': "P \ \INIT C ([C],False) \ unit,(ha, la, sha)\ \ \Val v',(h', l', sh')\" and addr: "new_Addr h' = None" then have i: "i = Processing" using iDP shC s\<^sub>1 by simp then have "(h', l', sh') = (ha, la, sha)" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shC by blast then show ?thesis using NewInitOOM NewInitRed.hyps e' addr s' s\<^sub>1a init by auto next fix sha ha la a assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and "\sfs. sha C \ \(sfs, Done)\" and init': "P \ \INIT C ([C],False) \ unit,(ha, la, sha)\ \ \throw a,s'\" then have i: "i = Processing" using iDP shC s\<^sub>1 by simp then show ?thesis using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shC by blast qed next fix e assume e': "e' = throw e" and init: "P \ \INIT C ([C],False) \ unit,(h, l, sh)\ \ \throw e,s'\" obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s') then obtain sfs i where shC: "sh' C = \(sfs, i)\" and iDP: "i = Error" using init_throw_PD[OF init] by auto then show ?thesis by (simp add: NewInitRed.hyps NewInitThrow e' init) qed next case CastRed then show ?case by(fastforce elim!: eval_cases intro: eval_evals.intros intro!: CastFail) next case RedCastNull then show ?case by simp (iprover elim: eval_cases intro: eval_evals.intros) next case RedCast then show ?case by (auto elim: eval_cases intro: eval_evals.intros) next case RedCastFail then show ?case by (auto elim!: eval_cases intro: eval_evals.intros) next case BinOpRed1 then show ?case by(fastforce elim!: eval_cases intro: eval_evals.intros simp: val_no_step) next case BinOpRed2 thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId) next case RedBinOp thus ?case by simp (iprover elim: eval_cases intro: eval_evals.intros) next case RedVar thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case LAssRed thus ?case by(fastforce elim: eval_cases intro: eval_evals.intros) next case RedLAss thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case FAccRed thus ?case by(fastforce elim: eval_cases intro: eval_evals.intros) next case RedFAcc then show ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedFAccNull then show ?case by (fastforce elim!: eval_cases intro: eval_evals.intros) next case RedFAccNone thus ?case by(fastforce elim: eval_cases intro: eval_evals.intros) next case RedFAccStatic thus ?case by(fastforce elim: eval_cases intro: eval_evals.intros) next case (RedSFAcc C F t D sh sfs i v h l) then have e':"e' = Val v" and s':"s' = (h, l, sh)" using eval_cases(3) by fastforce+ have "i = Done \ i = Processing" using RedSFAcc by (clarsimp simp: bconf_def initPD_def) then show ?case proof(cases i) case Done then show ?thesis using RedSFAcc e' s' SFAcc by simp next case Processing then have shC': "\sfs. sh D = Some(sfs,Done)" and shP: "sh D = Some(sfs,Processing)" using RedSFAcc by simp+ then have init: "P \ \INIT D ([D],False) \ unit,(h,l,sh)\ \ \unit,(h,l,sh)\" by(simp add: InitFinal InitProcessing Val) have "P \ \C\\<^sub>sF{D},(h, l, sh)\ \ \Val v,(h,l,sh)\" by(rule SFAccInit[OF RedSFAcc.hyps(1) shC' init shP RedSFAcc.hyps(3)]) then show ?thesis using e' s' by simp qed(auto) next case (SFAccInitRed C F t D sh h l) then have seq: "P \ \(INIT D ([D],False) \ unit);; C\\<^sub>sF{D},(h, l, sh)\ \ \e',s'\" using eval_init_seq by simp then show ?case proof(rule eval_cases(14)) \ \ Seq \ fix v s\<^sub>1 assume init: "P \ \INIT D ([D],False) \ unit,(h, l, sh)\ \ \Val v,s\<^sub>1\" and acc: "P \ \C\\<^sub>sF{D},s\<^sub>1\ \ \e',s'\" obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1) then obtain sfs i where shD: "sh\<^sub>1 D = \(sfs, i)\" and iDP: "i = Done \ i = Processing" using init_Val_PD[OF init] by auto show ?thesis proof(rule eval_cases(8)[OF acc]) \ \ SFAcc \ fix t sha sfs v ha la assume "s\<^sub>1 = (ha, la, sha)" and "e' = Val v" and "s' = (ha, la, sha)" and "P \ C has F,Static:t in D" and "sha D = \(sfs, Done)\" and "sfs F = \v\" then show ?thesis using SFAccInit SFAccInitRed.hyps(2) init by auto next fix t sha ha la v' h' l' sh' sfs i' v assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = Val v" and s': "s' = (h', l', sh')" and field: "P \ C has F,Static:t in D" and "\sfs. sha D \ \(sfs, Done)\" and init': "P \ \INIT D ([D],False) \ unit,(ha, la, sha)\ \ \Val v',(h', l', sh')\" and shD': "sh' D = \(sfs, i')\" and sfsF: "sfs F = \v\" then have i: "i = Processing" using iDP shD s\<^sub>1 by simp then have "(h', l', sh') = (ha, la, sha)" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast then show ?thesis using SFAccInit SFAccInitRed.hyps(2) e' s' field init s\<^sub>1a sfsF shD' by auto next fix t sha ha la a assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and "e' = throw a" and "P \ C has F,Static:t in D" and "\sfs. sha D \ \(sfs, Done)\" and init': "P \ \INIT D ([D],False) \ unit,(ha, la, sha)\ \ \throw a,s'\" then have i: "i = Processing" using iDP shD s\<^sub>1 by simp then show ?thesis using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast next assume "\b t. \ P \ C has F,b:t in D" then show ?thesis using SFAccInitRed.hyps(1) by blast next fix t assume field: "P \ C has F,NonStatic:t in D" then show ?thesis using has_field_fun[OF SFAccInitRed.hyps(1) field] by simp qed next fix e assume e': "e' = throw e" and init: "P \ \INIT D ([D],False) \ unit,(h, l, sh)\ \ \throw e,s'\" obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s') then obtain sfs i where shC: "sh' D = \(sfs, i)\" and iDP: "i = Error" using init_throw_PD[OF init] by auto then show ?thesis using SFAccInitRed.hyps(1) SFAccInitRed.hyps(2) SFAccInitThrow e' init by auto qed next case RedSFAccNone thus ?case by(fastforce elim: eval_cases intro: eval_evals.intros) next case RedSFAccNonStatic thus ?case by(fastforce elim: eval_cases intro: eval_evals.intros) next case (FAssRed1 e s b e1 s1 b1 F D e\<^sub>2) obtain h' l' sh' where "s'=(h',l',sh')" by(cases s') with FAssRed1 show ?case by(fastforce elim!: eval_cases(9)[where e\<^sub>1=e1] intro: eval_evals.intros simp: val_no_step intro!: FAss FAssNull FAssNone FAssStatic FAssThrow2) next case FAssRed2 obtain h' l' sh' where "s'=(h',l',sh')" by(cases s') with FAssRed2 show ?case by(auto elim!: eval_cases intro: eval_evals.intros intro!: FAss FAssNull FAssNone FAssStatic FAssThrow2 Val) next case RedFAss thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros) next case RedFAssNull thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros) next case RedFAssNone then show ?case by(auto elim!: eval_cases intro: eval_evals.intros eval_finalId) next case RedFAssStatic then show ?case by(auto elim!: eval_cases intro: eval_evals.intros eval_finalId) next case (SFAssRed e s b e'' s'' b'' C F D) obtain h l sh where [simp]: "s = (h,l,sh)" by(cases s) obtain h' l' sh' where [simp]: "s'=(h',l',sh')" by(cases s') have "val_of e = None" using val_no_step SFAssRed.hyps(1) by(meson option.exhaust) then have bconf: "P,sh \\<^sub>b (e,b) \" using SFAssRed by simp show ?case using SFAssRed.prems(2) SFAssRed bconf proof cases case 2 with SFAssRed bconf show ?thesis by(auto intro!: SFAssInit) next case 3 with SFAssRed bconf show ?thesis by(auto intro!: SFAssInitThrow) qed(auto intro: eval_evals.intros intro!: SFAss SFAssInit SFAssNone SFAssNonStatic) next case (RedSFAss C F t D sh sfs i sfs' v sh' h l) let ?sfs' = "sfs(F \ v)" have e':"e' = unit" and s':"s' = (h, l, sh(D \ (?sfs', i)))" using RedSFAss eval_cases(3) by fastforce+ have "i = Done \ i = Processing" using RedSFAss by(clarsimp simp: bconf_def initPD_def) then show ?case proof(cases i) case Done then show ?thesis using RedSFAss e' s' SFAss Val by auto next case Processing then have shC': "\sfs. sh D = Some(sfs,Done)" and shP: "sh D = Some(sfs,Processing)" using RedSFAss by simp+ then have init: "P \ \INIT D ([D],False) \ unit,(h,l,sh)\ \ \unit,(h,l,sh)\" by(simp add: InitFinal InitProcessing Val) have "P \ \C\\<^sub>sF{D} := Val v,(h, l, sh)\ \ \unit,(h,l,sh(D \ (?sfs', i)))\" using Processing by(auto intro: SFAssInit[OF Val RedSFAss.hyps(1) shC' init shP]) then show ?thesis using e' s' by simp qed(auto) next case (SFAssInitRed C F t D sh v h l) then have seq: "P \ \(INIT D ([D],False) \ unit);; C\\<^sub>sF{D} := Val v,(h, l, sh)\ \ \e',s'\" using eval_init_seq by simp then show ?case proof(rule eval_cases(14)) \ \ Seq \ fix v' s\<^sub>1 assume init: "P \ \INIT D ([D],False) \ unit,(h, l, sh)\ \ \Val v',s\<^sub>1\" and acc: "P \ \C\\<^sub>sF{D} := Val v,s\<^sub>1\ \ \e',s'\" obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1) then obtain sfs i where shD: "sh\<^sub>1 D = \(sfs, i)\" and iDP: "i = Done \ i = Processing" using init_Val_PD[OF init] by auto show ?thesis proof(rule eval_cases(10)[OF acc]) \ \ SFAss \ fix va h\<^sub>1 l\<^sub>1 sh\<^sub>1 t sfs assume e': "e' = unit" and s': "s' = (h\<^sub>1, l\<^sub>1, sh\<^sub>1(D \ (sfs(F \ va), Done)))" and val: "P \ \Val v,s\<^sub>1\ \ \Val va,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\" and field: "P \ C has F,Static:t in D" and shD': "sh\<^sub>1 D = \(sfs, Done)\" have "v = va" and "s\<^sub>1 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF val] by auto then show ?thesis using SFAssInit field SFAssInitRed.hyps(2) shD' e' s' init val by (metis eval_final eval_finalId) next fix va h\<^sub>1 l\<^sub>1 sh\<^sub>1 t v' h' l' sh' sfs i' assume e': "e' = unit" and s': "s' = (h', l', sh'(D \ (sfs(F \ va), i')))" and val: "P \ \Val v,s\<^sub>1\ \ \Val va,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\" and field: "P \ C has F,Static:t in D" and nDone: "\sfs. sh\<^sub>1 D \ \(sfs, Done)\" and init': "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\ \ \Val v',(h', l', sh')\" and shD': "sh' D = \(sfs, i')\" have v: "v = va" and s\<^sub>1a: "s\<^sub>1 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF val] by auto then have i: "i = Processing" using iDP shD s\<^sub>1 nDone by simp then have "(h\<^sub>1, l\<^sub>1, sh\<^sub>1) = (h', l', sh')" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast then show ?thesis using SFAssInit SFAssInitRed.hyps(2) e' s' field init v s\<^sub>1a shD' val by (metis eval_final eval_finalId) next fix va h\<^sub>1 l\<^sub>1 sh\<^sub>1 t a assume "e' = throw a" and val: "P \ \Val v,s\<^sub>1\ \ \Val va,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\" and "P \ C has F,Static:t in D" and nDone: "\sfs. sh\<^sub>1 D \ \(sfs, Done)\" and init': "P \ \INIT D ([D],False) \ unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\ \ \throw a,s'\" have v: "v = va" and s\<^sub>1a: "s\<^sub>1 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF val] by auto then have i: "i = Processing" using iDP shD s\<^sub>1 nDone by simp then have "(h\<^sub>1, l\<^sub>1, sh\<^sub>1) = s'" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast then show ?thesis using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD i by blast next fix e'' assume val:"P \ \Val v,s\<^sub>1\ \ \throw e'',s'\" then show ?thesis using eval_final_same[OF val] by simp next assume "\b t. \ P \ C has F,b:t in D" then show ?thesis using SFAssInitRed.hyps(1) by blast next fix t assume field: "P \ C has F,NonStatic:t in D" then show ?thesis using has_field_fun[OF SFAssInitRed.hyps(1) field] by simp qed next fix e assume e': "e' = throw e" and init: "P \ \INIT D ([D],False) \ unit,(h, l, sh)\ \ \throw e,s'\" obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s') then obtain sfs i where shC: "sh' D = \(sfs, i)\" and iDP: "i = Error" using init_throw_PD[OF init] by auto then show ?thesis using SFAssInitRed.hyps(1) SFAssInitRed.hyps(2) SFAssInitThrow Val by (metis e' init) qed next case (RedSFAssNone C F D v s b) then show ?case by(cases s) (auto elim!: eval_cases intro: eval_evals.intros eval_finalId) next case (RedSFAssNonStatic C F t D v s b) then show ?case by(cases s) (auto elim!: eval_cases intro: eval_evals.intros eval_finalId) next case CallObj note val_no_step[simp] from CallObj.prems(2) CallObj show ?case proof cases case 2 with CallObj show ?thesis by(fastforce intro!: CallParamsThrow) next case 3 with CallObj show ?thesis by(fastforce intro!: CallNull) next case 4 with CallObj show ?thesis by(fastforce intro!: CallNone) next case 5 with CallObj show ?thesis by(fastforce intro!: CallStatic) qed(fastforce intro!: CallObjThrow Call)+ next case (CallParams es s b es'' s'' b'' v M s') then obtain h' l' sh' where "s' = (h',l',sh')" by(cases s') with CallParams show ?case by(auto elim!: eval_cases intro!: CallNone eval_finalId CallStatic Val) (auto intro!: CallParamsThrow CallNull Call Val) next case (RedCall h a C fs M Ts T pns body D vs l sh b) have "P \ \addr a,(h,l,sh)\ \ \addr a,(h,l,sh)\" by (rule eval_evals.intros) moreover have finals: "finals(map Val vs)" by simp with finals have "P \ \map Val vs,(h,l,sh)\ [\] \map Val vs,(h,l,sh)\" by (iprover intro: eval_finalsId) moreover have "h a = Some (C, fs)" using RedCall by simp moreover have "method": "P \ C sees M,NonStatic: Ts\T = (pns, body) in D" by fact moreover have same_len\<^sub>1: "length Ts = length pns" and this_distinct: "this \ set pns" and fv: "fv (body) \ {this} \ set pns" using "method" wf by (fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ have same_len: "length vs = length pns" by fact moreover obtain l\<^sub>2' where l\<^sub>2': "l\<^sub>2' = [this\Addr a,pns[\]vs]" by simp moreover obtain h\<^sub>3 l\<^sub>3 sh\<^sub>3 where s': "s' = (h\<^sub>3,l\<^sub>3,sh\<^sub>3)" by (cases s') have eval_blocks: "P \ \(blocks (this # pns, Class D # Ts, Addr a # vs, body)),(h,l,sh)\ \ \e',s'\" by fact hence id: "l\<^sub>3 = l" using fv s' same_len\<^sub>1 same_len by(fastforce elim: eval_closed_lcl_unchanged) from eval_blocks obtain l\<^sub>3' where "P \ \body,(h,l\<^sub>2',sh)\ \ \e',(h\<^sub>3,l\<^sub>3',sh\<^sub>3)\" proof - from same_len\<^sub>1 have "length(this#pns) = length(Class D#Ts)" by simp moreover from same_len\<^sub>1 same_len have same_len\<^sub>2: "length (this#pns) = length (Addr a#vs)" by simp moreover from eval_blocks have "P \ \blocks (this#pns,Class D#Ts,Addr a#vs,body),(h,l,sh)\ \\e',(h\<^sub>3,l\<^sub>3,sh\<^sub>3)\" using s' same_len\<^sub>1 same_len\<^sub>2 by simp ultimately obtain l'' where "P \ \body,(h,l(this # pns[\]Addr a # vs),sh)\\\e',(h\<^sub>3, l'',sh\<^sub>3)\" by (blast dest:blocksEval) from eval_restrict_lcl[OF wf this fv] this_distinct same_len\<^sub>1 same_len have "P \ \body,(h,[this # pns[\]Addr a # vs],sh)\ \ \e',(h\<^sub>3, l''|`(set(this#pns)),sh\<^sub>3)\" using wf method by(simp add:subset_insert_iff insert_Diff_if) thus ?thesis by(fastforce intro!:that simp add: l\<^sub>2') qed ultimately have "P \ \(addr a)\M(map Val vs),(h,l,sh)\ \ \e',(h\<^sub>3,l,sh\<^sub>3)\" by (rule Call) with s' id show ?case by simp next case RedCallNull thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros eval_finalsId) next case (RedCallNone h a C fs M vs l sh b) then have tes: "THROW NoSuchMethodError = e' \ (h,l,sh) = s'" using eval_final_same by simp have "P \ \addr a,(h,l,sh)\ \ \addr a,(h,l,sh)\" and "P \ \map Val vs,(h,l,sh)\ [\] \map Val vs,(h,l,sh)\" using eval_finalId eval_finalsId by auto then show ?case using RedCallNone CallNone tes by auto next case (RedCallStatic h a C fs M Ts T m D vs l sh b) then have tes: "THROW IncompatibleClassChangeError = e' \ (h,l,sh) = s'" using eval_final_same by simp have "P \ \addr a,(h,l,sh)\ \ \addr a,(h,l,sh)\" and "P \ \map Val vs,(h,l,sh)\ [\] \map Val vs,(h,l,sh)\" using eval_finalId eval_finalsId by auto then show ?case using RedCallStatic CallStatic tes by fastforce next case (SCallParams es s b es'' s'' b' C M s') obtain h' l' sh' where s'[simp]: "s' = (h',l',sh')" by(cases s') obtain h l sh where s[simp]: "s = (h,l,sh)" by(cases s) have es: "map_vals_of es = None" using vals_no_step SCallParams.hyps(1) by (meson not_Some_eq) have bconf: "P,sh \\<^sub>b (es,b) \" using s SCallParams.prems(1) by (simp add: bconf_SCall[OF es]) from SCallParams.prems(2) SCallParams bconf show ?case proof cases case 2 with SCallParams bconf show ?thesis by(auto intro!: SCallNone) next case 3 with SCallParams bconf show ?thesis by(auto intro!: SCallNonStatic) next case 4 with SCallParams bconf show ?thesis by(auto intro!: SCallInitThrow) next case 5 with SCallParams bconf show ?thesis by(auto intro!: SCallInit) qed(auto intro!: SCallParamsThrow SCall) next case (RedSCall C M Ts T pns body D vs s) then obtain h l sh where s:"s = (h,l,sh)" by(cases s) then obtain sfs i where shC: "sh D = \(sfs, i)\" and "i = Done \ i = Processing" using RedSCall by(auto simp: bconf_def initPD_def dest: sees_method_fun) have finals: "finals(map Val vs)" by simp with finals have mVs: "P \ \map Val vs,(h,l,sh)\ [\] \map Val vs,(h,l,sh)\" by (iprover intro: eval_finalsId) obtain sfs i where shC: "sh D = \(sfs, i)\" using RedSCall s by(auto simp: bconf_def initPD_def dest: sees_method_fun) then have iDP: "i = Done \ i = Processing" using RedSCall s by (auto simp: bconf_def initPD_def dest: sees_method_fun[OF RedSCall.hyps(1)]) have "method": "P \ C sees M,Static: Ts\T = (pns, body) in D" by fact have same_len\<^sub>1: "length Ts = length pns" and fv: "fv (body) \ set pns" using "method" wf by (fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+ have same_len: "length vs = length pns" by fact obtain l\<^sub>2' where l\<^sub>2': "l\<^sub>2' = [pns[\]vs]" by simp obtain h\<^sub>3 l\<^sub>3 sh\<^sub>3 where s': "s' = (h\<^sub>3,l\<^sub>3,sh\<^sub>3)" by (cases s') have eval_blocks: "P \ \(blocks (pns, Ts, vs, body)),(h,l,sh)\ \ \e',s'\" using RedSCall.prems(2) s by simp hence id: "l\<^sub>3 = l" using fv s' same_len\<^sub>1 same_len by(fastforce elim: eval_closed_lcl_unchanged) from eval_blocks obtain l\<^sub>3' where body: "P \ \body,(h,l\<^sub>2',sh)\ \ \e',(h\<^sub>3,l\<^sub>3',sh\<^sub>3)\" proof - from eval_blocks have "P \ \blocks (pns,Ts,vs,body),(h,l,sh)\ \\e',(h\<^sub>3,l\<^sub>3,sh\<^sub>3)\" using s' same_len same_len\<^sub>1 by simp then obtain l'' where "P \ \body,(h,l(pns[\]vs),sh)\ \ \e',(h\<^sub>3, l'',sh\<^sub>3)\" by (blast dest:blocksEval[OF same_len\<^sub>1[THEN sym] same_len[THEN sym]]) from eval_restrict_lcl[OF wf this fv] same_len\<^sub>1 same_len have "P \ \body,(h,[pns[\]vs],sh)\ \ \e',(h\<^sub>3, l''|`(set(pns)),sh\<^sub>3)\" using wf method by(simp add:subset_insert_iff insert_Diff_if) thus ?thesis by(fastforce intro!:that simp add: l\<^sub>2') qed show ?case using iDP proof(cases i) case Done then have shC': "sh D = \(sfs, Done)\ \ M = clinit \ sh D = \(sfs, Processing)\" using shC by simp have "P \ \C\\<^sub>sM(map Val vs),(h,l,sh)\ \ \e',(h\<^sub>3,l,sh\<^sub>3)\" by (rule SCall[OF mVs method shC' same_len l\<^sub>2' body]) with s s' id show ?thesis by simp next case Processing then have shC': "\sfs. sh D = Some(sfs,Done)" and shP: "sh D = Some(sfs,Processing)" using shC by simp+ then have init: "P \ \INIT D ([D],False) \ unit,(h,l,sh)\ \ \unit,(h,l,sh)\" by(simp add: InitFinal InitProcessing Val) have "P \ \C\\<^sub>sM(map Val vs),(h,l,sh)\ \ \e',(h\<^sub>3,l,sh\<^sub>3)\" proof(cases "M = clinit") case False show ?thesis by(rule SCallInit[OF mVs method shC' False init same_len l\<^sub>2' body]) next case True then have shC': "sh D = \(sfs, Done)\ \ M = clinit \ sh D = \(sfs, Processing)\" using shC Processing by simp have "P \ \C\\<^sub>sM(map Val vs),(h,l,sh)\ \ \e',(h\<^sub>3,l,sh\<^sub>3)\" by (rule SCall[OF mVs method shC' same_len l\<^sub>2' body]) with s s' id show ?thesis by simp qed with s s' id show ?thesis by simp qed(auto) next case (SCallInitRed C M Ts T pns body D sh vs h l) then have seq: "P \ \(INIT D ([D],False) \ unit);; C\\<^sub>sM(map Val vs),(h, l, sh)\ \ \e',s'\" using eval_init_seq by simp then show ?case proof(rule eval_cases(14)) \ \ Seq \ fix v' s\<^sub>1 assume init: "P \ \INIT D ([D],False) \ unit,(h, l, sh)\ \ \Val v',s\<^sub>1\" and call: "P \ \C\\<^sub>sM(map Val vs),s\<^sub>1\ \ \e',s'\" obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1) then obtain sfs i where shD: "sh\<^sub>1 D = \(sfs, i)\" and iDP: "i = Done \ i = Processing" using init_Val_PD[OF init] by auto show ?thesis proof(rule eval_cases(12)[OF call]) \ \ SCall \ fix vsa ex es' assume "P \ \map Val vs,s\<^sub>1\ [\] \map Val vsa @ throw ex # es',s'\" then show ?thesis using evals_finals_same by (meson finals_def map_Val_nthrow_neq) next assume "\b Ts T a ba x. \ P \ C sees M, b : Ts\T = (a, ba) in x" then show ?thesis using SCallInitRed.hyps(1) by auto next fix Ts T m D assume "P \ C sees M, NonStatic : Ts\T = m in D" then show ?thesis using sees_method_fun[OF SCallInitRed.hyps(1)] by blast next fix vsa h1 l1 sh1 Ts T pns body D' a assume "e' = throw a" and vals: "P \ \map Val vs,s\<^sub>1\ [\] \map Val vsa,(h1, l1, sh1)\" and method: "P \ C sees M, Static : Ts\T = (pns, body) in D'" and nDone: "\sfs. sh1 D' \ \(sfs, Done)\" and init': "P \ \INIT D' ([D'],False) \ unit,(h1, l1, sh1)\ \ \throw a,s'\" have vs: "vs = vsa" and s\<^sub>1a: "s\<^sub>1 = (h1, l1, sh1)" using evals_finals_same[OF _ vals] map_Val_eq by auto have D: "D = D'" using sees_method_fun[OF SCallInitRed.hyps(1) method] by simp then have i: "i = Processing" using iDP shD s\<^sub>1 s\<^sub>1a nDone by simp then show ?thesis using D init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast next fix vsa h1 l1 sh1 Ts T pns' body' D' v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 h\<^sub>3 l\<^sub>3 sh\<^sub>3 assume s': "s' = (h\<^sub>3, l\<^sub>2, sh\<^sub>3)" and vals: "P \ \map Val vs,s\<^sub>1\ [\] \map Val vsa,(h1, l1, sh1)\" and method: "P \ C sees M, Static : Ts\T = (pns', body') in D'" and nDone: "\sfs. sh1 D' \ \(sfs, Done)\" and init': "P \ \INIT D' ([D'],False) \ unit,(h1, l1, sh1)\ \ \Val v',(h\<^sub>2, l\<^sub>2, sh\<^sub>2)\" and len: "length vsa = length pns'" and bstep: "P \ \body',(h\<^sub>2, [pns' [\] vsa], sh\<^sub>2)\ \ \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3)\" have vs: "vs = vsa" and s\<^sub>1a: "s\<^sub>1 = (h1, l1, sh1)" using evals_finals_same[OF _ vals] map_Val_eq by auto have D: "D = D'" and pns: "pns = pns'" and body: "body = body'" using sees_method_fun[OF SCallInitRed.hyps(1) method] by auto then have i: "i = Processing" using iDP shD s\<^sub>1 s\<^sub>1a nDone by simp then have s2: "(h\<^sub>2, l\<^sub>2, sh\<^sub>2) = s\<^sub>1" using D init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast then show ?thesis using eval_finalId SCallInit[OF eval_finalsId[of "map Val vs" P "(h,l,sh)"] SCallInitRed.hyps(1)] init init' len bstep nDone D pns body s' s\<^sub>1 s\<^sub>1a shD vals vs SCallInitRed.hyps(2-3) s2 by auto next fix vsa h\<^sub>2 l\<^sub>2 sh\<^sub>2 Ts T pns' body' D' sfs h\<^sub>3 l\<^sub>3 sh\<^sub>3 assume s': "s' = (h\<^sub>3, l\<^sub>2, sh\<^sub>3)" and vals: "P \ \map Val vs,s\<^sub>1\ [\] \map Val vsa,(h\<^sub>2, l\<^sub>2, sh\<^sub>2)\" and method: "P \ C sees M, Static : Ts\T = (pns', body') in D'" and "sh\<^sub>2 D' = \(sfs, Done)\ \ M = clinit \ sh\<^sub>2 D' = \(sfs, Processing)\" and len: "length vsa = length pns'" and bstep: "P \ \body',(h\<^sub>2, [pns' [\] vsa], sh\<^sub>2)\ \ \e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3)\" have vs: "vs = vsa" and s\<^sub>1a: "s\<^sub>1 = (h\<^sub>2, l\<^sub>2, sh\<^sub>2)" using evals_finals_same[OF _ vals] map_Val_eq by auto have D: "D = D'" and pns: "pns = pns'" and body: "body = body'" using sees_method_fun[OF SCallInitRed.hyps(1) method] by auto then show ?thesis using SCallInit[OF eval_finalsId[of "map Val vs" P "(h,l,sh)"] SCallInitRed.hyps(1)] bstep SCallInitRed.hyps(2-3) len s' s\<^sub>1a vals vs init by auto qed next fix e assume e': "e' = throw e" and init: "P \ \INIT D ([D],False) \ unit,(h, l, sh)\ \ \throw e,s'\" obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s') then obtain sfs i where shC: "sh' D = \(sfs, i)\" and iDP: "i = Error" using init_throw_PD[OF init] by auto then show ?thesis using SCallInitRed.hyps(2-3) init e' SCallInitThrow[OF eval_finalsId[of "map Val vs" _] SCallInitRed.hyps(1)] by auto qed next case (RedSCallNone C M vs s b) then have tes: "THROW NoSuchMethodError = e' \ s = s'" using eval_final_same by simp have "P \ \map Val vs,s\ [\] \map Val vs,s\" using eval_finalsId by simp then show ?case using RedSCallNone eval_evals.SCallNone tes by auto next case (RedSCallNonStatic C M Ts T m D vs s b) then have tes: "THROW IncompatibleClassChangeError = e' \ s = s'" using eval_final_same by simp have "P \ \map Val vs,s\ [\] \map Val vs,s\" using eval_finalsId by simp then show ?case using RedSCallNonStatic eval_evals.SCallNonStatic tes by auto next case InitBlockRed thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId simp: assigned_def map_upd_triv fun_upd_same) next case (RedInitBlock V T v u s b) then have "P \ \Val u,s\ \ \e',s'\" by simp then obtain s': "s'=s" and e': "e'=Val u" by cases simp obtain h l sh where s: "s=(h,l,sh)" by (cases s) have "P \ \{V:T :=Val v; Val u},(h,l,sh)\ \ \Val u,(h, (l(V\v))(V:=l V), sh)\" by (fastforce intro!: eval_evals.intros) then have "P \ \{V:T := Val v; Val u},s\ \ \e',s'\" using s s' e' by simp then show ?case by simp next case BlockRedNone thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros simp add: fun_upd_same fun_upd_idem) next case BlockRedSome thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros simp add: fun_upd_same fun_upd_idem) next case (RedBlock V T v s b) then have "P \ \Val v,s\ \ \e',s'\" by simp then obtain s': "s'=s" and e': "e'=Val v" by cases simp obtain h l sh where s: "s=(h,l,sh)" by (cases s) have "P \ \Val v,(h,l(V:=None),sh)\ \ \Val v,(h,l(V:=None),sh)\" by (rule eval_evals.intros) hence "P \ \{V:T;Val v},(h,l,sh)\ \ \Val v,(h,(l(V:=None))(V:=l V),sh)\" by (rule eval_evals.Block) then have "P \ \{V:T; Val v},s\ \ \e',s'\" using s s' e' by simp then show ?case by simp next case (SeqRed e s b e1 s1 b1 e\<^sub>2) show ?case proof(cases "val_of e") case None show ?thesis proof(cases "lass_val_of e") case lNone:None then have bconf: "P,shp s \\<^sub>b (e,b) \" using SeqRed.prems(1) None by simp then show ?thesis using SeqRed using seq_ext by fastforce next case (Some p) obtain V' v' where p: "p = (V',v')" and e: "e = V':=Val v'" using lass_val_of_spec[OF Some] by(cases p, auto) obtain h l sh h' l' sh' where s: "s = (h,l,sh)" and s1: "s1 = (h',l',sh')" by(cases s, cases s1) then have red: "P \ \e,(h,l,sh),b\ \ \e1,(h',l',sh'),b1\" using SeqRed.hyps(1) by simp then have s\<^sub>1': "e1 = unit \ h' = h \ l' = l(V' \ v') \ sh' = sh" using lass_val_of_red[OF Some red] p e by simp then have eval: "P \ \e,s\ \ \e1,s1\" using e s s1 LAss Val by auto then show ?thesis by (metis SeqRed.prems(2) eval_final eval_final_same seq_ext) qed next case (Some a) then show ?thesis using SeqRed.hyps(1) val_no_step by blast qed next case RedSeq thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case CondRed thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros simp: val_no_step) next case RedCondT thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedCondF thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedWhile thus ?case by (auto simp add: unfold_while intro:eval_evals.intros elim:eval_cases) next case ThrowRed then show ?case by(fastforce elim: eval_cases simp: eval_evals.intros) next case RedThrowNull thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case TryRed thus ?case by(fastforce elim: eval_cases intro: eval_evals.intros) next case RedTryCatch thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (RedTryFail s a D fs C V e\<^sub>2 b) thus ?case by (cases s)(auto elim!: eval_cases intro: eval_evals.intros) next case ListRed1 thus ?case by (fastforce elim: evals_cases intro: eval_evals.intros simp: val_no_step) next case ListRed2 thus ?case by (fastforce elim!: evals_cases eval_cases intro: eval_evals.intros eval_finalId) next case (RedInit e1 C b s1 b') then show ?case using InitFinal by simp next case (InitNoneRed sh C C' Cs e h l b) show ?case using InitNone InitNoneRed.hyps InitNoneRed.prems(2) by auto next case (RedInitDone sh C sfs C' Cs e h l b) show ?case using InitDone RedInitDone.hyps RedInitDone.prems(2) by auto next case (RedInitProcessing sh C sfs C' Cs e h l b) show ?case using InitProcessing RedInitProcessing.hyps RedInitProcessing.prems(2) by auto next case (RedInitError sh C sfs C' Cs e h l b) show ?case using InitError RedInitError.hyps RedInitError.prems(2) by auto next case (InitObjectRed sh C sfs sh' C' Cs e h l b) show ?case using InitObject InitObjectRed by auto next case (InitNonObjectSuperRed sh C sfs D r sh' C' Cs e h l b) show ?case using InitNonObject InitNonObjectSuperRed by auto next case (RedInitRInit C' C Cs e h l sh b) show ?case using InitRInit RedInitRInit by auto next case (RInitRed e s b e'' s'' b'' C Cs e\<^sub>0) then have IH: "\e' s'. P \ \e'',s''\ \ \e',s'\ \ P \ \e,s\ \ \e',s'\" by simp show ?case using RInitRed rinit_ext[OF IH] by simp next case (RedRInit sh C sfs i sh' C' Cs v e h l b s' e') then have init: "P \ \(INIT C' (Cs,True) \ e), (h, l, sh(C \ (sfs, Done)))\ \ \e',s'\" using RedRInit by simp then show ?case using RInit RedRInit.hyps(1) RedRInit.hyps(3) Val by fastforce next case BinOpThrow2 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case FAssThrow2 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case SFAssThrow then show ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (CallThrowParams es vs e es' v M s b) have val: "P \ \Val v,s\ \ \Val v,s\" by (rule eval_evals.intros) have eval_e: "P \ \throw (e),s\ \ \e',s'\" using CallThrowParams by simp then obtain xa where e': "e' = Throw xa" by (cases) (auto dest!: eval_final) with list_eval_Throw [OF eval_e] have vals: "P \ \es,s\ [\] \map Val vs @ Throw xa # es',s'\" using CallThrowParams.hyps(1) eval_e list_eval_Throw by blast then have "P \ \Val v\M(es),s\ \ \Throw xa,s'\" using eval_evals.CallParamsThrow[OF val vals] by simp thus ?case using e' by simp next case (SCallThrowParams es vs e es' C M s b) have eval_e: "P \ \throw (e),s\ \ \e',s'\" using SCallThrowParams by simp then obtain xa where e': "e' = Throw xa" by (cases) (auto dest!: eval_final) then have "P \ \es,s\ [\] \map Val vs @ Throw xa # es',s'\" using SCallThrowParams.hyps(1) eval_e list_eval_Throw by blast then have "P \ \C\\<^sub>sM(es),s\ \ \Throw xa,s'\" by (rule eval_evals.SCallParamsThrow) thus ?case using e' by simp next case (BlockThrow V T a s b) then have "P \ \Throw a, s\ \ \e',s'\" by simp then obtain s': "s' = s" and e': "e' = Throw a" by cases (auto elim!:eval_cases) obtain h l sh where s: "s=(h,l,sh)" by (cases s) have "P \ \Throw a, (h,l(V:=None),sh)\ \ \Throw a, (h,l(V:=None),sh)\" by (simp add:eval_evals.intros eval_finalId) hence "P\\{V:T;Throw a},(h,l,sh)\ \ \Throw a, (h,(l(V:=None))(V:=l V),sh)\" by (rule eval_evals.Block) then have "P \ \{V:T; Throw a},s\ \ \e',s'\" using s s' e' by simp then show ?case by simp next case (InitBlockThrow V T v a s b) then have "P \ \Throw a,s\ \ \e',s'\" by simp then obtain s': "s' = s" and e': "e' = Throw a" by cases (auto elim!:eval_cases) obtain h l sh where s: "s = (h,l,sh)" by (cases s) have "P \ \{V:T :=Val v; Throw a},(h,l,sh)\ \ \Throw a, (h, (l(V\v))(V:=l V),sh)\" by(fastforce intro:eval_evals.intros) then have "P \ \{V:T := Val v; Throw a},s\ \ \e',s'\" using s s' e' by simp then show ?case by simp next case (RInitInitThrow sh C sfs i sh' a D Cs e h l b) have IH: "\e' s'. P \ \RI (D,Throw a) ; Cs \ e,(h, l, sh(C \ (sfs, Error)))\ \ \e',s'\ \ P \ \RI (C,Throw a) ; D # Cs \ e,(h, l, sh)\ \ \e',s'\" using RInitInitFail[OF eval_finalId] RInitInitThrow by simp then show ?case using RInitInitThrow.hyps(2) RInitInitThrow.prems(2) by auto next case (RInitThrow sh C sfs i sh' a e h l b) then have e': "e' = Throw a" and s': "s' = (h,l,sh')" using eval_final_same final_def by fastforce+ show ?case using RInitFailFinal RInitThrow.hyps(1) RInitThrow.hyps(2) e' eval_finalId s' by auto qed(auto elim: eval_cases simp: eval_evals.intros) (*>*) (*<*) (* ... und wieder anschalten: *) declare split_paired_All [simp] split_paired_Ex [simp] (*>*) text \ Its extension to @{text"\*"}: \ lemma extend_eval: assumes wf: "wwf_J_prog P" and reds: "P \ \e,s,b\ \* \e'',s'',b''\" and eval_rest: "P \ \e'',s''\ \ \e',s'\" and iconf: "iconf (shp s) e" and bconf: "P,shp s \\<^sub>b (e::expr,b) \" shows "P \ \e,s\ \ \e',s'\" (*<*) using reds eval_rest iconf bconf proof (induct rule: converse_rtrancl_induct3) case refl then show ?case by simp next case (step e1 s1 b1 e2 s2 b2) then have ic: "iconf (shp s2) e2" using Red_preserves_iconf local.wf by blast then have ec: "P,shp s2 \\<^sub>b (e2,b2) \" using Red_preserves_bconf local.wf step.hyps(1) step.prems(2) step.prems(3) by blast show ?case using step ic ec extend_1_eval[OF wf step.hyps(1)] by simp qed (*>*) lemma extend_evals: assumes wf: "wwf_J_prog P" and reds: "P \ \es,s,b\ [\]* \es'',s'',b''\" and eval_rest: "P \ \es'',s''\ [\] \es',s'\" and iconf: "iconfs (shp s) es" and bconf: "P,shp s \\<^sub>b (es::expr list,b) \" shows "P \ \es,s\ [\] \es',s'\" (*<*) using reds eval_rest iconf bconf proof (induct rule: converse_rtrancl_induct3) case refl then show ?case by simp next case (step es1 s1 b1 es2 s2 b2) then have ic: "iconfs (shp s2) es2" using Reds_preserves_iconf local.wf by blast then have ec: "P,shp s2 \\<^sub>b (es2,b2) \" using Reds_preserves_bconf local.wf step.hyps(1) step.prems(2) step.prems(3) by blast show ?case using step ic ec extend_1_evals[OF wf step.hyps(1)] by simp qed (*>*) text \ Finally, small step semantics can be simulated by big step semantics: \ theorem assumes wf: "wwf_J_prog P" shows small_by_big: "\P \ \e,s,b\ \* \e',s',b'\; iconf (shp s) e; P,shp s \\<^sub>b (e,b) \; final e'\ \ P \ \e,s\ \ \e',s'\" and "\P \ \es,s,b\ [\]* \es',s',b'\; iconfs (shp s) es; P,shp s \\<^sub>b (es,b) \; finals es'\ \ P \ \es,s\ [\] \es',s'\" (*<*) proof - note wf moreover assume "P \ \e,s,b\ \* \e',s',b'\" moreover assume "final e'" then have "P \ \e',s'\ \ \e',s'\" by (simp add: eval_finalId) moreover assume "iconf (shp s) e" moreover assume "P,shp s \\<^sub>b (e,b) \" ultimately show "P \ \e,s\ \ \e',s'\" by (rule extend_eval) next assume fins: "finals es'" note wf moreover assume "P \ \es,s,b\ [\]* \es',s',b'\" moreover have "P \ \es',s'\ [\] \es',s'\" using fins by (rule eval_finalsId) moreover assume "iconfs (shp s) es" moreover assume "P,shp s \\<^sub>b (es,b) \" ultimately show "P \ \es,s\ [\] \es',s'\" by (rule extend_evals) qed (*>*) subsection "Equivalence" text\ And now, the crowning achievement: \ corollary big_iff_small: "\ wwf_J_prog P; iconf (shp s) e; P,shp s \\<^sub>b (e::expr,b) \ \ \ P \ \e,s\ \ \e',s'\ = (P \ \e,s,b\ \* \e',s',False\ \ final e')" (*<*)by(blast dest: big_by_small eval_final small_by_big)(*>*) corollary big_iff_small_WT: "wwf_J_prog P \ P,E \ e::T \ P,shp s \\<^sub>b (e,b) \ \ P \ \e,s\ \ \e',s'\ = (P \ \e,s,b\ \* \e',s',False\ \ final e')" (*<*)by(blast dest: big_iff_small WT_nsub_RI nsub_RI_iconf)(*>*) subsection \ Lifting type safety to @{text"\"} \ text\ \dots and now to the big step semantics, just for fun. \ lemma eval_preserves_sconf: fixes s::state and s'::state assumes "wf_J_prog P" and "P \ \e,s\ \ \e',s'\" and "iconf (shp s) e" and "P,E \ e::T" and "P,E \ s\" shows "P,E \ s'\" (*<*) proof - have "P,shp s \\<^sub>b (e,False) \" by(simp add: bconf_def) with assms show ?thesis by(blast intro:Red_preserves_sconf Red_preserves_iconf Red_preserves_bconf big_by_small WT_implies_WTrt wf_prog_wwf_prog) qed (*>*) lemma eval_preserves_type: fixes s::state assumes wf: "wf_J_prog P" and "P \ \e,s\ \ \e',s'\" and "P,E \ s\" and "iconf (shp s) e" and "P,E \ e::T" shows "\T'. P \ T' \ T \ P,E,hp s',shp s' \ e':T'" (*<*) proof - have "P,shp s \\<^sub>b (e,False) \" by(simp add: bconf_def) with assms show ?thesis by(blast dest:big_by_small[OF wf_prog_wwf_prog[OF wf]] WT_implies_WTrt Red_preserves_type[OF wf]) qed (*>*) end diff --git a/thys/JinjaDCI/J/TypeSafe.thy b/thys/JinjaDCI/J/TypeSafe.thy --- a/thys/JinjaDCI/J/TypeSafe.thy +++ b/thys/JinjaDCI/J/TypeSafe.thy @@ -1,1332 +1,1382 @@ (* Title: JinjaDCI/J/TypeSafe.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory J/TypeSafe.thy by Tobias Nipkow *) section \ Type Safety Proof \ theory TypeSafe imports Progress BigStep SmallStep JWellForm begin (* here because it requires well-typing def *) lemma red_shext_incr: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ (\E T. P,E,h,sh \ e : T \ sh \\<^sub>s sh')" and reds_shext_incr: "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ (\E Ts. P,E,h,sh \ es [:] Ts \ sh \\<^sub>s sh')" (*<*) proof(induct rule:red_reds_inducts) qed(auto simp: shext_def) (*>*) lemma wf_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 \ 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 pns body where sm: "(clinit, Static, [], Void, pns, body) \ set ms" by(unfold wf_clinit_def) auto then have "P \ C sees clinit,Static:[] \ Void = (pns,body) in C" using mdecl_visible[OF wf sP sm] by simp then show ?thesis using WTrtSCall proc by simp qed subsection\Basic preservation lemmas\ text\ First some easy preservation lemmas. \ theorem red_preserves_hconf: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ (\T E. \ P,E,h,sh \ e : T; P \ h \ \ \ P \ h' \)" and reds_preserves_hconf: "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ (\Ts E. \ P,E,h,sh \ es [:] Ts; P \ h \ \ \ P \ h' \)" (*<*) proof (induct rule:red_reds_inducts) case (RedNew h a C FDTs h' l sh es) have new: "new_Addr h = Some a" and fields: "P \ C has_fields FDTs" and h': "h' = h(a\blank P C)" and hconf: "P \ h \" by fact+ from new have None: "h a = None" by(rule new_Addr_SomeD) moreover have "P,h \ blank P C \" using fields by(rule oconf_blank) ultimately show "P \ h' \" using h' by(fast intro: hconf_new[OF hconf]) next case (RedFAss C F t D h a fs v l sh b') let ?fs' = "fs((F,D)\v)" have hconf: "P \ h \" and ha: "h a = Some(C,fs)" and wt: "P,E,h,sh \ addr a\F{D}:=Val v : T" by fact+ from wt ha obtain TF Tv where typeofv: "typeof\<^bsub>h\<^esub> v = Some Tv" and has: "P \ C has F,NonStatic:TF in D" and sub: "P \ Tv \ TF" by auto have "P,h \ (C, ?fs') \" proof (rule oconf_fupd[OF has]) show "P,h \ (C, fs) \" using hconf ha by(simp add:hconf_def) show "P,h \ v :\ TF" using sub typeofv by(simp add:conf_def) qed with hconf ha show "P \ h(a\(C, ?fs')) \" by (rule hconf_upd_obj) qed(auto elim: WTrt.cases) (*>*) theorem red_preserves_lconf: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ (\T E. \ P,E,h,sh \ e:T; P,h \ l (:\) E \ \ P,h' \ l' (:\) E)" and reds_preserves_lconf: "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ (\Ts E. \ P,E,h,sh \ es[:]Ts; P,h \ l (:\) E \ \ P,h' \ l' (:\) E)" (*<*) proof(induct rule:red_reds_inducts) case RedNew thus ?case by(fast intro:lconf_hext red_hext_incr[OF red_reds.RedNew]) next case RedLAss thus ?case by(fastforce elim: lconf_upd simp:conf_def) next case RedFAss thus ?case by(fast intro:lconf_hext red_hext_incr[OF red_reds.RedFAss]) next case (InitBlockRed e h l V v sh b e' h' l' sh' b' v' T T') have red: "P \ \e, (h, l(V\v),sh),b\ \ \e',(h', l',sh'),b'\" and IH: "\T E . \ P,E,h,sh \ e:T; P,h \ l(V\v) (:\) E \ \ P,h' \ l' (:\) E" and l'V: "l' V = Some v'" and lconf: "P,h \ l (:\) E" and wt: "P,E,h,sh \ {V:T := Val v; e} : T'" by fact+ from lconf_hext[OF lconf red_hext_incr[OF red]] have "P,h' \ l (:\) E" . moreover from IH lconf wt have "P,h' \ l' (:\) E(V\T)" by(auto simp del: fun_upd_apply simp: fun_upd_same lconf_upd2 conf_def) ultimately show "P,h' \ l'(V := l V) (:\) E" by (fastforce simp:lconf_def split:if_split_asm) next case (BlockRedNone e h l V sh b e' h' l' sh' b' T T') have red: "P \ \e,(h, l(V := None),sh),b\ \ \e',(h', l',sh'),b'\" and IH: "\E T. \ P,E,h,sh \ e : T; P,h \ l(V:=None) (:\) E \ \ P,h' \ l' (:\) E" and lconf: "P,h \ l (:\) E" and wt: "P,E,h,sh \ {V:T; e} : T'" by fact+ from lconf_hext[OF lconf red_hext_incr[OF red]] have "P,h' \ l (:\) E" . moreover have "P,h' \ l' (:\) E(V\T)" by(rule IH, insert lconf wt, auto simp:lconf_def) ultimately show "P,h' \ l'(V := l V) (:\) E" by (fastforce simp:lconf_def split:if_split_asm) next case (BlockRedSome e h l V sh b e' h' l' sh' b' v T T') have red: "P \ \e,(h, l(V := None),sh),b\ \ \e',(h', l',sh'),b'\" and IH: "\E T. \P,E,h,sh \ e : T; P,h \ l(V:=None) (:\) E\ \ P,h' \ l' (:\) E" and lconf: "P,h \ l (:\) E" and wt: "P,E,h,sh \ {V:T; e} : T'" by fact+ from lconf_hext[OF lconf red_hext_incr[OF red]] have "P,h' \ l (:\) E" . moreover have "P,h' \ l' (:\) E(V\T)" by(rule IH, insert lconf wt, auto simp:lconf_def) ultimately show "P,h' \ l'(V := l V) (:\) E" by (fastforce simp:lconf_def split:if_split_asm) qed(auto elim: WTrt.cases) (*>*) theorem red_preserves_shconf: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ (\T E. \ P,E,h,sh \ e : T; P,h \\<^sub>s sh \ \ \ P,h' \\<^sub>s sh' \)" and reds_preserves_shconf: "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ (\Ts E. \ P,E,h,sh \ es [:] Ts; P,h \\<^sub>s sh \ \ \ P,h' \\<^sub>s sh' \)" (*<*) proof (induct rule:red_reds_inducts) case (RedNew h a C FDTs h' l sh es) have new: "new_Addr h = Some a" and h': "h' = h(a\blank P C)" and shconf: "P,h \\<^sub>s sh \" by fact+ from new have None: "h a = None" by(rule new_Addr_SomeD) then show "P,h' \\<^sub>s sh \" using h' by(fast intro: shconf_hnew[OF shconf]) next case (RedFAss C F t D h a fs v l sh b) let ?fs' = "fs((F,D)\v)" have shconf: "P,h \\<^sub>s sh \" and ha: "h a = Some(C,fs)" by fact+ then show "P,h(a\(C, ?fs')) \\<^sub>s sh \" by (rule shconf_hupd_obj) next case (RedSFAss C F t D sh sfs i sfs' v sh' h l) let ?sfs' = "sfs(F\v)" have shconf: "P,h \\<^sub>s sh \" and shD: "sh D = \(sfs, i)\" and wt: "P,E,h,sh \ C\\<^sub>sF{D} := Val v : T" by fact+ from wt obtain TF Tv where typeofv: "typeof\<^bsub>h\<^esub> v = Some Tv" and has: "P \ C has F,Static:TF in D" and sub: "P \ Tv \ TF" by (auto elim: WTrt.cases) have has': "P \ D has F,Static:TF in D" using has by(rule has_field_idemp) have "P,h,D \\<^sub>s ?sfs' \" proof (rule soconf_fupd[OF has']) show "P,h,D \\<^sub>s sfs \" using shconf shD by(simp add:shconf_def) show "P,h \ v :\ TF" using sub typeofv by(simp add:conf_def) qed with shconf have "P,h \\<^sub>s sh(D\(?sfs',i)) \" by (rule shconf_upd_obj) then show ?case using RedSFAss.hyps(3) RedSFAss.hyps(4) by blast next case (InitNoneRed sh C C' Cs e h l) let ?sfs' = "sblank P C" have "P,h \\<^sub>s sh(C \ (?sfs', Prepared)) \" proof(rule shconf_upd_obj) show "P,h \\<^sub>s sh \" using InitNoneRed by simp show "P,h,C \\<^sub>s sblank P C \" by (metis has_field_def soconf_def soconf_sblank) qed then show ?case by blast next case (InitObjectRed sh C sfs sh' C' Cs e h l) have sh': "sh' = sh(C \ (sfs, Processing))" by fact have "P,h \\<^sub>s sh(C \ (sfs, Processing)) \" proof(rule shconf_upd_obj) show "P,h \\<^sub>s sh \" using InitObjectRed by simp moreover have "sh C = \(sfs, Prepared)\" using InitObjectRed by simp ultimately show "P,h,C \\<^sub>s sfs \" using shconfD by blast qed then show ?case using sh' by blast next case (InitNonObjectSuperRed sh C sfs D a b sh' C' Cs e h l) have sh': "sh' = sh(C \ (sfs, Processing))" by fact have "P,h \\<^sub>s sh(C \ (sfs, Processing)) \" proof(rule shconf_upd_obj) show "P,h \\<^sub>s sh \" using InitNonObjectSuperRed by simp moreover have "sh C = \(sfs, Prepared)\" using InitNonObjectSuperRed by simp ultimately show "P,h,C \\<^sub>s sfs \" using shconfD by blast qed then show ?case using sh' by blast next case (RedRInit sh C sfs i sh' C' Cs e v h l) have sh': "sh' = sh(C \ (sfs, Done))" by fact have "P,h \\<^sub>s sh(C \ (sfs, Done)) \" proof(rule shconf_upd_obj) show "P,h \\<^sub>s sh \" using RedRInit by simp moreover have "sh C = \(sfs, i)\" using RedRInit by simp ultimately show "P,h,C \\<^sub>s sfs \" using shconfD by blast qed then show ?case using sh' by blast next case (RInitInitThrow sh C sfs i sh' a D Cs e h l) have sh': "sh' = sh(C \ (sfs, Error))" by fact have "P,h \\<^sub>s sh(C \ (sfs, Error)) \" proof(rule shconf_upd_obj) show "P,h \\<^sub>s sh \" using RInitInitThrow by simp moreover have "sh C = \(sfs, i)\" using RInitInitThrow by simp ultimately show "P,h,C \\<^sub>s sfs \" using shconfD by blast qed then show ?case using sh' by blast next case (RInitThrow sh C sfs i sh' a e' h l) have sh': "sh' = sh(C \ (sfs, Error))" by fact have "P,h \\<^sub>s sh(C \ (sfs, Error)) \" proof(rule shconf_upd_obj) show "P,h \\<^sub>s sh \" using RInitThrow by simp moreover have "sh C = \(sfs, i)\" using RInitThrow by simp ultimately show "P,h,C \\<^sub>s sfs \" using shconfD by blast qed then show ?case using sh' by blast qed(auto elim: WTrt.cases) (*>*) theorem assumes wf: "wwf_J_prog P" shows red_preserves_iconf: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ iconf sh e \ iconf sh' e'" and reds_preserves_iconf: "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ iconfs sh es \ iconfs sh' es'" (*<*) proof (induct rule:red_reds_inducts) case (BinOpRed1 e h l sh b e' h' l' sh' b' bop e\<^sub>2) then show ?case using BinOpRed1 nsub_RI_iconf[of e\<^sub>2 sh'] val_of_spec proof(cases "val_of e") qed(simp,fast) next case (FAssRed1 e h l sh b e' h' l' sh' b' F D e\<^sub>2) then show ?case using FAssRed1 nsub_RI_iconf[of e\<^sub>2 sh'] val_of_spec proof(cases "val_of e") qed(simp,fast) next case (CallObj e h l sh b e' h' l' sh' b' M es) then show ?case using CallObj nsub_RIs_iconfs[of es sh'] val_of_spec proof(cases "val_of e") qed(simp,fast) next case RedCall then show ?case using sees_wwf_nsub_RI[OF wf RedCall.hyps(2)] by (clarsimp simp: assigned_def nsub_RI_iconf) next case (RedSCall C M Ts T pns body D vs a a b) then have "\sub_RI (blocks (pns, Ts, vs, body))" using sees_wwf_nsub_RI[OF wf RedSCall.hyps(1)] by simp then show ?case by(rule nsub_RI_iconf) next case SCallInitRed then show ?case by fastforce next case (BlockRedNone e h l V sh b e' h' l' sh' b' T) then show ?case by auto next case (SeqRed e h l sh b e' h' l' sh' b' e\<^sub>2) then show ?case proof(cases "lass_val_of e") case None then show ?thesis using SeqRed nsub_RI_iconf by(auto dest: val_of_spec) next case (Some a) have "e' = unit \ sh' = sh" by(simp add: lass_val_of_red[OF Some SeqRed(1)]) then show ?thesis using SeqRed Some by(auto dest: val_of_spec) qed next case (ListRed1 e h l sh b e' h' l' sh' b' es) then show ?case using ListRed1 nsub_RIs_iconfs[of es sh'] val_of_spec proof(cases "val_of e") qed(simp,fast) next case RedInit then show ?case by(auto simp: nsub_RI_iconf) next case (RedInitDone sh C sfs C' Cs e h l b) then show ?case proof(cases Cs) qed(auto simp: initPD_def) next case (RedInitProcessing sh C sfs C' Cs e h l b) then show ?case proof(cases Cs) qed(auto simp: initPD_def) next case (RedRInit sh C sfs i sh' C' Cs v e h l b) then show ?case proof(cases Cs) qed(auto simp: initPD_def) next case CallThrowParams then show ?case by(auto simp: iconfs_map_throw) next case SCallThrowParams then show ?case by(auto simp: iconfs_map_throw) qed(auto simp: nsub_RI_iconf lass_val_of_iconf) (*>*) lemma Seq_bconf_preserve_aux: assumes "P \ \e,(h, l, sh),b\ \ \e',(h', l', sh'),b'\" and "P,sh \\<^sub>b (e;; e\<^sub>2,b) \" and "P,sh \\<^sub>b (e::expr,b) \ \ P,sh' \\<^sub>b (e'::expr,b') \" shows "P,sh' \\<^sub>b (e';;e\<^sub>2,b') \" proof(cases "val_of e") case None show ?thesis proof(cases "lass_val_of e") case lNone: None show ?thesis proof(cases "lass_val_of e'") case lNone': None then show ?thesis using None assms lNone by(auto dest: val_of_spec simp: bconf_def option.distinct(1)) next case (Some a) then show ?thesis using None assms lNone by(auto dest: lass_val_of_spec simp: bconf_def) qed next case (Some a) then show ?thesis using None assms by(auto dest: lass_val_of_spec) qed next case (Some a) then show ?thesis using assms by(auto dest: val_of_spec) qed theorem red_preserves_bconf: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ iconf sh e \ P,sh \\<^sub>b (e,b) \ \ P,sh' \\<^sub>b (e',b') \" and reds_preserves_bconf: "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ iconfs sh es \ P,sh \\<^sub>b (es,b) \ \ P,sh' \\<^sub>b (es',b') \" (*<*) proof (induct rule:red_reds_inducts) case (CastRed e a a b b e' a a b b' C) then show ?case proof(cases b') qed(simp, simp add: bconf_def) next case (BinOpRed1 e h l sh b e' h' l' sh' b' bop e\<^sub>2) show ?case proof(cases b') case True with BinOpRed1 val_of_spec show ?thesis proof(cases "val_of e") qed(simp,fast) next case False then show ?thesis by (simp add: bconf_def) qed next case (BinOpRed2 e a a b b e' a a b b' v\<^sub>1 bop) then show ?case proof(cases b') qed(simp, simp add: bconf_def) next case (LAssRed e a a b b e' a a b b' V) then show ?case proof(cases b') qed(simp, simp add: bconf_def) next case (FAccRed e a a b b e' a a b b' F D) then show ?case proof(cases b') qed(simp, simp add: bconf_def) next case (RedSFAccNonStatic C F t D h l sh b) then show ?case using has_field_fun[of P C F NonStatic] by (auto simp: bconf_def) next case (FAssRed1 e h l sh b e' h' l' sh' b' F D e\<^sub>2) show ?case proof(cases b') case True with FAssRed1 val_of_spec show ?thesis proof(cases "val_of e'")qed((simp,fast)+) next case False then show ?thesis by(simp add: bconf_def) qed next case (FAssRed2 e a a b b e' a a b b' v F D) then show ?case proof(cases b') qed(simp, simp add: bconf_def) next case (SFAssRed e h l sh b e' h' l' sh' b' C F D) then show ?case proof(cases b') qed(fastforce simp: bconf_def val_no_step)+ next case (RedSFAssNonStatic C F t D v a a b b) then show ?case using has_field_fun[of P C F NonStatic] by (auto simp: bconf_def) next case (CallObj e h l sh b e' h' l' sh' b' M es) show ?case proof(cases b') case True with CallObj val_of_spec show ?thesis proof(cases "val_of e'")qed((simp,fast)+) next case False then show ?thesis by(simp add: bconf_def) qed next case (CallParams es a a b b es' a a b b' v M) then show ?case proof(cases b') qed(simp, simp add: bconf_def) next case (SCallParams es h l sh b es' h' l' sh' b' C M) show ?case proof(cases b') case b': True show ?thesis proof(cases "map_vals_of es'") case None then show ?thesis using SCallParams b' vals_no_step proof(cases "map_vals_of es")qed(clarsimp,fast) next case f: Some then show ?thesis using SCallParams b' vals_no_step proof(cases "map_vals_of es")qed(clarsimp,fast) qed next case False then show ?thesis by(simp add: bconf_def) qed next case (SCallInitDoneRed C M Ts T pns body D sh sfs vs h l) then show ?case by(auto simp: bconf_def initPD_def) (rule_tac x=D in exI, auto)+ next case (RedSCallNonStatic C M Ts T a b D vs h l sh b) then show ?case using sees_method_fun[of P C M NonStatic] by (auto simp: bconf_def) next case (BlockRedNone e h l V sh b e' h' l' sh' b' T) show ?case proof(cases "assigned V e'") case True then obtain v e2 where "e' = V := Val v;;e2" by(clarsimp simp: assigned_def) then show ?thesis using BlockRedNone by(clarsimp simp: bconf_def) next case False then show ?thesis using BlockRedNone by simp qed next case (BlockRedSome e h l V sh b e' h' l' sh' b' v T) then show ?case proof(cases b') qed(simp, simp add: bconf_def) next case (InitBlockRed e h l V v sh b e' h' l' sh' b' v' T) show ?case proof(cases b') case True then show ?thesis using InitBlockRed by (simp add: assigned_def) next case False then show ?thesis by(simp add: bconf_def) qed next case (RedBlock V T u) then have "\assigned V (Val u)" by(clarsimp simp: assigned_def) then show ?case using RedBlock by(simp add: bconf_def) next case (SeqRed e h l sh b e' h' l' sh' b' e\<^sub>2) have "iconf sh e" proof(cases "lass_val_of e") case (Some a) show ?thesis by(rule lass_val_of_iconf[OF Some]) next case None then show ?thesis using SeqRed by(auto dest: val_of_spec) qed then show ?case using SeqRed Seq_bconf_preserve_aux by simp next case (CondRed e a a b b e' a a b b' e\<^sub>1 e\<^sub>2) then show ?case proof(cases b') qed(simp, simp add: bconf_def) next case (ThrowRed e a a b b e' a a b b') then show ?case proof(cases b') qed(simp, simp add: bconf_def) next case (TryRed e a a b b e' a a b b' C V e\<^sub>2) then show ?case proof(cases b') qed(simp, simp add: bconf_def) next case (ListRed1 e h l sh b e' h' l' sh' b' es) with val_of_spec show ?case proof(cases b') qed((clarsimp,fast),simp add: bconfs_def) next case (RedInit C b' e X Y b b'') then show ?case by(auto simp: bconf_def icheck_ss_exp icheck_init_class icheck_curr_init) next case (RInitRed e a a b b e' a a b b' C Cs e\<^sub>0) then show ?case proof(cases b') qed(simp, simp add: bconf_def) next case (BlockThrow V T a) then have "\assigned V (Throw a)" by(simp add: assigned_def) then show ?case using BlockThrow by simp qed(simp_all, auto simp: bconf_def initPD_def) (*>*) text\ Preservation of definite assignment more complex and requires a few lemmas first. \ lemma [iff]: "\A. \ length Vs = length Ts; length vs = length Ts\ \ \ (blocks (Vs,Ts,vs,e)) A = \ e (A \ \set Vs\)" (*<*) -apply(induct Vs Ts vs e rule:blocks_induct) -apply(simp_all add:hyperset_defs) -done +by (induct Vs Ts vs e rule:blocks_induct) + (simp_all add:hyperset_defs) (*>*) lemma red_lA_incr: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ \dom l\ \ \ e \ \dom l'\ \ \ e'" and reds_lA_incr: "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ \dom l\ \ \s es \ \dom l'\ \ \s es'" (*<*) -apply(induct rule:red_reds_inducts) -apply(simp_all del:fun_upd_apply add:hyperset_defs) -apply auto -apply(blast dest:red_lcl_incr)+ -done +proof(induct rule:red_reds_inducts) + case TryRed then show ?case + by(simp del:fun_upd_apply add:hyperset_defs) + (blast dest:red_lcl_incr)+ +next + case BinOpRed1 then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case BinOpRed2 then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case LAssRed then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case FAssRed1 then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case FAssRed2 then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case CallObj then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case CallParams then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case BlockRedNone then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case BlockRedSome then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case InitBlockRed then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case SeqRed then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case CondRed then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case RedCondT then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case RedCondF then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case ListRed1 then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +next + case ListRed2 then show ?case + by(auto simp del:fun_upd_apply simp:hyperset_defs) +qed (simp_all del:fun_upd_apply add:hyperset_defs) (*>*) text\ Now preservation of definite assignment. \ lemma assumes wf: "wf_J_prog P" shows red_preserves_defass: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ \ e \dom l\ \ \ e' \dom l'\" and "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ \s es \dom l\ \ \s es' \dom l'\" (*<*) proof (induct rule:red_reds_inducts) case BinOpRed1 thus ?case by (auto elim!: D_mono[OF red_lA_incr]) next case FAssRed1 thus ?case by (auto elim!: D_mono[OF red_lA_incr]) next case CallObj thus ?case by (auto elim!: Ds_mono[OF red_lA_incr]) next case RedCall thus ?case by (auto dest!:sees_wf_mdecl[OF wf] simp:wf_mdecl_def hyperset_defs elim!:D_mono') next case RedSCall thus ?case by (auto dest!:sees_wf_mdecl[OF wf] simp:wf_mdecl_def hyperset_defs elim!:D_mono') next case SCallInitRed then show ?case by(auto simp:hyperset_defs Ds_Vals) next case InitBlockRed thus ?case by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply) next case BlockRedNone thus ?case by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply) next case BlockRedSome thus ?case by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply) next case SeqRed thus ?case by (auto elim!: D_mono[OF red_lA_incr]) next case CondRed thus ?case by (auto elim!: D_mono[OF red_lA_incr]) next case TryRed thus ?case by (fastforce dest:red_lcl_incr intro:D_mono' simp:hyperset_defs) next case ListRed1 thus ?case by (auto elim!: Ds_mono[OF red_lA_incr]) next case RedWhile thus ?case by(auto simp:hyperset_defs elim!:D_mono') next case RedInit then show ?case by (auto intro: D_mono' simp: hyperset_defs) next case (RInitRed e h l sh b e' h' l' sh' b' C Cs e\<^sub>0) then show ?case by(auto simp:hyperset_defs dest:red_lcl_incr elim!:D_mono') qed(auto simp:hyperset_defs) (*>*) text\ Combining conformance of heap, static heap, and local variables: \ definition sconf :: "J_prog \ env \ state \ bool" ("_,_ \ _ \" [51,51,51]50) where "P,E \ s \ \ let (h,l,sh) = s in P \ h \ \ P,h \ l (:\) E \ P,h \\<^sub>s sh \" lemma red_preserves_sconf: "\ P \ \e,s,b\ \ \e',s',b'\; P,E,hp s,shp s \ e : T; P,E \ s \ \ \ P,E \ s' \" (*<*) by(fastforce intro:red_preserves_hconf red_preserves_lconf red_preserves_shconf simp add:sconf_def) (*>*) lemma reds_preserves_sconf: "\ P \ \es,s,b\ [\] \es',s',b'\; P,E,hp s,shp s \ es [:] Ts; P,E \ s \ \ \ P,E \ s' \" (*<*) by(fastforce intro:reds_preserves_hconf reds_preserves_lconf reds_preserves_shconf simp add:sconf_def) (*>*) subsection "Subject reduction" lemma wt_blocks: "\E. \ length Vs = length Ts; length vs = length Ts \ \ (P,E,h,sh \ blocks(Vs,Ts,vs,e) : T) = (P,E(Vs[\]Ts),h,sh \ e:T \ (\Ts'. map (typeof\<^bsub>h\<^esub>) vs = map Some Ts' \ P \ Ts' [\] Ts))" (*<*) -apply(induct Vs Ts vs e rule:blocks_induct) -apply (force simp add:rel_list_all2_Cons2) -apply simp_all -done +proof(induct Vs Ts vs e rule:blocks_induct) + case (1 V Vs T Ts v vs e) + then show ?case by(force simp add:rel_list_all2_Cons2) +qed simp_all (*>*) theorem assumes wf: "wf_J_prog P" shows subject_reduction2: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ (\E T. \ P,E \ (h,l,sh) \; iconf sh e; P,E,h,sh \ e:T \ \ \T'. P,E,h',sh' \ e':T' \ P \ T' \ T)" and subjects_reduction2: "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ (\E Ts. \ P,E \ (h,l,sh) \; iconfs sh es; P,E,h,sh \ es [:] Ts \ \ \Ts'. P,E,h',sh' \ es' [:] Ts' \ P \ Ts' [\] Ts)" (*<*) proof (induct rule:red_reds_inducts) case RedNew then show ?case by (auto simp: blank_def) next case RedNewFail thus ?case by (unfold sconf_def hconf_def) (fastforce elim!:typeof_OutOfMemory) next case CastRed thus ?case by(clarsimp simp:is_refT_def) (blast intro: widens_trans dest!:widen_Class[THEN iffD1]) next case RedCastFail thus ?case by (unfold sconf_def hconf_def) (fastforce elim!:typeof_ClassCast) next case (BinOpRed1 e\<^sub>1 h l sh b e\<^sub>1' h' l' sh' b' bop e\<^sub>2) have red: "P \ \e\<^sub>1,(h,l,sh),b\ \ \e\<^sub>1',(h',l',sh'),b'\" and IH: "\E T. \P,E \ (h,l,sh) \; iconf sh e\<^sub>1; P,E,h,sh \ e\<^sub>1:T\ \ \U. P,E,h',sh' \ e\<^sub>1' : U \ P \ U \ T" and conf: "P,E \ (h,l,sh) \" and iconf: "iconf sh (e\<^sub>1 \bop\ e\<^sub>2)" and wt: "P,E,h,sh \ e\<^sub>1 \bop\ e\<^sub>2 : T" by fact+ have val: "val_of e\<^sub>1 = None" using red iconf val_no_step by auto then have iconf1: "iconf sh e\<^sub>1" and nsub_RI2: "\sub_RI e\<^sub>2" using iconf by simp+ have "P,E,h',sh' \ e\<^sub>1' \bop\ e\<^sub>2 : T" proof (cases bop) assume [simp]: "bop = Eq" from wt obtain T\<^sub>1 T\<^sub>2 where [simp]: "T = Boolean" and wt\<^sub>1: "P,E,h,sh \ e\<^sub>1 : T\<^sub>1" and wt\<^sub>2: "P,E,h,sh \ e\<^sub>2 : T\<^sub>2" by auto show ?thesis using WTrt_hext_shext_mono[OF wt\<^sub>2 red_hext_incr[OF red] red_shext_incr[OF red wt\<^sub>1] nsub_RI2] IH[OF conf iconf1 wt\<^sub>1] by auto next assume [simp]: "bop = Add" from wt have [simp]: "T = Integer" and wt\<^sub>1: "P,E,h,sh \ e\<^sub>1 : Integer" and wt\<^sub>2: "P,E,h,sh \ e\<^sub>2 : Integer" by auto show ?thesis using WTrt_hext_shext_mono[OF wt\<^sub>2 red_hext_incr[OF red] red_shext_incr[OF red wt\<^sub>1] nsub_RI2] IH[OF conf iconf1 wt\<^sub>1] by auto qed thus ?case by auto next case (BinOpRed2 e\<^sub>2 h l sh b e\<^sub>2' h' l' sh' b' v\<^sub>1 bop) have red: "P \ \e\<^sub>2,(h,l,sh),b\ \ \e\<^sub>2',(h',l',sh'),b'\" and IH: "\E T. \P,E \ (h,l,sh) \; iconf sh e\<^sub>2; P,E,h,sh \ e\<^sub>2:T\ \ \U. P,E,h',sh' \ e\<^sub>2' : U \ P \ U \ T" and conf: "P,E \ (h,l,sh) \" and iconf: "iconf sh (Val v\<^sub>1 \bop\ e\<^sub>2)" and wt: "P,E,h,sh \ (Val v\<^sub>1) \bop\ e\<^sub>2 : T" by fact+ have iconf2: "iconf sh e\<^sub>2" using iconf by simp have "P,E,h',sh' \ (Val v\<^sub>1) \bop\ e\<^sub>2' : T" proof (cases bop) assume [simp]: "bop = Eq" from wt obtain T\<^sub>1 T\<^sub>2 where [simp]: "T = Boolean" and wt\<^sub>1: "P,E,h,sh \ Val v\<^sub>1 : T\<^sub>1" and wt\<^sub>2: "P,E,h,sh \ e\<^sub>2:T\<^sub>2" by auto show ?thesis using IH[OF conf iconf2 wt\<^sub>2] WTrt_hext_mono[OF wt\<^sub>1 red_hext_incr[OF red]] by auto next assume [simp]: "bop = Add" from wt have [simp]: "T = Integer" and wt\<^sub>1: "P,E,h,sh \ Val v\<^sub>1 : Integer" and wt\<^sub>2: "P,E,h,sh \ e\<^sub>2 : Integer" by auto show ?thesis using IH[OF conf iconf2 wt\<^sub>2] WTrt_hext_mono[OF wt\<^sub>1 red_hext_incr[OF red]] by auto qed thus ?case by auto next case (RedBinOp bop) thus ?case proof (cases bop) case Eq thus ?thesis using RedBinOp by auto next case Add thus ?thesis using RedBinOp by auto qed next case RedVar thus ?case by (fastforce simp:sconf_def lconf_def conf_def) next case (LAssRed e h l sh b e' h' l' sh' b' V) obtain Te where Te: "P,E,h,sh \ e : Te \ P \ Te \ the(E V)" using LAssRed.prems(3) by auto then have wide: "P \ Te \ the(E V)" using LAssRed by simp then have "\T'. P,E,h',sh' \ e' : T' \ P \ T' \ Te" using LAssRed.hyps(2) LAssRed.prems(1,2) Te widen_trans[OF _ wide] by auto then obtain T' where wt: "P,E,h',sh' \ e' : T' \ P \ T' \ Te" by clarsimp have "P,E,h',sh' \ V:=e' : Void" using LAssRed wt widen_trans[OF _ wide] by auto then show ?case using LAssRed by(rule_tac x = Void in exI) auto next case (FAccRed e h l sh b e' h' l' sh' b' F D) have IH: "\E T. \P,E \ (h,l,sh) \; iconf sh e; P,E,h,sh \ e : T\ \ \U. P,E,h',sh' \ e' : U \ P \ U \ T" and conf: "P,E \ (h,l,sh) \" and iconf: "iconf sh (e\F{D})" and wt: "P,E,h,sh \ e\F{D} : T" by fact+ have iconf': "iconf sh e" using iconf by simp+ \ \The goal: ?case = @{prop ?case}\ \ \Now distinguish the two cases how wt can have arisen.\ { fix C assume wte: "P,E,h,sh \ e : Class C" and has: "P \ C has F,NonStatic:T in D" from IH[OF conf iconf' wte] obtain U where wte': "P,E,h',sh' \ e' : U" and UsubC: "P \ U \ Class C" by auto \ \Now distinguish what @{term U} can be.\ { assume "U = NT" hence ?case using wte' by(blast intro:WTrtFAccNT widen_refl) } moreover { fix C' assume U: "U = Class C'" and C'subC: "P \ C' \\<^sup>* C" from has_field_mono[OF has C'subC] wte' U have ?case by(blast intro:WTrtFAcc) } ultimately have ?case using UsubC by(simp add: widen_Class) blast } moreover { assume "P,E,h,sh \ e : NT" hence "P,E,h',sh' \ e' : NT" using IH[OF conf iconf'] by fastforce hence ?case by(fastforce intro:WTrtFAccNT widen_refl) } ultimately show ?case using wt by blast next case RedFAcc thus ?case by(fastforce simp:sconf_def hconf_def oconf_def conf_def has_field_def dest:has_fields_fun) next case RedFAccNull thus ?case by(fastforce intro: widen_refl WTThrow[OF WTVal] elim!: typeof_NullPointer simp: sconf_def hconf_def) next case RedSFAccNone then show ?case by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_NoSuchFieldError simp: sconf_def hconf_def) next case RedFAccStatic then show ?case by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError simp: sconf_def hconf_def) next case (RedSFAcc C F t D sh sfs i v h l es) then have "P \ C has F,Static:T in D" by fast then have dM: "P \ D has F,Static:T in D" by(rule has_field_idemp) then show ?case using RedSFAcc by(fastforce simp:sconf_def shconf_def soconf_def conf_def) next case SFAccInitDoneRed then show ?case by (meson widen_refl) next case (SFAccInitRed C F t D sh h l E T) have "is_class P D" using SFAccInitRed.hyps(1) by(rule has_field_is_class') then have "P,E,h,sh \ INIT D ([D],False) \ C\\<^sub>sF{D} : T \ P \ T \ T" using SFAccInitRed WTrtInit[OF SFAccInitRed.prems(3)] by clarsimp then show ?case by(rule exI) next case RedSFAccNonStatic then show ?case by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError simp: sconf_def hconf_def) next case (FAssRed1 e h l sh b e' h' l' sh' b' F D e\<^sub>2) have red: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\" and IH: "\E T. \P,E \ (h,l,sh) \; iconf sh e; P,E,h,sh \ e : T\ \ \U. P,E,h',sh' \ e' : U \ P \ U \ T" and conf: "P,E \ (h,l,sh) \" and iconf: "iconf sh (e\F{D} := e\<^sub>2)" and wt: "P,E,h,sh \ e\F{D}:=e\<^sub>2 : T" by fact+ have val: "val_of e = None" using red iconf val_no_step by auto then have iconf': "iconf sh e" and nsub_RI2: "\sub_RI e\<^sub>2" using iconf by simp+ from wt have void: "T = Void" by blast \ \We distinguish if @{term e} has type @{term NT} or a Class type\ \ \Remember ?case = @{term ?case}\ { assume wt':"P,E,h,sh \ e : NT" hence "P,E,h',sh' \ e' : NT" using IH[OF conf iconf'] by fastforce moreover obtain T\<^sub>2 where "P,E,h,sh \ e\<^sub>2 : T\<^sub>2" using wt by auto from this red_hext_incr[OF red] red_shext_incr[OF red wt'] nsub_RI2 have "P,E,h',sh' \ e\<^sub>2 : T\<^sub>2" by(rule WTrt_hext_shext_mono) ultimately have ?case using void by(blast intro!:WTrtFAssNT) } moreover { fix C TF T\<^sub>2 assume wt\<^sub>1: "P,E,h,sh \ e : Class C" and wt\<^sub>2: "P,E,h,sh \ e\<^sub>2 : T\<^sub>2" and has: "P \ C has F,NonStatic:TF in D" and sub: "P \ T\<^sub>2 \ TF" obtain U where wt\<^sub>1': "P,E,h',sh' \ e' : U" and UsubC: "P \ U \ Class C" using IH[OF conf iconf' wt\<^sub>1] by blast have wt\<^sub>2': "P,E,h',sh' \ e\<^sub>2 : T\<^sub>2" by(rule WTrt_hext_shext_mono[OF wt\<^sub>2 red_hext_incr[OF red] red_shext_incr[OF red wt\<^sub>1] nsub_RI2]) \ \Is @{term U} the null type or a class type?\ { assume "U = NT" with wt\<^sub>1' wt\<^sub>2' void have ?case by(blast intro!:WTrtFAssNT) } moreover { fix C' assume UClass: "U = Class C'" and "subclass": "P \ C' \\<^sup>* C" have "P,E,h',sh' \ e' : Class C'" using wt\<^sub>1' UClass by auto moreover have "P \ C' has F,NonStatic:TF in D" by(rule has_field_mono[OF has "subclass"]) ultimately have ?case using wt\<^sub>2' sub void by(blast intro:WTrtFAss) } ultimately have ?case using UsubC by(auto simp add:widen_Class) } ultimately show ?case using wt by blast next case (FAssRed2 e\<^sub>2 h l sh b e\<^sub>2' h' l' sh' b' v F D) have red: "P \ \e\<^sub>2,(h,l,sh),b\ \ \e\<^sub>2',(h',l',sh'),b'\" and IH: "\E T. \P,E \ (h,l,sh) \; iconf sh e\<^sub>2; P,E,h,sh \ e\<^sub>2 : T\ \ \U. P,E,h',sh' \ e\<^sub>2' : U \ P \ U \ T" and conf: "P,E \ (h,l,sh) \" and iconf: "iconf sh (Val v\F{D} := e\<^sub>2)" and wt: "P,E,h,sh \ Val v\F{D}:=e\<^sub>2 : T" by fact+ have iconf2: "iconf sh e\<^sub>2" using iconf by simp from wt have [simp]: "T = Void" by auto from wt show ?case proof (rule WTrt_elim_cases) fix C TF T\<^sub>2 assume wt\<^sub>1: "P,E,h,sh \ Val v : Class C" and has: "P \ C has F,NonStatic:TF in D" and wt\<^sub>2: "P,E,h,sh \ e\<^sub>2 : T\<^sub>2" and TsubTF: "P \ T\<^sub>2 \ TF" have wt\<^sub>1': "P,E,h',sh' \ Val v : Class C" using WTrt_hext_mono[OF wt\<^sub>1 red_hext_incr[OF red]] by auto obtain T\<^sub>2' where wt\<^sub>2': "P,E,h',sh' \ e\<^sub>2' : T\<^sub>2'" and T'subT: "P \ T\<^sub>2' \ T\<^sub>2" using IH[OF conf iconf2 wt\<^sub>2] by blast have "P,E,h',sh' \ Val v\F{D}:=e\<^sub>2' : Void" by(rule WTrtFAss[OF wt\<^sub>1' has wt\<^sub>2' widen_trans[OF T'subT TsubTF]]) thus ?case by auto next fix T\<^sub>2 assume null: "P,E,h,sh \ Val v : NT" and wt\<^sub>2: "P,E,h,sh \ e\<^sub>2 : T\<^sub>2" from null have "v = Null" by simp moreover obtain T\<^sub>2' where "P,E,h',sh' \ e\<^sub>2' : T\<^sub>2' \ P \ T\<^sub>2' \ T\<^sub>2" using IH[OF conf iconf2 wt\<^sub>2] by blast ultimately show ?thesis by(fastforce intro:WTrtFAssNT) qed next case RedFAss thus ?case by(auto simp del:fun_upd_apply) next case RedFAssNull thus ?case by(fastforce intro: WTThrow[OF WTVal] elim!:typeof_NullPointer simp:sconf_def hconf_def) next case RedFAssStatic then show ?case by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError simp: sconf_def hconf_def) next case (SFAssRed e h l sh b e' h' l' sh' b' C F D E T) have IH: "\E T. \P,E \ (h,l,sh) \; iconf sh e; P,E,h,sh \ e : T\ \ \U. P,E,h',sh' \ e' : U \ P \ U \ T" and conf: "P,E \ (h,l,sh) \" and iconf: "iconf sh (C\\<^sub>sF{D} := e)" and wt: "P,E,h,sh \ C\\<^sub>sF{D}:=e : T" by fact+ have iconf': "iconf sh e" using iconf by simp from wt have [simp]: "T = Void" by auto from wt show ?case proof (rule WTrt_elim_cases) fix TF T1 assume has: "P \ C has F,Static:TF in D" and wt1: "P,E,h,sh \ e : T1" and TsubTF: "P \ T1 \ TF" obtain T' where wt1': "P,E,h',sh' \ e' : T'" and T'subT: "P \ T' \ T1" using IH[OF conf iconf' wt1] by blast have "P,E,h',sh' \ C\\<^sub>sF{D}:=e' : Void" by(rule WTrtSFAss[OF wt1' has widen_trans[OF T'subT TsubTF]]) thus ?case by auto qed next case SFAssInitDoneRed then show ?case by (meson widen_refl) next case (SFAssInitRed C F t D sh v h l E T) have "is_class P D" using SFAssInitRed.hyps(1) by(rule has_field_is_class') then have "P,E,h,sh \ INIT D ([D],False) \ C\\<^sub>sF{D} := Val v : T \ P \ T \ T" using SFAssInitRed WTrtInit[OF SFAssInitRed.prems(3)] by clarsimp then show ?case by(rule exI) next case RedSFAssNone then show ?case by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_NoSuchFieldError simp: sconf_def hconf_def) next case RedSFAssNonStatic then show ?case by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError simp: sconf_def hconf_def) next case (CallObj e h l sh b e' h' l' sh' b' M es) have red: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\" and IH: "\E T. \P,E \ (h,l,sh) \; iconf sh e; P,E,h,sh \ e : T\ \ \U. P,E,h',sh' \ e' : U \ P \ U \ T" and conf: "P,E \ (h,l,sh) \" and iconf: "iconf sh (e\M(es))" and wt: "P,E,h,sh \ e\M(es) : T" by fact+ have val: "val_of e = None" using red iconf val_no_step by auto then have iconf': "iconf sh e" and nsub_RIs: "\sub_RIs es" using iconf by simp+ \ \We distinguish if @{term e} has type @{term NT} or a Class type\ \ \Remember ?case = @{term ?case}\ { assume wt':"P,E,h,sh \ e:NT" hence "P,E,h',sh' \ e' : NT" using IH[OF conf iconf'] by fastforce moreover fix Ts assume wtes: "P,E,h,sh \ es [:] Ts" have "P,E,h',sh' \ es [:] Ts" by(rule WTrts_hext_shext_mono[OF wtes red_hext_incr[OF red] red_shext_incr[OF red wt'] nsub_RIs]) ultimately have ?case by(blast intro!:WTrtCallNT) } moreover { fix C D Ts Us pns body assume wte: "P,E,h,sh \ e : Class C" and "method": "P \ C sees M,NonStatic:Ts\T = (pns,body) in D" and wtes: "P,E,h,sh \ es [:] Us" and subs: "P \ Us [\] Ts" obtain U where wte': "P,E,h',sh' \ e' : U" and UsubC: "P \ U \ Class C" using IH[OF conf iconf' wte] by blast \ \Is @{term U} the null type or a class type?\ { assume "U = NT" moreover have "P,E,h',sh' \ es [:] Us" by(rule WTrts_hext_shext_mono[OF wtes red_hext_incr[OF red] red_shext_incr[OF red wte] nsub_RIs]) ultimately have ?case using wte' by(blast intro!:WTrtCallNT) } moreover { fix C' assume UClass: "U = Class C'" and "subclass": "P \ C' \\<^sup>* C" have "P,E,h',sh' \ e' : Class C'" using wte' UClass by auto moreover obtain Ts' T' pns' body' D' where method': "P \ C' sees M,NonStatic:Ts'\T' = (pns',body') in D'" and subs': "P \ Ts [\] Ts'" and sub': "P \ T' \ T" using Call_lemma[OF "method" "subclass" wf] by fast moreover have "P,E,h',sh' \ es [:] Us" by(rule WTrts_hext_shext_mono[OF wtes red_hext_incr[OF red] red_shext_incr[OF red wte] nsub_RIs]) ultimately have ?case using subs by(blast intro:WTrtCall rtrancl_trans widens_trans) } ultimately have ?case using UsubC by(auto simp add:widen_Class) } ultimately show ?case using wt by auto next case (CallParams es h l sh b es' h' l' sh' b' v M) have reds: "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\" and IH: "\E Ts. \P,E \ (h,l,sh) \; iconfs sh es; P,E,h,sh \ es [:] Ts\ \ \Us. P,E,h',sh' \ es' [:] Us \ P \ Us [\] Ts" and conf: "P,E \ (h,l,sh) \" and iconf: "iconf sh (Val v\M(es))" and wt: "P,E,h,sh \ Val v\M(es) : T" by fact+ have iconfs: "iconfs sh es" using iconf by simp from wt show ?case proof (rule WTrt_elim_cases) fix C D Ts Us pns body assume wte: "P,E,h,sh \ Val v : Class C" and "P \ C sees M,NonStatic:Ts\T = (pns,body) in D" and wtes: "P,E,h,sh \ es [:] Us" and "P \ Us [\] Ts" moreover have "P,E,h',sh' \ Val v : Class C" using WTrt_hext_mono[OF wte reds_hext_incr[OF reds]] by auto moreover obtain Us' where "P,E,h',sh' \ es' [:] Us' \ P \ Us' [\] Us" using IH[OF conf iconfs wtes] by blast ultimately show ?thesis by(blast intro:WTrtCall widens_trans) next fix Us assume null: "P,E,h,sh \ Val v : NT" and wtes: "P,E,h,sh \ es [:] Us" from null have "v = Null" by simp moreover obtain Us' where "P,E,h',sh' \ es' [:] Us' \ P \ Us' [\] Us" using IH[OF conf iconfs wtes] by blast ultimately show ?thesis by(fastforce intro:WTrtCallNT) qed next case (RedCall h a C fs M Ts T pns body D vs l sh b E T') have hp: "h a = Some(C,fs)" and "method": "P \ C sees M,NonStatic: Ts\T = (pns,body) in D" and wt: "P,E,h,sh \ addr a\M(map Val vs) : T'" by fact+ obtain Ts' where wtes: "P,E,h,sh \ map Val vs [:] Ts'" and subs: "P \ Ts' [\] Ts" and T'isT: "T' = T" using wt "method" hp by (auto dest:sees_method_fun) from wtes subs have length_vs: "length vs = length Ts" by(fastforce simp:list_all2_iff dest!:WTrts_same_length) from sees_wf_mdecl[OF wf "method"] obtain T'' where wtabody: "P,[this#pns [\] Class D#Ts] \ body :: T''" and T''subT: "P \ T'' \ T" and length_pns: "length pns = length Ts" by(fastforce simp:wf_mdecl_def simp del:map_upds_twist) from wtabody have "P,Map.empty(this#pns [\] Class D#Ts),h,sh \ body : T''" by(rule WT_implies_WTrt) hence "P,E(this#pns [\] Class D#Ts),h,sh \ body : T''" by(rule WTrt_env_mono) simp hence "P,E,h,sh \ blocks(this#pns, Class D#Ts, Addr a#vs, body) : T''" using wtes subs hp sees_method_decl_above[OF "method"] length_vs length_pns by(fastforce simp add:wt_blocks rel_list_all2_Cons2) with T''subT T'isT show ?case by blast next case RedCallNull thus ?case by(fastforce intro: WTThrow[OF WTVal] elim!:typeof_NullPointer simp: sconf_def hconf_def) next case RedCallStatic then show ?case by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError simp: sconf_def hconf_def) next case (SCallParams es h l sh b es' h' l' sh' b' C M) have IH: "\E Ts. \P,E \ (h,l,sh) \; iconfs sh es; P,E,h,sh \ es [:] Ts\ \ \Us. P,E,h',sh' \ es' [:] Us \ P \ Us [\] Ts" and conf: "P,E \ (h,l,sh) \" and iconf: "iconf sh (C\\<^sub>sM(es))" and wt: "P,E,h,sh \ C\\<^sub>sM(es) : T" by fact+ have iconfs: "iconfs sh es" using iconf by simp from wt show ?case proof (rule WTrt_elim_cases) fix D Ts Us pns body sfs vs assume method: "P \ C sees M,Static:Ts\T = (pns,body) in D" and wtes: "P,E,h,sh \ es [:] Us" and us: "P \ Us [\] Ts" and clinit: "M = clinit \ sh D = \(sfs,Processing)\ \ es = map Val vs" obtain Us' where es': "P,E,h',sh' \ es' [:] Us'" and us': "P \ Us' [\] Us" using IH[OF conf iconfs wtes] by blast show ?thesis proof(cases "M = clinit") case True then show ?thesis using clinit SCallParams.hyps(1) by blast next case False then show ?thesis using es' method us us' by(blast intro:WTrtSCall widens_trans) qed qed next case (RedSCall C M Ts T pns body D vs h l sh E T') have "method": "P \ C sees M,Static: Ts\T = (pns,body) in D" and wt: "P,E,h,sh \ C\\<^sub>sM(map Val vs) : T'" by fact+ obtain Ts' where wtes: "P,E,h,sh \ map Val vs [:] Ts'" and subs: "P \ Ts' [\] Ts" and T'isT: "T' = T" using wt "method" map_Val_eq by(auto dest:sees_method_fun)+ from wtes subs have length_vs: "length vs = length Ts" by(fastforce simp:list_all2_iff dest!:WTrts_same_length) from sees_wf_mdecl[OF wf "method"] obtain T'' where wtabody: "P,[pns [\] Ts] \ body :: T''" and T''subT: "P \ T'' \ T" and length_pns: "length pns = length Ts" by(fastforce simp:wf_mdecl_def simp del:map_upds_twist) from wtabody have "P,Map.empty(pns [\] Ts),h,sh \ body : T''" by(rule WT_implies_WTrt) hence "P,E(pns [\] Ts),h,sh \ body : T''" by(rule WTrt_env_mono) simp hence "P,E,h,sh \ blocks(pns, Ts, vs, body) : T''" using wtes subs sees_method_decl_above[OF "method"] length_vs length_pns by(fastforce simp add:wt_blocks rel_list_all2_Cons2) with T''subT T'isT show ?case by blast next case SCallInitDoneRed then show ?case by (meson widen_refl) next case (SCallInitRed C F Ts t pns body D sh v h l E T) have "is_class P D" using SCallInitRed.hyps(1) by(rule sees_method_is_class') then have "P,E,h,sh \ INIT D ([D],False) \ C\\<^sub>sF(map Val v) : T \ P \ T \ T" using SCallInitRed WTrtInit[OF SCallInitRed.prems(3)] by clarsimp then show ?case by(rule exI) next case RedSCallNone then show ?case by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_NoSuchMethodError simp: sconf_def hconf_def) next case RedSCallNonStatic then show ?case by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError simp: sconf_def hconf_def) next case BlockRedNone thus ?case by(auto simp del:fun_upd_apply)(fastforce simp:sconf_def lconf_def) next case (BlockRedSome e h l V sh b e' h' l' sh' b' v T E Te) have red: "P \ \e,(h,l(V:=None),sh),b\ \ \e',(h',l',sh'),b'\" and IH: "\E T. \P,E \ (h,l(V:=None),sh) \; iconf sh e; P,E,h,sh \ e : T\ \ \T'. P,E,h',sh' \ e' : T' \ P \ T' \ T" and Some: "l' V = Some v" and conf: "P,E \ (h,l,sh) \" and iconf: "iconf sh {V:T; e}" and wt: "P,E,h,sh \ {V:T; e} : Te" by fact+ obtain Te' where IH': "P,E(V\T),h',sh' \ e' : Te' \ P \ Te' \ Te" using IH conf iconf wt by(fastforce simp:sconf_def lconf_def) have "P,h' \ l' (:\) E(V\T)" using conf wt by(fastforce intro:red_preserves_lconf[OF red] simp:sconf_def lconf_def) hence "P,h' \ v :\ T" using Some by(fastforce simp:lconf_def) with IH' show ?case by(fastforce simp:sconf_def conf_def fun_upd_same simp del:fun_upd_apply) next case (InitBlockRed e h l V v sh b e' h' l' sh' b' v' T E T') have red: "P \ \e, (h,l(V\v),sh),b\ \ \e',(h',l',sh'),b'\" and IH: "\E T. \P,E \ (h,l(V\v),sh) \; iconf sh e; P,E,h,sh \ e : T\ \ \U. P,E,h',sh' \ e' : U \ P \ U \ T" and v': "l' V = Some v'" and conf: "P,E \ (h,l,sh) \" and iconf: "iconf sh {V:T; V:=Val v;; e}" and wt: "P,E,h,sh \ {V:T := Val v; e} : T'" by fact+ from wt obtain T\<^sub>1 where wt\<^sub>1: "typeof\<^bsub>h\<^esub> v = Some T\<^sub>1" and T1subT: "P \ T\<^sub>1 \ T" and wt\<^sub>2: "P,E(V\T),h,sh \ e : T'" by auto have lconf\<^sub>2: "P,h \ l(V\v) (:\) E(V\T)" using conf wt\<^sub>1 T1subT by(simp add:sconf_def lconf_upd2 conf_def) have "\T\<^sub>1'. typeof\<^bsub>h'\<^esub> v' = Some T\<^sub>1' \ P \ T\<^sub>1' \ T" using v' red_preserves_lconf[OF red wt\<^sub>2 lconf\<^sub>2] by(fastforce simp:lconf_def conf_def) with IH conf iconf lconf\<^sub>2 wt\<^sub>2 show ?case by (fastforce simp add:sconf_def) next case (SeqRed e h l sh b e' h' l' sh' b' e\<^sub>2) then have val: "val_of e = None" by (simp add: val_no_step) show ?case proof(cases "lass_val_of e") case None then show ?thesis using SeqRed val by(auto elim: WTrt_hext_shext_mono[OF _ red_hext_incr red_shext_incr]) next case (Some a) have "sh = sh'" using SeqRed lass_val_of_spec[OF Some] by auto then show ?thesis using SeqRed val Some by(auto intro: lass_val_of_iconf[OF Some] elim: WTrt_hext_mono[OF _ red_hext_incr]) qed next case CondRed thus ?case by auto (blast intro:WTrt_hext_shext_mono[OF _ red_hext_incr red_shext_incr])+ next case ThrowRed thus ?case by(auto simp:is_refT_def)(blast dest:widen_Class[THEN iffD1])+ next case RedThrowNull thus ?case by(fastforce intro: WTThrow[OF WTVal] elim!:typeof_NullPointer simp:sconf_def hconf_def) next case TryRed thus ?case by auto (blast intro:widen_trans WTrt_hext_shext_mono[OF _ red_hext_incr red_shext_incr]) next case RedTryFail thus ?case by(fastforce intro: WTrtThrow[OF WTrtVal] simp:sconf_def hconf_def) next case (ListRed1 e h l sh b e' h' l' sh' b' es) then have val: "val_of e = None" by(simp add: val_no_step) obtain U Us where Ts: "Ts = U # Us" using ListRed1 by auto then have nsub_RI: "\ sub_RIs es" and wts: "P,E,h,sh \ es [:] Us" and wt: "P,E,h,sh \ e : U" and IH: "\E T. \P,E \ (h, l, sh) \; P,E,h,sh \ e : T\ \ \T'. P,E,h',sh' \ e' : T' \ P \ T' \ T" using ListRed1 val by auto obtain T' where "\E0 E1. (\T2. P,E1,h',sh' \ e' : T2 \ P \ T2 \ E0) = (P,E1,h',sh' \ e' : T' E0 E1 \ P \ T' E0 E1 \ E0)" by moura then have disj: "\E t. \ P,E \ (h, l, sh) \ \ \ P,E,h,sh \ e : t \ P,E,h',sh' \ e' : T' t E \ P \ T' t E \ t" using IH by presburger have "P,E,h',sh' \ es [:] Us" using nsub_RI wts wt by (metis (no_types) ListRed1.hyps(1) WTrts_hext_shext_mono red_hext_incr red_shext_incr) then have "\ts. (\t tsa. ts = t # tsa \ P,E,h',sh' \ e' : t \ P,E,h',sh' \ es [:] tsa) \ P \ ts [\] (U # Us)" using disj wt ListRed1.prems(1) by blast then show ?case using Ts by auto next case ListRed2 thus ?case by(fastforce dest: hext_typeof_mono[OF reds_hext_incr]) next case (InitNoneRed sh C C' Cs e h l b) then have sh: "sh \\<^sub>s sh(C \ (sblank P C, Prepared))" by(simp add: shext_def) have wt: "P,E,h,sh(C \ (sblank P C, Prepared)) \ INIT C' (C # Cs,False) \ e : T" using InitNoneRed WTrt_shext_mono[OF _ sh] by fastforce then show ?case by(rule_tac x = T in exI) (simp add: fun_upd_def) next case (RedInitDone sh C sfs C' Cs e h l b) then have "P,E,h,sh \ INIT C' (Cs,True) \ e : T" by auto (metis Nil_tl list.set_sel(2)) then show ?case by(rule_tac x = T in exI) simp next case (RedInitProcessing sh C sfs C' Cs e h l b) then have "P,E,h,sh \ INIT C' (Cs,True) \ e : T" by auto (metis Nil_tl list.set_sel(2))+ then show ?case by(rule_tac x = T in exI) simp next case RedInitError then show ?case by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_NoClassDefFoundError simp: sconf_def hconf_def) next case (InitObjectRed sh C sfs sh' C' Cs e h l b) then have sh: "sh \\<^sub>s sh(Object \ (sfs, Processing))" by(simp add: shext_def) have "P,E,h,sh' \ INIT C' (C # Cs,True) \ e : T" using InitObjectRed WTrt_shext_mono[OF _ sh] by auto then show ?case by(rule_tac x = T in exI) (simp add: fun_upd_def) next case (InitNonObjectSuperRed sh C sfs D fs ms sh' C' Cs e h l b) then have sh: "sh \\<^sub>s sh(C \ (sfs, Processing))" by(simp add: shext_def) then have cd: "is_class P D" using InitNonObjectSuperRed class_wf wf wf_cdecl_supD by blast have sup': "supercls_lst P (C # Cs)" using InitNonObjectSuperRed.prems(3) by auto then have sup: "supercls_lst P (D # C # Cs)" using supercls_lst_app[of P C Cs D] subcls1I[OF InitNonObjectSuperRed.hyps(3,2)] by auto have "distinct (C # Cs)" using InitNonObjectSuperRed.prems(3) by auto then have dist: "distinct (D # C # Cs)" using wf_supercls_distinct_app[OF wf InitNonObjectSuperRed.hyps(2-3) sup'] by simp have "P,E,h,sh' \ INIT C' (D # C # Cs,False) \ e : T" using InitNonObjectSuperRed WTrt_shext_mono[OF _ sh] cd sup dist by auto then show ?case by(rule_tac x = T in exI) simp next case (RedInitRInit C' C Cs e' h l sh b E T) then obtain a sfs where C: "class P C = \a\" and proc: "sh C = \(sfs, Processing)\" using WTrtInit by(auto simp: is_class_def) then have T': "P,E,h,sh \ C\\<^sub>sclinit([]) : Void" using wf_types_clinit[OF wf C] by simp have "P,E,h,sh \ RI (C,C\\<^sub>sclinit([])) ; Cs \ e' : T" using RedInitRInit by(auto intro: T') then show ?case by(rule_tac x = T in exI) simp next case (RInitRed e h l sh b e' h' l' sh' b' C Cs e\<^sub>0 E T) then have "(\E T. P,E \ (h, l, sh) \ \ P,E,h,sh \ e : T \ \T'. P,E,h',sh' \ e' : T' \ P \ T' \ T)" by auto then have "\T'. P,E,h',sh' \ e' : T'" using RInitRed by blast then obtain T' where e': "P,E,h',sh' \ e' : T'" by auto have wt\<^sub>0: "P,E,h',sh' \ e\<^sub>0 : T" using RInitRed by simp (auto intro: WTrt_hext_shext_mono[OF _ red_hext_incr red_shext_incr]) have nip: "\C' \ set (C#Cs). not_init C' e' \ (\sfs. sh' C' = \(sfs, Processing)\)" using RInitRed red_proc_pres[OF wf_prog_wwf_prog[OF wf]] by auto have shC: "\sfs. sh' C = \(sfs, Processing)\ \ sh' C = \(sfs, Error)\ \ e' = THROW NoClassDefFoundError" using RInitRed red_proc_pres[OF wf_prog_wwf_prog[OF wf] RInitRed.hyps(1)] by blast have "P,E,h',sh' \ RI (C,e') ; Cs \ e\<^sub>0 : T" using RInitRed e' wt\<^sub>0 nip shC by auto then show ?case by(rule_tac x = T in exI) simp next case (RedRInit sh C sfs i sh' C' Cs v e h l b) then have sh: "sh \\<^sub>s sh(C \ (sfs, Done))" by(auto simp: shext_def) have wt: "P,E,h,sh(C \ (sfs, Done)) \ e : T" using RedRInit WTrt_shext_mono[OF _ sh] by auto have shC: "\C' \ set(tl Cs). \sfs. sh C' = \(sfs, Processing)\" using RedRInit by(cases Cs, auto) have "P,E,h,sh' \ INIT C' (Cs,True) \ e : T" using RedRInit wt shC by(cases Cs, auto) then show ?case by(rule_tac x = T in exI) simp next case (SCallThrowParams es vs e es' C M h l sh b) then show ?case using map_Val_nthrow_neq[of _ vs e es'] by fastforce next case (RInitInitThrow sh C sfs i sh' a D Cs e h l b) then have sh: "sh \\<^sub>s sh(C \ (sfs, Error))" by(auto simp: shext_def) have wt: "P,E,h,sh(C \ (sfs, Error)) \ e : T" using RInitInitThrow WTrt_shext_mono[OF _ sh] by clarsimp then have "P,E,h,sh' \ RI (D,Throw a) ; Cs \ e : T" using RInitInitThrow by auto then show ?case by(rule_tac x = T in exI) simp qed fastforce+ (* esp all Throw propagation rules except RInitInit are dealt with here *) (*>*) corollary subject_reduction: "\ wf_J_prog P; P \ \e,s,b\ \ \e',s',b'\; P,E \ s \; iconf (shp s) e; P,E,hp s,shp s \ e:T \ \ \T'. P,E,hp s',shp s' \ e':T' \ P \ T' \ T" (*<*)by(cases s, cases s', fastforce dest:subject_reduction2)(*>*) corollary subjects_reduction: "\ wf_J_prog P; P \ \es,s,b\ [\] \es',s',b'\; P,E \ s \; iconfs (shp s) es; P,E,hp s,shp s \ es[:]Ts \ \ \Ts'. P,E,hp s',shp s' \ es'[:]Ts' \ P \ Ts' [\] Ts" (*<*)by(cases s, cases s', fastforce dest:subjects_reduction2)(*>*) subsection \ Lifting to @{text"\*"} \ text\ Now all these preservation lemmas are first lifted to the transitive closure \dots \ lemma Red_preserves_sconf: assumes wf: "wf_J_prog P" and Red: "P \ \e,s,b\ \* \e',s',b'\" shows "\T. \ P,E,hp s,shp s \ e : T; iconf (shp s) e; P,E \ s \ \ \ P,E \ s' \" (*<*) using Red proof (induct rule:converse_rtrancl_induct3) case refl show ?case by fact next case (step e s b e' s' b') obtain h l sh h' l' sh' where s:"s = (h,l,sh)" and s':"s' = (h',l',sh')" by(cases s, cases s') then have "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\" using step.hyps(1) by simp then have iconf': "iconf (shp s') e'" using red_preserves_iconf[OF wf_prog_wwf_prog[OF wf]] step.prems(2) s s' by simp thus ?case using step by(blast intro:red_preserves_sconf dest: subject_reduction[OF wf]) qed (*>*) lemma Red_preserves_iconf: assumes wf: "wwf_J_prog P" and Red: "P \ \e,s,b\ \* \e',s',b'\" shows "iconf (shp s) e \ iconf (shp s') e'" (*<*) using Red proof (induct rule:converse_rtrancl_induct3) case refl show ?case by fact next case (step e s b e' s' b') thus ?case using wf step by(cases s, cases s', simp) (blast intro:red_preserves_iconf) qed (*>*) lemma Reds_preserves_iconf: assumes wf: "wwf_J_prog P" and Red: "P \ \es,s,b\ [\]* \es',s',b'\" shows "iconfs (shp s) es \ iconfs (shp s') es'" (*<*) using Red proof (induct rule:converse_rtrancl_induct3) case refl show ?case by fact next case (step e s b e' s' b') thus ?case using wf step by(cases s, cases s', simp) (blast intro:reds_preserves_iconf) qed (*>*) lemma Red_preserves_bconf: assumes wf: "wwf_J_prog P" and Red: "P \ \e,s,b\ \* \e',s',b'\" shows "iconf (shp s) e \ P,(shp s) \\<^sub>b (e,b) \ \ P,(shp s') \\<^sub>b (e'::expr,b') \" (*<*) using Red proof (induct rule:converse_rtrancl_induct3) case refl show ?case by fact next case (step e s1 b e' s2 b') then have "iconf (shp s2) e'" using step red_preserves_iconf[OF wf] by(cases s1, cases s2) auto thus ?case using step by(cases s1, cases s2, simp) (blast intro:red_preserves_bconf) qed (*>*) lemma Reds_preserves_bconf: assumes wf: "wwf_J_prog P" and Red: "P \ \es,s,b\ [\]* \es',s',b'\" shows "iconfs (shp s) es \ P,(shp s) \\<^sub>b (es,b) \ \ P,(shp s') \\<^sub>b (es'::expr list,b') \" (*<*) using Red proof (induct rule:converse_rtrancl_induct3) case refl show ?case by fact next case (step es s1 b es' s2 b') then have "iconfs (shp s2) es'" using step reds_preserves_iconf[OF wf] by(cases s1, cases s2) auto thus ?case using step by(cases s1, cases s2, simp) (blast intro:reds_preserves_bconf) qed (*>*) lemma Red_preserves_defass: assumes wf: "wf_J_prog P" and reds: "P \ \e,s,b\ \* \e',s',b'\" shows "\ e \dom(lcl s)\ \ \ e' \dom(lcl s')\" using reds proof (induct rule:converse_rtrancl_induct3) case refl thus ?case . next case (step e s b e' s' b') thus ?case by(cases s,cases s')(auto dest:red_preserves_defass[OF wf]) qed lemma Red_preserves_type: assumes wf: "wf_J_prog P" and Red: "P \ \e,s,b\ \* \e',s',b'\" shows "!!T. \ P,E \ s\; iconf (shp s) e; P,E,hp s,shp s \ e:T \ \ \T'. P \ T' \ T \ P,E,hp s',shp s' \ e':T'" (*<*) using Red proof (induct rule:converse_rtrancl_induct3) case refl thus ?case by blast next case step thus ?case by(blast intro:widen_trans red_preserves_sconf Red_preserves_iconf[OF wf_prog_wwf_prog[OF wf]] dest:subject_reduction[OF wf]) qed (*>*) subsection "The final polish" text\ The above preservation lemmas are now combined and packed nicely. \ definition wf_config :: "J_prog \ env \ state \ expr \ ty \ bool" ("_,_,_ \ _ : _ \" [51,0,0,0,0]50) where "P,E,s \ e:T \ \ P,E \ s \ \ iconf (shp s) e \ P,E,hp s,shp s \ e:T" theorem Subject_reduction: assumes wf: "wf_J_prog P" shows "P \ \e,s,b\ \ \e',s',b'\ \ P,E,s \ e : T \ \ \T'. P,E,s' \ e' : T' \ \ P \ T' \ T" (*<*) by(cases s, cases s') (force simp: wf_config_def elim:red_preserves_sconf red_preserves_iconf[OF wf_prog_wwf_prog[OF wf]] dest:subject_reduction[OF wf]) (*>*) theorem Subject_reductions: assumes wf: "wf_J_prog P" and reds: "P \ \e,s,b\ \* \e',s',b'\" shows "\T. P,E,s \ e:T \ \ \T'. P,E,s' \ e':T' \ \ P \ T' \ T" (*<*) using reds proof (induct rule:converse_rtrancl_induct3) case refl thus ?case by blast next case step thus ?case by(blast dest:Subject_reduction[OF wf] intro:widen_trans) qed (*>*) corollary Progress: assumes wf: "wf_J_prog P" shows "\ P,E,s \ e : T \; \ e \dom(lcl s)\; P,shp s \\<^sub>b (e,b) \; \ final e \ \ \e' s' b'. P \ \e,s,b\ \ \e',s',b'\" (*<*) using progress[OF wf_prog_wwf_prog[OF wf]] by(cases b) (auto simp:wf_config_def sconf_def) (*>*) corollary TypeSafety: - "\ wf_J_prog P; P,E \ s \; P,E \ e::T; \ e \dom(lcl s)\; - iconf (shp s) e; P,(shp s) \\<^sub>b (e,b) \; - P \ \e,s,b\ \* \e',s',b'\; \(\e'' s'' b''. P \ \e',s',b'\ \ \e'',s'',b''\) \ - \ (\v. e' = Val v \ P,hp s' \ v :\ T) \ +fixes s::state and e::expr +assumes wf: "wf_J_prog P" and sconf: "P,E \ s \" and wt: "P,E \ e::T" + and \: "\ e \dom(lcl s)\" + and iconf: "iconf (shp s) e" and bconf: "P,(shp s) \\<^sub>b (e,b) \" + and steps: "P \ \e,s,b\ \* \e',s',b'\" + and nstep: "\(\e'' s'' b''. P \ \e',s',b'\ \ \e'',s'',b''\)" +shows "(\v. e' = Val v \ P,hp s' \ v :\ T) \ (\a. e' = Throw a \ a \ dom(hp s'))" (*<*) -apply(subgoal_tac "wwf_J_prog P") - prefer 2 apply(rule wf_prog_wwf_prog, simp) -apply(subgoal_tac " P,E,s \ e:T \") - prefer 2 - apply(fastforce simp:wf_config_def dest:WT_implies_WTrt) -apply(frule (2) Subject_reductions) -apply(erule exE conjE)+ -apply(frule (2) Red_preserves_defass) -apply(frule (3) Red_preserves_bconf) -apply(subgoal_tac "final e'") - prefer 2 - apply(blast dest: Progress) -apply (fastforce simp:wf_config_def final_def conf_def dest: Progress) -done +proof - + have wwf: "wwf_J_prog P" by(rule wf_prog_wwf_prog[OF wf]) + have wfc: "P,E,s \ e:T \" using WT_implies_WTrt[OF wt] sconf iconf + by(simp add:wf_config_def) + obtain T' where wfc': "P,E,s' \ e' : T' \" and T': "P \ T' \ T" + using Subject_reductions[OF wf steps wfc] by clarsimp + have \': "\ e' \dom (lcl s')\" + by(rule Red_preserves_defass[OF wf steps \]) + have bconf': "P,(shp s') \\<^sub>b (e',b') \" + by(rule Red_preserves_bconf[OF wwf steps iconf bconf]) + have fin': "final e'" using Progress[OF wf wfc' \' bconf'] nstep by blast + then show ?thesis using wfc wfc' T' + by(fastforce simp:wf_config_def final_def conf_def) +qed (*>*) end