diff --git a/thys/Jinja/BV/BVConform.thy b/thys/Jinja/BV/BVConform.thy --- a/thys/Jinja/BV/BVConform.thy +++ b/thys/Jinja/BV/BVConform.thy @@ -1,152 +1,154 @@ (* Title: HOL/MicroJava/BV/Correct.thy Author: Cornelia Pusch, Gerwin Klein Copyright 1999 Technische Universitaet Muenchen The invariant for the type safety proof. *) section \BV Type Safety Invariant\ theory BVConform imports BVSpec "../JVM/JVMExec" "../Common/Conform" begin definition confT :: "'c prog \ heap \ val \ ty err \ bool" ("_,_ \ _ :\\<^sub>\ _" [51,51,51,51] 50) where "P,h \ v :\\<^sub>\ E \ case E of Err \ True | OK T \ P,h \ v :\ T" notation (ASCII) confT ("_,_ |- _ :<=T _" [51,51,51,51] 50) abbreviation confTs :: "'c prog \ heap \ val list \ ty\<^sub>l \ bool" ("_,_ \ _ [:\\<^sub>\] _" [51,51,51,51] 50) where "P,h \ vs [:\\<^sub>\] Ts \ list_all2 (confT P h) vs Ts" notation (ASCII) confTs ("_,_ |- _ [:<=T] _" [51,51,51,51] 50) definition conf_f :: "jvm_prog \ heap \ ty\<^sub>i \ bytecode \ frame \ bool" where "conf_f P h \ \(ST,LT) is (stk,loc,C,M,pc). P,h \ stk [:\] ST \ P,h \ loc [:\\<^sub>\] LT \ pc < size is" lemma conf_f_def2: "conf_f P h (ST,LT) is (stk,loc,C,M,pc) \ P,h \ stk [:\] ST \ P,h \ loc [:\\<^sub>\] LT \ pc < size is" by (simp add: conf_f_def) primrec conf_fs :: "[jvm_prog,heap,ty\<^sub>P,mname,nat,ty,frame list] \ bool" where "conf_fs P h \ M\<^sub>0 n\<^sub>0 T\<^sub>0 [] = True" | "conf_fs P h \ M\<^sub>0 n\<^sub>0 T\<^sub>0 (f#frs) = (let (stk,loc,C,M,pc) = f in (\ST LT Ts T mxs mxl\<^sub>0 is xt. \ C M ! pc = Some (ST,LT) \ (P \ C sees M:Ts \ T = (mxs,mxl\<^sub>0,is,xt) in C) \ (\D Ts' T' m D'. is!pc = (Invoke M\<^sub>0 n\<^sub>0) \ ST!n\<^sub>0 = Class D \ P \ D sees M\<^sub>0:Ts' \ T' = m in D' \ P \ T\<^sub>0 \ T') \ conf_f P h (ST, LT) is f \ conf_fs P h \ M (size Ts) T frs))" definition correct_state :: "[jvm_prog,ty\<^sub>P,jvm_state] \ bool" ("_,_ \ _ \" [61,0,0] 61) where "correct_state P \ \ \(xp,h,frs). case xp of None \ (case frs of [] \ True | (f#fs) \ P\ h\ \ (let (stk,loc,C,M,pc) = f in \Ts T mxs mxl\<^sub>0 is xt \. (P \ C sees M:Ts\T = (mxs,mxl\<^sub>0,is,xt) in C) \ \ C M ! pc = Some \ \ conf_f P h \ is f \ conf_fs P h \ M (size Ts) T fs)) | Some x \ frs = []" notation correct_state ("_,_ |- _ [ok]" [61,0,0] 61) subsection \Values and \\\\ lemma confT_Err [iff]: "P,h \ x :\\<^sub>\ Err" by (simp add: confT_def) lemma confT_OK [iff]: "P,h \ x :\\<^sub>\ OK T = (P,h \ x :\ T)" by (simp add: confT_def) lemma confT_cases: "P,h \ x :\\<^sub>\ X = (X = Err \ (\T. X = OK T \ P,h \ x :\ T))" by (cases X) auto lemma confT_hext [intro?, trans]: "\ P,h \ x :\\<^sub>\ T; h \ h' \ \ P,h' \ x :\\<^sub>\ T" by (cases T) (blast intro: conf_hext)+ lemma confT_widen [intro?, trans]: "\ P,h \ x :\\<^sub>\ T; P \ T \\<^sub>\ T' \ \ P,h \ x :\\<^sub>\ T'" by (cases T', auto intro: conf_widen) subsection \Stack and Registers\ lemmas confTs_Cons1 [iff] = list_all2_Cons1 [of "confT P h"] for P h lemma confTs_confT_sup: - "\ P,h \ loc [:\\<^sub>\] LT; n < size LT; LT!n = OK T; P \ T \ T' \ - \ P,h \ (loc!n) :\ T'" +assumes confTs: "P,h \ loc [:\\<^sub>\] LT" and n: "n < size LT" and + LTn: "LT!n = OK T" and subtype: "P \ T \ T'" +shows "P,h \ (loc!n) :\ T'" (*<*) - apply (frule list_all2_lengthD) - apply (drule list_all2_nthD, simp) - apply simp - apply (erule conf_widen, assumption+) - done +proof - + have len: "n < length loc" using list_all2_lengthD[OF confTs] n + by simp + show ?thesis + using list_all2_nthD[OF confTs len] conf_widen[OF _ subtype] LTn + by simp +qed (*>*) lemma confTs_hext [intro?]: "P,h \ loc [:\\<^sub>\] LT \ h \ h' \ P,h' \ loc [:\\<^sub>\] LT" by (fast elim: list_all2_mono confT_hext) lemma confTs_widen [intro?, trans]: "P,h \ loc [:\\<^sub>\] LT \ P \ LT [\\<^sub>\] LT' \ P,h \ loc [:\\<^sub>\] LT'" by (rule list_all2_trans, rule confT_widen) lemma confTs_map [iff]: "\vs. (P,h \ vs [:\\<^sub>\] map OK Ts) = (P,h \ vs [:\] Ts)" by (induct Ts) (auto simp add: list_all2_Cons2) lemma reg_widen_Err [iff]: "\LT. (P \ replicate n Err [\\<^sub>\] LT) = (LT = replicate n Err)" by (induct n) (auto simp add: list_all2_Cons1) lemma confTs_Err [iff]: "P,h \ replicate n v [:\\<^sub>\] replicate n Err" by (induct n) auto subsection \correct-frames\ lemmas [simp del] = fun_upd_apply lemma conf_fs_hext: "\M n T\<^sub>r. \ conf_fs P h \ M n T\<^sub>r frs; h \ h' \ \ conf_fs P h' \ M n T\<^sub>r frs" (*<*) -apply (induct frs) - apply simp -apply clarify -apply (simp (no_asm_use)) -apply clarify -apply (unfold conf_f_def) -apply (simp (no_asm_use)) -apply clarify -apply (fast elim!: confs_hext confTs_hext) -done +proof(induct frs) + case (Cons fr frs) + obtain stk ls C M pc where fr: "fr = (stk, ls, C, M, pc)" by(cases fr) simp + moreover obtain ST LT where \: "\ C M ! pc = \(ST, LT)\" and + ST: "P,h \ stk [:\] ST" and LT: "P,h \ ls [:\\<^sub>\] LT" + using Cons.prems(1) fr by(auto simp: conf_f_def) + ultimately show ?case using Cons confs_hext[OF ST Cons(3)] confTs_hext[OF LT Cons(3)] + by (fastforce simp: conf_f_def) +qed simp (*>*) end diff --git a/thys/Jinja/BV/BVExec.thy b/thys/Jinja/BV/BVExec.thy --- a/thys/Jinja/BV/BVExec.thy +++ b/thys/Jinja/BV/BVExec.thy @@ -1,233 +1,247 @@ (* Title: HOL/MicroJava/BV/JVM.thy Author: Tobias Nipkow, Gerwin Klein Copyright 2000 TUM *) section \Kildall for the JVM \label{sec:JVM}\ theory BVExec imports "../DFA/Abstract_BV" TF_JVM begin definition kiljvm :: "jvm_prog \ nat \ nat \ ty \ instr list \ ex_table \ ty\<^sub>i' err list \ ty\<^sub>i' err list" where "kiljvm P mxs mxl T\<^sub>r is xt \ kildall (JVM_SemiType.le P mxs mxl) (JVM_SemiType.sup P mxs mxl) (exec P mxs T\<^sub>r xt is)" definition wt_kildall :: "jvm_prog \ cname \ ty list \ ty \ nat \ nat \ instr list \ ex_table \ bool" where "wt_kildall P C' Ts T\<^sub>r mxs mxl\<^sub>0 is xt \ 0 < size is \ (let first = Some ([],[OK (Class C')]@(map OK Ts)@(replicate mxl\<^sub>0 Err)); start = OK first#(replicate (size is - 1) (OK None)); result = kiljvm P mxs (1+size Ts+mxl\<^sub>0) T\<^sub>r is xt start in \n < size is. result!n \ Err)" definition wf_jvm_prog\<^sub>k :: "jvm_prog \ bool" where "wf_jvm_prog\<^sub>k P \ wf_prog (\P C' (M,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). wt_kildall P C' Ts T\<^sub>r mxs mxl\<^sub>0 is xt) P" theorem (in start_context) is_bcv_kiljvm: "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)" (*<*) - apply (insert wf) - apply (unfold kiljvm_def) - apply (fold r_def f_def step_def_exec) - apply (rule is_bcv_kildall) - apply simp apply (rule Semilat.intro) - apply (fold sl_def2) - apply (erule semilat_JVM) - apply simp - apply blast - apply (simp add: JVM_le_unfold) - apply (rule exec_pres_type) - apply (rule bounded_step) - apply (erule step_mono) - done +proof - + let ?n = "length is" + have "Semilat A r f" using semilat_JVM[OF wf] + by (simp add: Semilat.intro sl_def2) + moreover have "acc r" using wf by simp blast + moreover have "top r Err" by (simp add: JVM_le_unfold) + moreover have "pres_type step ?n A" by (rule exec_pres_type) + moreover have "bounded step ?n" by (rule bounded_step) + moreover have "mono r step ?n A" using step_mono[OF wf] by simp + ultimately have "is_bcv r Err step ?n A (kildall r f step)" + by(rule is_bcv_kildall) + moreover have kileq: "kiljvm P mxs mxl T\<^sub>r is xt = kildall r f step" + using f_def kiljvm_def r_def step_def_exec by blast + ultimately show ?thesis by simp +qed (*>*) (* FIXME: move? *) lemma subset_replicate [intro?]: "set (replicate n x) \ {x}" by (induct n) auto lemma in_set_replicate: assumes "x \ set (replicate n y)" shows "x = y" (*<*) proof - note assms also have "set (replicate n y) \ {y}" .. finally show ?thesis by simp qed (*>*) lemma (in start_context) start_in_A [intro?]: "0 < size is \ start \ list (size is) A" using Ts C (*<*) - apply (simp add: JVM_states_unfold) - apply (force intro!: listI list_appendI dest!: in_set_replicate) - done + by (force simp add: JVM_states_unfold intro!: listI list_appendI + dest!: in_set_replicate) (*>*) theorem (in start_context) wt_kil_correct: assumes wtk: "wt_kildall P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt" shows "\\s. wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" (*<*) proof - from wtk obtain res where result: "res = kiljvm P mxs mxl T\<^sub>r is xt start" and success: "\n < size is. res!n \ Err" and instrs: "0 < size is" by (unfold wt_kildall_def) simp have bcv: "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)" by (rule is_bcv_kiljvm) from instrs have "start \ list (size is) A" .. with bcv success result have "\ts\list (size is) A. start [\\<^sub>r] ts \ wt_step r Err step ts" by (unfold is_bcv_def) blast then obtain \s' where in_A: "\s' \ list (size is) A" and s: "start [\\<^sub>r] \s'" and w: "wt_step r Err step \s'" by blast hence wt_err_step: "wt_err_step (sup_state_opt P) step \s'" by (simp add: wt_err_step_def JVM_le_Err_conv) from in_A have l: "size \s' = size is" by simp moreover { from in_A have "check_types P mxs mxl \s'" by (simp add: check_types_def) also from w have "\x \ set \s'. x \ Err" by (auto simp add: wt_step_def all_set_conv_all_nth) hence [symmetric]: "map OK (map ok_val \s') = \s'" by (auto intro!: map_idI simp add: wt_step_def) finally have "check_types P mxs mxl (map OK (map ok_val \s'))" . } moreover { from s have "start!0 \\<^sub>r \s'!0" by (rule le_listD) simp moreover from instrs w l have "\s'!0 \ Err" by (unfold wt_step_def) simp then obtain \s0 where "\s'!0 = OK \s0" by auto ultimately have "wt_start P C Ts mxl\<^sub>0 (map ok_val \s')" using l instrs by (unfold wt_start_def) (simp add: lesub_def JVM_le_Err_conv Err.le_def) } moreover from in_A have "set \s' \ A" by simp with wt_err_step bounded_step have "wt_app_eff (sup_state_opt P) app eff (map ok_val \s')" by (auto intro: wt_err_imp_wt_app_eff simp add: l) ultimately have "wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (map ok_val \s')" using instrs by (simp add: wt_method_def2 check_types_def del: map_map) thus ?thesis by blast qed (*>*) theorem (in start_context) wt_kil_complete: assumes wtm: "wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" shows "wt_kildall P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt" (*<*) proof - from wtm obtain instrs: "0 < size is" and length: "length \s = length is" and ck_type: "check_types P mxs mxl (map OK \s)" and wt_start: "wt_start P C Ts mxl\<^sub>0 \s" and app_eff: "wt_app_eff (sup_state_opt P) app eff \s" by (simp add: wt_method_def2 check_types_def) from ck_type have in_A: "set (map OK \s) \ A" by (simp add: check_types_def) with app_eff in_A bounded_step have "wt_err_step (sup_state_opt P) (err_step (size \s) app eff) (map OK \s)" by - (erule wt_app_eff_imp_wt_err, auto simp add: exec_def length states_def) hence wt_err: "wt_err_step (sup_state_opt P) step (map OK \s)" by (simp add: length) have is_bcv: "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)" by (rule is_bcv_kiljvm) moreover from instrs have "start \ list (size is) A" .. moreover let ?\s = "map OK \s" have less_\s: "start [\\<^sub>r] ?\s" proof (rule le_listI) from length instrs show "length start = length (map OK \s)" by simp next fix n from wt_start have "P \ ok_val (start!0) \' \s!0" by (simp add: wt_start_def) moreover from instrs length have "0 < length \s" by simp ultimately have "start!0 \\<^sub>r ?\s!0" by (simp add: JVM_le_Err_conv lesub_def) moreover { fix n' have "OK None \\<^sub>r ?\s!n" by (auto simp add: JVM_le_Err_conv Err.le_def lesub_def split: err.splits) hence "\n = Suc n'; n < size start\ \ start!n \\<^sub>r ?\s!n" by simp } ultimately show "n < size start \ start!n \\<^sub>r ?\s!n" by (cases n, blast+) qed moreover from ck_type length have "?\s \ list (size is) A" by (auto intro!: listI simp add: check_types_def) moreover from wt_err have "wt_step r Err step ?\s" by (simp add: wt_err_step_def JVM_le_Err_conv) ultimately have "\p. p < size is \ kiljvm P mxs mxl T\<^sub>r is xt start ! p \ Err" by (unfold is_bcv_def) blast with instrs show "wt_kildall P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt" by (unfold wt_kildall_def) simp qed (*>*) theorem jvm_kildall_correct: "wf_jvm_prog\<^sub>k P = wf_jvm_prog P" (*<*) -proof +proof - let ?\ = "\C M. let (C,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)) = method P C M in SOME \s. wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" + let ?A = "\P C' (M, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). wt_kildall P C' Ts T\<^sub>r mxs mxl\<^sub>0 is xt" + let ?B\<^sub>\ = "\\. (\P C (M,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). + wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M))" - \ \soundness\ - assume wt: "wf_jvm_prog\<^sub>k P" - hence "wf_jvm_prog\<^bsub>?\\<^esub> P" - apply (unfold wf_jvm_prog_phi_def wf_jvm_prog\<^sub>k_def) - apply (erule wf_prog_lift) - apply (auto dest!: start_context.wt_kil_correct [OF start_context.intro] - intro: someI) - apply (erule sees_method_is_class) - done - thus "wf_jvm_prog P" by (unfold wf_jvm_prog_def) fast -next - \ \completeness\ - assume wt: "wf_jvm_prog P" - thus "wf_jvm_prog\<^sub>k P" - apply (unfold wf_jvm_prog_def wf_jvm_prog_phi_def wf_jvm_prog\<^sub>k_def) - apply (clarify) - apply (erule wf_prog_lift) - apply (auto intro!: start_context.wt_kil_complete start_context.intro) - apply (erule sees_method_is_class) - done + show ?thesis proof(rule iffI) + \ \soundness\ + assume wt: "wf_jvm_prog\<^sub>k P" + then have "wf_prog ?A P" by(simp add: wf_jvm_prog\<^sub>k_def) + moreover { + fix wf_md C M Ts Ca T m bd + assume "wf_prog wf_md P" and sees: "P \ Ca sees M : Ts\T = m in Ca" and + "set Ts \ types P" and "bd = (M, Ts, T, m)" and + "?A P Ca bd" + then have "(?B\<^sub>\ ?\) P Ca bd" using sees_method_is_class[OF sees] + by (auto dest!: start_context.wt_kil_correct [OF start_context.intro] + intro: someI) + } + ultimately have "wf_prog (?B\<^sub>\ ?\) P" by(rule wf_prog_lift) + then have "wf_jvm_prog\<^bsub>?\\<^esub> P" by (simp add: wf_jvm_prog_phi_def) + thus "wf_jvm_prog P" by (unfold wf_jvm_prog_def) fast + next + \ \completeness\ + assume wt: "wf_jvm_prog P" + then obtain \ where "wf_prog (?B\<^sub>\ \) P" by(clarsimp simp: wf_jvm_prog_def wf_jvm_prog_phi_def) + moreover { + fix wf_md C M b Ts Ca T m bd + assume "wf_prog wf_md P" and sees: "P \ Ca sees M: Ts\T = m in Ca" and + "set Ts \ types P" and "bd = (M, Ts, T, m)" and + "(?B\<^sub>\ \) P Ca bd" + then have "?A P Ca bd" using sees_method_is_class[OF sees] + by (auto intro!: start_context.wt_kil_complete start_context.intro) + } + ultimately have "wf_prog ?A P" by(rule wf_prog_lift) + thus "wf_jvm_prog\<^sub>k P" by (simp add: wf_jvm_prog\<^sub>k_def) + qed qed (*>*) end diff --git a/thys/Jinja/BV/BVNoTypeError.thy b/thys/Jinja/BV/BVNoTypeError.thy --- a/thys/Jinja/BV/BVNoTypeError.thy +++ b/thys/Jinja/BV/BVNoTypeError.thy @@ -1,312 +1,309 @@ (* Title: HOL/MicroJava/BV/BVNoTypeErrors.thy Author: Gerwin Klein Copyright GPL *) section \Welltyped Programs produce no Type Errors\ theory BVNoTypeError imports "../JVM/JVMDefensive" BVSpecTypeSafe begin lemma has_methodI: "P \ C sees M:Ts\T = m in D \ P \ C has M" by (unfold has_method_def) blast text \ Some simple lemmas about the type testing functions of the defensive JVM: \ lemma typeof_NoneD [simp,dest]: "typeof v = Some x \ \is_Addr v" by (cases v) auto lemma is_Ref_def2: "is_Ref v = (v = Null \ (\a. v = Addr a))" by (cases v) (auto simp add: is_Ref_def) lemma [iff]: "is_Ref Null" by (simp add: is_Ref_def2) lemma is_RefI [intro, simp]: "P,h \ v :\ T \ is_refT T \ is_Ref v" (*<*) - apply (cases T) - apply (auto simp add: is_refT_def is_Ref_def dest: conf_ClassD) - done +proof (cases T) +qed (auto simp add: is_refT_def is_Ref_def dest: conf_ClassD) (*>*) lemma is_IntgI [intro, simp]: "P,h \ v :\ Integer \ is_Intg v" -(*<*) - apply (unfold conf_def) - apply auto - done -(*>*) +(*<*)by (unfold conf_def) auto(*>*) lemma is_BoolI [intro, simp]: "P,h \ v :\ Boolean \ is_Bool v" -(*<*) - apply (unfold conf_def) - apply auto - done -(*>*) +(*<*)by (unfold conf_def) auto(*>*) declare defs1 [simp del] lemma wt_jvm_prog_states: - "\ wf_jvm_prog\<^bsub>\\<^esub> P; P \ C sees M: Ts\T = (mxs, mxl, ins, et) in C; - \ C M ! pc = \; pc < size ins \ - \ OK \ \ states P mxs (1+size Ts+mxl)" +assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" + and mC: "P \ C sees M: Ts\T = (mxs, mxl, ins, et) in C" + and \: "\ C M ! pc = \" and pc: "pc < size ins" +shows "OK \ \ states P mxs (1+size Ts+mxl)" (*<*) - apply (unfold wf_jvm_prog_phi_def) - apply (drule (1) sees_wf_mdecl) - apply (simp add: wf_mdecl_def wt_method_def check_types_def) - apply (blast intro: nth_in) - done +proof - + let ?wf_md = "(\P C (M, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). + wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M))" + have wfmd: "wf_prog ?wf_md P" using wf + by (unfold wf_jvm_prog_phi_def) assumption + show ?thesis using sees_wf_mdecl[OF wfmd mC] \ pc + by (simp add: wf_mdecl_def wt_method_def check_types_def) + (blast intro: nth_in) +qed (*>*) text \ The main theorem: welltyped programs do not produce type errors if they are started in a conformant state. \ theorem no_type_error: fixes \ :: jvm_state assumes welltyped: "wf_jvm_prog\<^bsub>\\<^esub> P" and conforms: "P,\ \ \ \" shows "exec_d P \ \ TypeError" (*<*) proof - from welltyped obtain mb where wf: "wf_prog mb P" by (fast dest: wt_jvm_progD) obtain xcp h frs where s [simp]: "\ = (xcp, h, frs)" by (cases \) from conforms have "xcp \ None \ frs = [] \ check P \" by (unfold correct_state_def check_def) auto moreover { assume "\(xcp \ None \ frs = [])" then obtain stk reg C M pc frs' where xcp [simp]: "xcp = None" and frs [simp]: "frs = (stk,reg,C,M,pc)#frs'" by (clarsimp simp add: neq_Nil_conv) from conforms obtain ST LT Ts T mxs mxl ins xt where hconf: "P \ h \" and meth: "P \ C sees M:Ts\T = (mxs, mxl, ins, xt) in C" and \: "\ C M ! pc = Some (ST,LT)" and frame: "conf_f P h (ST,LT) ins (stk,reg,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs'" by (fastforce simp add: correct_state_def dest: sees_method_fun) from frame obtain stk: "P,h \ stk [:\] ST" and reg: "P,h \ reg [:\\<^sub>\] LT" and pc: "pc < size ins" by (simp add: conf_f_def) from welltyped meth \ pc have "OK (Some (ST, LT)) \ states P mxs (1+size Ts+mxl)" by (rule wt_jvm_prog_states) hence "size ST \ mxs" by (auto simp add: JVM_states_unfold) with stk have mxs: "size stk \ mxs" by (auto dest: list_all2_lengthD) from welltyped meth pc have "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" by (rule wt_jvm_prog_impl_wt_instr) hence app\<^sub>0: "app (ins!pc) P mxs T pc (size ins) xt (\ C M!pc) " by (simp add: wt_instr_def) with \ have eff: "\(pc',s')\set (eff (ins ! pc) P pc xt (\ C M ! pc)). pc' < size ins" by (unfold app_def) simp from app\<^sub>0 \ have app: "xcpt_app (ins!pc) P pc mxs xt (ST,LT) \ app\<^sub>i (ins!pc, P, pc, mxs, T, (ST,LT))" by (clarsimp simp add: app_def) with eff stk reg have "check_instr (ins!pc) P h stk reg C M pc frs'" proof (cases "ins!pc") case (Getfield F C) with app stk reg \ obtain v vT stk' where field: "P \ C sees F:vT in C" and stk: "stk = v # stk'" and conf: "P,h \ v :\ Class C" by auto from conf have is_Ref: "is_Ref v" by auto moreover { assume "v \ Null" with conf field is_Ref wf have "\D vs. h (the_Addr v) = Some (D,vs) \ P \ D \\<^sup>* C" by (auto dest!: non_npD) } - ultimately show ?thesis using Getfield field stk hconf - apply clarsimp - apply (rule conjI, fastforce) - apply clarsimp - apply (drule has_visible_field) - apply (drule (1) has_field_mono) - apply (drule (1) hconfD) - apply (unfold oconf_def has_field_def) - apply clarsimp - apply (fastforce dest: has_fields_fun) - done + ultimately show ?thesis using Getfield field stk + has_field_mono[OF has_visible_field[OF field]] hconfD[OF hconf] + by (unfold oconf_def has_field_def) (fastforce dest: has_fields_fun) next case (Putfield F C) with app stk reg \ obtain v ref vT stk' where field: "P \ C sees F:vT in C" and stk: "stk = v # ref # stk'" and confv: "P,h \ v :\ vT" and confr: "P,h \ ref :\ Class C" by fastforce from confr have is_Ref: "is_Ref ref" by simp moreover { assume "ref \ Null" with confr field is_Ref wf have "\D vs. h (the_Addr ref) = Some (D,vs) \ P \ D \\<^sup>* C" by (auto dest: non_npD) } ultimately show ?thesis using Putfield field stk confv by fastforce next case (Invoke M' n) with app have n: "n < size ST" by simp from stk have [simp]: "size stk = size ST" by (rule list_all2_lengthD) { assume "stk!n = Null" with n Invoke have ?thesis by simp } moreover { assume "ST!n = NT" with n stk have "stk!n = Null" by (auto simp: list_all2_conv_all_nth) with n Invoke have ?thesis by simp } moreover { assume Null: "stk!n \ Null" and NT: "ST!n \ NT" from NT app Invoke obtain D D' Ts T m where D: "ST!n = Class D" and M': "P \ D sees M': Ts\T = m in D'" and Ts: "P \ rev (take n ST) [\] Ts" by auto from D stk n have "P,h \ stk!n :\ Class D" by (auto simp: list_all2_conv_all_nth) with Null obtain a C' fs where [simp]: "stk!n = Addr a" "h a = Some (C',fs)" and "P \ C' \\<^sup>* D" by (fastforce dest!: conf_ClassD) with M' wf obtain m' Ts' T' D'' where C': "P \ C' sees M': Ts'\T' = m' in D''" and Ts': "P \ Ts [\] Ts'" by (auto dest!: sees_method_mono) from stk have "P,h \ take n stk [:\] take n ST" .. hence "P,h \ rev (take n stk) [:\] rev (take n ST)" .. also note Ts also note Ts' finally have "P,h \ rev (take n stk) [:\] Ts'" . with Invoke Null n C' have ?thesis by (auto simp add: is_Ref_def2 has_methodI) } ultimately show ?thesis by blast next case Return with stk app \ meth frames show ?thesis by (auto simp add: has_methodI) qed (auto simp add: list_all2_lengthD) hence "check P \" using meth pc mxs by (simp add: check_def has_methodI) } ultimately have "check P \" by blast thus "exec_d P \ \ TypeError" .. qed (*>*) text \ The theorem above tells us that, in welltyped programs, the defensive machine reaches the same result as the aggressive one (after arbitrarily many steps). \ theorem welltyped_aggressive_imp_defensive: "wf_jvm_prog\<^bsub>\\<^esub> P \ P,\ \ \ \ \ P \ \ -jvm\ \' \ P \ (Normal \) -jvmd\ (Normal \')" (*<*) - apply (simp only: exec_all_def) - apply (erule rtrancl_induct) - apply (simp add: exec_all_d_def1) - apply simp - apply (simp only: exec_all_def [symmetric]) - apply (frule BV_correct, assumption+) - apply (drule no_type_error, assumption, drule no_type_error_commutes, simp) - apply (simp add: exec_all_d_def1) - apply (rule rtrancl_trans, assumption) - apply (drule exec_1_d_NormalI) - apply auto - done +proof - + assume wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and cf: "P,\ \ \ \" and exec: "P \ \ -jvm\ \'" + then have "(\, \') \ {(\, \'). exec (P, \) = \\'\}\<^sup>*" by(simp only: exec_all_def) + then show ?thesis proof(induct rule: rtrancl_induct) + case base + then show ?case by (simp add: exec_all_d_def1) + next + case (step y z) + then have \y: "P \ \ -jvm\ y" by (simp only: exec_all_def [symmetric]) + have exec_d: "exec_d P y = Normal \z\" using step + no_type_error_commutes[OF no_type_error[OF wf BV_correct[OF wf \y cf]]] + by (simp add: exec_all_d_def1) + show ?case using step.hyps(3) exec_1_d_NormalI[OF exec_d] + by (simp add: exec_all_d_def1) + qed +qed (*>*) text \ As corollary we get that the aggressive and the defensive machine are equivalent for welltyped programs (if started in a conformant state or in the canonical start state) \ corollary welltyped_commutes: fixes \ :: jvm_state assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and conforms: "P,\ \ \ \" shows "P \ (Normal \) -jvmd\ (Normal \') = P \ \ -jvm\ \'" - apply rule - apply (erule defensive_imp_aggressive) - apply (erule welltyped_aggressive_imp_defensive [OF wf conforms]) - done +proof(rule iffI) + assume "P \ Normal \ -jvmd\ Normal \'" then show "P \ \ -jvm\ \'" + by (rule defensive_imp_aggressive) +next + assume "P \ \ -jvm\ \'" then show "P \ Normal \ -jvmd\ Normal \'" + by (rule welltyped_aggressive_imp_defensive [OF wf conforms]) +qed corollary welltyped_initial_commutes: assumes wf: "wf_jvm_prog P" assumes meth: "P \ C sees M:[]\T = b in C" defines start: "\ \ start_state P C M" shows "P \ (Normal \) -jvmd\ (Normal \') = P \ \ -jvm\ \'" proof - from wf obtain \ where wf': "wf_jvm_prog\<^bsub>\\<^esub> P" by (auto simp: wf_jvm_prog_def) from this meth have "P,\ \ \ \" unfolding start by (rule BV_correct_initial) with wf' show ?thesis by (rule welltyped_commutes) qed lemma not_TypeError_eq [iff]: "x \ TypeError = (\t. x = Normal t)" by (cases x) auto locale cnf = fixes P and \ and \ assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes cnf: "correct_state P \ \" theorem (in cnf) no_type_errors: "P \ (Normal \) -jvmd\ \' \ \' \ TypeError" - apply (unfold exec_all_d_def1) - apply (erule rtrancl_induct) - apply simp - apply (fold exec_all_d_def1) - apply (insert cnf wf) - apply clarsimp - apply (drule defensive_imp_aggressive) - apply (frule (2) BV_correct) - apply (drule (1) no_type_error) back - apply (auto simp add: exec_1_d_eq) - done +proof - + assume "P \ (Normal \) -jvmd\ \'" + then have "(Normal \, \') \ (exec_1_d P)\<^sup>*" by (unfold exec_all_d_def1) simp + then show ?thesis proof(induct rule: rtrancl_induct) + case (step y z) + then obtain y\<^sub>n where [simp]: "y = Normal y\<^sub>n" by clarsimp + have n\y: "P \ Normal \ -jvmd\ Normal y\<^sub>n" using step.hyps(1) + by (fold exec_all_d_def1) simp + have \y: "P \ \ -jvm\ y\<^sub>n" using defensive_imp_aggressive[OF n\y] by simp + show ?case using step no_type_error[OF wf BV_correct[OF wf \y cnf]] + by (auto simp add: exec_1_d_eq) + qed simp +qed locale start = fixes P and C and M and \ and T and b assumes wf: "wf_jvm_prog P" assumes sees: "P \ C sees M:[]\T = b in C" defines "\ \ Normal (start_state P C M)" corollary (in start) bv_no_type_error: shows "P \ \ -jvmd\ \' \ \' \ TypeError" proof - from wf obtain \ where "wf_jvm_prog\<^bsub>\\<^esub> P" by (auto simp: wf_jvm_prog_def) moreover with sees have "correct_state P \ (start_state P C M)" by - (rule BV_correct_initial) ultimately have "cnf P \ (start_state P C M)" by (rule cnf.intro) moreover assume "P \ \ -jvmd\ \'" ultimately show ?thesis by (unfold \_def) (rule cnf.no_type_errors) qed end diff --git a/thys/Jinja/BV/BVSpec.thy b/thys/Jinja/BV/BVSpec.thy --- a/thys/Jinja/BV/BVSpec.thy +++ b/thys/Jinja/BV/BVSpec.thy @@ -1,93 +1,102 @@ (* Title: HOL/MicroJava/BV/BVSpec.thy Author: Cornelia Pusch, Gerwin Klein Copyright 1999 Technische Universitaet Muenchen *) section \The Bytecode Verifier \label{sec:BVSpec}\ theory BVSpec imports Effect begin text \ This theory contains a specification of the BV. The specification describes correct typings of method bodies; it corresponds to type \emph{checking}. \ definition \ \The method type only contains declared classes:\ check_types :: "'m prog \ nat \ nat \ ty\<^sub>i' err list \ bool" where "check_types P mxs mxl \s \ set \s \ states P mxs mxl" \ \An instruction is welltyped if it is applicable and its effect\ \ \is compatible with the type at all successor instructions:\ definition wt_instr :: "['m prog,ty,nat,pc,ex_table,instr,pc,ty\<^sub>m] \ bool" ("_,_,_,_,_ \ _,_ :: _" [60,0,0,0,0,0,0,61] 60) where "P,T,mxs,mpc,xt \ i,pc :: \s \ app i P mxs T pc mpc xt (\s!pc) \ (\(pc',\') \ set (eff i P pc xt (\s!pc)). P \ \' \' \s!pc')" \ \The type at @{text "pc=0"} conforms to the method calling convention:\ definition wt_start :: "['m prog,cname,ty list,nat,ty\<^sub>m] \ bool" where "wt_start P C Ts mxl\<^sub>0 \s \ P \ Some ([],OK (Class C)#map OK Ts@replicate mxl\<^sub>0 Err) \' \s!0" \ \A method is welltyped if the body is not empty,\ \ \if the method type covers all instructions and mentions\ \ \declared classes only, if the method calling convention is respected, and\ \ \if all instructions are welltyped.\ definition wt_method :: "['m prog,cname,ty list,ty,nat,nat,instr list, ex_table,ty\<^sub>m] \ bool" where "wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s \ 0 < size is \ size \s = size is \ check_types P mxs (1+size Ts+mxl\<^sub>0) (map OK \s) \ wt_start P C Ts mxl\<^sub>0 \s \ (\pc < size is. P,T\<^sub>r,mxs,size is,xt \ is!pc,pc :: \s)" \ \A program is welltyped if it is wellformed and all methods are welltyped\ definition wf_jvm_prog_phi :: "ty\<^sub>P \ jvm_prog \ bool" ("wf'_jvm'_prog\<^bsub>_\<^esub>") where "wf_jvm_prog\<^bsub>\\<^esub> \ wf_prog (\P C (M,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M))" definition wf_jvm_prog :: "jvm_prog \ bool" where "wf_jvm_prog P \ \\. wf_jvm_prog\<^bsub>\\<^esub> P" lemma wt_jvm_progD: "wf_jvm_prog\<^bsub>\\<^esub> P \ \wt. wf_prog wt P" (*<*) by (unfold wf_jvm_prog_phi_def, blast) (*>*) lemma wt_jvm_prog_impl_wt_instr: - "\ wf_jvm_prog\<^bsub>\\<^esub> P; - P \ C sees M:Ts \ T = (mxs,mxl\<^sub>0,ins,xt) in C; pc < size ins \ - \ P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" +assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and + sees: "P \ C sees M:Ts \ T = (mxs,mxl\<^sub>0,ins,xt) in C" and + pc: "pc < size ins" +shows "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" (*<*) - apply (unfold wf_jvm_prog_phi_def) - apply (drule (1) sees_wf_mdecl) - apply (simp add: wf_mdecl_def wt_method_def) - done +proof - + have wfm: "wf_prog + (\P C (M, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). + wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M)) P" using wf + by (unfold wf_jvm_prog_phi_def) + show ?thesis using sees_wf_mdecl[OF wfm sees] pc + by (simp add: wf_mdecl_def wt_method_def) +qed (*>*) lemma wt_jvm_prog_impl_wt_start: - "\ wf_jvm_prog\<^bsub>\\<^esub> P; - P \ C sees M:Ts \ T = (mxs,mxl\<^sub>0,ins,xt) in C \ \ - 0 < size ins \ wt_start P C Ts mxl\<^sub>0 (\ C M)" +assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and + sees: "P \ C sees M:Ts \ T = (mxs,mxl\<^sub>0,ins,xt) in C" +shows "0 < size ins \ wt_start P C Ts mxl\<^sub>0 (\ C M)" (*<*) - apply (unfold wf_jvm_prog_phi_def) - apply (drule (1) sees_wf_mdecl) - apply (simp add: wf_mdecl_def wt_method_def) - done +proof - + have wfm: "wf_prog + (\P C (M, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). + wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M)) P" using wf + by (unfold wf_jvm_prog_phi_def) + show ?thesis using sees_wf_mdecl[OF wfm sees] + by (simp add: wf_mdecl_def wt_method_def) +qed (*>*) end diff --git a/thys/Jinja/BV/BVSpecTypeSafe.thy b/thys/Jinja/BV/BVSpecTypeSafe.thy --- a/thys/Jinja/BV/BVSpecTypeSafe.thy +++ b/thys/Jinja/BV/BVSpecTypeSafe.thy @@ -1,1098 +1,1102 @@ (* Title: HOL/MicroJava/BV/BVSpecTypeSafe.thy Author: Cornelia Pusch, Gerwin Klein Copyright 1999 Technische Universitaet Muenchen *) section \BV Type Safety Proof \label{sec:BVSpecTypeSafe}\ theory BVSpecTypeSafe imports BVConform begin text \ This theory contains proof that the specification of the bytecode verifier only admits type safe programs. \ subsection \Preliminaries\ text \ Simp and intro setup for the type safety proof: \ lemmas defs1 = correct_state_def conf_f_def wt_instr_def eff_def norm_eff_def app_def xcpt_app_def lemmas widen_rules [intro] = conf_widen confT_widen confs_widens confTs_widen subsection \Exception Handling\ text \ For the \Invoke\ instruction the BV has checked all handlers that guard the current \pc\. \ lemma Invoke_handlers: "match_ex_table P C pc xt = Some (pc',d') \ \(f,t,D,h,d) \ set (relevant_entries P (Invoke n M) pc xt). P \ C \\<^sup>* D \ pc \ {f.. pc' = h \ d' = d" by (induct xt) (auto simp add: relevant_entries_def matches_ex_entry_def is_relevant_entry_def split: if_split_asm) text \ We can prove separately that the recursive search for exception handlers (\find_handler\) in the frame stack results in a conforming state (if there was no matching exception handler in the current frame). We require that the exception is a valid heap address, and that the state before the exception occurred conforms. \ term find_handler lemma uncaught_xcpt_correct: assumes wt: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes h: "h xcp = Some obj" shows "\f. P,\ \ (None, h, f#frs)\ \ P,\ \ (find_handler P xcp h frs) \" (is "\f. ?correct (None, h, f#frs) \ ?correct (?find frs)") (*<*) proof (induct frs) \ \the base case is trivial as it should be\ show "?correct (?find [])" by (simp add: correct_state_def) next \ \we will need both forms @{text wf_jvm_prog} and @{text wf_prog} later\ from wt obtain mb where wf: "wf_prog mb P" by (simp add: wf_jvm_prog_phi_def) \ \the assumption for the cons case:\ fix f f' frs' assume cr: "?correct (None, h, f#f'#frs')" \ \the induction hypothesis:\ assume IH: "\f. ?correct (None, h, f#frs') \ ?correct (?find frs')" from cr have cr': "?correct (None, h, f'#frs')" by (fastforce simp add: correct_state_def) obtain stk loc C M pc where [simp]: "f' = (stk,loc,C,M,pc)" by (cases f') from cr obtain Ts T mxs mxl\<^sub>0 ins xt where meth: "P \ C sees M:Ts \ T = (mxs,mxl\<^sub>0,ins,xt) in C" by (simp add: correct_state_def, blast) hence [simp]: "ex_table_of P C M = xt" by simp show "?correct (?find (f'#frs'))" proof (cases "match_ex_table P (cname_of h xcp) pc xt") case None with cr' IH [of f'] show ?thesis by fastforce next fix pc_d assume "match_ex_table P (cname_of h xcp) pc xt = Some pc_d" then obtain pc' d' where match: "match_ex_table P (cname_of h xcp) pc xt = Some (pc',d')" (is "?match (cname_of h xcp) = _") by (cases pc_d) auto from wt meth cr' [simplified] have wti: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" by (fastforce simp add: correct_state_def conf_f_def dest: sees_method_fun elim!: wt_jvm_prog_impl_wt_instr) from cr meth obtain n M' ST LT where ins: "ins!pc = Invoke n M'" (is "_ = ?i") and \: "\ C M ! pc = Some (ST, LT)" by (fastforce dest: sees_method_fun simp add: correct_state_def) from ins match obtain f t D where rel: "(f,t,D,pc',d') \ set (relevant_entries P (ins!pc) pc xt)" and D: "P \ cname_of h xcp \\<^sup>* D" by (fastforce dest: Invoke_handlers) from rel have "(pc', Some (Class D # drop (size ST - d') ST, LT)) \ set (xcpt_eff (ins!pc) P pc (ST,LT) xt)" (is "(_, Some (?ST',_)) \ _") by (force simp add: xcpt_eff_def image_def) with wti \ obtain pc: "pc' < size ins" and "P \ Some (?ST', LT) \' \ C M ! pc'" by (clarsimp simp add: defs1) blast then obtain ST' LT' where \': "\ C M ! pc' = Some (ST',LT')" and less: "P \ (?ST', LT) \\<^sub>i (ST',LT')" by (auto simp add: sup_state_opt_any_Some) from cr' \ meth have "conf_f P h (ST, LT) ins f'" by (unfold correct_state_def) (fastforce dest: sees_method_fun) hence loc: "P,h \ loc [:\\<^sub>\] LT" and stk: "P,h \ stk [:\] ST" by (unfold conf_f_def) auto hence [simp]: "size stk = size ST" by (simp add: list_all2_lengthD) let ?f = "(Addr xcp # drop (length stk - d') stk, loc, C, M, pc')" have "conf_f P h (ST',LT') ins ?f" proof - from wf less loc have "P,h \ loc [:\\<^sub>\] LT'" by simp blast moreover from D h have "P,h \ Addr xcp :\ Class D" by (simp add: conf_def obj_ty_def case_prod_unfold) with less stk have "P,h \ Addr xcp # drop (length stk - d') stk [:\] ST'" by (auto intro!: list_all2_dropI) ultimately show ?thesis using pc by (simp add: conf_f_def) qed with cr' match \' meth pc show ?thesis by (unfold correct_state_def) (fastforce dest: sees_method_fun) qed qed (*>*) text \ The requirement of lemma \uncaught_xcpt_correct\ (that the exception is a valid reference on the heap) is always met for welltyped instructions and conformant states: \ lemma exec_instr_xcpt_h: "\ fst (exec_instr (ins!pc) P h stk vars Cl M pc frs) = Some xcp; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ \obj. h xcp = Some obj" (is "\ ?xcpt; ?wt; ?correct \ \ ?thesis") (*<*) proof - note [simp] = split_beta note [split] = if_split_asm option.split_asm assume wt: ?wt ?correct hence pre: "preallocated h" by (simp add: correct_state_def hconf_def) assume xcpt: ?xcpt with pre show ?thesis proof (cases "ins!pc") case Throw with xcpt wt pre show ?thesis by (clarsimp iff: list_all2_Cons2 simp add: defs1) (auto dest: non_npD simp: is_refT_def elim: preallocatedE) qed (auto elim: preallocatedE) qed (*>*) lemma conf_sys_xcpt: "\preallocated h; C \ sys_xcpts\ \ P,h \ Addr (addr_of_sys_xcpt C) :\ Class C" by (auto elim: preallocatedE) lemma match_ex_table_SomeD: "match_ex_table P C pc xt = Some (pc',d') \ \(f,t,D,h,d) \ set xt. matches_ex_entry P C pc (f,t,D,h,d) \ h = pc' \ d=d'" by (induct xt) (auto split: if_split_asm) text \ Finally we can state that, whenever an exception occurs, the next state always conforms: \ lemma xcpt_correct: fixes \' :: jvm_state assumes wtp: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes meth: "P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes xp: "fst (exec_instr (ins!pc) P h stk loc C M pc frs) = Some xcp" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" assumes correct: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" shows "P,\ \ \'\" (*<*) proof - from wtp obtain wfmb where wf: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) note conf_sys_xcpt [elim!] note xp' = meth s' xp note wtp moreover from xp wt correct obtain obj where h: "h xcp = Some obj" by (blast dest: exec_instr_xcpt_h) moreover note correct ultimately have "P,\ \ find_handler P xcp h frs \" by (rule uncaught_xcpt_correct) with xp' have "match_ex_table P (cname_of h xcp) pc xt = None \ ?thesis" (is "?m (cname_of h xcp) = _ \ _" is "?match = _ \ _") by (simp add: split_beta) moreover { fix pc_d assume "?match = Some pc_d" then obtain pc' d' where some_handler: "?match = Some (pc',d')" by (cases pc_d) auto from correct meth obtain ST LT where h_ok: "P \ h \" and \_pc: "\ C M ! pc = Some (ST, LT)" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" by (unfold correct_state_def) (fastforce dest: sees_method_fun) from h_ok have preh: "preallocated h" by (simp add: hconf_def) from frame obtain stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" by (unfold conf_f_def) auto from stk have [simp]: "size stk = size ST" .. from wt \_pc have eff: "\(pc', s')\set (xcpt_eff (ins!pc) P pc (ST,LT) xt). pc' < size ins \ P \ s' \' \ C M!pc'" by (auto simp add: defs1) let ?stk' = "Addr xcp # drop (length stk - d') stk" let ?f = "(?stk', loc, C, M, pc')" from some_handler xp' have \': "\' = (None, h, ?f#frs)" by (simp add: split_beta) from some_handler obtain f t D where xt: "(f,t,D,pc',d') \ set xt" and "matches_ex_entry P (cname_of h xcp) pc (f,t,D,pc',d')" by (auto dest: match_ex_table_SomeD) hence match: "P \ cname_of h xcp \\<^sup>* D" "pc \ {f.. set (relevant_entries P (ins!pc) pc xt)" and conf: "P,h \ Addr xcp :\ Class D" proof (cases "ins!pc") case Return with xp have False by (auto simp: split_beta split: if_split_asm) thus ?thesis .. next case New with xp have [simp]: "xcp = addr_of_sys_xcpt OutOfMemory" by simp with New match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) with match preh xt show ?thesis by (fastforce simp add: relevant_entries_def intro: that) next case Getfield with xp have [simp]: "xcp = addr_of_sys_xcpt NullPointer" by (simp add: split_beta split: if_split_asm) with Getfield match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) with match preh xt show ?thesis by (fastforce simp add: relevant_entries_def intro: that) next case Putfield with xp have [simp]: "xcp = addr_of_sys_xcpt NullPointer" by (simp add: split_beta split: if_split_asm) with Putfield match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) with match preh xt show ?thesis by (fastforce simp add: relevant_entries_def intro: that) next case Checkcast with xp have [simp]: "xcp = addr_of_sys_xcpt ClassCast" by (simp add: split_beta split: if_split_asm) with Checkcast match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) with match preh xt show ?thesis by (fastforce simp add: relevant_entries_def intro: that) next case Invoke with xp match have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) moreover from xp wt correct obtain obj where xcp: "h xcp = Some obj" by (blast dest: exec_instr_xcpt_h) ultimately show ?thesis using xt match by (auto simp add: relevant_entries_def conf_def case_prod_unfold intro: that) next case Throw with xp match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) moreover from xp wt correct obtain obj where xcp: "h xcp = Some obj" by (blast dest: exec_instr_xcpt_h) ultimately show ?thesis using xt match by (auto simp add: relevant_entries_def conf_def case_prod_unfold intro: that) qed auto with eff obtain ST' LT' where \_pc': "\ C M ! pc' = Some (ST', LT')" and pc': "pc' < size ins" and less: "P \ (Class D # drop (size ST - d') ST, LT) \\<^sub>i (ST', LT')" by (fastforce simp add: xcpt_eff_def sup_state_opt_any_Some) with conf loc stk have "conf_f P h (ST',LT') ins ?f" by (auto simp add: defs1 intro: list_all2_dropI) with meth h_ok frames \_pc' \' have ?thesis by (unfold correct_state_def) (fastforce dest: sees_method_fun) } ultimately show ?thesis by (cases "?match") blast+ qed (*>*) subsection \Single Instructions\ text \ In this section we prove for each single (welltyped) instruction that the state after execution of the instruction still conforms. Since we have already handled exceptions above, we can now assume that no exception occurs in this step. \ declare defs1 [simp] lemma Invoke_correct: fixes \' :: jvm_state assumes wtprog: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes meth_C: "P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes ins: "ins ! pc = Invoke M' n" assumes wti: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes \': "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" assumes approx: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" assumes no_xcp: "fst (exec_instr (ins!pc) P h stk loc C M pc frs) = None" shows "P,\ \ \'\" (*<*) proof - note split_paired_Ex [simp del] from wtprog obtain wfmb where wfprog: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) from ins meth_C approx obtain ST LT where heap_ok: "P\ h\" and \_pc: "\ C M!pc = Some (ST,LT)" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" by (fastforce dest: sees_method_fun) from ins wti \_pc have n: "n < size ST" by simp { assume "stk!n = Null" with ins no_xcp have False by (simp add: split_beta) hence ?thesis .. } moreover { assume "ST!n = NT" moreover from frame have "P,h \ stk [:\] ST" by simp with n have "P,h \ stk!n :\ ST!n" by (simp add: list_all2_conv_all_nth) ultimately have "stk!n = Null" by simp with ins no_xcp have False by (simp add: split_beta) hence ?thesis .. } moreover { assume NT: "ST!n \ NT" and Null: "stk!n \ Null" from NT ins wti \_pc obtain D D' Ts T m ST' LT' where D: "ST!n = Class D" and pc': "pc+1 < size ins" and m_D: "P \ D sees M': Ts\T = m in D'" and Ts: "P \ rev (take n ST) [\] Ts" and \': "\ C M ! (pc+1) = Some (ST', LT')" and LT': "P \ LT [\\<^sub>\] LT'" and ST': "P \ (T # drop (n+1) ST) [\] ST'" by (clarsimp simp add: sup_state_opt_any_Some) from frame obtain stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" by simp from n stk D have "P,h \ stk!n :\ Class D" by (auto simp add: list_all2_conv_all_nth) with Null obtain a C' fs where Addr: "stk!n = Addr a" and obj: "h a = Some (C',fs)" and C'subD: "P \ C' \\<^sup>* D" by (fastforce dest!: conf_ClassD) with wfprog m_D obtain Ts' T' m' D'' mxs' mxl' ins' xt' where m_C': "P \ C' sees M': Ts'\T' = (mxs',mxl',ins',xt') in D''" and T': "P \ T' \ T" and Ts': "P \ Ts [\] Ts'" by (auto dest: sees_method_mono) let ?loc' = "Addr a # rev (take n stk) @ replicate mxl' undefined" let ?f' = "([], ?loc', D'', M', 0)" let ?f = "(stk, loc, C, M, pc)" from Addr obj m_C' ins \' meth_C have s': "\' = (None, h, ?f' # ?f # frs)" by simp from Ts n have [simp]: "size Ts = n" by (auto dest: list_all2_lengthD simp: min_def) with Ts' have [simp]: "size Ts' = n" by (auto dest: list_all2_lengthD) from m_C' wfprog obtain mD'': "P \ D'' sees M':Ts'\T'=(mxs',mxl',ins',xt') in D''" by (fast dest: sees_method_idemp) moreover with wtprog obtain start: "wt_start P D'' Ts' mxl' (\ D'' M')" and ins': "ins' \ []" by (auto dest: wt_jvm_prog_impl_wt_start) then obtain LT\<^sub>0 where LT\<^sub>0: "\ D'' M' ! 0 = Some ([], LT\<^sub>0)" by (clarsimp simp add: wt_start_def defs1 sup_state_opt_any_Some) moreover have "conf_f P h ([], LT\<^sub>0) ins' ?f'" proof - let ?LT = "OK (Class D'') # (map OK Ts') @ (replicate mxl' Err)" from stk have "P,h \ take n stk [:\] take n ST" .. hence "P,h \ rev (take n stk) [:\] rev (take n ST)" by simp also note Ts also note Ts' finally have "P,h \ rev (take n stk) [:\\<^sub>\] map OK Ts'" by simp also have "P,h \ replicate mxl' undefined [:\\<^sub>\] replicate mxl' Err" by simp also from m_C' have "P \ C' \\<^sup>* D''" by (rule sees_method_decl_above) with obj have "P,h \ Addr a :\ Class D''" by (simp add: conf_def) ultimately have "P,h \ ?loc' [:\\<^sub>\] ?LT" by simp also from start LT\<^sub>0 have "P \ \ [\\<^sub>\] LT\<^sub>0" by (simp add: wt_start_def) finally have "P,h \ ?loc' [:\\<^sub>\] LT\<^sub>0" . thus ?thesis using ins' by simp qed ultimately have ?thesis using s' \_pc approx meth_C m_D T' ins D by (fastforce dest: sees_method_fun [of _ C]) } ultimately show ?thesis by blast qed (*>*) declare list_all2_Cons2 [iff] lemma Return_correct: fixes \' :: jvm_state assumes wt_prog: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes meth: "P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes ins: "ins ! pc = Return" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" assumes correct: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" shows "P,\ \ \'\" (*<*) proof - from wt_prog obtain wfmb where wf: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) from meth ins s' have "frs = [] \ ?thesis" by (simp add: correct_state_def) moreover { fix f frs' assume frs': "frs = f#frs'" moreover obtain stk' loc' C' M' pc' where f: "f = (stk',loc',C',M',pc')" by (cases f) moreover note meth ins s' ultimately have \': "\' = (None,h,(hd stk#(drop (1+size Ts) stk'),loc',C',M',pc'+1)#frs')" (is "\' = (None,h,?f'#frs')") by simp from correct meth obtain ST LT where h_ok: "P \ h \" and \_pc: "\ C M ! pc = Some (ST, LT)" and frame: "conf_f P h (ST, LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" by (auto dest: sees_method_fun simp add: correct_state_def) from \_pc ins wt obtain U ST\<^sub>0 where "ST = U # ST\<^sub>0" "P \ U \ T" by (simp add: wt_instr_def app_def) blast with wf frame have hd_stk: "P,h \ hd stk :\ T" by (auto simp add: conf_f_def) from f frs' frames obtain ST' LT' Ts'' T'' mxs' mxl\<^sub>0' ins' xt' D Ts' T' m D' where \': "\ C' M' ! pc' = Some (ST', LT')" and meth_C': "P \ C' sees M':Ts''\T''=(mxs',mxl\<^sub>0',ins',xt') in C'" and ins': "ins' ! pc' = Invoke M (size Ts)" and D: "ST' ! (size Ts) = Class D" and meth_D: "P \ D sees M: Ts'\T' = m in D'" and T': "P \ T \ T'" and frame': "conf_f P h (ST',LT') ins' f" and conf_fs: "conf_fs P h \ M' (size Ts'') T'' frs'" by clarsimp from f frame' obtain stk': "P,h \ stk' [:\] ST'" and loc': "P,h \ loc' [:\\<^sub>\] LT'" and pc': "pc' < size ins'" by (simp add: conf_f_def) from wt_prog meth_C' pc' have "P,T'',mxs',size ins',xt' \ ins'!pc',pc' :: \ C' M'" by (rule wt_jvm_prog_impl_wt_instr) with ins' \' D meth_D obtain aTs ST'' LT'' where \_suc: "\ C' M' ! Suc pc' = Some (ST'', LT'')" and less: "P \ (T' # drop (size Ts+1) ST', LT') \\<^sub>i (ST'', LT'')" and suc_pc': "Suc pc' < size ins'" by (clarsimp simp add: sup_state_opt_any_Some) from hd_stk T' have hd_stk': "P,h \ hd stk :\ T'" .. have frame'': "conf_f P h (ST'',LT'') ins' ?f'" proof - from stk' have "P,h \ drop (1+size Ts) stk' [:\] drop (1+size Ts) ST'" .. moreover with hd_stk' less have "P,h \ hd stk # drop (1+size Ts) stk' [:\] ST''" by auto moreover from wf loc' less have "P,h \ loc' [:\\<^sub>\] LT''" by auto moreover note suc_pc' ultimately show ?thesis by (simp add: conf_f_def) qed with \' frs' f meth h_ok hd_stk \_suc frames meth_C' \' have ?thesis by (fastforce dest: sees_method_fun [of _ C']) } ultimately show ?thesis by (cases frs) blast+ qed (*>*) declare sup_state_opt_any_Some [iff] declare not_Err_eq [iff] lemma Load_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins!pc = Load idx; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs); P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" by (fastforce dest: sees_method_fun [of _ C] elim!: confTs_confT_sup) declare [[simproc del: list_to_set_comprehension]] lemma Store_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins!pc = Store idx; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs); P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" -(*<*) - apply clarsimp - apply (drule (1) sees_method_fun) - apply clarsimp - apply (blast intro!: list_all2_update_cong) - done -(*>*) +(*<*)by clarsimp (blast dest: sees_method_fun intro!: list_all2_update_cong)(*>*) lemma Push_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins!pc = Push v; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs); P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" -(*<*) - apply clarsimp - apply (drule (1) sees_method_fun) - apply clarsimp - apply (blast dest: typeof_lit_conf) - done -(*>*) +(*<*)by clarsimp (blast dest: sees_method_fun typeof_lit_conf)(*>*) lemma Cast_conf2: "\ wf_prog ok P; P,h \ v :\ T; is_refT T; cast_ok P C h v; P \ Class C \ T'; is_class P C\ \ P,h \ v :\ T'" (*<*) - apply (unfold cast_ok_def is_refT_def) - apply (frule Class_widen) - apply (elim exE disjE) - apply simp - apply simp - apply simp - apply (clarsimp simp add: conf_def obj_ty_def) - apply (cases v) - apply (auto intro: rtrancl_trans) - done +proof - + assume "wf_prog ok P" and "P,h \ v :\ T" and "is_refT T" and + "cast_ok P C h v" and wid: "P \ Class C \ T'" and "is_class P C" + then show "P,h \ v :\ T'" using Class_widen[OF wid] + by(cases v) + (auto simp: cast_ok_def is_refT_def conf_def obj_ty_def + intro: rtrancl_trans) +qed (*>*) lemma Checkcast_correct: "\ wf_jvm_prog\<^bsub>\\<^esub> P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins!pc = Checkcast D; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs) ; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\; fst (exec_instr (ins!pc) P h stk loc C M pc frs) = None \ \ P,\ \ \'\" (*<*) - apply (clarsimp simp add: wf_jvm_prog_phi_def split: if_split_asm) - apply (drule (1) sees_method_fun) - apply (blast intro: Cast_conf2) - done +by (clarsimp simp: wf_jvm_prog_phi_def split: if_split_asm) + (blast intro: Cast_conf2 dest: sees_method_fun) (*>*) declare split_paired_All [simp del] lemmas widens_Cons [iff] = list_all2_Cons1 [of "widen P"] for P lemma Getfield_correct: fixes \' :: jvm_state assumes wf: "wf_prog wt P" assumes mC: "P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes i: "ins!pc = Getfield F D" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" assumes cf: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" assumes xc: "fst (exec_instr (ins!pc) P h stk loc C M pc frs) = None" shows "P,\ \ \'\" (*<*) proof - from mC cf obtain ST LT where "h\": "P \ h \" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and fs: "conf_fs P h \ M (size Ts) T frs" by (fastforce dest: sees_method_fun) from i \ wt obtain oT ST'' vT ST' LT' vT' where oT: "P \ oT \ Class D" and ST: "ST = oT # ST''" and F: "P \ D sees F:vT in D" and pc': "pc+1 < size ins" and \': "\ C M ! (pc+1) = Some (vT'#ST', LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" and vT': "P \ vT \ vT'" by fastforce from stk ST obtain ref stk' where stk': "stk = ref#stk'" and ref: "P,h \ ref :\ oT" and ST'': "P,h \ stk' [:\] ST''" by auto from stk' i mC s' xc have "ref \ Null" by (simp add: split_beta split:if_split_asm) moreover from ref oT have "P,h \ ref :\ Class D" .. ultimately obtain a D' fs where a: "ref = Addr a" and h: "h a = Some (D', fs)" and D': "P \ D' \\<^sup>* D" by (blast dest: non_npD) from D' F have has_field: "P \ D' has F:vT in D" by (blast intro: has_field_mono has_visible_field) moreover from "h\" h have "P,h \ (D', fs) \" by (rule hconfD) ultimately obtain v where v: "fs (F, D) = Some v" "P,h \ v :\ vT" by (clarsimp simp add: oconf_def has_field_def) (blast dest: has_fields_fun) from a h i mC s' stk' v have "\' = (None, h, (v#stk',loc,C,M,pc+1)#frs)" by simp moreover from ST'' ST' have "P,h \ stk' [:\] ST'" .. moreover from v vT' have "P,h \ v :\ vT'" by blast moreover from loc LT' have "P,h \ loc [:\\<^sub>\] LT'" .. moreover note "h\" mC \' pc' v fs ultimately show "P,\ \ \' \" by fastforce qed (*>*) lemma Putfield_correct: fixes \' :: jvm_state assumes wf: "wf_prog wt P" assumes mC: "P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes i: "ins!pc = Putfield F D" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" assumes cf: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" assumes xc: "fst (exec_instr (ins!pc) P h stk loc C M pc frs) = None" shows "P,\ \ \'\" (*<*) proof - from mC cf obtain ST LT where "h\": "P \ h \" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and fs: "conf_fs P h \ M (size Ts) T frs" by (fastforce dest: sees_method_fun) from i \ wt obtain vT vT' oT ST'' ST' LT' where ST: "ST = vT # oT # ST''" and field: "P \ D sees F:vT' in D" and oT: "P \ oT \ Class D" and vT: "P \ vT \ vT'" and pc': "pc+1 < size ins" and \': "\ C M!(pc+1) = Some (ST',LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" by clarsimp from stk ST obtain v ref stk' where stk': "stk = v#ref#stk'" and v: "P,h \ v :\ vT" and ref: "P,h \ ref :\ oT" and ST'': "P,h \ stk' [:\] ST''" by auto from stk' i mC s' xc have "ref \ Null" by (auto simp add: split_beta) moreover from ref oT have "P,h \ ref :\ Class D" .. ultimately obtain a D' fs where a: "ref = Addr a" and h: "h a = Some (D', fs)" and D': "P \ D' \\<^sup>* D" by (blast dest: non_npD) from v vT have vT': "P,h \ v :\ vT'" .. from field D' have has_field: "P \ D' has F:vT' in D" by (blast intro: has_field_mono has_visible_field) let ?h' = "h(a\(D', fs((F, D)\v)))" and ?f' = "(stk',loc,C,M,pc+1)" from h have hext: "h \ ?h'" by (rule hext_upd_obj) from a h i mC s' stk' have "\' = (None, ?h', ?f'#frs)" by simp moreover from "h\" h have "P,h \ (D',fs)\" by (rule hconfD) with has_field vT' have "P,h \ (D',fs((F, D)\v))\" .. with "h\" h have "P \ ?h'\" by (rule hconf_upd_obj) moreover from ST'' ST' have "P,h \ stk' [:\] ST'" .. from this hext have "P,?h' \ stk' [:\] ST'" by (rule confs_hext) moreover from loc LT' have "P,h \ loc [:\\<^sub>\] LT'" .. from this hext have "P,?h' \ loc [:\\<^sub>\] LT'" by (rule confTs_hext) moreover from fs hext have "conf_fs P ?h' \ M (size Ts) T frs" by (rule conf_fs_hext) moreover note mC \' pc' ultimately show "P,\ \ \' \" by fastforce qed (*>*) (* FIXME: move *) lemma has_fields_b_fields: "P \ C has_fields FDTs \ fields P C = FDTs" (*<*) - apply (unfold fields_def) - apply (blast intro: the_equality has_fields_fun) - done +by (unfold fields_def) + (blast intro: the_equality has_fields_fun) (*>*) (* FIXME: move *) lemma oconf_blank [intro, simp]: "\is_class P C; wf_prog wt P\ \ P,h \ blank P C \" (*<*) by (fastforce simp add: blank_def has_fields_b_fields oconf_init_fields dest: wf_Fields_Ex) (*>*) lemma obj_ty_blank [iff]: "obj_ty (blank P C) = Class C" by (simp add: blank_def) lemma New_correct: fixes \' :: jvm_state assumes wf: "wf_prog wt P" assumes meth: "P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes ins: "ins!pc = New X" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes exec: "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" assumes conf: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" assumes no_x: "fst (exec_instr (ins!pc) P h stk loc C M pc frs) = None" shows "P,\ \ \'\" (*<*) proof - from ins conf meth obtain ST LT where heap_ok: "P\ h\" and \_pc: "\ C M!pc = Some (ST,LT)" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" by (auto dest: sees_method_fun) from \_pc ins wt obtain ST' LT' where is_class_X: "is_class P X" and mxs: "size ST < mxs" and suc_pc: "pc+1 < size ins" and \_suc: "\ C M!(pc+1) = Some (ST', LT')" and less: "P \ (Class X # ST, LT) \\<^sub>i (ST', LT')" by auto from ins no_x obtain oref where new_Addr: "new_Addr h = Some oref" by auto hence h: "h oref = None" by (rule new_Addr_SomeD) with exec ins meth new_Addr have \': "\' = (None, h(oref \ blank P X), (Addr oref#stk,loc,C,M,pc+1)#frs)" (is "\' = (None, ?h', ?f # frs)") by simp moreover from wf h heap_ok is_class_X have h': "P \ ?h' \" by (auto intro: hconf_new) moreover from h frame less suc_pc wf have "conf_f P ?h' (ST', LT') ins ?f" - apply (clarsimp simp add: fun_upd_apply conf_def blank_def split_beta) - apply (auto intro: confs_hext confTs_hext) - done + by (clarsimp simp add: fun_upd_apply conf_def blank_def split_beta) + (auto intro: confs_hext confTs_hext) moreover from h have "h \ ?h'" by simp with frames have "conf_fs P ?h' \ M (size Ts) T frs" by (rule conf_fs_hext) ultimately show ?thesis using meth \_suc by fastforce qed (*>*) lemma Goto_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins ! pc = Goto branch; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs) ; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" -(*<*) -apply clarsimp -apply (drule (1) sees_method_fun) -apply fastforce -done -(*>*) +(*<*)by clarsimp (fastforce dest: sees_method_fun)(*>*) lemma IfFalse_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins ! pc = IfFalse branch; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs) ; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" -(*<*) -apply clarsimp -apply (drule (1) sees_method_fun) -apply fastforce -done -(*>*) +(*<*)by clarsimp (fastforce dest: sees_method_fun)(*>*) lemma CmpEq_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins ! pc = CmpEq; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs) ; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" -(*<*) -apply clarsimp -apply (drule (1) sees_method_fun) -apply fastforce -done -(*>*) +(*<*)by clarsimp (fastforce dest: sees_method_fun)(*>*) lemma Pop_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins ! pc = Pop; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs) ; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" -(*<*) -apply clarsimp -apply (drule (1) sees_method_fun) -apply fastforce -done -(*>*) +(*<*)by clarsimp (fastforce dest: sees_method_fun)(*>*) lemma IAdd_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins ! pc = IAdd; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs) ; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" -(*<*) -apply (clarsimp simp add: conf_def) -apply (drule (1) sees_method_fun) -apply fastforce -done -(*>*) +(*<*)by clarsimp (fastforce dest: sees_method_fun)(*>*) lemma Throw_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins ! pc = Throw; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs) ; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\; fst (exec_instr (ins!pc) P h stk loc C M pc frs) = None \ \ P,\ \ \'\" by simp text \ The next theorem collects the results of the sections above, i.e.~exception handling and the execution step for each instruction. It states type safety for single step execution: in welltyped programs, a conforming state is transformed into another conforming state when one instruction is executed. \ theorem instr_correct: -"\ wf_jvm_prog\<^bsub>\\<^esub> P; - P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; - Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs); - P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ -\ P,\ \ \'\" +fixes \' :: jvm_state +assumes wtp: "wf_jvm_prog\<^bsub>\\<^esub> P" + and meth: "P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" + and exec: "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" + and conf: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" +shows "P,\ \ \'\" (*<*) -apply (subgoal_tac "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M") - prefer 2 - apply (erule wt_jvm_prog_impl_wt_instr, assumption) - apply clarsimp - apply (drule (1) sees_method_fun) - apply simp -apply (cases "fst (exec_instr (ins!pc) P h stk loc C M pc frs)") - prefer 2 - apply (erule xcpt_correct, assumption+) -apply (frule wt_jvm_progD, erule exE) -apply (cases "ins!pc") -apply (rule Load_correct, assumption+) -apply (rule Store_correct, assumption+) -apply (rule Push_correct, assumption+) -apply (rule New_correct, assumption+) -apply (rule Getfield_correct, assumption+) -apply (rule Putfield_correct, assumption+) -apply (rule Checkcast_correct, assumption+) -apply (rule Invoke_correct, assumption+) -apply (rule Return_correct, assumption+) -apply (rule Pop_correct, assumption+) -apply (rule IAdd_correct, assumption+) -apply (rule Goto_correct, assumption+) -apply (rule CmpEq_correct, assumption+) -apply (rule IfFalse_correct, assumption+) -apply (rule Throw_correct, assumption+) -done +proof - + from assms have pc: "pc < length ins" by(auto dest: sees_method_fun) + with wt_jvm_prog_impl_wt_instr[OF wtp meth] have wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" + by simp + + from conf obtain ST LT where \: "\ C M ! pc = Some(ST,LT)" by clarsimp + + show ?thesis + proof(cases "fst (exec_instr (ins!pc) P h stk loc C M pc frs)") + case Some show ?thesis by(rule xcpt_correct[OF wtp meth wt Some exec conf]) + next + case None + from wt_jvm_progD[OF wtp] obtain wf_md where wf: "wf_prog wf_md P" by clarify + + from exec conf None obtain + exec': "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" + and conf': "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" + and None': "fst (exec_instr (ins!pc) P h stk loc C M pc frs) = None" by simp + + show ?thesis + proof(cases "ins!pc") + case Load from Load_correct[OF wf meth this wt exec conf] show ?thesis by simp + next + case Store from Store_correct[OF wf meth this wt exec conf] show ?thesis by simp + next + case Push from Push_correct[OF wf meth this wt exec conf] show ?thesis by simp + next + case (New C) from New_correct[OF wf meth this wt exec conf None'] show ?thesis by simp + next + case Getfield from Getfield_correct[OF wf meth this wt exec conf None] + show ?thesis by simp + next + case Putfield from Putfield_correct[OF wf meth this wt exec conf None] + show ?thesis by simp + next + case Checkcast from Checkcast_correct[OF wtp meth this wt exec conf None] + show ?thesis by simp + next + case Invoke with Invoke_correct[OF wtp meth this wt exec conf None] show ?thesis by simp + next + case Return with Return_correct[OF wtp meth this wt exec conf] show ?thesis by simp + next + case Pop with Pop_correct[OF wf meth this wt exec conf] show ?thesis by simp + next + case IAdd with IAdd_correct[OF wf meth this wt exec conf] show ?thesis by simp + next + case Goto with Goto_correct[OF wf meth this wt exec conf] show ?thesis by simp + next + case CmpEq with CmpEq_correct[OF wf meth this wt exec conf] show ?thesis by simp + next + case IfFalse with IfFalse_correct[OF wf meth this wt exec conf] show ?thesis by simp + next + case Throw with Throw_correct[OF wf meth this exec conf None] show ?thesis by simp + qed + qed +qed (*>*) subsection \Main\ lemma correct_state_impl_Some_method: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \m Ts T. P \ C sees M:Ts\T = m in C" by fastforce lemma BV_correct_1 [rule_format]: "\\. \ wf_jvm_prog\<^bsub>\\<^esub> P; P,\ \ \\\ \ P \ \ -jvm\\<^sub>1 \' \ P,\ \ \'\" (*<*) -apply (simp only: split_tupled_all exec_1_iff) -apply (rename_tac xp h frs) -apply (case_tac xp) - apply (case_tac frs) - apply simp - apply (simp only: split_tupled_all) - apply hypsubst - apply (frule correct_state_impl_Some_method) - apply clarify - apply (rule instr_correct) - apply assumption+ - apply (rule sym) - apply assumption+ -apply (case_tac frs) -apply simp_all -done +proof - + fix \ assume wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and cf: "P,\ \ \\" + obtain xp h frs where \[simp]: "\ = (xp, h, frs)" by(cases \) simp + have "exec (P, xp, h, frs) = \\'\ \ P,\ |- \' [ok]" + proof(cases xp) + case None + with wf cf show ?thesis + proof(cases frs) + case (Cons fr frs') + obtain stk loc C M pc where [simp]: "fr = (stk,loc,C,M,pc)" by(cases fr) simp + then have cf': "P,\ |- (None, h, (stk,loc,C,M,pc) # frs') [ok]" + using Cons None cf by simp + then obtain mxs mxl\<^sub>0 ins xt Ts T + where mC: "P \ C sees M : Ts\T = (mxs, mxl\<^sub>0, ins, xt) in C" + using correct_state_impl_Some_method[OF cf'] by clarify + show ?thesis using Cons None instr_correct[OF wf mC _ cf'] by simp + qed simp + next + case (Some a) + then show ?thesis using wf cf by (case_tac frs) simp_all + qed + then show "P \ \ -jvm\\<^sub>1 \' \ P,\ \ \'\" by(simp add: exec_1_iff) +qed (*>*) theorem progress: "\ xp=None; frs\[] \ \ \\'. P \ (xp,h,frs) -jvm\\<^sub>1 \'" by (clarsimp simp add: exec_1_iff neq_Nil_conv split_beta simp del: split_paired_Ex) lemma progress_conform: "\wf_jvm_prog\<^bsub>\\<^esub> P; P,\ \ (xp,h,frs)\; xp=None; frs\[]\ \ \\'. P \ (xp,h,frs) -jvm\\<^sub>1 \' \ P,\ \ \'\" -(*<*) -apply (drule progress) -apply assumption -apply (fast intro: BV_correct_1) -done -(*>*) +(*<*)by (drule (1) progress) (fast intro: BV_correct_1)(*>*) theorem BV_correct [rule_format]: "\ wf_jvm_prog\<^bsub>\\<^esub> P; P \ \ -jvm\ \' \ \ P,\ \ \\ \ P,\ \ \'\" (*<*) -apply (simp only: exec_all_def1) -apply (erule rtrancl_induct) - apply simp -apply clarify -apply (erule (2) BV_correct_1) -done +proof - + assume wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and "P \ \ -jvm\ \'" + then have "(\, \') \ (exec_1 P)\<^sup>*" by (simp only: exec_all_def1) + then show ?thesis proof(induct rule: rtrancl_induct) + case (step y z) + then show ?case by clarify (erule (1) BV_correct_1[OF wf]) + qed simp +qed (*>*) lemma hconf_start: assumes wf: "wf_prog wf_mb P" shows "P \ (start_heap P) \" (*<*) - apply (unfold hconf_def) - apply (simp add: preallocated_start) - apply (clarify) - apply (drule sym) - apply (unfold start_heap_def) - apply (insert wf) - apply (auto simp add: fun_upd_apply is_class_xcpt split: if_split_asm) - done +proof - + { fix a obj assume assm: "start_heap P a = \obj\" + have "P,start_heap P \ obj \" using wf assm[THEN sym] + by (unfold start_heap_def) + (auto simp: fun_upd_apply is_class_xcpt split: if_split_asm) + } + then show ?thesis using preallocated_start[of P] + by (unfold hconf_def) simp +qed (*>*) lemma BV_correct_initial: - shows "\ wf_jvm_prog\<^bsub>\\<^esub> P; P \ C sees M:[]\T = m in C \ - \ P,\ \ start_state P C M \" +assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" + and mC: "P \ C sees M:[]\T = m in C" +shows "P,\ \ start_state P C M \" (*<*) - apply (cases m) - apply (unfold start_state_def) - apply (unfold correct_state_def) - apply (simp del: defs1) - apply (rule conjI) - apply (simp add: wf_jvm_prog_phi_def hconf_start) - apply (drule wt_jvm_prog_impl_wt_start, assumption+) - apply (unfold conf_f_def wt_start_def) - apply fastforce - done +proof - + obtain mxs mxl\<^sub>0 ins xt where + mC': "P \ C sees M:[]\T = (mxs,mxl\<^sub>0,ins,xt) in C" + using mC by (cases m) simp + then have method: "(C,[],T,mxs,mxl\<^sub>0,ins,xt) = method P C M" by simp + let ?h = "start_heap P" and ?pc = 0 + let ?f = "([], Null # replicate mxl\<^sub>0 undefined, C, M, ?pc)" and ?fs = "[]" + let ?frs = "?f#?fs" + have "is_class P Object" using wf by(simp add: wf_jvm_prog_phi_def) + then obtain Mm where Mm: "P \ Object sees_methods Mm" + by(fastforce simp: is_class_def dest: sees_methods_Object) + have "P\ ?h\" using wf by (simp add: wf_jvm_prog_phi_def hconf_start) + moreover have "\b Ts T mxs mxl\<^sub>0 is xt \. + (P \ C sees M:Ts\T = (mxs,mxl\<^sub>0,is,xt) in C) + \ \ C M ! ?pc = Some \ + \ conf_f P ?h \ is ?f \ conf_fs P ?h \ M (size Ts) T ?fs" + using wt_jvm_prog_impl_wt_start[OF wf mC'] mC' + by (unfold conf_f_def wt_start_def) fastforce + ultimately show ?thesis using method + by (fastforce simp del: defs1 simp: start_state_def correct_state_def) +qed declare [[simproc add: list_to_set_comprehension]] (*>*) theorem typesafe: assumes welltyped: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes main_method: "P \ C sees M:[]\T = m in C" shows "P \ start_state P C M -jvm\ \ \ P,\ \ \ \" (*<*) proof - from welltyped main_method have "P,\ \ start_state P C M \" by (rule BV_correct_initial) moreover assume "P \ start_state P C M -jvm\ \" ultimately show "P,\ \ \ \" using welltyped by - (rule BV_correct) qed (*>*) end diff --git a/thys/Jinja/BV/Effect.thy b/thys/Jinja/BV/Effect.thy --- a/thys/Jinja/BV/Effect.thy +++ b/thys/Jinja/BV/Effect.thy @@ -1,393 +1,395 @@ (* Title: HOL/MicroJava/BV/Effect.thy Author: Gerwin Klein Copyright 2000 Technische Universitaet Muenchen *) section \Effect of Instructions on the State Type\ theory Effect imports JVM_SemiType "../JVM/JVMExceptions" begin \ \FIXME\ locale prog = fixes P :: "'a prog" locale jvm_method = prog + fixes mxs :: nat fixes mxl\<^sub>0 :: nat fixes Ts :: "ty list" fixes T\<^sub>r :: ty fixes "is" :: "instr list" fixes xt :: ex_table fixes mxl :: nat defines mxl_def: "mxl \ 1+size Ts+mxl\<^sub>0" text \Program counter of successor instructions:\ primrec succs :: "instr \ ty\<^sub>i \ pc \ pc list" where "succs (Load idx) \ pc = [pc+1]" | "succs (Store idx) \ pc = [pc+1]" | "succs (Push v) \ pc = [pc+1]" | "succs (Getfield F C) \ pc = [pc+1]" | "succs (Putfield F C) \ pc = [pc+1]" | "succs (New C) \ pc = [pc+1]" | "succs (Checkcast C) \ pc = [pc+1]" | "succs Pop \ pc = [pc+1]" | "succs IAdd \ pc = [pc+1]" | "succs CmpEq \ pc = [pc+1]" | succs_IfFalse: "succs (IfFalse b) \ pc = [pc+1, nat (int pc + b)]" | succs_Goto: "succs (Goto b) \ pc = [nat (int pc + b)]" | succs_Return: "succs Return \ pc = []" | succs_Invoke: "succs (Invoke M n) \ pc = (if (fst \)!n = NT then [] else [pc+1])" | succs_Throw: "succs Throw \ pc = []" text "Effect of instruction on the state type:" fun the_class:: "ty \ cname" where "the_class (Class C) = C" fun eff\<^sub>i :: "instr \ 'm prog \ ty\<^sub>i \ ty\<^sub>i" where eff\<^sub>i_Load: "eff\<^sub>i (Load n, P, (ST, LT)) = (ok_val (LT ! n) # ST, LT)" | eff\<^sub>i_Store: "eff\<^sub>i (Store n, P, (T#ST, LT)) = (ST, LT[n:= OK T])" | eff\<^sub>i_Push: "eff\<^sub>i (Push v, P, (ST, LT)) = (the (typeof v) # ST, LT)" | eff\<^sub>i_Getfield: "eff\<^sub>i (Getfield F C, P, (T#ST, LT)) = (snd (field P C F) # ST, LT)" | eff\<^sub>i_Putfield: "eff\<^sub>i (Putfield F C, P, (T\<^sub>1#T\<^sub>2#ST, LT)) = (ST,LT)" | eff\<^sub>i_New: "eff\<^sub>i (New C, P, (ST,LT)) = (Class C # ST, LT)" | eff\<^sub>i_Checkcast: "eff\<^sub>i (Checkcast C, P, (T#ST,LT)) = (Class C # ST,LT)" | eff\<^sub>i_Pop: "eff\<^sub>i (Pop, P, (T#ST,LT)) = (ST,LT)" | eff\<^sub>i_IAdd: "eff\<^sub>i (IAdd, P,(T\<^sub>1#T\<^sub>2#ST,LT)) = (Integer#ST,LT)" | eff\<^sub>i_CmpEq: "eff\<^sub>i (CmpEq, P, (T\<^sub>1#T\<^sub>2#ST,LT)) = (Boolean#ST,LT)" | eff\<^sub>i_IfFalse: "eff\<^sub>i (IfFalse b, P, (T\<^sub>1#ST,LT)) = (ST,LT)" | eff\<^sub>i_Invoke: "eff\<^sub>i (Invoke M n, P, (ST,LT)) = (let C = the_class (ST!n); (D,Ts,T\<^sub>r,b) = method P C M in (T\<^sub>r # drop (n+1) ST, LT))" | eff\<^sub>i_Goto: "eff\<^sub>i (Goto n, P, s) = s" fun is_relevant_class :: "instr \ 'm prog \ cname \ bool" where rel_Getfield: "is_relevant_class (Getfield F D) = (\P C. P \ NullPointer \\<^sup>* C)" | rel_Putfield: "is_relevant_class (Putfield F D) = (\P C. P \ NullPointer \\<^sup>* C)" | rel_Checcast: "is_relevant_class (Checkcast D) = (\P C. P \ ClassCast \\<^sup>* C)" | rel_New: "is_relevant_class (New D) = (\P C. P \ OutOfMemory \\<^sup>* C)" | rel_Throw: "is_relevant_class Throw = (\P C. True)" | rel_Invoke: "is_relevant_class (Invoke M n) = (\P C. True)" | rel_default: "is_relevant_class i = (\P C. False)" definition is_relevant_entry :: "'m prog \ instr \ pc \ ex_entry \ bool" where "is_relevant_entry P i pc e \ (let (f,t,C,h,d) = e in is_relevant_class i P C \ pc \ {f.. instr \ pc \ ex_table \ ex_table" where "relevant_entries P i pc = filter (is_relevant_entry P i pc)" definition xcpt_eff :: "instr \ 'm prog \ pc \ ty\<^sub>i \ ex_table \ (pc \ ty\<^sub>i') list" where "xcpt_eff i P pc \ et = (let (ST,LT) = \ in map (\(f,t,C,h,d). (h, Some (Class C#drop (size ST - d) ST, LT))) (relevant_entries P i pc et))" definition norm_eff :: "instr \ 'm prog \ nat \ ty\<^sub>i \ (pc \ ty\<^sub>i') list" where "norm_eff i P pc \ = map (\pc'. (pc',Some (eff\<^sub>i (i,P,\)))) (succs i \ pc)" definition eff :: "instr \ 'm prog \ pc \ ex_table \ ty\<^sub>i' \ (pc \ ty\<^sub>i') list" where "eff i P pc et t = (case t of None \ [] | Some \ \ (norm_eff i P pc \) @ (xcpt_eff i P pc \ et))" lemma eff_None: "eff i P pc xt None = []" by (simp add: eff_def) lemma eff_Some: "eff i P pc xt (Some \) = norm_eff i P pc \ @ xcpt_eff i P pc \ xt" by (simp add: eff_def) (* FIXME: getfield, \T D. P \ C sees F:T in D \ .. *) text "Conditions under which eff is applicable:" fun app\<^sub>i :: "instr \ 'm prog \ pc \ nat \ ty \ ty\<^sub>i \ bool" where app\<^sub>i_Load: "app\<^sub>i (Load n, P, pc, mxs, T\<^sub>r, (ST,LT)) = (n < length LT \ LT ! n \ Err \ length ST < mxs)" | app\<^sub>i_Store: "app\<^sub>i (Store n, P, pc, mxs, T\<^sub>r, (T#ST, LT)) = (n < length LT)" | app\<^sub>i_Push: "app\<^sub>i (Push v, P, pc, mxs, T\<^sub>r, (ST,LT)) = (length ST < mxs \ typeof v \ None)" | app\<^sub>i_Getfield: "app\<^sub>i (Getfield F C, P, pc, mxs, T\<^sub>r, (T#ST, LT)) = (\T\<^sub>f. P \ C sees F:T\<^sub>f in C \ P \ T \ Class C)" | app\<^sub>i_Putfield: "app\<^sub>i (Putfield F C, P, pc, mxs, T\<^sub>r, (T\<^sub>1#T\<^sub>2#ST, LT)) = (\T\<^sub>f. P \ C sees F:T\<^sub>f in C \ P \ T\<^sub>2 \ (Class C) \ P \ T\<^sub>1 \ T\<^sub>f)" | app\<^sub>i_New: "app\<^sub>i (New C, P, pc, mxs, T\<^sub>r, (ST,LT)) = (is_class P C \ length ST < mxs)" | app\<^sub>i_Checkcast: "app\<^sub>i (Checkcast C, P, pc, mxs, T\<^sub>r, (T#ST,LT)) = (is_class P C \ is_refT T)" | app\<^sub>i_Pop: "app\<^sub>i (Pop, P, pc, mxs, T\<^sub>r, (T#ST,LT)) = True" | app\<^sub>i_IAdd: "app\<^sub>i (IAdd, P, pc, mxs, T\<^sub>r, (T\<^sub>1#T\<^sub>2#ST,LT)) = (T\<^sub>1 = T\<^sub>2 \ T\<^sub>1 = Integer)" | app\<^sub>i_CmpEq: "app\<^sub>i (CmpEq, P, pc, mxs, T\<^sub>r, (T\<^sub>1#T\<^sub>2#ST,LT)) = (T\<^sub>1 = T\<^sub>2 \ is_refT T\<^sub>1 \ is_refT T\<^sub>2)" | app\<^sub>i_IfFalse: "app\<^sub>i (IfFalse b, P, pc, mxs, T\<^sub>r, (Boolean#ST,LT)) = (0 \ int pc + b)" | app\<^sub>i_Goto: "app\<^sub>i (Goto b, P, pc, mxs, T\<^sub>r, s) = (0 \ int pc + b)" | app\<^sub>i_Return: "app\<^sub>i (Return, P, pc, mxs, T\<^sub>r, (T#ST,LT)) = (P \ T \ T\<^sub>r)" | app\<^sub>i_Throw: "app\<^sub>i (Throw, P, pc, mxs, T\<^sub>r, (T#ST,LT)) = is_refT T" | app\<^sub>i_Invoke: "app\<^sub>i (Invoke M n, P, pc, mxs, T\<^sub>r, (ST,LT)) = (n < length ST \ (ST!n \ NT \ (\C D Ts T m. ST!n = Class C \ P \ C sees M:Ts \ T = m in D \ P \ rev (take n ST) [\] Ts)))" | app\<^sub>i_default: "app\<^sub>i (i,P, pc,mxs,T\<^sub>r,s) = False" definition xcpt_app :: "instr \ 'm prog \ pc \ nat \ ex_table \ ty\<^sub>i \ bool" where "xcpt_app i P pc mxs xt \ \ (\(f,t,C,h,d) \ set (relevant_entries P i pc xt). is_class P C \ d \ size (fst \) \ d < mxs)" definition app :: "instr \ 'm prog \ nat \ ty \ nat \ nat \ ex_table \ ty\<^sub>i' \ bool" where "app i P mxs T\<^sub>r pc mpc xt t = (case t of None \ True | Some \ \ app\<^sub>i (i,P,pc,mxs,T\<^sub>r,\) \ xcpt_app i P pc mxs xt \ \ (\(pc',\') \ set (eff i P pc xt t). pc' < mpc))" lemma app_Some: "app i P mxs T\<^sub>r pc mpc xt (Some \) = (app\<^sub>i (i,P,pc,mxs,T\<^sub>r,\) \ xcpt_app i P pc mxs xt \ \ (\(pc',s') \ set (eff i P pc xt (Some \)). pc' < mpc))" by (simp add: app_def) locale eff = jvm_method + fixes eff\<^sub>i and app\<^sub>i and eff and app fixes norm_eff and xcpt_app and xcpt_eff fixes mpc defines "mpc \ size is" defines "eff\<^sub>i i \ \ Effect.eff\<^sub>i (i,P,\)" notes eff\<^sub>i_simps [simp] = Effect.eff\<^sub>i.simps [where P = P, folded eff\<^sub>i_def] defines "app\<^sub>i i pc \ \ Effect.app\<^sub>i (i, P, pc, mxs, T\<^sub>r, \)" notes app\<^sub>i_simps [simp] = Effect.app\<^sub>i.simps [where P=P and mxs=mxs and T\<^sub>r=T\<^sub>r, folded app\<^sub>i_def] defines "xcpt_eff i pc \ \ Effect.xcpt_eff i P pc \ xt" notes xcpt_eff = Effect.xcpt_eff_def [of _ P _ _ xt, folded xcpt_eff_def] defines "norm_eff i pc \ \ Effect.norm_eff i P pc \" notes norm_eff = Effect.norm_eff_def [of _ P, folded norm_eff_def eff\<^sub>i_def] defines "eff i pc \ Effect.eff i P pc xt" notes eff = Effect.eff_def [of _ P _ xt, folded eff_def norm_eff_def xcpt_eff_def] defines "xcpt_app i pc \ \ Effect.xcpt_app i P pc mxs xt \" notes xcpt_app = Effect.xcpt_app_def [of _ P _ mxs xt, folded xcpt_app_def] defines "app i pc \ Effect.app i P mxs T\<^sub>r pc mpc xt" notes app = Effect.app_def [of _ P mxs T\<^sub>r _ mpc xt, folded app_def xcpt_app_def app\<^sub>i_def eff_def] lemma length_cases2: assumes "\LT. P ([],LT)" assumes "\l ST LT. P (l#ST,LT)" shows "P s" by (cases s, cases "fst s") (auto intro!: assms) lemma length_cases3: assumes "\LT. P ([],LT)" assumes "\l LT. P ([l],LT)" assumes "\l ST LT. P (l#ST,LT)" shows "P s" (*<*) proof - obtain xs LT where s: "s = (xs,LT)" by (cases s) show ?thesis proof (cases xs) case Nil with assms s show ?thesis by simp next fix l xs' assume "xs = l#xs'" with assms s show ?thesis by simp qed qed (*>*) lemma length_cases4: assumes "\LT. P ([],LT)" assumes "\l LT. P ([l],LT)" assumes "\l l' LT. P ([l,l'],LT)" assumes "\l l' ST LT. P (l#l'#ST,LT)" shows "P s" (*<*) proof - obtain xs LT where s: "s = (xs,LT)" by (cases s) show ?thesis proof (cases xs) case Nil with assms s show ?thesis by simp next fix l xs' assume xs: "xs = l#xs'" thus ?thesis proof (cases xs') case Nil with assms s xs show ?thesis by simp next fix l' ST assume "xs' = l'#ST" with assms s xs show ?thesis by simp qed qed qed (*>*) text \ \medskip simp rules for @{term app} \ lemma appNone[simp]: "app i P mxs T\<^sub>r pc mpc et None = True" by (simp add: app_def) lemma appLoad[simp]: "app\<^sub>i (Load idx, P, T\<^sub>r, mxs, pc, s) = (\ST LT. s = (ST,LT) \ idx < length LT \ LT!idx \ Err \ length ST < mxs)" by (cases s, simp) lemma appStore[simp]: "app\<^sub>i (Store idx,P,pc,mxs,T\<^sub>r,s) = (\ts ST LT. s = (ts#ST,LT) \ idx < length LT)" by (rule length_cases2, auto) lemma appPush[simp]: "app\<^sub>i (Push v,P,pc,mxs,T\<^sub>r,s) = (\ST LT. s = (ST,LT) \ length ST < mxs \ typeof v \ None)" by (cases s, simp) lemma appGetField[simp]: "app\<^sub>i (Getfield F C,P,pc,mxs,T\<^sub>r,s) = (\ oT vT ST LT. s = (oT#ST, LT) \ P \ C sees F:vT in C \ P \ oT \ (Class C))" by (rule length_cases2 [of _ s]) auto lemma appPutField[simp]: "app\<^sub>i (Putfield F C,P,pc,mxs,T\<^sub>r,s) = (\ vT vT' oT ST LT. s = (vT#oT#ST, LT) \ P \ C sees F:vT' in C \ P \ oT \ (Class C) \ P \ vT \ vT')" by (rule length_cases4 [of _ s], auto) lemma appNew[simp]: "app\<^sub>i (New C,P,pc,mxs,T\<^sub>r,s) = (\ST LT. s=(ST,LT) \ is_class P C \ length ST < mxs)" by (cases s, simp) lemma appCheckcast[simp]: "app\<^sub>i (Checkcast C,P,pc,mxs,T\<^sub>r,s) = (\T ST LT. s = (T#ST,LT) \ is_class P C \ is_refT T)" by (cases s, cases "fst s", simp add: app_def) (cases "hd (fst s)", auto) lemma app\<^sub>iPop[simp]: "app\<^sub>i (Pop,P,pc,mxs,T\<^sub>r,s) = (\ts ST LT. s = (ts#ST,LT))" by (rule length_cases2, auto) lemma appIAdd[simp]: "app\<^sub>i (IAdd,P,pc,mxs,T\<^sub>r,s) = (\ST LT. s = (Integer#Integer#ST,LT))" (*<*) proof - obtain ST LT where [simp]: "s = (ST,LT)" by (cases s) have "ST = [] \ (\T. ST = [T]) \ (\T\<^sub>1 T\<^sub>2 ST'. ST = T\<^sub>1#T\<^sub>2#ST')" by (cases ST, auto, case_tac list, auto) moreover { assume "ST = []" hence ?thesis by simp } moreover { fix T assume "ST = [T]" hence ?thesis by (cases T, auto) } moreover { fix T\<^sub>1 T\<^sub>2 ST' assume "ST = T\<^sub>1#T\<^sub>2#ST'" hence ?thesis by (cases T\<^sub>1, auto) } ultimately show ?thesis by blast qed (*>*) lemma appIfFalse [simp]: "app\<^sub>i (IfFalse b,P,pc,mxs,T\<^sub>r,s) = (\ST LT. s = (Boolean#ST,LT) \ 0 \ int pc + b)" (*<*) - apply (rule length_cases2) - apply simp - apply (case_tac l) - apply auto - done + (is "?P s") +proof(rule length_cases2) + fix LT show "?P ([],LT)" by simp +next + fix l ST LT show "?P (l#ST,LT)" + by (case_tac l) auto +qed (*>*) lemma appCmpEq[simp]: "app\<^sub>i (CmpEq,P,pc,mxs,T\<^sub>r,s) = (\T\<^sub>1 T\<^sub>2 ST LT. s = (T\<^sub>1#T\<^sub>2#ST,LT) \ (\is_refT T\<^sub>1 \ T\<^sub>2 = T\<^sub>1 \ is_refT T\<^sub>1 \ is_refT T\<^sub>2))" by (rule length_cases4, auto) lemma appReturn[simp]: "app\<^sub>i (Return,P,pc,mxs,T\<^sub>r,s) = (\T ST LT. s = (T#ST,LT) \ P \ T \ T\<^sub>r)" by (rule length_cases2, auto) lemma appThrow[simp]: "app\<^sub>i (Throw,P,pc,mxs,T\<^sub>r,s) = (\T ST LT. s=(T#ST,LT) \ is_refT T)" by (rule length_cases2, auto) lemma effNone: "(pc', s') \ set (eff i P pc et None) \ s' = None" by (auto simp add: eff_def xcpt_eff_def norm_eff_def) text \some helpers to make the specification directly executable:\ lemma relevant_entries_append [simp]: "relevant_entries P i pc (xt @ xt') = relevant_entries P i pc xt @ relevant_entries P i pc xt'" by (unfold relevant_entries_def) simp lemma xcpt_app_append [iff]: "xcpt_app i P pc mxs (xt@xt') \ = (xcpt_app i P pc mxs xt \ \ xcpt_app i P pc mxs xt' \)" by (unfold xcpt_app_def) fastforce lemma xcpt_eff_append [simp]: "xcpt_eff i P pc \ (xt@xt') = xcpt_eff i P pc \ xt @ xcpt_eff i P pc \ xt'" by (unfold xcpt_eff_def, cases \) simp lemma app_append [simp]: "app i P pc T mxs mpc (xt@xt') \ = (app i P pc T mxs mpc xt \ \ app i P pc T mxs mpc xt' \)" by (unfold app_def eff_def) auto end diff --git a/thys/Jinja/BV/LBVJVM.thy b/thys/Jinja/BV/LBVJVM.thy --- a/thys/Jinja/BV/LBVJVM.thy +++ b/thys/Jinja/BV/LBVJVM.thy @@ -1,223 +1,242 @@ (* Title: HOL/MicroJava/BV/JVM.thy Author: Tobias Nipkow, Gerwin Klein Copyright 2000 TUM *) section \LBV for the JVM \label{sec:JVM}\ theory LBVJVM imports "../DFA/Abstract_BV" TF_JVM begin type_synonym prog_cert = "cname \ mname \ ty\<^sub>i' err list" definition check_cert :: "jvm_prog \ nat \ nat \ nat \ ty\<^sub>i' err list \ bool" where "check_cert P mxs mxl n cert \ check_types P mxs mxl cert \ size cert = n+1 \ (\i Err) \ cert!n = OK None" definition lbvjvm :: "jvm_prog \ nat \ nat \ ty \ ex_table \ ty\<^sub>i' err list \ instr list \ ty\<^sub>i' err \ ty\<^sub>i' err" where "lbvjvm P mxs maxr T\<^sub>r et cert bs \ wtl_inst_list bs cert (JVM_SemiType.sup P mxs maxr) (JVM_SemiType.le P mxs maxr) Err (OK None) (exec P mxs T\<^sub>r et bs) 0" definition wt_lbv :: "jvm_prog \ cname \ ty list \ ty \ nat \ nat \ ex_table \ ty\<^sub>i' err list \ instr list \ bool" where "wt_lbv P C Ts T\<^sub>r mxs mxl\<^sub>0 et cert ins \ check_cert P mxs (1+size Ts+mxl\<^sub>0) (size ins) cert \ 0 < size ins \ (let start = Some ([],(OK (Class C))#((map OK Ts))@(replicate mxl\<^sub>0 Err)); result = lbvjvm P mxs (1+size Ts+mxl\<^sub>0) T\<^sub>r et cert ins (OK start) in result \ Err)" definition wt_jvm_prog_lbv :: "jvm_prog \ prog_cert \ bool" where "wt_jvm_prog_lbv P cert \ wf_prog (\P C (mn,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,b,et)). wt_lbv P C Ts T\<^sub>r mxs mxl\<^sub>0 et (cert C mn) b) P" definition mk_cert :: "jvm_prog \ nat \ ty \ ex_table \ instr list \ ty\<^sub>m \ ty\<^sub>i' err list" where "mk_cert P mxs T\<^sub>r et bs phi \ make_cert (exec P mxs T\<^sub>r et bs) (map OK phi) (OK None)" definition prg_cert :: "jvm_prog \ ty\<^sub>P \ prog_cert" where "prg_cert P phi C mn \ let (C,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)) = method P C mn in mk_cert P mxs T\<^sub>r et ins (phi C mn)" lemma check_certD [intro?]: "check_cert P mxs mxl n cert \ cert_ok cert n Err (OK None) (states P mxs mxl)" by (unfold cert_ok_def check_cert_def check_types_def) auto lemma (in start_context) wt_lbv_wt_step: assumes lbv: "wt_lbv P C Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is" shows "\\s \ list (size is) A. wt_step r Err step \s \ OK first \\<^sub>r \s!0" (*<*) proof - from wf have "semilat (JVM_SemiType.sl P mxs mxl)" .. hence "semilat (A, r, f)" by (simp add: sl_def2) moreover have "top r Err" by (simp add: JVM_le_Err_conv) moreover have "Err \ A" by (simp add: JVM_states_unfold) moreover have "bottom r (OK None)" by (simp add: JVM_le_Err_conv bottom_def lesub_def Err.le_def split: err.split) moreover have "OK None \ A" by (simp add: JVM_states_unfold) moreover note bounded_step moreover from lbv have "cert_ok cert (size is) Err (OK None) A" by (unfold wt_lbv_def) (auto dest: check_certD) moreover note exec_pres_type moreover from lbv have "wtl_inst_list is cert f r Err (OK None) step 0 (OK first) \ Err" by (simp add: wt_lbv_def lbvjvm_def step_def_exec [symmetric]) moreover note first_in_A moreover from lbv have "0 < size is" by (simp add: wt_lbv_def) ultimately show ?thesis by (rule lbvs.wtl_sound_strong [OF lbvs.intro, OF lbv.intro lbvs_axioms.intro, OF Semilat.intro lbv_axioms.intro]) qed (*>*) lemma (in start_context) wt_lbv_wt_method: assumes lbv: "wt_lbv P C Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is" shows "\\s. wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" (*<*) proof - from lbv have l: "is \ []" by (simp add: wt_lbv_def) moreover from wf lbv C Ts obtain \s where list: "\s \ list (size is) A" and step: "wt_step r Err step \s" and start: "OK first \\<^sub>r \s!0" by (blast dest: wt_lbv_wt_step) from list have [simp]: "size \s = size is" by simp have "size (map ok_val \s) = size is" by simp moreover from l have 0: "0 < size \s" by simp with step obtain \s0 where "\s!0 = OK \s0" by (unfold wt_step_def) blast with start 0 have "wt_start P C Ts mxl\<^sub>0 (map ok_val \s)" by (simp add: wt_start_def JVM_le_Err_conv lesub_def Err.le_def) moreover { from list have "check_types P mxs mxl \s" by (simp add: check_types_def) also from step have "\x \ set \s. x \ Err" by (auto simp add: all_set_conv_all_nth wt_step_def) hence [symmetric]: "map OK (map ok_val \s) = \s" by (auto intro!: map_idI) finally have "check_types P mxs mxl (map OK (map ok_val \s))" . } moreover { note bounded_step moreover from list have "set \s \ A" by simp moreover from step have "wt_err_step (sup_state_opt P) step \s" by (simp add: wt_err_step_def JVM_le_Err_conv) ultimately have "wt_app_eff (sup_state_opt P) app eff (map ok_val \s)" by (auto intro: wt_err_imp_wt_app_eff simp add: exec_def states_def) } ultimately have "wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (map ok_val \s)" by (simp add: wt_method_def2 check_types_def del: map_map) thus ?thesis .. qed (*>*) lemma (in start_context) wt_method_wt_lbv: assumes wt: "wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" defines [simp]: "cert \ mk_cert P mxs T\<^sub>r xt is \s" shows "wt_lbv P C Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is" (*<*) proof - let ?\s = "map OK \s" let ?cert = "make_cert step ?\s (OK None)" from wt obtain 0: "0 < size is" and size: "size is = size ?\s" and ck_types: "check_types P mxs mxl ?\s" and wt_start: "wt_start P C Ts mxl\<^sub>0 \s" and app_eff: "wt_app_eff (sup_state_opt P) app eff \s" by (force simp add: wt_method_def2 check_types_def) from wf have "semilat (JVM_SemiType.sl P mxs mxl)" .. hence "semilat (A, r, f)" by (simp add: sl_def2) moreover have "top r Err" by (simp add: JVM_le_Err_conv) moreover have "Err \ A" by (simp add: JVM_states_unfold) moreover have "bottom r (OK None)" by (simp add: JVM_le_Err_conv bottom_def lesub_def Err.le_def split: err.split) moreover have "OK None \ A" by (simp add: JVM_states_unfold) moreover from wf have "mono r step (size is) A" by (rule step_mono) hence "mono r step (size ?\s) A" by (simp add: size) moreover from exec_pres_type have "pres_type step (size ?\s) A" by (simp add: size) moreover from ck_types have \s_in_A: "set ?\s \ A" by (simp add: check_types_def) hence "\pc. pc < size ?\s \ ?\s!pc \ A \ ?\s!pc \ Err" by auto moreover from bounded_step have "bounded step (size ?\s)" by (simp add: size) moreover have "OK None \ Err" by simp moreover from bounded_step size \s_in_A app_eff have "wt_err_step (sup_state_opt P) step ?\s" by (auto intro: wt_app_eff_imp_wt_err simp add: exec_def states_def) hence "wt_step r Err step ?\s" by (simp add: wt_err_step_def JVM_le_Err_conv) moreover from 0 size have "0 < size \s" by auto hence "?\s!0 = OK (\s!0)" by simp with wt_start have "OK first \\<^sub>r ?\s!0" by (clarsimp simp add: wt_start_def lesub_def Err.le_def JVM_le_Err_conv) moreover note first_in_A moreover have "OK first \ Err" by simp moreover note size ultimately have "wtl_inst_list is ?cert f r Err (OK None) step 0 (OK first) \ Err" by (rule lbvc.wtl_complete [OF lbvc.intro, OF lbv.intro lbvc_axioms.intro, OF Semilat.intro lbv_axioms.intro]) moreover from 0 size have "\s \ []" by auto moreover from ck_types have "check_types P mxs mxl ?cert" - apply (auto simp add: make_cert_def check_types_def JVM_states_unfold) - apply (subst Ok_in_err [symmetric]) - apply (drule nth_mem) - apply auto - done + by (fastforce simp: make_cert_def check_types_def JVM_states_unfold + dest!: nth_mem) moreover note 0 size ultimately show ?thesis by (simp add: wt_lbv_def lbvjvm_def mk_cert_def step_def_exec [symmetric] check_cert_def make_cert_def nth_append) qed (*>*) theorem jvm_lbv_correct: "wt_jvm_prog_lbv P Cert \ wf_jvm_prog P" (*<*) proof - let ?\ = "\C mn. let (C,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)) = method P C mn in SOME \s. wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" - + + let ?A = "\P C (mn,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)). wt_lbv P C Ts T\<^sub>r mxs mxl\<^sub>0 et (Cert C mn) ins" + let ?B = "\P C (M,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). + wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (?\ C M)" + assume wt: "wt_jvm_prog_lbv P Cert" - hence "wf_jvm_prog\<^bsub>?\\<^esub> P" - apply (unfold wf_jvm_prog_phi_def wt_jvm_prog_lbv_def) - apply (erule wf_prog_lift) - apply (auto dest!: start_context.wt_lbv_wt_method [OF start_context.intro] - intro: someI) - apply (erule sees_method_is_class) - done + then have "wf_prog ?A P" by(simp add: wt_jvm_prog_lbv_def) + moreover { + fix wf_md C M b Ts Ca T m bd + assume "wf_prog wf_md P" and sees: "P \ Ca sees M: Ts\T = m in Ca" and + "set Ts \ types P" and "bd = (M, Ts, T, m)" and + "?A P Ca bd" + then have "?B P Ca bd" using sees_method_is_class[OF sees] + by (auto dest!: start_context.wt_lbv_wt_method [OF start_context.intro] + intro: someI) + } + ultimately have "wf_prog ?B P" by(rule wf_prog_lift) + hence "wf_jvm_prog\<^bsub>?\\<^esub> P" by (simp add: wf_jvm_prog_phi_def) thus ?thesis by (unfold wf_jvm_prog_def) blast qed (*>*) theorem jvm_lbv_complete: assumes wt: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "wt_jvm_prog_lbv P (prg_cert P \)" (*<*) - using wt - apply (unfold wf_jvm_prog_phi_def wt_jvm_prog_lbv_def) - apply (erule wf_prog_lift) - apply (auto simp add: prg_cert_def - intro!: start_context.wt_method_wt_lbv start_context.intro) - apply (erule sees_method_is_class) - done +proof - + let ?cert = "prg_cert P \" + let ?A = "\P C (M,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). + wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M)" + let ?B = "\P C (mn,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)). + wt_lbv P C Ts T\<^sub>r mxs mxl\<^sub>0 et (?cert C mn) ins" + + from wt have "wf_prog ?A P" by(clarsimp simp: wf_jvm_prog_def wf_jvm_prog_phi_def) + moreover { + fix wf_md C M Ts Ca T m bd + assume "wf_prog wf_md P" and sees: "P \ Ca sees M: Ts\T = m in Ca" and + "set Ts \ types P" and "bd = (M, Ts, T, m)" and + "?A P Ca bd" + then have "?B P Ca bd" using sees_method_is_class[OF sees] + by (auto simp add: prg_cert_def + intro!: start_context.wt_method_wt_lbv start_context.intro) + } + ultimately have "wf_prog ?B P" by(rule wf_prog_lift) + thus "wt_jvm_prog_lbv P (prg_cert P \)" by (simp add: wt_jvm_prog_lbv_def) +qed (*>*) -end +end diff --git a/thys/Jinja/BV/TF_JVM.thy b/thys/Jinja/BV/TF_JVM.thy --- a/thys/Jinja/BV/TF_JVM.thy +++ b/thys/Jinja/BV/TF_JVM.thy @@ -1,275 +1,292 @@ (* Title: HOL/MicroJava/BV/JVM.thy Author: Tobias Nipkow, Gerwin Klein Copyright 2000 TUM *) section \The Typing Framework for the JVM \label{sec:JVM}\ theory TF_JVM imports "../DFA/Typing_Framework_err" EffectMono BVSpec begin definition exec :: "jvm_prog \ nat \ ty \ ex_table \ instr list \ ty\<^sub>i' err step_type" where "exec G maxs rT et bs \ err_step (size bs) (\pc. app (bs!pc) G maxs rT pc (size bs) et) (\pc. eff (bs!pc) G pc et)" locale JVM_sl = fixes P :: jvm_prog and mxs and mxl\<^sub>0 fixes Ts :: "ty list" and "is" and xt and T\<^sub>r fixes mxl and A and r and f and app and eff and step defines [simp]: "mxl \ 1+size Ts+mxl\<^sub>0" defines [simp]: "A \ states P mxs mxl" defines [simp]: "r \ JVM_SemiType.le P mxs mxl" defines [simp]: "f \ JVM_SemiType.sup P mxs mxl" defines [simp]: "app \ \pc. Effect.app (is!pc) P mxs T\<^sub>r pc (size is) xt" defines [simp]: "eff \ \pc. Effect.eff (is!pc) P pc xt" defines [simp]: "step \ err_step (size is) app eff" locale start_context = JVM_sl + fixes p and C assumes wf: "wf_prog p P" assumes C: "is_class P C" assumes Ts: "set Ts \ types P" fixes first :: ty\<^sub>i' and start defines [simp]: "first \ Some ([],OK (Class C) # map OK Ts @ replicate mxl\<^sub>0 Err)" defines [simp]: "start \ OK first # replicate (size is - 1) (OK None)" subsection \Connecting JVM and Framework\ lemma (in JVM_sl) step_def_exec: "step \ exec P mxs T\<^sub>r xt is" by (simp add: exec_def) lemma special_ex_swap_lemma [iff]: "(? X. (? n. X = A n & P n) & Q X) = (? n. Q(A n) & P n)" by blast lemma ex_in_list [iff]: "(\n. ST \ list n A \ n \ mxs) = (set ST \ A \ size ST \ mxs)" by (unfold list_def) auto lemma singleton_list: "(\n. [Class C] \ list n (types P) \ n \ mxs) = (is_class P C \ 0 < mxs)" by auto lemma set_drop_subset: "set xs \ A \ set (drop n xs) \ A" by (auto dest: in_set_dropD) lemma Suc_minus_minus_le: "n < mxs \ Suc (n - (n - b)) \ mxs" by arith lemma in_listE: "\ xs \ list n A; \size xs = n; set xs \ A\ \ P \ \ P" by (unfold list_def) blast declare is_relevant_entry_def [simp] declare set_drop_subset [simp] theorem (in start_context) exec_pres_type: "pres_type step (size is) A" (*<*) - apply (insert wf) - apply simp - apply (unfold JVM_states_unfold) - apply (rule pres_type_lift) - apply clarify - apply (rename_tac s pc pc' s') - apply (case_tac s) - apply simp - apply (drule effNone) - apply simp - apply (simp add: Effect.app_def xcpt_app_def Effect.eff_def - xcpt_eff_def norm_eff_def relevant_entries_def) - apply (case_tac "is!pc") - - \ \Load\ - apply clarsimp - apply (frule listE_nth_in, assumption) - apply fastforce - - \ \Store\ - apply fastforce - - \ \Push\ - apply (fastforce simp add: typeof_lit_is_type) - - \ \New\ - apply fastforce - - \ \Getfield\ - apply (fastforce dest: sees_field_is_type) - - \ \Putfield\ - apply fastforce - - \ \Checkcast\ - apply fastforce - - defer - - \ \Return\ - apply fastforce - - \ \Pop\ - apply fastforce - - \ \IAdd\ - apply fastforce - - \ \Goto\ - apply fastforce - - \ \CmpEq\ - apply fastforce - - \ \IfFalse\ - apply fastforce - - \ \Throw\ - apply fastforce - - \ \Invoke\ - apply (clarsimp split!: if_splits) - apply fastforce - apply (erule disjE) - prefer 2 - apply fastforce - apply clarsimp - apply (rule conjI) - apply (drule (1) sees_wf_mdecl) - apply (clarsimp simp add: wf_mdecl_def) - apply arith - done +proof - + let ?n = "size is" and ?app = app and ?step = eff + let ?mxl = "1 + length Ts + mxl\<^sub>0" + let ?A = "opt((Union {list n (types P) |n. n <= mxs}) \ + list ?mxl (err(types P)))" + have "pres_type (err_step ?n ?app ?step) ?n (err ?A)" + proof(rule pres_type_lift) + have "\s pc pc' s'. s\?A \ pc < ?n \ ?app pc s + \ (pc', s')\set (?step pc s) \ s' \ ?A" + proof - + fix s pc pc' s' + assume asms: "s\?A" "pc < ?n" "?app pc s" "(pc', s')\set (?step pc s)" + show "s' \ ?A" + proof(cases s) + case None + then show ?thesis using asms by (fastforce dest: effNone) + next + case (Some ab) + then show ?thesis using asms proof(cases "is!pc") + case Load + then show ?thesis using asms + by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def + xcpt_eff_def norm_eff_def + dest: listE_nth_in) + next + case Push + then show ?thesis using asms Some + by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def + xcpt_eff_def norm_eff_def typeof_lit_is_type) + next + case Getfield + then show ?thesis using asms wf + by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def + xcpt_eff_def norm_eff_def + dest: sees_field_is_type) + next + case (Invoke M n) + obtain a b where [simp]: "s = \(a,b)\" using Some asms(1) by blast + show ?thesis + proof(cases "a!n = NT") + case True + then show ?thesis using Invoke asms wf + by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def + xcpt_eff_def norm_eff_def) + next + case False + have "(pc', s') \ set (norm_eff (Invoke M n) P pc (a, b)) \ + (pc', s') \ set (xcpt_eff (Invoke M n) P pc (a, b) xt)" + using Invoke asms(4) by (simp add: Effect.eff_def) + then show ?thesis proof(rule disjE) + assume "(pc', s') \ set (xcpt_eff (Invoke M n) P pc (a, b) xt)" + then show ?thesis using Invoke asms(1-3) + by (fastforce simp: Effect.app_def xcpt_app_def xcpt_eff_def) + next + assume norm: "(pc', s') \ set (norm_eff (Invoke M n) P pc (a, b))" + also have "Suc (length a - Suc n) \ mxs" using Invoke asms(1,3) + by (simp add: Effect.app_def xcpt_app_def) arith + ultimately show ?thesis using False Invoke asms(1-3) wf + by (auto simp: Effect.app_def xcpt_app_def norm_eff_def wf_mdecl_def + dest!: sees_wf_mdecl) + qed + qed + qed (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def + xcpt_eff_def norm_eff_def)+ + qed + qed + then show "\s\?A. \p. p < ?n \ ?app p s \ (\(q, s')\set (?step p s). s' \ ?A)" + by clarsimp + qed + then show ?thesis by (simp add: JVM_states_unfold) +qed (*>*) declare is_relevant_entry_def [simp del] declare set_drop_subset [simp del] lemma lesubstep_type_simple: "xs [\\<^bsub>Product.le (=) r\<^esub>] ys \ set xs {\\<^bsub>r\<^esub>} set ys" (*<*) - apply (unfold lesubstep_type_def) - apply clarify - apply (simp add: set_conv_nth) - apply clarify - apply (drule le_listD, assumption) - apply (clarsimp simp add: lesub_def Product.le_def) - apply (rule exI) - apply (rule conjI) - apply (rule exI) - apply (rule conjI) - apply (rule sym) - apply assumption - apply assumption - apply assumption - done +proof - + assume assm: "xs [\\<^bsub>Product.le (=) r\<^esub>] ys" + have "\a b i y. (a, b) = xs ! i \ i < length xs + \ \\'. (\i. (a, \') = ys ! i \ i < length xs) \ b \\<^bsub>r\<^esub> \'" + proof - + fix a b i assume ith: "(a, b) = xs ! i" and len: "i < length xs" + obtain \ where "ys ! i = (a, \) \ r b \" + using le_listD[OF assm len] ith + by (clarsimp simp: lesub_def Product.le_def) + then have "(a, \) = ys ! i \ b \\<^bsub>r\<^esub> \" + by (clarsimp simp: lesub_def) + with len show "\\'. (\i. (a, \') = ys ! i \ i < length xs) \ b \\<^bsub>r\<^esub> \'" + by fastforce + qed + then show "set xs {\\<^bsub>r\<^esub>} set ys" using assm + by (clarsimp simp: lesubstep_type_def set_conv_nth) +qed (*>*) declare is_relevant_entry_def [simp del] lemma conjI2: "\ A; A \ B \ \ A \ B" by blast lemma (in JVM_sl) eff_mono: - "\wf_prog p P; pc < length is; s \\<^bsub>sup_state_opt P\<^esub> t; app pc t\ - \ set (eff pc s) {\\<^bsub>sup_state_opt P\<^esub>} set (eff pc t)" +assumes wf: "wf_prog p P" and "pc < length is" and + lesub: "s \\<^bsub>sup_state_opt P\<^esub> t" and app: "app pc t" +shows "set (eff pc s) {\\<^bsub>sup_state_opt P\<^esub>} set (eff pc t)" (*<*) - apply simp - apply (unfold Effect.eff_def) - apply (cases t) - apply (simp add: lesub_def) - apply (rename_tac a) - apply (cases s) - apply simp - apply (rename_tac b) - apply simp - apply (rule lesubstep_union) - prefer 2 - apply (rule lesubstep_type_simple) - apply (simp add: xcpt_eff_def) - apply (rule le_listI) - apply (simp add: split_beta) - apply (simp add: split_beta) - apply (simp add: lesub_def fun_of_def) - apply (case_tac a) - apply (case_tac b) - apply simp - apply (subgoal_tac "size ab = size aa") - prefer 2 - apply (clarsimp simp add: list_all2_lengthD) - apply simp - apply (clarsimp simp add: norm_eff_def lesubstep_type_def lesub_def iff del: sup_state_conv) - apply (rule exI) - apply (rule conjI2) - apply (rule imageI) - apply (clarsimp simp add: Effect.app_def iff del: sup_state_conv) - apply (drule (2) succs_mono) - apply blast - apply simp - apply (erule eff\<^sub>i_mono) - apply simp - apply assumption - apply clarsimp - apply clarsimp - done +proof(cases t) + case None then show ?thesis using lesub + by (simp add: Effect.eff_def lesub_def) +next + case tSome: (Some a) + show ?thesis proof(cases s) + case None then show ?thesis using lesub + by (simp add: Effect.eff_def lesub_def) + next + case (Some b) + let ?norm = "\x. norm_eff (is ! pc) P pc x" + let ?xcpt = "\x. xcpt_eff (is ! pc) P pc x xt" + let ?r = "Product.le (=) (sup_state_opt P)" + let ?\' = "\eff\<^sub>i (is ! pc, P, a)\" + { + fix x assume xb: "x \ set (succs (is ! pc) b pc)" + then have appi: "app\<^sub>i (is ! pc, P, pc, mxs, T\<^sub>r, a)" and + bia: "P \ b \\<^sub>i a" and appa: "app pc \a\" + using lesub app tSome Some by (auto simp add: lesub_def Effect.app_def) + have xa: "x \ set (succs (is ! pc) a pc)" + using xb succs_mono[OF wf appi bia] by auto + then have "(x, ?\') \ (\pc'. (pc', ?\')) ` set (succs (is ! pc) a pc)" + by (rule imageI) + moreover have "P \ \eff\<^sub>i (is ! pc, P, b)\ \' ?\'" + using xb xa eff\<^sub>i_mono[OF wf bia] appa by fastforce + ultimately have "\\'. (x, \') \ (\pc'. (pc', \eff\<^sub>i (is ! pc, P, a)\)) ` set (succs (is ! pc) a pc) \ + P \ \eff\<^sub>i (is ! pc, P, b)\ \' \'" by blast + } + then have norm: "set (?norm b) {\\<^bsub>sup_state_opt P\<^esub>} set (?norm a)" + using tSome Some by (clarsimp simp: norm_eff_def lesubstep_type_def lesub_def) + obtain a1 b1 a2 b2 where a: "a = (a1, b1)" and b: "b = (a2, b2)" + using tSome Some by fastforce + then have a12: "size a2 = size a1" using lesub tSome Some + by (clarsimp simp: lesub_def list_all2_lengthD) + have "length (?xcpt b) = length (?xcpt a)" + by (simp add: xcpt_eff_def split_beta) + moreover have "\n. n < length (?xcpt b) \ (?xcpt b) ! n \\<^bsub>?r\<^esub> (?xcpt a) ! n" + using lesub tSome Some a b a12 + by (simp add: xcpt_eff_def split_beta fun_of_def) (simp add: lesub_def) + ultimately have "?xcpt b [\\<^bsub>?r\<^esub>] ?xcpt a" + by(rule le_listI) + then have "set (?xcpt b) {\\<^bsub>sup_state_opt P\<^esub>} set (?xcpt a)" + by (rule lesubstep_type_simple) + moreover note norm + ultimately have + "set (?norm b) \ set (?xcpt b) {\\<^bsub>sup_state_opt P\<^esub>} set (?norm a) \ set (?xcpt a)" + by(intro lesubstep_union) + then show ?thesis using tSome Some by(simp add: Effect.eff_def) + qed +qed (*>*) lemma (in JVM_sl) bounded_step: "bounded step (size is)" (*<*) - apply simp - apply (unfold bounded_def err_step_def Effect.app_def Effect.eff_def) - apply (auto simp add: error_def map_snd_def split: err.splits option.splits) - done + by (auto simp: bounded_def err_step_def Effect.app_def Effect.eff_def + error_def map_snd_def + split: err.splits option.splits) (*>*) theorem (in JVM_sl) step_mono: "wf_prog wf_mb P \ mono r step (size is) A" (*<*) - apply (simp add: JVM_le_Err_conv) - apply (insert bounded_step) - apply (unfold JVM_states_unfold) - apply (rule mono_lift) - apply blast - apply (unfold app_mono_def lesub_def) - apply clarsimp - apply (erule (2) app_mono) - apply simp - apply clarify - apply (drule eff_mono) - apply (auto simp add: lesub_def) - done +proof - + assume wf: "wf_prog wf_mb P" + let ?r = "sup_state_opt P" and ?n = "length is" and ?app = app and ?step = eff + let ?A = "opt (\ {list n (types P) |n. n \ mxs} \ + list (1 + length Ts + mxl\<^sub>0) (err (types P)))" + have "order ?r" using wf by simp + moreover have "app_mono ?r ?app ?n ?A" using app_mono[OF wf] + by (clarsimp simp: app_mono_def lesub_def) + moreover have "bounded (err_step ?n ?app ?step) ?n" using bounded_step + by simp + moreover have "\s p t. s \ ?A \ p < ?n \ s \\<^bsub>?r\<^esub> t \ + ?app p t \ set (?step p s) {\\<^bsub>?r\<^esub>} set (?step p t)" + using eff_mono[OF wf] by simp + ultimately have "mono (Err.le ?r) (err_step ?n ?app ?step) ?n (err ?A)" + by(rule mono_lift) + then show "mono r step (size is) A" using bounded_step + by (simp add: JVM_le_Err_conv JVM_states_unfold) +qed (*>*) lemma (in start_context) first_in_A [iff]: "OK first \ A" using Ts C by (force intro!: list_appendI simp add: JVM_states_unfold) lemma (in JVM_sl) wt_method_def2: "wt_method P C' Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s = (is \ [] \ size \s = size is \ OK ` set \s \ states P mxs mxl \ wt_start P C' Ts mxl\<^sub>0 \s \ wt_app_eff (sup_state_opt P) app eff \s)" (*<*) - apply (unfold wt_method_def wt_app_eff_def wt_instr_def lesub_def check_types_def) - apply auto - done + by (unfold wt_method_def wt_app_eff_def wt_instr_def lesub_def + check_types_def) auto (*>*) end diff --git a/thys/Jinja/Common/Auxiliary.thy b/thys/Jinja/Common/Auxiliary.thy --- a/thys/Jinja/Common/Auxiliary.thy +++ b/thys/Jinja/Common/Auxiliary.thy @@ -1,125 +1,121 @@ (* Title: Jinja/Common/Basis.thy Author: David von Oheimb, Tobias Nipkow Copyright 1999 TU Muenchen *) chapter \Jinja Source Language \label{cha:j}\ section \Auxiliary Definitions\ theory Auxiliary imports Main begin (* FIXME move and possibly turn into a general simproc *) lemma nat_add_max_le[simp]: "((n::nat) + max i j \ m) = (n + i \ m \ n + j \ m)" (*<*)by arith(*>*) lemma Suc_add_max_le[simp]: "(Suc(n + max i j) \ m) = (Suc(n + i) \ m \ Suc(n + j) \ m)" (*<*)by arith(*>*) notation Some ("(\_\)") (*<*) declare option.splits[split] Let_def[simp] subset_insertI2 [simp] Cons_eq_map_conv [iff] (*>*) subsection \\distinct_fst\\ definition distinct_fst :: "('a \ 'b) list \ bool" where "distinct_fst \ distinct \ map fst" lemma distinct_fst_Nil [simp]: "distinct_fst []" (*<*) -apply (unfold distinct_fst_def) -apply (simp (no_asm)) -done +by (unfold distinct_fst_def) (simp (no_asm)) (*>*) lemma distinct_fst_Cons [simp]: "distinct_fst ((k,x)#kxs) = (distinct_fst kxs \ (\y. (k,y) \ set kxs))" (*<*) -apply (unfold distinct_fst_def) -apply (auto simp:image_def) -done +by (unfold distinct_fst_def) (auto simp:image_def) (*>*) (* lemma distinct_fst_append: "\ distinct_fst kxs'; distinct_fst kxs; \(k,x) \ set kxs. \(k',x') \ set kxs'. k' \ k \ \ distinct_fst(kxs @ kxs')" by (induct kxs) (auto dest: fst_in_set_lemma) lemma distinct_fst_map_inj: "\ distinct_fst kxs; inj f \ \ distinct_fst (map (\(k,x). (f k, g k x)) kxs)" by (induct kxs) (auto dest: fst_in_set_lemma simp: inj_eq) *) lemma map_of_SomeI: "\ distinct_fst kxs; (k,x) \ set kxs \ \ map_of kxs k = Some x" (*<*)by (induct kxs) (auto simp:fun_upd_apply)(*>*) subsection \Using @{term list_all2} for relations\ definition fun_of :: "('a \ 'b) set \ 'a \ 'b \ bool" where "fun_of S \ \x y. (x,y) \ S" text \Convenience lemmas\ (*<*) declare fun_of_def [simp] (*>*) lemma rel_list_all2_Cons [iff]: "list_all2 (fun_of S) (x#xs) (y#ys) = ((x,y) \ S \ list_all2 (fun_of S) xs ys)" (*<*)by simp(*>*) lemma rel_list_all2_Cons1: "list_all2 (fun_of S) (x#xs) ys = (\z zs. ys = z#zs \ (x,z) \ S \ list_all2 (fun_of S) xs zs)" (*<*)by (cases ys) auto(*>*) lemma rel_list_all2_Cons2: "list_all2 (fun_of S) xs (y#ys) = (\z zs. xs = z#zs \ (z,y) \ S \ list_all2 (fun_of S) zs ys)" (*<*)by (cases xs) auto(*>*) lemma rel_list_all2_refl: "(\x. (x,x) \ S) \ list_all2 (fun_of S) xs xs" (*<*)by (simp add: list_all2_refl)(*>*) lemma rel_list_all2_antisym: "\ (\x y. \(x,y) \ S; (y,x) \ T\ \ x = y); list_all2 (fun_of S) xs ys; list_all2 (fun_of T) ys xs \ \ xs = ys" (*<*)by (rule list_all2_antisym) auto(*>*) lemma rel_list_all2_trans: "\ \a b c. \(a,b) \ R; (b,c) \ S\ \ (a,c) \ T; list_all2 (fun_of R) as bs; list_all2 (fun_of S) bs cs\ \ list_all2 (fun_of T) as cs" (*<*)by (rule list_all2_trans) auto(*>*) lemma rel_list_all2_update_cong: "\ i S \ \ list_all2 (fun_of S) (xs[i:=x]) (ys[i:=y])" (*<*)by (simp add: list_all2_update_cong)(*>*) lemma rel_list_all2_nthD: "\ list_all2 (fun_of S) xs ys; p < size xs \ \ (xs!p,ys!p) \ S" (*<*)by (drule list_all2_nthD) auto(*>*) lemma rel_list_all2I: "\ length a = length b; \n. n < length a \ (a!n,b!n) \ S \ \ list_all2 (fun_of S) a b" (*<*)by (erule list_all2_all_nthI) simp(*>*) (*<*)declare fun_of_def [simp del](*>*) end diff --git a/thys/Jinja/Common/Conform.thy b/thys/Jinja/Common/Conform.thy --- a/thys/Jinja/Common/Conform.thy +++ b/thys/Jinja/Common/Conform.thy @@ -1,221 +1,147 @@ (* Title: Jinja/J/Conform.thy Author: David von Oheimb, Tobias Nipkow Copyright 1999 Technische Universitaet Muenchen *) section \Conformance Relations for Type Soundness Proofs\ theory Conform imports Exceptions begin definition conf :: "'m prog \ heap \ val \ ty \ bool" ("_,_ \ _ :\ _" [51,51,51,51] 50) where "P,h \ v :\ T \ \T'. typeof\<^bsub>h\<^esub> v = Some T' \ P \ T' \ T" definition oconf :: "'m prog \ heap \ obj \ bool" ("_,_ \ _ \" [51,51,51] 50) where "P,h \ obj \ \ let (C,fs) = obj in \F D T. P \ C has F:T in D \ (\v. fs(F,D) = Some v \ P,h \ v :\ T)" definition hconf :: "'m prog \ heap \ bool" ("_ \ _ \" [51,51] 50) where "P \ h \ \ (\a obj. h a = Some obj \ P,h \ obj \) \ preallocated h" definition lconf :: "'m prog \ heap \ (vname \ val) \ (vname \ ty) \ bool" ("_,_ \ _ '(:\') _" [51,51,51,51] 50) where "P,h \ l (:\) E \ \V v. l V = Some v \ (\T. E V = Some T \ P,h \ v :\ T)" abbreviation confs :: "'m prog \ heap \ val list \ ty list \ bool" ("_,_ \ _ [:\] _" [51,51,51,51] 50) where "P,h \ vs [:\] Ts \ list_all2 (conf P h) vs Ts" subsection\Value conformance \:\\\ lemma conf_Null [simp]: "P,h \ Null :\ T = P \ NT \ T" -(*<*) -apply (unfold conf_def) -apply (simp (no_asm)) -done -(*>*) +(*<*)by (simp (no_asm) add: conf_def)(*>*) lemma typeof_conf[simp]: "typeof\<^bsub>h\<^esub> v = Some T \ P,h \ v :\ T" -(*<*) -apply (unfold conf_def) -apply (induct v) -apply auto -done -(*>*) +(*<*)by (induct v) (auto simp: conf_def)(*>*) lemma typeof_lit_conf[simp]: "typeof v = Some T \ P,h \ v :\ T" (*<*)by (rule typeof_conf[OF typeof_lit_typeof])(*>*) lemma defval_conf[simp]: "P,h \ default_val T :\ T" -(*<*) -apply (unfold conf_def) -apply (cases T) -apply auto -done -(*>*) +(*<*)by (cases T) (auto simp: conf_def)(*>*) lemma conf_upd_obj: "h a = Some(C,fs) \ (P,h(a\(C,fs')) \ x :\ T) = (P,h \ x :\ T)" -(*<*) -apply (unfold conf_def) -apply (rule val.induct) -apply (auto simp:fun_upd_apply) -done -(*>*) +(*<*)by (rule val.induct) (auto simp:conf_def fun_upd_apply)(*>*) lemma conf_widen: "P,h \ v :\ T \ P \ T \ T' \ P,h \ v :\ T'" -(*<*) -apply (unfold conf_def) -apply (induct v) -apply (auto intro: widen_trans) -done -(*>*) +(*<*)by (induct v) (auto intro: widen_trans simp: conf_def)(*>*) lemma conf_hext: "h \ h' \ P,h \ v :\ T \ P,h' \ v :\ T" -(*<*) -apply (unfold conf_def) -apply (induct v) -apply (auto dest: hext_objD) -done -(*>*) +(*<*)by (induct v) (auto dest: hext_objD simp: conf_def)(*>*) lemma conf_ClassD: "P,h \ v :\ Class C \ v = Null \ (\a obj T. v = Addr a \ h a = Some obj \ obj_ty obj = T \ P \ T \ Class C)" -(*<*) -apply (unfold conf_def) -apply(induct "v") -apply(auto) -done -(*>*) +(*<*)by(induct v) (auto simp: conf_def)(*>*) lemma conf_NT [iff]: "P,h \ v :\ NT = (v = Null)" (*<*)by (auto simp add: conf_def)(*>*) lemma non_npD: "\ v \ Null; P,h \ v :\ Class C \ \ \a C' fs. v = Addr a \ h a = Some(C',fs) \ P \ C' \\<^sup>* C" -(*<*) -apply (drule conf_ClassD) -apply auto -done -(*>*) +(*<*)by (auto dest: conf_ClassD)(*>*) subsection\Value list conformance \[:\]\\ lemma confs_widens [trans]: "\P,h \ vs [:\] Ts; P \ Ts [\] Ts'\ \ P,h \ vs [:\] Ts'" -(*<*) - apply (rule list_all2_trans) - apply (rule conf_widen, assumption, assumption) - apply assumption - apply assumption - done -(*>*) +(*<*)by(auto intro: list_all2_trans conf_widen)(*>*) lemma confs_rev: "P,h \ rev s [:\] t = (P,h \ s [:\] rev t)" -(*<*) - apply rule - apply (rule subst [OF list_all2_rev]) - apply simp - apply (rule subst [OF list_all2_rev]) - apply simp - done -(*>*) +(*<*)by (simp add: list_all2_rev1)(*>*) lemma confs_conv_map: "\Ts'. P,h \ vs [:\] Ts' = (\Ts. map typeof\<^bsub>h\<^esub> vs = map Some Ts \ P \ Ts [\] Ts')" (*<*) -apply(induct vs) - apply simp -apply(case_tac Ts') -apply(auto simp add:conf_def) -done +proof(induct vs) + case (Cons a vs) + then show ?case by(case_tac Ts') (auto simp add:conf_def) +qed simp (*>*) lemma confs_hext: "P,h \ vs [:\] Ts \ h \ h' \ P,h' \ vs [:\] Ts" (*<*)by (erule list_all2_mono, erule conf_hext, assumption)(*>*) lemma confs_Cons2: "P,h \ xs [:\] y#ys = (\z zs. xs = z#zs \ P,h \ z :\ y \ P,h \ zs [:\] ys)" (*<*)by (rule list_all2_Cons2)(*>*) subsection "Object conformance" lemma oconf_hext: "P,h \ obj \ \ h \ h' \ P,h' \ obj \" -(*<*) -apply (unfold oconf_def) -apply (fastforce elim:conf_hext) -done -(*>*) +(*<*)by (fastforce elim:conf_hext simp: oconf_def)(*>*) lemma oconf_init_fields: "P \ C has_fields FDTs \ P,h \ (C, init_fields FDTs) \" by(fastforce simp add: has_field_def oconf_def init_fields_def map_of_map dest: has_fields_fun) lemma oconf_fupd [intro?]: "\ P \ C has F:T in D; P,h \ v :\ T; P,h \ (C,fs) \ \ \ P,h \ (C, fs((F,D)\v)) \" -(*<*) - apply (unfold oconf_def has_field_def) - apply clarsimp - apply (drule (1) has_fields_fun) - apply (auto simp add: fun_upd_apply) - done -(*>*) +(*<*)by (auto dest: has_fields_fun simp add: oconf_def has_field_def fun_upd_apply)(*>*) (*<*) lemmas oconf_new = oconf_hext [OF _ hext_new] lemmas oconf_upd_obj = oconf_hext [OF _ hext_upd_obj] (*>*) subsection "Heap conformance" lemma hconfD: "\ P \ h \; h a = Some obj \ \ P,h \ obj \" -(*<*) -apply (unfold hconf_def) -apply (fast) -done -(*>*) +(*<*)by (unfold hconf_def) fast(*>*) lemma hconf_new: "\ P \ h \; h a = None; P,h \ obj \ \ \ P \ h(a\obj) \" (*<*)by (unfold hconf_def) (auto intro: oconf_new preallocated_new)(*>*) lemma hconf_upd_obj: "\ P \ h\; h a = Some(C,fs); P,h \ (C,fs')\ \ \ P \ h(a\(C,fs'))\" (*<*)by (unfold hconf_def) (auto intro: oconf_upd_obj preallocated_upd_obj)(*>*) subsection "Local variable conformance" lemma lconf_hext: "\ P,h \ l (:\) E; h \ h' \ \ P,h' \ l (:\) E" -(*<*) -apply (unfold lconf_def) -apply (fast elim: conf_hext) -done -(*>*) +(*<*)by (unfold lconf_def) (fast elim: conf_hext)(*>*) lemma lconf_upd: "\ P,h \ l (:\) E; P,h \ v :\ T; E V = Some T \ \ P,h \ l(V\v) (:\) E" -(*<*) -apply (unfold lconf_def) -apply auto -done -(*>*) +(*<*)by (unfold lconf_def) auto(*>*) lemma lconf_empty[iff]: "P,h \ Map.empty (:\) E" (*<*)by(simp add:lconf_def)(*>*) lemma lconf_upd2: "\P,h \ l (:\) E; P,h \ v :\ T\ \ P,h \ l(V\v) (:\) E(V\T)" (*<*)by(simp add:lconf_def)(*>*) end diff --git a/thys/Jinja/Common/Decl.thy b/thys/Jinja/Common/Decl.thy --- a/thys/Jinja/Common/Decl.thy +++ b/thys/Jinja/Common/Decl.thy @@ -1,63 +1,64 @@ (* Title: HOL/MicroJava/J/Decl.thy Author: David von Oheimb Copyright 1999 Technische Universitaet Muenchen *) section \Class Declarations and Programs\ theory Decl imports Type begin type_synonym fdecl = "vname \ ty" \ \field declaration\ type_synonym 'm mdecl = "mname \ ty list \ ty \ 'm" \ \method = name, arg.\ types, return type, body\ type_synonym 'm "class" = "cname \ fdecl list \ 'm mdecl list" \ \class = superclass, fields, methods\ type_synonym 'm cdecl = "cname \ 'm class" \ \class declaration\ type_synonym 'm prog = "'m cdecl list" \ \program\ (*<*) translations (type) "fdecl" <= (type) "vname \ ty" (type) "'c mdecl" <= (type) "mname \ ty list \ ty \ 'c" (type) "'c class" <= (type) "cname \ fdecl list \ ('c mdecl) list" (type) "'c cdecl" <= (type) "cname \ ('c class)" (type) "'c prog" <= (type) "('c cdecl) list" (*>*) definition "class" :: "'m prog \ cname \ 'm class" where "class \ map_of" definition is_class :: "'m prog \ cname \ bool" where "is_class P C \ class P C \ None" lemma finite_is_class: "finite {C. is_class P C}" (*<*) -apply (unfold is_class_def class_def) -apply (fold dom_def) -apply (rule finite_dom_map_of) -done +proof - + have "{C. is_class P C} = dom (map_of P)" + by (simp add: is_class_def class_def dom_def) + thus ?thesis by (simp add: finite_dom_map_of) +qed (*>*) definition is_type :: "'m prog \ ty \ bool" where "is_type P T \ (case T of Void \ True | Boolean \ True | Integer \ True | NT \ True | Class C \ is_class P C)" lemma is_type_simps [simp]: "is_type P Void \ is_type P Boolean \ is_type P Integer \ is_type P NT \ is_type P (Class C) = is_class P C" (*<*)by(simp add:is_type_def)(*>*) abbreviation "types P == Collect (is_type P)" end diff --git a/thys/Jinja/Common/Objects.thy b/thys/Jinja/Common/Objects.thy --- a/thys/Jinja/Common/Objects.thy +++ b/thys/Jinja/Common/Objects.thy @@ -1,183 +1,156 @@ (* Title: HOL/MicroJava/J/State.thy Author: David von Oheimb Copyright 1999 Technische Universitaet Muenchen *) section \Objects and the Heap\ theory Objects imports TypeRel Value begin subsection\Objects\ type_synonym fields = "vname \ cname \ val" \ \field name, defining class, value\ type_synonym obj = "cname \ fields" \ \class instance with class name and fields\ definition obj_ty :: "obj \ ty" where "obj_ty obj \ Class (fst obj)" definition init_fields :: "((vname \ cname) \ ty) list \ fields" where "init_fields \ map_of \ map (\(F,T). (F,default_val T))" \ \a new, blank object with default values in all fields:\ definition blank :: "'m prog \ cname \ obj" where "blank P C \ (C,init_fields (fields P C))" lemma [simp]: "obj_ty (C,fs) = Class C" (*<*)by (simp add: obj_ty_def)(*>*) subsection\Heap\ type_synonym heap = "addr \ obj" abbreviation cname_of :: "heap \ addr \ cname" where "cname_of hp a == fst (the (hp a))" definition new_Addr :: "heap \ addr option" where "new_Addr h \ if \a. h a = None then Some(LEAST a. h a = None) else None" definition cast_ok :: "'m prog \ cname \ heap \ val \ bool" where "cast_ok P C h v \ v = Null \ P \ cname_of h (the_Addr v) \\<^sup>* C" definition hext :: "heap \ heap \ bool" ("_ \ _" [51,51] 50) where "h \ h' \ \a C fs. h a = Some(C,fs) \ (\fs'. h' a = Some(C,fs'))" primrec typeof_h :: "heap \ val \ ty option" ("typeof\<^bsub>_\<^esub>") where "typeof\<^bsub>h\<^esub> Unit = Some Void" | "typeof\<^bsub>h\<^esub> Null = Some NT" | "typeof\<^bsub>h\<^esub> (Bool b) = Some Boolean" | "typeof\<^bsub>h\<^esub> (Intg i) = Some Integer" | "typeof\<^bsub>h\<^esub> (Addr a) = (case h a of None \ None | Some(C,fs) \ Some(Class C))" lemma new_Addr_SomeD: "new_Addr h = Some a \ h a = None" (*<*)by(fastforce simp add:new_Addr_def split:if_splits intro:LeastI)(*>*) lemma [simp]: "(typeof\<^bsub>h\<^esub> v = Some Boolean) = (\b. v = Bool b)" (*<*)by(induct v) auto(*>*) lemma [simp]: "(typeof\<^bsub>h\<^esub> v = Some Integer) = (\i. v = Intg i)" (*<*)by(cases v) auto(*>*) lemma [simp]: "(typeof\<^bsub>h\<^esub> v = Some NT) = (v = Null)" (*<*)by(cases v) auto(*>*) lemma [simp]: "(typeof\<^bsub>h\<^esub> v = Some(Class C)) = (\a fs. v = Addr a \ h a = Some(C,fs))" (*<*)by(cases v) auto(*>*) lemma [simp]: "h a = Some(C,fs) \ typeof\<^bsub>(h(a\(C,fs')))\<^esub> v = typeof\<^bsub>h\<^esub> v" (*<*)by(induct v) (auto simp:fun_upd_apply)(*>*) text\For literal values the first parameter of @{term typeof} can be set to @{term Map.empty} because they do not contain addresses:\ abbreviation typeof :: "val \ ty option" where "typeof v == typeof_h Map.empty v" lemma typeof_lit_typeof: "typeof v = Some T \ typeof\<^bsub>h\<^esub> v = Some T" (*<*)by(cases v) auto(*>*) lemma typeof_lit_is_type: "typeof v = Some T \ is_type P T" (*<*)by (induct v) (auto simp:is_type_def)(*>*) subsection \Heap extension \\\\ lemma hextI: "\a C fs. h a = Some(C,fs) \ (\fs'. h' a = Some(C,fs')) \ h \ h'" -(*<*) -apply (unfold hext_def) -apply auto -done -(*>*) +(*<*)by(auto simp: hext_def)(*>*) lemma hext_objD: "\ h \ h'; h a = Some(C,fs) \ \ \fs'. h' a = Some(C,fs')" -(*<*) -apply (unfold hext_def) -apply (force) -done -(*>*) +(*<*)by(auto simp: hext_def)(*>*) lemma hext_refl [iff]: "h \ h" -(*<*) -apply (rule hextI) -apply (fast) -done -(*>*) +(*<*)by (rule hextI) fast(*>*) lemma hext_new [simp]: "h a = None \ h \ h(a\x)" -(*<*) -apply (rule hextI) -apply (auto simp:fun_upd_apply) -done -(*>*) +(*<*)by (rule hextI) (auto simp:fun_upd_apply)(*>*) lemma hext_trans: "\ h \ h'; h' \ h'' \ \ h \ h''" -(*<*) -apply (rule hextI) -apply (fast dest: hext_objD) -done -(*>*) +(*<*)by (rule hextI) (fast dest: hext_objD)(*>*) lemma hext_upd_obj: "h a = Some (C,fs) \ h \ h(a\(C,fs'))" -(*<*) -apply (rule hextI) -apply (auto simp:fun_upd_apply) -done -(*>*) +(*<*)by (rule hextI) (auto simp:fun_upd_apply)(*>*) lemma hext_typeof_mono: "\ h \ h'; typeof\<^bsub>h\<^esub> v = Some T \ \ typeof\<^bsub>h'\<^esub> v = Some T" (*<*) -apply(cases v) - apply simp - apply simp - apply simp - apply simp -apply(fastforce simp:hext_def) -done +proof(cases v) + case Addr assume "h \ h'" and "typeof\<^bsub>h\<^esub> v = \T\" + then show ?thesis using Addr by(fastforce simp:hext_def) +qed simp_all (*>*) text \Code generator setup for @{term "new_Addr"}\ definition gen_new_Addr :: "heap \ addr \ addr option" where "gen_new_Addr h n \ if \a. a \ n \ h a = None then Some(LEAST a. a \ n \ h a = None) else None" lemma new_Addr_code_code [code]: "new_Addr h = gen_new_Addr h 0" by(simp add: new_Addr_def gen_new_Addr_def split del: if_split cong: if_cong) lemma gen_new_Addr_code [code]: "gen_new_Addr h n = (if h n = None then Some n else gen_new_Addr h (Suc n))" apply(simp add: gen_new_Addr_def) apply(rule impI) apply(rule conjI) apply safe[1] apply(fastforce intro: Least_equality) apply(rule arg_cong[where f=Least]) apply(rule ext) apply(case_tac "n = ac") apply simp apply(auto)[1] apply clarify apply(subgoal_tac "a = n") apply simp apply(rule Least_equality) apply auto[2] apply(rule ccontr) apply(erule_tac x="a" in allE) apply simp done end diff --git a/thys/Jinja/Common/TypeRel.thy b/thys/Jinja/Common/TypeRel.thy --- a/thys/Jinja/Common/TypeRel.thy +++ b/thys/Jinja/Common/TypeRel.thy @@ -1,632 +1,679 @@ (* Title: Jinja/Common/TypeRel.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Relations between Jinja Types\ theory TypeRel imports "HOL-Library.Transitive_Closure_Table" Decl begin subsection\The subclass relations\ inductive_set subcls1 :: "'m prog \ (cname \ cname) set" and subcls1' :: "'m prog \ [cname, cname] \ bool" ("_ \ _ \\<^sup>1 _" [71,71,71] 70) for P :: "'m prog" where "P \ C \\<^sup>1 D \ (C,D) \ subcls1 P" | subcls1I: "\class P C = Some (D,rest); C \ Object\ \ P \ C \\<^sup>1 D" abbreviation subcls :: "'m prog \ [cname, cname] \ bool" ("_ \ _ \\<^sup>* _" [71,71,71] 70) where "P \ C \\<^sup>* D \ (C,D) \ (subcls1 P)\<^sup>*" lemma subcls1D: "P \ C \\<^sup>1 D \ C \ Object \ (\fs ms. class P C = Some (D,fs,ms))" (*<*)by(erule subcls1.induct)(fastforce simp add:is_class_def)(*>*) lemma [iff]: "\ P \ Object \\<^sup>1 C" (*<*)by(fastforce dest:subcls1D)(*>*) lemma [iff]: "(P \ Object \\<^sup>* C) = (C = Object)" (*<*) -apply(rule iffI) - apply(erule converse_rtranclE) - apply simp_all -done +proof(rule iffI) + assume "P \ Object \\<^sup>* C" then show "C = Object" + by(auto elim: converse_rtranclE) +qed simp (*>*) lemma subcls1_def2: "subcls1 P = (SIGMA C:{C. is_class P C}. {D. C\Object \ fst (the (class P C))=D})" (*<*) by (fastforce simp:is_class_def dest: subcls1D elim: subcls1I) (*>*) lemma finite_subcls1: "finite (subcls1 P)" (*<*) -apply (simp add: subcls1_def2) -apply(rule finite_SigmaI [OF finite_is_class]) -apply(rule_tac B = "{fst (the (class P C))}" in finite_subset) -apply auto -done +proof - + let ?SIG = "SIGMA C:{C. is_class P C}. {D. fst (the (class P C)) = D \ C \ Object}" + have "subcls1 P = ?SIG" by(simp add: subcls1_def2) + also have "finite ?SIG" + proof(rule finite_SigmaI [OF finite_is_class]) + fix C assume C_in: "C \ {C. is_class P C}" + then show "finite {D. fst (the (class P C)) = D \ C \ Object}" + by(rule_tac finite_subset[where B = "{fst (the (class P C))}"]) auto + qed + ultimately show ?thesis by simp +qed (*>*) (* lemma subcls_is_class: "(C,D) \ (subcls1 P)\<^sup>+ \ is_class P C" apply (unfold is_class_def) apply(erule trancl_trans_induct) apply (auto dest!: subcls1D) done lemma subcls_is_class: "P \ C \\<^sup>* D \ is_class P D \ is_class P C" apply (unfold is_class_def) apply (erule rtrancl_induct) apply (drule_tac [2] subcls1D) apply auto done *) subsection\The subtype relations\ inductive widen :: "'m prog \ ty \ ty \ bool" ("_ \ _ \ _" [71,71,71] 70) for P :: "'m prog" where widen_refl[iff]: "P \ T \ T" | widen_subcls: "P \ C \\<^sup>* D \ P \ Class C \ Class D" | widen_null[iff]: "P \ NT \ Class C" abbreviation widens :: "'m prog \ ty list \ ty list \ bool" ("_ \ _ [\] _" [71,71,71] 70) where "widens P Ts Ts' \ list_all2 (widen P) Ts Ts'" lemma [iff]: "(P \ T \ Void) = (T = Void)" (*<*)by (auto elim: widen.cases)(*>*) lemma [iff]: "(P \ T \ Boolean) = (T = Boolean)" (*<*)by (auto elim: widen.cases)(*>*) lemma [iff]: "(P \ T \ Integer) = (T = Integer)" (*<*)by (auto elim: widen.cases)(*>*) lemma [iff]: "(P \ Void \ T) = (T = Void)" (*<*)by (auto elim: widen.cases)(*>*) lemma [iff]: "(P \ Boolean \ T) = (T = Boolean)" (*<*)by (auto elim: widen.cases)(*>*) lemma [iff]: "(P \ Integer \ T) = (T = Integer)" (*<*)by (auto elim: widen.cases)(*>*) lemma Class_widen: "P \ Class C \ T \ \D. T = Class D" (*<*) -apply (ind_cases "P \ Class C \ T") -apply auto -done +by (ind_cases "P \ Class C \ T") auto (*>*) lemma [iff]: "(P \ T \ NT) = (T = NT)" (*<*) -apply(cases T) -apply(auto dest:Class_widen) -done +by(cases T) (auto dest:Class_widen) (*>*) lemma Class_widen_Class [iff]: "(P \ Class C \ Class D) = (P \ C \\<^sup>* D)" (*<*) -apply (rule iffI) -apply (ind_cases "P \ Class C \ Class D") -apply (auto elim: widen_subcls) -done +proof(rule iffI) + show "P \ Class C \ Class D \ P \ C \\<^sup>* D" + proof(ind_cases "P \ Class C \ Class D") qed(auto) +qed(auto elim: widen_subcls) (*>*) lemma widen_Class: "(P \ T \ Class C) = (T = NT \ (\D. T = Class D \ P \ D \\<^sup>* C))" (*<*)by(induct T, auto)(*>*) lemma widen_trans[trans]: "\P \ S \ U; P \ U \ T\ \ P \ S \ T" (*<*) proof - assume "P\S \ U" thus "\T. P \ U \ T \ P \ S \ T" proof induct case (widen_refl T T') thus "P \ T \ T'" . next case (widen_subcls C D T) then obtain E where "T = Class E" by (blast dest: Class_widen) with widen_subcls show "P \ Class C \ T" by (auto elim: rtrancl_trans) next case (widen_null C RT) then obtain D where "RT = Class D" by (blast dest: Class_widen) thus "P \ NT \ RT" by auto qed qed (*>*) lemma widens_trans [trans]: "\P \ Ss [\] Ts; P \ Ts [\] Us\ \ P \ Ss [\] Us" (*<*)by (rule list_all2_trans, rule widen_trans)(*>*) (*<*) lemmas widens_refl [iff] = list_all2_refl [of "widen P", OF widen_refl] for P lemmas widens_Cons [iff] = list_all2_Cons1 [of "widen P"] for P (*>*) subsection\Method lookup\ inductive Methods :: "['m prog, cname, mname \ (ty list \ ty \ 'm) \ cname] \ bool" ("_ \ _ sees'_methods _" [51,51,51] 50) for P :: "'m prog" where sees_methods_Object: "\ class P Object = Some(D,fs,ms); Mm = map_option (\m. (m,Object)) \ map_of ms \ \ P \ Object sees_methods Mm" | sees_methods_rec: "\ class P C = Some(D,fs,ms); C \ Object; P \ D sees_methods Mm; Mm' = Mm ++ (map_option (\m. (m,C)) \ map_of ms) \ \ P \ C sees_methods Mm'" lemma sees_methods_fun: assumes 1: "P \ C sees_methods Mm" shows "\Mm'. P \ C sees_methods Mm' \ Mm' = Mm" (*<*) using 1 proof induct case (sees_methods_rec C D fs ms Dres Cres Cres') have "class": "class P C = Some (D, fs, ms)" and notObj: "C \ Object" and Dmethods: "P \ D sees_methods Dres" and IH: "\Dres'. P \ D sees_methods Dres' \ Dres' = Dres" and Cres: "Cres = Dres ++ (map_option (\m. (m,C)) \ map_of ms)" and Cmethods': "P \ C sees_methods Cres'" by fact+ from Cmethods' notObj "class" obtain Dres' where Dmethods': "P \ D sees_methods Dres'" and Cres': "Cres' = Dres' ++ (map_option (\m. (m,C)) \ map_of ms)" by(auto elim: Methods.cases) from Cres Cres' IH[OF Dmethods'] show "Cres' = Cres" by simp next case sees_methods_Object thus ?case by(auto elim: Methods.cases) qed (*>*) lemma visible_methods_exist: "P \ C sees_methods Mm \ Mm M = Some(m,D) \ (\D' fs ms. class P D = Some(D',fs,ms) \ map_of ms M = Some m)" (*<*)by(induct rule:Methods.induct) auto(*>*) lemma sees_methods_decl_above: assumes Csees: "P \ C sees_methods Mm" shows "Mm M = Some(m,D) \ P \ C \\<^sup>* D" (*<*) using Csees proof induct case sees_methods_Object thus ?case by auto next case sees_methods_rec thus ?case by(fastforce simp:map_option_case split:option.splits elim:converse_rtrancl_into_rtrancl[OF subcls1I]) qed (*>*) lemma sees_methods_idemp: assumes Cmethods: "P \ C sees_methods Mm" shows "\m D. Mm M = Some(m,D) \ \Mm'. (P \ D sees_methods Mm') \ Mm' M = Some(m,D)" (*<*) using Cmethods proof induct case sees_methods_Object thus ?case by(fastforce dest: Methods.sees_methods_Object) next case sees_methods_rec thus ?case by(fastforce split:option.splits dest: Methods.sees_methods_rec) qed (*>*) (*FIXME something wrong with induct: need to attache [consumes 1] directly to proof of thm, declare does not work. *) lemma sees_methods_decl_mono: assumes sub: "P \ C' \\<^sup>* C" shows "P \ C sees_methods Mm \ \Mm' Mm\<^sub>2. P \ C' sees_methods Mm' \ Mm' = Mm ++ Mm\<^sub>2 \ (\M m D. Mm\<^sub>2 M = Some(m,D) \ P \ D \\<^sup>* C)" (*<*) (is "_ \ \Mm' Mm2. ?Q C' C Mm' Mm2") using sub proof (induct rule:converse_rtrancl_induct) assume "P \ C sees_methods Mm" hence "?Q C C Mm Map.empty" by simp thus "\Mm' Mm2. ?Q C C Mm' Mm2" by blast next fix C'' C' assume sub1: "P \ C'' \\<^sup>1 C'" and sub: "P \ C' \\<^sup>* C" and IH: "P \ C sees_methods Mm \ \Mm' Mm2. P \ C' sees_methods Mm' \ Mm' = Mm ++ Mm2 \ (\M m D. Mm2 M = Some(m,D) \ P \ D \\<^sup>* C)" and Csees: "P \ C sees_methods Mm" from IH[OF Csees] obtain Mm' Mm2 where C'sees: "P \ C' sees_methods Mm'" and Mm': "Mm' = Mm ++ Mm2" and subC:"\M m D. Mm2 M = Some(m,D) \ P \ D \\<^sup>* C" by blast obtain fs ms where "class": "class P C'' = Some(C',fs,ms)" "C'' \ Object" using subcls1D[OF sub1] by blast let ?Mm3 = "map_option (\m. (m,C'')) \ map_of ms" have "P \ C'' sees_methods (Mm ++ Mm2) ++ ?Mm3" using sees_methods_rec[OF "class" C'sees refl] Mm' by simp hence "?Q C'' C ((Mm ++ Mm2) ++ ?Mm3) (Mm2++?Mm3)" using converse_rtrancl_into_rtrancl[OF sub1 sub] by simp (simp add:map_add_def subC split:option.split) thus "\Mm' Mm2. ?Q C'' C Mm' Mm2" by blast qed (*>*) definition Method :: "'m prog \ cname \ mname \ ty list \ ty \ 'm \ cname \ bool" ("_ \ _ sees _: _\_ = _ in _" [51,51,51,51,51,51,51] 50) where "P \ C sees M: Ts\T = m in D \ \Mm. P \ C sees_methods Mm \ Mm M = Some((Ts,T,m),D)" definition has_method :: "'m prog \ cname \ mname \ bool" ("_ \ _ has _" [51,0,51] 50) where "P \ C has M \ \Ts T m D. P \ C sees M:Ts\T = m in D" lemma sees_method_fun: "\P \ C sees M:TS\T = m in D; P \ C sees M:TS'\T' = m' in D' \ \ TS' = TS \ T' = T \ m' = m \ D' = D" (*<*)by(fastforce dest: sees_methods_fun simp:Method_def)(*>*) lemma sees_method_decl_above: "P \ C sees M:Ts\T = m in D \ P \ C \\<^sup>* D" (*<*)by(clarsimp simp:Method_def sees_methods_decl_above)(*>*) lemma visible_method_exists: "P \ C sees M:Ts\T = m in D \ \D' fs ms. class P D = Some(D',fs,ms) \ map_of ms M = Some(Ts,T,m)" (*<*)by(fastforce simp:Method_def dest!: visible_methods_exist)(*>*) lemma sees_method_idemp: "P \ C sees M:Ts\T=m in D \ P \ D sees M:Ts\T=m in D" (*<*)by(fastforce simp: Method_def intro:sees_methods_idemp)(*>*) lemma sees_method_decl_mono: - "\ P \ C' \\<^sup>* C; P \ C sees M:Ts\T = m in D; - P \ C' sees M:Ts'\T' = m' in D' \ \ P \ D' \\<^sup>* D" +assumes sub: "P \ C' \\<^sup>* C" and + C_sees: "P \ C sees M:Ts\T=m in D" and + C'_sees: "P \ C' sees M:Ts'\T'=m' in D'" +shows "P \ D' \\<^sup>* D" (*<*) -apply(frule sees_method_decl_above) -apply(unfold Method_def) -apply clarsimp -apply(drule (1) sees_methods_decl_mono) -apply clarsimp -apply(drule (1) sees_methods_fun) -apply clarsimp -apply(blast intro:rtrancl_trans) -done +proof - + obtain Ms where Ms: "P \ C sees_methods Ms" + using C_sees by(auto simp: Method_def) + obtain Ms' Ms2 where Ms': "P \ C' sees_methods Ms'" and + Ms'_def: "Ms' = Ms ++ Ms2" and + Ms2_imp: "(\M m D. Ms2 M = \(m, D)\ \ P \ D \\<^sup>* C)" + using sees_methods_decl_mono[OF sub Ms] by clarsimp + have "(Ms ++ Ms2) M = \((Ts', T', m'), D')\" + using C'_sees sees_methods_fun[OF Ms'] Ms'_def by(clarsimp simp: Method_def) + then have "Ms2 M = \((Ts', T', m'), D')\ \ + Ms2 M = None \ Ts = Ts' \ T = T' \ m = m' \ D = D'" + using C_sees sees_methods_fun[OF Ms] by(clarsimp simp: Method_def) + also have "Ms2 M = \((Ts', T', m'), D')\ \ P \ D' \\<^sup>* C" + using Ms2_imp by simp + ultimately show ?thesis using sub sees_method_decl_above[OF C_sees] by auto +qed (*>*) lemma sees_method_is_class: "\ P \ C sees M:Ts\T = m in D \ \ is_class P C" (*<*)by (auto simp add: is_class_def Method_def elim: Methods.induct)(*>*) subsection\Field lookup\ inductive Fields :: "['m prog, cname, ((vname \ cname) \ ty) list] \ bool" ("_ \ _ has'_fields _" [51,51,51] 50) for P :: "'m prog" where has_fields_rec: "\ class P C = Some(D,fs,ms); C \ Object; P \ D has_fields FDTs; FDTs' = map (\(F,T). ((F,C),T)) fs @ FDTs \ \ P \ C has_fields FDTs'" | has_fields_Object: "\ class P Object = Some(D,fs,ms); FDTs = map (\(F,T). ((F,Object),T)) fs \ \ P \ Object has_fields FDTs" lemma has_fields_fun: assumes 1: "P \ C has_fields FDTs" shows "\FDTs'. P \ C has_fields FDTs' \ FDTs' = FDTs" (*<*) using 1 proof induct case (has_fields_rec C D fs ms Dres Cres Cres') have "class": "class P C = Some (D, fs, ms)" and notObj: "C \ Object" and DFields: "P \ D has_fields Dres" and IH: "\Dres'. P \ D has_fields Dres' \ Dres' = Dres" and Cres: "Cres = map (\(F,T). ((F,C),T)) fs @ Dres" and CFields': "P \ C has_fields Cres'" by fact+ from CFields' notObj "class" obtain Dres' where DFields': "P \ D has_fields Dres'" and Cres': "Cres' = map (\(F,T). ((F,C),T)) fs @ Dres'" by(auto elim: Fields.cases) from Cres Cres' IH[OF DFields'] show "Cres' = Cres" by simp next case has_fields_Object thus ?case by(auto elim: Fields.cases) qed (*>*) lemma all_fields_in_has_fields: assumes sub: "P \ C has_fields FDTs" shows "\ P \ C \\<^sup>* D; class P D = Some(D',fs,ms); (F,T) \ set fs \ \ ((F,D),T) \ set FDTs" (*<*) -using sub apply(induct) - apply(simp add:image_def) - apply(erule converse_rtranclE) - apply(force) - apply(drule subcls1D) - apply fastforce -apply(force simp:image_def) -done +using sub proof(induct) + case (has_fields_rec C D' fs ms FDTs FDTs') + then have C_D: "P \ C \\<^sup>* D" by simp + then show ?case proof(rule converse_rtranclE) + assume "C = D" + then show ?case using has_fields_rec by force + next + fix y assume sub1: "P \ C \\<^sup>1 y" and sub2: "P \ y \\<^sup>* D" + then show ?case using has_fields_rec subcls1D[OF sub1] by simp + qed +next + case (has_fields_Object D fs ms FDTs) + then show ?case by force +qed (*>*) lemma has_fields_decl_above: assumes fields: "P \ C has_fields FDTs" shows "((F,D),T) \ set FDTs \ P \ C \\<^sup>* D" (*<*) -using fields apply(induct) - prefer 2 apply fastforce -apply clarsimp -apply(erule disjE) - apply(clarsimp simp add:image_def) -apply simp -apply(blast dest:subcls1I converse_rtrancl_into_rtrancl) -done +using fields proof(induct) + case (has_fields_rec C D' fs ms FDTs FDTs') + then have "((F, D), T) \ (\x. case x of (F, x) \ ((F, C), x)) ` set fs \ + ((F, D), T) \ set FDTs" by clarsimp + then show ?case proof(rule disjE) + assume "((F, D), T) \ (\x. case x of (F, x) \ ((F, C), x)) ` set fs" + then show ?case using has_fields_rec by clarsimp + next + assume "((F, D), T) \ set FDTs" + then show ?case using has_fields_rec + by(blast dest:subcls1I converse_rtrancl_into_rtrancl) + qed +next + case (has_fields_Object D fs ms FDTs) + then show ?case by fastforce +qed (*>*) lemma subcls_notin_has_fields: assumes fields: "P \ C has_fields FDTs" shows "((F,D),T) \ set FDTs \ (D,C) \ (subcls1 P)\<^sup>+" (*<*) -using fields apply(induct) - prefer 2 apply(fastforce dest: tranclD) -apply clarsimp -apply(erule disjE) - apply(clarsimp simp add:image_def) - apply(drule tranclD) - apply clarify - apply(frule subcls1D) - apply(fastforce dest:all_fields_in_has_fields) -apply simp -apply(blast dest:subcls1I trancl_into_trancl) -done +using fields proof(induct) + case (has_fields_rec C D' fs ms FDTs FDTs') + then have "((F, D), T) \ (\x. case x of (F, x) \ ((F, C), x)) ` set fs + \ ((F, D), T) \ set FDTs" by clarsimp + then show ?case proof(rule disjE) + assume "((F, D), T) \ (\x. case x of (F, x) \ ((F, C), x)) ` set fs" + then have CD[simp]: "C = D" and fs: "(F, T) \ set fs" by clarsimp+ + then have "(D, D) \ (subcls1 P)\<^sup>+ \ False" proof - + assume DD: "(D, D) \ (subcls1 P)\<^sup>+" + obtain z where z1: "P \ D \\<^sup>1 z" and z_s: "P \ z \\<^sup>* D" + using tranclD[OF DD] by clarsimp + have [simp]: "z = D'" using subcls1D[OF z1] has_fields_rec.hyps(1) by clarsimp + then have "((F, D), T) \ set FDTs" + using z_s all_fields_in_has_fields[OF has_fields_rec.hyps(3) _ has_fields_rec.hyps(1) fs] + by simp + then have "(D, z) \ (subcls1 P)\<^sup>+" using has_fields_rec.hyps(4) by simp + then show False using z1 by auto + qed + then show ?case by clarsimp + next + assume "((F, D), T) \ set FDTs" + then show ?case using has_fields_rec by(blast dest:subcls1I trancl_into_trancl) + qed +next + case (has_fields_Object D fs ms FDTs) + then show ?case by(fastforce dest: tranclD) +qed (*>*) lemma has_fields_mono_lem: assumes sub: "P \ D \\<^sup>* C" shows "P \ C has_fields FDTs \ \pre. P \ D has_fields pre@FDTs \ dom(map_of pre) \ dom(map_of FDTs) = {}" (*<*) -using sub apply(induct rule:converse_rtrancl_induct) - apply(rule_tac x = "[]" in exI) - apply simp -apply clarsimp -apply(rename_tac D' D pre) -apply(subgoal_tac "(D',C) : (subcls1 P)^+") - prefer 2 apply(erule (1) rtrancl_into_trancl2) -apply(drule subcls1D) -apply clarsimp -apply(rename_tac fs ms) -apply(drule (2) has_fields_rec) - apply(rule refl) -apply(rule_tac x = "map (\(F,T). ((F,D'),T)) fs @ pre" in exI) -apply simp -apply(simp add:Int_Un_distrib2) -apply(rule equals0I) -apply(auto dest: subcls_notin_has_fields simp:dom_map_of_conv_image_fst image_def) -done +using sub proof(induct rule:converse_rtrancl_induct) + case base + then show ?case by(rule_tac x = "[]" in exI) simp +next + case (step D' D) + then obtain pre where D_flds: "P \ D has_fields pre @ FDTs" and + dom: "dom (map_of pre) \ dom (map_of FDTs) = {}" by clarsimp + have "(D',C) \ (subcls1 P)^+" by (rule rtrancl_into_trancl2[OF step.hyps(1,2)]) + obtain fs ms where D'_cls: "class P D' = \(D, fs, ms)\" "D' \ Object" + using subcls1D[OF step.hyps(1)] by clarsimp+ + have "P \ D' has_fields map (\(F, T). ((F, D'), T)) fs @ pre @ FDTs" + using has_fields_rec[OF D'_cls D_flds] by simp + also have "dom (map_of (map (\(F, T). ((F, D'), T)) fs @ pre)) + \ dom (map_of FDTs) = {}" + using dom subcls_notin_has_fields[OF D_flds, where D=D'] step.hyps(1) + by(auto simp:dom_map_of_conv_image_fst) fast + ultimately show ?case + by(rule_tac x = "map (\(F,T). ((F,D'),T)) fs @ pre" in exI) simp +qed (*>*) (* FIXME why is Field not displayed correctly? TypeRel qualifier seems to confuse printer*) definition has_field :: "'m prog \ cname \ vname \ ty \ cname \ bool" ("_ \ _ has _:_ in _" [51,51,51,51,51] 50) where "P \ C has F:T in D \ \FDTs. P \ C has_fields FDTs \ map_of FDTs (F,D) = Some T" lemma has_field_mono: - "\ P \ C has F:T in D; P \ C' \\<^sup>* C \ \ P \ C' has F:T in D" +assumes has: " P \ C has F:T in D" and sub: "P \ C' \\<^sup>* C" +shows "P \ C' has F:T in D" (*<*) -apply(clarsimp simp:has_field_def) -apply(drule (1) has_fields_mono_lem) -apply(fastforce simp: map_add_def split:option.splits) -done +proof - + obtain FDTs where FDTs:"P \ C has_fields FDTs" and "map_of FDTs (F, D) = \T\" + using has by(clarsimp simp: has_field_def) + also obtain pre where "P \ C' has_fields pre @ FDTs" + and "dom (map_of pre) \ dom (map_of FDTs) = {}" + using has_fields_mono_lem[OF sub FDTs] by clarify + ultimately show ?thesis by(fastforce simp: has_field_def map_add_def split:option.splits) +qed (*>*) definition sees_field :: "'m prog \ cname \ vname \ ty \ cname \ bool" ("_ \ _ sees _:_ in _" [51,51,51,51,51] 50) where "P \ C sees F:T in D \ \FDTs. P \ C has_fields FDTs \ map_of (map (\((F,D),T). (F,(D,T))) FDTs) F = Some(D,T)" lemma map_of_remap_SomeD: "map_of (map (\((k,k'),x). (k,(k',x))) t) k = Some (k',x) \ map_of t (k, k') = Some x" (*<*)by (induct t) (auto simp:fun_upd_apply split: if_split_asm)(*>*) lemma has_visible_field: "P \ C sees F:T in D \ P \ C has F:T in D" (*<*)by(auto simp add:has_field_def sees_field_def map_of_remap_SomeD)(*>*) lemma sees_field_fun: "\P \ C sees F:T in D; P \ C sees F:T' in D'\ \ T' = T \ D' = D" (*<*)by(fastforce simp:sees_field_def dest:has_fields_fun)(*>*) lemma sees_field_decl_above: "P \ C sees F:T in D \ P \ C \\<^sup>* D" (*<*) -apply(auto simp:sees_field_def) -apply(blast intro: has_fields_decl_above map_of_SomeD map_of_remap_SomeD) -done +by(auto simp:sees_field_def + intro: has_fields_decl_above map_of_SomeD map_of_remap_SomeD) (*>*) -(* FIXME ugly *) lemma sees_field_idemp: - "P \ C sees F:T in D \ P \ D sees F:T in D" +assumes sees: "P \ C sees F:T in D" +shows "P \ D sees F:T in D" (*<*) - apply (unfold sees_field_def) - apply clarsimp - apply (rule_tac P = "map_of xs F = y" for xs y in mp) - prefer 2 - apply assumption - apply (thin_tac "map_of xs F = y" for xs y) - apply (erule Fields.induct) - apply clarsimp - apply (frule map_of_SomeD) - apply clarsimp - apply (fastforce intro: has_fields_rec) - apply clarsimp - apply (frule map_of_SomeD) - apply clarsimp - apply (fastforce intro: has_fields_Object) - done +proof - + obtain FDTs where C_flds: "P \ C has_fields FDTs" + and FDTs: "map_of (map (\((F, D), T). (F, D, T)) FDTs) F = \(D, T)\" + (is "?FDTs") + using sees by(clarsimp simp: sees_field_def) + have map: "\C' fs. map_of (map ((\((F, D), a). (F, D, a)) \ (\(F, y). ((F, C'), y))) fs) F + = \(D, T)\ \ C' = D \ (F, T) \ set fs" + by(frule map_of_SomeD) clarsimp + have "?FDTs \ P \ D sees F:T in D" + using C_flds proof induct + case NObj: has_fields_rec + then show ?case using map by (fastforce intro: has_fields_rec simp: sees_field_def) + next + case Obj: has_fields_Object + then show ?case using map by(fastforce intro: has_fields_Object simp: sees_field_def) + qed + then show ?thesis using FDTs by(rule mp) +qed (*>*) subsection "Functional lookup" definition "method" :: "'m prog \ cname \ mname \ cname \ ty list \ ty \ 'm" where "method P C M \ THE (D,Ts,T,m). P \ C sees M:Ts \ T = m in D" definition field :: "'m prog \ cname \ vname \ cname \ ty" where "field P C F \ THE (D,T). P \ C sees F:T in D" definition fields :: "'m prog \ cname \ ((vname \ cname) \ ty) list" where "fields P C \ THE FDTs. P \ C has_fields FDTs" lemma fields_def2 [simp]: "P \ C has_fields FDTs \ fields P C = FDTs" (*<*)by (unfold fields_def) (auto dest: has_fields_fun)(*>*) lemma field_def2 [simp]: "P \ C sees F:T in D \ field P C F = (D,T)" (*<*)by (unfold field_def) (auto dest: sees_field_fun)(*>*) lemma method_def2 [simp]: "P \ C sees M: Ts\T = m in D \ method P C M = (D,Ts,T,m)" (*<*)by (unfold method_def) (auto dest: sees_method_fun)(*>*) subsection "Code generator setup" code_pred (modes: i \ i \ i \ bool, i \ i \ o \ bool) subcls1p . declare subcls1_def [code_pred_def] code_pred (modes: i \ i \ o \ bool, i \ i \ i \ bool) [inductify] subcls1 . definition subcls' where "subcls' G = (subcls1p G)^**" code_pred (modes: i \ i \ i \ bool, i \ i \ o \ bool) [inductify] subcls' . lemma subcls_conv_subcls' [code_unfold]: "(subcls1 G)^* = {(C, D). subcls' G C D}" by (simp add: subcls'_def subcls1_def rtrancl_def) code_pred (modes: i \ i \ i \ bool) widen . code_pred (modes: i \ i \ o \ bool) Fields . lemma has_field_code [code_pred_intro]: "\ P \ C has_fields FDTs; map_of FDTs (F, D) = \T\ \ \ P \ C has F:T in D" by(auto simp add: has_field_def) code_pred (modes: i \ i \ i \ o \ i \ bool, i \ i \ i \ i \ i \ bool) has_field by(auto simp add: has_field_def) lemma sees_field_code [code_pred_intro]: "\ P \ C has_fields FDTs; map_of (map (\((F, D), T). (F, D, T)) FDTs) F = \(D, T)\ \ \ P \ C sees F:T in D" by(auto simp add: sees_field_def) code_pred (modes: i \ i \ i \ o \ o \ bool, i \ i \ i \ o \ i \ bool, i \ i \ i \ i \ o \ bool, i \ i \ i \ i \ i \ bool) sees_field by(auto simp add: sees_field_def) code_pred (modes: i \ i \ o \ bool) Methods . lemma Method_code [code_pred_intro]: "\ P \ C sees_methods Mm; Mm M = \((Ts, T, m), D)\ \ \ P \ C sees M: Ts\T = m in D" by(auto simp add: Method_def) code_pred (modes: i \ i \ i \ o \ o \ o \ o \ bool, i \ i \ i \ i \ i \ i \ i \ bool) Method by(auto simp add: Method_def) lemma eval_Method_i_i_i_o_o_o_o_conv: "Predicate.eval (Method_i_i_i_o_o_o_o P C M) = (\(Ts, T, m, D). P \ C sees M:Ts\T=m in D)" by(auto intro: Method_i_i_i_o_o_o_oI elim: Method_i_i_i_o_o_o_oE intro!: ext) lemma method_code [code]: "method P C M = Predicate.the (Predicate.bind (Method_i_i_i_o_o_o_o P C M) (\(Ts, T, m, D). Predicate.single (D, Ts, T, m)))" apply (rule sym, rule the_eqI) apply (simp add: method_def eval_Method_i_i_i_o_o_o_o_conv) apply (rule arg_cong [where f=The]) apply (auto simp add: Sup_fun_def Sup_bool_def fun_eq_iff) done lemma eval_Fields_conv: "Predicate.eval (Fields_i_i_o P C) = (\FDTs. P \ C has_fields FDTs)" by(auto intro: Fields_i_i_oI elim: Fields_i_i_oE intro!: ext) lemma fields_code [code]: "fields P C = Predicate.the (Fields_i_i_o P C)" by(simp add: fields_def Predicate.the_def eval_Fields_conv) lemma eval_sees_field_i_i_i_o_o_conv: "Predicate.eval (sees_field_i_i_i_o_o P C F) = (\(T, D). P \ C sees F:T in D)" by(auto intro!: ext intro: sees_field_i_i_i_o_oI elim: sees_field_i_i_i_o_oE) lemma eval_sees_field_i_i_i_o_i_conv: "Predicate.eval (sees_field_i_i_i_o_i P C F D) = (\T. P \ C sees F:T in D)" by(auto intro!: ext intro: sees_field_i_i_i_o_iI elim: sees_field_i_i_i_o_iE) lemma field_code [code]: "field P C F = Predicate.the (Predicate.bind (sees_field_i_i_i_o_o P C F) (\(T, D). Predicate.single (D, T)))" apply (rule sym, rule the_eqI) apply (simp add: field_def eval_sees_field_i_i_i_o_o_conv) apply (rule arg_cong [where f=The]) apply (auto simp add: Sup_fun_def Sup_bool_def fun_eq_iff) done (*<*) end (*>*) diff --git a/thys/Jinja/Common/WellForm.thy b/thys/Jinja/Common/WellForm.thy --- a/thys/Jinja/Common/WellForm.thy +++ b/thys/Jinja/Common/WellForm.thy @@ -1,434 +1,423 @@ (* Title: Jinja/J/WellForm.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Generic Well-formedness of programs\ theory WellForm imports TypeRel SystemClasses begin text \\noindent This theory defines global well-formedness conditions for programs but does not look inside method bodies. Hence it works for both Jinja and JVM programs. Well-typing of expressions is defined elsewhere (in theory \WellType\). Because Jinja does not have method overloading, its policy for method overriding is the classical one: \emph{covariant in the result type but contravariant in the argument types.} This means the result type of the overriding method becomes more specific, the argument types become more general. \ type_synonym 'm wf_mdecl_test = "'m prog \ cname \ 'm mdecl \ bool" definition wf_fdecl :: "'m prog \ fdecl \ bool" where "wf_fdecl P \ \(F,T). is_type P T" definition wf_mdecl :: "'m wf_mdecl_test \ 'm wf_mdecl_test" where "wf_mdecl wf_md P C \ \(M,Ts,T,mb). (\T\set Ts. is_type P T) \ is_type P T \ wf_md P C (M,Ts,T,mb)" definition wf_cdecl :: "'m wf_mdecl_test \ 'm prog \ 'm cdecl \ bool" where "wf_cdecl wf_md P \ \(C,(D,fs,ms)). (\f\set fs. wf_fdecl P f) \ distinct_fst fs \ (\m\set ms. wf_mdecl wf_md P C m) \ distinct_fst ms \ (C \ Object \ is_class P D \ \ P \ D \\<^sup>* C \ (\(M,Ts,T,m)\set ms. \D' Ts' T' m'. P \ D sees M:Ts' \ T' = m' in D' \ P \ Ts' [\] Ts \ P \ T \ T'))" definition wf_syscls :: "'m prog \ bool" where "wf_syscls P \ {Object} \ sys_xcpts \ set(map fst P)" definition wf_prog :: "'m wf_mdecl_test \ 'm prog \ bool" where "wf_prog wf_md P \ wf_syscls P \ (\c \ set P. wf_cdecl wf_md P c) \ distinct_fst P" subsection\Well-formedness lemmas\ lemma class_wf: "\class P C = Some c; wf_prog wf_md P\ \ wf_cdecl wf_md P (C,c)" -(*<*) -apply (unfold wf_prog_def class_def) -apply (fast dest: map_of_SomeD) -done -(*>*) +(*<*)by (unfold wf_prog_def class_def) (fast dest: map_of_SomeD)(*>*) lemma class_Object [simp]: "wf_prog wf_md P \ \C fs ms. class P Object = Some (C,fs,ms)" -(*<*) -apply (unfold wf_prog_def wf_syscls_def class_def) -apply (auto simp: map_of_SomeI) -done +(*<*)by (unfold wf_prog_def wf_syscls_def class_def) + (auto simp: map_of_SomeI) (*>*) lemma is_class_Object [simp]: "wf_prog wf_md P \ is_class P Object" (*<*)by (simp add: is_class_def)(*>*) (* Unused lemma is_class_supclass: assumes wf: "wf_prog wf_md P" and sub: "P \ C \\<^sup>* D" shows "is_class P C \ is_class P D" -using sub apply(induct) - apply assumption -apply(fastforce simp:wf_cdecl_def subcls1_def is_class_def - dest:class_wf[OF _ wf]) -done +(*<*) +using sub proof(induct) + case step then show ?case + by(auto simp:wf_cdecl_def is_class_def dest!:class_wf[OF _ wf] subcls1D) +qed simp +(*>*) This is NOT true because P \ NT \ Class C for any Class C lemma is_type_suptype: "\ wf_prog p P; is_type P T; P \ T \ T' \ \ is_type P T'" *) lemma is_class_xcpt: "\ C \ sys_xcpts; wf_prog wf_md P \ \ is_class P C" (*<*) - apply (simp add: wf_prog_def wf_syscls_def is_class_def class_def) - apply (fastforce intro!: map_of_SomeI) - done +by (fastforce intro!: map_of_SomeI + simp add: wf_prog_def wf_syscls_def is_class_def class_def) (*>*) lemma subcls1_wfD: - "\ P \ C \\<^sup>1 D; wf_prog wf_md P \ \ D \ C \ (D,C) \ (subcls1 P)\<^sup>+" +assumes sub1: "P \ C \\<^sup>1 D" and wf: "wf_prog wf_md P" +shows "D \ C \ (D,C) \ (subcls1 P)\<^sup>+" (*<*) -apply( frule r_into_trancl) -apply( drule subcls1D) -apply(clarify) -apply( drule (1) class_wf) -apply( unfold wf_cdecl_def) -apply(force simp add: reflcl_trancl [THEN sym] simp del: reflcl_trancl) -done +proof - + obtain fs ms where "C \ Object" and cls: "class P C = \(D, fs, ms)\" + using subcls1D[OF sub1] by clarify + then show ?thesis using wf class_wf[OF cls wf] r_into_trancl[OF sub1] + by(force simp add: wf_cdecl_def reflcl_trancl [THEN sym] + simp del: reflcl_trancl) +qed (*>*) lemma wf_cdecl_supD: "\wf_cdecl wf_md P (C,D,r); C \ Object\ \ is_class P D" (*<*)by (auto simp: wf_cdecl_def)(*>*) lemma subcls_asym: "\ wf_prog wf_md P; (C,D) \ (subcls1 P)\<^sup>+ \ \ (D,C) \ (subcls1 P)\<^sup>+" -(*<*) -apply(erule tranclE) -apply(fast dest!: subcls1_wfD ) -apply(fast dest!: subcls1_wfD intro: trancl_trans) -done -(*>*) +(*<*)by(erule tranclE; fast dest!: subcls1_wfD intro: trancl_trans)(*>*) lemma subcls_irrefl: "\ wf_prog wf_md P; (C,D) \ (subcls1 P)\<^sup>+ \ \ C \ D" -(*<*) -apply (erule trancl_trans_induct) -apply (auto dest: subcls1_wfD subcls_asym) -done -(*>*) +(*<*)by (erule trancl_trans_induct) (auto dest: subcls1_wfD subcls_asym)(*>*) lemma acyclic_subcls1: "wf_prog wf_md P \ acyclic (subcls1 P)" -(*<*) -apply (unfold acyclic_def) -apply (fast dest: subcls_irrefl) -done -(*>*) +(*<*)by (unfold acyclic_def) (fast dest: subcls_irrefl)(*>*) lemma wf_subcls1: "wf_prog wf_md P \ wf ((subcls1 P)\)" (*<*) -apply (rule finite_acyclic_wf) -apply ( subst finite_converse) -apply ( rule finite_subcls1) -apply (subst acyclic_converse) -apply (erule acyclic_subcls1) -done +proof - + assume wf: "wf_prog wf_md P" + have "finite (subcls1 P)" by(rule finite_subcls1) + then have fin': "finite ((subcls1 P)\)" by(subst finite_converse) + + from wf have "acyclic (subcls1 P)" by(rule acyclic_subcls1) + then have acyc': "acyclic ((subcls1 P)\)" by (subst acyclic_converse) + + from fin' acyc' show ?thesis by (rule finite_acyclic_wf) +qed (*>*) lemma single_valued_subcls1: "wf_prog wf_md G \ single_valued (subcls1 G)" (*<*) by(auto simp:wf_prog_def distinct_fst_def single_valued_def dest!:subcls1D) (*>*) lemma subcls_induct: "\ wf_prog wf_md P; \C. \D. (C,D) \ (subcls1 P)\<^sup>+ \ Q D \ Q C \ \ Q C" (*<*) (is "?A \ PROP ?P \ _") proof - assume p: "PROP ?P" - assume ?A thus ?thesis apply - -apply(drule wf_subcls1) -apply(drule wf_trancl) -apply(simp only: trancl_converse) -apply(erule_tac a = C in wf_induct) -apply(rule p) -apply(auto) -done + assume ?A then have wf: "wf_prog wf_md P" by assumption + have wf':"wf (((subcls1 P)\<^sup>+)\)" using wf_trancl[OF wf_subcls1[OF wf]] + by(simp only: trancl_converse) + show ?thesis using wf_induct[where a = C and P = Q, OF wf' p] by simp qed (*>*) lemma subcls1_induct_aux: - "\ is_class P C; wf_prog wf_md P; Q Object; - \C D fs ms. +assumes "is_class P C" and wf: "wf_prog wf_md P" and QObj: "Q Object" +shows + "\ \C D fs ms. \ C \ Object; is_class P C; class P C = Some (D,fs,ms) \ wf_cdecl wf_md P (C,D,fs,ms) \ P \ C \\<^sup>1 D \ is_class P D \ Q D\ \ Q C \ \ Q C" (*<*) - (is "?A \ ?B \ ?C \ PROP ?P \ _") + (is "PROP ?P \ _") proof - assume p: "PROP ?P" - assume ?A ?B ?C thus ?thesis apply - -apply(unfold is_class_def) -apply( rule impE) -prefer 2 -apply( assumption) -prefer 2 -apply( assumption) -apply( erule thin_rl) -apply( rule subcls_induct) -apply( assumption) -apply( rule impI) -apply( case_tac "C = Object") -apply( fast) -apply safe -apply( frule (1) class_wf) -apply( frule (1) wf_cdecl_supD) - -apply( subgoal_tac "P \ C \\<^sup>1 a") -apply( erule_tac [2] subcls1I) -apply( rule p) -apply (unfold is_class_def) -apply auto -done + have "class P C \ None \ Q C" + proof(induct rule: subcls_induct[OF wf]) + case (1 C) + have "class P C \ None \ Q C" + proof(cases "C = Object") + case True + then show ?thesis using QObj by fast + next + case False + assume nNone: "class P C \ None" + then have is_cls: "is_class P C" by(simp add: is_class_def) + obtain D fs ms where cls: "class P C = \(D, fs, ms)\" using nNone by safe + also have wfC: "wf_cdecl wf_md P (C, D, fs, ms)" by(rule class_wf[OF cls wf]) + moreover have D: "is_class P D" by(rule wf_cdecl_supD[OF wfC False]) + moreover have "P \ C \\<^sup>1 D" by(rule subcls1I[OF cls False]) + moreover have "class P D \ None" using D by(simp add: is_class_def) + ultimately show ?thesis using 1 by (auto intro: p[OF False is_cls]) + qed + then show "class P C \ None \ Q C" by simp + qed + thus ?thesis using assms by(unfold is_class_def) simp qed (*>*) (* FIXME can't we prove this one directly?? *) lemma subcls1_induct [consumes 2, case_names Object Subcls]: "\ wf_prog wf_md P; is_class P C; Q Object; \C D. \C \ Object; P \ C \\<^sup>1 D; is_class P D; Q D\ \ Q C \ \ Q C" -(*<*) - apply (erule subcls1_induct_aux, assumption, assumption) - apply blast - done -(*>*) +(*<*)by (erule (2) subcls1_induct_aux) blast(*>*) lemma subcls_C_Object: - "\ is_class P C; wf_prog wf_md P \ \ P \ C \\<^sup>* Object" +assumes "class": "is_class P C" and wf: "wf_prog wf_md P" +shows "P \ C \\<^sup>* Object" (*<*) -apply(erule (1) subcls1_induct) - apply( fast) -apply(erule (1) converse_rtrancl_into_rtrancl) -done +using wf "class" +proof(induct rule: subcls1_induct) + case Subcls + then show ?case by(simp add: converse_rtrancl_into_rtrancl) +qed fast (*>*) lemma is_type_pTs: assumes "wf_prog wf_md P" and "(C,S,fs,ms) \ set P" and "(M,Ts,T,m) \ set ms" shows "set Ts \ types P" (*<*) proof from assms have "wf_mdecl wf_md P C (M,Ts,T,m)" by (unfold wf_prog_def wf_cdecl_def) auto hence "\t \ set Ts. is_type P t" by (unfold wf_mdecl_def) auto moreover fix t assume "t \ set Ts" ultimately have "is_type P t" by blast thus "t \ types P" .. qed (*>*) subsection\Well-formedness and method lookup\ lemma sees_wf_mdecl: - "\ wf_prog wf_md P; P \ C sees M:Ts\T = m in D \ \ wf_mdecl wf_md P D (M,Ts,T,m)" +assumes wf: "wf_prog wf_md P" and sees: "P \ C sees M:Ts\T = m in D" +shows "wf_mdecl wf_md P D (M,Ts,T,m)" (*<*) -apply(drule visible_method_exists) -apply(fastforce simp:wf_cdecl_def dest!:class_wf dest:map_of_SomeD) -done +using wf visible_method_exists[OF sees] +by(fastforce simp:wf_cdecl_def dest!:class_wf dest:map_of_SomeD) (*>*) lemma sees_method_mono [rule_format (no_asm)]: - "\ P \ C' \\<^sup>* C; wf_prog wf_md P \ \ - \D Ts T m. P \ C sees M:Ts\T = m in D \ +assumes sub: "P \ C' \\<^sup>* C" and wf: "wf_prog wf_md P" +shows "\D Ts T m. P \ C sees M:Ts\T = m in D \ (\D' Ts' T' m'. P \ C' sees M:Ts'\T' = m' in D' \ P \ Ts [\] Ts' \ P \ T' \ T)" (*<*) -apply( drule rtranclD) -apply( erule disjE) -apply( fastforce) -apply( erule conjE) -apply( erule trancl_trans_induct) -prefer 2 -apply( clarify) -apply( drule spec, drule spec, drule spec, drule spec, erule (1) impE) -apply clarify -apply( fast elim: widen_trans widens_trans) -apply( clarify) -apply( drule subcls1D) -apply( clarify) -apply(clarsimp simp:Method_def) -apply(frule (2) sees_methods_rec) -apply(rule refl) -apply(case_tac "map_of ms M") -apply(rule_tac x = D in exI) -apply(rule_tac x = Ts in exI) -apply(rule_tac x = T in exI) -apply simp -apply(rule_tac x = m in exI) -apply(fastforce simp add:map_add_def split:option.split) -apply clarsimp -apply(rename_tac Ts' T' m') -apply( drule (1) class_wf) -apply( unfold wf_cdecl_def Method_def) -apply( frule map_of_SomeD) -apply auto -apply(drule (1) bspec, simp) -apply(erule_tac x=D in allE, erule_tac x=Ts in allE, erule_tac x=T in allE) -apply(fastforce simp:map_add_def split:option.split) -done + (is "\D Ts T m. ?P C D Ts T m \ ?Q C' D Ts T m") +proof(rule disjE[OF rtranclD[OF sub]]) + assume "C' = C" + then show ?thesis using assms by fastforce +next + assume "C' \ C \ (C', C) \ (subcls1 P)\<^sup>+" + then have neq: "C' \ C" and subcls1: "(C', C) \ (subcls1 P)\<^sup>+" by simp+ + show ?thesis proof(induct rule: trancl_trans_induct[OF subcls1]) + case (2 x y z) + then have zy: "\D Ts T m. ?P z D Ts T m \ ?Q y D Ts T m" by blast + have "\D Ts T m. ?P z D Ts T m \ ?Q x D Ts T m" + proof - + fix D Ts T m assume P: "?P z D Ts T m" + then show "?Q x D Ts T m" using zy[OF P] 2(2) + by(fast elim: widen_trans widens_trans) + qed + then show ?case by blast + next + case (1 x y) + have "\D Ts T m. ?P y D Ts T m \ ?Q x D Ts T m" + proof - + fix D Ts T m assume P: "?P y D Ts T m" + then obtain Mm where sees: "P \ y sees_methods Mm" and + M: "Mm M = \((Ts, T, m), D)\" + by(clarsimp simp:Method_def) + obtain fs ms where nObj: "x \ Object" and + cls: "class P x = \(y, fs, ms)\" + using subcls1D[OF 1] by clarsimp + have x_meth: "P \ x sees_methods Mm ++ (map_option (\m. (m, x)) \ map_of ms)" + using sees_methods_rec[OF cls nObj sees] by simp + show "?Q x D Ts T m" proof(cases "map_of ms M") + case None + then have "\m'. P \ x sees M : Ts\T = m' in D" using M x_meth + by(fastforce simp add:Method_def map_add_def split:option.split) + then show ?thesis by auto + next + case (Some a) + then obtain Ts' T' m' where a: "a = (Ts',T',m')" by(cases a) + then have "(\m' Mm. P \ y sees_methods Mm \ Mm M = \((Ts, T, m'), D)\) + \ P \ Ts [\] Ts' \ P \ T' \ T" + using nObj class_wf[OF cls wf] map_of_SomeD[OF Some] + by(clarsimp simp: wf_cdecl_def Method_def) fast + then show ?thesis using Some a sees M x_meth + by(fastforce simp:Method_def map_add_def split:option.split) + qed + qed + then show ?case by simp + qed +qed (*>*) lemma sees_method_mono2: "\ P \ C' \\<^sup>* C; wf_prog wf_md P; P \ C sees M:Ts\T = m in D; P \ C' sees M:Ts'\T' = m' in D' \ \ P \ Ts [\] Ts' \ P \ T' \ T" (*<*)by(blast dest:sees_method_mono sees_method_fun)(*>*) lemma mdecls_visible: assumes wf: "wf_prog wf_md P" and "class": "is_class P C" shows "\D fs ms. class P C = Some(D,fs,ms) \ \Mm. P \ C sees_methods Mm \ (\(M,Ts,T,m) \ set ms. Mm M = Some((Ts,T,m),C))" (*<*) using wf "class" proof (induct rule:subcls1_induct) case Object with wf have "distinct_fst ms" by (unfold class_def wf_prog_def wf_cdecl_def) (fastforce dest:map_of_SomeD) with Object show ?case by(fastforce intro!: sees_methods_Object map_of_SomeI) next case Subcls with wf have "distinct_fst ms" by (unfold class_def wf_prog_def wf_cdecl_def) (fastforce dest:map_of_SomeD) with Subcls show ?case by(fastforce elim:sees_methods_rec dest:subcls1D map_of_SomeI simp:is_class_def) qed (*>*) lemma mdecl_visible: assumes wf: "wf_prog wf_md P" and C: "(C,S,fs,ms) \ set P" and m: "(M,Ts,T,m) \ set ms" shows "P \ C sees M:Ts\T = m in C" (*<*) proof - from wf C have "class": "class P C = Some (S,fs,ms)" by (auto simp add: wf_prog_def class_def is_class_def intro: map_of_SomeI) from "class" have "is_class P C" by(auto simp:is_class_def) with assms "class" show ?thesis by(bestsimp simp:Method_def dest:mdecls_visible) qed (*>*) lemma Call_lemma: - "\ P \ C sees M:Ts\T = m in D; P \ C' \\<^sup>* C; wf_prog wf_md P \ - \ \D' Ts' T' m'. +assumes sees: "P \ C sees M:Ts\T = m in D" and sub: "P \ C' \\<^sup>* C" and wf: "wf_prog wf_md P" +shows "\D' Ts' T' m'. P \ C' sees M:Ts'\T' = m' in D' \ P \ Ts [\] Ts' \ P \ T' \ T \ P \ C' \\<^sup>* D' \ is_type P T' \ (\T\set Ts'. is_type P T) \ wf_md P D' (M,Ts',T',m')" (*<*) -apply(frule (2) sees_method_mono) -apply(fastforce intro:sees_method_decl_above dest:sees_wf_mdecl - simp: wf_mdecl_def) -done +using assms sees_method_mono[OF sub wf sees] +by(fastforce intro:sees_method_decl_above dest:sees_wf_mdecl + simp: wf_mdecl_def) (*>*) lemma wf_prog_lift: assumes wf: "wf_prog (\P C bd. A P C bd) P" and rule: "\wf_md C M Ts C T m bd. wf_prog wf_md P \ P \ C sees M:Ts\T = m in C \ set Ts \ types P \ bd = (M,Ts,T,m) \ A P C bd \ B P C bd" shows "wf_prog (\P C bd. B P C bd) P" (*<*) proof - - from wf show ?thesis - apply (unfold wf_prog_def wf_cdecl_def) - apply clarsimp - apply (drule bspec, assumption) - apply (unfold wf_mdecl_def) - apply clarsimp - apply (drule bspec, assumption) - apply clarsimp - apply (frule mdecl_visible [OF wf], assumption+) - apply (frule is_type_pTs [OF wf], assumption+) - apply (drule rule [OF wf], assumption+) - apply auto - done + have "\c. c\set P \ wf_cdecl A P c \ wf_cdecl B P c" + proof - + fix c assume "c\set P" and "wf_cdecl A P c" + then show "wf_cdecl B P c" + using rule[OF wf mdecl_visible[OF wf] is_type_pTs[OF wf]] + by (auto simp: wf_cdecl_def wf_mdecl_def) + qed + then show ?thesis using wf by (clarsimp simp: wf_prog_def) qed (*>*) subsection\Well-formedness and field lookup\ lemma wf_Fields_Ex: - "\ wf_prog wf_md P; is_class P C \ \ \FDTs. P \ C has_fields FDTs" +assumes wf: "wf_prog wf_md P" and "is_class P C" +shows "\FDTs. P \ C has_fields FDTs" (*<*) -apply(frule class_Object) -apply(erule (1) subcls1_induct) - apply(blast intro:has_fields_Object) -apply(blast intro:has_fields_rec dest:subcls1D) -done +using assms proof(induct rule:subcls1_induct) + case Object + then show ?case using class_Object[OF wf] + by(blast intro:has_fields_Object) +next + case Subcls + then show ?case by(blast intro:has_fields_rec dest:subcls1D) +qed (*>*) lemma has_fields_types: "\ P \ C has_fields FDTs; (FD,T) \ set FDTs; wf_prog wf_md P \ \ is_type P T" (*<*) -apply(induct rule:Fields.induct) - apply(fastforce dest!: class_wf simp: wf_cdecl_def wf_fdecl_def) -apply(fastforce dest!: class_wf simp: wf_cdecl_def wf_fdecl_def) -done +proof(induct rule:Fields.induct) +qed(fastforce dest!: class_wf simp: wf_cdecl_def wf_fdecl_def)+ (*>*) lemma sees_field_is_type: "\ P \ C sees F:T in D; wf_prog wf_md P \ \ is_type P T" (*<*) by(fastforce simp: sees_field_def elim: has_fields_types map_of_SomeD[OF map_of_remap_SomeD]) (*>*) lemma wf_syscls: "set SystemClasses \ set P \ wf_syscls P" (*<*) -apply (simp add: image_def SystemClasses_def wf_syscls_def sys_xcpts_def +by (force simp add: image_def SystemClasses_def wf_syscls_def sys_xcpts_def ObjectC_def NullPointerC_def ClassCastC_def OutOfMemoryC_def) - apply force -done (*>*) end diff --git a/thys/Jinja/Compiler/Correctness1.thy b/thys/Jinja/Compiler/Correctness1.thy --- a/thys/Jinja/Compiler/Correctness1.thy +++ b/thys/Jinja/Compiler/Correctness1.thy @@ -1,760 +1,761 @@ (* Title: Jinja/Compiler/Correctness1.thy Author: Tobias Nipkow Copyright TUM 2003 *) section \Correctness of Stage 1\ theory Correctness1 imports J1WellForm Compiler1 begin subsection\Correctness of program compilation\ primrec unmod :: "expr\<^sub>1 \ nat \ bool" and unmods :: "expr\<^sub>1 list \ nat \ bool" where "unmod (new C) i = True" | "unmod (Cast C e) i = unmod e i" | "unmod (Val v) i = True" | "unmod (e\<^sub>1 \bop\ e\<^sub>2) i = (unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (Var i) j = True" | "unmod (i:=e) j = (i \ j \ unmod e j)" | "unmod (e\F{D}) i = unmod e i" | "unmod (e\<^sub>1\F{D}:=e\<^sub>2) i = (unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (e\M(es)) i = (unmod e i \ unmods es i)" | "unmod {j:T; e} i = unmod e i" | "unmod (e\<^sub>1;;e\<^sub>2) i = (unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (if (e) e\<^sub>1 else e\<^sub>2) i = (unmod e i \ unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (while (e) c) i = (unmod e i \ unmod c i)" | "unmod (throw e) i = unmod e i" | "unmod (try e\<^sub>1 catch(C i) e\<^sub>2) j = (unmod e\<^sub>1 j \ (if i=j then False else unmod e\<^sub>2 j))" | "unmods ([]) i = True" | "unmods (e#es) i = (unmod e i \ unmods es i)" lemma hidden_unmod: "\Vs. hidden Vs i \ unmod (compE\<^sub>1 Vs e) i" and "\Vs. hidden Vs i \ unmods (compEs\<^sub>1 Vs es) i" (*<*) -apply(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) -apply (simp_all add:hidden_inacc) -apply(auto simp add:hidden_def) -done +proof(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) + case TryCatch + then show ?case by(simp add:hidden_inacc)(auto simp add:hidden_def) +qed (simp_all add:hidden_inacc) (*>*) lemma eval\<^sub>1_preserves_unmod: "\ P \\<^sub>1 \e,(h,ls)\ \ \e',(h',ls')\; unmod e i; i < size ls \ \ ls ! i = ls' ! i" and "\ P \\<^sub>1 \es,(h,ls)\ [\] \es',(h',ls')\; unmods es i; i < size ls \ \ ls ! i = ls' ! i" (*<*) -apply(induct rule:eval\<^sub>1_evals\<^sub>1_inducts) -apply(auto dest!:eval\<^sub>1_preserves_len split:if_split_asm) -done +proof(induct rule:eval\<^sub>1_evals\<^sub>1_inducts) +qed(auto dest!:eval\<^sub>1_preserves_len evals\<^sub>1_preserves_len split:if_split_asm) (*>*) lemma LAss_lem: "\x \ set xs; size xs \ size ys \ \ m\<^sub>1 \\<^sub>m m\<^sub>2(xs[\]ys) \ m\<^sub>1(x\y) \\<^sub>m m\<^sub>2(xs[\]ys[last_index xs x := y])" (*<*) by(simp add:map_le_def fun_upds_apply eq_sym_conv) (*>*) lemma Block_lem: fixes l :: "'a \ 'b" assumes 0: "l \\<^sub>m [Vs [\] ls]" and 1: "l' \\<^sub>m [Vs [\] ls', V\v]" and hidden: "V \ set Vs \ ls ! last_index Vs V = ls' ! last_index Vs V" and size: "size ls = size ls'" "size Vs < size ls'" shows "l'(V := l V) \\<^sub>m [Vs [\] ls']" (*<*) proof - have "l'(V := l V) \\<^sub>m [Vs [\] ls', V\v](V := l V)" using 1 by(rule map_le_upd) also have "\ = [Vs [\] ls'](V := l V)" by simp also have "\ \\<^sub>m [Vs [\] ls']" proof (cases "l V") case None thus ?thesis by simp next case (Some w) hence "[Vs [\] ls] V = Some w" using 0 by(force simp add: map_le_def split:if_splits) hence VinVs: "V \ set Vs" and w: "w = ls ! last_index Vs V" using size by(auto simp add:fun_upds_apply split:if_splits) hence "w = ls' ! last_index Vs V" using hidden[OF VinVs] by simp hence "[Vs [\] ls'](V := l V) = [Vs [\] ls']" using Some size VinVs by(simp add: map_upds_upd_conv_last_index) thus ?thesis by simp qed finally show ?thesis . qed (*>*) (*<*) declare fun_upd_apply[simp del] (*>*) text\\noindent The main theorem:\ theorem assumes wf: "wwf_J_prog P" shows eval\<^sub>1_eval: "P \ \e,(h,l)\ \ \e',(h',l')\ \ (\Vs ls. \ fv e \ set Vs; l \\<^sub>m [Vs[\]ls]; size Vs + max_vars e \ size ls \ \ \ls'. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs e,(h,ls)\ \ \fin\<^sub>1 e',(h',ls')\ \ l' \\<^sub>m [Vs[\]ls'])" (*<*) (is "_ \ (\Vs ls. PROP ?P e h l e' h' l' Vs ls)" is "_ \ (\Vs ls. \ _; _; _ \ \ \ls'. ?Post e h l e' h' l' Vs ls ls')") (*>*) and evals\<^sub>1_evals: "P \ \es,(h,l)\ [\] \es',(h',l')\ \ (\Vs ls. \ fvs es \ set Vs; l \\<^sub>m [Vs[\]ls]; size Vs + max_varss es \ size ls \ \ \ls'. compP\<^sub>1 P \\<^sub>1 \compEs\<^sub>1 Vs es,(h,ls)\ [\] \compEs\<^sub>1 Vs es',(h',ls')\ \ l' \\<^sub>m [Vs[\]ls'])" (*<*) (is "_ \ (\Vs ls. PROP ?Ps es h l es' h' l' Vs ls)" is "_ \ (\Vs ls. \ _; _; _\ \ \ls'. ?Posts es h l es' h' l' Vs ls ls')") proof (induct rule:eval_evals_inducts) case Nil thus ?case by(fastforce intro!:Nil\<^sub>1) next case (Cons e h l v h' l' es es' h\<^sub>2 l\<^sub>2) have "PROP ?P e h l (Val v) h' l' Vs ls" by fact with Cons.prems obtain ls' where 1: "?Post e h l (Val v) h' l' Vs ls ls'" "size ls = size ls'" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h' l' es' h\<^sub>2 l\<^sub>2 Vs ls'" by fact with 1 Cons.prems obtain ls\<^sub>2 where 2: "?Posts es h' l' es' h\<^sub>2 l\<^sub>2 Vs ls' ls\<^sub>2" by(auto) from 1 2 Cons show ?case by(auto intro!:Cons\<^sub>1) next case ConsThrow thus ?case by(fastforce intro!:ConsThrow\<^sub>1 dest: eval_final) next case (Block e h l V e' h' l' T) let ?Vs = "Vs @ [V]" have IH: "\fv e \ set ?Vs; l(V := None) \\<^sub>m [?Vs [\] ls]; size ?Vs + max_vars e \ size ls\ \ \ls'. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e,(h,ls)\ \ \fin\<^sub>1 e',(h', ls')\ \ l' \\<^sub>m [?Vs [\] ls']" and fv: "fv {V:T; e} \ set Vs" and rel: "l \\<^sub>m [Vs [\] ls]" and len: "length Vs + max_vars {V:T; e} \ length ls" by fact+ have len': "length Vs < length ls" using len by auto have "fv e \ set ?Vs" using fv by auto moreover have "l(V := None) \\<^sub>m [?Vs [\] ls]" using rel len' by simp moreover have "size ?Vs + max_vars e \ size ls" using len by simp ultimately obtain ls' where 1: "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e,(h,ls)\ \ \fin\<^sub>1 e',(h',ls')\" and rel': "l' \\<^sub>m [?Vs [\] ls']" using IH by blast have [simp]: "length ls = length ls'" by(rule eval\<^sub>1_preserves_len[OF 1]) show "\ls'. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs {V:T; e},(h,ls)\ \ \fin\<^sub>1 e',(h',ls')\ \ l'(V := l V) \\<^sub>m [Vs [\] ls']" (is "\ls'. ?R ls'") proof show "?R ls'" proof show "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs {V:T; e},(h,ls)\ \ \fin\<^sub>1 e',(h',ls')\" using 1 by(simp add:Block\<^sub>1) next show "l'(V := l V) \\<^sub>m [Vs [\] ls']" proof - have "l' \\<^sub>m [Vs [\] ls', V \ ls' ! length Vs]" using len' rel' by simp moreover { assume VinVs: "V \ set Vs" hence "hidden (Vs @ [V]) (last_index Vs V)" by(rule hidden_last_index) hence "unmod (compE\<^sub>1 (Vs @ [V]) e) (last_index Vs V)" by(rule hidden_unmod) moreover have "last_index Vs V < length ls" using len' VinVs by simp ultimately have "ls ! last_index Vs V = ls' ! last_index Vs V" by(rule eval\<^sub>1_preserves_unmod[OF 1]) } ultimately show ?thesis using Block_lem[OF rel] len' by auto qed qed qed next case (TryThrow e' h l a h' l' D fs C V e\<^sub>2) have "PROP ?P e' h l (Throw a) h' l' Vs ls" by fact with TryThrow.prems obtain ls' where 1: "?Post e' h l (Throw a) h' l' Vs ls ls'" by(auto) show ?case using 1 TryThrow.hyps by(auto intro!:eval\<^sub>1_evals\<^sub>1.TryThrow\<^sub>1) next case (TryCatch e\<^sub>1 h l a h\<^sub>1 l\<^sub>1 D fs C e\<^sub>2 V e' h\<^sub>2 l\<^sub>2) let ?e = "try e\<^sub>1 catch(C V) e\<^sub>2" have IH\<^sub>1: "\fv e\<^sub>1 \ set Vs; l \\<^sub>m [Vs [\] ls]; size Vs + max_vars e\<^sub>1 \ length ls\ \ \ls\<^sub>1. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs e\<^sub>1,(h,ls)\ \ \fin\<^sub>1 (Throw a),(h\<^sub>1,ls\<^sub>1)\ \ l\<^sub>1 \\<^sub>m [Vs [\] ls\<^sub>1]" and fv: "fv ?e \ set Vs" and rel: "l \\<^sub>m [Vs [\] ls]" and len: "length Vs + max_vars ?e \ length ls" by fact+ have "fv e\<^sub>1 \ set Vs" using fv by auto moreover have "length Vs + max_vars e\<^sub>1 \ length ls" using len by(auto) ultimately obtain ls\<^sub>1 where 1: "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs e\<^sub>1,(h,ls)\ \ \Throw a,(h\<^sub>1,ls\<^sub>1)\" and rel\<^sub>1: "l\<^sub>1 \\<^sub>m [Vs [\] ls\<^sub>1]" using IH\<^sub>1 rel by fastforce from 1 have [simp]: "size ls = size ls\<^sub>1" by(rule eval\<^sub>1_preserves_len) let ?Vs = "Vs @ [V]" let ?ls = "(ls\<^sub>1[size Vs:=Addr a])" have IH\<^sub>2: "\fv e\<^sub>2 \ set ?Vs; l\<^sub>1(V \ Addr a) \\<^sub>m [?Vs [\] ?ls]; length ?Vs + max_vars e\<^sub>2 \ length ?ls\ \ \ls\<^sub>2. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e\<^sub>2,(h\<^sub>1,?ls)\ \ \fin\<^sub>1 e',(h\<^sub>2, ls\<^sub>2)\ \ l\<^sub>2 \\<^sub>m [?Vs [\] ls\<^sub>2]" by fact have len\<^sub>1: "size Vs < size ls\<^sub>1" using len by(auto) have "fv e\<^sub>2 \ set ?Vs" using fv by auto moreover have "l\<^sub>1(V \ Addr a) \\<^sub>m [?Vs [\] ?ls]" using rel\<^sub>1 len\<^sub>1 by simp moreover have "length ?Vs + max_vars e\<^sub>2 \ length ?ls" using len by(auto) ultimately obtain ls\<^sub>2 where 2: "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e\<^sub>2,(h\<^sub>1,?ls)\ \ \fin\<^sub>1 e',(h\<^sub>2, ls\<^sub>2)\" and rel\<^sub>2: "l\<^sub>2 \\<^sub>m [?Vs [\] ls\<^sub>2]" using IH\<^sub>2 by blast from 2 have [simp]: "size ls\<^sub>1 = size ls\<^sub>2" by(fastforce dest: eval\<^sub>1_preserves_len) show "\ls\<^sub>2. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs ?e,(h,ls)\ \ \fin\<^sub>1 e',(h\<^sub>2,ls\<^sub>2)\ \ l\<^sub>2(V := l\<^sub>1 V) \\<^sub>m [Vs [\] ls\<^sub>2]" (is "\ls\<^sub>2. ?R ls\<^sub>2") proof show "?R ls\<^sub>2" proof have hp: "h\<^sub>1 a = Some (D, fs)" by fact have "P \ D \\<^sup>* C" by fact hence caught: "compP\<^sub>1 P \ D \\<^sup>* C" by simp from TryCatch\<^sub>1[OF 1 _ caught len\<^sub>1 2, OF hp] show "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs ?e,(h,ls)\ \ \fin\<^sub>1 e',(h\<^sub>2,ls\<^sub>2)\" by simp next show "l\<^sub>2(V := l\<^sub>1 V) \\<^sub>m [Vs [\] ls\<^sub>2]" proof - have "l\<^sub>2 \\<^sub>m [Vs [\] ls\<^sub>2, V \ ls\<^sub>2 ! length Vs]" using len\<^sub>1 rel\<^sub>2 by simp moreover { assume VinVs: "V \ set Vs" hence "hidden (Vs @ [V]) (last_index Vs V)" by(rule hidden_last_index) hence "unmod (compE\<^sub>1 (Vs @ [V]) e\<^sub>2) (last_index Vs V)" by(rule hidden_unmod) moreover have "last_index Vs V < length ?ls" using len\<^sub>1 VinVs by simp ultimately have "?ls ! last_index Vs V = ls\<^sub>2 ! last_index Vs V" by(rule eval\<^sub>1_preserves_unmod[OF 2]) moreover have "last_index Vs V < size Vs" using VinVs by simp ultimately have "ls\<^sub>1 ! last_index Vs V = ls\<^sub>2 ! last_index Vs V" using len\<^sub>1 by(simp del:size_last_index_conv) } ultimately show ?thesis using Block_lem[OF rel\<^sub>1] len\<^sub>1 by simp qed qed qed next case Try thus ?case by(fastforce intro!:Try\<^sub>1) next case Throw thus ?case by(fastforce intro!:Throw\<^sub>1) next case ThrowNull thus ?case by(fastforce intro!:ThrowNull\<^sub>1) next case ThrowThrow thus ?case by(fastforce intro!:ThrowThrow\<^sub>1) next case (CondT e h l h\<^sub>1 l\<^sub>1 e\<^sub>1 e' h\<^sub>2 l\<^sub>2 e\<^sub>2) have "PROP ?P e h l true h\<^sub>1 l\<^sub>1 Vs ls" by fact with CondT.prems obtain ls\<^sub>1 where 1: "?Post e h l true h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 e' h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 CondT.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 e' h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 show ?case by(auto intro!:CondT\<^sub>1) next case (CondF e h l h\<^sub>1 l\<^sub>1 e\<^sub>2 e' h\<^sub>2 l\<^sub>2 e\<^sub>1 Vs ls) have "PROP ?P e h l false h\<^sub>1 l\<^sub>1 Vs ls" by fact with CondF.prems obtain ls\<^sub>1 where 1: "?Post e h l false h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 e' h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 CondF.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 e' h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 show ?case by(auto intro!:CondF\<^sub>1) next case CondThrow thus ?case by(fastforce intro!:CondThrow\<^sub>1) next case (Seq e h l v h\<^sub>1 l\<^sub>1 e\<^sub>1 e' h\<^sub>2 l\<^sub>2) have "PROP ?P e h l (Val v) h\<^sub>1 l\<^sub>1 Vs ls" by fact with Seq.prems obtain ls\<^sub>1 where 1: "?Post e h l (Val v) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 e' h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 Seq.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 e' h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 Seq show ?case by(auto intro!:Seq\<^sub>1) next case SeqThrow thus ?case by(fastforce intro!:SeqThrow\<^sub>1) next case WhileF thus ?case by(fastforce intro!:eval\<^sub>1_evals\<^sub>1.intros) next case (WhileT e h l h\<^sub>1 l\<^sub>1 c v h\<^sub>2 l\<^sub>2 e' h\<^sub>3 l\<^sub>3) have "PROP ?P e h l true h\<^sub>1 l\<^sub>1 Vs ls" by fact with WhileT.prems obtain ls\<^sub>1 where 1: "?Post e h l true h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P c h\<^sub>1 l\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 WhileT.prems obtain ls\<^sub>2 where 2: "?Post c h\<^sub>1 l\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" "size ls\<^sub>1 = size ls\<^sub>2" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P (While (e) c) h\<^sub>2 l\<^sub>2 e' h\<^sub>3 l\<^sub>3 Vs ls\<^sub>2" by fact with 1 2 WhileT.prems obtain ls\<^sub>3 where 3: "?Post (While (e) c) h\<^sub>2 l\<^sub>2 e' h\<^sub>3 l\<^sub>3 Vs ls\<^sub>2 ls\<^sub>3" by(auto) from 1 2 3 show ?case by(auto intro!:WhileT\<^sub>1) next case (WhileBodyThrow e h l h\<^sub>1 l\<^sub>1 c e' h\<^sub>2 l\<^sub>2) have "PROP ?P e h l true h\<^sub>1 l\<^sub>1 Vs ls" by fact with WhileBodyThrow.prems obtain ls\<^sub>1 where 1: "?Post e h l true h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P c h\<^sub>1 l\<^sub>1 (throw e') h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 WhileBodyThrow.prems obtain ls\<^sub>2 where 2: "?Post c h\<^sub>1 l\<^sub>1 (throw e') h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by auto from 1 2 show ?case by(auto intro!:WhileBodyThrow\<^sub>1) next case WhileCondThrow thus ?case by(fastforce intro!:WhileCondThrow\<^sub>1) next case New thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case NewFail thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case Cast thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case CastNull thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case CastThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (CastFail e h l a h\<^sub>1 l\<^sub>1 D fs C) have "PROP ?P e h l (addr a) h\<^sub>1 l\<^sub>1 Vs ls" by fact with CastFail.prems obtain ls\<^sub>1 where 1: "?Post e h l (addr a) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" by auto show ?case using 1 CastFail.hyps by(auto intro!:CastFail\<^sub>1[where D=D]) next case Val thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (BinOp e h l v\<^sub>1 h\<^sub>1 l\<^sub>1 e\<^sub>1 v\<^sub>2 h\<^sub>2 l\<^sub>2 bop v) have "PROP ?P e h l (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 Vs ls" by fact with BinOp.prems obtain ls\<^sub>1 where 1: "?Post e h l (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 (Val v\<^sub>2) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 BinOp.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 (Val v\<^sub>2) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 BinOp show ?case by(auto intro!:BinOp\<^sub>1) next case (BinOpThrow2 e\<^sub>0 h l v\<^sub>1 h\<^sub>1 l\<^sub>1 e\<^sub>1 e h\<^sub>2 l\<^sub>2 bop) have "PROP ?P e\<^sub>0 h l (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 Vs ls" by fact with BinOpThrow2.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>0 h l (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 BinOpThrow2.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 BinOpThrow2 show ?case by(auto intro!:BinOpThrow\<^sub>2\<^sub>1) next case BinOpThrow1 thus ?case by(fastforce intro!:eval\<^sub>1_evals\<^sub>1.intros) next case Var thus ?case by(force intro!:Var\<^sub>1 simp add: map_le_def fun_upds_apply) next case LAss thus ?case by(fastforce simp add: LAss_lem intro:eval\<^sub>1_evals\<^sub>1.intros dest:eval\<^sub>1_preserves_len) next case LAssThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case FAcc thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case FAccNull thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case FAccThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (FAss e\<^sub>1 h l a h\<^sub>1 l\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 C fs fs' F D h\<^sub>2') have "PROP ?P e\<^sub>1 h l (addr a) h\<^sub>1 l\<^sub>1 Vs ls" by fact with FAss.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l (addr a) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAss.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAss show ?case by(auto intro!:FAss\<^sub>1) next case (FAssNull e\<^sub>1 h l h\<^sub>1 l\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 F D) have "PROP ?P e\<^sub>1 h l null h\<^sub>1 l\<^sub>1 Vs ls" by fact with FAssNull.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l null h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAssNull.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAssNull show ?case by(auto intro!:FAssNull\<^sub>1) next case FAssThrow1 thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (FAssThrow2 e\<^sub>1 h l v h\<^sub>1 l\<^sub>1 e\<^sub>2 e h\<^sub>2 l\<^sub>2 F D) have "PROP ?P e\<^sub>1 h l (Val v) h\<^sub>1 l\<^sub>1 Vs ls" by fact with FAssThrow2.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l (Val v) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAssThrow2.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAssThrow2 show ?case by(auto intro!:FAssThrow\<^sub>2\<^sub>1) next case (CallNull e h l h\<^sub>1 l\<^sub>1 es vs h\<^sub>2 l\<^sub>2 M) have "PROP ?P e h l null h\<^sub>1 l\<^sub>1 Vs ls" by fact with CallNull.prems obtain ls\<^sub>1 where 1: "?Post e h l null h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h\<^sub>1 l\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 CallNull.prems obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 CallNull show ?case by (auto simp add: comp_def elim!: CallNull\<^sub>1) next case CallObjThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (CallParamsThrow e h l v h\<^sub>1 l\<^sub>1 es vs ex es' h\<^sub>2 l\<^sub>2 M) have "PROP ?P e h l (Val v) h\<^sub>1 l\<^sub>1 Vs ls" by fact with CallParamsThrow.prems obtain ls\<^sub>1 where 1: "?Post e h l (Val v) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h\<^sub>1 l\<^sub>1 (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 CallParamsThrow.prems obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 CallParamsThrow show ?case by (auto simp add: comp_def elim!: CallParamsThrow\<^sub>1 dest!:evals_final) next case (Call e h l a h\<^sub>1 l\<^sub>1 es vs h\<^sub>2 l\<^sub>2 C fs M Ts T pns body D l\<^sub>2' b' h\<^sub>3 l\<^sub>3) have "PROP ?P e h l (addr a) h\<^sub>1 l\<^sub>1 Vs ls" by fact with Call.prems obtain ls\<^sub>1 where 1: "?Post e h l (addr a) h\<^sub>1 l\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h\<^sub>1 l\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1" by fact with 1 Call.prems obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" "size ls\<^sub>1 = size ls\<^sub>2" by(auto intro!:evals\<^sub>1_preserves_len) let ?Vs = "this#pns" let ?ls = "Addr a # vs @ replicate (max_vars body) undefined" have mdecl: "P \ C sees M: Ts\T = (pns, body) in D" by fact have fv_body: "fv body \ set ?Vs" and wf_size: "size Ts = size pns" using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def) have mdecl\<^sub>1: "compP\<^sub>1 P \ C sees M: Ts\T = (compE\<^sub>1 ?Vs body) in D" using sees_method_compP[OF mdecl, of "\(pns,e). compE\<^sub>1 (this#pns) e"] by(simp) have [simp]: "l\<^sub>2' = [this \ Addr a, pns [\] vs]" by fact have Call_size: "size vs = size pns" by fact have "PROP ?P body h\<^sub>2 l\<^sub>2' b' h\<^sub>3 l\<^sub>3 ?Vs ?ls" by fact with 1 2 fv_body Call_size Call.prems obtain ls\<^sub>3 where 3: "?Post body h\<^sub>2 l\<^sub>2' b' h\<^sub>3 l\<^sub>3 ?Vs ?ls ls\<^sub>3" by(auto) have hp: "h\<^sub>2 a = Some (C, fs)" by fact from 1 2 3 hp mdecl\<^sub>1 wf_size Call_size show ?case by(fastforce simp add: comp_def intro!: Call\<^sub>1 dest!:evals_final) qed (*>*) subsection\Preservation of well-formedness\ text\The compiler preserves well-formedness. Is less trivial than it may appear. We start with two simple properties: preservation of well-typedness\ lemma compE\<^sub>1_pres_wt: "\Vs Ts U. \ P,[Vs[\]Ts] \ e :: U; size Ts = size Vs \ \ compP f P,Ts \\<^sub>1 compE\<^sub>1 Vs e :: U" and "\Vs Ts Us. \ P,[Vs[\]Ts] \ es [::] Us; size Ts = size Vs \ \ compP f P,Ts \\<^sub>1 compEs\<^sub>1 Vs es [::] Us" (*<*) -apply(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) -apply clarsimp -apply(fastforce) -apply clarsimp -apply(fastforce split:bop.splits) -apply (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) -apply (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) -apply (fastforce) -apply (fastforce) -apply (fastforce dest!: sees_method_compP[where f = f]) -apply (fastforce simp:nth_append) -apply (fastforce) -apply (fastforce) -apply (fastforce) -apply (fastforce) -apply (fastforce simp:nth_append) -apply simp -apply (fastforce) -done +proof(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) + case Var then show ?case + by (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) +next + case LAss then show ?case + by (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) +next + case Call then show ?case + by (fastforce dest!: sees_method_compP[where f = f]) +next + case Block then show ?case by (fastforce simp:nth_append) +next + case TryCatch then show ?case by (fastforce simp:nth_append) +qed fastforce+ (*>*) text\\noindent and the correct block numbering:\ lemma \: "\Vs n. size Vs = n \ \ (compE\<^sub>1 Vs e) n" and \s: "\Vs n. size Vs = n \ \s (compEs\<^sub>1 Vs es) n" -(*<*)apply (induction e and es rule: \.induct \s.induct) - apply (auto dest: sym) - apply (metis length_append_singleton) - apply (metis length_append_singleton) - done(*>*) +(*<*)by (induct e and es rule: \.induct \s.induct) + (force | simp,metis length_append_singleton)+(*>*) text\The main complication is preservation of definite assignment @{term"\"}.\ lemma image_last_index: "A \ set(xs@[x]) \ last_index (xs @ [x]) ` A = (if x \ A then insert (size xs) (last_index xs ` (A-{x})) else last_index xs ` A)" (*<*) by(auto simp:image_def) (*>*) lemma A_compE\<^sub>1_None[simp]: "\Vs. \ e = None \ \ (compE\<^sub>1 Vs e) = None" and "\Vs. \s es = None \ \s (compEs\<^sub>1 Vs es) = None" (*<*)by(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct)(auto simp:hyperset_defs)(*>*) lemma A_compE\<^sub>1: "\A Vs. \ \ e = \A\; fv e \ set Vs \ \ \ (compE\<^sub>1 Vs e) = \last_index Vs ` A\" and "\A Vs. \ \s es = \A\; fvs es \ set Vs \ \ \s (compEs\<^sub>1 Vs es) = \last_index Vs ` A\" (*<*) proof(induct e and es rule: fv.induct fvs.induct) case (Block V' T e) hence "fv e \ set (Vs@[V'])" by fastforce moreover obtain B where "\ e = \B\" using Block.prems by(simp add: hyperset_defs) moreover from calculation have "B \ set (Vs@[V'])" by(auto dest!:A_fv) ultimately show ?case using Block by(auto simp add: hyperset_defs image_last_index last_index_size_conv) next case (TryCatch e\<^sub>1 C V' e\<^sub>2) hence fve\<^sub>2: "fv e\<^sub>2 \ set (Vs@[V'])" by auto show ?case proof (cases "\ e\<^sub>1") assume A\<^sub>1: "\ e\<^sub>1 = None" then obtain A\<^sub>2 where A\<^sub>2: "\ e\<^sub>2 = \A\<^sub>2\" using TryCatch by(simp add:hyperset_defs) hence "A\<^sub>2 \ set (Vs@[V'])" using TryCatch.prems A_fv[OF A\<^sub>2] by simp blast thus ?thesis using TryCatch fve\<^sub>2 A\<^sub>1 A\<^sub>2 by(auto simp add:hyperset_defs image_last_index last_index_size_conv) next fix A\<^sub>1 assume A\<^sub>1: "\ e\<^sub>1 = \A\<^sub>1\" show ?thesis proof (cases "\ e\<^sub>2") assume A\<^sub>2: "\ e\<^sub>2 = None" then show ?case using TryCatch A\<^sub>1 by(simp add:hyperset_defs) next fix A\<^sub>2 assume A\<^sub>2: "\ e\<^sub>2 = \A\<^sub>2\" have "A\<^sub>1 \ set Vs" using TryCatch.prems A_fv[OF A\<^sub>1] by simp blast moreover have "A\<^sub>2 \ set (Vs@[V'])" using TryCatch.prems A_fv[OF A\<^sub>2] by simp blast ultimately show ?thesis using TryCatch A\<^sub>1 A\<^sub>2 by (auto simp add: Diff_subset_conv last_index_size_conv subsetD hyperset_defs dest!: sym [of _ A]) qed qed next case (Cond e e\<^sub>1 e\<^sub>2) { assume "\ e = None \ \ e\<^sub>1 = None \ \ e\<^sub>2 = None" hence ?case using Cond by (auto simp add: hyperset_defs) } moreover { fix A A\<^sub>1 A\<^sub>2 assume "\ e = \A\" and A\<^sub>1: "\ e\<^sub>1 = \A\<^sub>1\" and A\<^sub>2: "\ e\<^sub>2 = \A\<^sub>2\" moreover have "A\<^sub>1 \ set Vs" using Cond.prems A_fv[OF A\<^sub>1] by simp blast moreover have "A\<^sub>2 \ set Vs" using Cond.prems A_fv[OF A\<^sub>2] by simp blast ultimately have ?case using Cond by(auto simp add:hyperset_defs image_Un inj_on_image_Int[OF inj_on_last_index]) } ultimately show ?case by fastforce qed (auto simp add:hyperset_defs) (*>*) lemma D_None[iff]: "\ (e::'a exp) None" and [iff]: "\s (es::'a exp list) None" (*<*)by(induct e and es rule: \.induct \s.induct)(simp_all)(*>*) lemma D_last_index_compE\<^sub>1: "\A Vs. \ A \ set Vs; fv e \ set Vs \ \ \ e \A\ \ \ (compE\<^sub>1 Vs e) \last_index Vs ` A\" and "\A Vs. \ A \ set Vs; fvs es \ set Vs \ \ \s es \A\ \ \s (compEs\<^sub>1 Vs es) \last_index Vs ` A\" (*<*) proof(induct e and es rule: \.induct \s.induct) case (BinOp e\<^sub>1 bop e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using BinOp by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] BinOp.prems by auto have "A \ A\<^sub>1 \ set Vs" using BinOp.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using BinOp Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (FAss e\<^sub>1 F D e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using FAss by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] FAss.prems by auto have "A \ A\<^sub>1 \ set Vs" using FAss.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using FAss Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (Call e\<^sub>1 M es) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using Call by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] Call.prems by auto have "A \ A\<^sub>1 \ set Vs" using Call.prems A_fv[OF Some] by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` (A \ A\<^sub>1)\" using Call Some by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (TryCatch e\<^sub>1 C V e\<^sub>2) have "\ A\{V} \ set(Vs@[V]); fv e\<^sub>2 \ set(Vs@[V]); \ e\<^sub>2 \A\{V}\\ \ \ (compE\<^sub>1 (Vs@[V]) e\<^sub>2) \last_index (Vs@[V]) ` (A\{V})\" by fact hence "\ (compE\<^sub>1 (Vs@[V]) e\<^sub>2) \last_index (Vs@[V]) ` (A\{V})\" using TryCatch.prems by(simp add:Diff_subset_conv) moreover have "last_index (Vs@[V]) ` A \ last_index Vs ` A \ {size Vs}" using TryCatch.prems by(auto simp add: image_last_index split:if_split_asm) ultimately show ?case using TryCatch by(auto simp:hyperset_defs elim!:D_mono') next case (Seq e\<^sub>1 e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using Seq by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] Seq.prems by auto have "A \ A\<^sub>1 \ set Vs" using Seq.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using Seq Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (Cond e e\<^sub>1 e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e) \last_index Vs ` A\" by simp show ?case proof (cases "\ e") case None thus ?thesis using Cond by simp next case (Some B) have indexB: "\ (compE\<^sub>1 Vs e) = \last_index Vs ` B\" using A_compE\<^sub>1[OF Some] Cond.prems by auto have "A \ B \ set Vs" using Cond.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` (A \ B)\" and "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ B)\" using Cond Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A \ last_index Vs ` B\" and "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` B\" by(simp add: image_Un)+ thus ?thesis using IH\<^sub>1 indexB by auto qed next case (While e\<^sub>1 e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using While by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] While.prems by auto have "A \ A\<^sub>1 \ set Vs" using While.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using While Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (Block V T e) have "\ A-{V} \ set(Vs@[V]); fv e \ set(Vs@[V]); \ e \A-{V}\ \ \ \ (compE\<^sub>1 (Vs@[V]) e) \last_index (Vs@[V]) ` (A-{V})\" by fact hence "\ (compE\<^sub>1 (Vs@[V]) e) \last_index (Vs@[V]) ` (A-{V})\" using Block.prems by(simp add:Diff_subset_conv) moreover have "size Vs \ last_index Vs ` A" using Block.prems by(auto simp add:image_def size_last_index_conv) ultimately show ?case using Block by(auto simp add: image_last_index Diff_subset_conv hyperset_defs elim!: D_mono') next case (Cons_exp e\<^sub>1 es) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using Cons_exp by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] Cons_exp.prems by auto have "A \ A\<^sub>1 \ set Vs" using Cons_exp.prems A_fv[OF Some] by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` (A \ A\<^sub>1)\" using Cons_exp Some by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed qed (simp_all add:hyperset_defs) (*>*) lemma last_index_image_set: "distinct xs \ last_index xs ` set xs = {..*) lemma D_compE\<^sub>1: "\ \ e \set Vs\; fv e \ set Vs; distinct Vs \ \ \ (compE\<^sub>1 Vs e) \{.." (*<*)by(fastforce dest!: D_last_index_compE\<^sub>1[OF subset_refl] simp add:last_index_image_set)(*>*) lemma D_compE\<^sub>1': assumes "\ e \set(V#Vs)\" and "fv e \ set(V#Vs)" and "distinct(V#Vs)" shows "\ (compE\<^sub>1 (V#Vs) e) \{..length Vs}\" (*<*) proof - have "{..size Vs} = {..1) qed (*>*) lemma compP\<^sub>1_pres_wf: "wf_J_prog P \ wf_J\<^sub>1_prog (compP\<^sub>1 P)" (*<*) -apply simp -apply(rule wf_prog_compPI) - prefer 2 apply assumption -apply(case_tac m) -apply(simp add:wf_mdecl_def wf_J\<^sub>1_mdecl_def wf_J_mdecl) -apply(clarify) -apply(frule WT_fv) -apply(fastforce intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1' \) -done +proof - + assume wf: "wf_J_prog P" + let ?f = "(\(pns, body). compE\<^sub>1 (this # pns) body)" + let ?wf\<^sub>2 = "wf_J\<^sub>1_mdecl" + + { fix C M b Ts T m + assume cM: "P \ C sees M : Ts\T = m in C" + and wfm: "wf_mdecl wf_J_mdecl P C (M, Ts, T, m)" + obtain pns body where [simp]: "m = (pns, body)" by(cases m) simp + let ?E = "[pns [\] Ts, this \ Class C]" + obtain T' where WT: "P,?E \ body :: T'" and subT: "P \ T' \ T" + using wfm by(auto simp: wf_mdecl_def) + have fv: "fv body \ dom ?E" by(rule WT_fv[OF WT]) + have "wf_mdecl ?wf\<^sub>2 (compP ?f P) C (M, Ts, T, ?f m)" using cM wfm fv + by(clarsimp simp add:wf_mdecl_def wf_J\<^sub>1_mdecl_def) + (fastforce intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1' \) + } + then show ?thesis by simp (rule wf_prog_compPI[OF _ wf]) +qed (*>*) end diff --git a/thys/Jinja/Compiler/Correctness2.thy b/thys/Jinja/Compiler/Correctness2.thy --- a/thys/Jinja/Compiler/Correctness2.thy +++ b/thys/Jinja/Compiler/Correctness2.thy @@ -1,1090 +1,1145 @@ (* Title: Jinja/Compiler/Correctness2.thy Author: Tobias Nipkow Copyright TUM 2003 *) section \Correctness of Stage 2\ theory Correctness2 imports "HOL-Library.Sublist" Compiler2 begin (*<*)hide_const (open) Throw(*>*) subsection\Instruction sequences\ text\How to select individual instructions and subsequences of instructions from a program given the class, method and program counter.\ definition before :: "jvm_prog \ cname \ mname \ nat \ instr list \ bool" ("(_,_,_,_/ \ _)" [51,0,0,0,51] 50) where "P,C,M,pc \ is \ prefix is (drop pc (instrs_of P C M))" definition at :: "jvm_prog \ cname \ mname \ nat \ instr \ bool" ("(_,_,_,_/ \ _)" [51,0,0,0,51] 50) where "P,C,M,pc \ i \ (\is. drop pc (instrs_of P C M) = i#is)" lemma [simp]: "P,C,M,pc \ []" (*<*)by(simp add:before_def)(*>*) lemma [simp]: "P,C,M,pc \ (i#is) = (P,C,M,pc \ i \ P,C,M,pc + 1 \ is)" (*<*)by(fastforce simp add:before_def at_def prefix_def drop_Suc drop_tl)(*>*) (*<*) declare drop_drop[simp del] (*>*) lemma [simp]: "P,C,M,pc \ (is\<^sub>1 @ is\<^sub>2) = (P,C,M,pc \ is\<^sub>1 \ P,C,M,pc + size is\<^sub>1 \ is\<^sub>2)" (*<*) -apply(simp add:before_def prefix_def) -apply(subst add.commute) -apply(simp add: drop_drop[symmetric]) -apply fastforce -done +by(subst add.commute) + (fastforce simp add:before_def prefix_def drop_drop[symmetric]) (*>*) (*<*) declare drop_drop[simp] (*>*) lemma [simp]: "P,C,M,pc \ i \ instrs_of P C M ! pc = i" (*<*)by(clarsimp simp add:at_def strict_prefix_def nth_via_drop)(*>*) lemma beforeM: "P \ C sees M: Ts\T = body in D \ compP\<^sub>2 P,D,M,0 \ compE\<^sub>2 body @ [Return]" (*<*) -apply(drule sees_method_idemp) -apply(simp add:before_def compP\<^sub>2_def compMb\<^sub>2_def) -done +by(drule sees_method_idemp) + (simp add:before_def compP\<^sub>2_def compMb\<^sub>2_def) (*>*) text\This lemma executes a single instruction by rewriting:\ lemma [simp]: "P,C,M,pc \ instr \ (P \ (None, h, (vs,ls,C,M,pc) # frs) -jvm\ \') = ((None, h, (vs,ls,C,M,pc) # frs) = \' \ (\\. exec(P,(None, h, (vs,ls,C,M,pc) # frs)) = Some \ \ P \ \ -jvm\ \'))" (*<*) -apply(simp only: exec_all_def) -apply(blast intro: converse_rtranclE converse_rtrancl_into_rtrancl) -done +by(simp only: exec_all_def) + (blast intro: converse_rtranclE converse_rtrancl_into_rtrancl) (*>*) subsection\Exception tables\ definition pcs :: "ex_table \ nat set" where "pcs xt \ \(f,t,C,h,d) \ set xt. {f ..< t}" lemma pcs_subset: shows "\pc d. pcs(compxE\<^sub>2 e pc d) \ {pc..2 e)}" and "\pc d. pcs(compxEs\<^sub>2 es pc d) \ {pc..2 es)}" (*<*) -apply(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) -apply (simp_all add:pcs_def) -apply (fastforce split:bop.splits)+ -done +proof(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) + case Cast then show ?case by (fastforce simp:pcs_def) +next + case BinOp then show ?case by (fastforce simp:pcs_def split:bop.splits) +next + case LAss then show ?case by (fastforce simp: pcs_def) +next + case FAcc then show ?case by (fastforce simp: pcs_def) +next + case FAss then show ?case by (fastforce simp: pcs_def) +next + case Call then show ?case by (fastforce simp: pcs_def) +next + case Seq then show ?case by (fastforce simp: pcs_def) +next + case Cond then show ?case by (fastforce simp: pcs_def) +next + case While then show ?case by (fastforce simp: pcs_def) +next + case throw then show ?case by (fastforce simp: pcs_def) +next + case TryCatch then show ?case by (fastforce simp: pcs_def) +next + case Cons_exp then show ?case by (fastforce simp: pcs_def) +qed (simp_all add:pcs_def) (*>*) lemma [simp]: "pcs [] = {}" (*<*)by(simp add:pcs_def)(*>*) lemma [simp]: "pcs (x#xt) = {fst x ..< fst(snd x)} \ pcs xt" (*<*)by(auto simp add: pcs_def)(*>*) lemma [simp]: "pcs(xt\<^sub>1 @ xt\<^sub>2) = pcs xt\<^sub>1 \ pcs xt\<^sub>2" (*<*)by(simp add:pcs_def)(*>*) lemma [simp]: "pc < pc\<^sub>0 \ pc\<^sub>0+size(compE\<^sub>2 e) \ pc \ pc \ pcs(compxE\<^sub>2 e pc\<^sub>0 d)" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc < pc\<^sub>0 \ pc\<^sub>0+size(compEs\<^sub>2 es) \ pc \ pc \ pcs(compxEs\<^sub>2 es pc\<^sub>0 d)" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc\<^sub>1 + size(compE\<^sub>2 e\<^sub>1) \ pc\<^sub>2 \ pcs(compxE\<^sub>2 e\<^sub>1 pc\<^sub>1 d\<^sub>1) \ pcs(compxE\<^sub>2 e\<^sub>2 pc\<^sub>2 d\<^sub>2) = {}" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc\<^sub>1 + size(compE\<^sub>2 e) \ pc\<^sub>2 \ pcs(compxE\<^sub>2 e pc\<^sub>1 d\<^sub>1) \ pcs(compxEs\<^sub>2 es pc\<^sub>2 d\<^sub>2) = {}" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc \ pcs xt\<^sub>0 \ match_ex_table P C pc (xt\<^sub>0 @ xt\<^sub>1) = match_ex_table P C pc xt\<^sub>1" (*<*)by (induct xt\<^sub>0) (auto simp: matches_ex_entry_def)(*>*) lemma [simp]: "\ x \ set xt; pc \ pcs xt \ \ \ matches_ex_entry P D pc x" (*<*)by(auto simp:matches_ex_entry_def pcs_def)(*>*) lemma [simp]: assumes xe: "xe \ set(compxE\<^sub>2 e pc d)" and outside: "pc' < pc \ pc+size(compE\<^sub>2 e) \ pc'" shows "\ matches_ex_entry P C pc' xe" (*<*) proof assume "matches_ex_entry P C pc' xe" with xe have "pc' \ pcs(compxE\<^sub>2 e pc d)" by(force simp add:matches_ex_entry_def pcs_def) with outside show False by simp qed (*>*) lemma [simp]: assumes xe: "xe \ set(compxEs\<^sub>2 es pc d)" and outside: "pc' < pc \ pc+size(compEs\<^sub>2 es) \ pc'" shows "\ matches_ex_entry P C pc' xe" (*<*) proof assume "matches_ex_entry P C pc' xe" with xe have "pc' \ pcs(compxEs\<^sub>2 es pc d)" by(force simp add:matches_ex_entry_def pcs_def) with outside show False by simp qed (*>*) lemma match_ex_table_app[simp]: "\xte \ set xt\<^sub>1. \ matches_ex_entry P D pc xte \ match_ex_table P D pc (xt\<^sub>1 @ xt) = match_ex_table P D pc xt" (*<*)by(induct xt\<^sub>1) simp_all(*>*) lemma [simp]: "\x \ set xtab. \ matches_ex_entry P C pc x \ match_ex_table P C pc xtab = None" (*<*)using match_ex_table_app[where ?xt = "[]"] by fastforce(*>*) lemma match_ex_entry: "matches_ex_entry P C pc (start, end, catch_type, handler) = (start \ pc \ pc < end \ P \ C \\<^sup>* catch_type)" (*<*)by(simp add:matches_ex_entry_def)(*>*) definition caught :: "jvm_prog \ pc \ heap \ addr \ ex_table \ bool" where "caught P pc h a xt \ (\entry \ set xt. matches_ex_entry P (cname_of h a) pc entry)" definition beforex :: "jvm_prog \ cname \ mname \ ex_table \ nat set \ nat \ bool" ("(2_,/_,/_ \/ _ /'/ _,/_)" [51,0,0,0,0,51] 50) where "P,C,M \ xt / I,d \ (\xt\<^sub>0 xt\<^sub>1. ex_table_of P C M = xt\<^sub>0 @ xt @ xt\<^sub>1 \ pcs xt\<^sub>0 \ I = {} \ pcs xt \ I \ (\pc \ I. \C pc' d'. match_ex_table P C pc xt\<^sub>1 = \(pc',d')\ \ d' \ d))" definition dummyx :: "jvm_prog \ cname \ mname \ ex_table \ nat set \ nat \ bool" ("(2_,_,_ \/ _ '/_,_)" [51,0,0,0,0,51] 50) where "P,C,M \ xt/I,d \ P,C,M \ xt/I,d" +abbreviation +"beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1 + \ ex_table_of P C M = xt\<^sub>0 @ xt @ xt\<^sub>1 \ pcs xt\<^sub>0 \ I = {} + \ pcs xt \ I \ (\pc \ I. \C pc' d'. match_ex_table P C pc xt\<^sub>1 = \(pc',d')\ \ d' \ d)" + +lemma beforex_beforex\<^sub>0_eq: + "P,C,M \ xt / I,d \ \xt\<^sub>0 xt\<^sub>1. beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1" +using beforex_def by auto + + lemma beforexD1: "P,C,M \ xt / I,d \ pcs xt \ I" (*<*)by(auto simp add:beforex_def)(*>*) lemma beforex_mono: "\ P,C,M \ xt/I,d'; d' \ d \ \ P,C,M \ xt/I,d" (*<*)by(fastforce simp:beforex_def)(*>*) lemma [simp]: "P,C,M \ xt/I,d \ P,C,M \ xt/I,Suc d" (*<*)by(fastforce intro:beforex_mono)(*>*) lemma beforex_append[simp]: "pcs xt\<^sub>1 \ pcs xt\<^sub>2 = {} \ P,C,M \ xt\<^sub>1 @ xt\<^sub>2/I,d = (P,C,M \ xt\<^sub>1/I-pcs xt\<^sub>2,d \ P,C,M \ xt\<^sub>2/I-pcs xt\<^sub>1,d \ P,C,M \ xt\<^sub>1@xt\<^sub>2/I,d)" -(*<*) -apply(rule iffI) - prefer 2 - apply(simp add:dummyx_def) -apply(auto simp add: beforex_def dummyx_def) - apply(rule_tac x = xt\<^sub>0 in exI) - apply auto -apply(rule_tac x = "xt\<^sub>0@xt\<^sub>1" in exI) -apply auto -done +(*<*)(is "?Q \ ?P = (?P1 \ ?P2 \ ?P3)" is "?Q \ ?P = ?P123") +proof - + assume pcs: ?Q + show ?thesis proof(rule iffI) + assume "?P123" then show ?P by(simp add:dummyx_def) + next + assume hyp: ?P + let ?xt = "xt\<^sub>1 @ xt\<^sub>2" + let ?beforex = "beforex\<^sub>0 P C M d" + obtain xt\<^sub>0 xt\<^sub>1' where beforex: "?beforex I ?xt xt\<^sub>0 xt\<^sub>1'" + using hyp by(clarsimp simp: beforex_def) + have "\xt\<^sub>0 xt\<^sub>1'. ?beforex (I - pcs xt\<^sub>2) xt\<^sub>1 xt\<^sub>0 xt\<^sub>1'" \ \?P1\ + using pcs beforex by(rule_tac x=xt\<^sub>0 in exI) auto + moreover have "\xt\<^sub>0 xt\<^sub>1'. ?beforex (I - pcs xt\<^sub>1) xt\<^sub>2 xt\<^sub>0 xt\<^sub>1'" \ \?P2\ + using pcs beforex by(rule_tac x="xt\<^sub>0@xt\<^sub>1" in exI) auto + moreover have ?P3 using hyp by(simp add: dummyx_def) + ultimately show ?P123 by (simp add: beforex_def) + qed +qed (*>*) lemma beforex_appendD1: - "\ P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d; - pcs xt\<^sub>1 \ J; J \ I; J \ pcs xt\<^sub>2 = {} \ - \ P,C,M \ xt\<^sub>1 / J,d" +assumes bx: "P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d" + and pcs: "pcs xt\<^sub>1 \ J" and JI: "J \ I" and Jpcs: "J \ pcs xt\<^sub>2 = {}" +shows "P,C,M \ xt\<^sub>1 / J,d" (*<*) -apply(auto simp:beforex_def) -apply(rule exI,rule exI,rule conjI, rule refl) -apply(rule conjI, blast) -apply(auto) -apply(subgoal_tac "pc \ pcs xt\<^sub>2") - prefer 2 apply blast -apply (auto split:if_split_asm) -done +proof - + let ?beforex = "beforex\<^sub>0 P C M d" + obtain xt\<^sub>0 xt\<^sub>1' where bx': "?beforex I (xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)]) xt\<^sub>0 xt\<^sub>1'" + using bx by(clarsimp simp:beforex_def) + let ?xt0 = xt\<^sub>0 and ?xt1 = "xt\<^sub>2 @ (f, t, D, h, d) # xt\<^sub>1'" + have "pcs xt\<^sub>0 \ J = {}" using bx' JI by blast + moreover { + fix pc C pc' d' assume pcJ: "pc\J" + then have "pc \ pcs xt\<^sub>2" using bx' Jpcs by blast + then have "match_ex_table P C pc (xt\<^sub>2 @ (f, t, D, h, d) # xt\<^sub>1') + = \(pc', d')\ \ d' \ d" + using bx' JI pcJ by (auto split:if_split_asm) + } + ultimately have "?beforex J xt\<^sub>1 ?xt0 ?xt1" using bx' pcs by simp + then show ?thesis using beforex_def by blast +qed (*>*) lemma beforex_appendD2: - "\ P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d; - pcs xt\<^sub>2 \ J; J \ I; J \ pcs xt\<^sub>1 = {} \ - \ P,C,M \ xt\<^sub>2 / J,d" +assumes bx: "P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d" + and pcs: "pcs xt\<^sub>2 \ J" and JI: "J \ I" and Jpcs: "J \ pcs xt\<^sub>1 = {}" +shows "P,C,M \ xt\<^sub>2 / J,d" (*<*) -apply(auto simp:beforex_def) -apply(rule_tac x = "xt\<^sub>0 @ xt\<^sub>1" in exI) -apply fastforce -done +proof - + let ?beforex = "beforex\<^sub>0 P C M d" + obtain xt\<^sub>0 xt\<^sub>1' where bx': "?beforex I (xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)]) xt\<^sub>0 xt\<^sub>1'" + using bx by(clarsimp simp:beforex_def) + then have "\xt\<^sub>1''. beforex\<^sub>0 P C M d J xt\<^sub>2 (xt\<^sub>0 @ xt\<^sub>1) xt\<^sub>1''" + using assms by fastforce + then show ?thesis using beforex_def by blast +qed (*>*) lemma beforexM: "P \ C sees M: Ts\T = body in D \ compP\<^sub>2 P,D,M \ compxE\<^sub>2 body 0 0/{..2 body)},0" (*<*) -apply(drule sees_method_idemp) -apply(drule sees_method_compP[where f = compMb\<^sub>2]) -apply(simp add:beforex_def compP\<^sub>2_def compMb\<^sub>2_def) -apply(rule_tac x = "[]" in exI) -using pcs_subset apply fastforce -done +proof - + assume cM: "P \ C sees M: Ts\T = body in D" + let ?xt0 = "[]" + have "\xt1. beforex\<^sub>0 (compP\<^sub>2 P) D M 0 ({..2 body)}) (compxE\<^sub>2 body 0 0) ?xt0 xt1" + using sees_method_compP[where f = compMb\<^sub>2, OF sees_method_idemp[OF cM]] + pcs_subset by(fastforce simp add: compP\<^sub>2_def compMb\<^sub>2_def) + then show ?thesis using beforex_def by blast +qed (*>*) lemma match_ex_table_SomeD2: - "\ match_ex_table P D pc (ex_table_of P C M) = \(pc',d')\; - P,C,M \ xt/I,d; \x \ set xt. \ matches_ex_entry P D pc x; pc \ I \ - \ d' \ d" +assumes met: "match_ex_table P D pc (ex_table_of P C M) = \(pc',d')\" + and bx: "P,C,M \ xt/I,d" + and nmet: "\x \ set xt. \ matches_ex_entry P D pc x" and pcI: "pc \ I" +shows "d' \ d" (*<*) -apply(auto simp:beforex_def) -apply(subgoal_tac "pc \ pcs xt\<^sub>0") -apply auto -done +proof - + obtain xt\<^sub>0 xt\<^sub>1 where bx': "beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1" + using bx by(clarsimp simp:beforex_def) + then have "pc \ pcs xt\<^sub>0" using pcI by blast + then show ?thesis using bx' met nmet pcI by simp +qed (*>*) lemma match_ex_table_SomeD1: "\ match_ex_table P D pc (ex_table_of P C M) = \(pc',d')\; P,C,M \ xt / I,d; pc \ I; pc \ pcs xt \ \ d' \ d" (*<*)by(auto elim: match_ex_table_SomeD2)(*>*) subsection\The correctness proof\ (*<*) declare nat_add_distrib[simp] caught_def[simp] declare fun_upd_apply[simp del] (*>*) definition handle :: "jvm_prog \ cname \ mname \ addr \ heap \ val list \ val list \ nat \ frame list \ jvm_state" where "handle P C M a h vs ls pc frs = find_handler P a h ((vs,ls,C,M,pc) # frs)" lemma handle_Cons: "\ P,C,M \ xt/I,d; d \ size vs; pc \ I; \x \ set xt. \ matches_ex_entry P (cname_of h xa) pc x \ \ handle P C M xa h (v # vs) ls pc frs = handle P C M xa h vs ls pc frs" (*<*)by(auto simp:handle_def Suc_diff_le dest: match_ex_table_SomeD2)(*>*) lemma handle_append: - "\ P,C,M \ xt/I,d; d \ size vs; pc \ I; pc \ pcs xt \ \ - handle P C M xa h (ws @ vs) ls pc frs = handle P C M xa h vs ls pc frs" +assumes bx: "P,C,M \ xt/I,d" and d: "d \ size vs" + and pcI: "pc \ I" and pc_not: "pc \ pcs xt" +shows "handle P C M xa h (ws @ vs) ls pc frs = handle P C M xa h vs ls pc frs" (*<*) -apply(auto simp:handle_def) -apply(rename_tac pc' d') -apply(subgoal_tac "size ws \ length ws + length vs - d'") -apply(simp add:drop_all) -apply(fastforce dest:match_ex_table_SomeD2 split:nat_diff_split) -done +proof - + { fix pc' d' + assume "match_ex_table P (cname_of h xa) pc (ex_table_of P C M) = \(pc', d')\" + then have "length ws \ length ws + length vs - d'" + using assms by(fastforce dest:match_ex_table_SomeD2 split:nat_diff_split) + } + then show ?thesis by(simp add: handle_def) +qed (*>*) lemma aux_isin[simp]: "\ B \ A; a \ B \ \ a \ A" (*<*)by blast(*>*) lemma fixes P\<^sub>1 defines [simp]: "P \ compP\<^sub>2 P\<^sub>1" shows Jcc: "P\<^sub>1 \\<^sub>1 \e,(h\<^sub>0,ls\<^sub>0)\ \ \ef,(h\<^sub>1,ls\<^sub>1)\ \ (\C M pc v xa vs frs I. \ P,C,M,pc \ compE\<^sub>2 e; P,C,M \ compxE\<^sub>2 e pc (size vs)/I,size vs; {pc..2 e)} \ I \ \ (ef = Val v \ P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 e))#frs)) \ (ef = Throw xa \ (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compE\<^sub>2 e) \ \ caught P pc\<^sub>1 h\<^sub>1 xa (compxE\<^sub>2 e pc (size vs)) \ P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ handle P C M xa h\<^sub>1 vs ls\<^sub>1 pc\<^sub>1 frs)))" (*<*) (is "_ \ (\C M pc v xa vs frs I. PROP ?P e h\<^sub>0 ls\<^sub>0 ef h\<^sub>1 ls\<^sub>1 C M pc v xa vs frs I)") (*>*) and "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>0,ls\<^sub>0)\ [\] \fs,(h\<^sub>1,ls\<^sub>1)\ \ (\C M pc ws xa es' vs frs I. \ P,C,M,pc \ compEs\<^sub>2 es; P,C,M \ compxEs\<^sub>2 es pc (size vs)/I,size vs; {pc..2 es)} \ I \ \ (fs = map Val ws \ P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(rev ws @ vs,ls\<^sub>1,C,M,pc+size(compEs\<^sub>2 es))#frs)) \ (fs = map Val ws @ Throw xa # es' \ (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compEs\<^sub>2 es) \ \ caught P pc\<^sub>1 h\<^sub>1 xa (compxEs\<^sub>2 es pc (size vs)) \ P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ handle P C M xa h\<^sub>1 vs ls\<^sub>1 pc\<^sub>1 frs)))" (*<*) (is "_ \ (\C M pc ws xa es' vs frs I. PROP ?Ps es h\<^sub>0 ls\<^sub>0 fs h\<^sub>1 ls\<^sub>1 C M pc ws xa es' vs frs I)") proof (induct rule:eval\<^sub>1_evals\<^sub>1_inducts) case New\<^sub>1 thus ?case by (clarsimp simp add:blank_def fun_eq_iff) next case NewFail\<^sub>1 thus ?case by(auto simp: handle_def pcs_def) next case (Cast\<^sub>1 e h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1 D fs C') let ?pc = "pc + length(compE\<^sub>2 e)" have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc)#frs)" using Cast\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc+1)#frs)" using Cast\<^sub>1 by (auto simp add:cast_ok_def) finally show ?case by auto next case (CastNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 C') let ?pc = "pc + length(compE\<^sub>2 e)" have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc)#frs)" using CastNull\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc+1)#frs)" using CastNull\<^sub>1 by (auto simp add:cast_ok_def) finally show ?case by auto next case (CastFail\<^sub>1 e h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1 D fs C') let ?pc = "pc + length(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt ClassCast" have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc)#frs)" using CastFail\<^sub>1 by fastforce also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc frs" using CastFail\<^sub>1 by (auto simp add:handle_def cast_ok_def) also have "handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc frs = handle P C M ?xa h\<^sub>1 vs ls\<^sub>1 ?pc frs" using CastFail\<^sub>1.prems by(auto simp:handle_Cons) finally have exec: "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ \". show ?case (is "?N \ (?eq \ (\pc\<^sub>1. ?H pc\<^sub>1))") proof show ?N by simp next have "?eq \ ?H ?pc" using exec by auto thus "?eq \ (\pc\<^sub>1. ?H pc\<^sub>1)" by blast qed next case CastThrow\<^sub>1 thus ?case by fastforce next case Val\<^sub>1 thus ?case by simp next case Var\<^sub>1 thus ?case by auto next case (BinOp\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 e\<^sub>2 v\<^sub>2 h\<^sub>2 ls\<^sub>2 bop w) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 (Val v\<^sub>2) h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 v\<^sub>2 xa (v\<^sub>1#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using BinOp\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>2,(v\<^sub>2#v\<^sub>1#vs,ls\<^sub>2,C,M,?pc\<^sub>2)#frs)" using BinOp\<^sub>1.prems IH\<^sub>2 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>2,(w#vs,ls\<^sub>2,C,M,?pc\<^sub>2+1)#frs)" using BinOp\<^sub>1 by(cases bop) auto finally show ?case by (auto split: bop.splits simp:add.assoc) next case BinOpThrow\<^sub>1\<^sub>1 thus ?case by(fastforce) next case (BinOpThrow\<^sub>2\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 e\<^sub>2 e h\<^sub>2 ls\<^sub>2 bop) let ?pc = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?\\<^sub>1 = "(None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc)#frs)" have 1: "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ ?\\<^sub>1" using BinOpThrow\<^sub>2\<^sub>1 by fastforce show ?case (is "?N \ (?eq \ (\pc\<^sub>2. ?H pc\<^sub>2))") proof show ?N by simp next { assume ?eq moreover have "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 (throw e) h\<^sub>2 ls\<^sub>2 C M ?pc v xa (v\<^sub>1#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact ultimately obtain pc\<^sub>2 where pc\<^sub>2: "?pc \ pc\<^sub>2 \ pc\<^sub>2 < ?pc + size(compE\<^sub>2 e\<^sub>2) \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc (size vs + 1))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (v\<^sub>1#vs) ls\<^sub>2 pc\<^sub>2 frs" using BinOpThrow\<^sub>2\<^sub>1.prems by fastforce have 3: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 vs ls\<^sub>2 pc\<^sub>2 frs" using 2 BinOpThrow\<^sub>2\<^sub>1.prems pc\<^sub>2 by(auto simp:handle_Cons) have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF 1 3] by auto hence "\pc\<^sub>2. ?H pc\<^sub>2" by iprover } thus "?eq \ (\pc\<^sub>2. ?H pc\<^sub>2)" by iprover qed next case (FAcc\<^sub>1 e h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1 C fs F D w) let ?pc = "pc + length(compE\<^sub>2 e)" have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc)#frs)" using FAcc\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc+1)#frs)" using FAcc\<^sub>1 by auto finally show ?case by auto next case (FAccNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 F D) let ?pc = "pc + length(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt NullPointer" have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc)#frs)" using FAccNull\<^sub>1 by fastforce also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc frs" using FAccNull\<^sub>1.prems by(fastforce simp:split_beta handle_def simp del: split_paired_Ex) also have "handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc frs = handle P C M ?xa h\<^sub>1 vs ls\<^sub>1 ?pc frs" using FAccNull\<^sub>1.prems by(auto simp add:handle_Cons) finally show ?case by (auto intro: exI[where x = ?pc]) next case FAccThrow\<^sub>1 thus ?case by fastforce next case (LAss\<^sub>1 e h\<^sub>0 ls\<^sub>0 w h\<^sub>1 ls\<^sub>1 i ls\<^sub>2) let ?pc = "pc + length(compE\<^sub>2 e)" have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc)#frs)" using LAss\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>1,(Unit#vs,ls\<^sub>2,C,M,?pc+2)#frs)" using LAss\<^sub>1 by auto finally show ?case using LAss\<^sub>1 by auto next case LAssThrow\<^sub>1 thus ?case by fastforce next case (FAss\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 C fs fs' F D h\<^sub>2') let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 w xa (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using FAss\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>2,(w#Addr a#vs,ls\<^sub>2,C,M,?pc\<^sub>2)#frs)" using FAss\<^sub>1.prems IH\<^sub>2 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>2',(Unit#vs,ls\<^sub>2,C,M,?pc\<^sub>2+2)#frs)" using FAss\<^sub>1 by auto finally show ?case using FAss\<^sub>1 by (auto simp:add.assoc) next case (FAssNull\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 F D) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" let ?xa = "addr_of_sys_xcpt NullPointer" have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 w xa (Null#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using FAssNull\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>2,(w#Null#vs,ls\<^sub>2,C,M,?pc\<^sub>2)#frs)" using FAssNull\<^sub>1.prems IH\<^sub>2 by fastforce also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (w#Null#vs) ls\<^sub>2 ?pc\<^sub>2 frs" using FAssNull\<^sub>1.prems by(fastforce simp:split_beta handle_def simp del: split_paired_Ex) also have "handle P C M ?xa h\<^sub>2 (w#Null#vs) ls\<^sub>2 ?pc\<^sub>2 frs = handle P C M ?xa h\<^sub>2 vs ls\<^sub>2 ?pc\<^sub>2 frs" using FAssNull\<^sub>1.prems by(auto simp add:handle_Cons) finally show ?case by (auto intro: exI[where x = ?pc\<^sub>2]) next case (FAssThrow\<^sub>2\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 w h\<^sub>1 ls\<^sub>1 e\<^sub>2 e' h\<^sub>2 ls\<^sub>2 F D) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?\\<^sub>1 = "(None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" have 1: "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ ?\\<^sub>1" using FAssThrow\<^sub>2\<^sub>1 by fastforce show ?case (is "?N \ (?eq \ (\pc\<^sub>2. ?H pc\<^sub>2))") proof show ?N by simp next { assume ?eq moreover have "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 v xa (w#vs) frs (I - pcs (compxE\<^sub>2 e\<^sub>1 pc (length vs)))" by fact ultimately obtain pc\<^sub>2 where pc\<^sub>2: "?pc\<^sub>1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1 + size(compE\<^sub>2 e\<^sub>2) \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1 (size vs + 1))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (w#vs) ls\<^sub>2 pc\<^sub>2 frs" using FAssThrow\<^sub>2\<^sub>1.prems by fastforce have 3: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 vs ls\<^sub>2 pc\<^sub>2 frs" using 2 FAssThrow\<^sub>2\<^sub>1.prems pc\<^sub>2 by(auto simp:handle_Cons) have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF 1 3] by auto hence "\pc\<^sub>2. ?H pc\<^sub>2" by iprover } thus "?eq \ (\pc\<^sub>2. ?H pc\<^sub>2)" by iprover qed next case FAssThrow\<^sub>1\<^sub>1 thus ?case by fastforce next case (Call\<^sub>1 e h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 Ca fs M' Ts T body D ls\<^sub>2' f h\<^sub>3 ls\<^sub>3) have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) let ?\\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc)#frs)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>1 = "(None,h\<^sub>1,(Addr a # vs, ls\<^sub>1, C,M,?pc\<^sub>1)#frs)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" let ?frs\<^sub>2 = "(rev pvs @ Addr a # vs, ls\<^sub>2, C,M,?pc\<^sub>2)#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2)" let ?frs\<^sub>2' = "([], ls\<^sub>2', D,M',0) # ?frs\<^sub>2" let ?\\<^sub>2' = "(None, h\<^sub>2, ?frs\<^sub>2')" have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 pvs xa (map Val pvs) (Addr a # vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact have "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using Call\<^sub>1 by fastforce also have "P \ \ -jvm\ ?\\<^sub>2" using IH_es Call\<^sub>1.prems by fastforce also have "P \ \ -jvm\ ?\\<^sub>2'" using Call\<^sub>1 by(auto simp add: nth_append compMb\<^sub>2_def) finally have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>2'". have "P\<^sub>1 \ Ca sees M': Ts\T = body in D" by fact then have M'_in_D: "P\<^sub>1 \ D sees M': Ts\T = body in D" by(rule sees_method_idemp) hence M'_code: "compP\<^sub>2 P\<^sub>1,D,M',0 \ compE\<^sub>2 body @ [Return]" and M'_xtab: "compP\<^sub>2 P\<^sub>1,D,M' \ compxE\<^sub>2 body 0 0/{..2 body)},0" by(rule beforeM, rule beforexM) have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' f h\<^sub>3 ls\<^sub>3 D M' 0 v xa [] ?frs\<^sub>2 ({..2 body)})" by fact show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>2' -jvm\ (None,h\<^sub>3,([v],ls\<^sub>3,D,M',size(compE\<^sub>2 body))#?frs\<^sub>2)" using val IH_body Call\<^sub>1.prems M'_code M'_xtab by (fastforce simp del:split_paired_Ex) also have "P \ \ -jvm\ (None, h\<^sub>3, (v # vs, ls\<^sub>2, C,M,?pc\<^sub>2+1)#frs)" using Call\<^sub>1 M'_code M'_in_D by(auto simp: nth_append compMb\<^sub>2_def) finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof assume throw: ?throw with IH_body obtain pc\<^sub>2 where pc\<^sub>2: "0 \ pc\<^sub>2 \ pc\<^sub>2 < size(compE\<^sub>2 body) \ \ caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and 2: "P \ ?\\<^sub>2' -jvm\ handle P D M' xa h\<^sub>3 [] ls\<^sub>3 pc\<^sub>2 ?frs\<^sub>2" using Call\<^sub>1.prems M'_code M'_xtab by (fastforce simp del:split_paired_Ex) have "handle P D M' xa h\<^sub>3 [] ls\<^sub>3 pc\<^sub>2 ?frs\<^sub>2 = handle P C M xa h\<^sub>3 (rev pvs @ Addr a # vs) ls\<^sub>2 ?pc\<^sub>2 frs" using pc\<^sub>2 M'_in_D by(auto simp add:handle_def) also have "\ = handle P C M xa h\<^sub>3 vs ls\<^sub>2 ?pc\<^sub>2 frs" using Call\<^sub>1.prems by(auto simp add:handle_append handle_Cons) finally have "?H ?pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF 1 2] by auto thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case (CallParamsThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 w h\<^sub>1 ls\<^sub>1 es es' h\<^sub>2 ls\<^sub>2 pvs ex es'' M') let ?\\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc)#frs)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>1 = "(None,h\<^sub>1,(w # vs, ls\<^sub>1, C,M,?pc\<^sub>1)#frs)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using CallParamsThrow\<^sub>1 by fastforce show ?case (is "?N \ (?eq \ (\pc\<^sub>2. ?H pc\<^sub>2))") proof show ?N by simp next { assume ?eq moreover have "PROP ?Ps es h\<^sub>1 ls\<^sub>1 es' h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 pvs xa es'' (w#vs) frs (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact ultimately have "\pc\<^sub>2. (?pc\<^sub>1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1 + size(compEs\<^sub>2 es) \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es ?pc\<^sub>1 (size vs + 1))) \ P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (w#vs) ls\<^sub>2 pc\<^sub>2 frs" (is "\pc\<^sub>2. ?PC pc\<^sub>2 \ ?Exec pc\<^sub>2") using CallParamsThrow\<^sub>1 by force then obtain pc\<^sub>2 where pc\<^sub>2: "?PC pc\<^sub>2" and 2: "?Exec pc\<^sub>2" by iprover have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF 1 2] CallParamsThrow\<^sub>1 by(auto simp:handle_Cons) hence "\pc\<^sub>2. ?H pc\<^sub>2" by iprover } thus "?eq \ (\pc\<^sub>2. ?H pc\<^sub>2)" by iprover qed next case (CallNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 M') have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" let ?xa = "addr_of_sys_xcpt NullPointer" have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 pvs xa (map Val pvs) (Null#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using CallNull\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>2,(rev pvs@Null#vs,ls\<^sub>2,C,M,?pc\<^sub>2)#frs)" using CallNull\<^sub>1 IH_es by fastforce also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (rev pvs@Null#vs) ls\<^sub>2 ?pc\<^sub>2 frs" using CallNull\<^sub>1.prems by(auto simp:split_beta handle_def nth_append simp del: split_paired_Ex) also have "handle P C M ?xa h\<^sub>2 (rev pvs@Null#vs) ls\<^sub>2 ?pc\<^sub>2 frs = handle P C M ?xa h\<^sub>2 vs ls\<^sub>2 ?pc\<^sub>2 frs" using CallNull\<^sub>1.prems by(auto simp:handle_Cons handle_append) finally show ?case by (auto intro: exI[where x = ?pc\<^sub>2]) next case CallObjThrow\<^sub>1 thus ?case by fastforce next case Block\<^sub>1 thus ?case by auto next case (Seq\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 w h\<^sub>1 ls\<^sub>1 e\<^sub>2 e\<^sub>2' h\<^sub>2 ls\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1)#frs)" have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using Seq\<^sub>1 by fastforce also have "P \ \ -jvm\ ?\\<^sub>1" using Seq\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". let ?pc\<^sub>2 = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>2)" have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 e\<^sub>2' h\<^sub>2 ls\<^sub>2 C M (?pc\<^sub>1+1) v xa vs frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note eval\<^sub>1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2)#frs)" using val Seq\<^sub>1.prems IH\<^sub>2 by fastforce finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof assume throw: ?throw then obtain pc\<^sub>2 where pc\<^sub>2: "?pc\<^sub>1+1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2 \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 (?pc\<^sub>1+1) (size vs))" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 vs ls\<^sub>2 pc\<^sub>2 frs" using IH\<^sub>2 Seq\<^sub>1.prems by fastforce have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by auto thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case SeqThrow\<^sub>1 thus ?case by fastforce next case (CondT\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 e\<^sub>1 e' h\<^sub>2 ls\<^sub>2 e\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1)#frs)" have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool(True)#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using CondT\<^sub>1 by (fastforce simp add: Int_Un_distrib) also have "P \ \ -jvm\ ?\\<^sub>1" using CondT\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". let ?pc\<^sub>1' = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2' = "?pc\<^sub>1' + 1 + length(compE\<^sub>2 e\<^sub>2)" show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note eval\<^sub>1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>1')#frs)" using val CondT\<^sub>1 by(fastforce simp:Int_Un_distrib) also have "P \ \ -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2')#frs)" using CondT\<^sub>1 by(auto simp:add.assoc) finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof let ?d = "size vs" let ?I = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>2 (?pc\<^sub>1'+1) ?d)" assume throw: ?throw moreover have "PROP ?P e\<^sub>1 h\<^sub>1 ls\<^sub>1 e' h\<^sub>2 ls\<^sub>2 C M (?pc\<^sub>1+1) v xa vs frs ?I" by fact ultimately obtain pc\<^sub>2 where pc\<^sub>2: "?pc\<^sub>1+1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1' \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) (size vs))" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 vs ls\<^sub>2 pc\<^sub>2 frs" using CondT\<^sub>1.prems by (fastforce simp:Int_Un_distrib) have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by auto thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case (CondF\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 e\<^sub>2 e' h\<^sub>2 ls\<^sub>2 e\<^sub>1) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?pc\<^sub>2 = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>1)+ 1" let ?pc\<^sub>2' = "?pc\<^sub>2 + length(compE\<^sub>2 e\<^sub>2)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>2)#frs)" have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool(False)#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using CondF\<^sub>1 by (fastforce simp add: Int_Un_distrib) also have "P \ \ -jvm\ ?\\<^sub>1" using CondF\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note eval\<^sub>1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2')#frs)" using val CondF\<^sub>1 by(fastforce simp:Int_Un_distrib) finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof let ?d = "size vs" let ?I = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) ?d)" assume throw: ?throw moreover have "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 e' h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>2 v xa vs frs ?I" by fact ultimately obtain pc\<^sub>2 where pc\<^sub>2: "?pc\<^sub>2 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2' \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>2 ?d)" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 vs ls\<^sub>2 pc\<^sub>2 frs" using CondF\<^sub>1.prems by(fastforce simp:Int_Un_distrib) have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by auto thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case (CondThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 f h\<^sub>1 ls\<^sub>1 e\<^sub>1 e\<^sub>2) let ?d = "size vs" let ?xt\<^sub>1 = "compxE\<^sub>2 e\<^sub>1 (pc+size(compE\<^sub>2 e)+1) ?d" let ?xt\<^sub>2 = "compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e)+size(compE\<^sub>2 e\<^sub>1)+2) ?d" let ?I = "I - (pcs ?xt\<^sub>1 \ pcs ?xt\<^sub>2)" have "pcs(compxE\<^sub>2 e pc ?d) \ pcs(?xt\<^sub>1 @ ?xt\<^sub>2) = {}" using CondThrow\<^sub>1.prems by (simp add:Int_Un_distrib) moreover have "PROP ?P e h\<^sub>0 ls\<^sub>0 (throw f) h\<^sub>1 ls\<^sub>1 C M pc v xa vs frs ?I" by fact ultimately show ?case using CondThrow\<^sub>1.prems by fastforce next case (WhileF\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 c) let ?pc = "pc + length(compE\<^sub>2 e)" let ?pc' = "?pc + length(compE\<^sub>2 c) + 3" have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(Bool False#vs,ls\<^sub>1,C,M,?pc)#frs)" using WhileF\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc')#frs)" using WhileF\<^sub>1 by (auto simp:add.assoc) also have "P \ \ -jvm\ (None,h\<^sub>1,(Unit#vs,ls\<^sub>1,C,M,?pc'+1)#frs)" using WhileF\<^sub>1.prems by (auto simp:eval_nat_numeral) finally show ?case by (simp add:add.assoc eval_nat_numeral) next case (WhileT\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 c v\<^sub>1 h\<^sub>2 ls\<^sub>2 e\<^sub>3 h\<^sub>3 ls\<^sub>3) let ?pc = "pc + length(compE\<^sub>2 e)" let ?pc' = "?pc + length(compE\<^sub>2 c) + 1" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?\\<^sub>2 = "(None,h\<^sub>2,(vs,ls\<^sub>2,C,M,pc)#frs)" have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool True#vs,ls\<^sub>1,C,M,?pc)#frs)" using WhileT\<^sub>1 by fastforce also have "P \ \ -jvm\ (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc+1)#frs)" using WhileT\<^sub>1.prems by auto also have "P \ \ -jvm\ (None,h\<^sub>2,(v\<^sub>1#vs,ls\<^sub>2,C,M,?pc')#frs)" using WhileT\<^sub>1 by(fastforce) also have "P \ \ -jvm\ ?\\<^sub>2" using WhileT\<^sub>1.prems by auto finally have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>2". show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>2 -jvm\ (None,h\<^sub>3,(v#vs,ls\<^sub>3,C,M,?pc'+3)#frs)" using val WhileT\<^sub>1 by (auto simp add:add.assoc eval_nat_numeral) finally show ?trans by(simp add:add.assoc eval_nat_numeral) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof assume throw: ?throw moreover have "PROP ?P (while (e) c) h\<^sub>2 ls\<^sub>2 e\<^sub>3 h\<^sub>3 ls\<^sub>3 C M pc v xa vs frs I" by fact ultimately obtain pc\<^sub>2 where pc\<^sub>2: "pc \ pc\<^sub>2 \ pc\<^sub>2 < ?pc'+3 \ \ caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 (while (e) c) pc (size vs))" and 2: "P \ ?\\<^sub>2 -jvm\ handle P C M xa h\<^sub>3 vs ls\<^sub>3 pc\<^sub>2 frs" using WhileT\<^sub>1.prems by (auto simp:add.assoc eval_nat_numeral) have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF 1 2] by auto thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case WhileCondThrow\<^sub>1 thus ?case by fastforce next case (WhileBodyThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1 c e' h\<^sub>2 ls\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1)#frs)" have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool(True)#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using WhileBodyThrow\<^sub>1 by (fastforce simp add: Int_Un_distrib) also have "P \ \ -jvm\ ?\\<^sub>1" using WhileBodyThrow\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". let ?pc\<^sub>1' = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 c)" show ?case (is "?Norm \ ?Err") proof show ?Norm by simp next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof assume throw: ?throw moreover have "PROP ?P c h\<^sub>1 ls\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 C M (?pc\<^sub>1+1) v xa vs frs (I - pcs (compxE\<^sub>2 e pc (size vs)))" by fact ultimately obtain pc\<^sub>2 where pc\<^sub>2: "?pc\<^sub>1+1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1' \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 c (?pc\<^sub>1+1) (size vs))" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 vs ls\<^sub>2 pc\<^sub>2 frs" using WhileBodyThrow\<^sub>1.prems by (fastforce simp:Int_Un_distrib) have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by auto thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case (Throw\<^sub>1 e h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1) let ?pc = "pc + size(compE\<^sub>2 e)" show ?case (is "?Norm \ ?Err") proof show ?Norm by simp next show ?Err (is "?throw \ (\pc\<^sub>1. ?H pc\<^sub>1)") proof assume ?throw hence "P \ (None, h\<^sub>0, (vs, ls\<^sub>0, C, M, pc) # frs) -jvm\ (None, h\<^sub>1, (Addr xa#vs, ls\<^sub>1, C, M, ?pc) # frs)" using Throw\<^sub>1 by fastforce also have "P \ \ -jvm\ handle P C M xa h\<^sub>1 (Addr xa#vs) ls\<^sub>1 ?pc frs" using Throw\<^sub>1.prems by(auto simp add:handle_def) also have "handle P C M xa h\<^sub>1 (Addr xa#vs) ls\<^sub>1 ?pc frs = handle P C M xa h\<^sub>1 vs ls\<^sub>1 ?pc frs" using Throw\<^sub>1.prems by(auto simp add:handle_Cons) finally have "?H ?pc" by simp thus "\pc\<^sub>1. ?H pc\<^sub>1" by iprover qed qed next case (ThrowNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 h\<^sub>1 ls\<^sub>1) let ?pc = "pc + size(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt NullPointer" show ?case (is "?Norm \ ?Err") proof show ?Norm by simp next show ?Err (is "?throw \ (\pc\<^sub>1. ?H pc\<^sub>1)") proof assume throw: ?throw have "P \ (None, h\<^sub>0, (vs, ls\<^sub>0, C, M, pc) # frs) -jvm\ (None, h\<^sub>1, (Null#vs, ls\<^sub>1, C, M, ?pc) # frs)" using ThrowNull\<^sub>1 by fastforce also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc frs" using ThrowNull\<^sub>1.prems by(auto simp add:handle_def) also have "handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc frs = handle P C M ?xa h\<^sub>1 vs ls\<^sub>1 ?pc frs" using ThrowNull\<^sub>1.prems by(auto simp add:handle_Cons) finally have "?H ?pc" using throw by simp thus "\pc\<^sub>1. ?H pc\<^sub>1" by iprover qed qed next case ThrowThrow\<^sub>1 thus ?case by fastforce next case (Try\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 Ci i e\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>1' = "?pc\<^sub>1 + 2 + length(compE\<^sub>2 e\<^sub>2)" have "P,C,M \ compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs) / I,size vs" by fact hence "P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs) / {pc..2 e\<^sub>1)},size vs" using Try\<^sub>1.prems by (fastforce simp:beforex_def split:if_split_asm) hence "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs) -jvm\ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" using Try\<^sub>1 by auto also have "P \ \ -jvm\ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1')#frs)" using Try\<^sub>1.prems by auto finally show ?case by (auto simp:add.assoc) next case (TryCatch\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1 D fs Ci i e\<^sub>2 e\<^sub>2' h\<^sub>2 ls\<^sub>2) let ?e = "try e\<^sub>1 catch(Ci i) e\<^sub>2" let ?xt = "compxE\<^sub>2 ?e pc (size vs)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?ls\<^sub>1 = "ls\<^sub>1[i := Addr a]" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>1' = "?pc\<^sub>1 + 2" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,?ls\<^sub>1,C,M, ?pc\<^sub>1') # frs)" have I: "{pc..2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \ I" and beforex: "P,C,M \ ?xt/I,size vs" by fact+ have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,((Addr a)#vs,ls\<^sub>1,C,M, ?pc\<^sub>1+1) # frs)" proof - have "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 (Throw a) h\<^sub>1 ls\<^sub>1 C M pc w a vs frs {pc..2 e\<^sub>1)}" by fact moreover have "P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs)/{pc..1},size vs" using beforex I pcs_subset by(force elim!: beforex_appendD1) ultimately have "\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < ?pc\<^sub>1 \ \ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs)) \ P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 vs ls\<^sub>1 pc\<^sub>1 frs" using TryCatch\<^sub>1.prems by auto then obtain pc\<^sub>1 where pc\<^sub>1_in_e\<^sub>1: "pc \ pc\<^sub>1" "pc\<^sub>1 < ?pc\<^sub>1" and pc\<^sub>1_not_caught: "\ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs))" and 0: "P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 vs ls\<^sub>1 pc\<^sub>1 frs" by iprover from beforex obtain xt\<^sub>0 xt\<^sub>1 where ex_tab: "ex_table_of P C M = xt\<^sub>0 @ ?xt @ xt\<^sub>1" and disj: "pcs xt\<^sub>0 \ I = {}" by(auto simp:beforex_def) have hp: "h\<^sub>1 a = Some (D, fs)" "P\<^sub>1 \ D \\<^sup>* Ci" by fact+ have "pc\<^sub>1 \ pcs xt\<^sub>0" using pc\<^sub>1_in_e\<^sub>1 I disj by auto with pc\<^sub>1_in_e\<^sub>1 pc\<^sub>1_not_caught hp show ?thesis using ex_tab 0 by(simp add:handle_def matches_ex_entry_def) qed also have "P \ \ -jvm\ ?\\<^sub>1" using TryCatch\<^sub>1 by auto finally have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" . let ?pc\<^sub>2 = "?pc\<^sub>1' + length(compE\<^sub>2 e\<^sub>2)" let ?I\<^sub>2 = "{?pc\<^sub>1' ..< ?pc\<^sub>2}" have "P,C,M \ compxE\<^sub>2 ?e pc (size vs) / I,size vs" by fact hence beforex\<^sub>2: "P,C,M \ compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1' (size vs) / ?I\<^sub>2, size vs" using I pcs_subset[of _ ?pc\<^sub>1'] by(auto elim!:beforex_appendD2) have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ?ls\<^sub>1 e\<^sub>2' h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1' v xa vs frs ?I\<^sub>2" by fact show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2)#frs)" using val beforex\<^sub>2 IH\<^sub>2 TryCatch\<^sub>1.prems by auto finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof assume throw: ?throw then obtain pc\<^sub>2 where pc\<^sub>2: "?pc\<^sub>1+2 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2 \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1' (size vs))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 vs ls\<^sub>2 pc\<^sub>2 frs" using IH\<^sub>2 beforex\<^sub>2 TryCatch\<^sub>1.prems by auto have "?H pc\<^sub>2" using pc\<^sub>2 jvm_trans[OF 1 2] by (simp add:match_ex_entry) (fastforce) thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case (TryThrow\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 a h\<^sub>1 ls\<^sub>1 D fs Ci i e\<^sub>2) let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?e = "try e\<^sub>1 catch(Ci i) e\<^sub>2" let ?xt = "compxE\<^sub>2 ?e pc (size vs)" have I: "{pc..2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \ I" and beforex: "P,C,M \ ?xt/I,size vs" by fact+ have "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 (Throw a) h\<^sub>1 ls\<^sub>1 C M pc w a vs frs {pc..2 e\<^sub>1)}" by fact moreover have "P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs)/{pc..1},size vs" using beforex I pcs_subset by(force elim!: beforex_appendD1) ultimately have "\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < ?pc\<^sub>1 \ \ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs)) \ P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 vs ls\<^sub>1 pc\<^sub>1 frs" using TryThrow\<^sub>1.prems by auto then obtain pc\<^sub>1 where pc\<^sub>1_in_e\<^sub>1: "pc \ pc\<^sub>1" "pc\<^sub>1 < ?pc\<^sub>1" and pc\<^sub>1_not_caught: "\ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs))" and 0: "P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 vs ls\<^sub>1 pc\<^sub>1 frs" by iprover show ?case (is "?N \ (?eq \ (\pc\<^sub>2. ?H pc\<^sub>2))") proof show ?N by simp next { assume ?eq with TryThrow\<^sub>1 pc\<^sub>1_in_e\<^sub>1 pc\<^sub>1_not_caught 0 have "?H pc\<^sub>1" by (simp add:match_ex_entry) auto hence "\pc\<^sub>2. ?H pc\<^sub>2" by iprover } thus "?eq \ (\pc\<^sub>2. ?H pc\<^sub>2)" by iprover qed next case Nil\<^sub>1 thus ?case by simp next case (Cons\<^sub>1 e h\<^sub>0 ls\<^sub>0 v h\<^sub>1 ls\<^sub>1 es fs h\<^sub>2 ls\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc)#frs)" let ?\\<^sub>1 = "(None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,?pc\<^sub>1)#frs)" have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using Cons\<^sub>1 by fastforce let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" have IHs: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 fs h\<^sub>2 ls\<^sub>2 C M ?pc\<^sub>1 (tl ws) xa es' (v#vs) frs (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(rev(ws) @ vs,ls\<^sub>2,C,M,?pc\<^sub>2)#frs)" using val IHs Cons\<^sub>1.prems by fastforce finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof assume throw: ?throw then obtain pc\<^sub>2 where pc\<^sub>2: "?pc\<^sub>1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2 \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es ?pc\<^sub>1 (size vs + 1))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (v#vs) ls\<^sub>2 pc\<^sub>2 frs" using IHs Cons\<^sub>1.prems by(fastforce simp:Cons_eq_append_conv neq_Nil_conv) have "?H pc\<^sub>2" using Cons\<^sub>1.prems pc\<^sub>2 jvm_trans[OF 1 2] by (auto simp add: handle_Cons) thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case ConsThrow\<^sub>1 thus ?case by (fastforce simp:Cons_eq_append_conv) qed (*>*) (*FIXME move! *) lemma atLeast0AtMost[simp]: "{0::nat..n} = {..n}" by auto lemma atLeast0LessThan[simp]: "{0::nat.. addr option" where "exception (Throw a) = Some a" | "exception e = None" lemma comp\<^sub>2_correct: assumes "method": "P\<^sub>1 \ C sees M:Ts\T = body in C" and eval: "P\<^sub>1 \\<^sub>1 \body,(h,ls)\ \ \e',(h',ls')\" shows "compP\<^sub>2 P\<^sub>1 \ (None,h,[([],ls,C,M,0)]) -jvm\ (exception e',h',[])" (*<*) (is "_ \ ?\\<^sub>0 -jvm\ ?\\<^sub>1") proof - let ?P = "compP\<^sub>2 P\<^sub>1" have code: "?P,C,M,0 \ compE\<^sub>2 body" using beforeM[OF "method"] by auto have xtab: "?P,C,M \ compxE\<^sub>2 body 0 (size[])/{..2 body)},size[]" using beforexM[OF "method"] by auto \ \Distinguish if e' is a value or an exception\ { fix v assume [simp]: "e' = Val v" have "?P \ ?\\<^sub>0 -jvm\ (None,h',[([v],ls',C,M,size(compE\<^sub>2 body))])" using Jcc[OF eval code xtab] by auto also have "?P \ \ -jvm\ ?\\<^sub>1" using beforeM[OF "method"] by auto finally have ?thesis . } moreover { fix a assume [simp]: "e' = Throw a" obtain pc where pc: "0 \ pc \ pc < size(compE\<^sub>2 body) \ \ caught ?P pc h' a (compxE\<^sub>2 body 0 0)" and 1: "?P \ ?\\<^sub>0 -jvm\ handle ?P C M a h' [] ls' pc []" using Jcc[OF eval code xtab] by fastforce from pc have "handle ?P C M a h' [] ls' pc [] = ?\\<^sub>1" using xtab "method" by(auto simp:handle_def compMb\<^sub>2_def) with 1 have ?thesis by simp } ultimately show ?thesis using eval\<^sub>1_final[OF eval] by(auto simp:final_def) qed (*>*) end diff --git a/thys/Jinja/Compiler/J1.thy b/thys/Jinja/Compiler/J1.thy --- a/thys/Jinja/Compiler/J1.thy +++ b/thys/Jinja/Compiler/J1.thy @@ -1,225 +1,221 @@ (* Title: Jinja/Compiler/J1.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) chapter \Compilation \label{cha:comp}\ section \An Intermediate Language\ theory J1 imports "../J/BigStep" begin type_synonym expr\<^sub>1 = "nat exp" type_synonym J\<^sub>1_prog = "expr\<^sub>1 prog" type_synonym state\<^sub>1 = "heap \ (val list)" primrec max_vars :: "'a exp \ nat" and max_varss :: "'a exp list \ nat" where "max_vars(new C) = 0" | "max_vars(Cast C e) = max_vars e" | "max_vars(Val v) = 0" | "max_vars(e\<^sub>1 \bop\ e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)" | "max_vars(Var V) = 0" | "max_vars(V:=e) = max_vars e" | "max_vars(e\F{D}) = max_vars e" | "max_vars(FAss e\<^sub>1 F D e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)" | "max_vars(e\M(es)) = max (max_vars e) (max_varss es)" | "max_vars({V:T; e}) = max_vars e + 1" | "max_vars(e\<^sub>1;;e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)" | "max_vars(if (e) e\<^sub>1 else e\<^sub>2) = max (max_vars e) (max (max_vars e\<^sub>1) (max_vars e\<^sub>2))" | "max_vars(while (b) e) = max (max_vars b) (max_vars e)" | "max_vars(throw e) = max_vars e" | "max_vars(try e\<^sub>1 catch(C V) e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2 + 1)" | "max_varss [] = 0" | "max_varss (e#es) = max (max_vars e) (max_varss es)" inductive eval\<^sub>1 :: "J\<^sub>1_prog \ expr\<^sub>1 \ state\<^sub>1 \ expr\<^sub>1 \ state\<^sub>1 \ bool" ("_ \\<^sub>1 ((1\_,/_\) \/ (1\_,/_\))" [51,0,0,0,0] 81) and evals\<^sub>1 :: "J\<^sub>1_prog \ expr\<^sub>1 list \ state\<^sub>1 \ expr\<^sub>1 list \ state\<^sub>1 \ bool" ("_ \\<^sub>1 ((1\_,/_\) [\]/ (1\_,/_\))" [51,0,0,0,0] 81) for P :: J\<^sub>1_prog where New\<^sub>1: "\ new_Addr h = Some a; P \ C has_fields FDTs; h' = h(a\(C,init_fields FDTs)) \ \ P \\<^sub>1 \new C,(h,l)\ \ \addr a,(h',l)\" | NewFail\<^sub>1: "new_Addr h = None \ P \\<^sub>1 \new C, (h,l)\ \ \THROW OutOfMemory,(h,l)\" | Cast\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,l)\; h a = Some(D,fs); P \ D \\<^sup>* C \ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \addr a,(h,l)\" | CastNull\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \null,s\<^sub>1\" | CastFail\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,l)\; h a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \THROW ClassCast,(h,l)\" | CastThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | Val\<^sub>1: "P \\<^sub>1 \Val v,s\ \ \Val v,s\" | BinOp\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v\<^sub>2,s\<^sub>2\; binop(bop,v\<^sub>1,v\<^sub>2) = Some v \ \ P \\<^sub>1 \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\ \ \Val v,s\<^sub>2\" | BinOpThrow\<^sub>1\<^sub>1: "P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \throw e,s\<^sub>1\ \ P \\<^sub>1 \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0\ \ \throw e,s\<^sub>1\" | BinOpThrow\<^sub>2\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \throw e,s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\ \ \throw e,s\<^sub>2\" | Var\<^sub>1: "\ ls!i = v; i < size ls \ \ P \\<^sub>1 \Var i,(h,ls)\ \ \Val v,(h,ls)\" | LAss\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \Val v,(h,ls)\; i < size ls; ls' = ls[i := v] \ \ P \\<^sub>1 \i:= e,s\<^sub>0\ \ \unit,(h,ls')\" | LAssThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \i:= e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAcc\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,ls)\; h a = Some(C,fs); fs(F,D) = Some v \ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \Val v,(h,ls)\" | FAccNull\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \THROW NullPointer,s\<^sub>1\" | FAccThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAss\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v,(h\<^sub>2,l\<^sub>2)\; h\<^sub>2 a = Some(C,fs); fs' = fs((F,D)\v); h\<^sub>2' = h\<^sub>2(a\(C,fs')) \ \ P \\<^sub>1 \e\<^sub>1\F{D}:= e\<^sub>2,s\<^sub>0\ \ \unit,(h\<^sub>2',l\<^sub>2)\" | FAssNull\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \null,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v,s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>1\F{D}:= e\<^sub>2,s\<^sub>0\ \ \THROW NullPointer,s\<^sub>2\" | FAssThrow\<^sub>1\<^sub>1: "P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e\<^sub>1\F{D}:= e\<^sub>2,s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAssThrow\<^sub>2\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \throw e',s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>1\F{D}:= e\<^sub>2,s\<^sub>0\ \ \throw e',s\<^sub>2\" | CallObjThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \throw e',s\<^sub>1\" | CallNull\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \map Val vs,s\<^sub>2\ \ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \THROW NullPointer,s\<^sub>2\" | Call\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \map Val vs,(h\<^sub>2,ls\<^sub>2)\; h\<^sub>2 a = Some(C,fs); P \ C sees M:Ts\T = body in D; size vs = size Ts; ls\<^sub>2' = (Addr a) # vs @ replicate (max_vars body) undefined; P \\<^sub>1 \body,(h\<^sub>2,ls\<^sub>2')\ \ \e',(h\<^sub>3,ls\<^sub>3)\ \ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \e',(h\<^sub>3,ls\<^sub>2)\" | CallParamsThrow\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \es',s\<^sub>2\; es' = map Val vs @ throw ex # es\<^sub>2 \ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \throw ex,s\<^sub>2\" | Block\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \e',s\<^sub>1\ \ P \\<^sub>1 \Block i T e,s\<^sub>0\ \ \e',s\<^sub>1\" | Seq\<^sub>1: "\ P \\<^sub>1 \e\<^sub>0,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \e\<^sub>1,s\<^sub>1\ \ \e\<^sub>2,s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>0;;e\<^sub>1,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" | SeqThrow\<^sub>1: "P \\<^sub>1 \e\<^sub>0,s\<^sub>0\ \ \throw e,s\<^sub>1\ \ P \\<^sub>1 \e\<^sub>0;;e\<^sub>1,s\<^sub>0\ \ \throw e,s\<^sub>1\" | CondT\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \\<^sub>1 \e\<^sub>1,s\<^sub>1\ \ \e',s\<^sub>2\ \ \ P \\<^sub>1 \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \ \e',s\<^sub>2\" | CondF\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \false,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \e',s\<^sub>2\ \ \ P \\<^sub>1 \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \ \e',s\<^sub>2\" | CondThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\ \ \throw e',s\<^sub>1\" | WhileF\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \false,s\<^sub>1\ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \unit,s\<^sub>1\" | WhileT\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \\<^sub>1 \c,s\<^sub>1\ \ \Val v\<^sub>1,s\<^sub>2\; P \\<^sub>1 \while (e) c,s\<^sub>2\ \ \e\<^sub>3,s\<^sub>3\ \ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \e\<^sub>3,s\<^sub>3\" | WhileCondThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \throw e',s\<^sub>1\" | WhileBodyThrow\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \\<^sub>1 \c,s\<^sub>1\ \ \throw e',s\<^sub>2\\ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \throw e',s\<^sub>2\" | Throw\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,s\<^sub>1\ \ P \\<^sub>1 \throw e,s\<^sub>0\ \ \Throw a,s\<^sub>1\" | ThrowNull\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \\<^sub>1 \throw e,s\<^sub>0\ \ \THROW NullPointer,s\<^sub>1\" | ThrowThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \throw e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | Try\<^sub>1: "P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\ \ P \\<^sub>1 \try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\" | TryCatch\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Throw a,(h\<^sub>1,ls\<^sub>1)\; h\<^sub>1 a = Some(D,fs); P \ D \\<^sup>* C; i < length ls\<^sub>1; P \\<^sub>1 \e\<^sub>2,(h\<^sub>1,ls\<^sub>1[i:=Addr a])\ \ \e\<^sub>2',(h\<^sub>2,ls\<^sub>2)\ \ \ P \\<^sub>1 \try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',(h\<^sub>2,ls\<^sub>2)\" | TryThrow\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Throw a,(h\<^sub>1,ls\<^sub>1)\; h\<^sub>1 a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \\<^sub>1 \try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\ \ \Throw a,(h\<^sub>1,ls\<^sub>1)\" | Nil\<^sub>1: "P \\<^sub>1 \[],s\ [\] \[],s\" | Cons\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \es',s\<^sub>2\ \ \ P \\<^sub>1 \e#es,s\<^sub>0\ [\] \Val v # es',s\<^sub>2\" | ConsThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e#es,s\<^sub>0\ [\] \throw e' # es, s\<^sub>1\" (*<*) lemmas eval\<^sub>1_evals\<^sub>1_induct = eval\<^sub>1_evals\<^sub>1.induct [split_format (complete)] and eval\<^sub>1_evals\<^sub>1_inducts = eval\<^sub>1_evals\<^sub>1.inducts [split_format (complete)] (*>*) lemma eval\<^sub>1_preserves_len: "P \\<^sub>1 \e\<^sub>0,(h\<^sub>0,ls\<^sub>0)\ \ \e\<^sub>1,(h\<^sub>1,ls\<^sub>1)\ \ length ls\<^sub>0 = length ls\<^sub>1" and evals\<^sub>1_preserves_len: "P \\<^sub>1 \es\<^sub>0,(h\<^sub>0,ls\<^sub>0)\ [\] \es\<^sub>1,(h\<^sub>1,ls\<^sub>1)\ \ length ls\<^sub>0 = length ls\<^sub>1" (*<*)by (induct rule:eval\<^sub>1_evals\<^sub>1_inducts, simp_all)(*>*) lemma evals\<^sub>1_preserves_elen: "\es' s s'. P \\<^sub>1 \es,s\ [\] \es',s'\ \ length es = length es'" -(*<*) -apply(induct es type:list) -apply (auto elim:evals\<^sub>1.cases) -done -(*>*) +(*<*)by(induct es type:list) (auto elim:evals\<^sub>1.cases)(*>*) lemma eval\<^sub>1_final: "P \\<^sub>1 \e,s\ \ \e',s'\ \ final e'" and evals\<^sub>1_final: "P \\<^sub>1 \es,s\ [\] \es',s'\ \ finals es'" (*<*)by(induct rule:eval\<^sub>1_evals\<^sub>1.inducts, simp_all)(*>*) end diff --git a/thys/Jinja/Compiler/J1WellForm.thy b/thys/Jinja/Compiler/J1WellForm.thy --- a/thys/Jinja/Compiler/J1WellForm.thy +++ b/thys/Jinja/Compiler/J1WellForm.thy @@ -1,223 +1,209 @@ (* Title: Jinja/Compiler/WellType1.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Well-Formedness of Intermediate Language\ theory J1WellForm imports "../J/JWellForm" J1 begin subsection "Well-Typedness" type_synonym env\<^sub>1 = "ty list" \ \type environment indexed by variable number\ inductive WT\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1, expr\<^sub>1 , ty ] \ bool" ("(_,_ \\<^sub>1/ _ :: _)" [51,51,51]50) and WTs\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1, expr\<^sub>1 list, ty list] \ bool" ("(_,_ \\<^sub>1/ _ [::] _)" [51,51,51]50) for P :: J\<^sub>1_prog where WTNew\<^sub>1: "is_class P C \ P,E \\<^sub>1 new C :: Class C" | WTCast\<^sub>1: "\ P,E \\<^sub>1 e :: Class D; is_class P C; P \ C \\<^sup>* D \ P \ D \\<^sup>* C \ \ P,E \\<^sub>1 Cast C e :: Class C" | WTVal\<^sub>1: "typeof v = Some T \ P,E \\<^sub>1 Val v :: T" | WTVar\<^sub>1: "\ E!i = T; i < size E \ \ P,E \\<^sub>1 Var i :: T" | WTBinOp\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1 :: T\<^sub>1; P,E \\<^sub>1 e\<^sub>2 :: T\<^sub>2; case bop of Eq \ (P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1) \ T = Boolean | Add \ T\<^sub>1 = Integer \ T\<^sub>2 = Integer \ T = Integer \ \ P,E \\<^sub>1 e\<^sub>1 \bop\ e\<^sub>2 :: T" | WTLAss\<^sub>1: "\ E!i = T; i < size E; P,E \\<^sub>1 e :: T'; P \ T' \ T \ \ P,E \\<^sub>1 i:=e :: Void" | WTFAcc\<^sub>1: "\ P,E \\<^sub>1 e :: Class C; P \ C sees F:T in D \ \ P,E \\<^sub>1 e\F{D} :: T" | WTFAss\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1 :: Class C; P \ C sees F:T in D; P,E \\<^sub>1 e\<^sub>2 :: T'; P \ T' \ T \ \ P,E \\<^sub>1 e\<^sub>1\F{D} := e\<^sub>2 :: Void" | WTCall\<^sub>1: "\ P,E \\<^sub>1 e :: Class C; P \ C sees M:Ts' \ T = m in D; P,E \\<^sub>1 es [::] Ts; P \ Ts [\] Ts' \ \ P,E \\<^sub>1 e\M(es) :: T" | WTBlock\<^sub>1: "\ is_type P T; P,E@[T] \\<^sub>1 e::T' \ \ P,E \\<^sub>1 {i:T; e} :: T'" | WTSeq\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1::T\<^sub>1; P,E \\<^sub>1 e\<^sub>2::T\<^sub>2 \ \ P,E \\<^sub>1 e\<^sub>1;;e\<^sub>2 :: T\<^sub>2" | WTCond\<^sub>1: "\ P,E \\<^sub>1 e :: Boolean; P,E \\<^sub>1 e\<^sub>1::T\<^sub>1; P,E \\<^sub>1 e\<^sub>2::T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1; P \ T\<^sub>1 \ T\<^sub>2 \ T = T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ T = T\<^sub>1 \ \ P,E \\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 :: T" | WTWhile\<^sub>1: "\ P,E \\<^sub>1 e :: Boolean; P,E \\<^sub>1 c::T \ \ P,E \\<^sub>1 while (e) c :: Void" | WTThrow\<^sub>1: "P,E \\<^sub>1 e :: Class C \ P,E \\<^sub>1 throw e :: Void" | WTTry\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1 :: T; P,E@[Class C] \\<^sub>1 e\<^sub>2 :: T; is_class P C \ \ P,E \\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 :: T" | WTNil\<^sub>1: "P,E \\<^sub>1 [] [::] []" | WTCons\<^sub>1: "\ P,E \\<^sub>1 e :: T; P,E \\<^sub>1 es [::] Ts \ \ P,E \\<^sub>1 e#es [::] T#Ts" (*<*) declare WT\<^sub>1_WTs\<^sub>1.intros[intro!] declare WTNil\<^sub>1[iff] lemmas WT\<^sub>1_WTs\<^sub>1_induct = WT\<^sub>1_WTs\<^sub>1.induct [split_format (complete)] and WT\<^sub>1_WTs\<^sub>1_inducts = WT\<^sub>1_WTs\<^sub>1.inducts [split_format (complete)] inductive_cases eee[elim!]: "P,E \\<^sub>1 Val v :: T" "P,E \\<^sub>1 Var i :: T" "P,E \\<^sub>1 Cast D e :: T" "P,E \\<^sub>1 i:=e :: T" "P,E \\<^sub>1 {i:U; e} :: T" "P,E \\<^sub>1 e\<^sub>1;;e\<^sub>2 :: T" "P,E \\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 :: T" "P,E \\<^sub>1 while (e) c :: T" "P,E \\<^sub>1 throw e :: T" "P,E \\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 :: T" "P,E \\<^sub>1 e\F{D} :: T" "P,E \\<^sub>1 e\<^sub>1\F{D}:=e\<^sub>2 :: T" "P,E \\<^sub>1 e\<^sub>1 \bop\ e\<^sub>2 :: T" "P,E \\<^sub>1 new C :: T" "P,E \\<^sub>1 e\M(es) :: T" "P,E \\<^sub>1 [] [::] Ts" "P,E \\<^sub>1 e#es [::] Ts" (*>*) lemma WTs\<^sub>1_same_size: "\Ts. P,E \\<^sub>1 es [::] Ts \ size es = size Ts" (*<*)by (induct es type:list) auto(*>*) lemma WT\<^sub>1_unique: "P,E \\<^sub>1 e :: T\<^sub>1 \ (\T\<^sub>2. P,E \\<^sub>1 e :: T\<^sub>2 \ T\<^sub>1 = T\<^sub>2)" and "P,E \\<^sub>1 es [::] Ts\<^sub>1 \ (\Ts\<^sub>2. P,E \\<^sub>1 es [::] Ts\<^sub>2 \ Ts\<^sub>1 = Ts\<^sub>2)" (*<*) -apply(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) -apply blast -apply blast -apply clarsimp -apply blast -apply clarsimp -apply(case_tac bop) -apply clarsimp -apply clarsimp -apply blast -apply (blast dest:sees_field_idemp sees_field_fun) -apply blast -apply (blast dest:sees_method_idemp sees_method_fun) -apply blast -apply blast -apply blast -apply blast -apply clarify -apply blast -apply blast -apply blast -done +proof(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) + case WTVal\<^sub>1 then show ?case by clarsimp +next + case (WTBinOp\<^sub>1 E e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 bop T) + then show ?case by(case_tac bop) force+ +next + case WTFAcc\<^sub>1 then show ?case + by (blast dest:sees_field_idemp sees_field_fun) +next + case WTCall\<^sub>1 then show ?case + by (blast dest:sees_method_idemp sees_method_fun) +qed blast+ (*>*) lemma assumes wf: "wf_prog p P" shows WT\<^sub>1_is_type: "P,E \\<^sub>1 e :: T \ set E \ types P \ is_type P T" and "P,E \\<^sub>1 es [::] Ts \ True" (*<*) -apply(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) -apply simp -apply simp -apply (simp add:typeof_lit_is_type) -apply (blast intro:nth_mem) -apply(simp split:bop.splits) -apply simp -apply (simp add:sees_field_is_type[OF _ wf]) -apply simp -apply(fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def) -apply simp -apply simp -apply blast -apply simp -apply simp -apply simp -apply simp -apply simp -done +proof(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) + case WTVal\<^sub>1 then show ?case by (simp add:typeof_lit_is_type) +next + case WTVar\<^sub>1 then show ?case by (blast intro:nth_mem) +next + case WTBinOp\<^sub>1 then show ?case by (simp split:bop.splits) +next + case WTFAcc\<^sub>1 then show ?case + by (simp add:sees_field_is_type[OF _ wf]) +next + case WTCall\<^sub>1 then show ?case + by (fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def) +next + case WTCond\<^sub>1 then show ?case by blast +qed simp+ (*>*) subsection\Well-formedness\ \ \Indices in blocks increase by 1\ primrec \ :: "expr\<^sub>1 \ nat \ bool" and \s :: "expr\<^sub>1 list \ nat \ bool" where "\ (new C) i = True" | "\ (Cast C e) i = \ e i" | "\ (Val v) i = True" | "\ (e\<^sub>1 \bop\ e\<^sub>2) i = (\ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (Var j) i = True" | "\ (e\F{D}) i = \ e i" | "\ (j:=e) i = \ e i" | "\ (e\<^sub>1\F{D} := e\<^sub>2) i = (\ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (e\M(es)) i = (\ e i \ \s es i)" | "\ ({j:T ; e}) i = (i = j \ \ e (i+1))" | "\ (e\<^sub>1;;e\<^sub>2) i = (\ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (if (e) e\<^sub>1 else e\<^sub>2) i = (\ e i \ \ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (throw e) i = \ e i" | "\ (while (e) c) i = (\ e i \ \ c i)" | "\ (try e\<^sub>1 catch(C j) e\<^sub>2) i = (\ e\<^sub>1 i \ i=j \ \ e\<^sub>2 (i+1))" | "\s [] i = True" | "\s (e#es) i = (\ e i \ \s es i)" definition wf_J\<^sub>1_mdecl :: "J\<^sub>1_prog \ cname \ expr\<^sub>1 mdecl \ bool" where "wf_J\<^sub>1_mdecl P C \ \(M,Ts,T,body). (\T'. P,Class C#Ts \\<^sub>1 body :: T' \ P \ T' \ T) \ \ body \{..size Ts}\ \ \ body (size Ts + 1)" lemma wf_J\<^sub>1_mdecl[simp]: "wf_J\<^sub>1_mdecl P C (M,Ts,T,body) \ ((\T'. P,Class C#Ts \\<^sub>1 body :: T' \ P \ T' \ T) \ \ body \{..size Ts}\ \ \ body (size Ts + 1))" (*<*)by (simp add:wf_J\<^sub>1_mdecl_def)(*>*) abbreviation "wf_J\<^sub>1_prog == wf_prog wf_J\<^sub>1_mdecl" end diff --git a/thys/Jinja/Compiler/PCompiler.thy b/thys/Jinja/Compiler/PCompiler.thy --- a/thys/Jinja/Compiler/PCompiler.thy +++ b/thys/Jinja/Compiler/PCompiler.thy @@ -1,278 +1,350 @@ (* Title: Jinja/Compiler/PCompiler.thy Author: Tobias Nipkow Copyright TUM 2003 *) section \Program Compilation\ theory PCompiler imports "../Common/WellForm" begin definition compM :: "('a \ 'b) \ 'a mdecl \ 'b mdecl" where "compM f \ \(M, Ts, T, m). (M, Ts, T, f m)" definition compC :: "('a \ 'b) \ 'a cdecl \ 'b cdecl" where "compC f \ \(C,D,Fdecls,Mdecls). (C,D,Fdecls, map (compM f) Mdecls)" definition compP :: "('a \ 'b) \ 'a prog \ 'b prog" where "compP f \ map (compC f)" text\Compilation preserves the program structure. Therfore lookup functions either commute with compilation (like method lookup) or are preserved by it (like the subclass relation).\ lemma map_of_map4: "map_of (map (\(x,a,b,c).(x,a,b,f c)) ts) = map_option (\(a,b,c).(a,b,f c)) \ (map_of ts)" (*<*) -apply(induct ts) - apply simp -apply(rule ext) -apply fastforce -done +proof(induct ts) + case Nil then show ?case by simp +qed fastforce (*>*) lemma class_compP: "class P C = Some (D, fs, ms) \ class (compP f P) C = Some (D, fs, map (compM f) ms)" (*<*)by(simp add:class_def compP_def compC_def map_of_map4)(*>*) lemma class_compPD: "class (compP f P) C = Some (D, fs, cms) \ \ms. class P C = Some(D,fs,ms) \ cms = map (compM f) ms" (*<*)by(clarsimp simp add:class_def compP_def compC_def map_of_map4)(*>*) lemma [simp]: "is_class (compP f P) C = is_class P C" (*<*)by(auto simp:is_class_def dest: class_compP class_compPD)(*>*) lemma [simp]: "class (compP f P) C = map_option (\c. snd(compC f (C,c))) (class P C)" (*<*) -apply(simp add:compP_def compC_def class_def map_of_map4) -apply(simp add:split_def) -done +by(simp add:compP_def compC_def class_def map_of_map4) + (simp add:split_def) (*>*) lemma sees_methods_compP: "P \ C sees_methods Mm \ compP f P \ C sees_methods (map_option (\((Ts,T,m),D). ((Ts,T,f m),D)) \ Mm)" -(*<*) -apply(erule Methods.induct) - apply(rule sees_methods_Object) - apply(erule class_compP) - apply(rule ext) - apply(simp add:compM_def map_of_map4 option.map_comp) - apply(case_tac "map_of ms x") - apply simp - apply fastforce -apply(rule sees_methods_rec) - apply(erule class_compP) - apply assumption - apply assumption -apply(rule ext) -apply(simp add:map_add_def compM_def map_of_map4 option.map_comp split:option.split) -done +(*<*)(is "?P \ compP f P \ C sees_methods (?map Mm)") +proof(induct rule: Methods.induct) + case Object: (sees_methods_Object D fs ms Mm) + let ?Mm1 = "\x. map_option ((\m. (m, Object)) \ (\(Ts, T, m). (Ts, T, f m))) (map_of ms x)" + let ?Mm2 = "\x. map_option (case_prod (\(Ts, T, m). + Pair (Ts, T, f m)) \ (\m. (m, Object))) (map_of ms x)" + have Mm_eq: "\x. ?Mm1 x = ?Mm2 x" + proof - + fix x show "?Mm1 x = ?Mm2 x" + proof(cases "map_of ms x") + case None then show ?thesis by simp + qed fastforce + qed + + have Mm: "Mm = map_option (\m. (m, Object)) \ map_of ms" by fact + let ?Mm = "map_option (\m. (m, Object)) \ map_of (map (compM f) ms)" + let ?Mm' = "?map Mm" + have "?Mm' = ?Mm" + by(rule ext) (simp add:Mm Mm_eq compM_def map_of_map4 option.map_comp) + then show ?case by(rule sees_methods_Object[OF class_compP[OF Object(1)]]) +next + case rec: (sees_methods_rec C D fs ms Mm Mm') + have Mm': "Mm' = Mm ++ (map_option (\m. (m, C)) \ map_of ms)" by fact + let ?Mm' = "?map Mm'" + let ?Mm'' = "(?map Mm) ++ (map_option (\m. (m, C)) \ map_of (map (compM f) ms))" + have "?Mm' = ?Mm''" + by(rule ext) (simp add:Mm' map_add_def compM_def map_of_map4) + moreover have "compP f P \ C sees_methods ?Mm''" + using sees_methods_rec[OF class_compP[OF rec(1)] rec(2,4)] by fast + ultimately show "compP f P \ C sees_methods ?Mm'" by simp +qed (*>*) lemma sees_method_compP: "P \ C sees M: Ts\T = m in D \ compP f P \ C sees M: Ts\T = (f m) in D" (*<*)by(fastforce elim:sees_methods_compP simp add:Method_def)(*>*) lemma [simp]: "P \ C sees M: Ts\T = m in D \ method (compP f P) C M = (D,Ts,T,f m)" (*<*) -apply(drule sees_method_compP) -apply(simp add:method_def) -apply(rule the_equality) - apply simp -apply(fastforce dest:sees_method_fun) -done +proof - + let ?P = "\(D, Ts, T, m). compP f P \ C sees M: Ts\T = m in D" + let ?a = "(D, Ts, T, f m)" + assume cM: "P \ C sees M: Ts\T = m in D" + have compP_cM: "?P ?a" using sees_method_compP[OF cM] by simp + moreover { + fix x assume "?P x" then have "x = ?a" + using compP_cM by(fastforce dest:sees_method_fun) + } + ultimately have "(THE x. ?P x) = ?a" by(rule the_equality) + then show ?thesis by(simp add:method_def) +qed (*>*) lemma sees_methods_compPD: "\ cP \ C sees_methods Mm'; cP = compP f P \ \ \Mm. P \ C sees_methods Mm \ Mm' = (map_option (\((Ts,T,m),D). ((Ts,T,f m),D)) \ Mm)" -(*<*) -apply(erule Methods.induct) - apply(clarsimp simp:compC_def) - apply(rule exI) - apply(rule conjI, erule sees_methods_Object) - apply(rule refl) - apply(rule ext) - apply(simp add:compM_def map_of_map4 option.map_comp) - apply(case_tac "map_of b x") - apply simp - apply fastforce -apply(clarsimp simp:compC_def) -apply(rule exI, rule conjI) -apply(erule (2) sees_methods_rec) - apply(rule refl) -apply(rule ext) -apply(simp add:map_add_def compM_def map_of_map4 option.map_comp split:option.split) -done +(*<*)(is "\ ?P; ?Q \ \ \Mm. P \ C sees_methods Mm \ Mm' = (?map Mm)") +proof(induct rule: Methods.induct) + case Object: (sees_methods_Object D fs ms Mm) + then obtain ms' where P_Obj: "class P Object = \(D, fs, ms')\" + and ms: "ms = map (compM f) ms'" by(clarsimp simp:compC_def) + + let ?Mm1 = "\x. map_option ((\m. (m, Object)) \ (\(Ts, T, m). (Ts, T, f m))) (map_of ms' x)" + let ?Mm2 = "\x. map_option (case_prod (\(Ts, T, m). Pair (Ts, T, f m)) \ (\m. (m, Object))) + (map_of ms' x)" + have Mm_eq: "\x. ?Mm1 x = ?Mm2 x" + proof - + fix x show "?Mm1 x = ?Mm2 x" + proof(cases "map_of ms' x") + case None then show ?thesis by simp + qed fastforce + qed + + let ?Mm = "map_option (\m. (m,Object)) \ map_of ms'" + let ?Mm' = "?map ?Mm" + have Mm: "Mm = map_option (\m. (m, Object)) \ map_of ms" by fact + have "P \ Object sees_methods ?Mm" + using sees_methods_Object[OF P_Obj] by simp + moreover have "Mm = ?Mm'" + by(rule ext) (simp add:Mm_eq Mm ms compM_def map_of_map4 option.map_comp) + ultimately show ?case by fast +next + case rec: (sees_methods_rec C D fs ms Mm Mm') + then obtain ms' Mm\<^sub>D where P_D: "class P C = \(D, fs, ms')\" + and ms: "ms = map (compM f) ms'" and C_nObj: "C \ Object" + and Mm\<^sub>D: "P \ D sees_methods Mm\<^sub>D" + and Mm: "Mm = (\a. map_option (case_prod (\(Ts, T, m). Pair (Ts, T, f m))) (Mm\<^sub>D a))" + by(clarsimp simp:compC_def) + + let ?Mm = "Mm\<^sub>D ++ (map_option (\m. (m, C)) \ map_of ms')" + let ?Mm1 = "Mm ++ (map_option (\m. (m, C)) \ map_of ms)" + let ?Mm2 = "Mm ++ (map_option (\m. (m, C)) \ map_of (map (compM f) ms'))" + let ?Mm3 = "?map ?Mm" + have "Mm' = ?Mm1" by fact + also have "\ = ?Mm2" using ms by simp + also have "\ = ?Mm3" + by(rule ext)(simp add:Mm map_add_def compM_def map_of_map4) + moreover have "P \ C sees_methods ?Mm" + using sees_methods_rec[OF P_D C_nObj Mm\<^sub>D] by simp + ultimately show ?case by fast +qed (*>*) lemma sees_method_compPD: "compP f P \ C sees M: Ts\T = fm in D \ \m. P \ C sees M: Ts\T = m in D \ f m = fm" (*<*) -apply(simp add:Method_def) -apply clarify -apply(drule sees_methods_compPD[OF _ refl]) -apply clarsimp -apply blast -done +proof - + assume "compP f P \ C sees M: Ts\T = fm in D" + then obtain Mm where Mm: "compP f P \ C sees_methods Mm" + and MmM: "Mm M = \((Ts, T, fm), D)\" + by(clarsimp simp:Method_def) + show ?thesis using sees_methods_compPD[OF Mm refl] MmM + by(fastforce simp: Method_def) +qed (*>*) lemma [simp]: "subcls1(compP f P) = subcls1 P" (*<*) by(fastforce simp add: is_class_def compC_def intro:subcls1I order_antisym dest:subcls1D) (*>*) lemma compP_widen[simp]: "(compP f P \ T \ T') = (P \ T \ T')" (*<*)by(cases T')(simp_all add:widen_Class)(*>*) lemma [simp]: "(compP f P \ Ts [\] Ts') = (P \ Ts [\] Ts')" (*<*) -apply(induct Ts) - apply simp -apply(cases Ts') -apply(auto simp:fun_of_def) -done +proof(induct Ts) + case (Cons a Ts) + then show ?case by(cases Ts')(auto simp:fun_of_def) +qed simp (*>*) lemma [simp]: "is_type (compP f P) T = is_type P T" (*<*)by(cases T) simp_all(*>*) lemma [simp]: "(compP (f::'a\'b) P \ C has_fields FDTs) = (P \ C has_fields FDTs)" (*<*) (is "?A = ?B") proof { fix cP::"'b prog" assume "cP \ C has_fields FDTs" hence "cP = compP f P \ P \ C has_fields FDTs" proof induct case has_fields_Object thus ?case by(fast intro:Fields.has_fields_Object dest:class_compPD) next case has_fields_rec thus ?case by(fast intro:Fields.has_fields_rec dest:class_compPD) qed } note lem = this assume ?A with lem show ?B by blast next assume ?B thus ?A proof induct case has_fields_Object thus ?case by(fast intro:Fields.has_fields_Object class_compP) next case has_fields_rec thus ?case by(fast intro:Fields.has_fields_rec class_compP) qed qed (*>*) lemma [simp]: "fields (compP f P) C = fields P C" (*<*)by(simp add:fields_def)(*>*) lemma [simp]: "(compP f P \ C sees F:T in D) = (P \ C sees F:T in D)" (*<*)by(simp add:sees_field_def)(*>*) lemma [simp]: "field (compP f P) F D = field P F D" (*<*)by(simp add:field_def)(*>*) subsection\Invariance of @{term wf_prog} under compilation\ lemma [iff]: "distinct_fst (compP f P) = distinct_fst P" (*<*) -apply(simp add:distinct_fst_def compP_def compC_def) -apply(induct P) -apply (auto simp:image_iff) -done +by (induct P) + (auto simp:distinct_fst_def compP_def compC_def image_iff) (*>*) lemma [iff]: "distinct_fst (map (compM f) ms) = distinct_fst ms" (*<*) -apply(simp add:distinct_fst_def compM_def) -apply(induct ms) -apply (auto simp:image_iff) -done +by (induct ms) + (auto simp:distinct_fst_def compM_def image_iff) (*>*) lemma [iff]: "wf_syscls (compP f P) = wf_syscls P" (*<*)by(simp add:wf_syscls_def compP_def compC_def image_def Bex_def)(*>*) lemma [iff]: "wf_fdecl (compP f P) = wf_fdecl P" (*<*)by(simp add:wf_fdecl_def)(*>*) lemma set_compP: "((C,D,fs,ms') \ set(compP f P)) = (\ms. (C,D,fs,ms) \ set P \ ms' = map (compM f) ms)" (*<*)by(fastforce simp add:compP_def compC_def image_iff Bex_def)(*>*) lemma wf_cdecl_compPI: "\ \C M Ts T m. \ wf_mdecl wf\<^sub>1 P C (M,Ts,T,m); P \ C sees M:Ts\T = m in C \ \ wf_mdecl wf\<^sub>2 (compP f P) C (M,Ts,T, f m); \x\set P. wf_cdecl wf\<^sub>1 P x; x \ set (compP f P); wf_prog p P \ \ wf_cdecl wf\<^sub>2 (compP f P) x" (*<*) -apply(clarsimp simp add:wf_cdecl_def Ball_def set_compP) -apply(rename_tac C D fs ms) -apply(rule conjI) - apply (clarsimp simp:compM_def) - apply (drule (2) mdecl_visible) - apply simp -apply(clarify) -apply(drule sees_method_compPD[where f = f]) -apply clarsimp -apply(fastforce simp:image_iff compM_def) -done +proof - + assume + wfm: "\C M Ts T m. \ wf_mdecl wf\<^sub>1 P C (M,Ts,T,m); P \ C sees M:Ts\T = m in C \ + \ wf_mdecl wf\<^sub>2 (compP f P) C (M,Ts,T, f m)" + and wfc: "\x\set P. wf_cdecl wf\<^sub>1 P x" + and compP: "x \ set (compP f P)" and wf: "wf_prog p P" + obtain C D fs ms where x: "x = (C, D, fs, map (compM f) ms)" + and x_set: "(C, D, fs, ms) \ set P" + using compP by(case_tac x) (clarsimp simp: set_compP) + have wfc': "wf_cdecl wf\<^sub>1 P (C, D, fs, ms)" using wfc x_set by fast + let ?P = "compP f P" and ?ms = "compM f ` set ms" + { fix M Ts T m + assume M: "(M,Ts,T,m) \ set ms" + then have "wf_mdecl wf\<^sub>1 P C (M, Ts, T, m)" using wfc' + by(simp add:wf_cdecl_def) + moreover have cM: "P \ C sees M : Ts\T = m in C" using M + by(rule mdecl_visible[OF wf x_set]) + ultimately have "wf_mdecl wf\<^sub>2 (compP f P) C (M, Ts, T, f m)" + by(rule wfm) + } + then have "\m \ ?ms. wf_mdecl wf\<^sub>2 ?P C m" + by (clarsimp simp:compM_def) + moreover have "C \ Object \ + (\(M,Ts,T,m)\?ms. + \D' Ts' T' m'. ?P \ D sees M:Ts' \ T' = m' in D' \ + P \ Ts' [\] Ts \ P \ T \ T')" + proof - + { fix M Ts T m D' Ts' T' m' + assume "C \ Object" and "(M,Ts,T,m)\?ms" + and dM: "?P \ D sees M:Ts' \ T' = m' in D'" + then have "P \ Ts' [\] Ts \ P \ T \ T'" + using wfc' sees_method_compPD[OF dM] + by(fastforce simp:wf_cdecl_def image_iff compM_def) + } + then show ?thesis by fast + qed + moreover have "(\f\set fs. wf_fdecl P f) \ distinct_fst fs \ distinct_fst ms + \ (C \ Object \ is_class P D \ \ P \ D \\<^sup>* C)" using wfc' + by(simp add: wf_cdecl_def) + ultimately show ?thesis using x by(simp add:wf_cdecl_def) +qed (*>*) lemma wf_prog_compPI: assumes lift: "\C M Ts T m. \ P \ C sees M:Ts\T = m in C; wf_mdecl wf\<^sub>1 P C (M,Ts,T,m) \ \ wf_mdecl wf\<^sub>2 (compP f P) C (M,Ts,T, f m)" and wf: "wf_prog wf\<^sub>1 P" shows "wf_prog wf\<^sub>2 (compP f P)" (*<*) using wf by (simp add:wf_prog_def) (blast intro:wf_cdecl_compPI lift wf) (*>*) end diff --git a/thys/Jinja/Compiler/TypeComp.thy b/thys/Jinja/Compiler/TypeComp.thy --- a/thys/Jinja/Compiler/TypeComp.thy +++ b/thys/Jinja/Compiler/TypeComp.thy @@ -1,1308 +1,1445 @@ (* Title: Jinja/Compiler/TypeComp.thy Author: Tobias Nipkow Copyright TUM 2003 *) section \Preservation of Well-Typedness\ theory TypeComp imports Compiler "../BV/BVSpec" begin (*<*) declare nth_append[simp] (*>*) locale TC0 = fixes P :: "J\<^sub>1_prog" and mxl :: nat begin definition "ty E e = (THE T. P,E \\<^sub>1 e :: T)" definition "ty\<^sub>l E A' = map (\i. if i \ A' \ i < size E then OK(E!i) else Err) [0..i' ST E A = (case A of None \ None | \A'\ \ Some(ST, ty\<^sub>l E A'))" definition "after E A ST e = ty\<^sub>i' (ty E e # ST) E (A \ \ e)" end lemma (in TC0) ty_def2 [simp]: "P,E \\<^sub>1 e :: T \ ty E e = T" -(*<*) -apply (unfold ty_def) -apply(blast intro: the_equality WT\<^sub>1_unique) -done -(*>*) +(*<*)by(unfold ty_def) (blast intro: the_equality WT\<^sub>1_unique)(*>*) lemma (in TC0) [simp]: "ty\<^sub>i' ST E None = None" (*<*)by (simp add: ty\<^sub>i'_def)(*>*) lemma (in TC0) ty\<^sub>l_app_diff[simp]: "ty\<^sub>l (E@[T]) (A - {size E}) = ty\<^sub>l E A" (*<*)by(auto simp add:ty\<^sub>l_def hyperset_defs)(*>*) lemma (in TC0) ty\<^sub>i'_app_diff[simp]: "ty\<^sub>i' ST (E @ [T]) (A \ size E) = ty\<^sub>i' ST E A" (*<*)by(auto simp add:ty\<^sub>i'_def hyperset_defs)(*>*) lemma (in TC0) ty\<^sub>l_antimono: "A \ A' \ P \ ty\<^sub>l E A' [\\<^sub>\] ty\<^sub>l E A" (*<*)by(auto simp:ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>i'_antimono: "A \ A' \ P \ ty\<^sub>i' ST E \A'\ \' ty\<^sub>i' ST E \A\" (*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>l_env_antimono: "P \ ty\<^sub>l (E@[T]) A [\\<^sub>\] ty\<^sub>l E A" (*<*)by(auto simp:ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>i'_env_antimono: "P \ ty\<^sub>i' ST (E@[T]) A \' ty\<^sub>i' ST E A" (*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>i'_incr: "P \ ty\<^sub>i' ST (E @ [T]) \insert (size E) A\ \' ty\<^sub>i' ST E \A\" (*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>l_incr: "P \ ty\<^sub>l (E @ [T]) (insert (size E) A) [\\<^sub>\] ty\<^sub>l E A" (*<*)by(auto simp: hyperset_defs ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>l_in_types: "set E \ types P \ ty\<^sub>l E A \ list mxl (err (types P))" (*<*)by(auto simp add:ty\<^sub>l_def intro!:listI dest!: nth_mem)(*>*) locale TC1 = TC0 begin primrec compT :: "ty list \ nat hyperset \ ty list \ expr\<^sub>1 \ ty\<^sub>i' list" and compTs :: "ty list \ nat hyperset \ ty list \ expr\<^sub>1 list \ ty\<^sub>i' list" where "compT E A ST (new C) = []" | "compT E A ST (Cast C e) = compT E A ST e @ [after E A ST e]" | "compT E A ST (Val v) = []" | "compT E A ST (e\<^sub>1 \bop\ e\<^sub>2) = (let ST\<^sub>1 = ty E e\<^sub>1#ST; A\<^sub>1 = A \ \ e\<^sub>1 in compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @ compT E A\<^sub>1 ST\<^sub>1 e\<^sub>2 @ [after E A\<^sub>1 ST\<^sub>1 e\<^sub>2])" | "compT E A ST (Var i) = []" | "compT E A ST (i := e) = compT E A ST e @ [after E A ST e, ty\<^sub>i' ST E (A \ \ e \ \{i}\)]" | "compT E A ST (e\F{D}) = compT E A ST e @ [after E A ST e]" | "compT E A ST (e\<^sub>1\F{D} := e\<^sub>2) = (let ST\<^sub>1 = ty E e\<^sub>1#ST; A\<^sub>1 = A \ \ e\<^sub>1; A\<^sub>2 = A\<^sub>1 \ \ e\<^sub>2 in compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @ compT E A\<^sub>1 ST\<^sub>1 e\<^sub>2 @ [after E A\<^sub>1 ST\<^sub>1 e\<^sub>2] @ [ty\<^sub>i' ST E A\<^sub>2])" | "compT E A ST {i:T; e} = compT (E@[T]) (A\i) ST e" | "compT E A ST (e\<^sub>1;;e\<^sub>2) = (let A\<^sub>1 = A \ \ e\<^sub>1 in compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1, ty\<^sub>i' ST E A\<^sub>1] @ compT E A\<^sub>1 ST e\<^sub>2)" | "compT E A ST (if (e) e\<^sub>1 else e\<^sub>2) = (let A\<^sub>0 = A \ \ e; \ = ty\<^sub>i' ST E A\<^sub>0 in compT E A ST e @ [after E A ST e, \] @ compT E A\<^sub>0 ST e\<^sub>1 @ [after E A\<^sub>0 ST e\<^sub>1, \] @ compT E A\<^sub>0 ST e\<^sub>2)" | "compT E A ST (while (e) c) = (let A\<^sub>0 = A \ \ e; A\<^sub>1 = A\<^sub>0 \ \ c; \ = ty\<^sub>i' ST E A\<^sub>0 in compT E A ST e @ [after E A ST e, \] @ compT E A\<^sub>0 ST c @ [after E A\<^sub>0 ST c, ty\<^sub>i' ST E A\<^sub>1, ty\<^sub>i' ST E A\<^sub>0])" | "compT E A ST (throw e) = compT E A ST e @ [after E A ST e]" | "compT E A ST (e\M(es)) = compT E A ST e @ [after E A ST e] @ compTs E (A \ \ e) (ty E e # ST) es" | "compT E A ST (try e\<^sub>1 catch(C i) e\<^sub>2) = compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @ [ty\<^sub>i' (Class C#ST) E A, ty\<^sub>i' ST (E@[Class C]) (A \ \{i}\)] @ compT (E@[Class C]) (A \ \{i}\) ST e\<^sub>2" | "compTs E A ST [] = []" | "compTs E A ST (e#es) = compT E A ST e @ [after E A ST e] @ compTs E (A \ (\ e)) (ty E e # ST) es" definition compT\<^sub>a :: "ty list \ nat hyperset \ ty list \ expr\<^sub>1 \ ty\<^sub>i' list" where "compT\<^sub>a E A ST e = compT E A ST e @ [after E A ST e]" end lemma compE\<^sub>2_not_Nil[simp]: "compE\<^sub>2 e \ []" (*<*)by(induct e) auto(*>*) lemma (in TC1) compT_sizes[simp]: shows "\E A ST. size(compT E A ST e) = size(compE\<^sub>2 e) - 1" and "\E A ST. size(compTs E A ST es) = size(compEs\<^sub>2 es)" (*<*) -apply(induct e and es rule: compE\<^sub>2.induct compEs\<^sub>2.induct) -apply(auto split:bop.splits nat_diff_split) -done +by(induct e and es rule: compE\<^sub>2.induct compEs\<^sub>2.induct) + (auto split:bop.splits nat_diff_split) (*>*) lemma (in TC1) [simp]: "\ST E. \\\ \ set (compT E None ST e)" and [simp]: "\ST E. \\\ \ set (compTs E None ST es)" (*<*)by(induct e and es rule: compT.induct compTs.induct) (simp_all add:after_def)(*>*) lemma (in TC0) pair_eq_ty\<^sub>i'_conv: "(\(ST, LT)\ = ty\<^sub>i' ST\<^sub>0 E A) = (case A of None \ False | Some A \ (ST = ST\<^sub>0 \ LT = ty\<^sub>l E A))" (*<*)by(simp add:ty\<^sub>i'_def)(*>*) lemma (in TC0) pair_conv_ty\<^sub>i': "\(ST, ty\<^sub>l E A)\ = ty\<^sub>i' ST E \A\" (*<*)by(simp add:ty\<^sub>i'_def)(*>*) (*<*) declare (in TC0) ty\<^sub>i'_antimono [intro!] after_def[simp] pair_conv_ty\<^sub>i'[simp] pair_eq_ty\<^sub>i'_conv[simp] (*>*) lemma (in TC1) compT_LT_prefix: "\E A ST\<^sub>0. \ \(ST,LT)\ \ set(compT E A ST\<^sub>0 e); \ e (size E) \ \ P \ \(ST,LT)\ \' ty\<^sub>i' ST E A" and "\E A ST\<^sub>0. \ \(ST,LT)\ \ set(compTs E A ST\<^sub>0 es); \s es (size E) \ \ P \ \(ST,LT)\ \' ty\<^sub>i' ST E A" (*<*) proof(induct e and es rule: compT.induct compTs.induct) case FAss thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case BinOp thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans split:bop.splits) next case Seq thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case While thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case Cond thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case Block thus ?case by(force simp add:hyperset_defs ty\<^sub>i'_def simp del:pair_conv_ty\<^sub>i' elim!:sup_state_opt_trans) next case Call thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case Cons_exp thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case TryCatch thus ?case by(fastforce simp:hyperset_defs intro!:(* ty\<^sub>i'_env_antimono *) ty\<^sub>i'_incr elim!:sup_state_opt_trans) qed (auto simp:hyperset_defs) declare (in TC0) ty\<^sub>i'_antimono [rule del] after_def[simp del] pair_conv_ty\<^sub>i'[simp del] pair_eq_ty\<^sub>i'_conv[simp del] (*>*) lemma [iff]: "OK None \ states P mxs mxl" (*<*)by(simp add: JVM_states_unfold)(*>*) lemma (in TC0) after_in_states: - "\ wf_prog p P; P,E \\<^sub>1 e :: T; set E \ types P; set ST \ types P; - size ST + max_stack e \ mxs \ - \ OK (after E A ST e) \ states P mxs mxl" +assumes wf: "wf_prog p P" and wt: "P,E \\<^sub>1 e :: T" + and Etypes: "set E \ types P" and STtypes: "set ST \ types P" + and stack: "size ST + max_stack e \ mxs" +shows "OK (after E A ST e) \ states P mxs mxl" (*<*) -apply(subgoal_tac "size ST + 1 \ mxs") - apply(simp add: after_def ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types) - apply(blast intro!:listI WT\<^sub>1_is_type) -using max_stack1[of e] apply simp -done +proof - + have "size ST + 1 \ mxs" using max_stack1[of e] wt stack + by fastforce + then show ?thesis using assms + by(simp add: after_def ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types) + (blast intro!:listI WT\<^sub>1_is_type) +qed (*>*) lemma (in TC0) OK_ty\<^sub>i'_in_statesI[simp]: "\ set E \ types P; set ST \ types P; size ST \ mxs \ \ OK (ty\<^sub>i' ST E A) \ states P mxs mxl" (*<*) -apply(simp add:ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types) -apply(blast intro!:listI) -done +by(simp add:ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types) + (blast intro!:listI) (*>*) lemma is_class_type_aux: "is_class P C \ is_type P (Class C)" (*<*)by(simp)(*>*) (*<*) declare is_type_simps[simp del] subsetI[rule del] (*>*) theorem (in TC1) compT_states: assumes wf: "wf_prog p P" shows "\E T A ST. \ P,E \\<^sub>1 e :: T; set E \ types P; set ST \ types P; size ST + max_stack e \ mxs; size E + max_vars e \ mxl \ \ OK ` set(compT E A ST e) \ states P mxs mxl" (*<*)(is "\E T A ST. PROP ?P e E T A ST")(*>*) and "\E Ts A ST. \ P,E \\<^sub>1 es[::]Ts; set E \ types P; set ST \ types P; size ST + max_stacks es \ mxs; size E + max_varss es \ mxl \ \ OK ` set(compTs E A ST es) \ states P mxs mxl" (*<*)(is "\E Ts A ST. PROP ?Ps es E Ts A ST") proof(induct e and es rule: compT.induct compTs.induct) case new thus ?case by(simp) next case (Cast C e) thus ?case by (auto simp:after_in_states[OF wf]) next case Val thus ?case by(simp) next case Var thus ?case by(simp) next case LAss thus ?case by(auto simp:after_in_states[OF wf]) next case FAcc thus ?case by(auto simp:after_in_states[OF wf]) next case FAss thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case Seq thus ?case by(auto simp:image_Un after_in_states[OF wf]) next case BinOp thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case Cond thus ?case by(force simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case While thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case Block thus ?case by(auto) next case (TryCatch e\<^sub>1 C i e\<^sub>2) moreover have "size ST + 1 \ mxs" using TryCatch.prems max_stack1[of e\<^sub>1] by auto ultimately show ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf] is_class_type_aux) next case Nil_exp thus ?case by simp next case Cons_exp thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case throw thus ?case by(auto simp: WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case Call thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) qed declare is_type_simps[simp] subsetI[intro!] (*>*) definition shift :: "nat \ ex_table \ ex_table" where "shift n xt \ map (\(from,to,C,handler,depth). (from+n,to+n,C,handler+n,depth)) xt" lemma [simp]: "shift 0 xt = xt" (*<*)by(induct xt)(auto simp:shift_def)(*>*) lemma [simp]: "shift n [] = []" (*<*)by(simp add:shift_def)(*>*) lemma [simp]: "shift n (xt\<^sub>1 @ xt\<^sub>2) = shift n xt\<^sub>1 @ shift n xt\<^sub>2" (*<*)by(simp add:shift_def)(*>*) lemma [simp]: "shift m (shift n xt) = shift (m+n) xt" (*<*)by(induct xt)(auto simp:shift_def)(*>*) lemma [simp]: "pcs (shift n xt) = {pc+n|pc. pc \ pcs xt}" (*<*) -apply(auto simp:shift_def pcs_def) -apply(rule_tac x = "x-n" in exI) -apply (force split:nat_diff_split) -done +proof - + { fix x f t C h d + assume "(f,t,C,h,d) \ set xt" and "f + n \ x" + and "x < t + n" + then have "\pc. x = pc + n \ (\x\set xt. pc \ (case x of (f, t, C, h, d) \ {f..*) lemma shift_compxE\<^sub>2: shows "\pc pc' d. shift pc (compxE\<^sub>2 e pc' d) = compxE\<^sub>2 e (pc' + pc) d" and "\pc pc' d. shift pc (compxEs\<^sub>2 es pc' d) = compxEs\<^sub>2 es (pc' + pc) d" (*<*) -apply(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) -apply(auto simp:shift_def ac_simps) -done +by(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) + (auto simp:shift_def ac_simps) (*>*) lemma compxE\<^sub>2_size_convs[simp]: shows "n \ 0 \ compxE\<^sub>2 e n d = shift n (compxE\<^sub>2 e 0 d)" and "n \ 0 \ compxEs\<^sub>2 es n d = shift n (compxEs\<^sub>2 es 0 d)" (*<*)by(simp_all add:shift_compxE\<^sub>2)(*>*) locale TC2 = TC1 + fixes T\<^sub>r :: ty and mxs :: pc begin definition wt_instrs :: "instr list \ ex_table \ ty\<^sub>i' list \ bool" ("(\ _, _ /[::]/ _)" [0,0,51] 50) where "\ is,xt [::] \s \ size is < size \s \ pcs xt \ {0.. (\pc< size is. P,T\<^sub>r,mxs,size \s,xt \ is!pc,pc :: \s)" end notation TC2.wt_instrs ("(_,_,_ \/ _, _ /[::]/ _)" [50,50,50,50,50,51] 50) (*<*) lemmas (in TC2) wt_defs = wt_instrs_def wt_instr_def app_def eff_def norm_eff_def (*>*) lemma (in TC2) [simp]: "\s \ [] \ \ [],[] [::] \s" (*<*) by (simp add: wt_defs) (*>*) lemma [simp]: "eff i P pc et None = []" (*<*)by (simp add: Effect.eff_def)(*>*) (*<*) declare split_comp_eq[simp del] (*>*) lemma wt_instr_appR: "\ P,T,m,mpc,xt \ is!pc,pc :: \s; pc < size is; size is < size \s; mpc \ size \s; mpc \ mpc' \ \ P,T,m,mpc',xt \ is!pc,pc :: \s@\s'" (*<*)by (fastforce simp:wt_instr_def app_def)(*>*) lemma relevant_entries_shift [simp]: "relevant_entries P i (pc+n) (shift n xt) = shift n (relevant_entries P i pc xt)" (*<*) - apply (induct xt) - apply (unfold relevant_entries_def shift_def) - apply simp - apply (auto simp add: is_relevant_entry_def) - done +proof(induct xt) + case Nil + then show ?case by (simp add: relevant_entries_def shift_def) +next + case (Cons a xt) + then show ?case + by (auto simp add: relevant_entries_def shift_def is_relevant_entry_def) +qed (*>*) lemma [simp]: "xcpt_eff i P (pc+n) \ (shift n xt) = map (\(pc,\). (pc + n, \)) (xcpt_eff i P pc \ xt)" (*<*) -apply(simp add: xcpt_eff_def) -apply(cases \) -apply(auto simp add: shift_def) -done +proof - + obtain ST LT where "\ = (ST, LT)" by(cases \) simp + then show ?thesis by(simp add: xcpt_eff_def) (auto simp add: shift_def) +qed (*>*) lemma [simp]: "app\<^sub>i (i, P, pc, m, T, \) \ eff i P (pc+n) (shift n xt) (Some \) = map (\(pc,\). (pc+n,\)) (eff i P pc xt (Some \))" -(*<*) -apply(simp add:eff_def norm_eff_def) -apply(cases "i",auto) -done -(*>*) +(*<*)by(cases "i") (auto simp add:eff_def norm_eff_def)(*>*) lemma [simp]: "xcpt_app i P (pc+n) mxs (shift n xt) \ = xcpt_app i P pc mxs xt \" (*<*)by (simp add: xcpt_app_def) (auto simp add: shift_def)(*>*) lemma wt_instr_appL: - "\ P,T,m,mpc,xt \ i,pc :: \s; pc < size \s; mpc \ size \s \ - \ P,T,m,mpc + size \s',shift (size \s') xt \ i,pc+size \s' :: \s'@\s" +assumes "P,T,m,mpc,xt \ i,pc :: \s" and "pc < size \s" and "mpc \ size \s" +shows "P,T,m,mpc + size \s',shift (size \s') xt \ i,pc+size \s' :: \s'@\s" (*<*) -apply(auto simp:wt_instr_def app_def) -prefer 2 apply(fast) -prefer 2 apply(fast) -apply(cases "i",auto) -done +proof - + let ?t = "(\s'@\s)!(pc+size \s')" + show ?thesis + proof(cases ?t) + case (Some \) + obtain ST LT where [simp]: "\ = (ST, LT)" by(cases \) simp + have "app\<^sub>i (i, P, pc + length \s', m, T, \)" + using Some assms by(cases "i") (auto simp:wt_instr_def app_def) + moreover { + fix pc' \' assume "(pc',\') \ set (eff i P pc xt ?t)" + then have "P \ \' \' \s!pc'" and "pc' < mpc" + using Some assms by(auto simp:wt_instr_def app_def) + } + ultimately show ?thesis using Some assms + by(fastforce simp:wt_instr_def app_def) + qed (auto simp:wt_instr_def app_def) +qed (*>*) lemma wt_instr_Cons: - "\ P,T,m,mpc - 1,[] \ i,pc - 1 :: \s; - 0 < pc; 0 < mpc; pc < size \s + 1; mpc \ size \s + 1 \ - \ P,T,m,mpc,[] \ i,pc :: \#\s" +assumes wti: "P,T,m,mpc - 1,[] \ i,pc - 1 :: \s" + and pcl: "0 < pc" and mpcl: "0 < mpc" + and pcu: "pc < size \s + 1" and mpcu: "mpc \ size \s + 1" +shows "P,T,m,mpc,[] \ i,pc :: \#\s" (*<*) -apply(drule wt_instr_appL[where \s' = "[\]"]) -apply arith -apply arith -apply (simp split:nat_diff_split_asm) -done +proof - + have "pc - 1 < length \s" using pcl pcu by arith + moreover have "mpc - 1 \ length \s" using mpcl mpcu by arith + ultimately have + "P,T,m,mpc - 1 + length [\],shift (length [\]) [] \ i,pc - 1 + length [\] :: [\] @ \s" + by(rule wt_instr_appL[where \s' = "[\]", OF wti]) + then show ?thesis using pcl mpcl by (simp split:nat_diff_split_asm) +qed (*>*) lemma wt_instr_append: - "\ P,T,m,mpc - size \s',[] \ i,pc - size \s' :: \s; - size \s' \ pc; size \s' \ mpc; pc < size \s + size \s'; mpc \ size \s + size \s' \ - \ P,T,m,mpc,[] \ i,pc :: \s'@\s" +assumes wti: "P,T,m,mpc - size \s',[] \ i,pc - size \s' :: \s" + and pcl: "size \s' \ pc" and mpcl: "size \s' \ mpc" + and pcu: "pc < size \s + size \s'" and mpcu: "mpc \ size \s + size \s'" +shows "P,T,m,mpc,[] \ i,pc :: \s'@\s" (*<*) -apply(drule wt_instr_appL[where \s' = \s']) -apply arith -apply arith -apply (simp split:nat_diff_split_asm) -done +proof - + have "pc - length \s' < length \s" using pcl pcu by arith + moreover have "mpc - length \s' \ length \s" using mpcl mpcu by arith +thm wt_instr_appL[where \s' = "\s'", OF wti] + ultimately have "P,T,m,mpc - length \s' + length \s',shift (length \s') [] + \ i,pc - length \s' + length \s' :: \s' @ \s" + by(rule wt_instr_appL[where \s' = "\s'", OF wti]) + then show ?thesis using pcl mpcl by (simp split:nat_diff_split_asm) +qed (*>*) lemma xcpt_app_pcs: "pc \ pcs xt \ xcpt_app i P pc mxs xt \" (*<*) by (auto simp add: xcpt_app_def relevant_entries_def is_relevant_entry_def pcs_def) (*>*) lemma xcpt_eff_pcs: "pc \ pcs xt \ xcpt_eff i P pc \ xt = []" (*<*) by (cases \) (auto simp add: is_relevant_entry_def xcpt_eff_def relevant_entries_def pcs_def intro!: filter_False) (*>*) lemma pcs_shift: "pc < n \ pc \ pcs (shift n xt)" (*<*)by (auto simp add: shift_def pcs_def)(*>*) lemma wt_instr_appRx: "\ P,T,m,mpc,xt \ is!pc,pc :: \s; pc < size is; size is < size \s; mpc \ size \s \ \ P,T,m,mpc,xt @ shift (size is) xt' \ is!pc,pc :: \s" (*<*)by (auto simp:wt_instr_def eff_def app_def xcpt_app_pcs xcpt_eff_pcs)(*>*) lemma wt_instr_appLx: "\ P,T,m,mpc,xt \ i,pc :: \s; pc \ pcs xt' \ \ P,T,m,mpc,xt'@xt \ i,pc :: \s" (*<*)by (auto simp:wt_instr_def app_def eff_def xcpt_app_pcs xcpt_eff_pcs)(*>*) lemma (in TC2) wt_instrs_extR: "\ is,xt [::] \s \ \ is,xt [::] \s @ \s'" (*<*)by(auto simp add:wt_instrs_def wt_instr_appR)(*>*) lemma (in TC2) wt_instrs_ext: - "\ \ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@\s\<^sub>2; \ is\<^sub>2,xt\<^sub>2 [::] \s\<^sub>2; size \s\<^sub>1 = size is\<^sub>1 \ - \ \ is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" +assumes wt\<^sub>1: "\ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@\s\<^sub>2" and wt\<^sub>2: "\ is\<^sub>2,xt\<^sub>2 [::] \s\<^sub>2" + and \s_size: "size \s\<^sub>1 = size is\<^sub>1" + shows "\ is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" (*<*) -apply(clarsimp simp:wt_instrs_def) -apply(rule conjI, fastforce) -apply(rule conjI, fastforce) -apply clarsimp -apply(rule conjI, fastforce simp:wt_instr_appRx) -apply clarsimp -apply(erule_tac x = "pc - size is\<^sub>1" in allE)+ -apply(thin_tac "P \ Q" for P Q) -apply(erule impE, arith) -apply(drule_tac \s' = "\s\<^sub>1" in wt_instr_appL) - apply arith - apply simp -apply(fastforce simp add:add.commute intro!: wt_instr_appLx) -done +proof - + let ?is = "is\<^sub>1@is\<^sub>2" and ?xt = "xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2" + and ?\s = "\s\<^sub>1@\s\<^sub>2" + have "size ?is < size ?\s" using wt\<^sub>2 \s_size by(fastforce simp:wt_instrs_def) + moreover have "pcs ?xt \ {0..1 wt\<^sub>2 + by(fastforce simp:wt_instrs_def) + moreover { + fix pc assume pc: "pcr,mxs,size ?\s,?xt \ ?is!pc,pc :: ?\s" + proof(cases "pc < length is\<^sub>1") + case True then show ?thesis using wt\<^sub>1 pc + by(fastforce simp: wt_instrs_def wt_instr_appRx) + next + case False + then have "pc - length is\<^sub>1 < length is\<^sub>2" using pc by fastforce + then have "P,T\<^sub>r,mxs,length \s\<^sub>2,xt\<^sub>2 \ is\<^sub>2 ! (pc - length is\<^sub>1),pc - length is\<^sub>1 :: \s\<^sub>2" + using wt\<^sub>2 by(clarsimp simp: wt_instrs_def) + moreover have "pc - length is\<^sub>1 < length \s\<^sub>2" using pc wt\<^sub>2 + by(clarsimp simp: wt_instrs_def) arith + moreover have "length \s\<^sub>2 \ length \s\<^sub>2" by simp + moreover have "pc - length is\<^sub>1 + length \s\<^sub>1 \ pcs xt\<^sub>1" using wt\<^sub>1 \s_size + by(fastforce simp: wt_instrs_def) + ultimately have "P,T\<^sub>r,mxs,length \s\<^sub>2 + length \s\<^sub>1,xt\<^sub>1 @ shift (length \s\<^sub>1) xt\<^sub>2 + \ is\<^sub>2 ! (pc - length is\<^sub>1),pc - length is\<^sub>1 + length \s\<^sub>1 :: \s\<^sub>1 @ \s\<^sub>2" + by(rule wt_instr_appLx[OF wt_instr_appL[where \s' = "\s\<^sub>1"]]) + then show ?thesis using False \s_size by(simp add:add.commute) + qed + } + ultimately show ?thesis by(clarsimp simp:wt_instrs_def) +qed (*>*) corollary (in TC2) wt_instrs_ext2: "\ \ is\<^sub>2,xt\<^sub>2 [::] \s\<^sub>2; \ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@\s\<^sub>2; size \s\<^sub>1 = size is\<^sub>1 \ \ \ is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" (*<*)by(rule wt_instrs_ext)(*>*) corollary (in TC2) wt_instrs_ext_prefix [trans]: "\ \ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@\s\<^sub>2; \ is\<^sub>2,xt\<^sub>2 [::] \s\<^sub>3; size \s\<^sub>1 = size is\<^sub>1; prefix \s\<^sub>3 \s\<^sub>2 \ \ \ is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" (*<*)by(bestsimp simp:prefix_def elim: wt_instrs_ext dest:wt_instrs_extR)(*>*) corollary (in TC2) wt_instrs_app: assumes is\<^sub>1: "\ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@[\]" assumes is\<^sub>2: "\ is\<^sub>2,xt\<^sub>2 [::] \#\s\<^sub>2" assumes s: "size \s\<^sub>1 = size is\<^sub>1" shows "\ is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\#\s\<^sub>2" (*<*) proof - from is\<^sub>1 have "\ is\<^sub>1,xt\<^sub>1 [::] (\s\<^sub>1@[\])@\s\<^sub>2" by (rule wt_instrs_extR) hence "\ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@\#\s\<^sub>2" by simp from this is\<^sub>2 s show ?thesis by (rule wt_instrs_ext) qed (*>*) corollary (in TC2) wt_instrs_app_last[trans]: - "\ \ is\<^sub>2,xt\<^sub>2 [::] \#\s\<^sub>2; \ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1; - last \s\<^sub>1 = \; size \s\<^sub>1 = size is\<^sub>1+1 \ - \ \ is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" +assumes "\ is\<^sub>2,xt\<^sub>2 [::] \#\s\<^sub>2" "\ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1" + "last \s\<^sub>1 = \" "size \s\<^sub>1 = size is\<^sub>1+1" +shows "\ is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" (*<*) -apply(cases \s\<^sub>1 rule:rev_cases) - apply simp -apply(simp add:wt_instrs_app) -done +using assms proof(cases \s\<^sub>1 rule:rev_cases) + case (snoc ys y) + then show ?thesis using assms by(simp add:wt_instrs_app) +qed simp (*>*) corollary (in TC2) wt_instrs_append_last[trans]: - "\ \ is,xt [::] \s; P,T\<^sub>r,mxs,mpc,[] \ i,pc :: \s; - pc = size is; mpc = size \s; size is + 1 < size \s \ - \ \ is@[i],xt [::] \s" +assumes wtis: "\ is,xt [::] \s" and wti: "P,T\<^sub>r,mxs,mpc,[] \ i,pc :: \s" + and pc: "pc = size is" and mpc: "mpc = size \s" and is_\s: "size is + 1 < size \s" +shows "\ is@[i],xt [::] \s" (*<*) -apply(clarsimp simp add:wt_instrs_def) -apply(rule conjI, fastforce) -apply(fastforce intro!:wt_instr_appLx[where xt = "[]",simplified] - dest!:less_antisym) -done +proof - + have pc_xt: "pc \ pcs xt" using wtis pc by (fastforce simp:wt_instrs_def) + have "pcs xt \ {.. pc' < length is" "pc' < Suc (length is)" + have "P,T\<^sub>r,mxs,length \s,xt \ i,pc' :: \s" + using wt_instr_appLx[where xt = "[]",simplified,OF wti pc_xt] + less_antisym[OF pc'] pc mpc by simp + } + ultimately show ?thesis using wtis is_\s by(clarsimp simp add:wt_instrs_def) +qed (*>*) corollary (in TC2) wt_instrs_app2: "\ \ is\<^sub>2,xt\<^sub>2 [::] \'#\s\<^sub>2; \ is\<^sub>1,xt\<^sub>1 [::] \#\s\<^sub>1@[\']; xt' = xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2; size \s\<^sub>1+1 = size is\<^sub>1 \ \ \ is\<^sub>1@is\<^sub>2,xt' [::] \#\s\<^sub>1@\'#\s\<^sub>2" (*<*)using wt_instrs_app[where ?\s\<^sub>1.0 = "\ # \s\<^sub>1"] by simp (*>*) corollary (in TC2) wt_instrs_app2_simp[trans,simp]: "\ \ is\<^sub>2,xt\<^sub>2 [::] \'#\s\<^sub>2; \ is\<^sub>1,xt\<^sub>1 [::] \#\s\<^sub>1@[\']; size \s\<^sub>1+1 = size is\<^sub>1 \ \ \ is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \#\s\<^sub>1@\'#\s\<^sub>2" (*<*)using wt_instrs_app[where ?\s\<^sub>1.0 = "\ # \s\<^sub>1"] by simp(*>*) corollary (in TC2) wt_instrs_Cons[simp]: "\ \s \ []; \ [i],[] [::] [\,\']; \ is,xt [::] \'#\s \ \ \ i#is,shift 1 xt [::] \#\'#\s" (*<*) using wt_instrs_app2[where ?is\<^sub>1.0 = "[i]" and ?\s\<^sub>1.0 = "[]" and ?is\<^sub>2.0 = "is" and ?xt\<^sub>1.0 = "[]"] by simp corollary (in TC2) wt_instrs_Cons2[trans]: assumes \s: "\ is,xt [::] \s" assumes i: "P,T\<^sub>r,mxs,mpc,[] \ i,0 :: \#\s" assumes mpc: "mpc = size \s + 1" shows "\ i#is,shift 1 xt [::] \#\s" (*<*) proof - from \s have "\s \ []" by (auto simp: wt_instrs_def) with mpc i have "\ [i],[] [::] [\]@\s" by (simp add: wt_instrs_def) with \s show ?thesis by (fastforce dest: wt_instrs_ext) qed (*>*) lemma (in TC2) wt_instrs_last_incr[trans]: - "\ \ is,xt [::] \s@[\]; P \ \ \' \' \ \ \ is,xt [::] \s@[\']" +assumes wtis: "\ is,xt [::] \s@[\]" and ss: "P \ \ \' \'" +shows "\ is,xt [::] \s@[\']" (*<*) -apply(clarsimp simp add:wt_instrs_def wt_instr_def) -apply(rule conjI) -apply(fastforce) -apply(clarsimp) -apply(rename_tac pc' tau') -apply(erule allE, erule (1) impE) -apply(clarsimp) -apply(drule (1) bspec) -apply(clarsimp) -apply(subgoal_tac "pc' = size \s") -prefer 2 -apply(clarsimp simp:app_def) -apply(drule (1) bspec) -apply(clarsimp) -apply(auto elim!:sup_state_opt_trans) -done +proof - + let ?\s = "\s@[\]" and ?\s' = "\s@[\']" + { fix pc assume pc: "pc< size is" + let ?i = "is!pc" + have app_pc: "app (is ! pc) P mxs T\<^sub>r pc (length ?\s) xt (\s ! pc)" + using wtis pc by(clarsimp simp add:wt_instrs_def wt_instr_def) + then have Apc\': "\pc' \'. (pc',\') \ set (eff ?i P pc xt (?\s!pc)) + \ pc' < length ?\s" + using wtis pc by(fastforce simp add:wt_instrs_def app_def) + have Aepc\': "\pc' \'. (pc',\') \ set (eff ?i P pc xt (?\s!pc)) + \ P \ \' \' ?\s!pc'" + using wtis pc by(fastforce simp add:wt_instrs_def wt_instr_def) + { fix pc1 \1 assume pc\1: "(pc1,\1) \ set (eff ?i P pc xt (?\s'!pc))" + then have epc\': "(pc1,\1) \ set (eff ?i P pc xt (?\s!pc))" + using wtis pc by(simp add:wt_instrs_def) + have "P \ \1 \' ?\s'!pc1" + proof(cases "pc1 < length \s") + case True + then show ?thesis using wtis pc pc\1 + by(fastforce simp add:wt_instrs_def wt_instr_def) + next + case False + then have "pc1 < length ?\s" using Apc\'[OF epc\'] by simp + then have [simp]: "pc1 = size \s" using False by clarsimp + have "P \ \1 \' \" using Aepc\'[OF epc\'] by simp + then have "P \ \1 \' \'" by(rule sup_state_opt_trans[OF _ ss]) + then show ?thesis by simp + qed + } + then have "P,T\<^sub>r,mxs,size ?\s',xt \ is!pc,pc :: ?\s'" using wtis pc + by(clarsimp simp add:wt_instrs_def wt_instr_def) + } + then show ?thesis using wtis by(simp add:wt_instrs_def) +qed (*>*) lemma [iff]: "xcpt_app i P pc mxs [] \" (*<*)by (simp add: xcpt_app_def relevant_entries_def)(*>*) lemma [simp]: "xcpt_eff i P pc \ [] = []" (*<*)by (simp add: xcpt_eff_def relevant_entries_def)(*>*) lemma (in TC2) wt_New: "\ is_class P C; size ST < mxs \ \ \ [New C],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (Class C#ST) E A]" (*<*)by(simp add:wt_defs ty\<^sub>i'_def)(*>*) lemma (in TC2) wt_Cast: "is_class P C \ \ [Checkcast C],[] [::] [ty\<^sub>i' (Class D # ST) E A, ty\<^sub>i' (Class C # ST) E A]" (*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_Push: "\ size ST < mxs; typeof v = Some T \ \ \ [Push v],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (T#ST) E A]" (*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_Pop: "\ [Pop],[] [::] (ty\<^sub>i' (T#ST) E A # ty\<^sub>i' ST E A # \s)" (*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_CmpEq: "\ P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1\ \ \ [CmpEq],[] [::] [ty\<^sub>i' (T\<^sub>2 # T\<^sub>1 # ST) E A, ty\<^sub>i' (Boolean # ST) E A]" (*<*) by(auto simp:ty\<^sub>i'_def wt_defs elim!: refTE not_refTE) (*>*) lemma (in TC2) wt_IAdd: "\ [IAdd],[] [::] [ty\<^sub>i' (Integer#Integer#ST) E A, ty\<^sub>i' (Integer#ST) E A]" (*<*)by(simp add:ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_Load: "\ size ST < mxs; size E \ mxl; i \\ A; i < size E \ \ \ [Load i],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (E!i # ST) E A]" (*<*)by(auto simp add:ty\<^sub>i'_def wt_defs ty\<^sub>l_def hyperset_defs)(*>*) lemma (in TC2) wt_Store: "\ P \ T \ E!i; i < size E; size E \ mxl \ \ \ [Store i],[] [::] [ty\<^sub>i' (T#ST) E A, ty\<^sub>i' ST E (\{i}\ \ A)]" (*<*) by(auto simp:hyperset_defs nth_list_update ty\<^sub>i'_def wt_defs ty\<^sub>l_def intro:list_all2_all_nthI) (*>*) lemma (in TC2) wt_Get: "\ P \ C sees F:T in D \ \ \ [Getfield F D],[] [::] [ty\<^sub>i' (Class C # ST) E A, ty\<^sub>i' (T # ST) E A]" (*<*)by(auto simp: ty\<^sub>i'_def wt_defs dest: sees_field_idemp sees_field_decl_above)(*>*) lemma (in TC2) wt_Put: "\ P \ C sees F:T in D; P \ T' \ T \ \ \ [Putfield F D],[] [::] [ty\<^sub>i' (T' # Class C # ST) E A, ty\<^sub>i' ST E A]" (*<*)by(auto intro: sees_field_idemp sees_field_decl_above simp: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_Throw: "\ [Throw],[] [::] [ty\<^sub>i' (Class C # ST) E A, \']" (*<*)by(auto simp: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_IfFalse: "\ 2 \ i; nat i < size \s + 2; P \ ty\<^sub>i' ST E A \' \s ! nat(i - 2) \ \ \ [IfFalse i],[] [::] ty\<^sub>i' (Boolean # ST) E A # ty\<^sub>i' ST E A # \s" (*<*) by(simp add: ty\<^sub>i'_def wt_defs eval_nat_numeral nat_diff_distrib) (*>*) lemma wt_Goto: "\ 0 \ int pc + i; nat (int pc + i) < size \s; size \s \ mpc; P \ \s!pc \' \s ! nat (int pc + i) \ \ P,T,mxs,mpc,[] \ Goto i,pc :: \s" (*<*)by(clarsimp simp add: TC2.wt_defs)(*>*) lemma (in TC2) wt_Invoke: "\ size es = size Ts'; P \ C sees M: Ts\T = m in D; P \ Ts' [\] Ts \ \ \ [Invoke M (size es)],[] [::] [ty\<^sub>i' (rev Ts' @ Class C # ST) E A, ty\<^sub>i' (T#ST) E A]" (*<*)by(fastforce simp add: ty\<^sub>i'_def wt_defs)(*>*) corollary (in TC2) wt_instrs_app3[simp]: "\ \ is\<^sub>2,[] [::] (\' # \s\<^sub>2); \ is\<^sub>1,xt\<^sub>1 [::] \ # \s\<^sub>1 @ [\']; size \s\<^sub>1+1 = size is\<^sub>1\ \ \ (is\<^sub>1 @ is\<^sub>2),xt\<^sub>1 [::] \ # \s\<^sub>1 @ \' # \s\<^sub>2" (*<*)using wt_instrs_app2[where ?xt\<^sub>2.0 = "[]"] by (simp add:shift_def)(*>*) corollary (in TC2) wt_instrs_Cons3[simp]: "\ \s \ []; \ [i],[] [::] [\,\']; \ is,[] [::] \'#\s \ \ \ (i # is),[] [::] \ # \' # \s" (*<*) using wt_instrs_Cons[where ?xt = "[]"] by (simp add:shift_def) (*<*) declare nth_append[simp del] declare [[simproc del: list_to_set_comprehension]] (*>*) lemma (in TC2) wt_instrs_xapp[trans]: - "\ \ is\<^sub>1 @ is\<^sub>2, xt [::] \s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \s\<^sub>2; - \\ \ set \s\<^sub>1. \ST' LT'. \ = Some(ST',LT') \ - size ST \ size ST' \ P \ Some (drop (size ST' - size ST) ST',LT') \' ty\<^sub>i' ST E A; - size is\<^sub>1 = size \s\<^sub>1; is_class P C; size ST < mxs \ \ - \ is\<^sub>1 @ is\<^sub>2, xt @ [(0,size is\<^sub>1 - 1,C,size is\<^sub>1,size ST)] [::] \s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \s\<^sub>2" -(*<*) -apply(simp add:wt_instrs_def) -apply(rule conjI) - apply(clarsimp) - apply arith -apply clarsimp -apply(erule allE, erule (1) impE) -apply(clarsimp simp add: wt_instr_def app_def eff_def) -apply(rule conjI) - apply (thin_tac "\x\ A \ B. P x" for A B P) - apply (thin_tac "\x\ A \ B. P x" for A B P) - apply (clarsimp simp add: xcpt_app_def relevant_entries_def) - apply (simp add: nth_append is_relevant_entry_def split!: if_splits) - apply (drule_tac x="\s\<^sub>1!pc" in bspec) - apply (blast intro: nth_mem) - apply fastforce -apply (rule conjI) - apply clarsimp - apply (erule disjE, blast) - apply (erule disjE, blast) - apply (clarsimp simp add: xcpt_eff_def relevant_entries_def split: if_split_asm) -apply clarsimp -apply (erule disjE, blast) -apply (erule disjE, blast) -apply (clarsimp simp add: xcpt_eff_def relevant_entries_def split: if_split_asm) -apply (simp add: nth_append is_relevant_entry_def split: if_split_asm) - apply (drule_tac x = "\s\<^sub>1!pc" in bspec) - apply (blast intro: nth_mem) - apply (fastforce simp add: ty\<^sub>i'_def) -done +assumes wtis: "\ is\<^sub>1 @ is\<^sub>2, xt [::] \s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \s\<^sub>2" + and P_\s\<^sub>1: "\\ \ set \s\<^sub>1. \ST' LT'. \ = Some(ST',LT') \ + size ST \ size ST' \ P \ Some (drop (size ST' - size ST) ST',LT') \' ty\<^sub>i' ST E A" + and is_\s: "size is\<^sub>1 = size \s\<^sub>1" and PC: "is_class P C" and ST_mxs: "size ST < mxs" +shows "\ is\<^sub>1 @ is\<^sub>2, xt @ [(0,size is\<^sub>1 - 1,C,size is\<^sub>1,size ST)] [::] \s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \s\<^sub>2" +(*<*)(is "\ ?is, xt@[?xte] [::] ?\s" is "\ ?is, ?xt' [::] ?\s") +proof - + let ?is = "is\<^sub>1 @ is\<^sub>2" and ?\s = "\s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \s\<^sub>2" + let ?xte = "(0,size is\<^sub>1 - 1,C,size is\<^sub>1,size ST)" + let ?xt' = "xt @ [?xte]" + have P_\s\<^sub>1': "\\. \ \ set \s\<^sub>1 \ (\ST' LT'. \ = Some(ST',LT') \ + size ST \ size ST' \ P \ Some (drop (size ST' - size ST) ST',LT') \' ty\<^sub>i' ST E A)" + using P_\s\<^sub>1 by fast + have "size ?is < size ?\s" and "pcs ?xt' \ {0..pc. pc< size ?is \ P,T\<^sub>r,mxs,size ?\s,xt \ ?is!pc,pc :: ?\s" + using wtis by(simp_all add:wt_instrs_def) + moreover { + fix pc let ?mpc = "size ?\s" and ?i = "?is!pc" and ?t = "?\s!pc" + assume "pc< size ?is" + then have wti: "P,T\<^sub>r,mxs,?mpc,xt \ ?i,pc :: ?\s" by(rule P_pc) + then have app: "app ?i P mxs T\<^sub>r pc ?mpc xt ?t" + and eff_ss: "\pc' \'. (pc',\') \ set (eff ?i P pc xt (?\s!pc)) + \ P \ \' \' ?\s!pc'" + by(fastforce simp add: wt_instr_def)+ + have "app ?i P mxs T\<^sub>r pc ?mpc ?xt' ?t + \ (\(pc',\') \ set (eff ?i P pc ?xt' ?t). P \ \' \' ?\s!pc')" + proof (cases ?t) + case (Some \) + obtain ST' LT' where \: "\ = (ST', LT')" by(cases \) simp + have app\<^sub>i: "app\<^sub>i (?i,P,pc,mxs,T\<^sub>r,\)" and xcpt_app: "xcpt_app ?i P pc mxs xt \" + and eff_pc: "\pc' \'. (pc',\') \ set (eff ?i P pc xt ?t) \ pc' < ?mpc" + using Some app by(fastforce simp add: app_def)+ + have "xcpt_app ?i P pc mxs ?xt' \" + proof(cases "pc < length \s\<^sub>1 - 1") + case False then show ?thesis using Some \ is_\s xcpt_app + by (clarsimp simp: xcpt_app_def relevant_entries_def + is_relevant_entry_def) + next + case True + then have True': "pc < length \s\<^sub>1" by simp + then have "\s\<^sub>1 ! pc = ?t" by(fastforce simp: nth_append) + moreover have \s\<^sub>1_pc: "\s\<^sub>1 ! pc \ set \s\<^sub>1" by(rule nth_mem[OF True']) + ultimately show ?thesis + using Some \ PC ST_mxs xcpt_app P_\s\<^sub>1'[OF \s\<^sub>1_pc] + by (simp add: xcpt_app_def relevant_entries_def) + qed + moreover { + fix pc' \' assume efft: "(pc',\') \ set (eff ?i P pc ?xt' ?t)" + have "pc' < ?mpc \ P \ \' \' ?\s!pc'" (is "?P1 \ ?P2") + proof(cases "(pc',\') \ set (eff ?i P pc xt ?t)") + case True + have ?P1 using True by(rule eff_pc) + moreover have ?P2 using True by(rule eff_ss) + ultimately show ?thesis by simp + next + case False + then have xte: "(pc',\') \ set (xcpt_eff ?i P pc \ [?xte])" + using efft Some by(clarsimp simp: eff_def) + then have ?P1 using Some \ is_\s + by(clarsimp simp: xcpt_eff_def relevant_entries_def split: if_split_asm) + moreover have ?P2 + proof(cases "pc < length \s\<^sub>1 - 1") + case False': False + then show ?thesis using False Some \ xte is_\s + by(simp add: xcpt_eff_def relevant_entries_def is_relevant_entry_def) + next + case True + then have True': "pc < length \s\<^sub>1" by simp + have \s\<^sub>1_pc: "\s\<^sub>1 ! pc \ set \s\<^sub>1" by(rule nth_mem[OF True']) + have "P \ \(Class C # drop (length ST' - length ST) ST', LT')\ + \' ty\<^sub>i' (Class C # ST) E A" + using True' Some \ P_\s\<^sub>1'[OF \s\<^sub>1_pc] + by (fastforce simp: nth_append ty\<^sub>i'_def) + then show ?thesis using \ xte is_\s + by(simp add: xcpt_eff_def relevant_entries_def split: if_split_asm) + qed + ultimately show ?thesis by simp + qed + } + ultimately show ?thesis using Some app\<^sub>i by(fastforce simp add: app_def) + qed simp + then have "P,T\<^sub>r,mxs,size ?\s,?xt' \ ?is!pc,pc :: ?\s" + by(simp add: wt_instr_def) + } + ultimately show ?thesis by(simp add:wt_instrs_def) +qed declare [[simproc add: list_to_set_comprehension]] declare nth_append[simp] (*>*) lemma drop_Cons_Suc: "\xs. drop n xs = y#ys \ drop (Suc n) xs = ys" - apply (induct n) - apply simp - apply (simp add: drop_Suc) - done +proof(induct n) + case (Suc n) then show ?case by(simp add: drop_Suc) +qed simp lemma drop_mess: - "\Suc (length xs\<^sub>0) \ length xs; drop (length xs - Suc (length xs\<^sub>0)) xs = x # xs\<^sub>0\ - \ drop (length xs - length xs\<^sub>0) xs = xs\<^sub>0" -apply (cases xs) - apply simp -apply (simp add: Suc_diff_le) -apply (case_tac "length list - length xs\<^sub>0") - apply simp -apply (simp add: drop_Cons_Suc) -done +assumes "Suc (length xs\<^sub>0) \ length xs" + and "drop (length xs - Suc (length xs\<^sub>0)) xs = x # xs\<^sub>0" +shows "drop (length xs - length xs\<^sub>0) xs = xs\<^sub>0" +using assms proof(cases xs) + case (Cons a list) then show ?thesis using assms + proof(cases "length list - length xs\<^sub>0") + case Suc then show ?thesis using Cons assms + by (simp add: Suc_diff_le drop_Cons_Suc) + qed simp +qed simp (*<*) declare (in TC0) after_def[simp] pair_eq_ty\<^sub>i'_conv[simp] (*>*) lemma (in TC1) compT_ST_prefix: "\E A ST\<^sub>0. \(ST,LT)\ \ set(compT E A ST\<^sub>0 e) \ size ST\<^sub>0 \ size ST \ drop (size ST - size ST\<^sub>0) ST = ST\<^sub>0" and "\E A ST\<^sub>0. \(ST,LT)\ \ set(compTs E A ST\<^sub>0 es) \ size ST\<^sub>0 \ size ST \ drop (size ST - size ST\<^sub>0) ST = ST\<^sub>0" (*<*) proof(induct e and es rule: compT.induct compTs.induct) case (FAss e\<^sub>1 F D e\<^sub>2) moreover { let ?ST\<^sub>0 = "ty E e\<^sub>1 # ST\<^sub>0" fix A assume "\(ST, LT)\ \ set (compT E A ?ST\<^sub>0 e\<^sub>2)" with FAss have "length ?ST\<^sub>0 \ length ST \ drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast hence ?case by (clarsimp simp add: drop_mess) } ultimately show ?case by auto next case TryCatch thus ?case by auto next case Block thus ?case by auto next case Seq thus ?case by auto next case While thus ?case by auto next case Cond thus ?case by auto next case (Call e M es) moreover { let ?ST\<^sub>0 = "ty E e # ST\<^sub>0" fix A assume "\(ST, LT)\ \ set (compTs E A ?ST\<^sub>0 es)" with Call have "length ?ST\<^sub>0 \ length ST \ drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast hence ?case by (clarsimp simp add: drop_mess) } ultimately show ?case by auto next case (Cons_exp e es) moreover { let ?ST\<^sub>0 = "ty E e # ST\<^sub>0" fix A assume "\(ST, LT)\ \ set (compTs E A ?ST\<^sub>0 es)" with Cons_exp have "length ?ST\<^sub>0 \ length ST \ drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast hence ?case by (clarsimp simp add: drop_mess) } ultimately show ?case by auto next case (BinOp e\<^sub>1 bop e\<^sub>2) moreover { let ?ST\<^sub>0 = "ty E e\<^sub>1 # ST\<^sub>0" fix A assume "\(ST, LT)\ \ set (compT E A ?ST\<^sub>0 e\<^sub>2)" with BinOp have "length ?ST\<^sub>0 \ length ST \ drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast hence ?case by (clarsimp simp add: drop_mess) } ultimately show ?case by auto next case new thus ?case by auto next case Val thus ?case by auto next case Cast thus ?case by auto next case Var thus ?case by auto next case LAss thus ?case by auto next case throw thus ?case by auto next case FAcc thus ?case by auto next case Nil_exp thus ?case by auto qed declare (in TC0) after_def[simp del] pair_eq_ty\<^sub>i'_conv[simp del] (*>*) (* FIXME *) lemma fun_of_simp [simp]: "fun_of S x y = ((x,y) \ S)" (*<*) by (simp add: fun_of_def)(*>*) theorem (in TC2) compT_wt_instrs: "\E T A ST. \ P,E \\<^sub>1 e :: T; \ e A; \ e (size E); size ST + max_stack e \ mxs; size E + max_vars e \ mxl \ \ \ compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ty\<^sub>i' ST E A # compT E A ST e @ [after E A ST e]" (*<*)(is "\E T A ST. PROP ?P e E T A ST")(*>*) and "\E Ts A ST. \ P,E \\<^sub>1 es[::]Ts; \s es A; \s es (size E); size ST + max_stacks es \ mxs; size E + max_varss es \ mxl \ \ let \s = ty\<^sub>i' ST E A # compTs E A ST es in \ compEs\<^sub>2 es,compxEs\<^sub>2 es 0 (size ST) [::] \s \ last \s = ty\<^sub>i' (rev Ts @ ST) E (A \ \s es)" (*<*) (is "\E Ts A ST. PROP ?Ps es E Ts A ST") proof(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) case (TryCatch e\<^sub>1 C i e\<^sub>2) hence [simp]: "i = size E" by simp have wt\<^sub>1: "P,E \\<^sub>1 e\<^sub>1 :: T" and wt\<^sub>2: "P,E@[Class C] \\<^sub>1 e\<^sub>2 :: T" and "class": "is_class P C" using TryCatch by auto let ?A\<^sub>1 = "A \ \ e\<^sub>1" let ?A\<^sub>i = "A \ \{i}\" let ?E\<^sub>i = "E @ [Class C]" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>1 = "compT E A ST e\<^sub>1" let ?\\<^sub>1 = "ty\<^sub>i' (T#ST) E ?A\<^sub>1" let ?\\<^sub>2 = "ty\<^sub>i' (Class C#ST) E A" let ?\\<^sub>3 = "ty\<^sub>i' ST ?E\<^sub>i ?A\<^sub>i" let ?\s\<^sub>2 = "compT ?E\<^sub>i ?A\<^sub>i ST e\<^sub>2" let ?\\<^sub>2' = "ty\<^sub>i' (T#ST) ?E\<^sub>i (?A\<^sub>i \ \ e\<^sub>2)" let ?\' = "ty\<^sub>i' (T#ST) E (A \ \ e\<^sub>1 \ (\ e\<^sub>2 \ i))" let ?go = "Goto (int(size(compE\<^sub>2 e\<^sub>2)) + 2)" have "PROP ?P e\<^sub>2 ?E\<^sub>i T ?A\<^sub>i ST" by fact hence "\ compE\<^sub>2 e\<^sub>2,compxE\<^sub>2 e\<^sub>2 0 (size ST) [::] (?\\<^sub>3 # ?\s\<^sub>2) @ [?\\<^sub>2']" using TryCatch.prems by(auto simp:after_def) also have "?A\<^sub>i \ \ e\<^sub>2 = (A \ \ e\<^sub>2) \ \{size E}\" by(fastforce simp:hyperset_defs) also have "P \ ty\<^sub>i' (T#ST) ?E\<^sub>i \ \' ty\<^sub>i' (T#ST) E (A \ \ e\<^sub>2)" by(simp add:hyperset_defs ty\<^sub>l_incr ty\<^sub>i'_def) also have "P \ \ \' ty\<^sub>i' (T#ST) E (A \ \ e\<^sub>1 \ (\ e\<^sub>2 \ i))" by(auto intro!: ty\<^sub>l_antimono simp:hyperset_defs ty\<^sub>i'_def) also have "(?\\<^sub>3 # ?\s\<^sub>2) @ [?\'] = ?\\<^sub>3 # ?\s\<^sub>2 @ [?\']" by simp also have "\ [Store i],[] [::] ?\\<^sub>2 # [] @ [?\\<^sub>3]" using TryCatch.prems by(auto simp:nth_list_update wt_defs ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth hyperset_defs) also have "[] @ (?\\<^sub>3 # ?\s\<^sub>2 @ [?\']) = (?\\<^sub>3 # ?\s\<^sub>2 @ [?\'])" by simp also have "P,T\<^sub>r,mxs,size(compE\<^sub>2 e\<^sub>2)+3,[] \ ?go,0 :: ?\\<^sub>1#?\\<^sub>2#?\\<^sub>3#?\s\<^sub>2 @ [?\']" by (auto simp: hyperset_defs ty\<^sub>i'_def wt_defs nth_Cons nat_add_distrib fun_of_def intro: ty\<^sub>l_antimono list_all2_refl split:nat.split) also have "\ compE\<^sub>2 e\<^sub>1,compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\ # ?\s\<^sub>1 @ [?\\<^sub>1]" using TryCatch by(auto simp:after_def) also have "?\ # ?\s\<^sub>1 @ ?\\<^sub>1 # ?\\<^sub>2 # ?\\<^sub>3 # ?\s\<^sub>2 @ [?\'] = (?\ # ?\s\<^sub>1 @ [?\\<^sub>1]) @ ?\\<^sub>2 # ?\\<^sub>3 # ?\s\<^sub>2 @ [?\']" by simp also have "compE\<^sub>2 e\<^sub>1 @ ?go # [Store i] @ compE\<^sub>2 e\<^sub>2 = (compE\<^sub>2 e\<^sub>1 @ [?go]) @ (Store i # compE\<^sub>2 e\<^sub>2)" by simp also let "?Q \" = "\ST' LT'. \ = \(ST', LT')\ \ size ST \ size ST' \ P \ Some (drop (size ST' - size ST) ST',LT') \' ty\<^sub>i' ST E A" { have "?Q (ty\<^sub>i' ST E A)" by (clarsimp simp add: ty\<^sub>i'_def) moreover have "?Q (ty\<^sub>i' (T # ST) E ?A\<^sub>1)" by (fastforce simp add: ty\<^sub>i'_def hyperset_defs intro!: ty\<^sub>l_antimono) moreover have "\\. \ \ set (compT E A ST e\<^sub>1) \ ?Q \" using TryCatch.prems by clarsimp (frule compT_ST_prefix, fastforce dest!: compT_LT_prefix simp add: ty\<^sub>i'_def) ultimately have "\\\set (ty\<^sub>i' ST E A # compT E A ST e\<^sub>1 @ [ty\<^sub>i' (T # ST) E ?A\<^sub>1]). ?Q \" by auto } also from TryCatch.prems max_stack1[of e\<^sub>1] have "size ST + 1 \ mxs" by auto ultimately show ?case using wt\<^sub>1 wt\<^sub>2 TryCatch.prems "class" by (simp add:after_def) next case new thus ?case by(auto simp add:after_def wt_New) next case (BinOp e\<^sub>1 bop e\<^sub>2) let ?op = "case bop of Eq \ [CmpEq] | Add \ [IAdd]" have T: "P,E \\<^sub>1 e\<^sub>1 \bop\ e\<^sub>2 :: T" by fact then obtain T\<^sub>1 T\<^sub>2 where T\<^sub>1: "P,E \\<^sub>1 e\<^sub>1 :: T\<^sub>1" and T\<^sub>2: "P,E \\<^sub>1 e\<^sub>2 :: T\<^sub>2" and bopT: "case bop of Eq \ (P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1) \ T = Boolean | Add \ T\<^sub>1 = Integer \ T\<^sub>2 = Integer \ T = Integer" by auto let ?A\<^sub>1 = "A \ \ e\<^sub>1" let ?A\<^sub>2 = "?A\<^sub>1 \ \ e\<^sub>2" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>1 = "compT E A ST e\<^sub>1" let ?\\<^sub>1 = "ty\<^sub>i' (T\<^sub>1#ST) E ?A\<^sub>1" let ?\s\<^sub>2 = "compT E ?A\<^sub>1 (T\<^sub>1#ST) e\<^sub>2" let ?\\<^sub>2 = "ty\<^sub>i' (T\<^sub>2#T\<^sub>1#ST) E ?A\<^sub>2" let ?\' = "ty\<^sub>i' (T#ST) E ?A\<^sub>2" from bopT have "\ ?op,[] [::] [?\\<^sub>2,?\']" by (cases bop) (auto simp add: wt_CmpEq wt_IAdd) also have "PROP ?P e\<^sub>2 E T\<^sub>2 ?A\<^sub>1 (T\<^sub>1#ST)" by fact with BinOp.prems T\<^sub>2 have "\ compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size (T\<^sub>1#ST)) [::] ?\\<^sub>1#?\s\<^sub>2@[?\\<^sub>2]" by (auto simp: after_def) also from BinOp T\<^sub>1 have "\ compE\<^sub>2 e\<^sub>1, compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\#?\s\<^sub>1@[?\\<^sub>1]" by (auto simp: after_def) finally show ?case using T T\<^sub>1 T\<^sub>2 by (simp add: after_def hyperUn_assoc) next case (Cons_exp e es) have "P,E \\<^sub>1 e # es [::] Ts" by fact then obtain T\<^sub>e Ts' where T\<^sub>e: "P,E \\<^sub>1 e :: T\<^sub>e" and Ts': "P,E \\<^sub>1 es [::] Ts'" and Ts: "Ts = T\<^sub>e#Ts'" by auto let ?A\<^sub>e = "A \ \ e" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>e = "compT E A ST e" let ?\\<^sub>e = "ty\<^sub>i' (T\<^sub>e#ST) E ?A\<^sub>e" let ?\s' = "compTs E ?A\<^sub>e (T\<^sub>e#ST) es" let ?\s = "?\ # ?\s\<^sub>e @ (?\\<^sub>e # ?\s')" have Ps: "PROP ?Ps es E Ts' ?A\<^sub>e (T\<^sub>e#ST)" by fact with Cons_exp.prems T\<^sub>e Ts' have "\ compEs\<^sub>2 es, compxEs\<^sub>2 es 0 (size (T\<^sub>e#ST)) [::] ?\\<^sub>e#?\s'" by (simp add: after_def) also from Cons_exp T\<^sub>e have "\ compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ?\#?\s\<^sub>e@[?\\<^sub>e]" by (auto simp: after_def) moreover from Ps Cons_exp.prems T\<^sub>e Ts' Ts have "last ?\s = ty\<^sub>i' (rev Ts@ST) E (?A\<^sub>e \ \s es)" by simp ultimately show ?case using T\<^sub>e by (simp add: after_def hyperUn_assoc) next case (FAss e\<^sub>1 F D e\<^sub>2) hence Void: "P,E \\<^sub>1 e\<^sub>1\F{D} := e\<^sub>2 :: Void" by auto then obtain C T T' where C: "P,E \\<^sub>1 e\<^sub>1 :: Class C" and sees: "P \ C sees F:T in D" and T': "P,E \\<^sub>1 e\<^sub>2 :: T'" and T'_T: "P \ T' \ T" by auto let ?A\<^sub>1 = "A \ \ e\<^sub>1" let ?A\<^sub>2 = "?A\<^sub>1 \ \ e\<^sub>2" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>1 = "compT E A ST e\<^sub>1" let ?\\<^sub>1 = "ty\<^sub>i' (Class C#ST) E ?A\<^sub>1" let ?\s\<^sub>2 = "compT E ?A\<^sub>1 (Class C#ST) e\<^sub>2" let ?\\<^sub>2 = "ty\<^sub>i' (T'#Class C#ST) E ?A\<^sub>2" let ?\\<^sub>3 = "ty\<^sub>i' ST E ?A\<^sub>2" let ?\' = "ty\<^sub>i' (Void#ST) E ?A\<^sub>2" from FAss.prems sees T'_T have "\ [Putfield F D,Push Unit],[] [::] [?\\<^sub>2,?\\<^sub>3,?\']" by (fastforce simp add: wt_Push wt_Put) also have "PROP ?P e\<^sub>2 E T' ?A\<^sub>1 (Class C#ST)" by fact with FAss.prems T' have "\ compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size ST+1) [::] ?\\<^sub>1#?\s\<^sub>2@[?\\<^sub>2]" by (auto simp add: after_def hyperUn_assoc) also from FAss C have "\ compE\<^sub>2 e\<^sub>1, compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\#?\s\<^sub>1@[?\\<^sub>1]" by (auto simp add: after_def) finally show ?case using Void C T' by (simp add: after_def hyperUn_assoc) next case Val thus ?case by(auto simp:after_def wt_Push) next case Cast thus ?case by (auto simp:after_def wt_Cast) next case (Block i T\<^sub>i e) let ?\s = "ty\<^sub>i' ST E A # compT (E @ [T\<^sub>i]) (A\i) ST e" have IH: "PROP ?P e (E@[T\<^sub>i]) T (A\i) ST" by fact hence "\ compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ?\s @ [ty\<^sub>i' (T#ST) (E@[T\<^sub>i]) (A\(size E) \ \ e)]" using Block.prems by (auto simp add: after_def) also have "P \ ty\<^sub>i' (T # ST) (E@[T\<^sub>i]) (A \ size E \ \ e) \' ty\<^sub>i' (T # ST) (E@[T\<^sub>i]) ((A \ \ e) \ size E)" by(auto simp add:hyperset_defs intro: ty\<^sub>i'_antimono) also have "\ = ty\<^sub>i' (T # ST) E (A \ \ e)" by simp also have "P \ \ \' ty\<^sub>i' (T # ST) E (A \ (\ e \ i))" by(auto simp add:hyperset_defs intro: ty\<^sub>i'_antimono) finally show ?case using Block.prems by(simp add: after_def) next case Var thus ?case by(auto simp:after_def wt_Load) next case FAcc thus ?case by(auto simp:after_def wt_Get) next case (LAss i e) thus ?case using max_stack1[of e] by(auto simp: hyper_insert_comm after_def wt_Store wt_Push) next case Nil_exp thus ?case by auto next case throw thus ?case by(auto simp add: after_def wt_Throw) next case (While e c) obtain Tc where wte: "P,E \\<^sub>1 e :: Boolean" and wtc: "P,E \\<^sub>1 c :: Tc" and [simp]: "T = Void" using While by auto have [simp]: "ty E (while (e) c) = Void" using While by simp let ?A\<^sub>0 = "A \ \ e" let ?A\<^sub>1 = "?A\<^sub>0 \ \ c" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>e = "compT E A ST e" let ?\\<^sub>e = "ty\<^sub>i' (Boolean#ST) E ?A\<^sub>0" let ?\\<^sub>1 = "ty\<^sub>i' ST E ?A\<^sub>0" let ?\s\<^sub>c = "compT E ?A\<^sub>0 ST c" let ?\\<^sub>c = "ty\<^sub>i' (Tc#ST) E ?A\<^sub>1" let ?\\<^sub>2 = "ty\<^sub>i' ST E ?A\<^sub>1" let ?\' = "ty\<^sub>i' (Void#ST) E ?A\<^sub>0" let ?\s = "(?\ # ?\s\<^sub>e @ [?\\<^sub>e]) @ ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c, ?\\<^sub>2, ?\\<^sub>1, ?\']" have "\ [],[] [::] [] @ ?\s" by(simp add:wt_instrs_def) also have "PROP ?P e E Boolean A ST" by fact hence "\ compE\<^sub>2 e,compxE\<^sub>2 e 0 (size ST) [::] ?\ # ?\s\<^sub>e @ [?\\<^sub>e]" using While.prems by (auto simp:after_def) also have "[] @ ?\s = (?\ # ?\s\<^sub>e) @ ?\\<^sub>e # ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c,?\\<^sub>2,?\\<^sub>1,?\']" by simp also let ?n\<^sub>e = "size(compE\<^sub>2 e)" let ?n\<^sub>c = "size(compE\<^sub>2 c)" let ?if = "IfFalse (int ?n\<^sub>c + 3)" have "\ [?if],[] [::] ?\\<^sub>e # ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c, ?\\<^sub>2, ?\\<^sub>1, ?\']" by(simp add: wt_instr_Cons wt_instr_append wt_IfFalse nat_add_distrib split: nat_diff_split) also have "(?\ # ?\s\<^sub>e) @ (?\\<^sub>e # ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c, ?\\<^sub>2, ?\\<^sub>1, ?\']) = ?\s" by simp also have "PROP ?P c E Tc ?A\<^sub>0 ST" by fact hence "\ compE\<^sub>2 c,compxE\<^sub>2 c 0 (size ST) [::] ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c]" using While.prems wtc by (auto simp:after_def) also have "?\s = (?\ # ?\s\<^sub>e @ [?\\<^sub>e,?\\<^sub>1] @ ?\s\<^sub>c) @ [?\\<^sub>c,?\\<^sub>2,?\\<^sub>1,?\']" by simp also have "\ [Pop],[] [::] [?\\<^sub>c, ?\\<^sub>2]" by(simp add:wt_Pop) also have "(?\ # ?\s\<^sub>e @ [?\\<^sub>e,?\\<^sub>1] @ ?\s\<^sub>c) @ [?\\<^sub>c,?\\<^sub>2,?\\<^sub>1,?\'] = ?\s" by simp also let ?go = "Goto (-int(?n\<^sub>c+?n\<^sub>e+2))" have "P \ ?\\<^sub>2 \' ?\" by(fastforce intro: ty\<^sub>i'_antimono simp: hyperset_defs) hence "P,T\<^sub>r,mxs,size ?\s,[] \ ?go,?n\<^sub>e+?n\<^sub>c+2 :: ?\s" by(simp add: wt_Goto split: nat_diff_split) also have "?\s = (?\ # ?\s\<^sub>e @ [?\\<^sub>e,?\\<^sub>1] @ ?\s\<^sub>c @ [?\\<^sub>c, ?\\<^sub>2]) @ [?\\<^sub>1, ?\']" by simp also have "\ [Push Unit],[] [::] [?\\<^sub>1,?\']" using While.prems max_stack1[of c] by(auto simp add:wt_Push) finally show ?case using wtc wte by (simp add:after_def) next case (Cond e e\<^sub>1 e\<^sub>2) obtain T\<^sub>1 T\<^sub>2 where wte: "P,E \\<^sub>1 e :: Boolean" and wt\<^sub>1: "P,E \\<^sub>1 e\<^sub>1 :: T\<^sub>1" and wt\<^sub>2: "P,E \\<^sub>1 e\<^sub>2 :: T\<^sub>2" and sub\<^sub>1: "P \ T\<^sub>1 \ T" and sub\<^sub>2: "P \ T\<^sub>2 \ T" using Cond by auto have [simp]: "ty E (if (e) e\<^sub>1 else e\<^sub>2) = T" using Cond by simp let ?A\<^sub>0 = "A \ \ e" let ?A\<^sub>2 = "?A\<^sub>0 \ \ e\<^sub>2" let ?A\<^sub>1 = "?A\<^sub>0 \ \ e\<^sub>1" let ?A' = "?A\<^sub>0 \ \ e\<^sub>1 \ \ e\<^sub>2" let ?\\<^sub>2 = "ty\<^sub>i' ST E ?A\<^sub>0" let ?\' = "ty\<^sub>i' (T#ST) E ?A'" let ?\s\<^sub>2 = "compT E ?A\<^sub>0 ST e\<^sub>2" have "PROP ?P e\<^sub>2 E T\<^sub>2 ?A\<^sub>0 ST" by fact hence "\ compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size ST) [::] (?\\<^sub>2#?\s\<^sub>2) @ [ty\<^sub>i' (T\<^sub>2#ST) E ?A\<^sub>2]" using Cond.prems wt\<^sub>2 by(auto simp add:after_def) also have "P \ ty\<^sub>i' (T\<^sub>2#ST) E ?A\<^sub>2 \' ?\'" using sub\<^sub>2 by(auto simp add: hyperset_defs ty\<^sub>i'_def intro!: ty\<^sub>l_antimono) also let ?\\<^sub>3 = "ty\<^sub>i' (T\<^sub>1 # ST) E ?A\<^sub>1" let ?g\<^sub>2 = "Goto(int (size (compE\<^sub>2 e\<^sub>2) + 1))" from sub\<^sub>1 have "P,T\<^sub>r,mxs,size(compE\<^sub>2 e\<^sub>2)+2,[] \ ?g\<^sub>2,0 :: ?\\<^sub>3#(?\\<^sub>2#?\s\<^sub>2)@[?\']" by(auto simp: hyperset_defs wt_defs nth_Cons ty\<^sub>i'_def split:nat.split intro!: ty\<^sub>l_antimono) also let ?\s\<^sub>1 = "compT E ?A\<^sub>0 ST e\<^sub>1" have "PROP ?P e\<^sub>1 E T\<^sub>1 ?A\<^sub>0 ST" by fact hence "\ compE\<^sub>2 e\<^sub>1,compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\\<^sub>2 # ?\s\<^sub>1 @ [?\\<^sub>3]" using Cond.prems wt\<^sub>1 by(auto simp add:after_def) also let ?\s\<^sub>1\<^sub>2 = "?\\<^sub>2 # ?\s\<^sub>1 @ ?\\<^sub>3 # (?\\<^sub>2 # ?\s\<^sub>2) @ [?\']" let ?\\<^sub>1 = "ty\<^sub>i' (Boolean#ST) E ?A\<^sub>0" let ?g\<^sub>1 = "IfFalse(int (size (compE\<^sub>2 e\<^sub>1) + 2))" let ?code = "compE\<^sub>2 e\<^sub>1 @ ?g\<^sub>2 # compE\<^sub>2 e\<^sub>2" have "\ [?g\<^sub>1],[] [::] [?\\<^sub>1] @ ?\s\<^sub>1\<^sub>2" by(simp add: wt_IfFalse nat_add_distrib split:nat_diff_split) also (wt_instrs_ext2) have "[?\\<^sub>1] @ ?\s\<^sub>1\<^sub>2 = ?\\<^sub>1 # ?\s\<^sub>1\<^sub>2" by simp also let ?\ = "ty\<^sub>i' ST E A" have "PROP ?P e E Boolean A ST" by fact hence "\ compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ?\ # compT E A ST e @ [?\\<^sub>1]" using Cond.prems wte by(auto simp add:after_def) finally show ?case using wte wt\<^sub>1 wt\<^sub>2 by(simp add:after_def hyperUn_assoc) next case (Call e M es) obtain C D Ts m Ts' where C: "P,E \\<^sub>1 e :: Class C" and "method": "P \ C sees M:Ts \ T = m in D" and wtes: "P,E \\<^sub>1 es [::] Ts'" and subs: "P \ Ts' [\] Ts" using Call.prems by auto from wtes have same_size: "size es = size Ts'" by(rule WTs\<^sub>1_same_size) let ?A\<^sub>0 = "A \ \ e" let ?A\<^sub>1 = "?A\<^sub>0 \ \s es" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>e = "compT E A ST e" let ?\\<^sub>e = "ty\<^sub>i' (Class C # ST) E ?A\<^sub>0" let ?\s\<^sub>e\<^sub>s = "compTs E ?A\<^sub>0 (Class C # ST) es" let ?\\<^sub>1 = "ty\<^sub>i' (rev Ts' @ Class C # ST) E ?A\<^sub>1" let ?\' = "ty\<^sub>i' (T # ST) E ?A\<^sub>1" have "\ [Invoke M (size es)],[] [::] [?\\<^sub>1,?\']" by(rule wt_Invoke[OF same_size "method" subs]) also have "PROP ?Ps es E Ts' ?A\<^sub>0 (Class C # ST)" by fact hence "\ compEs\<^sub>2 es,compxEs\<^sub>2 es 0 (size ST+1) [::] ?\\<^sub>e # ?\s\<^sub>e\<^sub>s" "last (?\\<^sub>e # ?\s\<^sub>e\<^sub>s) = ?\\<^sub>1" using Call.prems wtes by(auto simp add:after_def) also have "(?\\<^sub>e # ?\s\<^sub>e\<^sub>s) @ [?\'] = ?\\<^sub>e # ?\s\<^sub>e\<^sub>s @ [?\']" by simp also have "\ compE\<^sub>2 e,compxE\<^sub>2 e 0 (size ST) [::] ?\ # ?\s\<^sub>e @ [?\\<^sub>e]" using Call C by(auto simp add:after_def) finally show ?case using Call.prems C by(simp add:after_def hyperUn_assoc) next case Seq thus ?case by(auto simp:after_def) (fastforce simp:wt_Push wt_Pop hyperUn_assoc intro:wt_instrs_app2 wt_instrs_Cons) qed (*>*) lemma [simp]: "types (compP f P) = types P" (*<*)by auto(*>*) lemma [simp]: "states (compP f P) mxs mxl = states P mxs mxl" (*<*)by (simp add: JVM_states_unfold)(*>*) lemma [simp]: "app\<^sub>i (i, compP f P, pc, mpc, T, \) = app\<^sub>i (i, P, pc, mpc, T, \)" -(*<*) - apply (cases \) - apply (cases i) - apply auto - apply (fastforce dest!: sees_method_compPD) - apply (force dest: sees_method_compP) - done +(*<*)(is "?A = ?B") +proof - + obtain ST LT where \: "\ = (ST, LT)" by(cases \) simp + then show ?thesis proof(cases i) + case Invoke show ?thesis + proof(rule iffI) + assume ?A then show ?B using Invoke \ + by auto (fastforce dest!: sees_method_compPD) + next + assume ?B then show ?A using Invoke \ + by auto (force dest: sees_method_compP) + qed + qed auto +qed (*>*) lemma [simp]: "is_relevant_entry (compP f P) i = is_relevant_entry P i" (*<*) - apply (rule ext)+ - apply (unfold is_relevant_entry_def) - apply (cases i) - apply auto - done +proof - + { fix pc e + have "is_relevant_entry (compP f P) i pc e = is_relevant_entry P i pc e" + by(cases i) (auto simp: is_relevant_entry_def) + } + then show ?thesis by fast +qed (*>*) lemma [simp]: "relevant_entries (compP f P) i pc xt = relevant_entries P i pc xt" (*<*) by (simp add: relevant_entries_def)(*>*) lemma [simp]: "app i (compP f P) mpc T pc mxl xt \ = app i P mpc T pc mxl xt \" (*<*) - apply (simp add: app_def xcpt_app_def eff_def xcpt_eff_def norm_eff_def) - apply (fastforce simp add: image_def) - done +by (simp add: app_def xcpt_app_def eff_def xcpt_eff_def norm_eff_def) + (fastforce simp add: image_def) (*>*) -lemma [simp]: "app i P mpc T pc mxl xt \ \ eff i (compP f P) pc xt \ = eff i P pc xt \" +lemma [simp]: +assumes "app i P mpc T pc mxl xt \" +shows "eff i (compP f P) pc xt \ = eff i P pc xt \" (*<*) - apply (clarsimp simp add: eff_def norm_eff_def xcpt_eff_def app_def) - apply (cases i) - apply auto - done +using assms +proof(clarsimp simp: eff_def norm_eff_def xcpt_eff_def app_def, cases i) +qed auto (*>*) lemma [simp]: "subtype (compP f P) = subtype P" -(*<*) - apply (rule ext)+ - apply (simp) - done -(*>*) +(*<*)by (rule ext)+ simp(*>*) lemma [simp]: "compP f P \ \ \' \' = P \ \ \' \'" (*<*) by (simp add: sup_state_opt_def sup_state_def sup_ty_opt_def)(*>*) lemma [simp]: "compP f P,T,mpc,mxl,xt \ i,pc :: \s = P,T,mpc,mxl,xt \ i,pc :: \s" (*<*)by (simp add: wt_instr_def cong: conj_cong)(*>*) declare TC1.compT_sizes[simp] TC0.ty_def2[simp] context TC2 begin lemma compT_method: fixes e and A and C and Ts and mxl\<^sub>0 defines [simp]: "E \ Class C # Ts" and [simp]: "A \ \{..size Ts}\" and [simp]: "A' \ A \ \ e" and [simp]: "mxl\<^sub>0 \ max_vars e" assumes mxs: "max_stack e = mxs" and mxl: "Suc (length Ts + max_vars e) = mxl" - assumes assm: "wf_prog p P" "P,E \\<^sub>1 e :: T" "\ e A" "\ e (size E)" - "set E \ types P" "P \ T \ T\<^sub>r" + assumes wf: "wf_prog p P" and wte: "P,E \\<^sub>1 e :: T" and \: "\ e A" + and \: "\ e (size E)" and E_P: "set E \ types P" and wid: "P \ T \ T\<^sub>r" shows "wt_method (compP\<^sub>2 P) C Ts T\<^sub>r mxs mxl\<^sub>0 (compE\<^sub>2 e @ [Return]) (compxE\<^sub>2 e 0 0) (ty\<^sub>i' [] E A # compT\<^sub>a E A [] e)" -(*<*) -using assms apply (simp add: wt_method_def compT\<^sub>a_def after_def mxl) -apply (rule conjI) -apply (simp add: check_types_def OK_ty\<^sub>i'_in_statesI) -apply (rule conjI) -apply (drule (1) WT\<^sub>1_is_type) -apply simp -apply (insert max_stack1 [of e]) -apply (rule OK_ty\<^sub>i'_in_statesI) apply (simp_all add: mxs)[3] -apply (erule compT_states(1)) -apply assumption -apply (simp_all add: mxs mxl)[4] -apply (rule conjI) -apply (auto simp add: wt_start_def ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth - nth_Cons mxl split: nat.split dest: less_antisym)[1] -apply (frule (1) TC2.compT_wt_instrs [of P _ _ _ _ "[]" "max_stack e" "Suc (length Ts + max_vars e)" T\<^sub>r]) -apply simp_all -apply (clarsimp simp: after_def) -apply hypsubst_thin -apply (rule conjI) -apply (clarsimp simp: wt_instrs_def after_def mxl mxs) -apply clarsimp -apply (drule (1) less_antisym) -apply (clarsimp simp: wt_defs xcpt_app_pcs xcpt_eff_pcs ty\<^sub>i'_def) -done +(*<*)(is "wt_method ?P C Ts T\<^sub>r mxs mxl\<^sub>0 ?is ?xt ?\s") +proof - + let ?n = "length E + mxl\<^sub>0" + have wt_compE: "P,T\<^sub>r,mxs \ compE\<^sub>2 e, compxE\<^sub>2 e 0 (length []) [::] + TC0.ty\<^sub>i' ?n [] E A # TC1.compT P ?n E A [] e @[TC0.after P ?n E A [] e]" + using mxs TC2.compT_wt_instrs [OF wte \ \, of "[]" mxs ?n T\<^sub>r] by simp + + have "OK (ty\<^sub>i' [T] E A') \ states P mxs mxl" + using mxs WT\<^sub>1_is_type[OF wf wte E_P] max_stack1[of e] OK_ty\<^sub>i'_in_statesI[OF E_P] + by simp + moreover have "OK ` set (compT E A [] e) \ states P mxs mxl" + using mxs mxl wid compT_states(1)[OF wf wte E_P] by simp + ultimately have "check_types ?P mxs ?n (map OK ?\s)" + using mxl wte E_P by (simp add: compT\<^sub>a_def after_def check_types_def) + moreover have "wt_start ?P C Ts mxl\<^sub>0 ?\s" using mxl + by (auto simp: wt_start_def ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth nth_Cons + split: nat.split) + moreover { + fix pc assume pc: "pc < size ?is" + then have "?P,T\<^sub>r,mxs,size ?is,?xt \ ?is!pc,pc :: ?\s" + proof(cases "pc < length (compE\<^sub>2 e)") + case True + then show ?thesis using mxs wte wt_compE + by (clarsimp simp: compT\<^sub>a_def mxl after_def wt_instrs_def) + next + case False + have "length (compE\<^sub>2 e) = pc" using less_antisym[OF False] pc by simp + then show ?thesis using mxl wte E_P wid + by (clarsimp simp: compT\<^sub>a_def after_def wt_defs xcpt_app_pcs xcpt_eff_pcs ty\<^sub>i'_def) + qed + } + moreover have "0 < size ?is" and "size ?\s = size ?is" + using wte by (simp_all add: compT\<^sub>a_def) + ultimately show ?thesis by(simp add: wt_method_def) +qed (*>*) end definition compTP :: "J\<^sub>1_prog \ ty\<^sub>P" where "compTP P C M = ( let (D,Ts,T,e) = method P C M; E = Class C # Ts; A = \{..size Ts}\; mxl = 1 + size Ts + max_vars e in (TC0.ty\<^sub>i' mxl [] E A # TC1.compT\<^sub>a P mxl E A [] e))" theorem wt_compP\<^sub>2: - "wf_J\<^sub>1_prog P \ wf_jvm_prog (compP\<^sub>2 P)" +assumes wf: "wf_J\<^sub>1_prog P" shows "wf_jvm_prog (compP\<^sub>2 P)" (*<*) - apply (simp add: wf_jvm_prog_def wf_jvm_prog_phi_def) - apply(rule_tac x = "compTP P" in exI) - apply (rule wf_prog_compPI) - prefer 2 apply assumption - apply (clarsimp simp add: wf_mdecl_def) - apply (simp add: compTP_def) - apply (rule TC2.compT_method [simplified]) - apply (rule refl) - apply (rule refl) - apply assumption - apply assumption - apply assumption - apply assumption - apply (drule (1) sees_wf_mdecl) - apply (simp add: wf_mdecl_def) - apply (blast intro: sees_method_is_class) - apply assumption - done +proof - + let ?\ = "compTP P" and ?f = compMb\<^sub>2 + let ?wf\<^sub>2 = "\P C (M, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). + wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (?\ C M)" + and ?P = "compP ?f P" + { fix C M Ts T m + assume cM: "P \ C sees M : Ts\T = m in C" + and wfm: "wf_mdecl wf_J\<^sub>1_mdecl P C (M,Ts,T,m)" + then have Ts_types: "\T'\set Ts. is_type P T'" + and T_type: "is_type P T" and wfm\<^sub>1: "wf_J\<^sub>1_mdecl P C (M,Ts,T,m)" + by(simp_all add: wf_mdecl_def) + then obtain T' where wte: "P,Class C#Ts \\<^sub>1 m :: T'" and wid: "P \ T' \ T" + and \: "\ m \{..size Ts}\" and \: "\ m (Suc (size Ts))" + by(auto simp: wf_mdecl_def) + have CTs_P: "is_class P C \ set Ts \ types P" + using sees_wf_mdecl[OF wf cM] sees_method_is_class[OF cM] + by (clarsimp simp: wf_mdecl_def) + + have "?wf\<^sub>2 ?P C (M,Ts,T,?f m)" + using cM TC2.compT_method [simplified, OF _ _ wf wte \ \ CTs_P wid] + by(simp add: compTP_def) + then have "wf_mdecl ?wf\<^sub>2 ?P C (M, Ts, T, ?f m)" + using Ts_types T_type by(simp add: wf_mdecl_def) + } + then have "wf_prog ?wf\<^sub>2 (compP ?f P)" by(rule wf_prog_compPI[OF _ wf]) + then show ?thesis by (simp add: wf_jvm_prog_def wf_jvm_prog_phi_def) fast +qed (*>*) theorem wt_J2JVM: "wf_J_prog P \ wf_jvm_prog (J2JVM P)" (*<*) -apply(simp only:o_def J2JVM_def) -apply(blast intro:wt_compP\<^sub>2 compP\<^sub>1_pres_wf) -done +by(simp only:o_def J2JVM_def) + (blast intro:wt_compP\<^sub>2 compP\<^sub>1_pres_wf) end diff --git a/thys/Jinja/J/BigStep.thy b/thys/Jinja/J/BigStep.thy --- a/thys/Jinja/J/BigStep.thy +++ b/thys/Jinja/J/BigStep.thy @@ -1,406 +1,422 @@ (* Title: Jinja/J/BigStep.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Big Step Semantics\ theory BigStep imports Expr State begin inductive eval :: "J_prog \ expr \ state \ expr \ state \ bool" ("_ \ ((1\_,/_\) \/ (1\_,/_\))" [51,0,0,0,0] 81) and evals :: "J_prog \ expr list \ state \ expr list \ state \ bool" ("_ \ ((1\_,/_\) [\]/ (1\_,/_\))" [51,0,0,0,0] 81) for P :: J_prog where New: "\ new_Addr h = Some a; P \ C has_fields FDTs; h' = h(a\(C,init_fields FDTs)) \ \ P \ \new C,(h,l)\ \ \addr a,(h',l)\" | NewFail: "new_Addr h = None \ P \ \new C, (h,l)\ \ \THROW OutOfMemory,(h,l)\" | Cast: "\ P \ \e,s\<^sub>0\ \ \addr a,(h,l)\; h a = Some(D,fs); P \ D \\<^sup>* C \ \ P \ \Cast C e,s\<^sub>0\ \ \addr a,(h,l)\" | CastNull: "P \ \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \ \Cast C e,s\<^sub>0\ \ \null,s\<^sub>1\" | CastFail: "\ P \ \e,s\<^sub>0\\ \addr a,(h,l)\; h a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \ \Cast C e,s\<^sub>0\ \ \THROW ClassCast,(h,l)\" | CastThrow: "P \ \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \ \Cast C e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | Val: "P \ \Val v,s\ \ \Val v,s\" | BinOp: "\ P \ \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1\ \ \Val v\<^sub>2,s\<^sub>2\; binop(bop,v\<^sub>1,v\<^sub>2) = Some v \ \ P \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\\\Val v,s\<^sub>2\" | BinOpThrow1: "P \ \e\<^sub>1,s\<^sub>0\ \ \throw e,s\<^sub>1\ \ P \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0\ \ \throw e,s\<^sub>1\" | BinOpThrow2: "\ P \ \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1\ \ \throw e,s\<^sub>2\ \ \ P \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\ \ \throw e,s\<^sub>2\" | Var: "l V = Some v \ P \ \Var V,(h,l)\ \ \Val v,(h,l)\" | LAss: "\ P \ \e,s\<^sub>0\ \ \Val v,(h,l)\; l' = l(V\v) \ \ P \ \V:=e,s\<^sub>0\ \ \unit,(h,l')\" | LAssThrow: "P \ \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \ \V:=e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAcc: "\ P \ \e,s\<^sub>0\ \ \addr a,(h,l)\; h a = Some(C,fs); fs(F,D) = Some v \ \ P \ \e\F{D},s\<^sub>0\ \ \Val v,(h,l)\" | FAccNull: "P \ \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \ \e\F{D},s\<^sub>0\ \ \THROW NullPointer,s\<^sub>1\" | FAccThrow: "P \ \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \ \e\F{D},s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAss: "\ P \ \e\<^sub>1,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1\ \ \Val v,(h\<^sub>2,l\<^sub>2)\; h\<^sub>2 a = Some(C,fs); fs' = fs((F,D)\v); h\<^sub>2' = h\<^sub>2(a\(C,fs')) \ \ P \ \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \unit,(h\<^sub>2',l\<^sub>2)\" | FAssNull: "\ P \ \e\<^sub>1,s\<^sub>0\ \ \null,s\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1\ \ \Val v,s\<^sub>2\ \ \ P \ \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \THROW NullPointer,s\<^sub>2\" | FAssThrow1: "P \ \e\<^sub>1,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \ \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAssThrow2: "\ P \ \e\<^sub>1,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1\ \ \throw e',s\<^sub>2\ \ \ P \ \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \throw e',s\<^sub>2\" | CallObjThrow: "P \ \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \ \e\M(ps),s\<^sub>0\ \ \throw e',s\<^sub>1\" | CallParamsThrow: "\ P \ \e,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \ \es,s\<^sub>1\ [\] \map Val vs @ throw ex # es',s\<^sub>2\ \ \ P \ \e\M(es),s\<^sub>0\ \ \throw ex,s\<^sub>2\" | CallNull: "\ P \ \e,s\<^sub>0\ \ \null,s\<^sub>1\; P \ \ps,s\<^sub>1\ [\] \map Val vs,s\<^sub>2\ \ \ P \ \e\M(ps),s\<^sub>0\ \ \THROW NullPointer,s\<^sub>2\" | Call: "\ P \ \e,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \ \ps,s\<^sub>1\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2)\; h\<^sub>2 a = Some(C,fs); P \ C sees M:Ts\T = (pns,body) in D; length vs = length pns; l\<^sub>2' = [this\Addr a, pns[\]vs]; P \ \body,(h\<^sub>2,l\<^sub>2')\ \ \e',(h\<^sub>3,l\<^sub>3)\ \ \ P \ \e\M(ps),s\<^sub>0\ \ \e',(h\<^sub>3,l\<^sub>2)\" | Block: "P \ \e\<^sub>0,(h\<^sub>0,l\<^sub>0(V:=None))\ \ \e\<^sub>1,(h\<^sub>1,l\<^sub>1)\ \ P \ \{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0)\ \ \e\<^sub>1,(h\<^sub>1,l\<^sub>1(V:=l\<^sub>0 V))\" | Seq: "\ P \ \e\<^sub>0,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \ \e\<^sub>1,s\<^sub>1\ \ \e\<^sub>2,s\<^sub>2\ \ \ P \ \e\<^sub>0;;e\<^sub>1,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" | SeqThrow: "P \ \e\<^sub>0,s\<^sub>0\ \ \throw e,s\<^sub>1\ \ P \ \e\<^sub>0;;e\<^sub>1,s\<^sub>0\\\throw e,s\<^sub>1\" | CondT: "\ P \ \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \ \e\<^sub>1,s\<^sub>1\ \ \e',s\<^sub>2\ \ \ P \ \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \ \e',s\<^sub>2\" | CondF: "\ P \ \e,s\<^sub>0\ \ \false,s\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1\ \ \e',s\<^sub>2\ \ \ P \ \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \ \e',s\<^sub>2\" | CondThrow: "P \ \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\ \ \throw e',s\<^sub>1\" | WhileF: "P \ \e,s\<^sub>0\ \ \false,s\<^sub>1\ \ P \ \while (e) c,s\<^sub>0\ \ \unit,s\<^sub>1\" | WhileT: "\ P \ \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \ \c,s\<^sub>1\ \ \Val v\<^sub>1,s\<^sub>2\; P \ \while (e) c,s\<^sub>2\ \ \e\<^sub>3,s\<^sub>3\ \ \ P \ \while (e) c,s\<^sub>0\ \ \e\<^sub>3,s\<^sub>3\" | WhileCondThrow: "P \ \e,s\<^sub>0\ \ \ throw e',s\<^sub>1\ \ P \ \while (e) c,s\<^sub>0\ \ \throw e',s\<^sub>1\" | WhileBodyThrow: "\ P \ \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \ \c,s\<^sub>1\ \ \throw e',s\<^sub>2\\ \ P \ \while (e) c,s\<^sub>0\ \ \throw e',s\<^sub>2\" | Throw: "P \ \e,s\<^sub>0\ \ \addr a,s\<^sub>1\ \ P \ \throw e,s\<^sub>0\ \ \Throw a,s\<^sub>1\" | ThrowNull: "P \ \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \ \throw e,s\<^sub>0\ \ \THROW NullPointer,s\<^sub>1\" | ThrowThrow: "P \ \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \ \throw e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | Try: "P \ \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\ \ P \ \try e\<^sub>1 catch(C V) e\<^sub>2,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\" | TryCatch: "\ P \ \e\<^sub>1,s\<^sub>0\ \ \Throw a,(h\<^sub>1,l\<^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))\ \ \e\<^sub>2',(h\<^sub>2,l\<^sub>2)\ \ \ P \ \try e\<^sub>1 catch(C V) e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',(h\<^sub>2,l\<^sub>2(V:=l\<^sub>1 V))\" | TryThrow: "\ P \ \e\<^sub>1,s\<^sub>0\ \ \Throw a,(h\<^sub>1,l\<^sub>1)\; h\<^sub>1 a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \ \try e\<^sub>1 catch(C V) e\<^sub>2,s\<^sub>0\ \ \Throw a,(h\<^sub>1,l\<^sub>1)\" | Nil: "P \ \[],s\ [\] \[],s\" | Cons: "\ P \ \e,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \ \es,s\<^sub>1\ [\] \es',s\<^sub>2\ \ \ P \ \e#es,s\<^sub>0\ [\] \Val v # es',s\<^sub>2\" | ConsThrow: "P \ \e, s\<^sub>0\ \ \throw e', s\<^sub>1\ \ P \ \e#es, s\<^sub>0\ [\] \throw e' # es, s\<^sub>1\" (*<*) lemmas eval_evals_induct = eval_evals.induct [split_format (complete)] and eval_evals_inducts = eval_evals.inducts [split_format (complete)] inductive_cases eval_cases [cases set]: "P \ \Cast C e,s\ \ \e',s'\" "P \ \Val v,s\ \ \e',s'\" "P \ \e\<^sub>1 \bop\ e\<^sub>2,s\ \ \e',s'\" "P \ \V:=e,s\ \ \e',s'\" "P \ \e\F{D},s\ \ \e',s'\" "P \ \e\<^sub>1\F{D}:=e\<^sub>2,s\ \ \e',s'\" "P \ \e\M{D}(es),s\ \ \e',s'\" "P \ \{V:T;e\<^sub>1},s\ \ \e',s'\" "P \ \e\<^sub>1;;e\<^sub>2,s\ \ \e',s'\" "P \ \if (e) e\<^sub>1 else e\<^sub>2,s\ \ \e',s'\" "P \ \while (b) c,s\ \ \e',s'\" "P \ \throw e,s\ \ \e',s'\" "P \ \try e\<^sub>1 catch(C V) e\<^sub>2,s\ \ \e',s'\" inductive_cases evals_cases [cases set]: "P \ \[],s\ [\] \e',s'\" "P \ \e#es,s\ [\] \e',s'\" (*>*) subsection"Final expressions" definition final :: "'a exp \ bool" where "final e \ (\v. e = Val v) \ (\a. e = Throw a)" definition finals:: "'a exp list \ bool" where "finals es \ (\vs. es = map Val vs) \ (\vs a es'. es = map Val vs @ Throw a # es')" lemma [simp]: "final(Val v)" (*<*)by(simp add:final_def)(*>*) lemma [simp]: "final(throw e) = (\a. e = addr a)" (*<*)by(simp add:final_def)(*>*) lemma finalE: "\ final e; \v. e = Val v \ R; \a. e = Throw a \ R \ \ R" (*<*)by(auto simp:final_def)(*>*) lemma [iff]: "finals []" (*<*)by(simp add:finals_def)(*>*) lemma [iff]: "finals (Val v # es) = finals es" (*<*) -apply(clarsimp simp add: finals_def) -apply(rule iffI) - apply(erule disjE) - apply simp - apply(rule disjI2) - apply clarsimp - apply(case_tac vs) - apply simp - apply fastforce -apply(erule disjE) - apply clarsimp -apply(rule disjI2) -apply clarsimp -apply(rule_tac x = "v#vs" in exI) -apply simp -done +proof(rule iffI) + assume "finals (Val v # es)" + moreover { + fix vs a es' + assume "\vs a es'. es \ map Val vs @ Throw a # es'" + and "Val v # es = map Val vs @ Throw a # es'" + then have "\vs. es = map Val vs" by(case_tac vs; simp) + } + ultimately show "finals es" by(clarsimp simp add: finals_def) +next + assume "finals es" + moreover { + fix vs a es' + assume "es = map Val vs @ Throw a # es'" + then have "\vs' a' es''. Val v # map Val vs @ Throw a # es' = map Val vs' @ Throw a' # es''" + by(rule_tac x = "v#vs" in exI) simp + } + ultimately show "finals (Val v # es)" by(clarsimp simp add: finals_def) +qed (*>*) lemma finals_app_map[iff]: "finals (map Val vs @ es) = finals es" (*<*)by(induct_tac vs, auto)(*>*) lemma [iff]: "finals (map Val vs)" (*<*)using finals_app_map[of vs "[]"]by(simp)(*>*) lemma [iff]: "finals (throw e # es) = (\a. e = addr a)" (*<*) -apply(simp add:finals_def) -apply(rule iffI) - apply clarsimp - apply(case_tac vs) - apply simp - apply fastforce -apply clarsimp -apply(rule_tac x = "[]" in exI) -apply simp -done +proof(rule iffI) + assume "finals (throw e # es)" + moreover { + fix vs a es' + assume "throw e # es = map Val vs @ Throw a # es'" + then have "\a. e = addr a" by(case_tac vs; simp) + } + ultimately show "\a. e = addr a" by(clarsimp simp add: finals_def) +next + assume "\a. e = addr a" + moreover { + fix vs a es' + assume "e = addr a" + then have "\vs aa es'. Throw a # es = map Val vs @ Throw aa # es'" + by(rule_tac x = "[]" in exI) simp + } + ultimately show "finals (throw e # es)" by(clarsimp simp add: finals_def) +qed (*>*) lemma not_finals_ConsI: "\ final e \ \ finals(e#es)" - (*<*) -apply(clarsimp simp add:finals_def final_def) -apply(case_tac vs) -apply auto -done +(*<*) +proof - + assume "\ final e" + moreover { + fix vs a es' + assume "\v. e \ Val v" and "\a. e \ Throw a" + then have "e # es \ map Val vs @ Throw a # es'" by(case_tac vs; simp) + } + ultimately show ?thesis by(clarsimp simp add:finals_def final_def) +qed (*>*) lemma eval_final: "P \ \e,s\ \ \e',s'\ \ final e'" and evals_final: "P \ \es,s\ [\] \es',s'\ \ finals es'" (*<*)by(induct rule:eval_evals.inducts, simp_all)(*>*) lemma eval_lcl_incr: "P \ \e,(h\<^sub>0,l\<^sub>0)\ \ \e',(h\<^sub>1,l\<^sub>1)\ \ dom l\<^sub>0 \ dom l\<^sub>1" and evals_lcl_incr: "P \ \es,(h\<^sub>0,l\<^sub>0)\ [\] \es',(h\<^sub>1,l\<^sub>1)\ \ dom l\<^sub>0 \ dom l\<^sub>1" (*<*) proof (induct rule: eval_evals_inducts) case BinOp show ?case by(rule subset_trans)(rule BinOp.hyps)+ next case Call thus ?case by(simp del: fun_upd_apply) next case Seq show ?case by(rule subset_trans)(rule Seq.hyps)+ next case CondT show ?case by(rule subset_trans)(rule CondT.hyps)+ next case CondF show ?case by(rule subset_trans)(rule CondF.hyps)+ next case WhileT thus ?case by(blast) next case TryCatch thus ?case by(clarsimp simp:dom_def split:if_split_asm) blast next case Cons show ?case by(rule subset_trans)(rule Cons.hyps)+ next case Block thus ?case by(auto simp del:fun_upd_apply) qed auto (*>*) text\Only used later, in the small to big translation, but is already a good sanity check:\ lemma eval_finalId: "final e \ P \ \e,s\ \ \e,s\" (*<*)by (erule finalE) (iprover intro: eval_evals.intros)+(*>*) lemma eval_finalsId: assumes finals: "finals es" shows "P \ \es,s\ [\] \es,s\" (*<*) using finals proof (induct es type: list) case Nil show ?case by (rule eval_evals.intros) next case (Cons e es) have hyp: "finals es \ P \ \es,s\ [\] \es,s\" and finals: "finals (e # es)" by fact+ show "P \ \e # es,s\ [\] \e # es,s\" proof cases assume "final e" thus ?thesis proof (cases rule: finalE) fix v assume e: "e = Val v" have "P \ \Val v,s\ \ \Val v,s\" by (simp add: eval_finalId) moreover from finals e have "P \ \es,s\ [\] \es,s\" by(fast intro:hyp) ultimately have "P \ \Val v#es,s\ [\] \Val v#es,s\" by (rule eval_evals.intros) with e show ?thesis by simp next fix a assume e: "e = Throw a" have "P \ \Throw a,s\ \ \Throw a,s\" by (simp add: eval_finalId) hence "P \ \Throw a#es,s\ [\] \Throw a#es,s\" by (rule eval_evals.intros) with e show ?thesis by simp qed next assume "\ final e" with not_finals_ConsI finals have False by blast thus ?thesis .. qed qed (*>*) theorem eval_hext: "P \ \e,(h,l)\ \ \e',(h',l')\ \ h \ h'" and evals_hext: "P \ \es,(h,l)\ [\] \es',(h',l')\ \ h \ h'" (*<*) proof (induct rule: eval_evals_inducts) case New thus ?case by(fastforce intro!: hext_new intro:LeastI simp:new_Addr_def split:if_split_asm simp del:fun_upd_apply) next case BinOp thus ?case by (fast elim!:hext_trans) next case BinOpThrow2 thus ?case by(fast elim!: hext_trans) next case FAss thus ?case by(auto simp:sym[THEN hext_upd_obj] simp del:fun_upd_apply elim!: hext_trans) next case FAssNull thus ?case by (fast elim!:hext_trans) next case FAssThrow2 thus ?case by (fast elim!:hext_trans) next case CallParamsThrow thus ?case by(fast elim!: hext_trans) next case CallNull thus ?case by(fast elim!: hext_trans) next case Call thus ?case by(fast elim!: hext_trans) next case Seq thus ?case by(fast elim!: hext_trans) next case CondT thus ?case by(fast elim!: hext_trans) next case CondF thus ?case by(fast elim!: hext_trans) next case WhileT thus ?case by(fast elim!: hext_trans) next case WhileBodyThrow thus ?case by (fast elim!: hext_trans) next case TryCatch thus ?case by(fast elim!: hext_trans) next case Cons thus ?case by (fast intro: hext_trans) qed auto (*>*) end diff --git a/thys/Jinja/J/DefAss.thy b/thys/Jinja/J/DefAss.thy --- a/thys/Jinja/J/DefAss.thy +++ b/thys/Jinja/J/DefAss.thy @@ -1,194 +1,194 @@ (* Title: Jinja/DefAss.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Definite assignment\ theory DefAss imports BigStep begin subsection "Hypersets" type_synonym 'a hyperset = "'a set option" definition hyperUn :: "'a hyperset \ 'a hyperset \ 'a hyperset" (infixl "\" 65) where "A \ B \ case A of None \ None | \A\ \ (case B of None \ None | \B\ \ \A \ B\)" definition hyperInt :: "'a hyperset \ 'a hyperset \ 'a hyperset" (infixl "\" 70) where "A \ B \ case A of None \ B | \A\ \ (case B of None \ \A\ | \B\ \ \A \ B\)" definition hyperDiff1 :: "'a hyperset \ 'a \ 'a hyperset" (infixl "\" 65) where "A \ a \ case A of None \ None | \A\ \ \A - {a}\" definition hyper_isin :: "'a \ 'a hyperset \ bool" (infix "\\" 50) where "a \\ A \ case A of None \ True | \A\ \ a \ A" definition hyper_subset :: "'a hyperset \ 'a hyperset \ bool" (infix "\" 50) where "A \ B \ case B of None \ True | \B\ \ (case A of None \ False | \A\ \ A \ B)" lemmas hyperset_defs = hyperUn_def hyperInt_def hyperDiff1_def hyper_isin_def hyper_subset_def lemma [simp]: "\{}\ \ A = A \ A \ \{}\ = A" (*<*)by(simp add:hyperset_defs)(*>*) lemma [simp]: "\A\ \ \B\ = \A \ B\ \ \A\ \ a = \A - {a}\" (*<*)by(simp add:hyperset_defs)(*>*) lemma [simp]: "None \ A = None \ A \ None = None" (*<*)by(simp add:hyperset_defs)(*>*) lemma [simp]: "a \\ None \ None \ a = None" (*<*)by(simp add:hyperset_defs)(*>*) lemma hyperUn_assoc: "(A \ B) \ C = A \ (B \ C)" (*<*)by(simp add:hyperset_defs Un_assoc)(*>*) lemma hyper_insert_comm: "A \ \{a}\ = \{a}\ \ A \ A \ (\{a}\ \ B) = \{a}\ \ (A \ B)" (*<*)by(simp add:hyperset_defs)(*>*) subsection "Definite assignment" primrec \ :: "'a exp \ 'a hyperset" and \s :: "'a exp list \ 'a hyperset" where "\ (new C) = \{}\" | "\ (Cast C e) = \ e" | "\ (Val v) = \{}\" | "\ (e\<^sub>1 \bop\ e\<^sub>2) = \ e\<^sub>1 \ \ e\<^sub>2" | "\ (Var V) = \{}\" | "\ (LAss V e) = \{V}\ \ \ e" | "\ (e\F{D}) = \ e" | "\ (e\<^sub>1\F{D}:=e\<^sub>2) = \ e\<^sub>1 \ \ e\<^sub>2" | "\ (e\M(es)) = \ e \ \s es" | "\ ({V:T; e}) = \ e \ V" | "\ (e\<^sub>1;;e\<^sub>2) = \ e\<^sub>1 \ \ e\<^sub>2" | "\ (if (e) e\<^sub>1 else e\<^sub>2) = \ e \ (\ e\<^sub>1 \ \ e\<^sub>2)" | "\ (while (b) e) = \ b" | "\ (throw e) = None" | "\ (try e\<^sub>1 catch(C V) e\<^sub>2) = \ e\<^sub>1 \ (\ e\<^sub>2 \ V)" | "\s ([]) = \{}\" | "\s (e#es) = \ e \ \s es" primrec \ :: "'a exp \ 'a hyperset \ bool" and \s :: "'a exp list \ 'a hyperset \ bool" where "\ (new C) A = True" | "\ (Cast C e) A = \ e A" | "\ (Val v) A = True" | "\ (e\<^sub>1 \bop\ e\<^sub>2) A = (\ e\<^sub>1 A \ \ e\<^sub>2 (A \ \ e\<^sub>1))" | "\ (Var V) A = (V \\ A)" | "\ (LAss V e) A = \ e A" | "\ (e\F{D}) A = \ e A" | "\ (e\<^sub>1\F{D}:=e\<^sub>2) A = (\ e\<^sub>1 A \ \ e\<^sub>2 (A \ \ e\<^sub>1))" | "\ (e\M(es)) A = (\ e A \ \s es (A \ \ e))" | "\ ({V:T; e}) A = \ e (A \ V)" | "\ (e\<^sub>1;;e\<^sub>2) A = (\ e\<^sub>1 A \ \ e\<^sub>2 (A \ \ e\<^sub>1))" | "\ (if (e) e\<^sub>1 else e\<^sub>2) A = (\ e A \ \ e\<^sub>1 (A \ \ e) \ \ e\<^sub>2 (A \ \ e))" | "\ (while (e) c) A = (\ e A \ \ c (A \ \ e))" | "\ (throw e) A = \ e A" | "\ (try e\<^sub>1 catch(C V) e\<^sub>2) A = (\ e\<^sub>1 A \ \ e\<^sub>2 (A \ \{V}\))" | "\s ([]) A = True" | "\s (e#es) A = (\ e A \ \s es (A \ \ e))" lemma As_map_Val[simp]: "\s (map Val vs) = \{}\" (*<*)by (induct vs) simp_all(*>*) lemma D_append[iff]: "\A. \s (es @ es') A = (\s es A \ \s es' (A \ \s es))" (*<*)by (induct es type:list) (auto simp:hyperUn_assoc)(*>*) lemma A_fv: "\A. \ e = \A\ \ A \ fv e" and "\A. \s es = \A\ \ A \ fvs es" (*<*) -apply(induct e and es rule: \.induct \s.induct) -apply (simp_all add:hyperset_defs) -apply blast+ -done +by (induct e and es rule: \.induct \s.induct) + (fastforce simp add:hyperset_defs)+ (*>*) lemma sqUn_lem: "A \ A' \ A \ B \ A' \ B" (*<*)by(simp add:hyperset_defs) blast(*>*) lemma diff_lem: "A \ A' \ A \ b \ A' \ b" (*<*)by(simp add:hyperset_defs) blast(*>*) (* This order of the premises avoids looping of the simplifier *) lemma D_mono: "\A A'. A \ A' \ \ e A \ \ (e::'a exp) A'" and Ds_mono: "\A A'. A \ A' \ \s es A \ \s (es::'a exp list) A'" (*<*) -apply(induct e and es rule: \.induct \s.induct) -apply simp -apply simp -apply simp -apply simp apply (iprover dest:sqUn_lem) -apply (fastforce simp add:hyperset_defs) -apply simp -apply simp -apply simp apply (iprover dest:sqUn_lem) -apply simp apply (iprover dest:sqUn_lem) -apply simp apply (iprover dest:diff_lem) -apply simp apply (iprover dest:sqUn_lem) -apply simp apply (iprover dest:sqUn_lem) -apply simp apply (iprover dest:sqUn_lem) -apply simp -apply simp apply (iprover dest:sqUn_lem) -apply simp -apply simp apply (iprover dest:sqUn_lem) -done +proof(induct e and es rule: \.induct \s.induct) + case BinOp then show ?case by simp (iprover dest:sqUn_lem) +next + case Var then show ?case by (fastforce simp add:hyperset_defs) +next + case FAss then show ?case by simp (iprover dest:sqUn_lem) +next + case Call then show ?case by simp (iprover dest:sqUn_lem) +next + case Block then show ?case by simp (iprover dest:diff_lem) +next + case Seq then show ?case by simp (iprover dest:sqUn_lem) +next + case Cond then show ?case by simp (iprover dest:sqUn_lem) +next + case While then show ?case by simp (iprover dest:sqUn_lem) +next + case TryCatch then show ?case by simp (iprover dest:sqUn_lem) +next + case Cons_exp then show ?case by simp (iprover dest:sqUn_lem) +qed simp_all (*>*) (* And this is the order of premises preferred during application: *) lemma D_mono': "\ e A \ A \ A' \ \ e A'" and Ds_mono': "\s es A \ A \ A' \ \s es A'" (*<*)by(blast intro:D_mono, blast intro:Ds_mono)(*>*) (* text{* @{term"\"} is sound w.r.t.\ the big step semantics: it computes a conservative approximation of the variables actually assigned to. *} lemma "P \ \e,(h\<^sub>0,l\<^sub>0)\ \ \e',(h\<^sub>1,l\<^sub>1)\ \ (!!A. \ e = \A\ \ A \ dom l\<^sub>1)" and "P \ \es,(h\<^sub>0,l\<^sub>0)\ [\] \es',(h\<^sub>1,l\<^sub>1)\ \ (!!A. \s es = \A\ \ A \ dom l\<^sub>1)" proof (induct rule:eval_evals_induct) case LAss thus ?case apply(simp add:dom_def hyperset_defs) apply blast apply blast next case FAss thus ?case by simp (blast dest:eval_lcl_incr) next case BinOp thus ?case by simp (blast dest:eval_lcl_incr) next case Call thus ?case by simp (blast dest:evals_lcl_incr) next case Cons thus ?case by simp (blast dest:evals_lcl_incr) next case Block thus ?case by(simp del: fun_upd_apply) blast next case Seq thus ?case by simp (blast dest:eval_lcl_incr) next case CondT thus ?case by simp (blast dest:eval_lcl_incr) next case CondF thus ?case by simp (blast dest:eval_lcl_incr) next case Try thus ?case by(fastforce dest!: eval_lcl_incr) next case TryCatch thus ?case by(fastforce dest!: eval_lcl_incr) qed auto *) end diff --git a/thys/Jinja/J/Equivalence.thy b/thys/Jinja/J/Equivalence.thy --- a/thys/Jinja/J/Equivalence.thy +++ b/thys/Jinja/J/Equivalence.thy @@ -1,1646 +1,1606 @@ (* Title: Jinja/J/Equivalence.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Equivalence of Big Step and Small Step Semantics\ theory Equivalence imports BigStep SmallStep WWellForm begin subsection\Small steps simulate big step\ subsubsection "Cast" lemma CastReds: "P \ \e,s\ \* \e',s'\ \ P \ \Cast C e,s\ \* \Cast C e',s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule CastRed) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) CastRed[OF step(2)]]) +qed (*>*) lemma CastRedsNull: "P \ \e,s\ \* \null,s'\ \ P \ \Cast C e,s\ \* \null,s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule CastReds) -apply(rule RedCastNull) -done -(*>*) +(*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCastNull])(*>*) lemma CastRedsAddr: "\ P \ \e,s\ \* \addr a,s'\; hp s' a = Some(D,fs); P \ D \\<^sup>* C \ \ P \ \Cast C e,s\ \* \addr a,s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule CastReds) -apply(erule (1) RedCast) -done -(*>*) +(*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCast])(*>*) lemma CastRedsFail: "\ P \ \e,s\ \* \addr a,s'\; hp s' a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \ \Cast C e,s\ \* \THROW ClassCast,s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule CastReds) -apply(erule (1) RedCastFail) -done -(*>*) +(*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCastFail])(*>*) lemma CastRedsThrow: "\ P \ \e,s\ \* \throw a,s'\ \ \ P \ \Cast C e,s\ \* \throw a,s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule CastReds) -apply(rule red_reds.CastThrow) -done -(*>*) +(*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ red_reds.CastThrow])(*>*) subsubsection "LAss" lemma LAssReds: "P \ \e,s\ \* \e',s'\ \ P \ \ V:=e,s\ \* \ V:=e',s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule LAssRed) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) LAssRed[OF step(2)]]) +qed (*>*) lemma LAssRedsVal: "\ P \ \e,s\ \* \Val v,(h',l')\ \ \ P \ \ V:=e,s\ \* \unit,(h',l'(V\v))\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule LAssReds) -apply(rule RedLAss) -done -(*>*) +(*<*)by(rule LAssReds[THEN rtrancl_into_rtrancl, OF _ RedLAss])(*>*) lemma LAssRedsThrow: "\ P \ \e,s\ \* \throw a,s'\ \ \ P \ \ V:=e,s\ \* \throw a,s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule LAssReds) -apply(rule red_reds.LAssThrow) -done -(*>*) +(*<*)by(rule LAssReds[THEN rtrancl_into_rtrancl, OF _ red_reds.LAssThrow])(*>*) subsubsection "BinOp" lemma BinOp1Reds: "P \ \e,s\ \* \e',s'\ \ P \ \ e \bop\ e\<^sub>2, s\ \* \e' \bop\ e\<^sub>2, s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule BinOpRed1) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) BinOpRed1[OF step(2)]]) +qed (*>*) lemma BinOp2Reds: "P \ \e,s\ \* \e',s'\ \ P \ \(Val v) \bop\ e, s\ \* \(Val v) \bop\ e', s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule BinOpRed2) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) BinOpRed2[OF step(2)]]) +qed (*>*) lemma BinOpRedsVal: - "\ P \ \e\<^sub>1,s\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1\ \* \Val v\<^sub>2,s\<^sub>2\; binop(bop,v\<^sub>1,v\<^sub>2) = Some v \ - \ P \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0\ \* \Val v,s\<^sub>2\" -(*<*) -apply(rule rtrancl_trans) - apply(erule BinOp1Reds) -apply(rule rtrancl_into_rtrancl) - apply(erule BinOp2Reds) -apply(rule RedBinOp) -apply simp -done +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1\ \* \Val v\<^sub>2,s\<^sub>2\" + and op: "binop(bop,v\<^sub>1,v\<^sub>2) = Some v" +shows "P \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0\ \* \Val v,s\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(Val v\<^sub>1 \bop\ e\<^sub>2, s\<^sub>1)" + let ?y' = "(Val v\<^sub>1 \bop\ Val v\<^sub>2, s\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule BinOp1Reds[OF e\<^sub>1_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule BinOp2Reds[OF e\<^sub>2_steps]) + also have "(?y', ?z) \ (red P)" by(rule RedBinOp[OF op]) + ultimately show ?thesis by simp +qed (*>*) lemma BinOpRedsThrow1: "P \ \e,s\ \* \throw e',s'\ \ P \ \e \bop\ e\<^sub>2, s\ \* \throw e', s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule BinOp1Reds) -apply(rule red_reds.BinOpThrow1) -done -(*>*) +(*<*)by(rule BinOp1Reds[THEN rtrancl_into_rtrancl, OF _ red_reds.BinOpThrow1])(*>*) lemma BinOpRedsThrow2: - "\ P \ \e\<^sub>1,s\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1\ \* \throw e,s\<^sub>2\\ - \ P \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0\ \* \throw e,s\<^sub>2\" -(*<*) -apply(rule rtrancl_trans) - apply(erule BinOp1Reds) -apply(rule rtrancl_into_rtrancl) - apply(erule BinOp2Reds) -apply(rule red_reds.BinOpThrow2) -done +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1\ \* \throw e,s\<^sub>2\" +shows "P \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0\ \* \throw e,s\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(Val v\<^sub>1 \bop\ e\<^sub>2, s\<^sub>1)" + let ?y' = "(Val v\<^sub>1 \bop\ throw e, s\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule BinOp1Reds[OF e\<^sub>1_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule BinOp2Reds[OF e\<^sub>2_steps]) + also have "(?y', ?z) \ (red P)" by(rule red_reds.BinOpThrow2) + ultimately show ?thesis by simp +qed (*>*) subsubsection "FAcc" lemma FAccReds: "P \ \e,s\ \* \e',s'\ \ P \ \e\F{D}, s\ \* \e'\F{D}, s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule FAccRed) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) FAccRed[OF step(2)]]) +qed (*>*) lemma FAccRedsVal: "\P \ \e,s\ \* \addr a,s'\; hp s' a = Some(C,fs); fs(F,D) = Some v \ \ P \ \e\F{D},s\ \* \Val v,s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule FAccReds) -apply(erule (1) RedFAcc) -done -(*>*) +(*<*)by(rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAcc])(*>*) lemma FAccRedsNull: "P \ \e,s\ \* \null,s'\ \ P \ \e\F{D},s\ \* \THROW NullPointer,s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule FAccReds) -apply(rule RedFAccNull) -done -(*>*) +(*<*)by(rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAccNull])(*>*) lemma FAccRedsThrow: "P \ \e,s\ \* \throw a,s'\ \ P \ \e\F{D},s\ \* \throw a,s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule FAccReds) -apply(rule red_reds.FAccThrow) -done -(*>*) +(*<*)by(rule FAccReds[THEN rtrancl_into_rtrancl, OF _ red_reds.FAccThrow])(*>*) subsubsection "FAss" lemma FAssReds1: "P \ \e,s\ \* \e',s'\ \ P \ \e\F{D}:=e\<^sub>2, s\ \* \e'\F{D}:=e\<^sub>2, s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule FAssRed1) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) FAssRed1[OF step(2)]]) +qed (*>*) lemma FAssReds2: "P \ \e,s\ \* \e',s'\ \ P \ \Val v\F{D}:=e, s\ \* \Val v\F{D}:=e', s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule FAssRed2) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) FAssRed2[OF step(2)]]) +qed (*>*) lemma FAssRedsVal: - "\ P \ \e\<^sub>1,s\<^sub>0\ \* \addr a,s\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1\ \* \Val v,(h\<^sub>2,l\<^sub>2)\; Some(C,fs) = h\<^sub>2 a \ \ - P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0\ \* \unit, (h\<^sub>2(a\(C,fs((F,D)\v))),l\<^sub>2)\" -(*<*) -apply(rule rtrancl_trans) - apply(erule FAssReds1) -apply(rule rtrancl_into_rtrancl) - apply(erule FAssReds2) -apply(rule RedFAss) -apply simp -done +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0\ \* \addr a,s\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1\ \* \Val v,(h\<^sub>2,l\<^sub>2)\" + and hC: "Some(C,fs) = h\<^sub>2 a" +shows "P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0\ \* \unit, (h\<^sub>2(a\(C,fs((F,D)\v))),l\<^sub>2)\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(addr a\F{D}:=e\<^sub>2, s\<^sub>1)" + let ?y' = "(addr a\F{D}:=Val v::expr,(h\<^sub>2,l\<^sub>2))" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) + also have "(?y', ?z) \ (red P)" using RedFAss hC by simp + ultimately show ?thesis by simp +qed (*>*) lemma FAssRedsNull: - "\ P \ \e\<^sub>1,s\<^sub>0\ \* \null,s\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1\ \* \Val v,s\<^sub>2\ \ \ - P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0\ \* \THROW NullPointer, s\<^sub>2\" -(*<*) -apply(rule rtrancl_trans) - apply(erule FAssReds1) -apply(rule rtrancl_into_rtrancl) - apply(erule FAssReds2) -apply(rule RedFAssNull) -done +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0\ \* \null,s\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1\ \* \Val v,s\<^sub>2\" +shows "P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0\ \* \THROW NullPointer, s\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(null\F{D}:=e\<^sub>2, s\<^sub>1)" + let ?y' = "(null\F{D}:=Val v::expr,s\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) + also have "(?y', ?z) \ (red P)" by(rule RedFAssNull) + ultimately show ?thesis by simp +qed (*>*) lemma FAssRedsThrow1: "P \ \e,s\ \* \throw e',s'\ \ P \ \e\F{D}:=e\<^sub>2, s\ \* \throw e', s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule FAssReds1) -apply(rule red_reds.FAssThrow1) -done -(*>*) +(*<*)by(rule FAssReds1[THEN rtrancl_into_rtrancl, OF _ red_reds.FAssThrow1])(*>*) lemma FAssRedsThrow2: - "\ P \ \e\<^sub>1,s\<^sub>0\ \* \Val v,s\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1\ \* \throw e,s\<^sub>2\ \ - \ P \ \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \* \throw e,s\<^sub>2\" -(*<*) -apply(rule rtrancl_trans) - apply(erule FAssReds1) -apply(rule rtrancl_into_rtrancl) - apply(erule FAssReds2) -apply(rule red_reds.FAssThrow2) -done +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0\ \* \Val v,s\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1\ \* \throw e,s\<^sub>2\" +shows "P \ \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \* \throw e,s\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(Val v\F{D}:=e\<^sub>2,s\<^sub>1)" + let ?y' = "(Val v\F{D}:=throw e,s\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) + also have "(?y', ?z) \ (red P)" by(rule red_reds.FAssThrow2) + ultimately show ?thesis by simp +qed (*>*) subsubsection";;" lemma SeqReds: "P \ \e,s\ \* \e',s'\ \ P \ \e;;e\<^sub>2, s\ \* \e';;e\<^sub>2, s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule SeqRed) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) SeqRed[OF step(2)]]) +qed (*>*) lemma SeqRedsThrow: "P \ \e,s\ \* \throw e',s'\ \ P \ \e;;e\<^sub>2, s\ \* \throw e', s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule SeqReds) -apply(rule red_reds.SeqThrow) -done -(*>*) +(*<*)by(rule SeqReds[THEN rtrancl_into_rtrancl, OF _ red_reds.SeqThrow])(*>*) lemma SeqReds2: - "\ P \ \e\<^sub>1,s\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1\; P \ \e\<^sub>2,s\<^sub>1\ \* \e\<^sub>2',s\<^sub>2\ \ \ P \ \e\<^sub>1;;e\<^sub>2, s\<^sub>0\ \* \e\<^sub>2',s\<^sub>2\" -(*<*) -apply(rule rtrancl_trans) - apply(erule SeqReds) -apply(rule converse_rtrancl_into_rtrancl) - apply(rule RedSeq) -apply assumption -done +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1\ \* \e\<^sub>2',s\<^sub>2\" +shows "P \ \e\<^sub>1;;e\<^sub>2, s\<^sub>0\ \* \e\<^sub>2',s\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(Val v\<^sub>1;; e\<^sub>2,s\<^sub>1)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule SeqReds[OF e\<^sub>1_steps]) + also have "(?y, ?z) \ (red P)\<^sup>*" + by(rule RedSeq[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>2_steps]) + ultimately show ?thesis by simp +qed (*>*) subsubsection"If" lemma CondReds: "P \ \e,s\ \* \e',s'\ \ P \ \if (e) e\<^sub>1 else e\<^sub>2,s\ \* \if (e') e\<^sub>1 else e\<^sub>2,s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule CondRed) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) CondRed[OF step(2)]]) +qed (*>*) lemma CondRedsThrow: "P \ \e,s\ \* \throw a,s'\ \ P \ \if (e) e\<^sub>1 else e\<^sub>2, s\ \* \throw a,s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule CondReds) -apply(rule red_reds.CondThrow) -done -(*>*) +(*<*)by(rule CondReds[THEN rtrancl_into_rtrancl, OF _ red_reds.CondThrow])(*>*) lemma CondReds2T: - "\P \ \e,s\<^sub>0\ \* \true,s\<^sub>1\; P \ \e\<^sub>1, s\<^sub>1\ \* \e',s\<^sub>2\ \ \ P \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\ \* \e',s\<^sub>2\" -(*<*) -apply(rule rtrancl_trans) - apply(erule CondReds) -apply(rule converse_rtrancl_into_rtrancl) - apply(rule RedCondT) -apply assumption -done +assumes e_steps: "P \ \e,s\<^sub>0\ \* \true,s\<^sub>1\" + and e\<^sub>1_steps: "P \ \e\<^sub>1, s\<^sub>1\ \* \e',s\<^sub>2\" +shows "P \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\ \* \e',s\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(if (true) e\<^sub>1 else e\<^sub>2,s\<^sub>1)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF e_steps]) + also have "(?y, ?z) \ (red P)\<^sup>*" + by(rule RedCondT[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>1_steps]) + ultimately show ?thesis by simp +qed (*>*) lemma CondReds2F: - "\P \ \e,s\<^sub>0\ \* \false,s\<^sub>1\; P \ \e\<^sub>2, s\<^sub>1\ \* \e',s\<^sub>2\ \ \ P \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\ \* \e',s\<^sub>2\" -(*<*) -apply(rule rtrancl_trans) - apply(erule CondReds) -apply(rule converse_rtrancl_into_rtrancl) - apply(rule RedCondF) -apply assumption -done +assumes e_steps: "P \ \e,s\<^sub>0\ \* \false,s\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2, s\<^sub>1\ \* \e',s\<^sub>2\" +shows "P \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\ \* \e',s\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(if (false) e\<^sub>1 else e\<^sub>2,s\<^sub>1)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF e_steps]) + also have "(?y, ?z) \ (red P)\<^sup>*" + by(rule RedCondF[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>2_steps]) + ultimately show ?thesis by simp +qed (*>*) subsubsection "While" lemma WhileFReds: - "P \ \b,s\ \* \false,s'\ \ P \ \while (b) c,s\ \* \unit,s'\" +assumes b_steps: "P \ \b,s\ \* \false,s'\" +shows "P \ \while (b) c,s\ \* \unit,s'\" (*<*) -apply(rule converse_rtrancl_into_rtrancl) - apply(rule RedWhile) -apply(rule rtrancl_into_rtrancl) - apply(erule CondReds) -apply(rule RedCondF) -done +by(rule RedWhile[THEN converse_rtrancl_into_rtrancl, + OF CondReds[THEN rtrancl_into_rtrancl, + OF b_steps RedCondF]]) (*>*) lemma WhileRedsThrow: - "P \ \b,s\ \* \throw e,s'\ \ P \ \while (b) c,s\ \* \throw e,s'\" +assumes b_steps: "P \ \b,s\ \* \throw e,s'\" +shows "P \ \while (b) c,s\ \* \throw e,s'\" (*<*) -apply(rule converse_rtrancl_into_rtrancl) - apply(rule RedWhile) -apply(rule rtrancl_into_rtrancl) - apply(erule CondReds) -apply(rule red_reds.CondThrow) -done +by(rule RedWhile[THEN converse_rtrancl_into_rtrancl, + OF CondReds[THEN rtrancl_into_rtrancl, + OF b_steps red_reds.CondThrow]]) (*>*) lemma WhileTReds: - "\ P \ \b,s\<^sub>0\ \* \true,s\<^sub>1\; P \ \c,s\<^sub>1\ \* \Val v\<^sub>1,s\<^sub>2\; P \ \while (b) c,s\<^sub>2\ \* \e,s\<^sub>3\ \ - \ P \ \while (b) c,s\<^sub>0\ \* \e,s\<^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 +assumes b_steps: "P \ \b,s\<^sub>0\ \* \true,s\<^sub>1\" + and c_steps: "P \ \c,s\<^sub>1\ \* \Val v\<^sub>1,s\<^sub>2\" + and while_steps: "P \ \while (b) c,s\<^sub>2\ \* \e,s\<^sub>3\" +shows "P \ \while (b) c,s\<^sub>0\ \* \e,s\<^sub>3\" +(*<*)(is "(?a, ?c) \ (red P)\<^sup>*") +proof - + let ?b = "(if (b) (c;; while (b) c) else unit,s\<^sub>0)" + let ?y = "(if (true) (c;; while (b) c) else unit,s\<^sub>1)" + let ?b' = "(c;; while (b) c,s\<^sub>1)" + let ?y' = "(Val v\<^sub>1;; while (b) c,s\<^sub>2)" + have "(?a, ?b) \ (red P)\<^sup>*" + using RedWhile[THEN converse_rtrancl_into_rtrancl] by simp + also have "(?b, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF b_steps]) + also have "(?y, ?b') \ (red P)\<^sup>*" + using RedCondT[THEN converse_rtrancl_into_rtrancl] by simp + also have "(?b', ?y') \ (red P)\<^sup>*" by(rule SeqReds[OF c_steps]) + also have "(?y', ?c) \ (red P)\<^sup>*" + by(rule RedSeq[THEN converse_rtrancl_into_rtrancl, OF while_steps]) + ultimately show ?thesis by simp +qed (*>*) lemma WhileTRedsThrow: - "\ P \ \b,s\<^sub>0\ \* \true,s\<^sub>1\; P \ \c,s\<^sub>1\ \* \throw e,s\<^sub>2\ \ - \ P \ \while (b) c,s\<^sub>0\ \* \throw e,s\<^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 +assumes b_steps: "P \ \b,s\<^sub>0\ \* \true,s\<^sub>1\" + and c_steps: "P \ \c,s\<^sub>1\ \* \throw e,s\<^sub>2\" +shows "P \ \while (b) c,s\<^sub>0\ \* \throw e,s\<^sub>2\" +(*<*)(is "(?a, ?c) \ (red P)\<^sup>*") +proof - + let ?b = "(if (b) (c;; while (b) c) else unit,s\<^sub>0)" + let ?y = "(if (true) (c;; while (b) c) else unit,s\<^sub>1)" + let ?b' = "(c;; while (b) c,s\<^sub>1)" + let ?y' = "(throw e;; while (b) c,s\<^sub>2)" + have "(?a, ?b) \ (red P)\<^sup>*" + using RedWhile[THEN converse_rtrancl_into_rtrancl] by simp + also have "(?b, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF b_steps]) + also have "(?y, ?b') \ (red P)\<^sup>*" + using RedCondT[THEN converse_rtrancl_into_rtrancl] by simp + also have "(?b', ?y') \ (red P)\<^sup>*" by(rule SeqReds[OF c_steps]) + also have "(?y', ?c) \ (red P)" by(rule red_reds.SeqThrow) + ultimately show ?thesis by simp +qed (*>*) subsubsection"Throw" lemma ThrowReds: "P \ \e,s\ \* \e',s'\ \ P \ \throw e,s\ \* \throw e',s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule ThrowRed) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) ThrowRed[OF step(2)]]) +qed (*>*) lemma ThrowRedsNull: "P \ \e,s\ \* \null,s'\ \ P \ \throw e,s\ \* \THROW NullPointer,s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule ThrowReds) -apply(rule RedThrowNull) -done -(*>*) +(*<*)by(rule ThrowReds[THEN rtrancl_into_rtrancl, OF _ RedThrowNull])(*>*) lemma ThrowRedsThrow: "P \ \e,s\ \* \throw a,s'\ \ P \ \throw e,s\ \* \throw a,s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule ThrowReds) -apply(rule red_reds.ThrowThrow) -done -(*>*) +(*<*)by(rule ThrowReds[THEN rtrancl_into_rtrancl, OF _ red_reds.ThrowThrow])(*>*) subsubsection "InitBlock" lemma InitBlockReds_aux: "P \ \e,s\ \* \e',s'\ \ \h l h' l' v. s = (h,l(V\v)) \ s' = (h',l') \ P \ \{V:T := Val v; e},(h,l)\ \* \{V:T := Val(the(l' V)); e'},(h',l'(V:=(l V)))\" (*<*) -apply(erule converse_rtrancl_induct2) - apply(fastforce simp: fun_upd_same simp del:fun_upd_apply) -apply clarify -apply(rename_tac e0 X Y e1 h1 l1 h0 l0 h2 l2 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 +proof(induct rule: converse_rtrancl_induct2) + case refl then show ?case + by(fastforce simp: fun_upd_same simp del:fun_upd_apply) +next + case (step e0 s0 e1 s1) + obtain h1 l1 where s1[simp]: "s1 = (h1, l1)" by(cases s1) simp + { fix h0 l0 h2 l2 v0 + assume [simp]: "s0 = (h0, l0(V \ v0))" and s'[simp]: "s' = (h2, l2)" + then have "V \ dom l1" using step(1) by(auto dest!: red_lcl_incr) + then obtain v1 where l1V[simp]: "l1 V = \v1\" by blast + + let ?a = "({V:T; V:=Val v0;; e0},(h0, l0))" + let ?b = "({V:T; V:=Val v1;; e1},(h1, l1(V := l0 V)))" + let ?c = "({V:T; V:=Val (the (l2 V));; e'},(h2, l2(V := l0 V)))" + let ?l = "l1(V := l0 V)" and ?v = v1 + + have e0_steps: "P \ \e0,(h0, l0(V \ v0))\ \ \e1,(h1, l1)\" + using step(1) by simp + + have lv: "\l v. l1 = l(V \ v) \ + P \ \{V:T; V:=Val v;; e1},(h1, l)\ \* + \{V:T; V:=Val (the (l2 V));; e'},(h2, l2(V := l V))\" + using step(3) s' s1 by blast + moreover have "l1 = ?l(V \ ?v)" by(rule ext) (simp add:fun_upd_def) + ultimately have "(?b, ?c) \ (red P)\<^sup>*" using lv[of ?l ?v] by simp + then have "(?a, ?c) \ (red P)\<^sup>*" + by(rule InitBlockRed[THEN converse_rtrancl_into_rtrancl, OF e0_steps l1V]) + } + then show ?case by blast +qed (*>*) lemma InitBlockReds: "P \ \e, (h,l(V\v))\ \* \e', (h',l')\ \ P \ \{V:T := Val v; e}, (h,l)\ \* \{V:T := Val(the(l' V)); e'}, (h',l'(V:=(l V)))\" (*<*)by(blast dest:InitBlockReds_aux)(*>*) lemma InitBlockRedsFinal: "\ P \ \e,(h,l(V\v))\ \* \e',(h',l')\; final e' \ \ P \ \{V:T := Val v; e},(h,l)\ \* \e',(h', l'(V := l V))\" (*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule InitBlockReds) -apply(fast elim!:finalE intro:RedInitBlock InitBlockThrow) -done +by(fast elim!:InitBlockReds[THEN rtrancl_into_rtrancl] finalE + intro:RedInitBlock InitBlockThrow) (*>*) subsubsection "Block" lemma BlockRedsFinal: assumes reds: "P \ \e\<^sub>0,s\<^sub>0\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2)\" and fin: "final e\<^sub>2" shows "\h\<^sub>0 l\<^sub>0. s\<^sub>0 = (h\<^sub>0,l\<^sub>0(V:=None)) \ P \ \{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0)\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2(V:=l\<^sub>0 V))\" (*<*) using reds proof (induct rule:converse_rtrancl_induct2) 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 e\<^sub>1 s\<^sub>1) have red: "P \ \e\<^sub>0,s\<^sub>0\ \ \e\<^sub>1,s\<^sub>1\" and reds: "P \ \e\<^sub>1,s\<^sub>1\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2)\" and IH: "\h l. s\<^sub>1 = (h,l(V := None)) \ P \ \{V:T; e\<^sub>1},(h,l)\ \* \e\<^sub>2,(h\<^sub>2, l\<^sub>2(V := l V))\" and s\<^sub>0: "s\<^sub>0 = (h\<^sub>0, l\<^sub>0(V := None))" by fact+ obtain h\<^sub>1 l\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1)" by fastforce 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))" by auto from e\<^sub>1 fin have "e\<^sub>1 \ e\<^sub>2" by (auto simp:final_def) then obtain e' s' where red1: "P \ \e\<^sub>1,s\<^sub>1\ \ \e',s'\" and reds': "P \ \e',s'\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2)\" using converse_rtranclE2[OF reds] by blast from red1 e\<^sub>1 have es': "e' = e" "s' = s\<^sub>1" by auto show ?case using e\<^sub>0 s\<^sub>1 es' reds' by(fastforce 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)\ \ \{V:T; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V))\" 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))\ \* \e\<^sub>2,(h\<^sub>2, l\<^sub>2(V := l\<^sub>0 V))\" 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)\ \ \{V:T := Val v; e\<^sub>1},(h\<^sub>1,l\<^sub>1(V := l\<^sub>0 V))\" 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))\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2(V:=l\<^sub>0 V))\" 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\ \* \e',s'\ \ P \ \try e catch(C V) e\<^sub>2,s\ \* \try e' catch(C V) e\<^sub>2,s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule TryRed) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) TryRed[OF step(2)]]) +qed (*>*) lemma TryRedsVal: "P \ \e,s\ \* \Val v,s'\ \ P \ \try e catch(C V) e\<^sub>2,s\ \* \Val v,s'\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule TryReds) -apply(rule RedTry) -done -(*>*) +(*<*)by(rule TryReds[THEN rtrancl_into_rtrancl, OF _ RedTry])(*>*) lemma TryCatchRedsFinal: - "\ P \ \e\<^sub>1,s\<^sub>0\ \* \Throw a,(h\<^sub>1,l\<^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))\ \* \e\<^sub>2', (h\<^sub>2,l\<^sub>2)\; final e\<^sub>2' \ - \ P \ \try e\<^sub>1 catch(C V) e\<^sub>2, s\<^sub>0\ \* \e\<^sub>2', (h\<^sub>2, l\<^sub>2(V := l\<^sub>1 V))\" -(*<*) -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 +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0\ \* \Throw a,(h\<^sub>1,l\<^sub>1)\" + and h\<^sub>1a: "h\<^sub>1 a = Some(D,fs)" and sub: "P \ D \\<^sup>* C" + and e\<^sub>2_steps: "P \ \e\<^sub>2, (h\<^sub>1, l\<^sub>1(V \ Addr a))\ \* \e\<^sub>2', (h\<^sub>2,l\<^sub>2)\" + and final: "final e\<^sub>2'" +shows "P \ \try e\<^sub>1 catch(C V) e\<^sub>2, s\<^sub>0\ \* \e\<^sub>2', (h\<^sub>2, (l\<^sub>2::locals)(V := l\<^sub>1 V))\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(try Throw a catch(C V) e\<^sub>2,(h\<^sub>1, l\<^sub>1))" + let ?b = "({V:Class C; V:=addr a;; e\<^sub>2},(h\<^sub>1, l\<^sub>1))" + have bz: "(?b, ?z) \ (red P)\<^sup>*" + by(rule InitBlockRedsFinal[OF e\<^sub>2_steps final]) + have hp: "hp (h\<^sub>1, l\<^sub>1) a = \(D, fs)\" using h\<^sub>1a by simp + have "(?x, ?y) \ (red P)\<^sup>*" by(rule TryReds[OF e\<^sub>1_steps]) + also have "(?y, ?z) \ (red P)\<^sup>*" + by(rule RedTryCatch[THEN converse_rtrancl_into_rtrancl, OF hp sub bz]) + ultimately show ?thesis by simp +qed (*>*) lemma TryRedsFail: "\ P \ \e\<^sub>1,s\ \* \Throw a,(h,l)\; h a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \ \try e\<^sub>1 catch(C V) e\<^sub>2,s\ \* \Throw a,(h,l)\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule TryReds) -apply(fastforce intro!: RedTryFail) -done -(*>*) +(*<*)by(fastforce intro!: TryReds[THEN rtrancl_into_rtrancl, OF _ RedTryFail])(*>*) subsubsection "List" lemma ListReds1: "P \ \e,s\ \* \e',s'\ \ P \ \e#es,s\ [\]* \e' # es,s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule ListRed1) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) ListRed1[OF step(2)]]) +qed (*>*) lemma ListReds2: "P \ \es,s\ [\]* \es',s'\ \ P \ \Val v # es,s\ [\]* \Val v # es',s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule ListRed2) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) ListRed2[OF step(2)]]) +qed (*>*) lemma ListRedsVal: "\ P \ \e,s\<^sub>0\ \* \Val v,s\<^sub>1\; P \ \es,s\<^sub>1\ [\]* \es',s\<^sub>2\ \ \ P \ \e#es,s\<^sub>0\ [\]* \Val v # es',s\<^sub>2\" -(*<*) -apply(rule rtrancl_trans) - apply(erule ListReds1) -apply(erule ListReds2) -done -(*>*) +(*<*)by(rule rtrancl_trans[OF ListReds1 ListReds2])(*>*) 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)\ \ \e',(h',l')\ \ fv e' \ fv e" and "P \ \es,(h,l)\ [\] \es',(h',l')\ \ fvs es' \ fvs es" (*<*) proof (induct rule:red_reds_inducts) case (RedCall h l a C fs M Ts T pns body D vs) 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 qed auto (*>*) lemma Red_dom_lcl: "P \ \e,(h,l)\ \ \e',(h',l')\ \ dom l' \ dom l \ fv e" and "P \ \es,(h,l)\ [\] \es',(h',l')\ \ 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)\ \* \e',(h',l')\ \ \ dom l' \ dom l \ fv e" +assumes wf: "wwf_J_prog P" +shows "P \ \e,(h,l)\ \* \e',(h',l')\ \ dom l' \ dom l \ fv e" (*<*) -apply(erule converse_rtrancl_induct_red) - apply blast -apply(blast dest: Red_fv Red_dom_lcl) -done +proof(induct rule: converse_rtrancl_induct_red) + case 1 then show ?case by blast +next + case 2 then show ?case using wf by(blast dest: Red_fv Red_dom_lcl) +qed (*>*) text\Now a few lemmas on the behaviour of blocks during reduction.\ (* If you want to avoid the premise "distinct" further down \ consts upd_vals :: "locals \ vname list \ val list \ val list" primrec "upd_vals l [] vs = []" "upd_vals l (V#Vs) vs = (if V \ set Vs then hd vs else the(l V)) # upd_vals l Vs (tl vs)" lemma [simp]: "\vs. length(upd_vals l Vs vs) = length Vs" by(induct Vs, auto) *) 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 +(*<*)by(rule ext) (simp add:override_on_def) 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))\ \* \e', (h',l')\ \ \ P \ \blocks(Vs,Ts,vs,e), (h,l)\ \* \blocks(Vs,Ts,map (the \ l') Vs,e'), (h',override_on l' l (set Vs))\" (*<*) 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)\ \* \e, (h,l)\" (*<*) 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))\ \* \e', (h',l')\" and fin: "final e'" and l'': "l'' = override_on l' l (set Vs)" shows "P \ \blocks(Vs,Ts,vs,e), (h,l)\ \* \e', (h',l'')\" (*<*) proof - let ?bv = "blocks(Vs,Ts,map (the o l') Vs,e')" have "P \ \blocks(Vs,Ts,vs,e), (h,l)\ \* \?bv, (h',l'')\" using l'' by simp (rule blocksReds[OF wf reds]) also have "P \ \?bv, (h',l'')\ \* \e', (h',l'')\" 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\ \* \e',s'\ \ P \ \e\M(es),s\ \* \e'\M(es),s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule CallObj) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) CallObj[OF step(2)]]) +qed (*>*) lemma CallRedsParams: "P \ \es,s\ [\]* \es',s'\ \ P \ \(Val v)\M(es),s\ \* \(Val v)\M(es'),s'\" (*<*) -apply(erule rtrancl_induct2) - apply blast -apply(erule rtrancl_into_rtrancl) -apply(erule CallParams) -done +proof(induct rule: rtrancl_induct2) + case refl show ?case by blast +next + case step show ?case + by(rule rtrancl_into_rtrancl[OF step(3) CallParams[OF step(2)]]) +qed (*>*) lemma CallRedsFinal: assumes wwf: "wwf_J_prog P" and "P \ \e,s\<^sub>0\ \* \addr a,s\<^sub>1\" "P \ \es,s\<^sub>1\ [\]* \map Val vs,(h\<^sub>2,l\<^sub>2)\" "h\<^sub>2 a = Some(C,fs)" "P \ C sees M: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')\ \* \ef,(h\<^sub>3,l\<^sub>3)\" and "final ef" shows "P \ \e\M(es), s\<^sub>0\ \* \ef,(h\<^sub>3,l\<^sub>2)\" (*<*) 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))\ \* \ef,(h\<^sub>3,l\<^sub>2++l\<^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\ \* \(addr a)\M(es),s\<^sub>1\" by(rule CallRedsObj)(rule assms(2)) also have "P \ \(addr a)\M(es),s\<^sub>1\ \* \(addr a)\M(map Val vs),(h\<^sub>2,l\<^sub>2)\" by(rule CallRedsParams)(rule assms(3)) also have "P \ \(addr a)\M(map Val vs), (h\<^sub>2,l\<^sub>2)\ \ \blocks(this#pns, Class D#Ts, Addr a#vs, body), (h\<^sub>2,l\<^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)\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns))\" by(rule blocksRedsFinal, insert assms wf body', simp_all) finally show ?thesis using eql\<^sub>2 by simp qed (*>*) lemma CallRedsThrowParams: - "\ P \ \e,s0\ \* \Val v,s\<^sub>1\; P \ \es,s\<^sub>1\ [\]* \map Val vs\<^sub>1 @ throw a # es\<^sub>2,s\<^sub>2\ \ - \ P \ \e\M(es),s0\ \* \throw a,s\<^sub>2\" -(*<*) -apply(rule rtrancl_trans) - apply(erule CallRedsObj) -apply(rule rtrancl_into_rtrancl) - apply(erule CallRedsParams) -apply(rule CallThrowParams) -apply simp -done +assumes e_steps: "P \ \e,s\<^sub>0\ \* \Val v,s\<^sub>1\" + and es_steps: "P \ \es,s\<^sub>1\ [\]* \map Val vs\<^sub>1 @ throw a # es\<^sub>2,s\<^sub>2\" +shows "P \ \e\M(es),s\<^sub>0\ \* \throw a,s\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(Val v\M(es),s\<^sub>1)" + let ?y' = "(Val v\M(map Val vs\<^sub>1 @ throw a # es\<^sub>2),s\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps]) + also have "(?y', ?z) \ (red P)\<^sup>*" using CallThrowParams by fast + ultimately show ?thesis by simp +qed (*>*) lemma CallRedsThrowObj: "P \ \e,s0\ \* \throw a,s\<^sub>1\ \ P \ \e\M(es),s0\ \* \throw a,s\<^sub>1\" -(*<*) -apply(rule rtrancl_into_rtrancl) - apply(erule CallRedsObj) -apply(rule CallThrowObj) -done -(*>*) +(*<*)by(rule CallRedsObj[THEN rtrancl_into_rtrancl, OF _ CallThrowObj])(*>*) lemma CallRedsNull: - "\ P \ \e,s\<^sub>0\ \* \null,s\<^sub>1\; P \ \es,s\<^sub>1\ [\]* \map Val vs,s\<^sub>2\ \ - \ P \ \e\M(es),s\<^sub>0\ \* \THROW NullPointer,s\<^sub>2\" -(*<*) -apply(rule rtrancl_trans) - apply(erule CallRedsObj) -apply(rule rtrancl_into_rtrancl) - apply(erule CallRedsParams) -apply(rule RedCallNull) -done +assumes e_steps: "P \ \e,s\<^sub>0\ \* \null,s\<^sub>1\" + and es_steps: "P \ \es,s\<^sub>1\ [\]* \map Val vs,s\<^sub>2\" +shows "P \ \e\M(es),s\<^sub>0\ \* \THROW NullPointer,s\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(null\M(es),s\<^sub>1)" + let ?y' = "(null\M(map Val vs),s\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps]) + also have "(?y', ?z) \ (red P)" by(rule RedCallNull) + ultimately show ?thesis by simp +qed (*>*) subsubsection "The main Theorem" lemma assumes wwf: "wwf_J_prog P" shows big_by_small: "P \ \e,s\ \ \e',s'\ \ P \ \e,s\ \* \e',s'\" and bigs_by_smalls: "P \ \es,s\ [\] \es',s'\ \ P \ \es,s\ [\]* \es',s'\" (*<*) proof (induct rule: eval_evals.inducts) case New thus ?case by (auto simp:RedNew) next case NewFail thus ?case by (auto simp:RedNewFail) next case Cast thus ?case by(fastforce intro:CastRedsAddr) next case CastNull thus ?case by(simp add:CastRedsNull) next case CastFail thus ?case by(fastforce intro!:CastRedsFail) next case CastThrow thus ?case by(auto dest!:eval_final simp:CastRedsThrow) next case Val thus ?case by simp next case BinOp thus ?case by(auto simp:BinOpRedsVal) next case BinOpThrow1 thus ?case by(auto dest!:eval_final simp: BinOpRedsThrow1) next case BinOpThrow2 thus ?case by(auto dest!:eval_final simp: BinOpRedsThrow2) next case Var thus ?case by (auto simp:RedVar) next case LAss thus ?case by(auto simp: LAssRedsVal) next case LAssThrow thus ?case by(auto dest!:eval_final simp: LAssRedsThrow) next case FAcc thus ?case by(auto intro:FAccRedsVal) next case FAccNull thus ?case by(simp add:FAccRedsNull) next case FAccThrow thus ?case by(auto dest!:eval_final simp:FAccRedsThrow) next case FAss thus ?case by(auto simp:FAssRedsVal) next case FAssNull thus ?case by(auto simp:FAssRedsNull) next case FAssThrow1 thus ?case by(auto dest!:eval_final simp:FAssRedsThrow1) next case FAssThrow2 thus ?case by(auto dest!:eval_final simp:FAssRedsThrow2) next case CallObjThrow thus ?case by(auto dest!:eval_final simp:CallRedsThrowObj) next case CallNull thus ?case by(simp add:CallRedsNull) next case CallParamsThrow thus ?case by(auto dest!:evals_final simp:CallRedsThrowParams) next case (Call e s\<^sub>0 a s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 C fs M Ts T pns body D l\<^sub>2' e' h\<^sub>3 l\<^sub>3) have IHe: "P \ \e,s\<^sub>0\ \* \addr a,s\<^sub>1\" and IHes: "P \ \ps,s\<^sub>1\ [\]* \map Val vs,(h\<^sub>2,l\<^sub>2)\" and h\<^sub>2a: "h\<^sub>2 a = Some(C,fs)" and "method": "P \ C sees M:Ts\T = (pns,body) in D" and same_length: "length vs = length pns" and l\<^sub>2': "l\<^sub>2' = [this \ Addr a, pns[\]vs]" and eval_body: "P \ \body,(h\<^sub>2, l\<^sub>2')\ \ \e',(h\<^sub>3, l\<^sub>3)\" and IHbody: "P \ \body,(h\<^sub>2,l\<^sub>2')\ \* \e',(h\<^sub>3,l\<^sub>3)\" by fact+ show "P \ \e\M(ps),s\<^sub>0\ \* \e',(h\<^sub>3, l\<^sub>2)\" using "method" same_length l\<^sub>2' h\<^sub>2a IHbody eval_final[OF eval_body] by(fastforce intro:CallRedsFinal[OF wwf IHe IHes]) next case Block thus ?case by(auto simp: BlockRedsFinal dest:eval_final) next case Seq thus ?case by(auto simp:SeqReds2) next case SeqThrow thus ?case by(auto dest!:eval_final simp: SeqRedsThrow) next case CondT thus ?case by(auto simp:CondReds2T) next case CondF thus ?case by(auto simp:CondReds2F) next case CondThrow thus ?case by(auto dest!:eval_final simp:CondRedsThrow) next case WhileF thus ?case by(auto simp:WhileFReds) next case WhileT thus ?case by(auto simp: WhileTReds) next case WhileCondThrow thus ?case by(auto dest!:eval_final simp: WhileRedsThrow) next case WhileBodyThrow thus ?case by(auto dest!:eval_final simp: WhileTRedsThrow) next case Throw thus ?case by(auto simp:ThrowReds) next case ThrowNull thus ?case by(auto simp:ThrowRedsNull) next case ThrowThrow thus ?case by(auto dest!:eval_final simp:ThrowRedsThrow) next case Try thus ?case by(simp add:TryRedsVal) next case TryCatch thus ?case by(fast intro!: TryCatchRedsFinal dest!:eval_final) next case TryThrow thus ?case by(fastforce intro!:TryRedsFail) next case Nil thus ?case by simp next case Cons thus ?case by(fastforce intro!:Cons_eq_appendI[OF refl refl] ListRedsVal) next case ConsThrow thus ?case by(fastforce elim: ListReds1) qed (*>*) subsection\Big steps simulates small step\ text\This direction was carried out by Norbert Schirmer and Daniel Wasserrab.\ text \The big step equivalent of \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)\ \ \e',(h',l')\ \ \ \ l''. P \ \e,(h,l(ps[\]vs))\ \ \e',(h',l'')\" (*<*) 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)\ \ \e',(h', l')\" by fact with Ts vs have "P \ \{p:T := Val v; blocks (ps', Ts', vs', e)},(h,l)\ \ \e',(h', l')\" by simp then obtain l''' where eval_ps': "P \ \blocks (ps', Ts', vs', e),(h, l(p\v))\ \ \e',(h', l''')\" 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'))\ \ \e',(h', l'')\" using length_eqs Ts vs Cons.hyps [OF _ _ eval_ps'] by auto from hyp show "\l''. P \ \e,(h, l(p # ps'[\]vs))\ \ \e',(h', l'')\" using Ts vs by auto qed (*>*) (* FIXME exercise: show precise relationship between l' and l'': lemma blocksEval: "\ Ts vs l l'. \length ps = length Ts; length ps = length vs; P\ \blocks(ps,Ts,vs,e),(h,l)\ \ \e',(h',l')\ \ \ \ l''. P \ \e,(h,l(ps[\]vs))\ \ \e',(h',l'')\ \ l'=l''(l|set ps)" proof (induct ps) case Nil then show ?case by simp next case (Cons p ps') have length_eqs: "length (p # ps') = length Ts" "length (p # ps') = length vs" . 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)\ \ \e',(h', l')\". with Ts vs have "P \ \{p:T := Val v; blocks (ps', Ts', vs', e)},(h,l)\ \ \e',(h', l')\" by simp then obtain l''' where eval_ps': "P \ \blocks (ps', Ts', vs', e),(h, l(p\v))\ \ \e',(h', l''')\" and l''': "l'=l'''(p:=l p)" by (cases) (auto elim: eval_cases) then obtain l'' where hyp: "P \ \e,(h, l(p\v)(ps'[\]vs'))\ \ \e',(h', l'')\" and l'': "l''' = l''(l(p\v)|set ps')" using length_eqs Ts vs Cons.hyps [OF _ _ eval_ps'] by auto have "l' = l''(l|set (p # ps'))" proof - have "(l''(l(p\v)|set ps'))(p := l p) = l''(l|insert p (set ps'))" by (induct ps') (auto intro: ext simp add: fun_upd_def override_on_def) with l''' l'' show ?thesis by simp qed with hyp show "\l''. P \ \e,(h, l(p # ps'[\]vs))\ \ \e',(h', l'')\ \ l' = l''(l|set (p # ps'))" using Ts vs by auto qed *) lemma assumes wf: "wwf_J_prog P" shows eval_restrict_lcl: "P \ \e,(h,l)\ \ \e',(h',l')\ \ (\W. fv e \ W \ P \ \e,(h,l|`W)\ \ \e',(h',l'|`W)\)" and "P \ \es,(h,l)\ [\] \es',(h',l')\ \ (\W. fvs es \ W \ P \ \es,(h,l|`W)\ [\] \es',(h',l'|`W)\)" (*<*) proof(induct rule:eval_evals_inducts) case (Block e\<^sub>0 h\<^sub>0 l\<^sub>0 V e\<^sub>1 h\<^sub>1 l\<^sub>1 T) have IH: "\W. fv e\<^sub>0 \ W \ P \ \e\<^sub>0,(h\<^sub>0,l\<^sub>0(V:=None)|`W)\ \ \e\<^sub>1,(h\<^sub>1,l\<^sub>1|`W)\" 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))\ \ \e\<^sub>1,(h\<^sub>1, l\<^sub>1|`insert V W)\" 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 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 v h l l' V) have IH: "\W. fv e \ W \ P \ \e,(h\<^sub>0,l\<^sub>0|`W)\ \ \Val v,(h,l|`W)\" 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 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 CallObjThrow thus ?case by simp (blast intro: eval_evals.intros) next case CallNull thus ?case by simp (blast intro: eval_evals.CallNull) next case CallParamsThrow thus ?case by simp (blast intro: eval_evals.CallParamsThrow) next case (Call e h\<^sub>0 l\<^sub>0 a h\<^sub>1 l\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 C fs M Ts T pns body D l\<^sub>2' e' h\<^sub>3 l\<^sub>3) have IHe: "\W. fv e \ W \ P \ \e,(h\<^sub>0,l\<^sub>0|`W)\ \ \addr a,(h\<^sub>1,l\<^sub>1|`W)\" and IHps: "\W. fvs ps \ W \ P \ \ps,(h\<^sub>1,l\<^sub>1|`W)\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2|`W)\" and IHbd: "\W. fv body \ W \ P \ \body,(h\<^sub>2,l\<^sub>2'|`W)\ \ \e',(h\<^sub>3,l\<^sub>3|`W)\" and h\<^sub>2a: "h\<^sub>2 a = Some (C, fs)" and "method": "P \ C sees M: 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 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 a h\<^sub>1 l\<^sub>1 D fs C e\<^sub>2 V e\<^sub>2' h\<^sub>2 l\<^sub>2) have IH\<^sub>1: "\W. fv e\<^sub>1 \ W \ P \ \e\<^sub>1,(h\<^sub>0,l\<^sub>0|`W)\ \ \Throw a,(h\<^sub>1,l\<^sub>1|`W)\" and IH\<^sub>2: "\W. fv e\<^sub>2 \ W \ P \ \e\<^sub>2,(h\<^sub>1,l\<^sub>1(V\Addr a)|`W)\ \ \e\<^sub>2',(h\<^sub>2,l\<^sub>2|`W)\" 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))\ \ \e\<^sub>2',(h\<^sub>2,l\<^sub>2|`insert V W)\" 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) qed (*>*) lemma eval_notfree_unchanged: "P \ \e,(h,l)\ \ \e',(h',l')\ \ (\V. V \ fv e \ l' V = l V)" and "P \ \es,(h,l)\ [\] \es',(h',l')\ \ (\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 qed simp_all (*>*) lemma eval_closed_lcl_unchanged: "\ P \ \e,(h,l)\ \ \e',(h',l')\; 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 (*>*) (* 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] (*>*) (* FIXME exercise 1: define a big step semantics where the body of a procedure can access not juts this and pns but all of the enclosing l. What exactly is fed in? What exactly is returned at the end? Notion: "dynamic binding" excercise 2: the semantics of exercise 1 is closer to the small step semantics. Reformulate equivalence proof by modifying call lemmas. *) text \The key lemma:\ lemma assumes wf: "wwf_J_prog P" shows extend_1_eval: "P \ \e,s\ \ \e'',s''\ \ (\s' e'. P \ \e'',s''\ \ \e',s'\ \ P \ \e,s\ \ \e',s'\)" and extend_1_evals: "P \ \es,t\ [\] \es'',t''\ \ (\t' es'. P \ \es'',t''\ [\] \es',t'\ \ P \ \es,t\ [\] \es',t'\)" (*<*) proof (induct rule: red_reds.inducts) case (RedCall s a C fs M Ts T pns body D vs s' e') have "P \ \addr a,s\ \ \addr a,s\" by (rule eval_evals.intros) moreover have finals: "finals(map Val vs)" by simp obtain h\<^sub>2 l\<^sub>2 where s: "s = (h\<^sub>2,l\<^sub>2)" by (cases s) with finals have "P \ \map Val vs,s\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2)\" by (iprover intro: eval_finalsId) moreover from s have "h\<^sub>2 a = Some (C, fs)" using RedCall by simp moreover have "method": "P \ C sees M: 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 where s': "s' = (h\<^sub>3,l\<^sub>3)" by (cases s') have eval_blocks: "P \ \blocks (this # pns, Class D # Ts, Addr a # vs, body),s\ \ \e',s'\" by fact hence id: "l\<^sub>3 = l\<^sub>2" using fv s 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\<^sub>2,l\<^sub>2')\ \ \e',(h\<^sub>3,l\<^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 "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\<^sub>2,l\<^sub>2)\ \\e',(h\<^sub>3,l\<^sub>3)\" using s s' by simp ultimately obtain l'' where "P \ \body,(h\<^sub>2,l\<^sub>2(this # pns[\]Addr a # vs))\\\e',(h\<^sub>3, l'')\" by (blast dest:blocksEval) from eval_restrict_lcl[OF wf this fv] this_distinct same_len\<^sub>1 same_len have "P \ \body,(h\<^sub>2,[this # pns[\]Addr a # vs])\ \ \e',(h\<^sub>3, l''|`(set(this#pns)))\" 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),s\ \ \e',(h\<^sub>3,l\<^sub>2)\" by (rule Call) with s' id show ?case by simp next case RedNew thus ?case by (iprover elim: eval_cases intro: eval_evals.intros) next case RedNewFail thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (CastRed e s e'' s'' C s' e') thus ?case by(cases s, cases s') (erule eval_cases, auto intro: eval_evals.intros) next case RedCastNull thus ?case by (iprover elim: eval_cases intro: eval_evals.intros) next case (RedCast s a D fs C s'' e'') thus ?case by (cases s) (auto elim: eval_cases intro: eval_evals.intros) next case (RedCastFail s a D fs C s'' e'') thus ?case by (cases s) (auto elim!: eval_cases intro: eval_evals.intros) next case (BinOpRed1 e s e' s' bop e\<^sub>2 s'' e') thus ?case by (cases s'')(erule eval_cases,auto intro: eval_evals.intros) next case BinOpRed2 thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId) next case RedBinOp thus ?case by (iprover elim: eval_cases intro: eval_evals.intros) next case (RedVar s V v s' e') thus ?case by (cases s)(fastforce elim: eval_cases intro: eval_evals.intros) next case (LAssRed e s e' s' V s'') thus ?case by (cases s'')(erule eval_cases,auto intro: eval_evals.intros) next case RedLAss thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (FAccRed e s e' s' F D s'') thus ?case by (cases s'')(erule eval_cases,auto intro: eval_evals.intros) next case (RedFAcc s a C fs F D v s' e') thus ?case by (cases s)(fastforce elim: eval_cases intro: eval_evals.intros) next case RedFAccNull thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros) next case (FAssRed1 e s e' s'' F D e\<^sub>2 s' e') thus ?case by (cases s')(erule eval_cases, auto intro: eval_evals.intros) next case (FAssRed2 e s e' s'' v F D s' e') thus ?case by (cases s) (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId) 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 CallObj thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros) next case CallParams thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId) next case RedCallNull thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros eval_finalsId) next case InitBlockRed thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId simp add: map_upd_triv fun_upd_same) next case (RedInitBlock V T v u s s' e') have "P \ \Val u,s\ \ \e',s'\" by fact then obtain s': "s'=s" and e': "e'=Val u" by cases simp obtain h l where s: "s=(h,l)" by (cases s) have "P \ \{V:T :=Val v; Val u},(h,l)\ \ \Val u,(h, (l(V\v))(V:=l V))\" by (fastforce intro!: eval_evals.intros) thus "P \ \{V:T := Val v; Val u},s\ \ \e',s'\" using s s' e' 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 s' e') have "P \ \Val v,s\ \ \e',s'\" by fact then obtain s': "s'=s" and e': "e'=Val v" by cases simp obtain h l where s: "s=(h,l)" by (cases s) have "P \ \Val v,(h,l(V:=None))\ \ \Val v,(h,l(V:=None))\" by (rule eval_evals.intros) hence "P \ \{V:T;Val v},(h,l)\ \ \Val v,(h,(l(V:=None))(V:=l V))\" by (rule eval_evals.Block) thus "P \ \{V:T; Val v},s\ \ \e',s'\" using s s' e' by simp next case SeqRed thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) 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) 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 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedThrowNull thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (TryRed e s e' s' C V e\<^sub>2 s'' e') thus ?case by (cases s, cases s'', auto elim!: eval_cases intro: eval_evals.intros) next case RedTry 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 s' e') 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) next case ListRed2 thus ?case by (fastforce elim!: evals_cases eval_cases intro: eval_evals.intros eval_finalId) next case CastThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case BinOpThrow1 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case BinOpThrow2 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case LAssThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case FAccThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case FAssThrow1 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 CallThrowObj thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (CallThrowParams es vs e es' v M s s' e') have "P \ \Val v,s\ \ \Val v,s\" by (rule eval_evals.intros) moreover have es: "es = map Val vs @ throw e # es'" by fact have eval_e: "P \ \throw e,s\ \ \e',s'\" by fact then obtain xa where e': "e' = Throw xa" by (cases) (auto dest!: eval_final) with list_eval_Throw [OF eval_e] es have "P \ \es,s\ [\] \map Val vs @ Throw xa # es',s'\" by simp ultimately have "P \ \Val v\M(es),s\ \ \Throw xa,s'\" by (rule eval_evals.CallParamsThrow) thus ?case using e' by simp next case (InitBlockThrow V T v a s s' e') have "P \ \Throw a,s\ \ \e',s'\" by fact then obtain s': "s' = s" and e': "e' = Throw a" by cases (auto elim!:eval_cases) obtain h l where s: "s = (h,l)" by (cases s) have "P \ \{V:T :=Val v; Throw a},(h,l)\ \ \Throw a, (h, (l(V\v))(V:=l V))\" by(fastforce intro:eval_evals.intros) thus "P \ \{V:T := Val v; Throw a},s\ \ \e',s'\" using s s' e' by simp next case (BlockThrow V T a s s' e') have "P \ \Throw a, s\ \ \e',s'\" by fact then obtain s': "s' = s" and e': "e' = Throw a" by cases (auto elim!:eval_cases) obtain h l where s: "s=(h,l)" by (cases s) have "P \ \Throw a, (h,l(V:=None))\ \ \Throw a, (h,l(V:=None))\" by (simp add:eval_evals.intros eval_finalId) hence "P\\{V:T;Throw a},(h,l)\\\Throw a, (h,(l(V:=None))(V:=l V))\" by (rule eval_evals.Block) thus "P \ \{V:T; Throw a},s\ \ \e',s'\" using s s' e' by simp next case SeqThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case CondThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case ThrowThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) qed (*>*) (*<*) (* ... und wieder anschalten: *) declare split_paired_All [simp] split_paired_Ex [simp] (*>*) text \Its extension to \\*\:\ lemma extend_eval: assumes wf: "wwf_J_prog P" and reds: "P \ \e,s\ \* \e'',s''\" and eval_rest: "P \ \e'',s''\ \ \e',s'\" shows "P \ \e,s\ \ \e',s'\" (*<*) -using reds eval_rest -apply (induct rule: converse_rtrancl_induct2) -apply simp -apply simp -apply (rule extend_1_eval) -apply (rule wf) -apply assumption -apply assumption -done +using reds eval_rest +proof (induct rule: converse_rtrancl_induct2) + case refl then show ?case by simp +next + case step + show ?case using step 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\ [\]* \es'',s''\" and eval_rest: "P \ \es'',s''\ [\] \es',s'\" shows "P \ \es,s\ [\] \es',s'\" (*<*) -using reds eval_rest -apply (induct rule: converse_rtrancl_induct2) -apply simp -apply simp -apply (rule extend_1_evals) -apply (rule wf) -apply assumption -apply assumption -done +using reds eval_rest +proof (induct rule: converse_rtrancl_induct2) + case refl then show ?case by simp +next + case step + show ?case using step 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\ \* \e',s'\; final e'\ \ P \ \e,s\ \ \e',s'\" and "\P \ \es,s\ [\]* \es',s'\; finals es'\ \ P \ \es,s\ [\] \es',s'\" (*<*) proof - note wf moreover assume "P \ \e,s\ \* \e',s'\" moreover assume "final e'" then have "P \ \e',s'\ \ \e',s'\" by (rule eval_finalId) ultimately show "P \ \e,s\\\e',s'\" by (rule extend_eval) next note wf moreover assume "P \ \es,s\ [\]* \es',s'\" moreover assume "finals es'" then have "P \ \es',s'\ [\] \es',s'\" by (rule eval_finalsId) 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 \ P \ \e,s\ \ \e',s'\ = (P \ \e,s\ \* \e',s'\ \ final e')" (*<*)by(blast dest: big_by_small eval_final small_by_big)(*>*) end diff --git a/thys/Jinja/J/Expr.thy b/thys/Jinja/J/Expr.thy --- a/thys/Jinja/J/Expr.thy +++ b/thys/Jinja/J/Expr.thy @@ -1,104 +1,98 @@ (* Title: Jinja/J/Expr.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Expressions\ theory Expr imports "../Common/Exceptions" begin datatype bop = Eq | Add \ \names of binary operations\ datatype 'a exp = new cname \ \class instance creation\ | Cast cname "('a exp)" \ \type cast\ | Val val \ \value\ | BinOp "('a exp)" bop "('a exp)" ("_ \_\ _" [80,0,81] 80) \ \binary operation\ | Var 'a \ \local variable (incl. parameter)\ | LAss 'a "('a exp)" ("_:=_" [90,90]90) \ \local assignment\ | FAcc "('a exp)" vname cname ("_\_{_}" [10,90,99]90) \ \field access\ | FAss "('a exp)" vname cname "('a exp)" ("_\_{_} := _" [10,90,99,90]90) \ \field assignment\ | Call "('a exp)" mname "('a exp list)" ("_\_'(_')" [90,99,0] 90) \ \method call\ | Block 'a ty "('a exp)" ("'{_:_; _}") | Seq "('a exp)" "('a exp)" ("_;;/ _" [61,60]60) | Cond "('a exp)" "('a exp)" "('a exp)" ("if '(_') _/ else _" [80,79,79]70) | While "('a exp)" "('a exp)" ("while '(_') _" [80,79]70) | throw "('a exp)" | TryCatch "('a exp)" cname 'a "('a exp)" ("try _/ catch'(_ _') _" [0,99,80,79] 70) type_synonym expr = "vname exp" \ \Jinja expression\ type_synonym J_mb = "vname list \ expr" \ \Jinja method body: parameter names and expression\ type_synonym J_prog = "J_mb prog" \ \Jinja program\ text\The semantics of binary operators:\ fun binop :: "bop \ val \ val \ val option" where "binop(Eq,v\<^sub>1,v\<^sub>2) = Some(Bool (v\<^sub>1 = v\<^sub>2))" | "binop(Add,Intg i\<^sub>1,Intg i\<^sub>2) = Some(Intg(i\<^sub>1+i\<^sub>2))" | "binop(bop,v\<^sub>1,v\<^sub>2) = None" lemma [simp]: "(binop(Add,v\<^sub>1,v\<^sub>2) = Some v) = (\i\<^sub>1 i\<^sub>2. v\<^sub>1 = Intg i\<^sub>1 \ v\<^sub>2 = Intg i\<^sub>2 \ v = Intg(i\<^sub>1+i\<^sub>2))" -(*<*) -apply(cases v\<^sub>1) -apply auto -apply(cases v\<^sub>2) -apply auto -done -(*>*) +(*<*)by(cases v\<^sub>1; cases v\<^sub>2) auto(*>*) subsection "Syntactic sugar" abbreviation (input) InitBlock:: "'a \ ty \ 'a exp \ 'a exp \ 'a exp" ("(1'{_:_ := _;/ _})") where "InitBlock V T e1 e2 == {V:T; V := e1;; e2}" abbreviation unit where "unit == Val Unit" abbreviation null where "null == Val Null" abbreviation "addr a == Val(Addr a)" abbreviation "true == Val(Bool True)" abbreviation "false == Val(Bool False)" abbreviation Throw :: "addr \ 'a exp" where "Throw a == throw(Val(Addr a))" abbreviation THROW :: "cname \ 'a exp" where "THROW xc == Throw(addr_of_sys_xcpt xc)" subsection\Free Variables\ primrec fv :: "expr \ vname set" and fvs :: "expr list \ vname set" where "fv(new C) = {}" | "fv(Cast C e) = fv e" | "fv(Val v) = {}" | "fv(e\<^sub>1 \bop\ e\<^sub>2) = fv e\<^sub>1 \ fv e\<^sub>2" | "fv(Var V) = {V}" | "fv(LAss V e) = {V} \ fv e" | "fv(e\F{D}) = fv e" | "fv(e\<^sub>1\F{D}:=e\<^sub>2) = fv e\<^sub>1 \ fv e\<^sub>2" | "fv(e\M(es)) = fv e \ fvs es" | "fv({V:T; e}) = fv e - {V}" | "fv(e\<^sub>1;;e\<^sub>2) = fv e\<^sub>1 \ fv e\<^sub>2" | "fv(if (b) e\<^sub>1 else e\<^sub>2) = fv b \ fv e\<^sub>1 \ fv e\<^sub>2" | "fv(while (b) e) = fv b \ fv e" | "fv(throw e) = fv e" | "fv(try e\<^sub>1 catch(C V) e\<^sub>2) = fv e\<^sub>1 \ (fv e\<^sub>2 - {V})" | "fvs([]) = {}" | "fvs(e#es) = fv e \ fvs es" lemma [simp]: "fvs(es\<^sub>1 @ es\<^sub>2) = fvs es\<^sub>1 \ fvs es\<^sub>2" (*<*)by (induct es\<^sub>1 type:list) auto(*>*) lemma [simp]: "fvs(map Val vs) = {}" (*<*)by (induct vs) auto(*>*) end diff --git a/thys/Jinja/J/JWellForm.thy b/thys/Jinja/J/JWellForm.thy --- a/thys/Jinja/J/JWellForm.thy +++ b/thys/Jinja/J/JWellForm.thy @@ -1,64 +1,53 @@ (* Title: Jinja/J/JWellForm.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Well-formedness Constraints\ theory JWellForm imports "../Common/WellForm" WWellForm WellType DefAss begin definition wf_J_mdecl :: "J_prog \ cname \ J_mb mdecl \ bool" where "wf_J_mdecl P C \ \(M,Ts,T,(pns,body)). length Ts = length pns \ distinct pns \ this \ set pns \ (\T'. P,[this\Class C,pns[\]Ts] \ body :: T' \ P \ T' \ T) \ \ body \{this} \ set pns\" lemma wf_J_mdecl[simp]: "wf_J_mdecl P C (M,Ts,T,pns,body) \ (length Ts = length pns \ distinct pns \ this \ set pns \ (\T'. P,[this\Class C,pns[\]Ts] \ body :: T' \ P \ T' \ T) \ \ body \{this} \ set pns\)" (*<*)by(simp add:wf_J_mdecl_def)(*>*) abbreviation wf_J_prog :: "J_prog \ bool" where "wf_J_prog == wf_prog wf_J_mdecl" lemma wf_J_prog_wf_J_mdecl: "\ wf_J_prog P; (C, D, fds, mths) \ set P; jmdcl \ set mths \ \ wf_J_mdecl P C jmdcl" -(*<*) -apply (simp add: wf_prog_def) -apply (simp add: wf_cdecl_def) -apply (erule conjE)+ -apply (drule bspec, assumption) -apply simp -apply (erule conjE)+ -apply (drule bspec, assumption) -apply (simp add: wf_mdecl_def split_beta) -done -(*>*) +(*<*)by(fastforce simp: wf_prog_def wf_cdecl_def wf_mdecl_def)(*>*) lemma wf_mdecl_wwf_mdecl: "wf_J_mdecl P C Md \ wwf_J_mdecl P C Md" (*<*)by(fastforce simp:wwf_J_mdecl_def dest!:WT_fv)(*>*) lemma wf_prog_wwf_prog: "wf_J_prog P \ wwf_J_prog P" (*<*) -apply(simp add:wf_prog_def wf_cdecl_def wf_mdecl_def) -apply(fast intro:wf_mdecl_wwf_mdecl) -done +by (simp add:wf_prog_def wf_cdecl_def wf_mdecl_def) + (fast intro:wf_mdecl_wwf_mdecl) (*>*) end diff --git a/thys/Jinja/J/Progress.thy b/thys/Jinja/J/Progress.thy --- a/thys/Jinja/J/Progress.thy +++ b/thys/Jinja/J/Progress.thy @@ -1,626 +1,619 @@ (* Title: Jinja/J/SmallProgress.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Progress of Small Step Semantics\ theory Progress imports Equivalence WellTypeRT DefAss "../Common/Conform" begin lemma final_addrE: "\ P,E,h \ e : Class C; final e; \a. e = addr a \ R; \a. e = Throw a \ R \ \ R" (*<*)by(auto simp:final_def)(*>*) lemma finalRefE: "\ P,E,h \ e : T; is_refT T; final e; e = null \ R; \a C. \ e = addr a; T = Class C \ \ R; \a. e = Throw a \ R \ \ R" (*<*)by(auto simp:final_def is_refT_def)(*>*) text\Derivation of new induction scheme for well typing:\ inductive WTrt' :: "[J_prog,heap,env,expr,ty] \ bool" and WTrts' :: "[J_prog,heap,env,expr list, ty list] \ bool" and WTrt2' :: "[J_prog,env,heap,expr,ty] \ bool" ("_,_,_ \ _ :'' _" [51,51,51]50) and WTrts2' :: "[J_prog,env,heap,expr list, ty list] \ bool" ("_,_,_ \ _ [:''] _" [51,51,51]50) for P :: J_prog and h :: heap where "P,E,h \ e :' T \ WTrt' P h E e T" | "P,E,h \ es [:'] Ts \ WTrts' P h E es Ts" | "is_class P C \ P,E,h \ new C :' Class C" | "\ P,E,h \ e :' T; is_refT T; is_class P C \ \ P,E,h \ Cast C e :' Class C" | "typeof\<^bsub>h\<^esub> v = Some T \ P,E,h \ Val v :' T" | "E v = Some T \ P,E,h \ Var v :' T" | "\ P,E,h \ e\<^sub>1 :' T\<^sub>1; P,E,h \ e\<^sub>2 :' T\<^sub>2 \ \ P,E,h \ e\<^sub>1 \Eq\ e\<^sub>2 :' Boolean" | "\ P,E,h \ e\<^sub>1 :' Integer; P,E,h \ e\<^sub>2 :' Integer \ \ P,E,h \ e\<^sub>1 \Add\ e\<^sub>2 :' Integer" | "\ P,E,h \ Var V :' T; P,E,h \ e :' T'; P \ T' \ T \<^cancel>\V \ This\ \ \ P,E,h \ V:=e :' Void" | "\ P,E,h \ e :' Class C; P \ C has F:T in D \ \ P,E,h \ e\F{D} :' T" | "P,E,h \ e :' NT \ P,E,h \ e\F{D} :' T" | "\ P,E,h \ e\<^sub>1 :' Class C; P \ C has F:T in D; P,E,h \ e\<^sub>2 :' T\<^sub>2; P \ T\<^sub>2 \ T \ \ P,E,h \ e\<^sub>1\F{D}:=e\<^sub>2 :' Void" | "\ P,E,h \ e\<^sub>1:'NT; P,E,h \ e\<^sub>2 :' T\<^sub>2 \ \ P,E,h \ e\<^sub>1\F{D}:=e\<^sub>2 :' Void" | "\ P,E,h \ e :' Class C; P \ C sees M:Ts \ T = (pns,body) in D; P,E,h \ es [:'] Ts'; P \ Ts' [\] Ts \ \ P,E,h \ e\M(es) :' T" | "\ P,E,h \ e :' NT; P,E,h \ es [:'] Ts \ \ P,E,h \ e\M(es) :' T" | "P,E,h \ [] [:'] []" | "\ P,E,h \ e :' T; P,E,h \ es [:'] Ts \ \ P,E,h \ e#es [:'] T#Ts" | "\ typeof\<^bsub>h\<^esub> v = Some T\<^sub>1; P \ T\<^sub>1 \ T; P,E(V\T),h \ e\<^sub>2 :' T\<^sub>2 \ \ P,E,h \ {V:T := Val v; e\<^sub>2} :' T\<^sub>2" | "\ P,E(V\T),h \ e :' T'; \ assigned V e \ \ P,E,h \ {V:T; e} :' T'" | "\ P,E,h \ e\<^sub>1:' T\<^sub>1; P,E,h \ e\<^sub>2:'T\<^sub>2 \ \ P,E,h \ e\<^sub>1;;e\<^sub>2 :' T\<^sub>2" | "\ P,E,h \ e :' Boolean; P,E,h \ e\<^sub>1:' T\<^sub>1; P,E,h \ e\<^sub>2:' T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1; P \ T\<^sub>1 \ T\<^sub>2 \ T = T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ T = T\<^sub>1 \ \ P,E,h \ if (e) e\<^sub>1 else e\<^sub>2 :' T" (* "\ P,E,h \ e :' Boolean; P,E,h \ e\<^sub>1:' T\<^sub>1; P,E,h \ e\<^sub>2:' T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ \ P,E,h \ if (e) e\<^sub>1 else e\<^sub>2 :' T\<^sub>2" "\ P,E,h \ e :' Boolean; P,E,h \ e\<^sub>1:' T\<^sub>1; P,E,h \ e\<^sub>2:' T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ \ P,E,h \ if (e) e\<^sub>1 else e\<^sub>2 :' T\<^sub>1" *) | "\ P,E,h \ e :' Boolean; P,E,h \ c:' T \ \ P,E,h \ while(e) c :' Void" | "\ P,E,h \ e :' T\<^sub>r; is_refT T\<^sub>r \ \ P,E,h \ throw e :' T" | "\ P,E,h \ e\<^sub>1 :' T\<^sub>1; P,E(V \ Class C),h \ e\<^sub>2 :' T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ \ P,E,h \ try e\<^sub>1 catch(C V) e\<^sub>2 :' T\<^sub>2" (*<*) lemmas WTrt'_induct = WTrt'_WTrts'.induct [split_format (complete)] and WTrt'_inducts = WTrt'_WTrts'.inducts [split_format (complete)] inductive_cases WTrt'_elim_cases[elim!]: "P,E,h \ V :=e :' T" (*>*) lemma [iff]: "P,E,h \ e\<^sub>1;;e\<^sub>2 :' T\<^sub>2 = (\T\<^sub>1. P,E,h \ e\<^sub>1:' T\<^sub>1 \ P,E,h \ e\<^sub>2:' T\<^sub>2)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros) -done -(*>*) +(*<*)by(rule iffI) (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros)(*>*) lemma [iff]: "P,E,h \ Val v :' T = (typeof\<^bsub>h\<^esub> v = Some T)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros) -done -(*>*) +(*<*)by(rule iffI) (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros)(*>*) lemma [iff]: "P,E,h \ Var v :' T = (E v = Some T)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros) -done -(*>*) +(*<*)by(rule iffI) (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros)(*>*) lemma wt_wt': "P,E,h \ e : T \ P,E,h \ e :' T" and wts_wts': "P,E,h \ es [:] Ts \ P,E,h \ es [:'] Ts" (*<*) -apply (induct rule:WTrt_inducts) -prefer 14 -apply(case_tac "assigned V e") -apply(clarsimp simp add:fun_upd_same assigned_def simp del:fun_upd_apply) -apply(erule (2) WTrt'_WTrts'.intros) -apply(erule (1) WTrt'_WTrts'.intros) -apply(blast intro:WTrt'_WTrts'.intros)+ -done +proof(induct rule:WTrt_inducts) + case (WTrtBlock E V T e T') + then show ?case + proof(cases "assigned V e") + case True then show ?thesis using WTrtBlock.hyps(2) + by(clarsimp simp add:fun_upd_same assigned_def WTrt'_WTrts'.intros + simp del:fun_upd_apply) + next + case False then show ?thesis + by (simp add: WTrtBlock.hyps(2) WTrt'_WTrts'.intros) + qed +qed (blast intro:WTrt'_WTrts'.intros)+ (*>*) lemma wt'_wt: "P,E,h \ e :' T \ P,E,h \ e : T" and wts'_wts: "P,E,h \ es [:'] Ts \ P,E,h \ es [:] Ts" (*<*) -apply (induct rule:WTrt'_inducts) -prefer 16 -apply(rule WTrt_WTrts.intros) -apply(rule WTrt_WTrts.intros) -apply(rule WTrt_WTrts.intros) -apply simp -apply(erule (2) WTrt_WTrts.intros) -apply(blast intro:WTrt_WTrts.intros)+ -done +proof(induct rule:WTrt'_inducts) + case Block: (16 v T\<^sub>1 T E V e\<^sub>2 T\<^sub>2) + let ?E = "E(V \ T)" + have "P,?E,h \ Val v : T\<^sub>1" using Block.hyps(1) by simp + moreover have "P \ T\<^sub>1 \ T" by(rule Block.hyps(2)) + ultimately have "P,?E,h \ V:=Val v : Void" using WTrtLAss by simp + moreover have "P,?E,h \ e\<^sub>2 : T\<^sub>2" by(rule Block.hyps(4)) + ultimately have "P,?E,h \ V:=Val v;; e\<^sub>2 : T\<^sub>2" by blast + then show ?case by simp +qed (blast intro:WTrt_WTrts.intros)+ (*>*) corollary wt'_iff_wt: "(P,E,h \ e :' T) = (P,E,h \ e : T)" (*<*)by(blast intro:wt_wt' wt'_wt)(*>*) corollary wts'_iff_wts: "(P,E,h \ es [:'] Ts) = (P,E,h \ es [:] Ts)" (*<*)by(blast intro:wts_wts' wts'_wts)(*>*) (*<*) lemmas WTrt_inducts2 = WTrt'_inducts [unfolded wt'_iff_wt wts'_iff_wts, case_names WTrtNew WTrtCast WTrtVal WTrtVar WTrtBinOpEq WTrtBinOpAdd WTrtLAss WTrtFAcc WTrtFAccNT WTrtFAss WTrtFAssNT WTrtCall WTrtCallNT WTrtNil WTrtCons WTrtInitBlock WTrtBlock WTrtSeq WTrtCond WTrtWhile WTrtThrow WTrtTry, consumes 1] (*>*) theorem assumes wf: "wwf_J_prog P" and hconf: "P \ h \" shows progress: "P,E,h \ e : T \ (\l. \ \ e \dom l\; \ final e \ \ \e' s'. P \ \e,(h,l)\ \ \e',s'\)" and "P,E,h \ es [:] Ts \ (\l. \ \s es \dom l\; \ finals es \ \ \es' s'. P \ \es,(h,l)\ [\] \es',s'\)" (*<*) proof (induct rule:WTrt_inducts2) case WTrtNew show ?case proof cases assume "\a. h a = None" with assms WTrtNew show ?thesis by (fastforce del:exE intro!:RedNew simp add:new_Addr_def elim!:wf_Fields_Ex[THEN exE]) next assume "\(\a. h a = None)" with assms WTrtNew show ?thesis by(fastforce intro:RedNewFail simp add:new_Addr_def) qed next case (WTrtCast E e T C) have wte: "P,E,h \ e : T" and ref: "is_refT T" and IH: "\l. \\ e \dom l\; \ final e\ \ \e' s'. P \ \e,(h,l)\ \ \e',s'\" and D: "\ (Cast C e) \dom l\" by fact+ from D have De: "\ e \dom l\" by auto show ?case proof cases assume "final e" with wte ref show ?thesis proof (rule finalRefE) assume "e = null" thus ?case by(fastforce intro:RedCastNull) next fix D a assume A: "T = Class D" "e = addr a" show ?thesis proof cases assume "P \ D \\<^sup>* C" thus ?thesis using A wte by(fastforce intro:RedCast) next assume "\ P \ D \\<^sup>* C" thus ?thesis using A wte by(force intro!:RedCastFail) qed next fix a assume "e = Throw a" thus ?thesis by(blast intro!:red_reds.CastThrow) qed next assume nf: "\ final e" from IH[OF De nf] show ?thesis by (blast intro:CastRed) qed next case WTrtVal thus ?case by(simp add:final_def) next case WTrtVar thus ?case by(fastforce intro:RedVar simp:hyper_isin_def) next case (WTrtBinOpEq E e1 T1 e2 T2) show ?case proof cases assume "final e1" thus ?thesis proof (rule finalE) fix v1 assume [simp]: "e1 = Val v1" show ?thesis proof cases assume "final e2" thus ?thesis proof (rule finalE) fix v2 assume "e2 = Val v2" thus ?thesis using WTrtBinOpEq by(fastforce intro:RedBinOp) next fix a assume "e2 = Throw a" thus ?thesis by(auto intro:red_reds.BinOpThrow2) qed next assume "\ final e2" with WTrtBinOpEq show ?thesis by simp (fast intro!:BinOpRed2) qed next fix a assume "e1 = Throw a" thus ?thesis by simp (fast intro:red_reds.BinOpThrow1) qed next assume "\ final e1" with WTrtBinOpEq show ?thesis by simp (fast intro:BinOpRed1) qed next case (WTrtBinOpAdd E e1 e2) show ?case proof cases assume "final e1" thus ?thesis proof (rule finalE) fix v1 assume [simp]: "e1 = Val v1" show ?thesis proof cases assume "final e2" thus ?thesis proof (rule finalE) fix v2 assume "e2 = Val v2" thus ?thesis using WTrtBinOpAdd by(fastforce intro:RedBinOp) next fix a assume "e2 = Throw a" thus ?thesis by(auto intro:red_reds.BinOpThrow2) qed next assume "\ final e2" with WTrtBinOpAdd show ?thesis by simp (fast intro!:BinOpRed2) qed next fix a assume "e1 = Throw a" thus ?thesis by simp (fast intro:red_reds.BinOpThrow1) qed next assume "\ final e1" with WTrtBinOpAdd show ?thesis by simp (fast intro:BinOpRed1) qed next case (WTrtLAss E V T e T') show ?case proof cases assume "final e" with WTrtLAss show ?thesis by(auto simp:final_def intro!:RedLAss red_reds.LAssThrow) next assume "\ final e" with WTrtLAss show ?thesis by simp (fast intro:LAssRed) qed next case (WTrtFAcc E e C F T D) have wte: "P,E,h \ e : Class C" and field: "P \ C has F:T in D" by fact+ show ?case proof cases assume "final e" with wte show ?thesis proof (rule final_addrE) fix a assume e: "e = addr a" with wte obtain fs where hp: "h a = Some(C,fs)" by auto with hconf have "P,h \ (C,fs) \" using hconf_def by fastforce then obtain v where "fs(F,D) = Some v" using field by(fastforce dest:has_fields_fun simp:oconf_def has_field_def) with hp e show ?thesis by(fastforce intro:RedFAcc) next fix a assume "e = Throw a" thus ?thesis by(fastforce intro:red_reds.FAccThrow) qed next assume "\ final e" with WTrtFAcc show ?thesis by(fastforce intro!:FAccRed) qed next case (WTrtFAccNT E e F D T) show ?case proof cases assume "final e" \ \@{term e} is @{term null} or @{term throw}\ with WTrtFAccNT show ?thesis by(fastforce simp:final_def intro: RedFAccNull red_reds.FAccThrow) next assume "\ final e" \ \@{term e} reduces by IH\ with WTrtFAccNT show ?thesis by simp (fast intro:FAccRed) qed next case (WTrtFAss E e1 C F T D e2 T2) have wte1: "P,E,h \ e1 : Class C" by fact show ?case proof cases assume "final e1" with wte1 show ?thesis proof (rule final_addrE) fix a assume e1: "e1 = addr a" show ?thesis proof cases assume "final e2" thus ?thesis proof (rule finalE) fix v assume "e2 = Val v" thus ?thesis using e1 wte1 by(fastforce intro:RedFAss) next fix a assume "e2 = Throw a" thus ?thesis using e1 by(fastforce intro:red_reds.FAssThrow2) qed next assume "\ final e2" with WTrtFAss e1 show ?thesis by simp (fast intro!:FAssRed2) qed next fix a assume "e1 = Throw a" thus ?thesis by(fastforce intro:red_reds.FAssThrow1) qed next assume "\ final e1" with WTrtFAss show ?thesis by simp (blast intro!:FAssRed1) qed next case (WTrtFAssNT E e\<^sub>1 e\<^sub>2 T\<^sub>2 F D) show ?case proof cases assume e1: "final e\<^sub>1" \ \@{term e\<^sub>1} is @{term null} or @{term throw}\ show ?thesis proof cases assume "final e\<^sub>2" \ \@{term e\<^sub>2} is @{term Val} or @{term throw}\ with WTrtFAssNT e1 show ?thesis by(fastforce simp:final_def intro: RedFAssNull red_reds.FAssThrow1 red_reds.FAssThrow2) next assume "\ final e\<^sub>2" \ \@{term e\<^sub>2} reduces by IH\ with WTrtFAssNT e1 show ?thesis by (fastforce simp:final_def intro!:red_reds.FAssRed2 red_reds.FAssThrow1) qed next assume "\ final e\<^sub>1" \ \@{term e\<^sub>1} reduces by IH\ with WTrtFAssNT show ?thesis by (fastforce intro:FAssRed1) qed next case (WTrtCall E e C M Ts T pns body D es Ts') have 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 [:] Ts'"and sub: "P \ Ts' [\] Ts" and IHes: "\l. \\s es \dom l\; \ finals es\ \ \es' s'. P \ \es,(h,l)\ [\] \es',s'\" and D: "\ (e\M(es)) \dom l\" by fact+ show ?case proof cases assume "final e" with wte show ?thesis proof (rule final_addrE) fix a assume e_addr: "e = addr a" show ?thesis proof cases assume es: "\vs. es = map Val vs" from wte e_addr obtain fs where ha: "h a = Some(C,fs)" by auto show ?thesis using e_addr ha "method" WTrts_same_length[OF wtes] sub es sees_wf_mdecl[OF wf "method"] by (fastforce intro!: RedCall simp:list_all2_iff wf_mdecl_def) next assume "\(\vs. es = map Val vs)" hence not_all_Val: "\(\e \ set es. \v. e = Val v)" by(simp add:ex_map_conv) let ?ves = "takeWhile (\e. \v. e = Val v) es" let ?rest = "dropWhile (\e. \v. e = Val v) es" let ?ex = "hd ?rest" let ?rst = "tl ?rest" from not_all_Val have nonempty: "?rest \ []" by auto hence es: "es = ?ves @ ?ex # ?rst" by simp have "\e \ set ?ves. \v. e = Val v" by(fastforce dest:set_takeWhileD) then obtain vs where ves: "?ves = map Val vs" using ex_map_conv by blast show ?thesis proof cases assume "final ?ex" moreover from nonempty have "\(\v. ?ex = Val v)" by(auto simp:neq_Nil_conv simp del:dropWhile_eq_Nil_conv) (simp add:dropWhile_eq_Cons_conv) ultimately obtain b where ex_Throw: "?ex = Throw b" by(fast elim!:finalE) show ?thesis using e_addr es ex_Throw ves by(fastforce intro:CallThrowParams) next assume not_fin: "\ final ?ex" have "finals es = finals(?ves @ ?ex # ?rst)" using es by(rule arg_cong) also have "\ = finals(?ex # ?rst)" using ves by simp finally have "finals es = finals(?ex # ?rst)" . hence "\ finals es" using not_finals_ConsI[OF not_fin] by blast thus ?thesis using e_addr D IHes by(fastforce intro!:CallParams) qed qed next fix a assume "e = Throw a" with WTrtCall.prems show ?thesis by(fast intro!:CallThrowObj) qed next assume "\ final e" with WTrtCall show ?thesis by simp (blast intro!:CallObj) qed next case (WTrtCallNT E e es Ts M T) show ?case proof cases assume "final e" moreover { fix v assume e: "e = Val v" hence "e = null" using WTrtCallNT by simp have ?case proof cases assume "finals es" moreover { fix vs assume "es = map Val vs" with WTrtCallNT e have ?thesis by(fastforce intro: RedCallNull) } moreover { fix vs a es' assume "es = map Val vs @ Throw a # es'" with WTrtCallNT e have ?thesis by(fastforce intro: CallThrowParams) } ultimately show ?thesis by(fastforce simp:finals_def) next assume "\ finals es" \ \@{term es} reduces by IH\ with WTrtCallNT e show ?thesis by(fastforce intro: CallParams) qed } moreover { fix a assume "e = Throw a" with WTrtCallNT have ?case by(fastforce intro: CallThrowObj) } ultimately show ?thesis by(fastforce simp:final_def) next assume "\ final e" \ \@{term e} reduces by IH\ with WTrtCallNT show ?thesis by (fastforce intro:CallObj) qed next case WTrtNil thus ?case by simp next case (WTrtCons E e T es Ts) have IHe: "\l. \\ e \dom l\; \ final e\ \ \e' s'. P \ \e,(h,l)\ \ \e',s'\" and IHes: "\l. \\s es \dom l\; \ finals es\ \ \es' s'. P \ \es,(h,l)\ [\] \es',s'\" and D: "\s (e#es) \dom l\" and not_fins: "\ finals(e # es)" by fact+ have De: "\ e \dom l\" and Des: "\s es (\dom l\ \ \ e)" using D by auto show ?case proof cases assume "final e" thus ?thesis proof (rule finalE) fix v assume e: "e = Val v" hence Des': "\s es \dom l\" using De Des by auto have not_fins_tl: "\ finals es" using not_fins e by simp show ?thesis using e IHes[OF Des' not_fins_tl] by (blast intro!:ListRed2) next fix a assume "e = Throw a" hence False using not_fins by simp thus ?thesis .. qed next assume "\ final e" with IHe[OF De] show ?thesis by(fast intro!:ListRed1) qed next case (WTrtInitBlock v T\<^sub>1 T E V e\<^sub>2 T\<^sub>2) have IH2: "\l. \\ e\<^sub>2 \dom l\; \ final e\<^sub>2\ \ \e' s'. P \ \e\<^sub>2,(h,l)\ \ \e',s'\" and D: "\ {V:T := Val v; e\<^sub>2} \dom l\" by fact+ show ?case proof cases assume "final e\<^sub>2" then show ?thesis proof (rule finalE) fix v\<^sub>2 assume "e\<^sub>2 = Val v\<^sub>2" thus ?thesis by(fast intro:RedInitBlock) next fix a assume "e\<^sub>2 = Throw a" thus ?thesis by(fast intro:red_reds.InitBlockThrow) qed next assume not_fin2: "\ final e\<^sub>2" from D have D2: "\ e\<^sub>2 \dom(l(V\v))\" by (auto simp:hyperset_defs) from IH2[OF D2 not_fin2] obtain h' l' e' where red2: "P \ \e\<^sub>2,(h, l(V\v))\ \ \e',(h', l')\" by auto from red_lcl_incr[OF red2] have "V \ dom l'" by auto with red2 show ?thesis by(fastforce intro:InitBlockRed) qed next case (WTrtBlock E V T e T') have IH: "\l. \\ e \dom l\; \ final e\ \ \e' s'. P \ \e,(h,l)\ \ \e',s'\" and unass: "\ assigned V e" and D: "\ {V:T; e} \dom l\" by fact+ show ?case proof cases assume "final e" thus ?thesis proof (rule finalE) fix v assume "e = Val v" thus ?thesis by(fast intro:RedBlock) next fix a assume "e = Throw a" thus ?thesis by(fast intro:red_reds.BlockThrow) qed next assume not_fin: "\ final e" from D have De: "\ e \dom(l(V:=None))\" by(simp add:hyperset_defs) from IH[OF De not_fin] obtain h' l' e' where red: "P \ \e,(h,l(V:=None))\ \ \e',(h',l')\" by auto show ?thesis proof (cases "l' V") assume "l' V = None" with red unass show ?thesis by(blast intro: BlockRedNone) next fix v assume "l' V = Some v" with red unass show ?thesis by(blast intro: BlockRedSome) qed qed next case (WTrtSeq E e1 T1 e2 T2) show ?case proof cases assume "final e1" thus ?thesis by(fast elim:finalE intro:RedSeq red_reds.SeqThrow) next assume "\ final e1" with WTrtSeq show ?thesis by simp (blast intro:SeqRed) qed next case (WTrtCond E e e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 T) have wt: "P,E,h \ e : Boolean" by fact show ?case proof cases assume "final e" thus ?thesis proof (rule finalE) fix v assume val: "e = Val v" then obtain b where v: "v = Bool b" using wt by auto show ?thesis proof (cases b) case True with val v show ?thesis by(auto intro:RedCondT) next case False with val v show ?thesis by(auto intro:RedCondF) qed next fix a assume "e = Throw a" thus ?thesis by(fast intro:red_reds.CondThrow) qed next assume "\ final e" with WTrtCond show ?thesis by simp (fast intro:CondRed) qed next case WTrtWhile show ?case by(fast intro:RedWhile) next case (WTrtThrow E e T\<^sub>r T) show ?case proof cases assume "final e" \ \Then @{term e} must be @{term throw} or @{term null}\ with WTrtThrow show ?thesis by(fastforce simp:final_def is_refT_def intro:red_reds.ThrowThrow red_reds.RedThrowNull) next assume "\ final e" \ \Then @{term e} must reduce\ with WTrtThrow show ?thesis by simp (blast intro:ThrowRed) qed next case (WTrtTry E e1 T1 V C e2 T2) have wt1: "P,E,h \ e1 : T1" by fact show ?case proof cases assume "final e1" thus ?thesis proof (rule finalE) fix v assume "e1 = Val v" thus ?thesis by(fast intro:RedTry) next fix a assume e1_Throw: "e1 = Throw a" with wt1 obtain D fs where ha: "h a = Some(D,fs)" by fastforce show ?thesis proof cases assume "P \ D \\<^sup>* C" with e1_Throw ha show ?thesis by(fastforce intro!:RedTryCatch) next assume "\ P \ D \\<^sup>* C" with e1_Throw ha show ?thesis by(force intro!:RedTryFail) qed qed next assume "\ final e1" with WTrtTry show ?thesis by simp (fast intro:TryRed) qed qed (*>*) end 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/Jinja/J/WellType.thy b/thys/Jinja/J/WellType.thy --- a/thys/Jinja/J/WellType.thy +++ b/thys/Jinja/J/WellType.thy @@ -1,235 +1,201 @@ (* Title: Jinja/J/WellType.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Well-typedness of Jinja expressions\ theory WellType imports "../Common/Objects" Expr begin type_synonym env = "vname \ ty" inductive WT :: "[J_prog,env, expr , ty ] \ bool" ("_,_ \ _ :: _" [51,51,51]50) and WTs :: "[J_prog,env, expr list, ty list] \ bool" ("_,_ \ _ [::] _" [51,51,51]50) for P :: J_prog where WTNew: "is_class P C \ P,E \ new C :: Class C" | WTCast: "\ P,E \ e :: Class D; is_class P C; P \ C \\<^sup>* D \ P \ D \\<^sup>* C \ \ P,E \ Cast C e :: Class C" | WTVal: "typeof v = Some T \ P,E \ Val v :: T" | WTVar: "E V = Some T \ P,E \ Var V :: T" (* WTBinOp: "\ P,E \ e\<^sub>1 :: T\<^sub>1; P,E \ e\<^sub>2 :: T\<^sub>2; case bop of Eq \ (P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1) \ T = Boolean | Add \ T\<^sub>1 = Integer \ T\<^sub>2 = Integer \ T = Integer \ \ P,E \ e\<^sub>1 \bop\ e\<^sub>2 :: T" *) | WTBinOpEq: "\ P,E \ e\<^sub>1 :: T\<^sub>1; P,E \ e\<^sub>2 :: T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1 \ \ P,E \ e\<^sub>1 \Eq\ e\<^sub>2 :: Boolean" | WTBinOpAdd: "\ P,E \ e\<^sub>1 :: Integer; P,E \ e\<^sub>2 :: Integer \ \ P,E \ e\<^sub>1 \Add\ e\<^sub>2 :: Integer" | WTLAss: "\ E V = Some T; P,E \ e :: T'; P \ T' \ T; V \ this \ \ P,E \ V:=e :: Void" | WTFAcc: "\ P,E \ e :: Class C; P \ C sees F:T in D \ \ P,E \ e\F{D} :: T" | WTFAss: "\ P,E \ e\<^sub>1 :: Class C; P \ C sees F:T in D; P,E \ e\<^sub>2 :: T'; P \ T' \ T \ \ P,E \ e\<^sub>1\F{D}:=e\<^sub>2 :: Void" | WTCall: "\ P,E \ e :: Class C; P \ C sees M:Ts \ T = (pns,body) in D; P,E \ es [::] Ts'; P \ Ts' [\] Ts \ \ P,E \ e\M(es) :: T" | WTBlock: "\ is_type P T; P,E(V \ T) \ e :: T' \ \ P,E \ {V:T; e} :: T'" | WTSeq: "\ P,E \ e\<^sub>1::T\<^sub>1; P,E \ e\<^sub>2::T\<^sub>2 \ \ P,E \ e\<^sub>1;;e\<^sub>2 :: T\<^sub>2" | WTCond: "\ P,E \ e :: Boolean; P,E \ e\<^sub>1::T\<^sub>1; P,E \ e\<^sub>2::T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1; P \ T\<^sub>1 \ T\<^sub>2 \ T = T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ T = T\<^sub>1 \ \ P,E \ if (e) e\<^sub>1 else e\<^sub>2 :: T" | WTWhile: "\ P,E \ e :: Boolean; P,E \ c::T \ \ P,E \ while (e) c :: Void" | WTThrow: "P,E \ e :: Class C \ P,E \ throw e :: Void" | WTTry: "\ P,E \ e\<^sub>1 :: T; P,E(V \ Class C) \ e\<^sub>2 :: T; is_class P C \ \ P,E \ try e\<^sub>1 catch(C V) e\<^sub>2 :: T" \ \well-typed expression lists\ | WTNil: "P,E \ [] [::] []" | WTCons: "\ P,E \ e :: T; P,E \ es [::] Ts \ \ P,E \ e#es [::] T#Ts" (*<*) (* lemmas [intro!] = WTNew WTCast WTVal WTVar WTBinOp WTLAss WTFAcc WTFAss WTCall WTBlock WTSeq WTWhile WTThrow WTTry WTNil WTCons lemmas [intro] = WTCond1 WTCond2 *) declare WT_WTs.intros[intro!] (* WTNil[iff] *) lemmas WT_WTs_induct = WT_WTs.induct [split_format (complete)] and WT_WTs_inducts = WT_WTs.inducts [split_format (complete)] (*>*) lemma [iff]: "(P,E \ [] [::] Ts) = (Ts = [])" -(*<*) -apply(rule iffI) -apply (auto elim: WTs.cases) -done -(*>*) +(*<*)by (rule iffI) (auto elim: WTs.cases)(*>*) lemma [iff]: "(P,E \ e#es [::] T#Ts) = (P,E \ e :: T \ P,E \ es [::] Ts)" -(*<*) -apply(rule iffI) -apply (auto elim: WTs.cases) -done -(*>*) +(*<*)by (rule iffI) (auto elim: WTs.cases)(*>*) lemma [iff]: "(P,E \ (e#es) [::] Ts) = (\U Us. Ts = U#Us \ P,E \ e :: U \ P,E \ es [::] Us)" -(*<*) -apply(rule iffI) -apply (auto elim: WTs.cases) -done -(*>*) +(*<*)by (rule iffI) (auto elim: WTs.cases)(*>*) lemma [iff]: "\Ts. (P,E \ es\<^sub>1 @ es\<^sub>2 [::] Ts) = (\Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \ P,E \ es\<^sub>1 [::] Ts\<^sub>1 \ P,E \ es\<^sub>2[::]Ts\<^sub>2)" (*<*) -apply(induct es\<^sub>1 type:list) - apply simp -apply clarsimp -apply(erule thin_rl) -apply (rule iffI) - apply clarsimp - apply(rule exI)+ - apply(rule conjI) - prefer 2 apply blast - apply simp -apply fastforce -done +proof(induct es\<^sub>1 type:list) + case (Cons a list) + let ?lhs = "(\U Us. Ts = U # Us \ P,E \ a :: U \ + (\Ts\<^sub>1 Ts\<^sub>2. Us = Ts\<^sub>1 @ Ts\<^sub>2 \ P,E \ list [::] Ts\<^sub>1 \ P,E \ es\<^sub>2 [::] Ts\<^sub>2))" + let ?rhs = "(\Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \ + (\U Us. Ts\<^sub>1 = U # Us \ P,E \ a :: U \ P,E \ list [::] Us) \ P,E \ es\<^sub>2 [::] Ts\<^sub>2)" + { assume ?lhs + then have ?rhs by (auto intro: Cons_eq_appendI) + } + moreover { + assume ?rhs + then have ?lhs by fastforce + } + ultimately have "?lhs = ?rhs" by(rule iffI) + then show ?case by (clarsimp simp: Cons) +qed simp (*>*) lemma [iff]: "P,E \ Val v :: T = (typeof v = Some T)" -(*<*) -apply(rule iffI) -apply (auto elim: WT.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*) lemma [iff]: "P,E \ Var V :: T = (E V = Some T)" -(*<*) -apply(rule iffI) -apply (auto elim: WT.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*) lemma [iff]: "P,E \ e\<^sub>1;;e\<^sub>2 :: T\<^sub>2 = (\T\<^sub>1. P,E \ e\<^sub>1::T\<^sub>1 \ P,E \ e\<^sub>2::T\<^sub>2)" -(*<*) -apply(rule iffI) -apply (auto elim: WT.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*) lemma [iff]: "(P,E \ {V:T; e} :: T') = (is_type P T \ P,E(V\T) \ e :: T')" -(*<*) -apply(rule iffI) -apply (auto elim: WT.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*) (*<*) inductive_cases WT_elim_cases[elim!]: "P,E \ V :=e :: T" "P,E \ if (e) e\<^sub>1 else e\<^sub>2 :: T" "P,E \ while (e) c :: T" "P,E \ throw e :: T" "P,E \ try e\<^sub>1 catch(C V) e\<^sub>2 :: T" "P,E \ Cast D e :: T" "P,E \ a\F{D} :: T" "P,E \ a\F{D} := v :: T" "P,E \ e\<^sub>1 \bop\ e\<^sub>2 :: T" "P,E \ new C :: T" "P,E \ e\M(ps) :: T" (*>*) lemma wt_env_mono: "P,E \ e :: T \ (\E'. E \\<^sub>m E' \ P,E' \ e :: T)" and "P,E \ es [::] Ts \ (\E'. E \\<^sub>m E' \ P,E' \ es [::] Ts)" (*<*) -apply(induct rule: WT_WTs_inducts) -apply(simp add: WTNew) -apply(fastforce simp: WTCast) -apply(fastforce simp: WTVal) -apply(simp add: WTVar map_le_def dom_def) -apply(fastforce simp: WTBinOpEq) -apply(fastforce simp: WTBinOpAdd) -apply(force simp:map_le_def) -apply(fastforce simp: WTFAcc) -apply(fastforce simp: WTFAss del:WT_WTs.intros WT_elim_cases) -apply(fastforce simp: WTCall) -apply(fastforce simp: map_le_def WTBlock) -apply(fastforce simp: WTSeq) -apply(fastforce simp: WTCond) -apply(fastforce simp: WTWhile) -apply(fastforce simp: WTThrow) -apply(fastforce simp: WTTry map_le_def dom_def) -apply(simp add: WTNil) -apply(simp add: WTCons) -done +proof(induct rule: WT_WTs_inducts) + case WTVar then show ?case by(simp add: map_le_def dom_def) +next + case WTLAss then show ?case by(force simp:map_le_def) +qed fastforce+ (*>*) lemma WT_fv: "P,E \ e :: T \ fv e \ dom E" and "P,E \ es [::] Ts \ fvs es \ dom E" (*<*) -apply(induct rule:WT_WTs.inducts) -apply(simp_all del: fun_upd_apply) -apply fast+ -done +proof(induct rule:WT_WTs.inducts) + case WTVar then show ?case by fastforce +next + case WTLAss then show ?case by fastforce +next + case WTBlock then show ?case by fastforce +next + case WTTry then show ?case by fastforce +qed simp_all end (*>*) diff --git a/thys/Jinja/J/WellTypeRT.thy b/thys/Jinja/J/WellTypeRT.thy --- a/thys/Jinja/J/WellTypeRT.thy +++ b/thys/Jinja/J/WellTypeRT.thy @@ -1,317 +1,243 @@ (* Title: Jinja/J/WellTypeRT.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Runtime Well-typedness\ theory WellTypeRT imports WellType begin inductive WTrt :: "J_prog \ heap \ env \ expr \ ty \ bool" and WTrts :: "J_prog \ heap \ env \ expr list \ ty list \ bool" and WTrt2 :: "[J_prog,env,heap,expr,ty] \ bool" ("_,_,_ \ _ : _" [51,51,51]50) and WTrts2 :: "[J_prog,env,heap,expr list, ty list] \ bool" ("_,_,_ \ _ [:] _" [51,51,51]50) for P :: J_prog and h :: heap where "P,E,h \ e : T \ WTrt P h E e T" | "P,E,h \ es[:]Ts \ WTrts P h E es Ts" | WTrtNew: "is_class P C \ P,E,h \ new C : Class C" | WTrtCast: "\ P,E,h \ e : T; is_refT T; is_class P C \ \ P,E,h \ Cast C e : Class C" | WTrtVal: "typeof\<^bsub>h\<^esub> v = Some T \ P,E,h \ Val v : T" | WTrtVar: "E V = Some T \ P,E,h \ Var V : T" (* WTrtBinOp: "\ P,E,h \ e\<^sub>1 : T\<^sub>1; P,E,h \ e\<^sub>2 : T\<^sub>2; case bop of Eq \ T = Boolean | Add \ T\<^sub>1 = Integer \ T\<^sub>2 = Integer \ T = Integer \ \ P,E,h \ e\<^sub>1 \bop\ e\<^sub>2 : T" *) | WTrtBinOpEq: "\ P,E,h \ e\<^sub>1 : T\<^sub>1; P,E,h \ e\<^sub>2 : T\<^sub>2 \ \ P,E,h \ e\<^sub>1 \Eq\ e\<^sub>2 : Boolean" | WTrtBinOpAdd: "\ P,E,h \ e\<^sub>1 : Integer; P,E,h \ e\<^sub>2 : Integer \ \ P,E,h \ e\<^sub>1 \Add\ e\<^sub>2 : Integer" | WTrtLAss: "\ E V = Some T; P,E,h \ e : T'; P \ T' \ T \ \ P,E,h \ V:=e : Void" | WTrtFAcc: "\ P,E,h \ e : Class C; P \ C has F:T in D \ \ P,E,h \ e\F{D} : T" | WTrtFAccNT: "P,E,h \ e : NT \ P,E,h \ e\F{D} : T" | WTrtFAss: "\ P,E,h \ e\<^sub>1 : Class C; P \ C has F:T in D; P,E,h \ e\<^sub>2 : T\<^sub>2; P \ T\<^sub>2 \ T \ \ P,E,h \ e\<^sub>1\F{D}:=e\<^sub>2 : Void" | WTrtFAssNT: "\ P,E,h \ e\<^sub>1:NT; P,E,h \ e\<^sub>2 : T\<^sub>2 \ \ P,E,h \ e\<^sub>1\F{D}:=e\<^sub>2 : Void" | WTrtCall: "\ P,E,h \ e : Class C; P \ C sees M:Ts \ T = (pns,body) in D; P,E,h \ es [:] Ts'; P \ Ts' [\] Ts \ \ P,E,h \ e\M(es) : T" | WTrtCallNT: "\ P,E,h \ e : NT; P,E,h \ es [:] Ts \ \ P,E,h \ e\M(es) : T" | WTrtBlock: "P,E(V\T),h \ e : T' \ P,E,h \ {V:T; e} : T'" | WTrtSeq: "\ P,E,h \ e\<^sub>1:T\<^sub>1; P,E,h \ e\<^sub>2:T\<^sub>2 \ \ P,E,h \ e\<^sub>1;;e\<^sub>2 : T\<^sub>2" | WTrtCond: "\ P,E,h \ e : Boolean; P,E,h \ e\<^sub>1:T\<^sub>1; P,E,h \ e\<^sub>2:T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1; P \ T\<^sub>1 \ T\<^sub>2 \ T = T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ T = T\<^sub>1 \ \ P,E,h \ if (e) e\<^sub>1 else e\<^sub>2 : T" | WTrtWhile: "\ P,E,h \ e : Boolean; P,E,h \ c:T \ \ P,E,h \ while(e) c : Void" | WTrtThrow: "\ P,E,h \ e : T\<^sub>r; is_refT T\<^sub>r \ \ P,E,h \ throw e : T" | WTrtTry: "\ P,E,h \ e\<^sub>1 : T\<^sub>1; P,E(V \ Class C),h \ e\<^sub>2 : T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ \ P,E,h \ try e\<^sub>1 catch(C V) e\<^sub>2 : T\<^sub>2" \ \well-typed expression lists\ | WTrtNil: "P,E,h \ [] [:] []" | WTrtCons: "\ P,E,h \ e : T; P,E,h \ es [:] Ts \ \ P,E,h \ e#es [:] T#Ts" (*<*) declare WTrt_WTrts.intros[intro!] WTrtNil[iff] declare WTrtFAcc[rule del] WTrtFAccNT[rule del] WTrtFAss[rule del] WTrtFAssNT[rule del] WTrtCall[rule del] WTrtCallNT[rule del] lemmas WTrt_induct = WTrt_WTrts.induct [split_format (complete)] and WTrt_inducts = WTrt_WTrts.inducts [split_format (complete)] (*>*) subsection\Easy consequences\ lemma [iff]: "(P,E,h \ [] [:] Ts) = (Ts = [])" -(*<*) -apply(rule iffI) -apply (auto elim: WTrts.cases) -done -(*>*) +(*<*)by(rule iffI) (auto elim: WTrts.cases)(*>*) lemma [iff]: "(P,E,h \ e#es [:] T#Ts) = (P,E,h \ e : T \ P,E,h \ es [:] Ts)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrts.cases) -done -(*>*) +(*<*)by(rule iffI) (auto elim: WTrts.cases)(*>*) lemma [iff]: "(P,E,h \ (e#es) [:] Ts) = (\U Us. Ts = U#Us \ P,E,h \ e : U \ P,E,h \ es [:] Us)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrts.cases) -done -(*>*) +(*<*)by(rule iffI) (auto elim: WTrts.cases)(*>*) lemma [simp]: "\Ts. (P,E,h \ es\<^sub>1 @ es\<^sub>2 [:] Ts) = (\Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \ P,E,h \ es\<^sub>1 [:] Ts\<^sub>1 & P,E,h \ es\<^sub>2[:]Ts\<^sub>2)" (*<*) -apply(induct_tac es\<^sub>1) - apply simp -apply clarsimp -apply(erule thin_rl) -apply (rule iffI) - apply clarsimp - apply(rule exI)+ - apply(rule conjI) - prefer 2 apply blast - apply simp -apply fastforce -done +proof(induct es\<^sub>1) + case (Cons a list) + let ?lhs = "\Ts. (\U Us. Ts = U # Us \ P,E,h \ a : U \ + (\Ts\<^sub>1 Ts\<^sub>2. Us = Ts\<^sub>1 @ Ts\<^sub>2 \ P,E,h \ list [:] Ts\<^sub>1 \ P,E,h \ es\<^sub>2 [:] Ts\<^sub>2))" + let ?rhs = "\Ts. (\Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \ + (\U Us. Ts\<^sub>1 = U # Us \ P,E,h \ a : U \ P,E,h \ list [:] Us) \ P,E,h \ es\<^sub>2 [:] Ts\<^sub>2)" + { fix Ts assume "?lhs Ts" + then have "?rhs Ts" by (auto intro: Cons_eq_appendI) + } + moreover { + fix Ts assume "?rhs Ts" + then have "?lhs Ts" by fastforce + } + ultimately have "\Ts. ?lhs Ts = ?rhs Ts" by(rule iffI) + then show ?case by (clarsimp simp: Cons) +qed simp (*>*) lemma [iff]: "P,E,h \ Val v : T = (typeof\<^bsub>h\<^esub> v = Some T)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrt.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*) lemma [iff]: "P,E,h \ Var v : T = (E v = Some T)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrt.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*) lemma [iff]: "P,E,h \ e\<^sub>1;;e\<^sub>2 : T\<^sub>2 = (\T\<^sub>1. P,E,h \ e\<^sub>1:T\<^sub>1 \ P,E,h \ e\<^sub>2:T\<^sub>2)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrt.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*) lemma [iff]: "P,E,h \ {V:T; e} : T' = (P,E(V\T),h \ e : T')" -(*<*) -apply(rule iffI) -apply (auto elim: WTrt.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*) + (*<*) inductive_cases WTrt_elim_cases[elim!]: "P,E,h \ v :=e : T" "P,E,h \ if (e) e\<^sub>1 else e\<^sub>2 : T" "P,E,h \ while(e) c : T" "P,E,h \ throw e : T" "P,E,h \ try e\<^sub>1 catch(C V) e\<^sub>2 : T" "P,E,h \ Cast D e : T" "P,E,h \ e\F{D} : T" "P,E,h \ e\F{D} := v : T" "P,E,h \ e\<^sub>1 \bop\ e\<^sub>2 : T" "P,E,h \ new C : T" "P,E,h \ e\M{D}(es) : T" (*>*) subsection\Some interesting lemmas\ lemma WTrts_Val[simp]: "\Ts. (P,E,h \ map Val vs [:] Ts) = (map (typeof\<^bsub>h\<^esub>) vs = map Some Ts)" (*<*) -apply(induct vs) - apply simp -apply(case_tac Ts) - apply simp -apply simp -done +proof(induct vs) + case (Cons a vs) + then show ?case by(case_tac Ts; simp) +qed simp (*>*) lemma WTrts_same_length: "\Ts. P,E,h \ es [:] Ts \ length es = length Ts" (*<*)by(induct es type:list)auto(*>*) lemma WTrt_env_mono: "P,E,h \ e : T \ (\E'. E \\<^sub>m E' \ P,E',h \ e : T)" and "P,E,h \ es [:] Ts \ (\E'. E \\<^sub>m E' \ P,E',h \ es [:] Ts)" (*<*) -apply(induct rule: WTrt_inducts) -apply(simp add: WTrtNew) -apply(fastforce simp: WTrtCast) -apply(fastforce simp: WTrtVal) -apply(simp add: WTrtVar map_le_def dom_def) -apply(fastforce simp add: WTrtBinOpEq) -apply(fastforce simp add: WTrtBinOpAdd) -apply(force simp: map_le_def) -apply(fastforce simp: WTrtFAcc) -apply(simp add: WTrtFAccNT) -apply(fastforce simp: WTrtFAss) -apply(fastforce simp: WTrtFAssNT) -apply(fastforce simp: WTrtCall) -apply(fastforce simp: WTrtCallNT) -apply(simp add: WTrtNil) -apply(simp add: WTrtCons) -apply(fastforce simp: map_le_def) -apply(fastforce) -apply(fastforce simp: WTrtSeq) -apply(fastforce simp: WTrtWhile) -apply(fastforce simp: WTrtThrow) -apply(auto simp: WTrtTry map_le_def dom_def) -done +proof(induct rule: WTrt_inducts) + case WTrtVar then show ?case by(simp add: map_le_def dom_def) +next + case WTrtLAss then show ?case by(force simp:map_le_def) +qed(fastforce intro: WTrt_WTrts.intros)+ (*>*) lemma WTrt_hext_mono: "P,E,h \ e : T \ h \ h' \ P,E,h' \ e : T" and WTrts_hext_mono: "P,E,h \ es [:] Ts \ h \ h' \ P,E,h' \ es [:] Ts" (*<*) -apply(induct rule: WTrt_inducts) -apply(simp add: WTrtNew) -apply(fastforce simp: WTrtCast) -apply(fastforce simp: WTrtVal dest:hext_typeof_mono) -apply(simp add: WTrtVar) -apply(fastforce simp add: WTrtBinOpEq) -apply(fastforce simp add: WTrtBinOpAdd) -apply(fastforce simp add: WTrtLAss) -apply(fast intro: WTrtFAcc) -apply(simp add: WTrtFAccNT) -apply(fastforce simp: WTrtFAss del:WTrt_WTrts.intros WTrt_elim_cases) -apply(fastforce simp: WTrtFAssNT) -apply(fastforce simp: WTrtCall) -apply(fastforce simp: WTrtCallNT) -apply(fastforce) -apply(fastforce simp add: WTrtSeq) -apply(fastforce simp add: WTrtCond) -apply(fastforce simp add: WTrtWhile) -apply(fastforce simp add: WTrtThrow) -apply(fastforce simp: WTrtTry) -apply(simp add: WTrtNil) -apply(simp add: WTrtCons) -done +proof(induct rule: WTrt_inducts) + case WTrtVal then show ?case by(simp add: hext_typeof_mono) +qed(fastforce intro: WTrt_WTrts.intros)+ (*>*) lemma WT_implies_WTrt: "P,E \ e :: T \ P,E,h \ e : T" and WTs_implies_WTrts: "P,E \ es [::] Ts \ P,E,h \ es [:] Ts" (*<*) -apply(induct rule: WT_WTs_inducts) -apply fast -apply (fast) -apply(fastforce dest:typeof_lit_typeof) -apply(simp) -apply(fastforce) -apply(fastforce) -apply(fastforce) -apply(fastforce simp: WTrtFAcc has_visible_field) -apply(fastforce simp: WTrtFAss dest: has_visible_field) -apply(fastforce simp: WTrtCall) -apply(fastforce) -apply(fastforce) -apply(fastforce simp: WTrtCond) -apply(fastforce) -apply(fastforce) -apply(fastforce) -apply(simp) -apply(simp) -done +proof(induct rule: WT_WTs_inducts) + case WTVal then show ?case by(fastforce dest: typeof_lit_typeof) +next + case WTFAcc + then show ?case by(fastforce simp: WTrtFAcc has_visible_field) +next + case WTFAss + then show ?case by(fastforce simp: WTrtFAss dest: has_visible_field) +qed(fastforce intro: WTrt_WTrts.intros)+ (*>*) end diff --git a/thys/Jinja/JVM/JVMDefensive.thy b/thys/Jinja/JVM/JVMDefensive.thy --- a/thys/Jinja/JVM/JVMDefensive.thy +++ b/thys/Jinja/JVM/JVMDefensive.thy @@ -1,190 +1,196 @@ (* Title: HOL/MicroJava/JVM/JVMDefensive.thy Author: Gerwin Klein Copyright GPL *) section \A Defensive JVM\ theory JVMDefensive imports JVMExec "../Common/Conform" begin text \ Extend the state space by one element indicating a type error (or other abnormal termination)\ datatype 'a type_error = TypeError | Normal 'a fun is_Addr :: "val \ bool" where "is_Addr (Addr a) \ True" | "is_Addr v \ False" fun is_Intg :: "val \ bool" where "is_Intg (Intg i) \ True" | "is_Intg v \ False" fun is_Bool :: "val \ bool" where "is_Bool (Bool b) \ True" | "is_Bool v \ False" definition is_Ref :: "val \ bool" where "is_Ref v \ v = Null \ is_Addr v" primrec check_instr :: "[instr, jvm_prog, heap, val list, val list, cname, mname, pc, frame list] \ bool" where check_instr_Load: "check_instr (Load n) P h stk loc C M\<^sub>0 pc frs = (n < length loc)" | check_instr_Store: "check_instr (Store n) P h stk loc C\<^sub>0 M\<^sub>0 pc frs = (0 < length stk \ n < length loc)" | check_instr_Push: "check_instr (Push v) P h stk loc C\<^sub>0 M\<^sub>0 pc frs = (\is_Addr v)" | check_instr_New: "check_instr (New C) P h stk loc C\<^sub>0 M\<^sub>0 pc frs = is_class P C" | check_instr_Getfield: "check_instr (Getfield F C) P h stk loc C\<^sub>0 M\<^sub>0 pc frs = (0 < length stk \ (\C' T. P \ C sees F:T in C') \ (let (C', T) = field P C F; ref = hd stk in C' = C \ is_Ref ref \ (ref \ Null \ h (the_Addr ref) \ None \ (let (D,vs) = the (h (the_Addr ref)) in P \ D \\<^sup>* C \ vs (F,C) \ None \ P,h \ the (vs (F,C)) :\ T))))" | check_instr_Putfield: "check_instr (Putfield F C) P h stk loc C\<^sub>0 M\<^sub>0 pc frs = (1 < length stk \ (\C' T. P \ C sees F:T in C') \ (let (C', T) = field P C F; v = hd stk; ref = hd (tl stk) in C' = C \ is_Ref ref \ (ref \ Null \ h (the_Addr ref) \ None \ (let D = fst (the (h (the_Addr ref))) in P \ D \\<^sup>* C \ P,h \ v :\ T))))" | check_instr_Checkcast: "check_instr (Checkcast C) P h stk loc C\<^sub>0 M\<^sub>0 pc frs = (0 < length stk \ is_class P C \ is_Ref (hd stk))" | check_instr_Invoke: "check_instr (Invoke M n) P h stk loc C\<^sub>0 M\<^sub>0 pc frs = (n < length stk \ is_Ref (stk!n) \ (stk!n \ Null \ (let a = the_Addr (stk!n); C = cname_of h a; Ts = fst (snd (method P C M)) in h a \ None \ P \ C has M \ P,h \ rev (take n stk) [:\] Ts)))" | check_instr_Return: "check_instr Return P h stk loc C\<^sub>0 M\<^sub>0 pc frs = (0 < length stk \ ((0 < length frs) \ (P \ C\<^sub>0 has M\<^sub>0) \ (let v = hd stk; T = fst (snd (snd (method P C\<^sub>0 M\<^sub>0))) in P,h \ v :\ T)))" | check_instr_Pop: "check_instr Pop P h stk loc C\<^sub>0 M\<^sub>0 pc frs = (0 < length stk)" | check_instr_IAdd: "check_instr IAdd P h stk loc C\<^sub>0 M\<^sub>0 pc frs = (1 < length stk \ is_Intg (hd stk) \ is_Intg (hd (tl stk)))" | check_instr_IfFalse: "check_instr (IfFalse b) P h stk loc C\<^sub>0 M\<^sub>0 pc frs = (0 < length stk \ is_Bool (hd stk) \ 0 \ int pc+b)" | check_instr_CmpEq: "check_instr CmpEq P h stk loc C\<^sub>0 M\<^sub>0 pc frs = (1 < length stk)" | check_instr_Goto: "check_instr (Goto b) P h stk loc C\<^sub>0 M\<^sub>0 pc frs = (0 \ int pc+b)" | check_instr_Throw: "check_instr Throw P h stk loc C\<^sub>0 M\<^sub>0 pc frs = (0 < length stk \ is_Ref (hd stk))" definition check :: "jvm_prog \ jvm_state \ bool" where "check P \ = (let (xcpt, h, frs) = \ in (case frs of [] \ True | (stk,loc,C,M,pc)#frs' \ P \ C has M \ (let (C',Ts,T,mxs,mxl\<^sub>0,ins,xt) = method P C M; i = ins!pc in pc < size ins \ size stk \ mxs \ check_instr i P h stk loc C M pc frs')))" definition exec_d :: "jvm_prog \ jvm_state \ jvm_state option type_error" where "exec_d P \ = (if check P \ then Normal (exec (P, \)) else TypeError)" inductive_set exec_1_d :: "jvm_prog \ (jvm_state type_error \ jvm_state type_error) set" and exec_1_d' :: "jvm_prog \ jvm_state type_error \ jvm_state type_error \ bool" ("_ \ _ -jvmd\\<^sub>1 _" [61,61,61]60) for P :: jvm_prog where "P \ \ -jvmd\\<^sub>1 \' \ (\,\') \ exec_1_d P" | exec_1_d_ErrorI: "exec_d P \ = TypeError \ P \ Normal \ -jvmd\\<^sub>1 TypeError" | exec_1_d_NormalI: "exec_d P \ = Normal (Some \') \ P \ Normal \ -jvmd\\<^sub>1 Normal \'" \ \reflexive transitive closure:\ definition exec_all_d :: "jvm_prog \ jvm_state type_error \ jvm_state type_error \ bool" ("_ \ _ -jvmd\ _" [61,61,61]60) where exec_all_d_def1: "P \ \ -jvmd\ \' \ (\,\') \ (exec_1_d P)\<^sup>*" notation (ASCII) "exec_all_d" ("_ |- _ -jvmd-> _" [61,61,61]60) lemma exec_1_d_eq: "exec_1_d P = {(s,t). \\. s = Normal \ \ t = TypeError \ exec_d P \ = TypeError} \ {(s,t). \\ \'. s = Normal \ \ t = Normal \' \ exec_d P \ = Normal (Some \')}" by (auto elim!: exec_1_d.cases intro!: exec_1_d.intros) declare split_paired_All [simp del] declare split_paired_Ex [simp del] lemma if_neq [dest!]: "(if P then A else B) \ B \ P" by (cases P, auto) lemma exec_d_no_errorI [intro]: "check P \ \ exec_d P \ \ TypeError" by (unfold exec_d_def) simp theorem no_type_error_commutes: "exec_d P \ \ TypeError \ exec_d P \ = Normal (exec (P, \))" by (unfold exec_d_def, auto) lemma defensive_imp_aggressive: "P \ (Normal \) -jvmd\ (Normal \') \ P \ \ -jvm\ \'" (*<*) proof - have "\x y. P \ x -jvmd\ y \ \\ \'. x = Normal \ \ y = Normal \' \ P \ \ -jvm\ \'" - apply (unfold exec_all_d_def1) - apply (erule rtrancl_induct) - apply (simp add: exec_all_def) - apply (fold exec_all_d_def1) - apply simp - apply (intro allI impI) - apply (erule exec_1_d.cases, simp) - apply (simp add: exec_all_def exec_d_def split: type_error.splits if_split_asm) - apply (rule rtrancl_trans, assumption) - apply blast - done + proof - + fix x y assume xy: "P \ x -jvmd\ y" + then have "(x, y) \ (exec_1_d P)\<^sup>*" by (unfold exec_all_d_def1) + then show "\\ \'. x = Normal \ \ y = Normal \' \ P \ \ -jvm\ \'" + proof(induct rule: rtrancl_induct) + case base + then show ?case by (simp add: exec_all_def) + next + case (step y' z') + show ?case proof(induct rule: exec_1_d.cases[OF step.hyps(2)]) + case (2 \ \') + then have "(\, \') \ {(\, \'). exec (P, \) = \\'\}\<^sup>*" using step + by(fastforce simp: exec_d_def split: type_error.splits if_split_asm) + then show ?case using step 2 by (auto simp: exec_all_def) + qed simp + qed + qed moreover assume "P \ (Normal \) -jvmd\ (Normal \')" ultimately show "P \ \ -jvm\ \'" by blast qed (*>*) end diff --git a/thys/Jinja/JVM/JVMExec.thy b/thys/Jinja/JVM/JVMExec.thy --- a/thys/Jinja/JVM/JVMExec.thy +++ b/thys/Jinja/JVM/JVMExec.thy @@ -1,111 +1,114 @@ (* Title: HOL/MicroJava/JVM/JVMExec.thy Author: Cornelia Pusch, Gerwin Klein Copyright 1999 Technische Universitaet Muenchen *) section \Program Execution in the JVM\ theory JVMExec imports JVMExecInstr JVMExceptions begin abbreviation instrs_of :: "jvm_prog \ cname \ mname \ instr list" where "instrs_of P C M == fst(snd(snd(snd(snd(snd(method P C M))))))" fun exec :: "jvm_prog \ jvm_state => jvm_state option" where \ \single step execution\ "exec (P, xp, h, []) = None" | "exec (P, None, h, (stk,loc,C,M,pc)#frs) = (let i = instrs_of P C M ! pc; (xcpt', h', frs') = exec_instr i P h stk loc C M pc frs in Some(case xcpt' of None \ (None,h',frs') | Some a \ find_handler P a h ((stk,loc,C,M,pc)#frs)))" | "exec (P, Some xa, h, frs) = None" \ \relational view\ inductive_set exec_1 :: "jvm_prog \ (jvm_state \ jvm_state) set" and exec_1' :: "jvm_prog \ jvm_state \ jvm_state \ bool" ("_ \/ _ -jvm\\<^sub>1/ _" [61,61,61] 60) for P :: jvm_prog where "P \ \ -jvm\\<^sub>1 \' \ (\,\') \ exec_1 P" | exec_1I: "exec (P,\) = Some \' \ P \ \ -jvm\\<^sub>1 \'" \ \reflexive transitive closure:\ definition exec_all :: "jvm_prog \ jvm_state \ jvm_state \ bool" ("(_ \/ _ -jvm\/ _)" [61,61,61]60) where (* FIXME exec_all \ exec_star, also in Def.JVM *) exec_all_def1: "P \ \ -jvm\ \' \ (\,\') \ (exec_1 P)\<^sup>*" notation (ASCII) exec_all ("_ |-/ _ -jvm->/ _" [61,61,61]60) lemma exec_1_eq: "exec_1 P = {(\,\'). exec (P,\) = Some \'}" (*<*)by (auto intro: exec_1I elim: exec_1.cases)(*>*) lemma exec_1_iff: "P \ \ -jvm\\<^sub>1 \' = (exec (P,\) = Some \')" (*<*)by (simp add: exec_1_eq)(*>*) lemma exec_all_def: "P \ \ -jvm\ \' = ((\,\') \ {(\,\'). exec (P,\) = Some \'}\<^sup>*)" (*<*)by (simp add: exec_all_def1 exec_1_eq)(*>*) lemma jvm_refl[iff]: "P \ \ -jvm\ \" (*<*)by(simp add: exec_all_def)(*>*) lemma jvm_trans[trans]: "\ P \ \ -jvm\ \'; P \ \' -jvm\ \'' \ \ P \ \ -jvm\ \''" (*<*)by(simp add: exec_all_def)(*>*) lemma jvm_one_step1[trans]: "\ P \ \ -jvm\\<^sub>1 \'; P \ \' -jvm\ \'' \ \ P \ \ -jvm\ \''" (*<*) by (simp add: exec_all_def1) (*>*) lemma jvm_one_step2[trans]: "\ P \ \ -jvm\ \'; P \ \' -jvm\\<^sub>1 \'' \ \ P \ \ -jvm\ \''" (*<*) by (simp add: exec_all_def1) (*>*) lemma exec_all_conf: "\ P \ \ -jvm\ \'; P \ \ -jvm\ \'' \ \ P \ \' -jvm\ \'' \ P \ \'' -jvm\ \'" (*<*)by(simp add: exec_all_def single_valued_def single_valued_confluent)(*>*) lemma exec_all_finalD: "P \ (x, h, []) -jvm\ \ \ \ = (x, h, [])" (*<*) -apply(simp only: exec_all_def) -apply(erule converse_rtranclE) - apply simp -apply simp -done +proof - + assume "P \ (x, h, []) -jvm\ \" + then have "((x, h, []), \) \ {(\, \'). exec (P, \) = \\'\}\<^sup>*" + by(simp only: exec_all_def) + then show ?thesis proof(rule converse_rtranclE) qed simp+ +qed (*>*) lemma exec_all_deterministic: "\ P \ \ -jvm\ (x,h,[]); P \ \ -jvm\ \' \ \ P \ \' -jvm\ (x,h,[])" (*<*) -apply(drule (1) exec_all_conf) -apply(blast dest!: exec_all_finalD) -done +proof - + assume assms: "P \ \ -jvm\ (x,h,[])" "P \ \ -jvm\ \'" + show ?thesis using exec_all_conf[OF assms] + by(blast dest!: exec_all_finalD) +qed (*>*) text \ The start configuration of the JVM: in the start heap, we call a method \m\ of class \C\ in program \P\. The \this\ pointer of the frame is set to \Null\ to simulate a static method invokation. \ definition start_state :: "jvm_prog \ cname \ mname \ jvm_state" where "start_state P C M = (let (D,Ts,T,mxs,mxl\<^sub>0,b) = method P C M in (None, start_heap P, [([], Null # replicate mxl\<^sub>0 undefined, C, M, 0)]))" end diff --git a/thys/JinjaDCI/BV/BVConform.thy b/thys/JinjaDCI/BV/BVConform.thy --- a/thys/JinjaDCI/BV/BVConform.thy +++ b/thys/JinjaDCI/BV/BVConform.thy @@ -1,431 +1,433 @@ (* Title: JinjaDCI/BV/BVConform.thy Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory BV/BVConform.thy by Cornelia Pusch and Gerwin Klein The invariant for the type safety proof. *) section \ BV Type Safety Invariant \ theory BVConform imports BVSpec "../JVM/JVMExec" "../Common/Conform" begin subsection \ @{text "correct_state"} definitions \ definition confT :: "'c prog \ heap \ val \ ty err \ bool" ("_,_ \ _ :\\<^sub>\ _" [51,51,51,51] 50) where "P,h \ v :\\<^sub>\ E \ case E of Err \ True | OK T \ P,h \ v :\ T" notation (ASCII) confT ("_,_ |- _ :<=T _" [51,51,51,51] 50) abbreviation confTs :: "'c prog \ heap \ val list \ ty\<^sub>l \ bool" ("_,_ \ _ [:\\<^sub>\] _" [51,51,51,51] 50) where "P,h \ vs [:\\<^sub>\] Ts \ list_all2 (confT P h) vs Ts" notation (ASCII) confTs ("_,_ |- _ [:<=T] _" [51,51,51,51] 50) fun Called_context :: "jvm_prog \ cname \ instr \ bool" where "Called_context P C\<^sub>0 (New C') = (C\<^sub>0=C')" | "Called_context P C\<^sub>0 (Getstatic C F D) = ((C\<^sub>0=D) \ (\t. P \ C has F,Static:t in D))" | "Called_context P C\<^sub>0 (Putstatic C F D) = ((C\<^sub>0=D) \ (\t. P \ C has F,Static:t in D))" | "Called_context P C\<^sub>0 (Invokestatic C M n) = (\Ts T m D. (C\<^sub>0=D) \ P \ C sees M,Static:Ts \ T = m in D)" | "Called_context P _ _ = False" abbreviation Called_set :: "instr set" where "Called_set \ {i. \C. i = New C} \ {i. \C M n. i = Invokestatic C M n} \ {i. \C F D. i = Getstatic C F D} \ {i. \C F D. i = Putstatic C F D}" lemma Called_context_Called_set: "Called_context P D i \ i \ Called_set" by(cases i, auto) fun valid_ics :: "jvm_prog \ heap \ sheap \ cname \ mname \ pc \ init_call_status \ bool" ("_,_,_ \\<^sub>i _" [51,51,51,51] 50) where "valid_ics P h sh (C,M,pc,Calling C' Cs) = (let ins = instrs_of P C M in Called_context P (last (C'#Cs)) (ins!pc) \ is_class P C')" | "valid_ics P h sh (C,M,pc,Throwing Cs a) =(let ins = instrs_of P C M in \C1. Called_context P C1 (ins!pc) \ (\obj. h a = Some obj))" | "valid_ics P h sh (C,M,pc,Called Cs) = (let ins = instrs_of P C M in \C1 sobj. Called_context P C1 (ins!pc) \ sh C1 = Some sobj)" | "valid_ics P _ _ _ = True" definition conf_f :: "jvm_prog \ heap \ sheap \ ty\<^sub>i \ bytecode \ frame \ bool" where "conf_f P h sh \ \(ST,LT) is (stk,loc,C,M,pc,ics). P,h \ stk [:\] ST \ P,h \ loc [:\\<^sub>\] LT \ pc < size is \ P,h,sh \\<^sub>i (C,M,pc,ics)" lemma conf_f_def2: "conf_f P h sh (ST,LT) is (stk,loc,C,M,pc,ics) \ P,h \ stk [:\] ST \ P,h \ loc [:\\<^sub>\] LT \ pc < size is \ P,h,sh \\<^sub>i (C,M,pc,ics)" by (simp add: conf_f_def) primrec conf_fs :: "[jvm_prog,heap,sheap,ty\<^sub>P,cname,mname,nat,ty,frame list] \ bool" where "conf_fs P h sh \ C\<^sub>0 M\<^sub>0 n\<^sub>0 T\<^sub>0 [] = True" | "conf_fs P h sh \ C\<^sub>0 M\<^sub>0 n\<^sub>0 T\<^sub>0 (f#frs) = (let (stk,loc,C,M,pc,ics) = f in (\ST LT b Ts T mxs mxl\<^sub>0 is xt. \ C M ! pc = Some (ST,LT) \ (P \ C sees M,b:Ts \ T = (mxs,mxl\<^sub>0,is,xt) in C) \ ((\D Ts' T' m D'. M\<^sub>0 \ clinit \ ics = No_ics \ is!pc = Invoke M\<^sub>0 n\<^sub>0 \ ST!n\<^sub>0 = Class D \ P \ D sees M\<^sub>0,NonStatic:Ts' \ T' = m in D' \ P \ C\<^sub>0 \\<^sup>* D' \ P \ T\<^sub>0 \ T') \ (\D Ts' T' m. M\<^sub>0 \ clinit \ ics = No_ics \ is!pc = Invokestatic D M\<^sub>0 n\<^sub>0 \ P \ D sees M\<^sub>0,Static:Ts' \ T' = m in C\<^sub>0 \ P \ T\<^sub>0 \ T') \ (M\<^sub>0 = clinit \ (\Cs. ics = Called Cs))) \ conf_f P h sh (ST, LT) is f \ conf_fs P h sh \ C M (size Ts) T frs))" fun ics_classes :: "init_call_status \ cname list" where "ics_classes (Calling C Cs) = Cs" | "ics_classes (Throwing Cs a) = Cs" | "ics_classes (Called Cs) = Cs" | "ics_classes _ = []" fun frame_clinit_classes :: "frame \ cname list" where "frame_clinit_classes (stk,loc,C,M,pc,ics) = (if M=clinit then [C] else []) @ ics_classes ics" abbreviation clinit_classes :: "frame list \ cname list" where "clinit_classes frs \ concat (map frame_clinit_classes frs)" definition distinct_clinit :: "frame list \ bool" where "distinct_clinit frs \ distinct (clinit_classes frs)" definition conf_clinit :: "jvm_prog \ sheap \ frame list \ bool" where "conf_clinit P sh frs \ distinct_clinit frs \ (\C \ set(clinit_classes frs). is_class P C \ (\sfs. sh C = Some(sfs, Processing)))" (*************************) definition correct_state :: "[jvm_prog,ty\<^sub>P,jvm_state] \ bool" ("_,_ \ _ \" [61,0,0] 61) where "correct_state P \ \ \(xp,h,frs,sh). case xp of None \ (case frs of [] \ True | (f#fs) \ P\ h\ \ P,h\\<^sub>s sh\ \ conf_clinit P sh frs \ (let (stk,loc,C,M,pc,ics) = f in \b Ts T mxs mxl\<^sub>0 is xt \. (P \ C sees M,b:Ts\T = (mxs,mxl\<^sub>0,is,xt) in C) \ \ C M ! pc = Some \ \ conf_f P h sh \ is f \ conf_fs P h sh \ C M (size Ts) T fs)) | Some x \ frs = []" notation correct_state ("_,_ |- _ [ok]" [61,0,0] 61) subsection \ Values and @{text "\"} \ lemma confT_Err [iff]: "P,h \ x :\\<^sub>\ Err" by (simp add: confT_def) lemma confT_OK [iff]: "P,h \ x :\\<^sub>\ OK T = (P,h \ x :\ T)" by (simp add: confT_def) lemma confT_cases: "P,h \ x :\\<^sub>\ X = (X = Err \ (\T. X = OK T \ P,h \ x :\ T))" by (cases X) auto lemma confT_hext [intro?, trans]: "\ P,h \ x :\\<^sub>\ T; h \ h' \ \ P,h' \ x :\\<^sub>\ T" by (cases T) (blast intro: conf_hext)+ lemma confT_widen [intro?, trans]: "\ P,h \ x :\\<^sub>\ T; P \ T \\<^sub>\ T' \ \ P,h \ x :\\<^sub>\ T'" by (cases T', auto intro: conf_widen) subsection \ Stack and Registers \ lemmas confTs_Cons1 [iff] = list_all2_Cons1 [of "confT P h"] for P h lemma confTs_confT_sup: - "\ P,h \ loc [:\\<^sub>\] LT; n < size LT; LT!n = OK T; P \ T \ T' \ - \ P,h \ (loc!n) :\ T'" +assumes confTs: "P,h \ loc [:\\<^sub>\] LT" and n: "n < size LT" and + LTn: "LT!n = OK T" and subtype: "P \ T \ T'" +shows "P,h \ (loc!n) :\ T'" (*<*) - apply (frule list_all2_lengthD) - apply (drule list_all2_nthD, simp) - apply simp - apply (erule conf_widen, assumption+) - done +proof - + have len: "n < length loc" using list_all2_lengthD[OF confTs] n + by simp + show ?thesis + using list_all2_nthD[OF confTs len] conf_widen[OF _ subtype] LTn + by simp +qed (*>*) lemma confTs_hext [intro?]: "P,h \ loc [:\\<^sub>\] LT \ h \ h' \ P,h' \ loc [:\\<^sub>\] LT" by (fast elim: list_all2_mono confT_hext) lemma confTs_widen [intro?, trans]: "P,h \ loc [:\\<^sub>\] LT \ P \ LT [\\<^sub>\] LT' \ P,h \ loc [:\\<^sub>\] LT'" by (rule list_all2_trans, rule confT_widen) lemma confTs_map [iff]: "\vs. (P,h \ vs [:\\<^sub>\] map OK Ts) = (P,h \ vs [:\] Ts)" by (induct Ts) (auto simp: list_all2_Cons2) lemma reg_widen_Err [iff]: "\LT. (P \ replicate n Err [\\<^sub>\] LT) = (LT = replicate n Err)" by (induct n) (auto simp: list_all2_Cons1) lemma confTs_Err [iff]: "P,h \ replicate n v [:\\<^sub>\] replicate n Err" by (induct n) auto subsection \ valid @{text "init_call_status"} \ lemma valid_ics_shupd: assumes "P,h,sh \\<^sub>i (C, M, pc, ics)" and "distinct (C'#ics_classes ics)" shows "P,h,sh(C' \ (sfs, i')) \\<^sub>i (C, M, pc, ics)" using assms by(cases ics; clarsimp simp: fun_upd_apply) fastforce subsection \ correct-frame \ lemma conf_f_Throwing: assumes "conf_f P h sh (ST, LT) is (stk, loc, C, M, pc, Called Cs)" and "is_class P C'" and "h xcp = Some obj" and "sh C' = Some(sfs,Processing)" shows "conf_f P h sh (ST, LT) is (stk, loc, C, M, pc, Throwing (C' # Cs) xcp)" using assms by(auto simp: conf_f_def2) lemma conf_f_shupd: assumes "conf_f P h sh (ST,LT) ins f" and "i = Processing \ (distinct (C#ics_classes (ics_of f)) \ (curr_method f = clinit \ C \ curr_class f))" shows "conf_f P h (sh(C \ (sfs, i))) (ST,LT) ins f" using assms by(cases f, cases "ics_of f"; clarsimp simp: conf_f_def2 fun_upd_apply) fastforce+ lemma conf_f_shupd': assumes "conf_f P h sh (ST,LT) ins f" and "sh C = Some(sfs,i)" shows "conf_f P h (sh(C \ (sfs', i))) (ST,LT) ins f" using assms by(cases f, cases "ics_of f"; clarsimp simp: conf_f_def2 fun_upd_apply) fastforce+ subsection \ correct-frames \ lemmas [simp del] = fun_upd_apply lemma conf_fs_hext: "\C M n T\<^sub>r. \ conf_fs P h sh \ C M n T\<^sub>r frs; h \ h' \ \ conf_fs P h' sh \ C M n T\<^sub>r frs" (*<*) -apply (induct frs) - apply simp -apply clarify -apply (simp (no_asm_use)) -apply clarify -apply (unfold conf_f_def) -apply (simp (no_asm_use)) -apply clarify -apply (fastforce elim!: confs_hext confTs_hext) -done +proof(induct frs) + case (Cons fr frs) + obtain stk ls C M pc ics where fr: "fr = (stk, ls, C, M, pc, ics)" by(cases fr) simp + moreover obtain ST LT where \: "\ C M ! pc = \(ST, LT)\" and + ST: "P,h \ stk [:\] ST" and LT: "P,h \ ls [:\\<^sub>\] LT" + using Cons.prems(1) fr by(auto simp: conf_f_def) + ultimately show ?case using Cons confs_hext[OF ST Cons(3)] confTs_hext[OF LT Cons(3)] + by (fastforce simp: conf_f_def) +qed simp (*>*) lemma conf_fs_shupd: assumes "conf_fs P h sh \ C\<^sub>0 M n T frs" and dist: "distinct (C#clinit_classes frs)" shows "conf_fs P h (sh(C \ (sfs, i))) \ C\<^sub>0 M n T frs" using assms proof(induct frs arbitrary: C\<^sub>0 C M n T) case (Cons f' frs') then obtain stk' loc' C' M' pc' ics' where f': "f' = (stk',loc',C',M',pc',ics')" by(cases f') with assms Cons obtain ST LT b Ts T1 mxs mxl\<^sub>0 ins xt where ty: "\ C' M' ! pc' = Some (ST,LT)" and meth: "P \ C' sees M',b:Ts \ T1 = (mxs,mxl\<^sub>0,ins,xt) in C'" and conf: "conf_f P h sh (ST, LT) ins f'" and confs: "conf_fs P h sh \ C' M' (size Ts) T1 frs'" by clarsimp from f' Cons.prems(2) have "distinct (C#ics_classes (ics_of f')) \ (curr_method f' = clinit \ C \ curr_class f')" by fastforce with conf_f_shupd[where C=C, OF conf] have conf': "conf_f P h (sh(C \ (sfs, i))) (ST, LT) ins f'" by simp from Cons.prems(2) have dist': "distinct (C # clinit_classes frs')" by(auto simp: distinct_length_2_or_more) from Cons.hyps[OF confs dist'] have confs': "conf_fs P h (sh(C \ (sfs, i))) \ C' M' (length Ts) T1 frs'" by simp from conf' confs' ty meth f' Cons.prems show ?case by(fastforce dest: sees_method_fun) qed(simp) lemma conf_fs_shupd': assumes "conf_fs P h sh \ C\<^sub>0 M n T frs" and shC: "sh C = Some(sfs,i)" shows "conf_fs P h (sh(C \ (sfs', i))) \ C\<^sub>0 M n T frs" using assms proof(induct frs arbitrary: C\<^sub>0 C M n T sfs i sfs') case (Cons f' frs') then obtain stk' loc' C' M' pc' ics' where f': "f' = (stk',loc',C',M',pc',ics')" by(cases f') with assms Cons obtain ST LT b Ts T1 mxs mxl\<^sub>0 ins xt where ty: "\ C' M' ! pc' = Some (ST,LT)" and meth: "P \ C' sees M',b:Ts \ T1 = (mxs,mxl\<^sub>0,ins,xt) in C'" and conf: "conf_f P h sh (ST, LT) ins f'" and confs: "conf_fs P h sh \ C' M' (size Ts) T1 frs'" and shC': "sh C = Some(sfs,i)" by clarsimp have conf': "conf_f P h (sh(C \ (sfs', i))) (ST, LT) ins f'" by(rule conf_f_shupd'[OF conf shC']) from Cons.hyps[OF confs shC'] have confs': "conf_fs P h (sh(C \ (sfs', i))) \ C' M' (length Ts) T1 frs'" by simp from conf' confs' ty meth f' Cons.prems show ?case by(fastforce dest: sees_method_fun) qed(simp) subsection \ correctness wrt @{term clinit} use \ lemma conf_clinit_Cons: assumes "conf_clinit P sh (f#frs)" shows "conf_clinit P sh frs" proof - from assms have dist: "distinct_clinit (f#frs)" by(cases "curr_method f = clinit", auto simp: conf_clinit_def) then have dist': "distinct_clinit frs" by(simp add: distinct_clinit_def) with assms show ?thesis by(cases frs; fastforce simp: conf_clinit_def) qed lemma conf_clinit_Cons_Cons: "conf_clinit P sh (f'#f#frs) \ conf_clinit P sh (f'#frs)" by(auto simp: conf_clinit_def distinct_clinit_def) lemma conf_clinit_diff: assumes "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" shows "conf_clinit P sh ((stk',loc',C,M,pc',ics)#frs)" using assms by(cases "M = clinit", simp_all add: conf_clinit_def distinct_clinit_def) lemma conf_clinit_diff': assumes "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" shows "conf_clinit P sh ((stk',loc',C,M,pc',No_ics)#frs)" using assms by(cases "M = clinit", simp_all add: conf_clinit_def distinct_clinit_def) lemma conf_clinit_Called_Throwing: "conf_clinit P sh ((stk', loc', C', clinit, pc', ics') # (stk, loc, C, M, pc, Called Cs) # fs) \ conf_clinit P sh ((stk, loc, C, M, pc, Throwing (C' # Cs) xcp) # fs)" by(simp add: conf_clinit_def distinct_clinit_def) lemma conf_clinit_Throwing: "conf_clinit P sh ((stk, loc, C, M, pc, Throwing (C'#Cs) xcp) # fs) \ conf_clinit P sh ((stk, loc, C, M, pc, Throwing Cs xcp) # fs)" by(simp add: conf_clinit_def distinct_clinit_def) lemma conf_clinit_Called: "\ conf_clinit P sh ((stk, loc, C, M, pc, Called (C'#Cs)) # frs); P \ C' sees clinit,Static: [] \ Void=(mxs',mxl',ins',xt') in C' \ \ conf_clinit P sh (create_init_frame P C' # (stk, loc, C, M, pc, Called Cs) # frs)" by(simp add: conf_clinit_def distinct_clinit_def) lemma conf_clinit_Cons_nclinit: assumes "conf_clinit P sh frs" and nclinit: "M \ clinit" shows "conf_clinit P sh ((stk, loc, C, M, pc, No_ics) # frs)" proof - from nclinit have "clinit_classes ((stk, loc, C, M, pc, No_ics) # frs) = clinit_classes frs" by simp with assms show ?thesis by(simp add: conf_clinit_def distinct_clinit_def) qed lemma conf_clinit_Invoke: assumes "conf_clinit P sh ((stk, loc, C, M, pc, ics) # frs)" and "M' \ clinit" shows "conf_clinit P sh ((stk', loc', C', M', pc', No_ics) # (stk, loc, C, M, pc, No_ics) # frs)" using assms conf_clinit_Cons_nclinit conf_clinit_diff' by auto lemma conf_clinit_nProc_dist: assumes "conf_clinit P sh frs" and "\sfs. sh C \ Some(sfs,Processing)" shows "distinct (C # clinit_classes frs)" using assms by(auto simp: conf_clinit_def distinct_clinit_def) lemma conf_clinit_shupd: assumes "conf_clinit P sh frs" and dist: "distinct (C#clinit_classes frs)" shows "conf_clinit P (sh(C \ (sfs, i))) frs" using assms by(simp add: conf_clinit_def fun_upd_apply) lemma conf_clinit_shupd': assumes "conf_clinit P sh frs" and "sh C = Some(sfs,i)" shows "conf_clinit P (sh(C \ (sfs', i))) frs" using assms by(fastforce simp: conf_clinit_def fun_upd_apply) lemma conf_clinit_shupd_Called: assumes "conf_clinit P sh ((stk,loc,C,M,pc,Calling C' Cs)#frs)" and dist: "distinct (C'#clinit_classes ((stk,loc,C,M,pc,Calling C' Cs)#frs))" and cls: "is_class P C'" shows "conf_clinit P (sh(C' \ (sfs, Processing))) ((stk,loc,C,M,pc,Called (C'#Cs))#frs)" using assms by(clarsimp simp: conf_clinit_def fun_upd_apply distinct_clinit_def) lemma conf_clinit_shupd_Calling: assumes "conf_clinit P sh ((stk,loc,C,M,pc,Calling C' Cs)#frs)" and dist: "distinct (C'#clinit_classes ((stk,loc,C,M,pc,Calling C' Cs)#frs))" and cls: "is_class P C'" shows "conf_clinit P (sh(C' \ (sfs, Processing))) ((stk,loc,C,M,pc,Calling (fst(the(class P C'))) (C'#Cs))#frs)" using assms by(clarsimp simp: conf_clinit_def fun_upd_apply distinct_clinit_def) subsection \ correct state \ lemma correct_state_Cons: assumes cr: "P,\ |- (xp,h,f#frs,sh) [ok]" shows "P,\ |- (xp,h,frs,sh) [ok]" proof - from cr have dist: "conf_clinit P sh (f#frs)" by(simp add: correct_state_def) then have "conf_clinit P sh frs" by(rule conf_clinit_Cons) with cr show ?thesis by(cases frs; fastforce simp: correct_state_def) qed lemma correct_state_shupd: assumes cs: "P,\ |- (xp,h,frs,sh) [ok]" and shC: "sh C = Some(sfs,i)" and dist: "distinct (C#clinit_classes frs)" shows "P,\ |- (xp,h,frs,sh(C \ (sfs, i'))) [ok]" using assms proof(cases xp) case None with assms show ?thesis proof(cases frs) case (Cons f' frs') let ?sh = "sh(C \ (sfs, i'))" obtain stk' loc' C' M' pc' ics' where f': "f' = (stk',loc',C',M',pc',ics')" by(cases f') with cs Cons None obtain b Ts T mxs mxl\<^sub>0 ins xt ST LT where meth: "P \ C' sees M',b:Ts\T = (mxs,mxl\<^sub>0,ins,xt) in C'" and ty: "\ C' M' ! pc' = Some (ST,LT)" and conf: "conf_f P h sh (ST,LT) ins f'" and confs: "conf_fs P h sh \ C' M' (size Ts) T frs'" and confc: "conf_clinit P sh frs" and h_ok: "P\ h\" and sh_ok: "P,h \\<^sub>s sh \" by(auto simp: correct_state_def) from Cons dist have dist': "distinct (C#clinit_classes frs')" by(auto simp: distinct_length_2_or_more) from shconf_upd_obj[OF sh_ok shconfD[OF sh_ok shC]] have sh_ok': "P,h \\<^sub>s ?sh \" by simp from conf f' valid_ics_shupd Cons dist have conf': "conf_f P h ?sh (ST,LT) ins f'" by(auto simp: conf_f_def2 fun_upd_apply) have confs': "conf_fs P h ?sh \ C' M' (size Ts) T frs'" by(rule conf_fs_shupd[OF confs dist']) have confc': "conf_clinit P ?sh frs" by(rule conf_clinit_shupd[OF confc dist]) with h_ok sh_ok' meth ty conf' confs' f' Cons None show ?thesis by(fastforce simp: correct_state_def) qed(simp add: correct_state_def) qed(simp add: correct_state_def) lemma correct_state_Throwing_ex: assumes correct: "P,\ \ (xp,h,(stk,loc,C,M,pc,ics)#frs,sh)\" shows "\Cs a. ics = Throwing Cs a \ \obj. h a = Some obj" using correct by(clarsimp simp: correct_state_def conf_f_def) end diff --git a/thys/JinjaDCI/BV/BVExec.thy b/thys/JinjaDCI/BV/BVExec.thy --- a/thys/JinjaDCI/BV/BVExec.thy +++ b/thys/JinjaDCI/BV/BVExec.thy @@ -1,238 +1,253 @@ (* Title: JinjaDCI/BV/BVExec.thy Author: Tobias Nipkow, Gerwin Klein, Susannah Mansky Copyright 2000 TUM, 2020 UIUC Based on the Jinja theory BV/BVExec.thy by Tobias Nipkow and Gerwin Klein *) section \ Kildall for the JVM \label{sec:JVM} \ theory BVExec imports Jinja.Abstract_BV TF_JVM begin definition kiljvm :: "jvm_prog \ nat \ nat \ ty \ instr list \ ex_table \ ty\<^sub>i' err list \ ty\<^sub>i' err list" where "kiljvm P mxs mxl T\<^sub>r is xt \ kildall (JVM_SemiType.le P mxs mxl) (JVM_SemiType.sup P mxs mxl) (exec P mxs T\<^sub>r xt is)" definition wt_kildall :: "jvm_prog \ cname \ staticb \ ty list \ ty \ nat \ nat \ instr list \ ex_table \ bool" where "wt_kildall P C' b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \ 0 < size is \ (let first = Some ([],(case b of Static \ [] | NonStatic \ [OK (Class C')]) @(map OK Ts)@(replicate mxl\<^sub>0 Err)); start = (OK first)#(replicate (size is - 1) (OK None)); result = kiljvm P mxs ((case b of Static \ 0 | NonStatic \ 1)+size Ts+mxl\<^sub>0) T\<^sub>r is xt start in \n < size is. result!n \ Err)" definition wf_jvm_prog\<^sub>k :: "jvm_prog \ bool" where "wf_jvm_prog\<^sub>k P \ wf_prog (\P C' (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). wt_kildall P C' b Ts T\<^sub>r mxs mxl\<^sub>0 is xt) P" theorem (in start_context) is_bcv_kiljvm: "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)" (*<*) - apply (insert wf) - apply (unfold kiljvm_def) - apply (fold r_def f_def step_def_exec) - apply (rule is_bcv_kildall) - apply simp apply (rule Semilat.intro) - apply (fold sl_def2) - apply (erule semilat_JVM) - apply simp - apply blast - apply (simp add: JVM_le_unfold) - apply (rule exec_pres_type) - apply (rule bounded_step) - apply (erule step_mono) - done +proof - + let ?n = "length is" + have "Semilat A r f" using semilat_JVM[OF wf] + by (simp add: Semilat.intro sl_def2) + moreover have "acc r" using wf by simp blast + moreover have "top r Err" by (simp add: JVM_le_unfold) + moreover have "pres_type step ?n A" by (rule exec_pres_type) + moreover have "bounded step ?n" by (rule bounded_step) + moreover have "mono r step ?n A" using step_mono[OF wf] by simp + ultimately have "is_bcv r Err step ?n A (kildall r f step)" + by(rule is_bcv_kildall) + moreover have kileq: "kiljvm P mxs mxl T\<^sub>r is xt = kildall r f step" + using f_def kiljvm_def r_def step_def_exec by blast + ultimately show ?thesis by simp +qed (*>*) (* FIXME: move? *) lemma subset_replicate [intro?]: "set (replicate n x) \ {x}" by (induct n) auto lemma in_set_replicate: assumes "x \ set (replicate n y)" shows "x = y" (*<*) proof - note assms also have "set (replicate n y) \ {y}" .. finally show ?thesis by simp qed (*>*) lemma (in start_context) start_in_A [intro?]: "0 < size is \ start \ list (size is) A" using Ts C (*<*) - apply (simp add: JVM_states_unfold) - apply (cases b; force intro!: listI list_appendI dest!: in_set_replicate) - done +proof(cases b) +qed (force simp: JVM_states_unfold intro!: listI list_appendI + dest!: in_set_replicate)+ (*>*) theorem (in start_context) wt_kil_correct: assumes wtk: "wt_kildall P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt" shows "\\s. wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" (*<*) proof - from wtk obtain res where result: "res = kiljvm P mxs mxl T\<^sub>r is xt start" and success: "\n < size is. res!n \ Err" and instrs: "0 < size is" by (unfold wt_kildall_def) simp have bcv: "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)" by (rule is_bcv_kiljvm) from instrs have "start \ list (size is) A" .. with bcv success result have "\ts\list (size is) A. start [\\<^sub>r] ts \ wt_step r Err step ts" by (unfold is_bcv_def) blast then obtain \s' where in_A: "\s' \ list (size is) A" and s: "start [\\<^sub>r] \s'" and w: "wt_step r Err step \s'" by blast hence wt_err_step: "wt_err_step (sup_state_opt P) step \s'" by (simp add: wt_err_step_def JVM_le_Err_conv) from in_A have l: "size \s' = size is" by simp moreover { from in_A have "check_types P mxs mxl \s'" by (simp add: check_types_def) also from w have "\x \ set \s'. x \ Err" by (auto simp add: wt_step_def all_set_conv_all_nth) hence [symmetric]: "map OK (map ok_val \s') = \s'" by (auto intro!: map_idI simp add: wt_step_def) finally have "check_types P mxs mxl (map OK (map ok_val \s'))" . } moreover { from s have "start!0 \\<^sub>r \s'!0" by (rule le_listD) simp moreover from instrs w l have "\s'!0 \ Err" by (unfold wt_step_def) simp then obtain \s0 where "\s'!0 = OK \s0" by auto ultimately have "wt_start P C b Ts mxl\<^sub>0 (map ok_val \s')" using l instrs by (unfold wt_start_def) (cases b; simp add: lesub_def JVM_le_Err_conv Err.le_def) } moreover from in_A have "set \s' \ A" by simp with wt_err_step bounded_step have "wt_app_eff (sup_state_opt P) app eff (map ok_val \s')" by (auto intro: wt_err_imp_wt_app_eff simp add: l) ultimately have "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (map ok_val \s')" using instrs by (simp add: wt_method_def2 check_types_def del: map_map) thus ?thesis by blast qed (*>*) theorem (in start_context) wt_kil_complete: assumes wtm: "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" shows "wt_kildall P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt" (*<*) proof - from wtm obtain instrs: "0 < size is" and length: "length \s = length is" and ck_type: "check_types P mxs mxl (map OK \s)" and wt_start: "wt_start P C b Ts mxl\<^sub>0 \s" and app_eff: "wt_app_eff (sup_state_opt P) app eff \s" by (simp add: wt_method_def2 check_types_def) from ck_type have in_A: "set (map OK \s) \ A" by (simp add: check_types_def) with app_eff in_A bounded_step have "wt_err_step (sup_state_opt P) (err_step (size \s) app eff) (map OK \s)" by - (erule wt_app_eff_imp_wt_err, auto simp add: exec_def length states_def) hence wt_err: "wt_err_step (sup_state_opt P) step (map OK \s)" by (simp add: length) have is_bcv: "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)" by (rule is_bcv_kiljvm) moreover from instrs have "start \ list (size is) A" .. moreover let ?\s = "map OK \s" have less_\s: "start [\\<^sub>r] ?\s" proof (rule le_listI) from length instrs show "length start = length (map OK \s)" by simp next fix n from wt_start have "P \ ok_val (start!0) \' \s!0" by (cases b; simp add: wt_start_def) moreover from instrs length have "0 < length \s" by simp ultimately have "start!0 \\<^sub>r ?\s!0" by (simp add: JVM_le_Err_conv lesub_def) moreover { fix n' have "OK None \\<^sub>r ?\s!n" by (auto simp add: JVM_le_Err_conv Err.le_def lesub_def split: err.splits) hence "\n = Suc n'; n < size start\ \ start!n \\<^sub>r ?\s!n" by simp } ultimately show "n < size start \ start!n \\<^sub>r ?\s!n" by (cases n, blast+) qed moreover from ck_type length have "?\s \ list (size is) A" by (auto intro!: listI simp add: check_types_def) moreover from wt_err have "wt_step r Err step ?\s" by (simp add: wt_err_step_def JVM_le_Err_conv) ultimately have "\p. p < size is \ kiljvm P mxs mxl T\<^sub>r is xt start ! p \ Err" by (unfold is_bcv_def) blast with instrs show "wt_kildall P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt" by (unfold wt_kildall_def) simp qed (*>*) theorem jvm_kildall_correct: "wf_jvm_prog\<^sub>k P = wf_jvm_prog P" (*<*) -proof +proof - let ?\ = "\C M. let (C,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)) = method P C M in SOME \s. wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" + let ?A = "\P C' (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). wt_kildall P C' b Ts T\<^sub>r mxs mxl\<^sub>0 is xt" + let ?B\<^sub>\ = "\\. (\P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). + wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M))" - \ \soundness\ - assume wt: "wf_jvm_prog\<^sub>k P" - hence "wf_jvm_prog\<^bsub>?\\<^esub> P" - apply (unfold wf_jvm_prog_phi_def wf_jvm_prog\<^sub>k_def) - apply (erule wf_prog_lift) - apply (auto dest!: start_context.wt_kil_correct [OF start_context.intro] - intro: someI) - apply (erule sees_method_is_class) - done - thus "wf_jvm_prog P" by (unfold wf_jvm_prog_def) fast -next - \ \completeness\ - assume wt: "wf_jvm_prog P" - thus "wf_jvm_prog\<^sub>k P" - apply (unfold wf_jvm_prog_def wf_jvm_prog_phi_def wf_jvm_prog\<^sub>k_def) - apply (clarify) - apply (erule wf_prog_lift) - apply (auto intro!: start_context.wt_kil_complete start_context.intro) - apply (erule sees_method_is_class) - done + show ?thesis proof(rule iffI) + \ \soundness\ + assume wt: "wf_jvm_prog\<^sub>k P" + then have "wf_prog ?A P" by(simp add: wf_jvm_prog\<^sub>k_def) + moreover { + fix wf_md C M b Ts Ca T m bd + assume "wf_prog wf_md P" and sees: "P \ Ca sees M, b : Ts\T = m in Ca" and + "set Ts \ types P" and "bd = (M, b, Ts, T, m)" and + "?A P Ca bd" + then have "(?B\<^sub>\ ?\) P Ca bd" using sees_method_is_class[OF sees] + by (auto dest!: start_context.wt_kil_correct [OF start_context.intro] + intro: someI) + } + ultimately have "wf_prog (?B\<^sub>\ ?\) P" by(rule wf_prog_lift) + then have "wf_jvm_prog\<^bsub>?\\<^esub> P" by (simp add: wf_jvm_prog_phi_def) + thus "wf_jvm_prog P" by (unfold wf_jvm_prog_def) fast + next + \ \completeness\ + assume wt: "wf_jvm_prog P" + then obtain \ where "wf_prog (?B\<^sub>\ \) P" by(clarsimp simp: wf_jvm_prog_def wf_jvm_prog_phi_def) + moreover { + fix wf_md C M b Ts Ca T m bd + assume "wf_prog wf_md P" and sees: "P \ Ca sees M, b : Ts\T = m in Ca" and + "set Ts \ types P" and "bd = (M, b, Ts, T, m)" and + "(?B\<^sub>\ \) P Ca bd" + then have "?A P Ca bd" using sees_method_is_class[OF sees] + by (auto intro!: start_context.wt_kil_complete start_context.intro) + } + ultimately have "wf_prog ?A P" by(rule wf_prog_lift) + thus "wf_jvm_prog\<^sub>k P" by (simp add: wf_jvm_prog\<^sub>k_def) + qed qed (*>*) end diff --git a/thys/JinjaDCI/BV/BVNoTypeError.thy b/thys/JinjaDCI/BV/BVNoTypeError.thy --- a/thys/JinjaDCI/BV/BVNoTypeError.thy +++ b/thys/JinjaDCI/BV/BVNoTypeError.thy @@ -1,394 +1,389 @@ (* Title: JinjaDCI/BV/BVNoTypeErrors.thy Author: Gerwin Klein, Susannah Mansky Copyright GPL Based on the Jinja theory BV/BVNoTypeErrors.thy by Gerwin Klein *) section \ Welltyped Programs produce no Type Errors \ theory BVNoTypeError imports "../JVM/JVMDefensive" BVSpecTypeSafe begin lemma has_methodI: "P \ C sees M,b:Ts\T = m in D \ P \ C has M,b" by (unfold has_method_def) blast text \ Some simple lemmas about the type testing functions of the defensive JVM: \ lemma typeof_NoneD [simp,dest]: "typeof v = Some x \ \is_Addr v" by (cases v) auto lemma is_Ref_def2: "is_Ref v = (v = Null \ (\a. v = Addr a))" by (cases v) (auto simp add: is_Ref_def) lemma [iff]: "is_Ref Null" by (simp add: is_Ref_def2) lemma is_RefI [intro, simp]: "P,h \ v :\ T \ is_refT T \ is_Ref v" (*<*) - apply (cases T) - apply (auto simp add: is_refT_def is_Ref_def dest: conf_ClassD) - done +proof(cases T) +qed (auto simp add: is_refT_def is_Ref_def dest: conf_ClassD) (*>*) lemma is_IntgI [intro, simp]: "P,h \ v :\ Integer \ is_Intg v" -(*<*) - apply (unfold conf_def) - apply auto - done -(*>*) +(*<*)by (unfold conf_def) auto(*>*) lemma is_BoolI [intro, simp]: "P,h \ v :\ Boolean \ is_Bool v" -(*<*) - apply (unfold conf_def) - apply auto - done -(*>*) +(*<*)by (unfold conf_def) auto(*>*) declare defs1 [simp del] lemma wt_jvm_prog_states_NonStatic: - "\ wf_jvm_prog\<^bsub>\\<^esub> P; P \ C sees M,NonStatic: Ts\T = (mxs, mxl, ins, et) in C; - \ C M ! pc = \; pc < size ins \ - \ OK \ \ states P mxs (1+size Ts+mxl)" +assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" + and mC: "P \ C sees M,NonStatic: Ts\T = (mxs, mxl, ins, et) in C" + and \: "\ C M ! pc = \" and pc: "pc < size ins" +shows "OK \ \ states P mxs (1+size Ts+mxl)" (*<*) - apply (unfold wf_jvm_prog_phi_def) - apply (drule (1) sees_wf_mdecl) - apply (simp add: wf_mdecl_def wt_method_def check_types_def) - apply (blast intro: nth_in) - done +proof - + let ?wf_md = "(\P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). + wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M))" + have wfmd: "wf_prog ?wf_md P" using wf + by (unfold wf_jvm_prog_phi_def) assumption + show ?thesis using sees_wf_mdecl[OF wfmd mC] \ pc + by (simp add: wf_mdecl_def wt_method_def check_types_def) + (blast intro: nth_in) +qed (*>*) lemma wt_jvm_prog_states_Static: - "\ wf_jvm_prog\<^bsub>\\<^esub> P; P \ C sees M,Static: Ts\T = (mxs, mxl, ins, et) in C; - \ C M ! pc = \; pc < size ins \ - \ OK \ \ states P mxs (size Ts+mxl)" +assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" + and mC: "P \ C sees M,Static: Ts\T = (mxs, mxl, ins, et) in C" + and \: "\ C M ! pc = \" and pc: "pc < size ins" +shows "OK \ \ states P mxs (size Ts+mxl)" (*<*) - apply (unfold wf_jvm_prog_phi_def) - apply (drule (1) sees_wf_mdecl) - apply (simp add: wf_mdecl_def wt_method_def check_types_def) - apply (blast intro: nth_in) - done +proof - + let ?wf_md = "(\P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). + wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M))" + have wfmd: "wf_prog ?wf_md P" using wf + by (unfold wf_jvm_prog_phi_def) assumption + show ?thesis using sees_wf_mdecl[OF wfmd mC] \ pc + by (simp add: wf_mdecl_def wt_method_def check_types_def) + (blast intro: nth_in) +qed (*>*) text \ The main theorem: welltyped programs do not produce type errors if they are started in a conformant state. \ theorem no_type_error: fixes \ :: jvm_state assumes welltyped: "wf_jvm_prog\<^bsub>\\<^esub> P" and conforms: "P,\ \ \ \" shows "exec_d P \ \ TypeError" (*<*) proof - from welltyped obtain mb where wf: "wf_prog mb P" by (fast dest: wt_jvm_progD) obtain xcp h frs sh where s [simp]: "\ = (xcp, h, frs, sh)" by (cases \) from conforms have "xcp \ None \ frs = [] \ check P \" by (unfold correct_state_def check_def) auto moreover { assume "\(xcp \ None \ frs = [])" then obtain stk reg C M pc ics frs' where xcp [simp]: "xcp = None" and frs [simp]: "frs = (stk,reg,C,M,pc,ics)#frs'" by (clarsimp simp add: neq_Nil_conv) from conforms obtain ST LT b Ts T mxs mxl ins xt where hconf: "P \ h \" and shconf: "P,h \\<^sub>s sh \" and meth: "P \ C sees M,b:Ts\T = (mxs, mxl, ins, xt) in C" and \: "\ C M ! pc = Some (ST,LT)" and frame: "conf_f P h sh (ST,LT) ins (stk,reg,C,M,pc,ics)" and frames: "conf_fs P h sh \ C M (size Ts) T frs'" by (fastforce simp add: correct_state_def dest: sees_method_fun) from frame obtain stk: "P,h \ stk [:\] ST" and reg: "P,h \ reg [:\\<^sub>\] LT" and pc: "pc < size ins" by (simp add: conf_f_def) from welltyped meth \ pc have "OK (Some (ST, LT)) \ states P mxs (1+size Ts+mxl) \ OK (Some (ST, LT)) \ states P mxs (size Ts+mxl)" by (cases b, auto dest: wt_jvm_prog_states_NonStatic wt_jvm_prog_states_Static) hence "size ST \ mxs" by (auto simp add: JVM_states_unfold) with stk have mxs: "size stk \ mxs" by (auto dest: list_all2_lengthD) from welltyped meth pc have "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" by (rule wt_jvm_prog_impl_wt_instr) hence app\<^sub>0: "app (ins!pc) P mxs T pc (size ins) xt (\ C M!pc) " by (simp add: wt_instr_def) with \ have eff: "\(pc',s')\set (eff (ins ! pc) P pc xt (\ C M ! pc)). pc' < size ins" by (unfold app_def) simp from app\<^sub>0 \ have app: "xcpt_app (ins!pc) P pc mxs xt (ST,LT) \ app\<^sub>i (ins!pc, P, pc, mxs, T, (ST,LT))" by (clarsimp simp add: app_def) with eff stk reg have "check_instr (ins!pc) P h stk reg C M pc frs' sh" proof (cases "ins!pc") case (Getfield F C) with app stk reg \ obtain v vT stk' where field: "P \ C sees F,NonStatic:vT in C" and stk: "stk = v # stk'" and conf: "P,h \ v :\ Class C" by auto from conf have is_Ref: "is_Ref v" by auto moreover { assume "v \ Null" with conf field is_Ref wf have "\D vs. h (the_Addr v) = Some (D,vs) \ P \ D \\<^sup>* C" by (auto dest!: non_npD) } - ultimately show ?thesis using Getfield field stk hconf - apply clarsimp - apply (rule conjI, fastforce) - apply clarsimp - apply (drule has_visible_field) - apply (drule (1) has_field_mono) - apply (drule (1) hconfD) - apply (unfold oconf_def has_field_def) - apply clarsimp - apply (fastforce dest: has_fields_fun) - done + ultimately show ?thesis using Getfield field stk + has_field_mono[OF has_visible_field[OF field]] hconfD[OF hconf] + by (unfold oconf_def has_field_def) (fastforce dest: has_fields_fun) next case (Getstatic C F D) with app stk reg \ obtain vT where field: "P \ C sees F,Static:vT in D" by auto - then show ?thesis using Getstatic field stk shconf - apply clarsimp - apply (rule conjI, fastforce) - apply clarsimp - apply (drule has_visible_field) - apply (drule has_field_idemp) - apply (drule (1) shconfD) - apply (unfold soconf_def has_field_def) - apply clarsimp - apply (fastforce dest: has_fields_fun) - done + then show ?thesis using Getstatic stk + has_field_idemp[OF has_visible_field[OF field]] shconfD[OF shconf] + by (unfold soconf_def has_field_def) (fastforce dest: has_fields_fun) next case (Putfield F C) with app stk reg \ obtain v ref vT stk' where field: "P \ C sees F,NonStatic:vT in C" and stk: "stk = v # ref # stk'" and confv: "P,h \ v :\ vT" and confr: "P,h \ ref :\ Class C" by fastforce from confr have is_Ref: "is_Ref ref" by simp moreover { assume "ref \ Null" with confr field is_Ref wf have "\D vs. h (the_Addr ref) = Some (D,vs) \ P \ D \\<^sup>* C" by (auto dest: non_npD) } ultimately show ?thesis using Putfield field stk confv by fastforce next case (Invoke M' n) with app have n: "n < size ST" by simp from stk have [simp]: "size stk = size ST" by (rule list_all2_lengthD) { assume "stk!n = Null" with n Invoke have ?thesis by simp } moreover { assume "ST!n = NT" with n stk have "stk!n = Null" by (auto simp: list_all2_conv_all_nth) with n Invoke have ?thesis by simp } moreover { assume Null: "stk!n \ Null" and NT: "ST!n \ NT" from NT app Invoke obtain D D' Ts T m where D: "ST!n = Class D" and M': "P \ D sees M',NonStatic: Ts\T = m in D'" and Ts: "P \ rev (take n ST) [\] Ts" by auto from D stk n have "P,h \ stk!n :\ Class D" by (auto simp: list_all2_conv_all_nth) with Null obtain a C' fs where [simp]: "stk!n = Addr a" "h a = Some (C',fs)" and "P \ C' \\<^sup>* D" by (fastforce dest!: conf_ClassD) with M' wf obtain m' Ts' T' D'' where C': "P \ C' sees M',NonStatic: Ts'\T' = m' in D''" and Ts': "P \ Ts [\] Ts'" by (auto dest!: sees_method_mono) from stk have "P,h \ take n stk [:\] take n ST" .. hence "P,h \ rev (take n stk) [:\] rev (take n ST)" .. also note Ts also note Ts' finally have "P,h \ rev (take n stk) [:\] Ts'" . with Invoke Null n C' have ?thesis by (auto simp add: is_Ref_def2 has_methodI) } ultimately show ?thesis by blast next case (Invokestatic C M' n) with app have n: "n \ size ST" by simp from stk have [simp]: "size stk = size ST" by (rule list_all2_lengthD) from app Invokestatic obtain D Ts T m where M': "P \ C sees M',Static: Ts\T = m in D" and Ts: "P \ rev (take n ST) [\] Ts" by auto from stk have "P,h \ take n stk [:\] take n ST" .. hence "P,h \ rev (take n stk) [:\] rev (take n ST)" .. also note Ts finally have "P,h \ rev (take n stk) [:\] Ts" . with Invokestatic n M' show ?thesis by (auto simp add: is_Ref_def2 has_methodI) next case Return show ?thesis proof(cases "M = clinit") case True have "is_class P C" by(rule sees_method_is_class[OF meth]) with wf_sees_clinit[OF wf] obtain m where "P \ C sees clinit,Static: [] \ Void = m in C" by(fastforce simp: is_class_def) with stk app \ meth frames True Return show ?thesis by (auto simp add: has_methodI) next case False with stk app \ meth frames Return show ?thesis by (auto intro: has_methodI) qed qed (auto simp add: list_all2_lengthD) hence "check P \" using meth pc mxs by (auto simp: check_def intro: has_methodI) } ultimately have "check P \" by blast thus "exec_d P \ \ TypeError" .. qed (*>*) text \ The theorem above tells us that, in welltyped programs, the defensive machine reaches the same result as the aggressive one (after arbitrarily many steps). \ theorem welltyped_aggressive_imp_defensive: "wf_jvm_prog\<^bsub>\\<^esub> P \ P,\ \ \ \ \ P \ \ -jvm\ \' \ P \ (Normal \) -jvmd\ (Normal \')" (*<*) - apply (simp only: exec_all_def) - apply (erule rtrancl_induct) - apply (simp add: exec_all_d_def1) - apply simp - apply (simp only: exec_all_def [symmetric]) - apply (frule BV_correct, assumption+) - apply (drule no_type_error, assumption, drule no_type_error_commutes, simp) - apply (simp add: exec_all_d_def1) - apply (rule rtrancl_trans, assumption) - apply (drule exec_1_d_NormalI) - apply auto - done +proof - + assume wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and cf: "P,\ \ \ \" and exec: "P \ \ -jvm\ \'" + then have "(\, \') \ {(\, \'). exec (P, \) = \\'\}\<^sup>*" by(simp only: exec_all_def) + then show ?thesis proof(induct rule: rtrancl_induct) + case base + then show ?case by (simp add: exec_all_d_def1) + next + case (step y z) + then have \y: "P \ \ -jvm\ y" by (simp only: exec_all_def [symmetric]) + have exec_d: "exec_d P y = Normal \z\" using step + no_type_error_commutes[OF no_type_error[OF wf BV_correct[OF wf \y cf]]] + by (simp add: exec_all_d_def1) + show ?case using step.hyps(3) exec_1_d_NormalI[OF exec_d] + by (simp add: exec_all_d_def1) + qed +qed (*>*) text \ As corollary we get that the aggressive and the defensive machine are equivalent for welltyped programs (if started in a conformant state or in the canonical start state) \ corollary welltyped_commutes: fixes \ :: jvm_state assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and conforms: "P,\ \ \ \" shows "P \ (Normal \) -jvmd\ (Normal \') = P \ \ -jvm\ \'" - apply rule - apply (erule defensive_imp_aggressive) - apply (erule welltyped_aggressive_imp_defensive [OF wf conforms]) - done +proof(rule iffI) + assume "P \ Normal \ -jvmd\ Normal \'" then show "P \ \ -jvm\ \'" + by (rule defensive_imp_aggressive) +next + assume "P \ \ -jvm\ \'" then show "P \ Normal \ -jvmd\ Normal \'" + by (rule welltyped_aggressive_imp_defensive [OF wf conforms]) +qed corollary welltyped_initial_commutes: assumes wf: "wf_jvm_prog P" assumes nstart: "\ is_class P Start" assumes meth: "P \ C sees M,Static:[]\Void = b in C" assumes nclinit: "M \ clinit" assumes Obj_start_m: "(\b' Ts' T' m' D'. P \ Object sees start_m, b' : Ts'\T' = m' in D' \ b' = Static \ Ts' = [] \ T' = Void)" defines start: "\ \ start_state P" shows "start_prog P C M \ (Normal \) -jvmd\ (Normal \') = start_prog P C M \ \ -jvm\ \'" proof - from wf obtain \ where wf': "wf_jvm_prog\<^bsub>\\<^esub> P" by (auto simp: wf_jvm_prog_def) let ?\ = "\_start \" from start_prog_wf_jvm_prog_phi[where \'="?\", OF wf' nstart meth nclinit \_start Obj_start_m] have "wf_jvm_prog\<^bsub>?\\<^esub>(start_prog P C M)" by simp moreover from wf' nstart meth nclinit \_start(2) have "start_prog P C M,?\ \ \ \" unfolding start by (rule BV_correct_initial) ultimately show ?thesis by (rule welltyped_commutes) qed lemma not_TypeError_eq [iff]: "x \ TypeError = (\t. x = Normal t)" by (cases x) auto locale cnf = fixes P and \ and \ assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes cnf: "correct_state P \ \" theorem (in cnf) no_type_errors: "P \ (Normal \) -jvmd\ \' \ \' \ TypeError" - apply (unfold exec_all_d_def1) - apply (erule rtrancl_induct) - apply simp - apply (fold exec_all_d_def1) - apply (insert cnf wf) - apply clarsimp - apply (drule defensive_imp_aggressive) - apply (frule (2) BV_correct) - apply (auto simp add: exec_1_d_eq dest: no_type_error) - done +proof - + assume "P \ (Normal \) -jvmd\ \'" + then have "(Normal \, \') \ (exec_1_d P)\<^sup>*" by (unfold exec_all_d_def1) simp + then show ?thesis proof(induct rule: rtrancl_induct) + case (step y z) + then obtain y\<^sub>n where [simp]: "y = Normal y\<^sub>n" by clarsimp + have n\y: "P \ Normal \ -jvmd\ Normal y\<^sub>n" using step.hyps(1) + by (fold exec_all_d_def1) simp + have \y: "P \ \ -jvm\ y\<^sub>n" using defensive_imp_aggressive[OF n\y] by simp + show ?case using step no_type_error[OF wf BV_correct[OF wf \y cnf]] + by (auto simp add: exec_1_d_eq) + qed simp +qed locale start = fixes P and C and M and \ and T and b and P\<^sub>0 assumes wf: "wf_jvm_prog P" assumes nstart: "\ is_class P Start" assumes sees: "P \ C sees M,Static:[]\Void = b in C" assumes nclinit: "M \ clinit" assumes Obj_start_m: "(\b' Ts' T' m' D'. P \ Object sees start_m, b' : Ts'\T' = m' in D' \ b' = Static \ Ts' = [] \ T' = Void)" defines "\ \ Normal (start_state P)" defines [simp]: "P\<^sub>0 \ start_prog P C M" corollary (in start) bv_no_type_error: shows "P\<^sub>0 \ \ -jvmd\ \' \ \' \ TypeError" proof - from wf obtain \ where wf': "wf_jvm_prog\<^bsub>\\<^esub> P" by (auto simp: wf_jvm_prog_def) let ?\ = "\_start \" from start_prog_wf_jvm_prog_phi[where \'="?\", OF wf' nstart sees nclinit \_start Obj_start_m] have "wf_jvm_prog\<^bsub>?\\<^esub>P\<^sub>0" by simp moreover from BV_correct_initial[where \'="?\", OF wf' nstart sees nclinit \_start(2)] have "correct_state P\<^sub>0 ?\ (start_state P)" by simp ultimately have "cnf P\<^sub>0 ?\ (start_state P)" by (rule cnf.intro) moreover assume "P\<^sub>0 \ \ -jvmd\ \'" ultimately show ?thesis by (unfold \_def) (rule cnf.no_type_errors) qed end diff --git a/thys/JinjaDCI/BV/BVSpec.thy b/thys/JinjaDCI/BV/BVSpec.thy --- a/thys/JinjaDCI/BV/BVSpec.thy +++ b/thys/JinjaDCI/BV/BVSpec.thy @@ -1,104 +1,113 @@ (* Title: JinjaDCI/BV/BVSpec.thy Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory BV/BVSpec.thy by Tobias Nipkow *) section \ The Bytecode Verifier \label{sec:BVSpec} \ theory BVSpec imports Effect begin text \ This theory contains a specification of the BV. The specification describes correct typings of method bodies; it corresponds to type \emph{checking}. \ definition \ \The method type only contains declared classes:\ check_types :: "'m prog \ nat \ nat \ ty\<^sub>i' err list \ bool" where "check_types P mxs mxl \s \ set \s \ states P mxs mxl" \ \An instruction is welltyped if it is applicable and its effect\ \ \is compatible with the type at all successor instructions:\ definition wt_instr :: "['m prog,ty,nat,pc,ex_table,instr,pc,ty\<^sub>m] \ bool" ("_,_,_,_,_ \ _,_ :: _" [60,0,0,0,0,0,0,61] 60) where "P,T,mxs,mpc,xt \ i,pc :: \s \ app i P mxs T pc mpc xt (\s!pc) \ (\(pc',\') \ set (eff i P pc xt (\s!pc)). P \ \' \' \s!pc')" \ \The type at @{text "pc=0"} conforms to the method calling convention:\ definition wt_start :: "['m prog,cname,staticb,ty list,nat,ty\<^sub>m] \ bool" where "wt_start P C b Ts mxl\<^sub>0 \s \ case b of NonStatic \ P \ Some ([],OK (Class C)#map OK Ts@replicate mxl\<^sub>0 Err) \' \s!0 | Static \ P \ Some ([],map OK Ts@replicate mxl\<^sub>0 Err) \' \s!0" \ \A method is welltyped if the body is not empty,\ \ \if the method type covers all instructions and mentions\ \ \declared classes only, if the method calling convention is respected, and\ \ \if all instructions are welltyped.\ definition wt_method :: "['m prog,cname,staticb,ty list,ty,nat,nat,instr list, ex_table,ty\<^sub>m] \ bool" where "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s \ 0 < size is \ size \s = size is \ check_types P mxs ((case b of Static \ 0 | NonStatic \ 1)+size Ts+mxl\<^sub>0) (map OK \s) \ wt_start P C b Ts mxl\<^sub>0 \s \ (\pc < size is. P,T\<^sub>r,mxs,size is,xt \ is!pc,pc :: \s)" \ \A program is welltyped if it is wellformed and all methods are welltyped\ definition wf_jvm_prog_phi :: "ty\<^sub>P \ jvm_prog \ bool" ("wf'_jvm'_prog\<^bsub>_\<^esub>") where "wf_jvm_prog\<^bsub>\\<^esub> \ wf_prog (\P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M))" definition wf_jvm_prog :: "jvm_prog \ bool" where "wf_jvm_prog P \ \\. wf_jvm_prog\<^bsub>\\<^esub> P" lemma wt_jvm_progD: "wf_jvm_prog\<^bsub>\\<^esub> P \ \wt. wf_prog wt P" (*<*) by (unfold wf_jvm_prog_phi_def, blast) (*>*) lemma wt_jvm_prog_impl_wt_instr: - "\ wf_jvm_prog\<^bsub>\\<^esub> P; - P \ C sees M,b:Ts \ T = (mxs,mxl\<^sub>0,ins,xt) in C; pc < size ins \ - \ P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" +assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and + sees: "P \ C sees M,b:Ts \ T = (mxs,mxl\<^sub>0,ins,xt) in C" and + pc: "pc < size ins" +shows "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" (*<*) - apply (unfold wf_jvm_prog_phi_def) - apply (drule (1) sees_wf_mdecl) - apply (simp add: wf_mdecl_def wt_method_def) - done +proof - + have wfm: "wf_prog + (\P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). + wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M)) P" using wf + by (unfold wf_jvm_prog_phi_def) + show ?thesis using sees_wf_mdecl[OF wfm sees] pc + by (simp add: wf_mdecl_def wt_method_def) +qed (*>*) lemma wt_jvm_prog_impl_wt_start: - "\ wf_jvm_prog\<^bsub>\\<^esub> P; - P \ C sees M,b:Ts \ T = (mxs,mxl\<^sub>0,ins,xt) in C \ \ - 0 < size ins \ wt_start P C b Ts mxl\<^sub>0 (\ C M)" +assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and + sees: "P \ C sees M,b:Ts \ T = (mxs,mxl\<^sub>0,ins,xt) in C" +shows "0 < size ins \ wt_start P C b Ts mxl\<^sub>0 (\ C M)" (*<*) - apply (unfold wf_jvm_prog_phi_def) - apply (drule (1) sees_wf_mdecl) - apply (simp add: wf_mdecl_def wt_method_def) - done +proof - + have wfm: "wf_prog + (\P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). + wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M)) P" using wf + by (unfold wf_jvm_prog_phi_def) + show ?thesis using sees_wf_mdecl[OF wfm sees] + by (simp add: wf_mdecl_def wt_method_def) +qed (*>*) lemma wf_jvm_prog_nclinit: assumes wtp: "wf_jvm_prog\<^bsub>\\<^esub> P" and meth: "P \ C sees M, b : Ts\T = (mxs, mxl\<^sub>0, ins, xt) in D" and wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" and pc: "pc < length ins" and \: "\ C M ! pc = Some(ST,LT)" and ins: "ins ! pc = Invokestatic C\<^sub>0 M\<^sub>0 n" shows "M\<^sub>0 \ clinit" using assms by(simp add: wf_jvm_prog_phi_def wt_instr_def app_def) end diff --git a/thys/JinjaDCI/BV/BVSpecTypeSafe.thy b/thys/JinjaDCI/BV/BVSpecTypeSafe.thy --- a/thys/JinjaDCI/BV/BVSpecTypeSafe.thy +++ b/thys/JinjaDCI/BV/BVSpecTypeSafe.thy @@ -1,2341 +1,2345 @@ (* Title: JinjaDCI/BV/BVSpecTypeSafe.thy Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory BV/BVSpecTypeSafe.thy by Cornelia Pusch and Gerwin Klein *) section \ BV Type Safety Proof \label{sec:BVSpecTypeSafe} \ theory BVSpecTypeSafe imports BVConform StartProg begin text \ This theory contains proof that the specification of the bytecode verifier only admits type safe programs. \ subsection \ Preliminaries \ text \ Simp and intro setup for the type safety proof: \ lemmas defs1 = correct_state_def conf_f_def wt_instr_def eff_def norm_eff_def app_def xcpt_app_def lemmas widen_rules [intro] = conf_widen confT_widen confs_widens confTs_widen subsection \ Exception Handling \ text \ For the @{text Invoke} instruction the BV has checked all handlers that guard the current @{text pc}. \ lemma Invoke_handlers: "match_ex_table P C pc xt = Some (pc',d') \ \(f,t,D,h,d) \ set (relevant_entries P (Invoke n M) pc xt). P \ C \\<^sup>* D \ pc \ {f.. pc' = h \ d' = d" by (induct xt) (auto simp: relevant_entries_def matches_ex_entry_def is_relevant_entry_def split: if_split_asm) text \ For the @{text Invokestatic} instruction the BV has checked all handlers that guard the current @{text pc}. \ lemma Invokestatic_handlers: "match_ex_table P C pc xt = Some (pc',d') \ \(f,t,D,h,d) \ set (relevant_entries P (Invokestatic C\<^sub>0 n M) pc xt). P \ C \\<^sup>* D \ pc \ {f.. pc' = h \ d' = d" by (induct xt) (auto simp: relevant_entries_def matches_ex_entry_def is_relevant_entry_def split: if_split_asm) text \ For the instrs in @{text Called_set} the BV has checked all handlers that guard the current @{text pc}. \ lemma Called_set_handlers: "match_ex_table P C pc xt = Some (pc',d') \ i \ Called_set \ \(f,t,D,h,d) \ set (relevant_entries P i pc xt). P \ C \\<^sup>* D \ pc \ {f.. pc' = h \ d' = d" by (induct xt) (auto simp: relevant_entries_def matches_ex_entry_def is_relevant_entry_def split: if_split_asm) text \ We can prove separately that the recursive search for exception handlers (@{text find_handler}) in the frame stack results in a conforming state (if there was no matching exception handler in the current frame). We require that the exception is a valid heap address, and that the state before the exception occurred conforms. \ lemma uncaught_xcpt_correct: assumes wt: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes h: "h xcp = Some obj" shows "\f. P,\ \ (None, h, f#frs, sh)\ \ curr_method f \ clinit \ P,\ \ find_handler P xcp h frs sh \" (is "\f. ?correct (None, h, f#frs, sh) \ ?prem f \ ?correct (?find frs)") (*<*) proof (induct frs) \ \the base case is trivial as it should be\ show "?correct (?find [])" by (simp add: correct_state_def) next \ \we will need both forms @{text wf_jvm_prog} and @{text wf_prog} later\ from wt obtain mb where wf: "wf_prog mb P" by (simp add: wf_jvm_prog_phi_def) \ \the assumptions for the cons case:\ fix f f' frs' assume cr: "?correct (None, h, f#f'#frs', sh)" assume pr: "?prem f" \ \the induction hypothesis:\ assume IH: "\f. ?correct (None, h, f#frs', sh) \ ?prem f \ ?correct (?find frs')" from cr pr conf_clinit_Cons[where frs="f'#frs'" and f=f] obtain confc: "conf_clinit P sh (f'#frs')" and cr': "?correct (None, h, f'#frs', sh)" by(fastforce simp: correct_state_def) obtain stk loc C M pc ics where [simp]: "f' = (stk,loc,C,M,pc,ics)" by (cases f') from cr' obtain b Ts T mxs mxl\<^sub>0 ins xt where meth: "P \ C sees M,b:Ts \ T = (mxs,mxl\<^sub>0,ins,xt) in C" by (simp add: correct_state_def, blast) hence xt[simp]: "ex_table_of P C M = xt" by simp have cls: "is_class P C" by(rule sees_method_is_class'[OF meth]) from cr' obtain sfs where sfs: "M = clinit \ sh C = Some(sfs,Processing)" by(fastforce simp: defs1 conf_clinit_def) show "?correct (?find (f'#frs'))" proof (cases "match_ex_table P (cname_of h xcp) pc xt") case None with cr' IH[of f'] show ?thesis proof(cases "M=clinit") case True then show ?thesis using xt cr' IH[of f'] None h conf_clinit_Called_Throwing conf_f_Throwing[where h=h and sh=sh, OF _ cls h sfs] by(cases frs', auto simp: correct_state_def image_iff) fastforce qed(auto) next fix pc_d assume "match_ex_table P (cname_of h xcp) pc xt = Some pc_d" then obtain pc' d' where match: "match_ex_table P (cname_of h xcp) pc xt = Some (pc',d')" (is "?match (cname_of h xcp) = _") by (cases pc_d) auto from wt meth cr' [simplified] have wti: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" by (fastforce simp: correct_state_def conf_f_def dest: sees_method_fun elim!: wt_jvm_prog_impl_wt_instr) from cr' obtain ST LT where \: "\ C M ! pc = Some (ST, LT)" by(fastforce dest: sees_method_fun simp: correct_state_def) from cr' \ meth have conf': "conf_f P h sh (ST, LT) ins f'" by (unfold correct_state_def) (fastforce dest: sees_method_fun) hence loc: "P,h \ loc [:\\<^sub>\] LT" and stk: "P,h \ stk [:\] ST" by (unfold conf_f_def) auto hence [simp]: "size stk = size ST" by (simp add: list_all2_lengthD) from cr meth pr obtain D n M' where ins: "ins!pc = Invoke n M' \ ins!pc = Invokestatic D n M'" (is "_ = ?i \ _ = ?i'") by(fastforce dest: sees_method_fun simp: correct_state_def) with match obtain f1 t D where rel: "(f1,t,D,pc',d') \ set (relevant_entries P (ins!pc) pc xt)" and D: "P \ cname_of h xcp \\<^sup>* D" by(fastforce dest: Invoke_handlers Invokestatic_handlers) from rel have "(pc', Some (Class D # drop (size ST - d') ST, LT)) \ set (xcpt_eff (ins!pc) P pc (ST,LT) xt)" (is "(_, Some (?ST',_)) \ _") by (force simp: xcpt_eff_def image_def) with wti \ obtain pc: "pc' < size ins" and "P \ Some (?ST', LT) \' \ C M ! pc'" by (clarsimp simp: defs1) blast then obtain ST' LT' where \': "\ C M ! pc' = Some (ST',LT')" and less: "P \ (?ST', LT) \\<^sub>i (ST',LT')" by (auto simp: sup_state_opt_any_Some) let ?f = "(Addr xcp # drop (length stk - d') stk, loc, C, M, pc',No_ics)" have "conf_f P h sh (ST',LT') ins ?f" proof - from wf less loc have "P,h \ loc [:\\<^sub>\] LT'" by simp blast moreover from D h have "P,h \ Addr xcp :\ Class D" by (simp add: conf_def obj_ty_def case_prod_unfold) with less stk have "P,h \ Addr xcp # drop (length stk - d') stk [:\] ST'" by (auto intro!: list_all2_dropI) ultimately show ?thesis using pc conf' by(auto simp: conf_f_def) qed with cr' match \' meth pc show ?thesis by (unfold correct_state_def) (cases "M=clinit"; fastforce dest: sees_method_fun simp: conf_clinit_def distinct_clinit_def) qed qed (*>*) text \ The requirement of lemma @{text uncaught_xcpt_correct} (that the exception is a valid reference on the heap) is always met for welltyped instructions and conformant states: \ lemma exec_instr_xcpt_h: "\ fst (exec_instr (ins!pc) P h stk vars C M pc ics frs sh) = Some xcp; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\ \ \ \obj. h xcp = Some obj" (is "\ ?xcpt; ?wt; ?correct \ \ ?thesis") (*<*) proof - note [simp] = split_beta note [split] = if_split_asm option.split_asm assume wt: ?wt ?correct hence pre: "preallocated h" by (simp add: correct_state_def hconf_def) assume xcpt: ?xcpt with exec_instr_xcpts have opt: "ins!pc = Throw \ xcp \ {a. \x \ sys_xcpts. a = addr_of_sys_xcpt x}" by simp with pre show ?thesis proof (cases "ins!pc") case Throw with xcpt wt pre show ?thesis by (clarsimp iff: list_all2_Cons2 simp: defs1) (auto dest: non_npD simp: is_refT_def elim: preallocatedE) qed (auto elim: preallocatedE) qed (*>*) lemma exec_step_xcpt_h: assumes xcpt: "fst (exec_step P h stk vars C M pc ics frs sh) = Some xcp" and ins: "instrs_of P C M = ins" and wti: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" and correct: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" shows "\obj. h xcp = Some obj" proof - from correct have pre: "preallocated h" by(simp add: defs1 hconf_def) { fix C' Cs assume ics[simp]: "ics = Calling C' Cs" with xcpt have "xcp = addr_of_sys_xcpt NoClassDefFoundError" by(cases ics, auto simp: split_beta split: init_state.splits if_split_asm) with pre have ?thesis using preallocated_def by force } moreover { fix Cs a assume [simp]: "ics = Throwing Cs a" with xcpt have eq: "a = xcp" by(cases Cs; simp) from correct have "P,h,sh \\<^sub>i (C,M,pc,ics)" by(auto simp: defs1) with eq have ?thesis by simp } moreover { fix Cs assume ics: "ics = No_ics \ ics = Called Cs" with exec_instr_xcpt_h[OF _ wti correct] xcpt ins have ?thesis by(cases Cs, auto) } ultimately show ?thesis by(cases ics, auto) qed lemma conf_sys_xcpt: "\preallocated h; C \ sys_xcpts\ \ P,h \ Addr (addr_of_sys_xcpt C) :\ Class C" by (auto elim: preallocatedE) lemma match_ex_table_SomeD: "match_ex_table P C pc xt = Some (pc',d') \ \(f,t,D,h,d) \ set xt. matches_ex_entry P C pc (f,t,D,h,d) \ h = pc' \ d=d'" by (induct xt) (auto split: if_split_asm) text \ Finally we can state that, whenever an exception occurs, the next state always conforms: \ lemma xcpt_correct: fixes \' :: jvm_state assumes wtp: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes meth: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes xp: "fst (exec_step P h stk loc C M pc ics frs sh) = Some xcp" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" assumes correct: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" shows "P,\ \ \'\" (*<*) proof - from wtp obtain wfmb where wf: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) from meth have ins[simp]: "instrs_of P C M = ins" by simp have cls: "is_class P C" by(rule sees_method_is_class[OF meth]) from correct obtain sfs where sfs: "M = clinit \ sh C = Some(sfs,Processing)" by(auto simp: correct_state_def conf_clinit_def conf_f_def2) note conf_sys_xcpt [elim!] note xp' = meth s' xp from correct meth obtain ST LT where h_ok: "P \ h \" and sh_ok: "P,h \\<^sub>s sh \" and \_pc: "\ C M ! pc = Some (ST, LT)" and frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,ics)" and frames: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and vics: "P,h,sh \\<^sub>i (C,M,pc,ics)" by(auto simp: defs1 dest: sees_method_fun) from frame obtain stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" by (unfold conf_f_def) auto from h_ok have preh: "preallocated h" by (simp add: hconf_def) note wtp moreover from exec_step_xcpt_h[OF xp ins wt correct] obtain obj where h: "h xcp = Some obj" by clarify moreover note correct ultimately have fh: "curr_method (stk,loc,C,M,pc,ics) \ clinit \ P,\ \ find_handler P xcp h frs sh \" by (rule uncaught_xcpt_correct) with xp' have "M \ clinit \ \Cs a. ics \ Throwing Cs a \ match_ex_table P (cname_of h xcp) pc xt = None \ ?thesis" (is "?nc \ ?t \ ?m (cname_of h xcp) = _ \ _" is "?nc \ ?t \ ?match = _ \ _") by(cases ics; simp add: split_beta) moreover from correct xp' conf_clinit_Called_Throwing conf_f_Throwing[where h=h and sh=sh, OF _ cls h sfs] have "M = clinit \ \Cs a. ics \ Throwing Cs a \ match_ex_table P (cname_of h xcp) pc xt = None \ ?thesis" by(cases frs, auto simp: correct_state_def image_iff split_beta) fastforce moreover { fix pc_d assume "?match = Some pc_d" then obtain pc' d' where some_handler: "?match = Some (pc',d')" by (cases pc_d) auto from stk have [simp]: "size stk = size ST" .. from wt \_pc have eff: "\(pc', s')\set (xcpt_eff (ins!pc) P pc (ST,LT) xt). pc' < size ins \ P \ s' \' \ C M!pc'" by (auto simp: defs1) from some_handler obtain f t D where xt: "(f,t,D,pc',d') \ set xt" and "matches_ex_entry P (cname_of h xcp) pc (f,t,D,pc',d')" by (auto dest: match_ex_table_SomeD) hence match: "P \ cname_of h xcp \\<^sup>* D" "pc \ {f.. ics = Called (C'#Cs)" let ?stk' = "Addr xcp # drop (length stk - d') stk" let ?f = "(?stk', loc, C, M, pc', No_ics)" from some_handler xp' ics have \': "\' = (None, h, ?f#frs, sh)" by (cases ics; simp add: split_beta) from xp ics have "xcp = addr_of_sys_xcpt NoClassDefFoundError" by(cases ics, auto simp: split_beta split: init_state.splits if_split_asm) with match preh have conf: "P,h \ Addr xcp :\ Class D" by fastforce from correct ics obtain C1 where "Called_context P C1 (ins!pc)" by(fastforce simp: correct_state_def conf_f_def2) then have "ins!pc \ Called_set" by(rule Called_context_Called_set) with xt match have "(f,t,D,pc',d') \ set (relevant_entries P (ins!pc) pc xt)" by(auto simp: relevant_entries_def is_relevant_entry_def) with eff obtain ST' LT' where \_pc': "\ C M ! pc' = Some (ST', LT')" and pc': "pc' < size ins" and less: "P \ (Class D # drop (size ST - d') ST, LT) \\<^sub>i (ST', LT')" by (fastforce simp: xcpt_eff_def sup_state_opt_any_Some) with conf loc stk conf_f_def2 frame ics have "conf_f P h sh (ST',LT') ins ?f" by (auto simp: defs1 intro: list_all2_dropI) with meth h_ok frames \_pc' \' sh_ok confc ics have ?thesis by (unfold correct_state_def) (auto dest: sees_method_fun conf_clinit_diff' sees_method_is_class; fastforce) } moreover { assume ics: "ics = No_ics \ ics = Called []" let ?stk' = "Addr xcp # drop (length stk - d') stk" let ?f = "(?stk', loc, C, M, pc', No_ics)" from some_handler xp' ics have \': "\' = (None, h, ?f#frs, sh)" by (cases ics; simp add: split_beta) from xp ics obtain "(f,t,D,pc',d') \ set (relevant_entries P (ins!pc) pc xt)" and conf: "P,h \ Addr xcp :\ Class D" proof (cases "ins!pc") case Return with xp ics have False by(cases ics; cases frs, auto simp: split_beta split: if_split_asm) then show ?thesis by simp next case New with xp match have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) moreover from xp wt correct obtain obj where xcp: "h xcp = Some obj" by (blast dest: exec_step_xcpt_h[OF _ ins]) ultimately show ?thesis using xt match by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that) next case Getfield with xp ics have xcp: "xcp = addr_of_sys_xcpt NullPointer \ xcp = addr_of_sys_xcpt NoSuchFieldError \ xcp = addr_of_sys_xcpt IncompatibleClassChangeError" by (cases ics; simp add: split_beta split: if_split_asm staticb.splits) with Getfield match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (fastforce simp: is_relevant_entry_def) with match preh xt xcp show ?thesis by(fastforce simp: relevant_entries_def intro: that) next case Getstatic with xp match have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) moreover from xp wt correct obtain obj where xcp: "h xcp = Some obj" by (blast dest: exec_step_xcpt_h[OF _ ins]) ultimately show ?thesis using xt match by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that) next case Putfield with xp ics have xcp: "xcp = addr_of_sys_xcpt NullPointer \ xcp = addr_of_sys_xcpt NoSuchFieldError \ xcp = addr_of_sys_xcpt IncompatibleClassChangeError" by (cases ics; simp add: split_beta split: if_split_asm staticb.splits) with Putfield match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (fastforce simp: is_relevant_entry_def) with match preh xt xcp show ?thesis by (fastforce simp: relevant_entries_def intro: that) next case Putstatic with xp match have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) moreover from xp wt correct obtain obj where xcp: "h xcp = Some obj" by (blast dest: exec_step_xcpt_h[OF _ ins]) ultimately show ?thesis using xt match by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that) next case Checkcast with xp ics have [simp]: "xcp = addr_of_sys_xcpt ClassCast" by (cases ics; simp add: split_beta split: if_split_asm) with Checkcast match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) with match preh xt show ?thesis by (fastforce simp: relevant_entries_def intro: that) next case Invoke with xp match have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) moreover from xp wt correct obtain obj where xcp: "h xcp = Some obj" by (blast dest: exec_step_xcpt_h[OF _ ins]) ultimately show ?thesis using xt match by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that) next case Invokestatic with xp match have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) moreover from xp wt correct obtain obj where xcp: "h xcp = Some obj" by (blast dest: exec_step_xcpt_h[OF _ ins]) ultimately show ?thesis using xt match by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that) next case Throw with xp match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) moreover from xp wt correct obtain obj where xcp: "h xcp = Some obj" by (blast dest: exec_step_xcpt_h[OF _ ins]) ultimately show ?thesis using xt match by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that) qed(cases ics, (auto)[5])+ with eff obtain ST' LT' where \_pc': "\ C M ! pc' = Some (ST', LT')" and pc': "pc' < size ins" and less: "P \ (Class D # drop (size ST - d') ST, LT) \\<^sub>i (ST', LT')" by (fastforce simp: xcpt_eff_def sup_state_opt_any_Some) with conf loc stk conf_f_def2 frame ics have "conf_f P h sh (ST',LT') ins ?f" by (auto simp: defs1 intro: list_all2_dropI) with meth h_ok frames \_pc' \' sh_ok confc ics have ?thesis by (unfold correct_state_def) (auto dest: sees_method_fun conf_clinit_diff'; fastforce) } ultimately have "\Cs a. ics \ Throwing Cs a \ ?thesis" by(cases ics; metis list.exhaust) } moreover { fix Cs a assume "ics = Throwing Cs a" with xp' have ics: "ics = Throwing [] xcp" by(cases Cs; clarsimp) let ?frs = "(stk,loc,C,M,pc,No_ics)#frs" have eT: "exec_step P h stk loc C M pc (Throwing [] xcp) frs sh = (Some xcp, h, ?frs, sh)" by auto with xp' ics have \'_fh: "\' = find_handler P xcp h ?frs sh" by simp from meth have [simp]: "xt = ex_table_of P C M" by simp let ?match = "match_ex_table P (cname_of h xcp) pc xt" { assume clinit: "M = clinit" and None: "?match = None" note asms = clinit None have "P,\ |- find_handler P xcp h ?frs sh [ok]" proof(cases frs) case Nil with h_ok sh_ok asms show "P,\ |- find_handler P xcp h ?frs sh [ok]" by(simp add: correct_state_def) next case [simp]: (Cons f' frs') obtain stk' loc' C' M' pc' ics' where [simp]: "f' = (stk',loc',C',M',pc',ics')" by(cases f') have cls: "is_class P C" by(rule sees_method_is_class[OF meth]) have shC: "sh C = Some(sfs,Processing)" by(rule sfs[OF clinit]) from correct obtain b Ts T mxs' mxl\<^sub>0' ins' xt' ST' LT' where meth': "P \ C' sees M', b : Ts\T = (mxs', mxl\<^sub>0', ins', xt') in C'" and \_pc': "\ C' M' ! pc' = \(ST', LT')\" and frame': "conf_f P h sh (ST',LT') ins' (stk', loc', C', M', pc', ics')" and frames': "conf_fs P h sh \ C' M' (length Ts) T frs'" and confc': "conf_clinit P sh ((stk',loc',C',M',pc',ics')#frs')" by(auto dest: conf_clinit_Cons simp: correct_state_def) from meth' have ins'[simp]: "instrs_of P C' M' = ins'" and [simp]: "xt' = ex_table_of P C' M'" by simp+ let ?f' = "case ics' of Called Cs' \ (stk',loc',C',M',pc',Throwing (C#Cs') xcp) | _ \ (stk',loc',C',M',pc',ics')" from asms confc have confc_T: "conf_clinit P sh (?f'#frs')" by(cases ics', auto simp: conf_clinit_def distinct_clinit_def) from asms conf_f_Throwing[where h=h and sh=sh, OF _ cls h shC] frame' have frame_T: "conf_f P h sh (ST', LT') ins' ?f'" by(cases ics'; simp) with h_ok sh_ok meth' \_pc' confc_T frames' have "P,\ |- (None, h, ?f'#frs', sh) [ok]" by(cases ics') (fastforce simp: correct_state_def)+ with asms show ?thesis by(cases ics'; simp) qed } moreover { assume asms: "M \ clinit" "?match = None" from asms uncaught_xcpt_correct[OF wtp h correct] have "P,\ |- find_handler P xcp h frs sh [ok]" by simp with asms have "P,\ |- find_handler P xcp h ?frs sh [ok]" by auto } moreover { fix pc_d assume some_handler: "?match = \pc_d\" (is "?match = \pc_d\") then obtain pc1 d1 where sh': "?match = Some(pc1,d1)" by(cases pc_d, simp) let ?stk' = "Addr xcp # drop (length stk - d1) stk" let ?f = "(?stk', loc, C, M, pc1, No_ics)" from stk have [simp]: "size stk = size ST" .. from wt \_pc have eff: "\(pc1, s')\set (xcpt_eff (ins!pc) P pc (ST,LT) xt). pc1 < size ins \ P \ s' \' \ C M!pc1" by (auto simp: defs1) from match_ex_table_SomeD[OF sh'] obtain f t D where xt: "(f,t,D,pc1,d1) \ set xt" and "matches_ex_entry P (cname_of h xcp) pc (f,t,D,pc1,d1)" by auto hence match: "P \ cname_of h xcp \\<^sup>* D" "pc \ {f.. Called_set" by(rule Called_context_Called_set) with match xt xp ics obtain res: "(f,t,D,pc1,d1) \ set (relevant_entries P (ins!pc) pc xt)" by(auto simp: relevant_entries_def is_relevant_entry_def) with h match xt xp ics have conf: "P,h \ Addr xcp :\ Class D" by (auto simp: relevant_entries_def conf_def case_prod_unfold) with eff res obtain ST1 LT1 where \_pc1: "\ C M ! pc1 = Some (ST1, LT1)" and pc1: "pc1 < size ins" and less1: "P \ (Class D # drop (size ST - d1) ST, LT) \\<^sub>i (ST1, LT1)" by (fastforce simp: xcpt_eff_def sup_state_opt_any_Some) with conf loc stk conf_f_def2 frame ics have frame1: "conf_f P h sh (ST1,LT1) ins ?f" by (auto simp: defs1 intro: list_all2_dropI) from \_pc1 h_ok sh_ok meth frame1 frames conf_clinit_diff'[OF confc] have "P,\ |- (None, h, ?f # frs, sh) [ok]" by(fastforce simp: correct_state_def) with sh' have "P,\ |- find_handler P xcp h ?frs sh [ok]" by auto } ultimately have cr': "P,\ |- find_handler P xcp h ?frs sh [ok]" by(cases "?match") blast+ with \'_fh have ?thesis by simp } ultimately show ?thesis by (cases "?match") blast+ qed (*>*) (**********Non-exception Single-step correctness*************************) declare defs1 [simp] subsection \ Initialization procedure steps \ text \ In this section we prove that, for states that result in a step of the initialization procedure rather than an instruction execution, the state after execution of the step still conforms. \ lemma Calling_correct: fixes \' :: jvm_state assumes wtprog: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" assumes cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None" assumes ics: "ics = Calling C' Cs" shows "P,\ \ \'\" proof - from wtprog obtain wfmb where wf: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) from mC cf obtain ST LT where h_ok: "P \ h \" and sh_ok: "P,h \\<^sub>s sh \" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and frame: "conf_f P h sh (ST, LT) ins (stk,loc,C,M,pc,ics)" and fs: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and vics: "P,h,sh \\<^sub>i (C,M,pc,ics)" by (fastforce dest: sees_method_fun) with ics have confc\<^sub>0: "conf_clinit P sh ((stk,loc,C,M,pc,Calling C' Cs)#frs)" by simp from vics ics have cls': "is_class P C'" by auto { assume None: "sh C' = None" let ?sh = "sh(C' \ (sblank P C', Prepared))" obtain FDTs where flds: "P \ C' has_fields FDTs" using wf_Fields_Ex[OF wf cls'] by clarsimp from shconf_upd_obj[where C=C', OF sh_ok soconf_sblank[OF flds]] have sh_ok': "P,h \\<^sub>s ?sh \" by simp from None have "\sfs. sh C' \ Some(sfs,Processing)" by simp with conf_clinit_nProc_dist[OF confc] have dist': "distinct (C' # clinit_classes ((stk, loc, C, M, pc, ics) # frs))" by simp then have dist'': "distinct (C' # clinit_classes frs)" by simp have confc': "conf_clinit P ?sh ((stk, loc, C, M, pc, ics) # frs)" by(rule conf_clinit_shupd[OF confc dist']) have fs': "conf_fs P h ?sh \ C M (size Ts) T frs" by(rule conf_fs_shupd[OF fs dist'']) from vics ics have vics': "P,h,?sh \\<^sub>i (C, M, pc, ics)" by auto from s' ics None have "\' = (None, h, (stk, loc, C, M, pc, ics)#frs, ?sh)" by auto with mC h_ok sh_ok' \ stk loc pc fs' confc vics' confc' frame None have ?thesis by fastforce } moreover { fix a assume "sh C' = Some a" then obtain sfs i where shC'[simp]: "sh C' = Some(sfs,i)" by(cases a, simp) from confc ics have last: "\sobj. sh (last(C'#Cs)) = Some sobj" by(fastforce simp: conf_clinit_def) let "?f" = "\ics'. (stk, loc, C, M, pc, ics'::init_call_status)" { assume i: "i = Done \ i = Processing" let ?ics = "Called Cs" from last vics ics have vics': "P,h,sh \\<^sub>i (C, M, pc, ?ics)" by auto from confc ics have confc': "conf_clinit P sh (?f ?ics#frs)" by(cases "M=clinit"; clarsimp simp: conf_clinit_def distinct_clinit_def) from i s' ics have "\' = (None, h, ?f ?ics#frs, sh)" by auto with mC h_ok sh_ok \ stk loc pc fs confc' vics' frame ics have ?thesis by fastforce } moreover { assume i[simp]: "i = Error" let ?a = "addr_of_sys_xcpt NoClassDefFoundError" let ?ics = "Throwing Cs ?a" from h_ok have preh: "preallocated h" by (simp add: hconf_def) then obtain obj where ha: "h ?a = Some obj" by(clarsimp simp: preallocated_def sys_xcpts_def) with vics ics have vics': "P,h,sh \\<^sub>i (C, M, pc, ?ics)" by auto from confc ics have confc'': "conf_clinit P sh (?f ?ics#frs)" by(cases "M=clinit"; clarsimp simp: conf_clinit_def distinct_clinit_def) from s' ics have \': "\' = (None, h, ?f ?ics#frs, sh)" by auto from mC h_ok sh_ok \ stk loc pc fs confc'' vics \' ics ha have ?thesis by fastforce } moreover { assume i[simp]: "i = Prepared" let ?sh = "sh(C' \ (sfs,Processing))" let ?D = "fst(the(class P C'))" let ?ics = "if C' = Object then Called (C'#Cs) else Calling ?D (C'#Cs)" from shconf_upd_obj[where C=C', OF sh_ok shconfD[OF sh_ok shC']] have sh_ok': "P,h \\<^sub>s ?sh \" by simp from cls' have "C' \ Object \ P \ C' \\<^sup>* ?D" by(auto simp: is_class_def intro!: subcls1I) with is_class_supclass[OF wf _ cls'] have D: "C' \ Object \ is_class P ?D" by simp from i have "\sfs. sh C' \ Some(sfs,Processing)" by simp with conf_clinit_nProc_dist[OF confc\<^sub>0] have dist': "distinct (C' # clinit_classes ((stk, loc, C, M, pc, Calling C' Cs) # frs))" by fast then have dist'': "distinct (C' # clinit_classes frs)" by simp from conf_clinit_shupd_Calling[OF confc\<^sub>0 dist' cls'] conf_clinit_shupd_Called[OF confc\<^sub>0 dist' cls'] have confc': "conf_clinit P ?sh (?f ?ics#frs)" by clarsimp with last ics have "\sobj. ?sh (last(C'#Cs)) = Some sobj" by(auto simp: conf_clinit_def fun_upd_apply) with D vics ics have vics': "P,h,?sh \\<^sub>i (C, M, pc, ?ics)" by auto have fs': "conf_fs P h ?sh \ C M (size Ts) T frs" by(rule conf_fs_shupd[OF fs dist'']) from frame vics' have frame': "conf_f P h ?sh (ST, LT) ins (?f ?ics)" by simp from i s' ics have "\' = (None, h, ?f ?ics#frs, ?sh)" by(auto simp: if_split_asm) with mC h_ok sh_ok' \ stk loc pc fs' confc' frame' ics have ?thesis by fastforce } ultimately have ?thesis by(cases i, auto) } ultimately show ?thesis by(cases "sh C'", auto) qed lemma Throwing_correct: fixes \' :: jvm_state assumes wtprog: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" assumes cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None" assumes ics: "ics = Throwing (C'#Cs) a" shows "P,\ \ \'\" proof - from wtprog obtain wfmb where wf: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) from mC cf obtain ST LT where h_ok: "P \ h \" and sh_ok: "P,h \\<^sub>s sh \" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and frame: "conf_f P h sh (ST, LT) ins (stk,loc,C,M,pc,ics)" and fs: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and vics: "P,h,sh \\<^sub>i (C,M,pc,ics)" by (fastforce dest: sees_method_fun) with ics have confc\<^sub>0: "conf_clinit P sh ((stk,loc,C,M,pc,Throwing (C'#Cs) a)#frs)" by simp from frame ics mC have cc: "\C1. Called_context P C1 (ins ! pc)" by(clarsimp simp: conf_f_def2) from frame ics obtain obj where ha: "h a = Some obj" by(auto simp: conf_f_def2) from confc ics obtain sfs i where shC': "sh C' = Some(sfs,i)" by(clarsimp simp: conf_clinit_def) then have sfs: "P,h,C' \\<^sub>s sfs \" by(rule shconfD[OF sh_ok]) from s' ics have \': "\' = (None, h, (stk,loc,C,M,pc,Throwing Cs a)#frs, sh(C' \ (fst(the(sh C')), Error)))" (is "\' = (None, h, ?f'#frs, ?sh')") by simp from confc ics have dist: "distinct (C' # clinit_classes (?f' # frs))" by (simp add: conf_clinit_def distinct_clinit_def) then have dist': "distinct (C' # clinit_classes frs)" by simp from conf_clinit_Throwing confc ics have confc': "conf_clinit P sh (?f' # frs)" by simp from shconf_upd_obj[OF sh_ok sfs] shC' have "P,h \\<^sub>s ?sh' \" by simp moreover have "conf_fs P h ?sh' \ C M (length Ts) T frs" by(rule conf_fs_shupd[OF fs dist']) moreover have "conf_clinit P ?sh' (?f' # frs)" by(rule conf_clinit_shupd[OF confc' dist]) moreover note \' h_ok mC \ pc stk loc ha cc ultimately show "P,\ \ \' \" by fastforce qed lemma Called_correct: fixes \' :: jvm_state assumes wtprog: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" assumes cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None" assumes ics[simp]: "ics = Called (C'#Cs)" shows "P,\ \ \'\" proof - from wtprog obtain wfmb where wf: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) from mC cf obtain ST LT where h_ok: "P \ h \" and sh_ok: "P,h \\<^sub>s sh \" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and frame: "conf_f P h sh (ST, LT) ins (stk,loc,C,M,pc,ics)" and fs: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and vics: "P,h,sh \\<^sub>i (C,M,pc,ics)" by (fastforce dest: sees_method_fun) then have confc\<^sub>0: "conf_clinit P sh ((stk,loc,C,M,pc,Called (C'#Cs))#frs)" by simp from frame mC obtain C1 sobj where ss: "Called_context P C1 (ins ! pc)" and shC1: "sh C1 = Some sobj" by(clarsimp simp: conf_f_def2) from confc wf_sees_clinit[OF wf] obtain mxs' mxl' ins' xt' where clinit: "P \ C' sees clinit,Static: [] \ Void=(mxs',mxl',ins',xt') in C'" by(fastforce simp: conf_clinit_def is_class_def) let ?loc' = "replicate mxl' undefined" from s' clinit have \': "\' = (None, h, ([],?loc',C',clinit,0,No_ics)#(stk,loc,C,M,pc,Called Cs)#frs, sh)" (is "\' = (None, h, ?if#?f'#frs, sh)") by simp with wtprog clinit obtain start: "wt_start P C' Static [] mxl' (\ C' clinit)" and ins': "ins' \ []" by (auto dest: wt_jvm_prog_impl_wt_start) then obtain LT\<^sub>0 where LT\<^sub>0: "\ C' clinit ! 0 = Some ([], LT\<^sub>0)" by (clarsimp simp: wt_start_def defs1 sup_state_opt_any_Some split: staticb.splits) moreover have "conf_f P h sh ([], LT\<^sub>0) ins' ?if" proof - let ?LT = "replicate mxl' Err" have "P,h \ ?loc' [:\\<^sub>\] ?LT" by simp also from start LT\<^sub>0 have "P \ \ [\\<^sub>\] LT\<^sub>0" by (simp add: wt_start_def) finally have "P,h \ ?loc' [:\\<^sub>\] LT\<^sub>0" . thus ?thesis using ins' by simp qed moreover from conf_clinit_Called confc clinit have "conf_clinit P sh (?if # ?f' # frs)" by simp moreover note \' h_ok sh_ok mC \ pc stk loc clinit ss shC1 fs ultimately show "P,\ \ \' \" by fastforce qed subsection \ Single Instructions \ text \ In this section we prove for each single (welltyped) instruction that the state after execution of the instruction still conforms. Since we have already handled exceptions above, we can now assume that no exception occurs in this step. For instructions that may call the initialization procedure, we cover the calling and non-calling cases separately. \ lemma Invoke_correct: fixes \' :: jvm_state assumes wtprog: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes meth_C: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes ins: "ins ! pc = Invoke M' n" assumes wti: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes \': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" assumes approx: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" assumes no_xcp: "fst (exec_step P h stk loc C M pc ics frs sh) = None" shows "P,\ \ \'\" (*<*) proof - from meth_C approx ins have [simp]: "ics = No_ics" by(cases ics, auto) note split_paired_Ex [simp del] from wtprog obtain wfmb where wfprog: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) from ins meth_C approx obtain ST LT where heap_ok: "P\ h\" and \_pc: "\ C M!pc = Some (ST,LT)" and frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,ics)" and frames: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" by (fastforce dest: sees_method_fun) from ins wti \_pc have n: "n < size ST" by simp { assume "stk!n = Null" with ins no_xcp meth_C have False by (simp add: split_beta) hence ?thesis .. } moreover { assume "ST!n = NT" moreover from frame have "P,h \ stk [:\] ST" by simp with n have "P,h \ stk!n :\ ST!n" by (simp add: list_all2_conv_all_nth) ultimately have "stk!n = Null" by simp with ins no_xcp meth_C have False by (simp add: split_beta) hence ?thesis .. } moreover { assume NT: "ST!n \ NT" and Null: "stk!n \ Null" from NT ins wti \_pc obtain D D' b Ts T m ST' LT' where D: "ST!n = Class D" and pc': "pc+1 < size ins" and m_D: "P \ D sees M',b: Ts\T = m in D'" and Ts: "P \ rev (take n ST) [\] Ts" and \': "\ C M ! (pc+1) = Some (ST', LT')" and LT': "P \ LT [\\<^sub>\] LT'" and ST': "P \ (T # drop (n+1) ST) [\] ST'" and b[simp]: "b = NonStatic" by (clarsimp simp: sup_state_opt_any_Some) from frame obtain stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" by simp from n stk D have "P,h \ stk!n :\ Class D" by (auto simp: list_all2_conv_all_nth) with Null obtain a C' fs where Addr: "stk!n = Addr a" and obj: "h a = Some (C',fs)" and C'subD: "P \ C' \\<^sup>* D" by (fastforce dest!: conf_ClassD) with wfprog m_D no_xcp obtain Ts' T' D'' mxs' mxl' ins' xt' where m_C': "P \ C' sees M',NonStatic: Ts'\T' = (mxs',mxl',ins',xt') in D''" and T': "P \ T' \ T" and Ts': "P \ Ts [\] Ts'" by (auto dest: sees_method_mono) with wf_NonStatic_nclinit wtprog have nclinit: "M' \ clinit" by(simp add: wf_jvm_prog_phi_def) have D''subD': "P \ D'' \\<^sup>* D'" by(rule sees_method_decl_mono[OF C'subD m_D m_C']) let ?loc' = "Addr a # rev (take n stk) @ replicate mxl' undefined" let ?f' = "([], ?loc', D'', M', 0, No_ics)" let ?f = "(stk, loc, C, M, pc, ics)" from Addr obj m_C' ins \' meth_C no_xcp have s': "\' = (None, h, ?f' # ?f # frs, sh)" by simp from Ts n have [simp]: "size Ts = n" by (auto dest: list_all2_lengthD simp: min_def) with Ts' have [simp]: "size Ts' = n" by (auto dest: list_all2_lengthD) from m_C' wfprog obtain mD'': "P \ D'' sees M',NonStatic:Ts'\T'=(mxs',mxl',ins',xt') in D''" by (fast dest: sees_method_idemp) moreover with wtprog obtain start: "wt_start P D'' NonStatic Ts' mxl' (\ D'' M')" and ins': "ins' \ []" by (auto dest: wt_jvm_prog_impl_wt_start) then obtain LT\<^sub>0 where LT\<^sub>0: "\ D'' M' ! 0 = Some ([], LT\<^sub>0)" by (clarsimp simp: wt_start_def defs1 sup_state_opt_any_Some split: staticb.splits) moreover have "conf_f P h sh ([], LT\<^sub>0) ins' ?f'" proof - let ?LT = "OK (Class D'') # (map OK Ts') @ (replicate mxl' Err)" from stk have "P,h \ take n stk [:\] take n ST" .. hence "P,h \ rev (take n stk) [:\] rev (take n ST)" by simp also note Ts also note Ts' finally have "P,h \ rev (take n stk) [:\\<^sub>\] map OK Ts'" by simp also have "P,h \ replicate mxl' undefined [:\\<^sub>\] replicate mxl' Err" by simp also from m_C' have "P \ C' \\<^sup>* D''" by (rule sees_method_decl_above) with obj have "P,h \ Addr a :\ Class D''" by (simp add: conf_def) ultimately have "P,h \ ?loc' [:\\<^sub>\] ?LT" by simp also from start LT\<^sub>0 have "P \ \ [\\<^sub>\] LT\<^sub>0" by (simp add: wt_start_def) finally have "P,h \ ?loc' [:\\<^sub>\] LT\<^sub>0" . thus ?thesis using ins' nclinit by simp qed moreover have "conf_clinit P sh (?f'#?f#frs)" using conf_clinit_Invoke[OF confc nclinit] by simp ultimately have ?thesis using s' \_pc approx meth_C m_D T' ins D nclinit D''subD' by(fastforce dest: sees_method_fun [of _ C]) } ultimately show ?thesis by blast qed (*>*) lemma Invokestatic_nInit_correct: fixes \' :: jvm_state assumes wtprog: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes meth_C: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes ins: "ins ! pc = Invokestatic D M' n" and nclinit: "M' \ clinit" assumes wti: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes \': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" assumes approx: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" assumes no_xcp: "fst (exec_step P h stk loc C M pc ics frs sh) = None" assumes cs: "ics = Called [] \ (ics = No_ics \ (\sfs. sh (fst(method P D M')) = Some(sfs, Done)))" shows "P,\ \ \'\" (*<*) proof - note split_paired_Ex [simp del] from wtprog obtain wfmb where wfprog: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) from ins meth_C approx obtain ST LT where heap_ok: "P\ h\" and \_pc: "\ C M!pc = Some (ST,LT)" and frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,ics)" and frames: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" by (fastforce dest: sees_method_fun) from ins wti \_pc have n: "n \ size ST" by simp from ins wti \_pc obtain D' b Ts T mxs' mxl' ins' xt' ST' LT' where pc': "pc+1 < size ins" and m_D: "P \ D sees M',b: Ts\T = (mxs',mxl',ins',xt') in D'" and Ts: "P \ rev (take n ST) [\] Ts" and \': "\ C M ! (pc+1) = Some (ST', LT')" and LT': "P \ LT [\\<^sub>\] LT'" and ST': "P \ (T # drop n ST) [\] ST'" and b[simp]: "b = Static" by (clarsimp simp: sup_state_opt_any_Some) from frame obtain stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" by simp let ?loc' = "rev (take n stk) @ replicate mxl' undefined" let ?f' = "([], ?loc', D', M', 0, No_ics)" let ?f = "(stk, loc, C, M, pc, No_ics)" from m_D ins \' meth_C no_xcp cs have s': "\' = (None, h, ?f' # ?f # frs, sh)" by auto from Ts n have [simp]: "size Ts = n" by (auto dest: list_all2_lengthD) from m_D wfprog b obtain mD': "P \ D' sees M',Static:Ts\T=(mxs',mxl',ins',xt') in D'" by (fast dest: sees_method_idemp) moreover with wtprog obtain start: "wt_start P D' Static Ts mxl' (\ D' M')" and ins': "ins' \ []" by (auto dest: wt_jvm_prog_impl_wt_start) then obtain LT\<^sub>0 where LT\<^sub>0: "\ D' M' ! 0 = Some ([], LT\<^sub>0)" by (clarsimp simp: wt_start_def defs1 sup_state_opt_any_Some split: staticb.splits) moreover have "conf_f P h sh ([], LT\<^sub>0) ins' ?f'" proof - let ?LT = "(map OK Ts) @ (replicate mxl' Err)" from stk have "P,h \ take n stk [:\] take n ST" .. hence "P,h \ rev (take n stk) [:\] rev (take n ST)" by simp also note Ts finally have "P,h \ rev (take n stk) [:\\<^sub>\] map OK Ts" by simp also have "P,h \ replicate mxl' undefined [:\\<^sub>\] replicate mxl' Err" by simp also from m_D have "P \ D \\<^sup>* D'" by (rule sees_method_decl_above) ultimately have "P,h \ ?loc' [:\\<^sub>\] ?LT" by simp also from start LT\<^sub>0 have "P \ \ [\\<^sub>\] LT\<^sub>0" by (simp add: wt_start_def) finally have "P,h \ ?loc' [:\\<^sub>\] LT\<^sub>0" . thus ?thesis using ins' by simp qed moreover have "conf_clinit P sh (?f'#?f#frs)" by(rule conf_clinit_Invoke[OF confc nclinit]) ultimately show ?thesis using s' \_pc approx meth_C m_D ins nclinit by (fastforce dest: sees_method_fun [of _ C]) qed (*>*) lemma Invokestatic_Init_correct: fixes \' :: jvm_state assumes wtprog: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes meth_C: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes ins: "ins ! pc = Invokestatic D M' n" and nclinit: "M' \ clinit" assumes wti: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes \': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)" assumes approx: "P,\ \ (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\" assumes no_xcp: "fst (exec_step P h stk loc C M pc No_ics frs sh) = None" assumes nDone: "\sfs. sh (fst(method P D M')) \ Some(sfs, Done)" shows "P,\ \ \'\" (*<*) proof - note split_paired_Ex [simp del] from wtprog obtain wfmb where wfprog: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) from ins meth_C approx obtain ST LT where heap_ok: "P\ h\" and \_pc: "\ C M!pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and frames: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,No_ics)#frs)" and pc: "pc < size ins" by (fastforce dest: sees_method_fun) from ins wti \_pc obtain D' b Ts T mxs' mxl' ins' xt' where m_D: "P \ D sees M',b: Ts\T = (mxs',mxl',ins',xt') in D'" and b[simp]: "b = Static" by clarsimp let ?f = "(stk, loc, C, M, pc, Calling D' [])" from m_D ins \' meth_C no_xcp nDone have s': "\' = (None, h, ?f # frs, sh)" by(auto split: init_state.splits) have cls: "is_class P D'" by(rule sees_method_is_class'[OF m_D]) from confc have confc': "conf_clinit P sh (?f#frs)" by(auto simp: conf_clinit_def distinct_clinit_def split: if_split_asm) with s' \_pc approx meth_C m_D ins nclinit stk loc pc cls frames show ?thesis by(fastforce dest: sees_method_fun [of _ C]) qed (*>*) declare list_all2_Cons2 [iff] lemma Return_correct: fixes \' :: jvm_state assumes wt_prog: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes meth: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes ins: "ins ! pc = Return" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" assumes correct: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" shows "P,\ \ \'\" (*<*) proof - from meth correct ins have [simp]: "ics = No_ics" by(cases ics, auto) from wt_prog obtain wfmb where wf: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) from meth ins s' have "frs = [] \ ?thesis" by (simp add: correct_state_def) moreover { fix f frs' assume frs': "frs = f#frs'" then obtain stk' loc' C' M' pc' ics' where f: "f = (stk',loc',C',M',pc',ics')" by (cases f) from correct meth obtain ST LT where h_ok: "P \ h \" and sh_ok: "P,h \\<^sub>s sh \" and \_pc: "\ C M ! pc = Some (ST, LT)" and frame: "conf_f P h sh (ST, LT) ins (stk,loc,C,M,pc,ics)" and frames: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh frs" by (auto dest: sees_method_fun conf_clinit_Cons simp: correct_state_def) from \_pc ins wt obtain U ST\<^sub>0 where "ST = U # ST\<^sub>0" "P \ U \ T" by (simp add: wt_instr_def app_def) blast with wf frame have hd_stk: "P,h \ hd stk :\ T" by (auto simp: conf_f_def) from f frs' frames meth obtain ST' LT' b' Ts'' T'' mxs' mxl\<^sub>0' ins' xt' where \': "\ C' M' ! pc' = Some (ST', LT')" and meth_C': "P \ C' sees M',b':Ts''\T''=(mxs',mxl\<^sub>0',ins',xt') in C'" and frame': "conf_f P h sh (ST',LT') ins' f" and conf_fs: "conf_fs P h sh \ C' M' (size Ts'') T'' frs'" by clarsimp from f frame' obtain stk': "P,h \ stk' [:\] ST'" and loc': "P,h \ loc' [:\\<^sub>\] LT'" and pc': "pc' < size ins'" by (simp add: conf_f_def) { assume b[simp]: "b = NonStatic" from wf_NonStatic_nclinit[OF wf] meth have nclinit[simp]: "M \ clinit" by simp from f frs' meth ins s' have \': "\' = (None,h,(hd stk#(drop (1+size Ts) stk'),loc',C',M',pc'+1,ics')#frs',sh)" (is "\' = (None,h,?f'#frs',sh)") by simp from f frs' confc conf_clinit_diff have confc'': "conf_clinit P sh (?f'#frs')" by blast with \' meth_C' f frs' frames meth obtain D Ts' T' m D' where ins': "ins' ! pc' = Invoke M (size Ts)" and D: "ST' ! (size Ts) = Class D" and meth_D: "P \ D sees M,b: Ts'\T' = m in D'" and T': "P \ T \ T'" and CsubD': "P \ C \\<^sup>* D'" by(auto dest: sees_method_fun sees_method_fun[OF sees_method_idemp]) from wt_prog meth_C' pc' have "P,T'',mxs',size ins',xt' \ ins'!pc',pc' :: \ C' M'" by (rule wt_jvm_prog_impl_wt_instr) with ins' \' D meth_D obtain ST'' LT'' where \_suc: "\ C' M' ! Suc pc' = Some (ST'', LT'')" and less: "P \ (T' # drop (size Ts+1) ST', LT') \\<^sub>i (ST'', LT'')" and suc_pc': "Suc pc' < size ins'" by (clarsimp simp: sup_state_opt_any_Some) from hd_stk T' have hd_stk': "P,h \ hd stk :\ T'" .. have frame'': "conf_f P h sh (ST'',LT'') ins' ?f'" proof - from stk' have "P,h \ drop (1+size Ts) stk' [:\] drop (1+size Ts) ST'" .. moreover with hd_stk' less have "P,h \ hd stk # drop (1+size Ts) stk' [:\] ST''" by auto moreover from wf loc' less have "P,h \ loc' [:\\<^sub>\] LT''" by auto moreover note suc_pc' moreover from f frs' frames (* ics' = No_ics *) have "P,h,sh \\<^sub>i (C', M', Suc pc', ics')" by auto ultimately show ?thesis by (simp add: conf_f_def) qed with \' frs' f meth h_ok sh_ok hd_stk \_suc frames confc'' meth_C' \' have ?thesis by(fastforce dest: sees_method_fun [of _ C']) } moreover { assume b[simp]: "b = Static" and nclinit[simp]: "M \ clinit" from f frs' meth ins s' have \': "\' = (None,h,(hd stk#(drop (size Ts) stk'),loc',C',M',pc'+1,ics')#frs',sh)" (is "\' = (None,h,?f'#frs',sh)") by simp from f frs' confc conf_clinit_diff have confc'': "conf_clinit P sh (?f'#frs')" by blast with \' meth_C' f frs' frames meth obtain D Ts' T' m where ins': "ins' ! pc' = Invokestatic D M (size Ts)" and meth_D: "P \ D sees M,b: Ts'\T' = m in C" and T': "P \ T \ T'" by(auto dest: sees_method_fun sees_method_mono2[OF _ wf sees_method_idemp]) from wt_prog meth_C' pc' have "P,T'',mxs',size ins',xt' \ ins'!pc',pc' :: \ C' M'" by (rule wt_jvm_prog_impl_wt_instr) with ins' \' meth_D obtain ST'' LT'' where \_suc: "\ C' M' ! Suc pc' = Some (ST'', LT'')" and less: "P \ (T' # drop (size Ts) ST', LT') \\<^sub>i (ST'', LT'')" and suc_pc': "Suc pc' < size ins'" by (clarsimp simp: sup_state_opt_any_Some) from hd_stk T' have hd_stk': "P,h \ hd stk :\ T'" .. have frame'': "conf_f P h sh (ST'',LT'') ins' ?f'" proof - from stk' have "P,h \ drop (size Ts) stk' [:\] drop (size Ts) ST'" .. moreover with hd_stk' less have "P,h \ hd stk # drop (size Ts) stk' [:\] ST''" by auto moreover from wf loc' less have "P,h \ loc' [:\\<^sub>\] LT''" by auto moreover note suc_pc' moreover from f frs' frames (* ics' = No_ics *) have "P,h,sh \\<^sub>i (C', M', Suc pc', ics')" by auto ultimately show ?thesis by (simp add: conf_f_def) qed with \' frs' f meth h_ok sh_ok hd_stk \_suc frames confc'' meth_C' \' have ?thesis by(fastforce dest: sees_method_fun [of _ C']) } moreover { assume b[simp]: "b = Static" and clinit[simp]: "M = clinit" from frs' meth ins s' have \': "\' = (None,h,frs,sh(C\(fst(the(sh C)), Done)))" (is "\' = (None,h,frs,?sh)") by simp from correct have dist': "distinct (C # clinit_classes frs)" by(simp add: conf_clinit_def distinct_clinit_def) from f frs' correct have confc1: "conf_clinit P sh ((stk, loc, C, clinit, pc, No_ics) # (stk',loc',C',M',pc',ics') # frs')" by simp then have ics_dist: "distinct (C # ics_classes ics')" by(simp add: conf_clinit_def distinct_clinit_def) from conf_clinit_Cons_Cons[OF confc1] have dist'': "distinct (C # clinit_classes frs')" by(simp add: conf_clinit_def distinct_clinit_def) from correct shconf_upd_obj[OF sh_ok _ [OF shconfD[OF sh_ok]]] have sh'_ok: "P,h \\<^sub>s ?sh \" by(clarsimp simp: conf_clinit_def) have frame'': "conf_f P h ?sh (ST',LT') ins' f" proof - note stk' loc' pc' f valid_ics_shupd[OF _ ics_dist] moreover from f frs' frames have "P,h,sh \\<^sub>i (C', M', pc', ics')" by auto ultimately show ?thesis by (simp add: conf_f_def2) qed have conf_fs': "conf_fs P h ?sh \ C' M' (length Ts'') T'' frs'" by(rule conf_fs_shupd[OF conf_fs dist'']) have confc'': "conf_clinit P ?sh frs" by(rule conf_clinit_shupd[OF confc dist']) from \' f frs' h_ok sh'_ok conf_fs' frame'' \' stk' loc' pc' meth_C' confc'' have ?thesis by(fastforce dest: sees_method_fun) } ultimately have ?thesis by (cases b) blast+ } ultimately show ?thesis by (cases frs) blast+ qed (*>*) declare sup_state_opt_any_Some [iff] declare not_Err_eq [iff] lemma Load_correct: -"\ wf_prog wt P; - P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; - ins!pc = Load idx; - P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; - Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh); - P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\ \ -\ P,\ \ \'\" +assumes "wf_prog wt P" and + mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" and + i: "ins!pc = Load idx" and + "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" and + "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and + cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" +shows "P,\ |- \' [ok]" (*<*) - apply(subgoal_tac "ics = No_ics") - prefer 2 apply(cases ics, (auto)[4]) - apply clarsimp - apply (drule (1) sees_method_fun) - apply(fastforce elim!: confTs_confT_sup conf_clinit_diff) - done +proof - + have "ics = No_ics" using mC i cf by(cases ics) auto + then show ?thesis using assms(1,3-6) sees_method_fun[OF mC] + by(fastforce elim!: confTs_confT_sup conf_clinit_diff) +qed (*>*) declare [[simproc del: list_to_set_comprehension]] lemma Store_correct: -"\ wf_prog wt P; - P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; - ins!pc = Store idx; - P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; - Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh); - P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\ \ -\ P,\ \ \'\" +assumes "wf_prog wt P" and + mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" and + i: "ins!pc = Store idx" and + "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" and + "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and + cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" +shows "P,\ |- \' [ok]" (*<*) - apply(subgoal_tac "ics = No_ics") - prefer 2 apply(cases ics, (auto)[4]) - apply clarsimp - apply (drule (1) sees_method_fun) - apply (blast intro!: list_all2_update_cong conf_clinit_diff)+ - done +proof - + have "ics = No_ics" using mC i cf by(cases ics) auto + then show ?thesis using assms(1,3-6) sees_method_fun[OF mC] + by clarsimp (blast intro!: list_all2_update_cong conf_clinit_diff) +qed (*>*) lemma Push_correct: -"\ wf_prog wt P; - P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; - ins!pc = Push v; - P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; - Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh); - P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\ \ -\ P,\ \ \'\" +assumes "wf_prog wt P" and + mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" and + i: "ins!pc = Push v" and + "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" and + "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and + cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" +shows "P,\ |- \' [ok]" (*<*) - apply(subgoal_tac "ics = No_ics") - prefer 2 apply(cases ics, (auto)[4]) - apply clarsimp - apply (drule (1) sees_method_fun) - apply (blast dest: typeof_lit_conf conf_clinit_diff)+ - done -(*>*) +proof - + have "ics = No_ics" using mC i cf by(cases ics) auto + then show ?thesis using assms(1,3-6) sees_method_fun[OF mC] + by clarsimp (blast dest: typeof_lit_conf conf_clinit_diff) +qed lemma Cast_conf2: "\ wf_prog ok P; P,h \ v :\ T; is_refT T; cast_ok P C h v; P \ Class C \ T'; is_class P C\ \ P,h \ v :\ T'" (*<*) - apply (unfold cast_ok_def is_refT_def) - apply (frule Class_widen) - apply (elim exE disjE) - apply simp - apply simp - apply simp - apply (clarsimp simp: conf_def obj_ty_def) - apply (cases v) - apply (auto intro: rtrancl_trans) - done +proof - + assume "wf_prog ok P" and "P,h \ v :\ T" and "is_refT T" and + "cast_ok P C h v" and wid: "P \ Class C \ T'" and "is_class P C" + then show "P,h \ v :\ T'" using Class_widen[OF wid] + by(cases v) + (auto simp: cast_ok_def is_refT_def conf_def obj_ty_def + intro: rtrancl_trans) +qed (*>*) lemma Checkcast_correct: -"\ wf_jvm_prog\<^bsub>\\<^esub> P; - P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; - ins!pc = Checkcast D; - P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; - Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh) ; - P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\; - fst (exec_step P h stk loc C M pc ics frs sh) = None \ -\ P,\ \ \'\" +assumes "wf_jvm_prog\<^bsub>\\<^esub> P" and + mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" and + i: "ins!pc = Checkcast D" and + "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" and + "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and + cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" and + "fst (exec_step P h stk loc C M pc ics frs sh) = None" +shows "P,\ |- \' [ok]" (*<*) - apply(subgoal_tac "ics = No_ics") - prefer 2 apply(cases ics, (auto)[4]) - apply (clarsimp simp: wf_jvm_prog_phi_def split: if_split_asm) - apply (drule (1) sees_method_fun) - apply (blast intro: Cast_conf2 dest: sees_method_fun conf_clinit_diff) - done +proof - + have "ics = No_ics" using mC i cf by(cases ics) auto + then show ?thesis using assms + by (clarsimp simp: wf_jvm_prog_phi_def split: if_split_asm) + (blast intro: Cast_conf2 dest: sees_method_fun conf_clinit_diff) +qed (*>*) declare split_paired_All [simp del] lemmas widens_Cons [iff] = list_all2_Cons1 [of "widen P"] for P lemma Getfield_correct: fixes \' :: jvm_state assumes wf: "wf_prog wt P" assumes mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes i: "ins!pc = Getfield F D" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" assumes cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None" shows "P,\ \ \'\" (*<*) proof - from mC cf i have [simp]: "ics = No_ics" by(cases ics, auto) from mC cf obtain ST LT where "h\": "P \ h \" and "sh\": "P,h \\<^sub>s sh \" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and fs: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" by (fastforce dest: sees_method_fun) from i \ wt obtain oT ST'' vT ST' LT' vT' where oT: "P \ oT \ Class D" and ST: "ST = oT # ST''" and F: "P \ D sees F,NonStatic:vT in D" and pc': "pc+1 < size ins" and \': "\ C M ! (pc+1) = Some (vT'#ST', LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" and vT': "P \ vT \ vT'" by fastforce from stk ST obtain ref stk' where stk': "stk = ref#stk'" and ref: "P,h \ ref :\ oT" and ST'': "P,h \ stk' [:\] ST''" by auto from stk' i mC s' xc have "ref \ Null" by (simp add: split_beta split:if_split_asm) moreover from ref oT have "P,h \ ref :\ Class D" .. ultimately obtain a D' fs where a: "ref = Addr a" and h: "h a = Some (D', fs)" and D': "P \ D' \\<^sup>* D" by (blast dest: non_npD) from D' F have has_field: "P \ D' has F,NonStatic:vT in D" by (blast intro: has_field_mono has_visible_field) moreover from "h\" h have "P,h \ (D', fs) \" by (rule hconfD) ultimately obtain v where v: "fs (F, D) = Some v" "P,h \ v :\ vT" by (clarsimp simp: oconf_def has_field_def) (blast dest: has_fields_fun) from conf_clinit_diff[OF confc] have confc': "conf_clinit P sh ((v#stk',loc,C,M,pc+1,ics)#frs)" by simp from a h i mC s' stk' v xc has_field have "\' = (None, h, (v#stk',loc,C,M,pc+1,ics)#frs, sh)" by(simp add: split_beta split: if_split_asm) moreover from ST'' ST' have "P,h \ stk' [:\] ST'" .. moreover from v vT' have "P,h \ v :\ vT'" by blast moreover from loc LT' have "P,h \ loc [:\\<^sub>\] LT'" .. moreover note "h\" "sh\" mC \' pc' v fs confc' ultimately show "P,\ \ \' \" by fastforce qed (*>*) lemma Getstatic_nInit_correct: fixes \' :: jvm_state assumes wf: "wf_prog wt P" assumes mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes i: "ins!pc = Getstatic C' F D" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" assumes cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None" assumes cs: "ics = Called [] \ (ics = No_ics \ (\sfs. sh (fst(field P D F)) = Some(sfs, Done)))" shows "P,\ \ \'\" (*<*) proof - from mC cf obtain ST LT where "h\": "P \ h \" and "sh\": "P,h \\<^sub>s sh \" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and fs: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and vics: "P,h,sh \\<^sub>i (C,M,pc,ics)" by (fastforce dest: sees_method_fun) from i \ wt cs obtain vT ST' LT' vT' where F: "P \ C' sees F,Static:vT in D" and pc': "pc+1 < size ins" and \': "\ C M ! (pc+1) = Some (vT'#ST', LT')" and ST': "P \ ST [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" and vT': "P \ vT \ vT'" by fastforce with mC i vics obtain sobj where cc': "ics = Called [] \ Called_context P D (ins!pc) \ sh D = Some sobj" by(fastforce dest: has_visible_field) from field_def2[OF sees_field_idemp[OF F]] have D[simp]: "fst(field P D F) = D" by simp from cs cc' obtain sfs i where shD: "sh D = Some(sfs,i)" by(cases sobj, auto) note has_field_idemp[OF has_visible_field[OF F]] moreover from "sh\" shD have "P,h,D \\<^sub>s sfs \" by (rule shconfD) ultimately obtain v where v: "sfs F = Some v" "P,h \ v :\ vT" by (clarsimp simp: soconf_def has_field_def) blast from i mC s' v xc F cs cc' shD have "\' = (None, h, (v#stk,loc,C,M,pc+1,No_ics)#frs, sh)" by(fastforce simp: split_beta split: if_split_asm init_call_status.splits) moreover from stk ST' have "P,h \ stk [:\] ST'" .. moreover from v vT' have "P,h \ v :\ vT'" by blast moreover from loc LT' have "P,h \ loc [:\\<^sub>\] LT'" .. moreover have "conf_clinit P sh ((v#stk,loc,C,M,pc+1,No_ics)#frs)" by(rule conf_clinit_diff'[OF confc]) moreover note "h\" "sh\" mC \' pc' v fs ultimately show "P,\ \ \' \" by fastforce qed (*>*) lemma Getstatic_Init_correct: fixes \' :: jvm_state assumes wf: "wf_prog wt P" assumes mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes i: "ins!pc = Getstatic C' F D" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)" assumes cf: "P,\ \ (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\" assumes xc: "fst (exec_step P h stk loc C M pc No_ics frs sh) = None" assumes nDone: "\sfs. sh (fst(field P D F)) \ Some(sfs, Done)" shows "P,\ \ \'\" (*<*) proof - from mC cf obtain ST LT where "h\": "P \ h \" and "sh\": "P,h \\<^sub>s sh \" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and fs: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,No_ics)#frs)" by (fastforce dest: sees_method_fun) from i \ wt nDone obtain vT where F: "P \ C' sees F,Static:vT in D" by fastforce then have has_field: "P \ C' has F,Static:vT in D" by(rule has_visible_field) from field_def2[OF sees_field_idemp[OF F]] has_field_is_class'[OF has_field] obtain D[simp]: "fst(field P D F) = D" and cls: "is_class P D" by simp from i mC s' xc F nDone have "\' = (None, h, (stk,loc,C,M,pc,Calling D [])#frs, sh)" by(auto simp: split_beta split: if_split_asm init_state.splits) moreover from confc have "conf_clinit P sh ((stk,loc,C,M,pc,Calling D [])#frs)" by(auto simp: conf_clinit_def distinct_clinit_def split: if_split_asm) moreover note loc stk "h\" "sh\" mC \ pc fs i has_field cls ultimately show "P,\ \ \' \" by fastforce qed (*>*) lemma Putfield_correct: fixes \' :: jvm_state assumes wf: "wf_prog wt P" assumes mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes i: "ins!pc = Putfield F D" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" assumes cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None" shows "P,\ \ \'\" (*<*) proof - from mC cf i have [simp]: "ics = No_ics" by(cases ics, auto) from mC cf obtain ST LT where "h\": "P \ h \" and "sh\": "P,h \\<^sub>s sh \" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and fs: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics) # frs)" by (fastforce dest: sees_method_fun) from i \ wt obtain vT vT' oT ST'' ST' LT' where ST: "ST = vT # oT # ST''" and field: "P \ D sees F,NonStatic:vT' in D" and oT: "P \ oT \ Class D" and vT: "P \ vT \ vT'" and pc': "pc+1 < size ins" and \': "\ C M!(pc+1) = Some (ST',LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" by clarsimp from stk ST obtain v ref stk' where stk': "stk = v#ref#stk'" and v: "P,h \ v :\ vT" and ref: "P,h \ ref :\ oT" and ST'': "P,h \ stk' [:\] ST''" by auto from stk' i mC s' xc have "ref \ Null" by (auto simp: split_beta) moreover from ref oT have "P,h \ ref :\ Class D" .. ultimately obtain a D' fs where a: "ref = Addr a" and h: "h a = Some (D', fs)" and D': "P \ D' \\<^sup>* D" by (blast dest: non_npD) from v vT have vT': "P,h \ v :\ vT'" .. from field D' have has_field: "P \ D' has F,NonStatic:vT' in D" by (blast intro: has_field_mono has_visible_field) let ?h' = "h(a\(D', fs((F, D)\v)))" and ?f' = "(stk',loc,C,M,pc+1,ics)" from h have hext: "h \ ?h'" by (rule hext_upd_obj) have "sh\'": "P,?h' \\<^sub>s sh \" by(rule shconf_hupd_obj[OF "sh\" h]) from a h i mC s' stk' has_field field have "\' = (None, ?h', ?f'#frs, sh)" by(simp split: if_split_asm) moreover from "h\" h have "P,h \ (D',fs)\" by (rule hconfD) with has_field vT' have "P,h \ (D',fs((F, D)\v))\" .. with "h\" h have "P \ ?h'\" by (rule hconf_upd_obj) moreover from ST'' ST' have "P,h \ stk' [:\] ST'" .. from this hext have "P,?h' \ stk' [:\] ST'" by (rule confs_hext) moreover from loc LT' have "P,h \ loc [:\\<^sub>\] LT'" .. from this hext have "P,?h' \ loc [:\\<^sub>\] LT'" by (rule confTs_hext) moreover from fs hext have "conf_fs P ?h' sh \ C M (size Ts) T frs" by (rule conf_fs_hext) moreover have "conf_clinit P sh (?f' # frs)" by(rule conf_clinit_diff[OF confc]) moreover note mC \' pc' "sh\'" ultimately show "P,\ \ \' \" by fastforce qed (*>*) lemma Putstatic_nInit_correct: fixes \' :: jvm_state assumes wf: "wf_prog wt P" assumes mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes i: "ins!pc = Putstatic C' F D" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" assumes cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None" assumes cs: "ics = Called [] \ (ics = No_ics \ (\sfs. sh (fst(field P D F)) = Some(sfs, Done)))" shows "P,\ \ \'\" (*<*) proof - from mC cf obtain ST LT where "h\": "P \ h \" and "sh\": "P,h \\<^sub>s sh \" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and fs: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and vics: "P,h,sh \\<^sub>i (C,M,pc,ics)" by (fastforce dest: sees_method_fun) from i \ wt cs obtain vT vT' ST'' ST' LT' where ST: "ST = vT # ST''" and F: "P \ C' sees F,Static:vT' in D" and vT: "P \ vT \ vT'" and pc': "pc+1 < size ins" and \': "\ C M ! (pc+1) = Some (ST', LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" by fastforce from stk ST obtain v stk' where stk': "stk = v#stk'" and v: "P,h \ v :\ vT" and ST'': "P,h \ stk' [:\] ST''" by auto from v vT have vT': "P,h \ v :\ vT'" .. with mC i vics obtain sobj where cc': "ics = Called [] \ Called_context P D (ins!pc) \ sh D = Some sobj" by(fastforce dest: has_visible_field) from field_def2[OF sees_field_idemp[OF F]] have D[simp]: "fst(field P D F) = D" by simp from cs cc' obtain sfs i where shD: "sh D = Some(sfs,i)" by(cases sobj, auto) let ?sh' = "sh(D\(sfs(F\v),i))" and ?f' = "(stk',loc,C,M,pc+1,No_ics)" have m_D: "P \ D has F,Static:vT' in D" by (rule has_field_idemp[OF has_visible_field[OF F]]) from "sh\" shD have sfs: "P,h,D \\<^sub>s sfs \" by (rule shconfD) have "sh'\": "P,h \\<^sub>s ?sh' \" by (rule shconf_upd_obj[OF "sh\" soconf_fupd[OF m_D vT' sfs]]) from i mC s' v xc F cs cc' shD stk' have "\' = (None, h, (stk',loc,C,M,pc+1,No_ics)#frs, ?sh')" by(fastforce simp: split_beta split: if_split_asm init_call_status.splits) moreover from ST'' ST' have "P,h \ stk' [:\] ST'" .. moreover from loc LT' have "P,h \ loc [:\\<^sub>\] LT'" .. moreover have "conf_fs P h ?sh' \ C M (size Ts) T frs" by (rule conf_fs_shupd'[OF fs shD]) moreover have "conf_clinit P ?sh' ((stk',loc,C,M,pc+1,No_ics)#frs)" by(rule conf_clinit_diff'[OF conf_clinit_shupd'[OF confc shD]]) moreover note "h\" "sh'\" mC \' pc' v vT' ultimately show "P,\ \ \' \" by fastforce qed (*>*) lemma Putstatic_Init_correct: fixes \' :: jvm_state assumes wf: "wf_prog wt P" assumes mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes i: "ins!pc = Putstatic C' F D" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)" assumes cf: "P,\ \ (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\" assumes xc: "fst (exec_step P h stk loc C M pc No_ics frs sh) = None" assumes nDone: "\sfs. sh (fst(field P D F)) \ Some(sfs, Done)" shows "P,\ \ \'\" (*<*) proof - from mC cf obtain ST LT where "h\": "P \ h \" and "sh\": "P,h \\<^sub>s sh \" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and fs: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,No_ics)#frs)" by (fastforce dest: sees_method_fun) from i \ wt nDone obtain vT where F: "P \ C' sees F,Static:vT in D" by fastforce then have has_field: "P \ C' has F,Static:vT in D" by(rule has_visible_field) from field_def2[OF sees_field_idemp[OF F]] has_field_is_class'[OF has_field] obtain D[simp]: "fst(field P D F) = D" and cls: "is_class P D" by simp from i mC s' xc F nDone have "\' = (None, h, (stk,loc,C,M,pc,Calling D [])#frs, sh)" by(auto simp: split_beta split: if_split_asm init_state.splits) moreover from confc have "conf_clinit P sh ((stk,loc,C,M,pc,Calling D [])#frs)" by(auto simp: conf_clinit_def distinct_clinit_def split: if_split_asm) moreover note loc stk "h\" "sh\" mC \ pc fs i has_field cls ultimately show "P,\ \ \' \" by fastforce qed (*>*) (* FIXME: move *) lemma oconf_blank2 [intro, simp]: "\is_class P C; wf_prog wt P\ \ P,h \ blank P C \" (*<*) by (fastforce simp: oconf_blank dest: wf_Fields_Ex) (*>*) lemma obj_ty_blank [iff]: "obj_ty (blank P C) = Class C" by (simp add: blank_def) lemma New_nInit_correct: fixes \' :: jvm_state assumes wf: "wf_prog wt P" assumes meth: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes ins: "ins!pc = New X" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes exec: "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" assumes conf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" assumes no_x: "fst (exec_step P h stk loc C M pc ics frs sh) = None" assumes cs: "ics = Called [] \ (ics = No_ics \ (\sfs. sh X = Some(sfs, Done)))" shows "P,\ \ \'\" (*<*) proof - from ins conf meth obtain ST LT where heap_ok: "P\ h\" and sheap_ok: "P,h \\<^sub>s sh \" and \_pc: "\ C M!pc = Some (ST,LT)" and frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,ics)" and frames: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics) # frs)" by (auto dest: sees_method_fun) from \_pc ins wt obtain ST' LT' where is_class_X: "is_class P X" and mxs: "size ST < mxs" and suc_pc: "pc+1 < size ins" and \_suc: "\ C M!(pc+1) = Some (ST', LT')" and less: "P \ (Class X # ST, LT) \\<^sub>i (ST', LT')" by auto from ins no_x cs meth obtain oref where new_Addr: "new_Addr h = Some oref" by auto hence h: "h oref = None" by (rule new_Addr_SomeD) with exec ins meth new_Addr cs have \': "\' = (None, h(oref \ blank P X), (Addr oref#stk,loc,C,M,pc+1,No_ics)#frs, sh)" (is "\' = (None, ?h', ?f # frs, sh)") by auto moreover from wf h heap_ok is_class_X have h': "P \ ?h' \" by (auto intro: hconf_new) moreover from h frame less suc_pc wf have "conf_f P ?h' sh (ST', LT') ins ?f" - apply (clarsimp simp: fun_upd_apply conf_def blank_def split_beta) - apply (auto intro: confs_hext confTs_hext) - done + by (clarsimp simp: fun_upd_apply conf_def blank_def split_beta) + (auto intro: confs_hext confTs_hext) moreover from h have "h \ ?h'" by simp with frames have "conf_fs P ?h' sh \ C M (size Ts) T frs" by (rule conf_fs_hext) moreover have "P,?h' \\<^sub>s sh \" by(rule shconf_hnew[OF sheap_ok h]) moreover have "conf_clinit P sh (?f # frs)" by(rule conf_clinit_diff'[OF confc]) ultimately show ?thesis using meth \_suc by fastforce qed (*>*) lemma New_Init_correct: fixes \' :: jvm_state assumes wf: "wf_prog wt P" assumes meth: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes ins: "ins!pc = New X" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes exec: "Some \' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)" assumes conf: "P,\ \ (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\" assumes no_x: "fst (exec_step P h stk loc C M pc No_ics frs sh) = None" assumes nDone: "\sfs. sh X \ Some(sfs, Done)" shows "P,\ \ \'\" (*<*) proof - from ins conf meth obtain ST LT where heap_ok: "P\ h\" and sheap_ok: "P,h \\<^sub>s sh \" and \_pc: "\ C M!pc = Some (ST,LT)" and frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,No_ics)" and frames: "conf_fs P h sh \ C M (size Ts) T frs" and confc: "conf_clinit P sh ((stk,loc,C,M,pc,No_ics) # frs)" by (auto dest: sees_method_fun) from \_pc ins wt obtain ST' LT' where is_class_X: "is_class P X" and mxs: "size ST < mxs" and suc_pc: "pc+1 < size ins" and \_suc: "\ C M!(pc+1) = Some (ST', LT')" and less: "P \ (Class X # ST, LT) \\<^sub>i (ST', LT')" by auto with exec ins meth nDone have \': "\' = (None, h, (stk,loc,C,M,pc,Calling X [])#frs, sh)" (is "\' = (None, h, ?f # frs, sh)") by(auto split: init_state.splits) moreover from meth frame is_class_X ins have "conf_f P h sh (ST, LT) ins ?f" by auto moreover note heap_ok sheap_ok frames moreover from confc have "conf_clinit P sh (?f # frs)" by(auto simp: conf_clinit_def distinct_clinit_def split: if_split_asm) ultimately show ?thesis using meth \_pc by fastforce qed (*>*) lemma Goto_correct: -"\ wf_prog wt P; - P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; - ins ! pc = Goto branch; - P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; - Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh) ; - P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\ \ -\ P,\ \ \'\" +assumes "wf_prog wt P" and + mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" and + i: "ins!pc = Goto branch" and + "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" and + "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and + cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" +shows "P,\ |- \' [ok]" (*<*) -apply(subgoal_tac "ics = No_ics") - prefer 2 apply(cases ics, (auto)[4]) -apply clarsimp -apply (drule (1) sees_method_fun) -apply (fastforce elim!: conf_clinit_diff) -done +proof - + have "ics = No_ics" using mC i cf by(cases ics) auto + then show ?thesis using assms(1,3-6) sees_method_fun[OF mC] + by(fastforce elim!: conf_clinit_diff) +qed (*>*) lemma IfFalse_correct: -"\ wf_prog wt P; - P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; - ins ! pc = IfFalse branch; - P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; - Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh) ; - P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\ \ -\ P,\ \ \'\" +assumes "wf_prog wt P" and + mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" and + i: "ins!pc = IfFalse branch" and + "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" and + "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and + cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" +shows "P,\ |- \' [ok]" (*<*) -apply(subgoal_tac "ics = No_ics") - prefer 2 apply(cases ics, (auto)[4]) -apply clarsimp -apply (drule (1) sees_method_fun) -apply (fastforce elim!: conf_clinit_diff) -done +proof - + have "ics = No_ics" using mC i cf by(cases ics) auto + then show ?thesis using assms(1,3-6) sees_method_fun[OF mC] + by(fastforce elim!: conf_clinit_diff) +qed (*>*) lemma CmpEq_correct: -"\ wf_prog wt P; - P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; - ins ! pc = CmpEq; - P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; - Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh) ; - P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\ \ -\ P,\ \ \'\" +assumes "wf_prog wt P" and + mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" and + i: "ins!pc = CmpEq" and + "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" and + "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and + cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" +shows "P,\ |- \' [ok]" (*<*) -apply(subgoal_tac "ics = No_ics") - prefer 2 apply(cases ics, (auto)[4]) -apply clarsimp -apply (drule (1) sees_method_fun) -apply (fastforce elim!: conf_clinit_diff) -done +proof - + have "ics = No_ics" using mC i cf by(cases ics) auto + then show ?thesis using assms(1,3-6) sees_method_fun[OF mC] + by(fastforce elim!: conf_clinit_diff) +qed (*>*) lemma Pop_correct: -"\ wf_prog wt P; - P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; - ins ! pc = Pop; - P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; - Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh) ; - P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\ \ -\ P,\ \ \'\" +assumes "wf_prog wt P" and + mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" and + i: "ins!pc = Pop" and + "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" and + "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and + cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" +shows "P,\ |- \' [ok]" (*<*) -apply(subgoal_tac "ics = No_ics") - prefer 2 apply(cases ics, (auto)[4]) -apply clarsimp -apply (drule (1) sees_method_fun) -apply (fastforce elim!: conf_clinit_diff) -done +proof - + have "ics = No_ics" using mC i cf by(cases ics) auto + then show ?thesis using assms(1,3-6) sees_method_fun[OF mC] + by(fastforce elim!: conf_clinit_diff) +qed (*>*) lemma IAdd_correct: -"\ wf_prog wt P; - P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; - ins ! pc = IAdd; - P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; - Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh) ; - P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\ \ -\ P,\ \ \'\" +assumes "wf_prog wt P" and + mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" and + i: "ins!pc = IAdd" and + "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" and + "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and + cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" +shows "P,\ |- \' [ok]" (*<*) -apply(subgoal_tac "ics = No_ics") - prefer 2 apply(cases ics, (auto)[4]) -apply (clarsimp simp: conf_def) -apply (drule (1) sees_method_fun) -apply (fastforce elim!: conf_clinit_diff) -done +proof - + have "ics = No_ics" using mC i cf by(cases ics) auto + then show ?thesis using assms(1,3-6) sees_method_fun[OF mC] + by(fastforce elim!: conf_clinit_diff) +qed (*>*) lemma Throw_correct: -"\ wf_prog wt P; - P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; - ins ! pc = Throw; - Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh) ; - P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\; - fst (exec_step P h stk loc C M pc ics frs sh) = None \ -\ P,\ \ \'\" -apply(subgoal_tac "ics = No_ics") - prefer 2 apply(cases ics, (auto)[4]) -apply simp -done +assumes "wf_prog wt P" and + mC: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" and + i: "ins!pc = Throw" and + "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and + cf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" and + "fst (exec_step P h stk loc C M pc ics frs sh) = None" +shows "P,\ |- \' [ok]" +(*<*) +proof - + have "ics = No_ics" using mC i cf by(cases ics) auto + then show ?thesis using assms by simp +qed +(*>*) text \ The next theorem collects the results of the sections above, i.e.~exception handling, initialization procedure steps, and the execution step for each instruction. It states type safety for single step execution: in welltyped programs, a conforming state is transformed into another conforming state when one step of execution is performed. \ lemma step_correct: fixes \' :: jvm_state assumes wtp: "wf_jvm_prog\<^bsub>\\<^esub> P" and meth: "P \ C sees M,b:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" and exec: "Some \' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and conf: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\" shows "P,\ \ \'\" (*<*) proof - from assms have pc: "pc < length ins" by(auto dest: sees_method_fun) with wt_jvm_prog_impl_wt_instr[OF wtp meth] have wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" by simp from conf obtain ST LT where \: "\ C M ! pc = Some(ST,LT)" by clarsimp show ?thesis proof(cases "fst (exec_step P h stk loc C M pc ics frs sh)") case Some show ?thesis by(rule xcpt_correct[OF wtp meth wt Some exec conf]) next case None from wt_jvm_progD[OF wtp] obtain wf_md where wf: "wf_prog wf_md P" by clarify { assume [simp]: "ics = No_ics" from exec conf None obtain exec': "Some \' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)" and conf': "P,\ \ (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\" and None': "fst (exec_step P h stk loc C M pc No_ics frs sh) = None" by simp have ?thesis proof(cases "ins!pc") case Load from Load_correct[OF wf meth this wt exec conf] show ?thesis by simp next case Store from Store_correct[OF wf meth this wt exec conf] show ?thesis by simp next case Push from Push_correct[OF wf meth this wt exec conf] show ?thesis by simp next case (New C) show ?thesis proof(cases "\sfs. sh C = Some(sfs, Done)") case True with New_nInit_correct[OF wf meth New wt exec conf None] show ?thesis by simp next case False with New_Init_correct[OF wf meth New wt exec' conf' None'] show ?thesis by simp qed next case Getfield from Getfield_correct[OF wf meth this wt exec conf None] show ?thesis by simp next case (Getstatic C F D) show ?thesis proof(cases "\sfs. sh (fst (field P D F)) = Some(sfs, Done)") case True with Getstatic_nInit_correct[OF wf meth Getstatic wt exec conf None] show ?thesis by simp next case False with Getstatic_Init_correct[OF wf meth Getstatic wt exec' conf' None'] show ?thesis by simp qed next case Putfield from Putfield_correct[OF wf meth this wt exec conf None] show ?thesis by simp next case (Putstatic C F D) show ?thesis proof(cases "\sfs. sh (fst (field P D F)) = Some(sfs, Done)") case True with Putstatic_nInit_correct[OF wf meth Putstatic wt exec conf None] show ?thesis by simp next case False with Putstatic_Init_correct[OF wf meth Putstatic wt exec' conf' None'] show ?thesis by simp qed next case Checkcast from Checkcast_correct[OF wtp meth this wt exec conf None] show ?thesis by simp next case Invoke with Invoke_correct[OF wtp meth this wt exec conf None] show ?thesis by simp next case (Invokestatic C M n) from wf_jvm_prog_nclinit[OF wtp meth wt pc \ this] have ncl: "M \ clinit" by simp show ?thesis proof(cases "\sfs. sh (fst (method P C M)) = Some(sfs, Done)") case True with Invokestatic_nInit_correct[OF wtp meth Invokestatic ncl wt exec conf None] show ?thesis by simp next case False with Invokestatic_Init_correct[OF wtp meth Invokestatic ncl wt exec' conf' None'] show ?thesis by simp qed next case Return with Return_correct[OF wtp meth this wt exec conf] show ?thesis by simp next case Pop with Pop_correct[OF wf meth this wt exec conf] show ?thesis by simp next case IAdd with IAdd_correct[OF wf meth this wt exec conf] show ?thesis by simp next case Goto with Goto_correct[OF wf meth this wt exec conf] show ?thesis by simp next case CmpEq with CmpEq_correct[OF wf meth this wt exec conf] show ?thesis by simp next case IfFalse with IfFalse_correct[OF wf meth this wt exec conf] show ?thesis by simp next case Throw with Throw_correct[OF wf meth this exec conf None] show ?thesis by simp qed } moreover { fix Cs assume [simp]: "ics = Called Cs" have ?thesis proof(cases Cs) case [simp]: Nil from conf meth obtain C1 where "Called_context P C1 (ins ! pc)" by(clarsimp simp: conf_f_def2 intro!: Called_context_Called_set) then have "ins!pc \ Called_set" by(rule Called_context_Called_set) then show ?thesis proof(cases "ins!pc") case (New C) from New_nInit_correct[OF wf meth this wt exec conf None] show ?thesis by simp next case (Getstatic C F D) from Getstatic_nInit_correct[OF wf meth this wt exec conf None] show ?thesis by simp next case (Putstatic C F D) from Putstatic_nInit_correct[OF wf meth this wt exec conf None] show ?thesis by simp next case (Invokestatic C M n) from wf_jvm_prog_nclinit[OF wtp meth wt pc \ this] have ncl: "M \ clinit" by simp with Invokestatic_nInit_correct[OF wtp meth Invokestatic ncl wt exec conf None] show ?thesis by simp qed(simp_all) next case (Cons C' Cs') with Called_correct[OF wtp meth exec conf None] show ?thesis by simp qed } moreover { fix C' Cs assume [simp]: "ics = Calling C' Cs" with Calling_correct[OF wtp meth exec conf None] have ?thesis by simp } moreover { fix Cs a assume [simp]: "ics = Throwing Cs a" have ?thesis proof(cases Cs) case Nil with exec None show ?thesis by simp next case (Cons C' Cs') with Throwing_correct[OF wtp meth exec conf None] show ?thesis by simp qed } ultimately show ?thesis by(cases ics) auto qed qed (*>*) subsection \ Main \ lemma correct_state_impl_Some_method: "P,\ \ (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\ \ \b m Ts T. P \ C sees M,b:Ts\T = m in C" by fastforce lemma BV_correct_1 [rule_format]: "\\. \ wf_jvm_prog\<^bsub>\\<^esub> P; P,\ \ \\\ \ P \ \ -jvm\\<^sub>1 \' \ P,\ \ \'\" (*<*) -apply (simp only: split_tupled_all exec_1_iff) -apply (rename_tac xp h frs sh) -apply (case_tac xp) - apply (case_tac frs) - apply simp - apply (simp only: split_tupled_all) - apply hypsubst - apply (frule correct_state_impl_Some_method) - apply clarify - apply (rule step_correct) - apply assumption+ - apply (rule sym) - apply assumption+ -apply (case_tac frs) - apply simp_all -done +proof - + fix \ assume wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and cf: "P,\ \ \\" + obtain xp h frs sh where \[simp]: "\ = (xp, h, frs, sh)" by(cases \) simp + have "exec (P, xp, h, frs, sh) = \\'\ \ P,\ |- \' [ok]" + proof(cases xp) + case None + with wf cf show ?thesis + proof(cases frs) + case (Cons fr frs') + obtain stk loc C M pc ics where [simp]: "fr = (stk,loc,C,M,pc,ics)" by(cases fr) simp + then have cf': "P,\ |- (None, h, (stk,loc,C,M,pc,ics) # frs', sh) [ok]" + using Cons None cf by simp + then obtain b mxs mxl\<^sub>0 ins xt Ts T + where mC: "P \ C sees M, b : Ts\T = (mxs, mxl\<^sub>0, ins, xt) in C" + using correct_state_impl_Some_method[OF cf'] by clarify + show ?thesis using Cons None step_correct[OF wf mC _ cf'] by simp + qed simp + next + case (Some a) + then show ?thesis using wf cf by (case_tac frs) simp_all + qed + then show "P \ \ -jvm\\<^sub>1 \' \ P,\ \ \'\" by(simp add: exec_1_iff) +qed (*>*) theorem progress: "\ xp=None; frs\[] \ \ \\'. P \ (xp,h,frs,sh) -jvm\\<^sub>1 \'" by (clarsimp simp: exec_1_iff neq_Nil_conv split_beta simp del: split_paired_Ex) lemma progress_conform: "\wf_jvm_prog\<^bsub>\\<^esub> P; P,\ \ (xp,h,frs,sh)\; xp=None; frs\[]\ \ \\'. P \ (xp,h,frs,sh) -jvm\\<^sub>1 \' \ P,\ \ \'\" -(*<*) -apply (drule progress) - apply assumption -apply (fast intro: BV_correct_1) -done -(*>*) +(*<*)by (drule (1) progress) (fast intro: BV_correct_1)(*>*) theorem BV_correct [rule_format]: "\ wf_jvm_prog\<^bsub>\\<^esub> P; P \ \ -jvm\ \' \ \ P,\ \ \\ \ P,\ \ \'\" (*<*) -apply (simp only: exec_all_def1) -apply (erule rtrancl_induct) - apply simp -apply clarify -apply (erule (2) BV_correct_1) -done +proof - + assume wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and "P \ \ -jvm\ \'" + then have "(\, \') \ (exec_1 P)\<^sup>*" by (simp only: exec_all_def1) + then show ?thesis proof(induct rule: rtrancl_induct) + case (step y z) + then show ?case by clarify (erule (1) BV_correct_1[OF wf]) + qed simp +qed (*>*) lemma hconf_start: assumes wf: "wf_prog wf_mb P" shows "P \ (start_heap P) \" (*<*) - apply (unfold hconf_def) - apply (simp add: preallocated_start) - apply (clarify) - apply (drule sym) - apply (unfold start_heap_def) - apply (insert wf) - apply (auto simp: fun_upd_apply is_class_xcpt split: if_split_asm) - done +proof - + { fix a obj assume assm: "start_heap P a = \obj\" + have "P,start_heap P \ obj \" using wf assm[THEN sym] + by (unfold start_heap_def) + (auto simp: fun_upd_apply is_class_xcpt split: if_split_asm) + } + then show ?thesis using preallocated_start[of P] + by (unfold hconf_def) simp +qed (*>*) lemma shconf_start: "\ is_class P Start \ P,start_heap P \\<^sub>s start_sheap \" (*<*) - apply (unfold shconf_def) - apply (clarsimp simp: preallocated_start fun_upd_apply soconf_def has_field_is_class) - done + by (unfold shconf_def) + (clarsimp simp: preallocated_start fun_upd_apply soconf_def has_field_is_class) (*>*) lemma BV_correct_initial: - shows "\ wf_jvm_prog\<^bsub>\\<^esub> P; \is_class P Start; - P \ C sees M,Static:[]\Void = m in C; M \ clinit; - \' Start start_m = start_\\<^sub>m \ - \ start_prog P C M,\' \ start_state P \" +assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" + and nStart: "\is_class P Start" + and mC: "P \ C sees M,Static:[]\Void = m in C" + and nclinit: "M \ clinit" + and \: "\' Start start_m = start_\\<^sub>m" +shows "start_prog P C M,\' \ start_state P \" (*<*) - apply(subgoal_tac "is_class P Object") - prefer 2 apply(simp add: wf_jvm_prog_phi_def) - apply(subgoal_tac "\Mm. P \ Object sees_methods Mm") - prefer 2 apply(fastforce simp: is_class_def dest: sees_methods_Object) - apply (cases m) - apply (unfold start_state_def) - apply (unfold correct_state_def) - apply (simp del: defs1) - apply (rule conjI) - apply (simp add: wf_jvm_prog_phi_def class_add_hconf_wf[OF _ hconf_start] start_heap_nStart) - apply (rule conjI) - using start_prog_start_shconf apply(simp add: wf_jvm_prog_phi_def) - apply (rule conjI) - apply(simp add: conf_clinit_def distinct_clinit_def) - apply (drule wt_jvm_prog_impl_wt_start, assumption+) - apply (unfold conf_f_def wt_start_def) - apply (fastforce dest: start_prog_Start_sees_start_method) - done +proof - + let ?P = "class_add P (start_class C M)" + let ?h = "start_heap P" and ?sh = "[Start \ (Map.empty, Done)]" + let ?C = Start and ?M = start_m and ?pc = 0 and ?ics = No_ics + let ?f = "([], [], ?C, ?M, ?pc, ?ics)" and ?fs = "[]" + let ?frs = "?f#?fs" + have "is_class P Object" using wf by(simp add: wf_jvm_prog_phi_def) + then obtain Mm where Mm: "P \ Object sees_methods Mm" + by(fastforce simp: is_class_def dest: sees_methods_Object) + obtain mxs mxl\<^sub>0 ins xt where + mC': "P \ C sees M,Static:[]\Void = (mxs,mxl\<^sub>0,ins,xt) in C" + using mC by (cases m) simp + have "?P\ ?h\" using wf nStart class_add_hconf_wf[OF _ hconf_start] + by (simp add: wf_jvm_prog_phi_def start_heap_nStart) + moreover have "?P,?h\\<^sub>s ?sh\" by(rule start_prog_start_shconf) + moreover have "conf_clinit ?P ?sh ?frs" + by(simp add: conf_clinit_def distinct_clinit_def) + moreover have "\b Ts T mxs mxl\<^sub>0 is xt \. + (?P \ ?C sees ?M,b:Ts\T = (mxs,mxl\<^sub>0,is,xt) in ?C) + \ \' ?C ?M ! ?pc = Some \ + \ conf_f ?P ?h ?sh \ is ?f \ conf_fs ?P ?h ?sh \' ?C ?M (size Ts) T ?fs" + using \ start_prog_Start_sees_start_method[OF Mm] + by (unfold conf_f_def wt_start_def) fastforce + ultimately show ?thesis + by (simp del: defs1 add: start_state_def correct_state_def) +qed declare [[simproc add: list_to_set_comprehension]] (*>*) theorem typesafe: assumes welltyped: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes nstart: "\ is_class P Start" assumes main_method: "P \ C sees M,Static:[]\Void = m in C" assumes nclinit: "M \ clinit" assumes \: "\C. C \ Start \ \' C = \ C" assumes \': "\' Start start_m = start_\\<^sub>m" "\' Start clinit = start_\\<^sub>m" assumes Obj_start_m: "(\b' Ts' T' m' D'. P \ Object sees start_m, b' : Ts'\T' = m' in D' \ b' = Static \ Ts' = [] \ T' = Void)" shows "start_prog P C M \ start_state P -jvm\ \ \ start_prog P C M,\' \ \ \" (*<*) proof - from welltyped nstart main_method nclinit \'(1) have "start_prog P C M,\' \ start_state P \" by (rule BV_correct_initial) moreover assume "start_prog P C M \ start_state P -jvm\ \" moreover from start_prog_wf_jvm_prog_phi[OF welltyped nstart main_method nclinit \ \' Obj_start_m] have "wf_jvm_prog\<^bsub>\'\<^esub>(start_prog P C M)" by simp ultimately show "start_prog P C M,\' \ \ \" using welltyped by - (rule BV_correct) qed (*>*) end diff --git a/thys/JinjaDCI/BV/Effect.thy b/thys/JinjaDCI/BV/Effect.thy --- a/thys/JinjaDCI/BV/Effect.thy +++ b/thys/JinjaDCI/BV/Effect.thy @@ -1,441 +1,443 @@ (* Title: JinjaDCI/BV/Effect.thy Author: Gerwin Klein, Susannah Mansky Copyright 2000 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory BV/Effect.thy by Gerwin Klein *) section \Effect of Instructions on the State Type\ theory Effect imports JVM_SemiType "../JVM/JVMExceptions" begin \ \FIXME\ locale prog = fixes P :: "'a prog" locale jvm_method = prog + fixes mxs :: nat fixes mxl\<^sub>0 :: nat fixes Ts :: "ty list" fixes T\<^sub>r :: ty fixes "is" :: "instr list" fixes xt :: ex_table fixes mxl :: nat defines mxl_def: "mxl \ 1+size Ts+mxl\<^sub>0" text \ Program counter of successor instructions: \ primrec succs :: "instr \ ty\<^sub>i \ pc \ pc list" where "succs (Load idx) \ pc = [pc+1]" | "succs (Store idx) \ pc = [pc+1]" | "succs (Push v) \ pc = [pc+1]" | "succs (Getfield F C) \ pc = [pc+1]" | "succs (Getstatic C F D) \ pc = [pc+1]" | "succs (Putfield F C) \ pc = [pc+1]" | "succs (Putstatic C F D) \ pc = [pc+1]" | "succs (New C) \ pc = [pc+1]" | "succs (Checkcast C) \ pc = [pc+1]" | "succs Pop \ pc = [pc+1]" | "succs IAdd \ pc = [pc+1]" | "succs CmpEq \ pc = [pc+1]" | succs_IfFalse: "succs (IfFalse b) \ pc = [pc+1, nat (int pc + b)]" | succs_Goto: "succs (Goto b) \ pc = [nat (int pc + b)]" | succs_Return: "succs Return \ pc = []" | succs_Invoke: "succs (Invoke M n) \ pc = (if (fst \)!n = NT then [] else [pc+1])" | succs_Invokestatic: "succs (Invokestatic C M n) \ pc = [pc+1]" | succs_Throw: "succs Throw \ pc = []" text "Effect of instruction on the state type:" fun the_class:: "ty \ cname" where "the_class (Class C) = C" fun eff\<^sub>i :: "instr \ 'm prog \ ty\<^sub>i \ ty\<^sub>i" where eff\<^sub>i_Load: "eff\<^sub>i (Load n, P, (ST, LT)) = (ok_val (LT ! n) # ST, LT)" | eff\<^sub>i_Store: "eff\<^sub>i (Store n, P, (T#ST, LT)) = (ST, LT[n:= OK T])" | eff\<^sub>i_Push: "eff\<^sub>i (Push v, P, (ST, LT)) = (the (typeof v) # ST, LT)" | eff\<^sub>i_Getfield: "eff\<^sub>i (Getfield F C, P, (T#ST, LT)) = (snd (snd (field P C F)) # ST, LT)" | eff\<^sub>i_Getstatic: "eff\<^sub>i (Getstatic C F D, P, (ST, LT)) = (snd (snd (field P C F)) # ST, LT)" | eff\<^sub>i_Putfield: "eff\<^sub>i (Putfield F C, P, (T\<^sub>1#T\<^sub>2#ST, LT)) = (ST,LT)" | eff\<^sub>i_Putstatic: "eff\<^sub>i (Putstatic C F D, P, (T#ST, LT)) = (ST,LT)" | eff\<^sub>i_New: "eff\<^sub>i (New C, P, (ST,LT)) = (Class C # ST, LT)" | eff\<^sub>i_Checkcast: "eff\<^sub>i (Checkcast C, P, (T#ST,LT)) = (Class C # ST,LT)" | eff\<^sub>i_Pop: "eff\<^sub>i (Pop, P, (T#ST,LT)) = (ST,LT)" | eff\<^sub>i_IAdd: "eff\<^sub>i (IAdd, P,(T\<^sub>1#T\<^sub>2#ST,LT)) = (Integer#ST,LT)" | eff\<^sub>i_CmpEq: "eff\<^sub>i (CmpEq, P, (T\<^sub>1#T\<^sub>2#ST,LT)) = (Boolean#ST,LT)" | eff\<^sub>i_IfFalse: "eff\<^sub>i (IfFalse b, P, (T\<^sub>1#ST,LT)) = (ST,LT)" | eff\<^sub>i_Invoke: "eff\<^sub>i (Invoke M n, P, (ST,LT)) = (let C = the_class (ST!n); (D,b,Ts,T\<^sub>r,m) = method P C M in (T\<^sub>r # drop (n+1) ST, LT))" | eff\<^sub>i_Invokestatic: "eff\<^sub>i (Invokestatic C M n, P, (ST,LT)) = (let (D,b,Ts,T\<^sub>r,m) = method P C M in (T\<^sub>r # drop n ST, LT))" | eff\<^sub>i_Goto: "eff\<^sub>i (Goto n, P, s) = s" fun is_relevant_class :: "instr \ 'm prog \ cname \ bool" where rel_Getfield: "is_relevant_class (Getfield F D) = (\P C. P \ NullPointer \\<^sup>* C \ P \ NoSuchFieldError \\<^sup>* C \ P \ IncompatibleClassChangeError \\<^sup>* C)" | rel_Getstatic: "is_relevant_class (Getstatic C F D) = (\P C. True)" | rel_Putfield: "is_relevant_class (Putfield F D) = (\P C. P \ NullPointer \\<^sup>* C \ P \ NoSuchFieldError \\<^sup>* C \ P \ IncompatibleClassChangeError \\<^sup>* C)" | rel_Putstatic: "is_relevant_class (Putstatic C F D) = (\P C. True)" | rel_Checkcast: "is_relevant_class (Checkcast D) = (\P C. P \ ClassCast \\<^sup>* C)" | rel_New: "is_relevant_class (New D) = (\P C. True)" | rel_Throw: "is_relevant_class Throw = (\P C. True)" | rel_Invoke: "is_relevant_class (Invoke M n) = (\P C. True)" | rel_Invokestatic: "is_relevant_class (Invokestatic C M n) = (\P C. True)" | rel_default: "is_relevant_class i = (\P C. False)" definition is_relevant_entry :: "'m prog \ instr \ pc \ ex_entry \ bool" where "is_relevant_entry P i pc e \ (let (f,t,C,h,d) = e in is_relevant_class i P C \ pc \ {f.. instr \ pc \ ex_table \ ex_table" where "relevant_entries P i pc = filter (is_relevant_entry P i pc)" definition xcpt_eff :: "instr \ 'm prog \ pc \ ty\<^sub>i \ ex_table \ (pc \ ty\<^sub>i') list" where "xcpt_eff i P pc \ et = (let (ST,LT) = \ in map (\(f,t,C,h,d). (h, Some (Class C#drop (size ST - d) ST, LT))) (relevant_entries P i pc et))" definition norm_eff :: "instr \ 'm prog \ nat \ ty\<^sub>i \ (pc \ ty\<^sub>i') list" where "norm_eff i P pc \ = map (\pc'. (pc',Some (eff\<^sub>i (i,P,\)))) (succs i \ pc)" definition eff :: "instr \ 'm prog \ pc \ ex_table \ ty\<^sub>i' \ (pc \ ty\<^sub>i') list" where "eff i P pc et t = (case t of None \ [] | Some \ \ (norm_eff i P pc \) @ (xcpt_eff i P pc \ et))" lemma eff_None: "eff i P pc xt None = []" by (simp add: eff_def) lemma eff_Some: "eff i P pc xt (Some \) = norm_eff i P pc \ @ xcpt_eff i P pc \ xt" by (simp add: eff_def) (* FIXME: getfield, \T D. P \ C sees F:T in D \ .. *) text "Conditions under which eff is applicable:" fun app\<^sub>i :: "instr \ 'm prog \ pc \ nat \ ty \ ty\<^sub>i \ bool" where app\<^sub>i_Load: "app\<^sub>i (Load n, P, pc, mxs, T\<^sub>r, (ST,LT)) = (n < length LT \ LT ! n \ Err \ length ST < mxs)" | app\<^sub>i_Store: "app\<^sub>i (Store n, P, pc, mxs, T\<^sub>r, (T#ST, LT)) = (n < length LT)" | app\<^sub>i_Push: "app\<^sub>i (Push v, P, pc, mxs, T\<^sub>r, (ST,LT)) = (length ST < mxs \ typeof v \ None)" | app\<^sub>i_Getfield: "app\<^sub>i (Getfield F C, P, pc, mxs, T\<^sub>r, (T#ST, LT)) = (\T\<^sub>f. P \ C sees F,NonStatic:T\<^sub>f in C \ P \ T \ Class C)" | app\<^sub>i_Getstatic: "app\<^sub>i (Getstatic C F D, P, pc, mxs, T\<^sub>r, (ST, LT)) = (length ST < mxs \ (\T\<^sub>f. P \ C sees F,Static:T\<^sub>f in D))" | app\<^sub>i_Putfield: "app\<^sub>i (Putfield F C, P, pc, mxs, T\<^sub>r, (T\<^sub>1#T\<^sub>2#ST, LT)) = (\T\<^sub>f. P \ C sees F,NonStatic:T\<^sub>f in C \ P \ T\<^sub>2 \ (Class C) \ P \ T\<^sub>1 \ T\<^sub>f)" | app\<^sub>i_Putstatic: "app\<^sub>i (Putstatic C F D, P, pc, mxs, T\<^sub>r, (T#ST, LT)) = (\T\<^sub>f. P \ C sees F,Static:T\<^sub>f in D \ P \ T \ T\<^sub>f)" | app\<^sub>i_New: "app\<^sub>i (New C, P, pc, mxs, T\<^sub>r, (ST,LT)) = (is_class P C \ length ST < mxs)" | app\<^sub>i_Checkcast: "app\<^sub>i (Checkcast C, P, pc, mxs, T\<^sub>r, (T#ST,LT)) = (is_class P C \ is_refT T)" | app\<^sub>i_Pop: "app\<^sub>i (Pop, P, pc, mxs, T\<^sub>r, (T#ST,LT)) = True" | app\<^sub>i_IAdd: "app\<^sub>i (IAdd, P, pc, mxs, T\<^sub>r, (T\<^sub>1#T\<^sub>2#ST,LT)) = (T\<^sub>1 = T\<^sub>2 \ T\<^sub>1 = Integer)" | app\<^sub>i_CmpEq: "app\<^sub>i (CmpEq, P, pc, mxs, T\<^sub>r, (T\<^sub>1#T\<^sub>2#ST,LT)) = (T\<^sub>1 = T\<^sub>2 \ is_refT T\<^sub>1 \ is_refT T\<^sub>2)" | app\<^sub>i_IfFalse: "app\<^sub>i (IfFalse b, P, pc, mxs, T\<^sub>r, (Boolean#ST,LT)) = (0 \ int pc + b)" | app\<^sub>i_Goto: "app\<^sub>i (Goto b, P, pc, mxs, T\<^sub>r, s) = (0 \ int pc + b)" | app\<^sub>i_Return: "app\<^sub>i (Return, P, pc, mxs, T\<^sub>r, (T#ST,LT)) = (P \ T \ T\<^sub>r)" | app\<^sub>i_Throw: "app\<^sub>i (Throw, P, pc, mxs, T\<^sub>r, (T#ST,LT)) = is_refT T" | app\<^sub>i_Invoke: "app\<^sub>i (Invoke M n, P, pc, mxs, T\<^sub>r, (ST,LT)) = (n < length ST \ (ST!n \ NT \ (\C D Ts T m. ST!n = Class C \ P \ C sees M,NonStatic:Ts \ T = m in D \ P \ rev (take n ST) [\] Ts)))" | app\<^sub>i_Invokestatic: "app\<^sub>i (Invokestatic C M n, P, pc, mxs, T\<^sub>r, (ST,LT)) = (length ST - n < mxs \ n \ length ST \ M \ clinit \ (\D Ts T m. P \ C sees M,Static:Ts \ T = m in D \ P \ rev (take n ST) [\] Ts))" | app\<^sub>i_default: "app\<^sub>i (i,P, pc,mxs,T\<^sub>r,s) = False" definition xcpt_app :: "instr \ 'm prog \ pc \ nat \ ex_table \ ty\<^sub>i \ bool" where "xcpt_app i P pc mxs xt \ \ (\(f,t,C,h,d) \ set (relevant_entries P i pc xt). is_class P C \ d \ size (fst \) \ d < mxs)" definition app :: "instr \ 'm prog \ nat \ ty \ nat \ nat \ ex_table \ ty\<^sub>i' \ bool" where "app i P mxs T\<^sub>r pc mpc xt t = (case t of None \ True | Some \ \ app\<^sub>i (i,P,pc,mxs,T\<^sub>r,\) \ xcpt_app i P pc mxs xt \ \ (\(pc',\') \ set (eff i P pc xt t). pc' < mpc))" lemma app_Some: "app i P mxs T\<^sub>r pc mpc xt (Some \) = (app\<^sub>i (i,P,pc,mxs,T\<^sub>r,\) \ xcpt_app i P pc mxs xt \ \ (\(pc',s') \ set (eff i P pc xt (Some \)). pc' < mpc))" by (simp add: app_def) locale eff = jvm_method + fixes eff\<^sub>i and app\<^sub>i and eff and app fixes norm_eff and xcpt_app and xcpt_eff fixes mpc defines "mpc \ size is" defines "eff\<^sub>i i \ \ Effect.eff\<^sub>i (i,P,\)" notes eff\<^sub>i_simps [simp] = Effect.eff\<^sub>i.simps [where P = P, folded eff\<^sub>i_def] defines "app\<^sub>i i pc \ \ Effect.app\<^sub>i (i, P, pc, mxs, T\<^sub>r, \)" notes app\<^sub>i_simps [simp] = Effect.app\<^sub>i.simps [where P=P and mxs=mxs and T\<^sub>r=T\<^sub>r, folded app\<^sub>i_def] defines "xcpt_eff i pc \ \ Effect.xcpt_eff i P pc \ xt" notes xcpt_eff = Effect.xcpt_eff_def [of _ P _ _ xt, folded xcpt_eff_def] defines "norm_eff i pc \ \ Effect.norm_eff i P pc \" notes norm_eff = Effect.norm_eff_def [of _ P, folded norm_eff_def eff\<^sub>i_def] defines "eff i pc \ Effect.eff i P pc xt" notes eff = Effect.eff_def [of _ P _ xt, folded eff_def norm_eff_def xcpt_eff_def] defines "xcpt_app i pc \ \ Effect.xcpt_app i P pc mxs xt \" notes xcpt_app = Effect.xcpt_app_def [of _ P _ mxs xt, folded xcpt_app_def] defines "app i pc \ Effect.app i P mxs T\<^sub>r pc mpc xt" notes app = Effect.app_def [of _ P mxs T\<^sub>r _ mpc xt, folded app_def xcpt_app_def app\<^sub>i_def eff_def] lemma length_cases2: assumes "\LT. P ([],LT)" assumes "\l ST LT. P (l#ST,LT)" shows "P s" by (cases s, cases "fst s") (auto intro!: assms) lemma length_cases3: assumes "\LT. P ([],LT)" assumes "\l LT. P ([l],LT)" assumes "\l ST LT. P (l#ST,LT)" shows "P s" (*<*) proof - obtain xs LT where s: "s = (xs,LT)" by (cases s) show ?thesis proof (cases xs) case Nil with assms s show ?thesis by simp next fix l xs' assume "xs = l#xs'" with assms s show ?thesis by simp qed qed (*>*) lemma length_cases4: assumes "\LT. P ([],LT)" assumes "\l LT. P ([l],LT)" assumes "\l l' LT. P ([l,l'],LT)" assumes "\l l' ST LT. P (l#l'#ST,LT)" shows "P s" (*<*) proof - obtain xs LT where s: "s = (xs,LT)" by (cases s) show ?thesis proof (cases xs) case Nil with assms s show ?thesis by simp next fix l xs' assume xs: "xs = l#xs'" thus ?thesis proof (cases xs') case Nil with assms s xs show ?thesis by simp next fix l' ST assume "xs' = l'#ST" with assms s xs show ?thesis by simp qed qed qed (*>*) text \ \medskip simp rules for @{term app} \ lemma appNone[simp]: "app i P mxs T\<^sub>r pc mpc et None = True" by (simp add: app_def) lemma appLoad[simp]: "app\<^sub>i (Load idx, P, T\<^sub>r, mxs, pc, s) = (\ST LT. s = (ST,LT) \ idx < length LT \ LT!idx \ Err \ length ST < mxs)" by (cases s, simp) lemma appStore[simp]: "app\<^sub>i (Store idx,P,pc,mxs,T\<^sub>r,s) = (\ts ST LT. s = (ts#ST,LT) \ idx < length LT)" by (rule length_cases2, auto) lemma appPush[simp]: "app\<^sub>i (Push v,P,pc,mxs,T\<^sub>r,s) = (\ST LT. s = (ST,LT) \ length ST < mxs \ typeof v \ None)" by (cases s, simp) lemma appGetField[simp]: "app\<^sub>i (Getfield F C,P,pc,mxs,T\<^sub>r,s) = (\ oT vT ST LT. s = (oT#ST, LT) \ P \ C sees F,NonStatic:vT in C \ P \ oT \ (Class C))" by (rule length_cases2 [of _ s]) auto lemma appGetStatic[simp]: "app\<^sub>i (Getstatic C F D,P,pc,mxs,T\<^sub>r,s) = (\ vT ST LT. s = (ST, LT) \ length ST < mxs \ P \ C sees F,Static:vT in D)" by (rule length_cases2 [of _ s]) auto lemma appPutField[simp]: "app\<^sub>i (Putfield F C,P,pc,mxs,T\<^sub>r,s) = (\ vT vT' oT ST LT. s = (vT#oT#ST, LT) \ P \ C sees F,NonStatic:vT' in C \ P \ oT \ (Class C) \ P \ vT \ vT')" by (rule length_cases4 [of _ s], auto) lemma appPutstatic[simp]: "app\<^sub>i (Putstatic C F D,P,pc,mxs,T\<^sub>r,s) = (\ vT vT' ST LT. s = (vT#ST, LT) \ P \ C sees F,Static:vT' in D \ P \ vT \ vT')" by (rule length_cases4 [of _ s], auto) lemma appNew[simp]: "app\<^sub>i (New C,P,pc,mxs,T\<^sub>r,s) = (\ST LT. s=(ST,LT) \ is_class P C \ length ST < mxs)" by (cases s, simp) lemma appCheckcast[simp]: "app\<^sub>i (Checkcast C,P,pc,mxs,T\<^sub>r,s) = (\T ST LT. s = (T#ST,LT) \ is_class P C \ is_refT T)" by (cases s, cases "fst s", simp add: app_def) (cases "hd (fst s)", auto) lemma app\<^sub>iPop[simp]: "app\<^sub>i (Pop,P,pc,mxs,T\<^sub>r,s) = (\ts ST LT. s = (ts#ST,LT))" by (rule length_cases2, auto) lemma appIAdd[simp]: "app\<^sub>i (IAdd,P,pc,mxs,T\<^sub>r,s) = (\ST LT. s = (Integer#Integer#ST,LT))" (*<*) proof - obtain ST LT where [simp]: "s = (ST,LT)" by (cases s) have "ST = [] \ (\T. ST = [T]) \ (\T\<^sub>1 T\<^sub>2 ST'. ST = T\<^sub>1#T\<^sub>2#ST')" by (cases ST, auto, case_tac list, auto) moreover { assume "ST = []" hence ?thesis by simp } moreover { fix T assume "ST = [T]" hence ?thesis by (cases T, auto) } moreover { fix T\<^sub>1 T\<^sub>2 ST' assume "ST = T\<^sub>1#T\<^sub>2#ST'" hence ?thesis by (cases T\<^sub>1, auto) } ultimately show ?thesis by blast qed (*>*) lemma appIfFalse [simp]: "app\<^sub>i (IfFalse b,P,pc,mxs,T\<^sub>r,s) = (\ST LT. s = (Boolean#ST,LT) \ 0 \ int pc + b)" (*<*) - apply (rule length_cases2) - apply simp - apply (case_tac l) - apply auto - done + (is "?P s") +proof(rule length_cases2) + fix LT show "?P ([],LT)" by simp +next + fix l ST LT show "?P (l#ST,LT)" + by (case_tac l) auto +qed (*>*) lemma appCmpEq[simp]: "app\<^sub>i (CmpEq,P,pc,mxs,T\<^sub>r,s) = (\T\<^sub>1 T\<^sub>2 ST LT. s = (T\<^sub>1#T\<^sub>2#ST,LT) \ (\is_refT T\<^sub>1 \ T\<^sub>2 = T\<^sub>1 \ is_refT T\<^sub>1 \ is_refT T\<^sub>2))" by (rule length_cases4, auto) lemma appReturn[simp]: "app\<^sub>i (Return,P,pc,mxs,T\<^sub>r,s) = (\T ST LT. s = (T#ST,LT) \ P \ T \ T\<^sub>r)" by (rule length_cases2, auto) lemma appThrow[simp]: "app\<^sub>i (Throw,P,pc,mxs,T\<^sub>r,s) = (\T ST LT. s=(T#ST,LT) \ is_refT T)" by (rule length_cases2, auto) lemma effNone: "(pc', s') \ set (eff i P pc et None) \ s' = None" by (auto simp add: eff_def xcpt_eff_def norm_eff_def) text \ some helpers to make the specification directly executable: \ lemma relevant_entries_append [simp]: "relevant_entries P i pc (xt @ xt') = relevant_entries P i pc xt @ relevant_entries P i pc xt'" by (unfold relevant_entries_def) simp lemma xcpt_app_append [iff]: "xcpt_app i P pc mxs (xt@xt') \ = (xcpt_app i P pc mxs xt \ \ xcpt_app i P pc mxs xt' \)" by (unfold xcpt_app_def) fastforce lemma xcpt_eff_append [simp]: "xcpt_eff i P pc \ (xt@xt') = xcpt_eff i P pc \ xt @ xcpt_eff i P pc \ xt'" by (unfold xcpt_eff_def, cases \) simp lemma app_append [simp]: "app i P pc T mxs mpc (xt@xt') \ = (app i P pc T mxs mpc xt \ \ app i P pc T mxs mpc xt' \)" by (unfold app_def eff_def) auto end diff --git a/thys/JinjaDCI/BV/LBVJVM.thy b/thys/JinjaDCI/BV/LBVJVM.thy --- a/thys/JinjaDCI/BV/LBVJVM.thy +++ b/thys/JinjaDCI/BV/LBVJVM.thy @@ -1,226 +1,245 @@ (* Title: JinjaDCI/BV/LBVJVM.thy Author: Tobias Nipkow, Gerwin Klein, Susannah Mansky Copyright 2000 TUM, 2020 UIUC Based on the Jinja theory BV/LBVJVM.thy by Tobias Nipkow and Gerwin Klein *) section \ LBV for the JVM \label{sec:JVM} \ theory LBVJVM imports Jinja.Abstract_BV TF_JVM begin type_synonym prog_cert = "cname \ mname \ ty\<^sub>i' err list" definition check_cert :: "jvm_prog \ nat \ nat \ nat \ ty\<^sub>i' err list \ bool" where "check_cert P mxs mxl n cert \ check_types P mxs mxl cert \ size cert = n+1 \ (\i Err) \ cert!n = OK None" definition lbvjvm :: "jvm_prog \ nat \ nat \ ty \ ex_table \ ty\<^sub>i' err list \ instr list \ ty\<^sub>i' err \ ty\<^sub>i' err" where "lbvjvm P mxs maxr T\<^sub>r et cert bs \ wtl_inst_list bs cert (JVM_SemiType.sup P mxs maxr) (JVM_SemiType.le P mxs maxr) Err (OK None) (exec P mxs T\<^sub>r et bs) 0" definition wt_lbv :: "jvm_prog \ cname \ staticb \ ty list \ ty \ nat \ nat \ ex_table \ ty\<^sub>i' err list \ instr list \ bool" where "wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 et cert ins \ check_cert P mxs ((case b of Static \ 0 | NonStatic \ 1)+size Ts+mxl\<^sub>0) (size ins) cert \ 0 < size ins \ (let start = Some ([],(case b of Static \ [] | NonStatic \ [OK (Class C)]) @((map OK Ts))@(replicate mxl\<^sub>0 Err)); result = lbvjvm P mxs ((case b of Static \ 0 | NonStatic \ 1)+size Ts+mxl\<^sub>0) T\<^sub>r et cert ins (OK start) in result \ Err)" definition wt_jvm_prog_lbv :: "jvm_prog \ prog_cert \ bool" where "wt_jvm_prog_lbv P cert \ wf_prog (\P C (mn,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)). wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 et (cert C mn) ins) P" definition mk_cert :: "jvm_prog \ nat \ ty \ ex_table \ instr list \ ty\<^sub>m \ ty\<^sub>i' err list" where "mk_cert P mxs T\<^sub>r et bs phi \ make_cert (exec P mxs T\<^sub>r et bs) (map OK phi) (OK None)" definition prg_cert :: "jvm_prog \ ty\<^sub>P \ prog_cert" where "prg_cert P phi C mn \ let (C,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)) = method P C mn in mk_cert P mxs T\<^sub>r et ins (phi C mn)" lemma check_certD [intro?]: "check_cert P mxs mxl n cert \ cert_ok cert n Err (OK None) (states P mxs mxl)" by (unfold cert_ok_def check_cert_def check_types_def) auto lemma (in start_context) wt_lbv_wt_step: assumes lbv: "wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is" shows "\\s \ list (size is) A. wt_step r Err step \s \ OK first \\<^sub>r \s!0" (*<*) proof - from wf have "semilat (JVM_SemiType.sl P mxs mxl)" .. hence "semilat (A, r, f)" by (simp add: sl_def2) moreover have "top r Err" by (simp add: JVM_le_Err_conv) moreover have "Err \ A" by (simp add: JVM_states_unfold) moreover have "bottom r (OK None)" by (simp add: JVM_le_Err_conv bottom_def lesub_def Err.le_def split: err.split) moreover have "OK None \ A" by (simp add: JVM_states_unfold) moreover note bounded_step moreover from lbv have "cert_ok cert (size is) Err (OK None) A" by (unfold wt_lbv_def) (auto dest: check_certD) moreover note exec_pres_type moreover from lbv have "wtl_inst_list is cert f r Err (OK None) step 0 (OK first) \ Err" by (cases b; simp add: wt_lbv_def lbvjvm_def step_def_exec [symmetric]) moreover note first_in_A moreover from lbv have "0 < size is" by (simp add: wt_lbv_def) ultimately show ?thesis by (rule lbvs.wtl_sound_strong [OF lbvs.intro, OF lbv.intro lbvs_axioms.intro, OF Semilat.intro lbv_axioms.intro]) qed (*>*) lemma (in start_context) wt_lbv_wt_method: assumes lbv: "wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is" shows "\\s. wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" (*<*) proof - from lbv have l: "is \ []" by (simp add: wt_lbv_def) moreover from wf lbv C Ts obtain \s where list: "\s \ list (size is) A" and step: "wt_step r Err step \s" and start: "OK first \\<^sub>r \s!0" by (blast dest: wt_lbv_wt_step) from list have [simp]: "size \s = size is" by simp have "size (map ok_val \s) = size is" by simp moreover from l have 0: "0 < size \s" by simp with step obtain \s0 where "\s!0 = OK \s0" by (unfold wt_step_def) blast with start 0 have "wt_start P C b Ts mxl\<^sub>0 (map ok_val \s)" by (cases b; simp add: wt_start_def JVM_le_Err_conv lesub_def Err.le_def) moreover { from list have "check_types P mxs mxl \s" by (simp add: check_types_def) also from step have "\x \ set \s. x \ Err" by (auto simp add: all_set_conv_all_nth wt_step_def) hence [symmetric]: "map OK (map ok_val \s) = \s" by (auto intro!: map_idI) finally have "check_types P mxs mxl (map OK (map ok_val \s))" . } moreover { note bounded_step moreover from list have "set \s \ A" by simp moreover from step have "wt_err_step (sup_state_opt P) step \s" by (simp add: wt_err_step_def JVM_le_Err_conv) ultimately have "wt_app_eff (sup_state_opt P) app eff (map ok_val \s)" by (auto intro: wt_err_imp_wt_app_eff simp add: exec_def states_def) } ultimately have "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (map ok_val \s)" by (simp add: wt_method_def2 check_types_def del: map_map) thus ?thesis .. qed (*>*) lemma (in start_context) wt_method_wt_lbv: assumes wt: "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" defines [simp]: "cert \ mk_cert P mxs T\<^sub>r xt is \s" shows "wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is" (*<*) proof - let ?\s = "map OK \s" let ?cert = "make_cert step ?\s (OK None)" from wt obtain 0: "0 < size is" and size: "size is = size ?\s" and ck_types: "check_types P mxs mxl ?\s" and wt_start: "wt_start P C b Ts mxl\<^sub>0 \s" and app_eff: "wt_app_eff (sup_state_opt P) app eff \s" by (force simp add: wt_method_def2 check_types_def) from wf have "semilat (JVM_SemiType.sl P mxs mxl)" .. hence "semilat (A, r, f)" by (simp add: sl_def2) moreover have "top r Err" by (simp add: JVM_le_Err_conv) moreover have "Err \ A" by (simp add: JVM_states_unfold) moreover have "bottom r (OK None)" by (simp add: JVM_le_Err_conv bottom_def lesub_def Err.le_def split: err.split) moreover have "OK None \ A" by (simp add: JVM_states_unfold) moreover from wf have "mono r step (size is) A" by (rule step_mono) hence "mono r step (size ?\s) A" by (simp add: size) moreover from exec_pres_type have "pres_type step (size ?\s) A" by (simp add: size) moreover from ck_types have \s_in_A: "set ?\s \ A" by (simp add: check_types_def) hence "\pc. pc < size ?\s \ ?\s!pc \ A \ ?\s!pc \ Err" by auto moreover from bounded_step have "bounded step (size ?\s)" by (simp add: size) moreover have "OK None \ Err" by simp moreover from bounded_step size \s_in_A app_eff have "wt_err_step (sup_state_opt P) step ?\s" by (auto intro: wt_app_eff_imp_wt_err simp add: exec_def states_def) hence "wt_step r Err step ?\s" by (simp add: wt_err_step_def JVM_le_Err_conv) moreover from 0 size have "0 < size \s" by auto hence "?\s!0 = OK (\s!0)" by simp with wt_start have "OK first \\<^sub>r ?\s!0" by (cases b; clarsimp simp add: wt_start_def lesub_def Err.le_def JVM_le_Err_conv) moreover note first_in_A moreover have "OK first \ Err" by simp moreover note size ultimately have "wtl_inst_list is ?cert f r Err (OK None) step 0 (OK first) \ Err" by (rule lbvc.wtl_complete [OF lbvc.intro, OF lbv.intro lbvc_axioms.intro, OF Semilat.intro lbv_axioms.intro]) moreover from 0 size have "\s \ []" by auto moreover from ck_types have "check_types P mxs mxl ?cert" - apply (auto simp add: make_cert_def check_types_def JVM_states_unfold) - apply (subst Ok_in_err [symmetric]) - apply (drule nth_mem) - apply auto - done + by (fastforce simp: make_cert_def check_types_def JVM_states_unfold + dest!: nth_mem) moreover note 0 size ultimately show ?thesis by (simp add: wt_lbv_def lbvjvm_def mk_cert_def step_def_exec [symmetric] check_cert_def make_cert_def nth_append) qed (*>*) theorem jvm_lbv_correct: "wt_jvm_prog_lbv P Cert \ wf_jvm_prog P" (*<*) proof - let ?\ = "\C mn. let (C,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)) = method P C mn in SOME \s. wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" - + + let ?A = "\P C (mn,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)). wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 et (Cert C mn) ins" + let ?B = "\P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). + wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (?\ C M)" + assume wt: "wt_jvm_prog_lbv P Cert" - hence "wf_jvm_prog\<^bsub>?\\<^esub> P" - apply (unfold wf_jvm_prog_phi_def wt_jvm_prog_lbv_def) - apply (erule wf_prog_lift) - apply (auto dest!: start_context.wt_lbv_wt_method [OF start_context.intro] - intro: someI) - apply (erule sees_method_is_class) - done + then have "wf_prog ?A P" by(simp add: wt_jvm_prog_lbv_def) + moreover { + fix wf_md C M b Ts Ca T m bd + assume "wf_prog wf_md P" and sees: "P \ Ca sees M, b : Ts\T = m in Ca" and + "set Ts \ types P" and "bd = (M, b, Ts, T, m)" and + "?A P Ca bd" + then have "?B P Ca bd" using sees_method_is_class[OF sees] + by (auto dest!: start_context.wt_lbv_wt_method [OF start_context.intro] + intro: someI) + } + ultimately have "wf_prog ?B P" by(rule wf_prog_lift) + hence "wf_jvm_prog\<^bsub>?\\<^esub> P" by (simp add: wf_jvm_prog_phi_def) thus ?thesis by (unfold wf_jvm_prog_def) blast qed (*>*) theorem jvm_lbv_complete: assumes wt: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "wt_jvm_prog_lbv P (prg_cert P \)" (*<*) - using wt - apply (unfold wf_jvm_prog_phi_def wt_jvm_prog_lbv_def) - apply (erule wf_prog_lift) - apply (auto simp add: prg_cert_def - intro!: start_context.wt_method_wt_lbv start_context.intro) - apply (erule sees_method_is_class) - done +proof - + let ?cert = "prg_cert P \" + let ?A = "\P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). + wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M)" + let ?B = "\P C (mn,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)). + wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 et (?cert C mn) ins" + + from wt have "wf_prog ?A P" by(clarsimp simp: wf_jvm_prog_def wf_jvm_prog_phi_def) + moreover { + fix wf_md C M b Ts Ca T m bd + assume "wf_prog wf_md P" and sees: "P \ Ca sees M, b : Ts\T = m in Ca" and + "set Ts \ types P" and "bd = (M, b, Ts, T, m)" and + "?A P Ca bd" + then have "?B P Ca bd" using sees_method_is_class[OF sees] + by (auto simp add: prg_cert_def + intro!: start_context.wt_method_wt_lbv start_context.intro) + } + ultimately have "wf_prog ?B P" by(rule wf_prog_lift) + thus "wt_jvm_prog_lbv P (prg_cert P \)" by (simp add: wt_jvm_prog_lbv_def) +qed (*>*) -end +end diff --git a/thys/JinjaDCI/BV/TF_JVM.thy b/thys/JinjaDCI/BV/TF_JVM.thy --- a/thys/JinjaDCI/BV/TF_JVM.thy +++ b/thys/JinjaDCI/BV/TF_JVM.thy @@ -1,292 +1,319 @@ (* Title: JinjaDCI/BV/TF_JVM.thy Author: Tobias Nipkow, Gerwin Klein, Susannah Mansky Copyright 2000 TUM, 2019-20 UIUC Based on the Jinja theory BV/TF_JVM.thy by Tobias Nipkow and Gerwin Klein *) section \ The Typing Framework for the JVM \label{sec:JVM} \ theory TF_JVM imports Jinja.Typing_Framework_err EffectMono BVSpec begin definition exec :: "jvm_prog \ nat \ ty \ ex_table \ instr list \ ty\<^sub>i' err step_type" where "exec G maxs rT et bs \ err_step (size bs) (\pc. app (bs!pc) G maxs rT pc (size bs) et) (\pc. eff (bs!pc) G pc et)" locale JVM_sl = fixes P :: jvm_prog and mxs and mxl\<^sub>0 fixes b and Ts :: "ty list" and "is" and xt and T\<^sub>r fixes mxl and A and r and f and app and eff and step defines [simp]: "mxl \ (case b of Static \ 0 | NonStatic \ 1)+size Ts+mxl\<^sub>0" defines [simp]: "A \ states P mxs mxl" defines [simp]: "r \ JVM_SemiType.le P mxs mxl" defines [simp]: "f \ JVM_SemiType.sup P mxs mxl" defines [simp]: "app \ \pc. Effect.app (is!pc) P mxs T\<^sub>r pc (size is) xt" defines [simp]: "eff \ \pc. Effect.eff (is!pc) P pc xt" defines [simp]: "step \ err_step (size is) app eff" locale start_context = JVM_sl + fixes p and C assumes wf: "wf_prog p P" assumes C: "is_class P C" assumes Ts: "set Ts \ types P" fixes first :: ty\<^sub>i' and start defines [simp]: "first \ Some ([],(case b of Static \ [] | NonStatic \ [OK (Class C)]) @ map OK Ts @ replicate mxl\<^sub>0 Err)" defines [simp]: "start \ (OK first) # replicate (size is - 1) (OK None)" subsection \ Connecting JVM and Framework \ lemma (in JVM_sl) step_def_exec: "step \ exec P mxs T\<^sub>r xt is" by (simp add: exec_def) lemma special_ex_swap_lemma [iff]: "(? X. (? n. X = A n & P n) & Q X) = (? n. Q(A n) & P n)" by blast lemma ex_in_list [iff]: "(\n. ST \ list n A \ n \ mxs) = (set ST \ A \ size ST \ mxs)" by (unfold list_def) auto lemma singleton_list: "(\n. [Class C] \ list n (types P) \ n \ mxs) = (is_class P C \ 0 < mxs)" by auto lemma set_drop_subset: "set xs \ A \ set (drop n xs) \ A" by (auto dest: in_set_dropD) lemma Suc_minus_minus_le: "n < mxs \ Suc (n - (n - b)) \ mxs" by arith lemma in_listE: "\ xs \ list n A; \size xs = n; set xs \ A\ \ P \ \ P" by (unfold list_def) blast declare is_relevant_entry_def [simp] declare set_drop_subset [simp] theorem (in start_context) exec_pres_type: "pres_type step (size is) A" (*<*) - apply (insert wf) - apply simp - apply (unfold JVM_states_unfold) - apply (rule pres_type_lift) - apply clarify - apply (rename_tac s pc pc' s') - apply (case_tac s) - apply simp - apply (drule effNone) - apply simp - apply (simp add: Effect.app_def xcpt_app_def Effect.eff_def - xcpt_eff_def norm_eff_def relevant_entries_def) - apply (case_tac "is!pc") - - \ \Load\ - apply clarsimp - apply (frule listE_nth_in, assumption) - apply fastforce - - \ \Store\ - apply fastforce - - \ \Push\ - apply (fastforce simp add: typeof_lit_is_type) - - \ \New\ - apply fastforce - - \ \Getfield\ - apply (fastforce dest: sees_field_is_type) - - \ \Getstatic\ - apply (fastforce dest: sees_field_is_type) - - \ \Putfield\ - apply fastforce - - \ \Putstatic\ - apply fastforce - - \ \Checkcast\ - apply fastforce - - defer defer \ \Invoke and Invokestatic deferred\ - - \ \Return\ - apply fastforce - - \ \Pop\ - apply fastforce - - \ \IAdd\ - apply fastforce - - \ \Goto\ - apply fastforce - - \ \CmpEq\ - apply fastforce - - \ \IfFalse\ - apply fastforce - - \ \Throw\ - apply fastforce - - \ \Invoke\ - apply (clarsimp split!: if_splits) - apply fastforce - apply (erule disjE) - prefer 2 - apply fastforce - apply clarsimp - apply (rule conjI) - apply (drule (1) sees_wf_mdecl) - apply (clarsimp simp add: wf_mdecl_def) - apply arith - - \ \Invokestatic\ - apply (clarsimp split!: if_splits) - apply (erule disjE) - prefer 2 - apply fastforce - apply clarsimp - apply (drule (1) sees_wf_mdecl) - apply (clarsimp simp add: wf_mdecl_def) - done +proof - + let ?n = "size is" and ?app = app and ?step = eff + let ?mxl = "(case b of Static \ 0 | NonStatic \ 1) + length Ts + mxl\<^sub>0" + let ?A = "opt((Union {list n (types P) |n. n <= mxs}) \ + list ?mxl (err(types P)))" + have "pres_type (err_step ?n ?app ?step) ?n (err ?A)" + proof(rule pres_type_lift) + have "\s pc pc' s'. s\?A \ pc < ?n \ ?app pc s + \ (pc', s')\set (?step pc s) \ s' \ ?A" + proof - + fix s pc pc' s' + assume asms: "s\?A" "pc < ?n" "?app pc s" "(pc', s')\set (?step pc s)" + show "s' \ ?A" + proof(cases s) + case None + then show ?thesis using asms by (fastforce dest: effNone) + next + case (Some ab) + then show ?thesis using asms proof(cases "is!pc") + case Load + then show ?thesis using asms + by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def + xcpt_eff_def norm_eff_def + dest: listE_nth_in) + next + case Push + then show ?thesis using asms Some + by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def + xcpt_eff_def norm_eff_def typeof_lit_is_type) + next + case Getfield + then show ?thesis using asms wf + by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def + xcpt_eff_def norm_eff_def + dest: sees_field_is_type) + next + case Getstatic + then show ?thesis using asms wf + by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def + xcpt_eff_def norm_eff_def + dest: sees_field_is_type) + next + case (Invoke M n) + obtain a b where [simp]: "s = \(a,b)\" using Some asms(1) by blast + show ?thesis + proof(cases "a!n = NT") + case True + then show ?thesis using Invoke asms wf + by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def + xcpt_eff_def norm_eff_def) + next + case False + have "(pc', s') \ set (norm_eff (Invoke M n) P pc (a, b)) \ + (pc', s') \ set (xcpt_eff (Invoke M n) P pc (a, b) xt)" + using Invoke asms(4) by (simp add: Effect.eff_def) + then show ?thesis proof(rule disjE) + assume "(pc', s') \ set (xcpt_eff (Invoke M n) P pc (a, b) xt)" + then show ?thesis using Invoke asms(1-3) + by (fastforce simp: Effect.app_def xcpt_app_def xcpt_eff_def) + next + assume norm: "(pc', s') \ set (norm_eff (Invoke M n) P pc (a, b))" + also have "Suc (length a - Suc n) \ mxs" using Invoke asms(1,3) + by (simp add: Effect.app_def xcpt_app_def) arith + ultimately show ?thesis using False Invoke asms(1-3) wf + by (auto simp: Effect.app_def xcpt_app_def norm_eff_def wf_mdecl_def + dest!: sees_wf_mdecl) + qed + qed + next + case (Invokestatic C M n) + obtain a b where [simp]: "s = \(a,b)\" using Some asms(1) by blast + have "(pc', s') \ set (norm_eff (Invokestatic C M n) P pc (a, b)) \ + (pc', s') \ set (xcpt_eff (Invokestatic C M n) P pc (a, b) xt)" + using Invokestatic asms(4) by (simp add: Effect.eff_def) + then show ?thesis proof(rule disjE) + assume "(pc', s') \ set (xcpt_eff (Invokestatic C M n) P pc (a, b) xt)" + then show ?thesis using Invokestatic asms(1-3) + by (fastforce simp: Effect.app_def xcpt_app_def xcpt_eff_def) + next + assume norm: "(pc', s') \ set (norm_eff (Invokestatic C M n) P pc (a, b))" + also have "Suc (length a - Suc n) \ mxs" using Invokestatic asms(1,3) + by (simp add: Effect.app_def xcpt_app_def) arith + ultimately show ?thesis using Invokestatic asms(1-3) wf + by (auto simp: Effect.app_def xcpt_app_def norm_eff_def wf_mdecl_def + dest!: sees_wf_mdecl) + qed + qed (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def + xcpt_eff_def norm_eff_def)+ + qed + qed + then show "\s\?A. \p. p < ?n \ ?app p s \ (\(q, s')\set (?step p s). s' \ ?A)" + by clarsimp + qed + then show ?thesis by (simp add: JVM_states_unfold) +qed (*>*) declare is_relevant_entry_def [simp del] declare set_drop_subset [simp del] lemma lesubstep_type_simple: "xs [\\<^bsub>Product.le (=) r\<^esub>] ys \ set xs {\\<^bsub>r\<^esub>} set ys" (*<*) - apply (unfold lesubstep_type_def) - apply clarify - apply (simp add: set_conv_nth) - apply clarify - apply (drule le_listD, assumption) - apply (clarsimp simp add: lesub_def Product.le_def) - apply (rule exI) - apply (rule conjI) - apply (rule exI) - apply (rule conjI) - apply (rule sym) - apply assumption - apply assumption - apply assumption - done +proof - + assume assm: "xs [\\<^bsub>Product.le (=) r\<^esub>] ys" + have "\a b i y. (a, b) = xs ! i \ i < length xs + \ \\'. (\i. (a, \') = ys ! i \ i < length xs) \ b \\<^bsub>r\<^esub> \'" + proof - + fix a b i assume ith: "(a, b) = xs ! i" and len: "i < length xs" + obtain \ where "ys ! i = (a, \) \ r b \" + using le_listD[OF assm len] ith + by (clarsimp simp: lesub_def Product.le_def) + then have "(a, \) = ys ! i \ b \\<^bsub>r\<^esub> \" + by (clarsimp simp: lesub_def) + with len show "\\'. (\i. (a, \') = ys ! i \ i < length xs) \ b \\<^bsub>r\<^esub> \'" + by fastforce + qed + then show "set xs {\\<^bsub>r\<^esub>} set ys" using assm + by (clarsimp simp: lesubstep_type_def set_conv_nth) +qed (*>*) declare is_relevant_entry_def [simp del] lemma conjI2: "\ A; A \ B \ \ A \ B" by blast lemma (in JVM_sl) eff_mono: - "\wf_prog p P; pc < length is; s \\<^bsub>sup_state_opt P\<^esub> t; app pc t\ - \ set (eff pc s) {\\<^bsub>sup_state_opt P\<^esub>} set (eff pc t)" +assumes wf: "wf_prog p P" and "pc < length is" and + lesub: "s \\<^bsub>sup_state_opt P\<^esub> t" and app: "app pc t" +shows "set (eff pc s) {\\<^bsub>sup_state_opt P\<^esub>} set (eff pc t)" (*<*) - apply simp - apply (unfold Effect.eff_def) - apply (cases t) - apply (simp add: lesub_def) - apply (rename_tac a) - apply (cases s) - apply simp - apply (rename_tac b) - apply simp - apply (rule lesubstep_union) - prefer 2 - apply (rule lesubstep_type_simple) - apply (simp add: xcpt_eff_def) - apply (rule le_listI) - apply (simp add: split_beta) - apply (simp add: split_beta) - apply (simp add: lesub_def fun_of_def) - apply (case_tac a) - apply (case_tac b) - apply simp - apply (subgoal_tac "size ab = size aa") - prefer 2 - apply (clarsimp simp add: list_all2_lengthD) - apply simp - apply (clarsimp simp add: norm_eff_def lesubstep_type_def lesub_def iff del: sup_state_conv) - apply (rule exI) - apply (rule conjI2) - apply (rule imageI) - apply (clarsimp simp add: Effect.app_def iff del: sup_state_conv) - apply (drule (2) succs_mono) - apply blast - apply simp - apply (erule eff\<^sub>i_mono) - apply simp - apply assumption - apply clarsimp - apply clarsimp - done +proof(cases t) + case None then show ?thesis using lesub + by (simp add: Effect.eff_def lesub_def) +next + case tSome: (Some a) + show ?thesis proof(cases s) + case None then show ?thesis using lesub + by (simp add: Effect.eff_def lesub_def) + next + case (Some b) + let ?norm = "\x. norm_eff (is ! pc) P pc x" + let ?xcpt = "\x. xcpt_eff (is ! pc) P pc x xt" + let ?r = "Product.le (=) (sup_state_opt P)" + let ?\' = "\eff\<^sub>i (is ! pc, P, a)\" + { + fix x assume xb: "x \ set (succs (is ! pc) b pc)" + then have appi: "app\<^sub>i (is ! pc, P, pc, mxs, T\<^sub>r, a)" and + bia: "P \ b \\<^sub>i a" and appa: "app pc \a\" + using lesub app tSome Some by (auto simp add: lesub_def Effect.app_def) + have xa: "x \ set (succs (is ! pc) a pc)" + using xb succs_mono[OF wf appi bia] by auto + then have "(x, ?\') \ (\pc'. (pc', ?\')) ` set (succs (is ! pc) a pc)" + by (rule imageI) + moreover have "P \ \eff\<^sub>i (is ! pc, P, b)\ \' ?\'" + using xb xa eff\<^sub>i_mono[OF wf bia] appa by fastforce + ultimately have "\\'. (x, \') \ (\pc'. (pc', \eff\<^sub>i (is ! pc, P, a)\)) ` set (succs (is ! pc) a pc) \ + P \ \eff\<^sub>i (is ! pc, P, b)\ \' \'" by blast + } + then have norm: "set (?norm b) {\\<^bsub>sup_state_opt P\<^esub>} set (?norm a)" + using tSome Some by (clarsimp simp: norm_eff_def lesubstep_type_def lesub_def) + obtain a1 b1 a2 b2 where a: "a = (a1, b1)" and b: "b = (a2, b2)" + using tSome Some by fastforce + then have a12: "size a2 = size a1" using lesub tSome Some + by (clarsimp simp: lesub_def list_all2_lengthD) + have "length (?xcpt b) = length (?xcpt a)" + by (simp add: xcpt_eff_def split_beta) + moreover have "\n. n < length (?xcpt b) \ (?xcpt b) ! n \\<^bsub>?r\<^esub> (?xcpt a) ! n" + using lesub tSome Some a b a12 + by (simp add: xcpt_eff_def split_beta fun_of_def) (simp add: lesub_def) + ultimately have "?xcpt b [\\<^bsub>?r\<^esub>] ?xcpt a" + by(rule le_listI) + then have "set (?xcpt b) {\\<^bsub>sup_state_opt P\<^esub>} set (?xcpt a)" + by (rule lesubstep_type_simple) + moreover note norm + ultimately have + "set (?norm b) \ set (?xcpt b) {\\<^bsub>sup_state_opt P\<^esub>} set (?norm a) \ set (?xcpt a)" + by(intro lesubstep_union) + then show ?thesis using tSome Some by(simp add: Effect.eff_def) + qed +qed (*>*) lemma (in JVM_sl) bounded_step: "bounded step (size is)" (*<*) - apply simp - apply (unfold bounded_def err_step_def Effect.app_def Effect.eff_def) - apply (auto simp add: error_def map_snd_def split: err.splits option.splits) - done + by (auto simp: bounded_def err_step_def Effect.app_def Effect.eff_def + error_def map_snd_def + split: err.splits option.splits) (*>*) theorem (in JVM_sl) step_mono: "wf_prog wf_mb P \ mono r step (size is) A" (*<*) - apply (simp add: JVM_le_Err_conv) - apply (insert bounded_step) - apply (unfold JVM_states_unfold) - apply (rule mono_lift) - apply blast - apply (unfold app_mono_def lesub_def) - apply clarsimp - apply (erule (2) app_mono) - apply simp - apply clarify - apply (drule eff_mono) - apply (auto simp add: lesub_def) - done +proof - + assume wf: "wf_prog wf_mb P" + let ?r = "sup_state_opt P" and ?n = "length is" and ?app = app and ?step = eff + let ?A = "opt (\ {list n (types P) |n. n \ mxs} \ + list ((case b of Static \ 0 | NonStatic \ 1) + length Ts + mxl\<^sub>0) + (err (types P)))" + have "order ?r" using wf by simp + moreover have "app_mono ?r ?app ?n ?A" using app_mono[OF wf] + by (clarsimp simp: app_mono_def lesub_def) + moreover have "bounded (err_step ?n ?app ?step) ?n" using bounded_step + by simp + moreover have "\s p t. s \ ?A \ p < ?n \ s \\<^bsub>?r\<^esub> t \ + ?app p t \ set (?step p s) {\\<^bsub>?r\<^esub>} set (?step p t)" + using eff_mono[OF wf] by simp + ultimately have "mono (Err.le ?r) (err_step ?n ?app ?step) ?n (err ?A)" + by(rule mono_lift) + then show "mono r step (size is) A" using bounded_step + by (simp add: JVM_le_Err_conv JVM_states_unfold) +qed (*>*) lemma (in start_context) first_in_A [iff]: "OK first \ A" using Ts C by (cases b; force intro!: list_appendI simp add: JVM_states_unfold) lemma (in JVM_sl) wt_method_def2: "wt_method P C' b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s = (is \ [] \ size \s = size is \ OK ` set \s \ states P mxs mxl \ wt_start P C' b Ts mxl\<^sub>0 \s \ wt_app_eff (sup_state_opt P) app eff \s)" (*<*) - apply (unfold wt_method_def wt_app_eff_def wt_instr_def lesub_def check_types_def) - apply auto - done + by (unfold wt_method_def wt_app_eff_def wt_instr_def lesub_def + check_types_def) auto (*>*) end diff --git a/thys/JinjaDCI/Compiler/Correctness1.thy b/thys/JinjaDCI/Compiler/Correctness1.thy --- a/thys/JinjaDCI/Compiler/Correctness1.thy +++ b/thys/JinjaDCI/Compiler/Correctness1.thy @@ -1,1058 +1,1064 @@ (* Title: JinjaDCI/Compiler/Correctness1.thy Author: Tobias Nipkow, Susannah Mansky Copyright TUM 2003, UIUC 2019-20 Based on the Jinja theory Compiler/Correctness1.thy by Tobias Nipkow *) section \ Correctness of Stage 1 \ theory Correctness1 imports J1WellForm Compiler1 begin subsection\Correctness of program compilation \ primrec unmod :: "expr\<^sub>1 \ nat \ bool" and unmods :: "expr\<^sub>1 list \ nat \ bool" where "unmod (new C) i = True" | "unmod (Cast C e) i = unmod e i" | "unmod (Val v) i = True" | "unmod (e\<^sub>1 \bop\ e\<^sub>2) i = (unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (Var i) j = True" | "unmod (i:=e) j = (i \ j \ unmod e j)" | "unmod (e\F{D}) i = unmod e i" | "unmod (C\\<^sub>sF{D}) i = True" | "unmod (e\<^sub>1\F{D}:=e\<^sub>2) i = (unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (C\\<^sub>sF{D}:=e\<^sub>2) i = unmod e\<^sub>2 i" | "unmod (e\M(es)) i = (unmod e i \ unmods es i)" | "unmod (C\\<^sub>sM(es)) i = unmods es i" | "unmod {j:T; e} i = unmod e i" | "unmod (e\<^sub>1;;e\<^sub>2) i = (unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (if (e) e\<^sub>1 else e\<^sub>2) i = (unmod e i \ unmod e\<^sub>1 i \ unmod e\<^sub>2 i)" | "unmod (while (e) c) i = (unmod e i \ unmod c i)" | "unmod (throw e) i = unmod e i" | "unmod (try e\<^sub>1 catch(C i) e\<^sub>2) j = (unmod e\<^sub>1 j \ (if i=j then False else unmod e\<^sub>2 j))" | "unmod (INIT C (Cs,b) \ e) i = unmod e i" | "unmod (RI(C,e);Cs \ e') i = (unmod e i \ unmod e' i)" | "unmods ([]) i = True" | "unmods (e#es) i = (unmod e i \ unmods es i)" lemma hidden_unmod: "\Vs. hidden Vs i \ unmod (compE\<^sub>1 Vs e) i" and "\Vs. hidden Vs i \ unmods (compEs\<^sub>1 Vs es) i" (*<*) -apply(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) -apply (simp_all add:hidden_inacc) -apply(auto simp add:hidden_def) -done +proof(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) + case TryCatch + then show ?case by(simp add:hidden_inacc)(auto simp add:hidden_def) +qed (simp_all add:hidden_inacc) (*>*) lemma eval\<^sub>1_preserves_unmod: "\ P \\<^sub>1 \e,(h,ls,sh)\ \ \e',(h',ls',sh')\; unmod e i; i < size ls \ \ ls ! i = ls' ! i" and "\ P \\<^sub>1 \es,(h,ls,sh)\ [\] \es',(h',ls',sh')\; unmods es i; i < size ls \ \ ls ! i = ls' ! i" (*<*) proof(induct rule:eval\<^sub>1_evals\<^sub>1_inducts) case (RInitInitFail\<^sub>1 e h l sh a h' l' sh' C sfs i sh'' D Cs e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1) have "final (throw a)" using eval\<^sub>1_final[OF RInitInitFail\<^sub>1.hyps(1)] by simp then show ?case using RInitInitFail\<^sub>1 by(auto simp: eval\<^sub>1_preserves_len) qed(auto dest!:eval\<^sub>1_preserves_len evals\<^sub>1_preserves_len split:if_split_asm) (*>*) lemma LAss_lem: "\x \ set xs; size xs \ size ys \ \ m\<^sub>1 \\<^sub>m m\<^sub>2(xs[\]ys) \ m\<^sub>1(x\y) \\<^sub>m m\<^sub>2(xs[\]ys[last_index xs x := y])" (*<*) by(simp add:map_le_def fun_upds_apply eq_sym_conv) (*>*) lemma Block_lem: fixes l :: "'a \ 'b" assumes 0: "l \\<^sub>m [Vs [\] ls]" and 1: "l' \\<^sub>m [Vs [\] ls', V\v]" and hidden: "V \ set Vs \ ls ! last_index Vs V = ls' ! last_index Vs V" and size: "size ls = size ls'" "size Vs < size ls'" shows "l'(V := l V) \\<^sub>m [Vs [\] ls']" (*<*) proof - have "l'(V := l V) \\<^sub>m [Vs [\] ls', V\v](V := l V)" using 1 by(rule map_le_upd) also have "\ = [Vs [\] ls'](V := l V)" by simp also have "\ \\<^sub>m [Vs [\] ls']" proof (cases "l V") case None thus ?thesis by simp next case (Some w) hence "[Vs [\] ls] V = Some w" using 0 by(force simp add: map_le_def split:if_splits) hence VinVs: "V \ set Vs" and w: "w = ls ! last_index Vs V" using size by(auto simp add:fun_upds_apply split:if_splits) hence "w = ls' ! last_index Vs V" using hidden[OF VinVs] by simp hence "[Vs [\] ls'](V := l V) = [Vs [\] ls']" using Some size VinVs by(simp add: map_upds_upd_conv_last_index) thus ?thesis by simp qed finally show ?thesis . qed (*>*) (*<*) declare fun_upd_apply[simp del] (*>*) text\\noindent The main theorem: \ theorem assumes wf: "wwf_J_prog P" shows eval\<^sub>1_eval: "P \ \e,(h,l,sh)\ \ \e',(h',l',sh')\ \ (\Vs ls. \ fv e \ set Vs; l \\<^sub>m [Vs[\]ls]; size Vs + max_vars e \ size ls \ \ \ls'. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs e,(h,ls,sh)\ \ \fin\<^sub>1 e',(h',ls',sh')\ \ l' \\<^sub>m [Vs[\]ls'])" (*<*) (is "_ \ (\Vs ls. PROP ?P e h l sh e' h' l' sh' Vs ls)" is "_ \ (\Vs ls. \ _; _; _ \ \ \ls'. ?Post e h l sh e' h' l' sh' Vs ls ls')") (*>*) and evals\<^sub>1_evals: "P \ \es,(h,l,sh)\ [\] \es',(h',l',sh')\ \ (\Vs ls. \ fvs es \ set Vs; l \\<^sub>m [Vs[\]ls]; size Vs + max_varss es \ size ls \ \ \ls'. compP\<^sub>1 P \\<^sub>1 \compEs\<^sub>1 Vs es,(h,ls,sh)\ [\] \compEs\<^sub>1 Vs es',(h',ls',sh')\ \ l' \\<^sub>m [Vs[\]ls'])" (*<*) (is "_ \ (\Vs ls. PROP ?Ps es h l sh es' h' l' sh' Vs ls)" is "_ \ (\Vs ls. \ _; _; _\ \ \ls'. ?Posts es h l sh es' h' l' sh' Vs ls ls')") proof (induct rule:eval_evals_inducts) case Nil thus ?case by(fastforce intro!:Nil\<^sub>1) next case (Cons e h l sh v h' l' sh' es es' h\<^sub>2 l\<^sub>2 sh\<^sub>2) have "PROP ?P e h l sh (Val v) h' l' sh' Vs ls" by fact with Cons.prems obtain ls' where 1: "?Post e h l sh (Val v) h' l' sh' Vs ls ls'" "size ls = size ls'" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h' l' sh' es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls'" by fact with 1 Cons.prems obtain ls\<^sub>2 where 2: "?Posts es h' l' sh' es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls' ls\<^sub>2" by(auto) from 1 2 Cons show ?case by(auto intro!:Cons\<^sub>1) next case ConsThrow thus ?case by(fastforce intro!:ConsThrow\<^sub>1 dest: eval_final) next case (Block e h l V sh e' h' l' sh' T) let ?Vs = "Vs @ [V]" have IH: "\fv e \ set ?Vs; l(V := None) \\<^sub>m [?Vs [\] ls]; size ?Vs + max_vars e \ size ls\ \ \ls'. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e,(h,ls,sh)\ \ \fin\<^sub>1 e',(h', ls',sh')\ \ l' \\<^sub>m [?Vs [\] ls']" and fv: "fv {V:T; e} \ set Vs" and rel: "l \\<^sub>m [Vs [\] ls]" and len: "length Vs + max_vars {V:T; e} \ length ls" by fact+ have len': "length Vs < length ls" using len by auto have "fv e \ set ?Vs" using fv by auto moreover have "l(V := None) \\<^sub>m [?Vs [\] ls]" using rel len' by simp moreover have "size ?Vs + max_vars e \ size ls" using len by simp ultimately obtain ls' where 1: "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e,(h,ls,sh)\ \ \fin\<^sub>1 e',(h',ls',sh')\" and rel': "l' \\<^sub>m [?Vs [\] ls']" using IH by blast have [simp]: "length ls = length ls'" by(rule eval\<^sub>1_preserves_len[OF 1]) show "\ls'. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs {V:T; e},(h,ls,sh)\ \ \fin\<^sub>1 e',(h',ls',sh')\ \ l'(V := l V) \\<^sub>m [Vs [\] ls']" (is "\ls'. ?R ls'") proof show "?R ls'" proof show "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs {V:T; e},(h,ls,sh)\ \ \fin\<^sub>1 e',(h',ls',sh')\" using 1 by(simp add:Block\<^sub>1) next show "l'(V := l V) \\<^sub>m [Vs [\] ls']" proof - have "l' \\<^sub>m [Vs [\] ls', V \ ls' ! length Vs]" using len' rel' by simp moreover { assume VinVs: "V \ set Vs" hence "hidden (Vs @ [V]) (last_index Vs V)" by(rule hidden_last_index) hence "unmod (compE\<^sub>1 (Vs @ [V]) e) (last_index Vs V)" by(rule hidden_unmod) moreover have "last_index Vs V < length ls" using len' VinVs by simp ultimately have "ls ! last_index Vs V = ls' ! last_index Vs V" by(rule eval\<^sub>1_preserves_unmod[OF 1]) } ultimately show ?thesis using Block_lem[OF rel] len' by auto qed qed qed next case (TryThrow e' h l sh a h' l' sh' D fs C V e\<^sub>2) have "PROP ?P e' h l sh (Throw a) h' l' sh' Vs ls" by fact with TryThrow.prems obtain ls' where 1: "?Post e' h l sh (Throw a) h' l' sh' Vs ls ls'" by(auto) show ?case using 1 TryThrow.hyps by(auto intro!:eval\<^sub>1_evals\<^sub>1.TryThrow\<^sub>1) next case (TryCatch e\<^sub>1 h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C e\<^sub>2 V e' h\<^sub>2 l\<^sub>2 sh\<^sub>2) let ?e = "try e\<^sub>1 catch(C V) e\<^sub>2" have IH\<^sub>1: "\fv e\<^sub>1 \ set Vs; l \\<^sub>m [Vs [\] ls]; size Vs + max_vars e\<^sub>1 \ length ls\ \ \ls\<^sub>1. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs e\<^sub>1,(h,ls,sh)\ \ \fin\<^sub>1 (Throw a),(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\ \ l\<^sub>1 \\<^sub>m [Vs [\] ls\<^sub>1]" and fv: "fv ?e \ set Vs" and rel: "l \\<^sub>m [Vs [\] ls]" and len: "length Vs + max_vars ?e \ length ls" by fact+ have "fv e\<^sub>1 \ set Vs" using fv by auto moreover have "length Vs + max_vars e\<^sub>1 \ length ls" using len by(auto) ultimately obtain ls\<^sub>1 where 1: "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs e\<^sub>1,(h,ls,sh)\ \ \Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\" and rel\<^sub>1: "l\<^sub>1 \\<^sub>m [Vs [\] ls\<^sub>1]" using IH\<^sub>1 rel by fastforce from 1 have [simp]: "size ls = size ls\<^sub>1" by(rule eval\<^sub>1_preserves_len) let ?Vs = "Vs @ [V]" let ?ls = "(ls\<^sub>1[size Vs:=Addr a])" have IH\<^sub>2: "\fv e\<^sub>2 \ set ?Vs; l\<^sub>1(V \ Addr a) \\<^sub>m [?Vs [\] ?ls]; length ?Vs + max_vars e\<^sub>2 \ length ?ls\ \ \ls\<^sub>2. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e\<^sub>2,(h\<^sub>1,?ls,sh\<^sub>1)\ \ \fin\<^sub>1 e',(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\ \ l\<^sub>2 \\<^sub>m [?Vs [\] ls\<^sub>2]" by fact have len\<^sub>1: "size Vs < size ls\<^sub>1" using len by(auto) have "fv e\<^sub>2 \ set ?Vs" using fv by auto moreover have "l\<^sub>1(V \ Addr a) \\<^sub>m [?Vs [\] ?ls]" using rel\<^sub>1 len\<^sub>1 by simp moreover have "length ?Vs + max_vars e\<^sub>2 \ length ?ls" using len by(auto) ultimately obtain ls\<^sub>2 where 2: "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 ?Vs e\<^sub>2,(h\<^sub>1,?ls,sh\<^sub>1)\ \ \fin\<^sub>1 e',(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" and rel\<^sub>2: "l\<^sub>2 \\<^sub>m [?Vs [\] ls\<^sub>2]" using IH\<^sub>2 by blast from 2 have [simp]: "size ls\<^sub>1 = size ls\<^sub>2" by(fastforce dest: eval\<^sub>1_preserves_len) show "\ls\<^sub>2. compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs ?e,(h,ls,sh)\ \ \fin\<^sub>1 e',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\ \ l\<^sub>2(V := l\<^sub>1 V) \\<^sub>m [Vs [\] ls\<^sub>2]" (is "\ls\<^sub>2. ?R ls\<^sub>2") proof show "?R ls\<^sub>2" proof have hp: "h\<^sub>1 a = Some (D, fs)" by fact have "P \ D \\<^sup>* C" by fact hence caught: "compP\<^sub>1 P \ D \\<^sup>* C" by simp from TryCatch\<^sub>1[OF 1 _ caught len\<^sub>1 2, OF hp] show "compP\<^sub>1 P \\<^sub>1 \compE\<^sub>1 Vs ?e,(h,ls,sh)\ \ \fin\<^sub>1 e',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\" by simp next show "l\<^sub>2(V := l\<^sub>1 V) \\<^sub>m [Vs [\] ls\<^sub>2]" proof - have "l\<^sub>2 \\<^sub>m [Vs [\] ls\<^sub>2, V \ ls\<^sub>2 ! length Vs]" using len\<^sub>1 rel\<^sub>2 by simp moreover { assume VinVs: "V \ set Vs" hence "hidden (Vs @ [V]) (last_index Vs V)" by(rule hidden_last_index) hence "unmod (compE\<^sub>1 (Vs @ [V]) e\<^sub>2) (last_index Vs V)" by(rule hidden_unmod) moreover have "last_index Vs V < length ?ls" using len\<^sub>1 VinVs by simp ultimately have "?ls ! last_index Vs V = ls\<^sub>2 ! last_index Vs V" by(rule eval\<^sub>1_preserves_unmod[OF 2]) moreover have "last_index Vs V < size Vs" using VinVs by simp ultimately have "ls\<^sub>1 ! last_index Vs V = ls\<^sub>2 ! last_index Vs V" using len\<^sub>1 by(simp del:size_last_index_conv) } ultimately show ?thesis using Block_lem[OF rel\<^sub>1] len\<^sub>1 by simp qed qed qed next case Try thus ?case by(fastforce intro!:Try\<^sub>1) next case Throw thus ?case by(fastforce intro!:Throw\<^sub>1) next case ThrowNull thus ?case by(fastforce intro!:ThrowNull\<^sub>1) next case ThrowThrow thus ?case by(fastforce intro!:ThrowThrow\<^sub>1) next case (CondT e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 e\<^sub>2) have "PROP ?P e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with CondT.prems obtain ls\<^sub>1 where 1: "?Post e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 CondT.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 show ?case by(auto intro!:CondT\<^sub>1) next case (CondF e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 e\<^sub>1 Vs ls) have "PROP ?P e h l sh false h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with CondF.prems obtain ls\<^sub>1 where 1: "?Post e h l sh false h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 CondF.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 show ?case by(auto intro!:CondF\<^sub>1) next case CondThrow thus ?case by(fastforce intro!:CondThrow\<^sub>1) next case (Seq e h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2) have "PROP ?P e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with Seq.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 Seq.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 Seq show ?case by(auto intro!:Seq\<^sub>1) next case SeqThrow thus ?case by(fastforce intro!:SeqThrow\<^sub>1) next case WhileF thus ?case by(fastforce intro!:eval\<^sub>1_evals\<^sub>1.intros) next case (WhileT e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 c v h\<^sub>2 l\<^sub>2 sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) have "PROP ?P e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with WhileT.prems obtain ls\<^sub>1 where 1: "?Post e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 WhileT.prems obtain ls\<^sub>2 where 2: "?Post c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" "size ls\<^sub>1 = size ls\<^sub>2" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P (While (e) c) h\<^sub>2 l\<^sub>2 sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 Vs ls\<^sub>2" by fact with 1 2 WhileT.prems obtain ls\<^sub>3 where 3: "?Post (While (e) c) h\<^sub>2 l\<^sub>2 sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 Vs ls\<^sub>2 ls\<^sub>3" by(auto) from 1 2 3 show ?case by(auto intro!:WhileT\<^sub>1) next case (WhileBodyThrow e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 c e' h\<^sub>2 l\<^sub>2 sh\<^sub>2) have "PROP ?P e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with WhileBodyThrow.prems obtain ls\<^sub>1 where 1: "?Post e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 WhileBodyThrow.prems obtain ls\<^sub>2 where 2: "?Post c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by auto from 1 2 show ?case by(auto intro!:WhileBodyThrow\<^sub>1) next case WhileCondThrow thus ?case by(fastforce intro!:WhileCondThrow\<^sub>1) next case New thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case NewFail thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case NewInit then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case NewInitOOM then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case NewInitThrow then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case Cast thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case CastNull thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case CastThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (CastFail e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C) have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with CastFail.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" by auto show ?case using 1 CastFail.hyps by(auto intro!:CastFail\<^sub>1[where D=D]) next case Val thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (BinOp e h l sh v\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 v\<^sub>2 h\<^sub>2 l\<^sub>2 sh\<^sub>2 bop v) have "PROP ?P e h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with BinOp.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v\<^sub>2) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 BinOp.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v\<^sub>2) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 BinOp show ?case by(auto intro!:BinOp\<^sub>1) next case (BinOpThrow2 e\<^sub>0 h l sh v\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 e h\<^sub>2 l\<^sub>2 sh\<^sub>2 bop) have "PROP ?P e\<^sub>0 h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with BinOpThrow2.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>0 h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 BinOpThrow2.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 BinOpThrow2 show ?case by(auto intro!:BinOpThrow\<^sub>2\<^sub>1) next case BinOpThrow1 thus ?case by(fastforce intro!:eval\<^sub>1_evals\<^sub>1.intros) next case Var thus ?case by(force intro!:Var\<^sub>1 simp add: map_le_def fun_upds_apply) next case LAss thus ?case by(fastforce simp add: LAss_lem intro:eval\<^sub>1_evals\<^sub>1.intros dest:eval\<^sub>1_preserves_len) next case LAssThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case FAcc thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case FAccNull thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case FAccThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (FAccNone e h l sh a h' l' sh' C fs F D) have "PROP ?P e h l sh (addr a) h' l' sh' Vs ls" by fact with FAccNone.prems obtain ls\<^sub>2 where 2: "?Post e h l sh (addr a) h' l' sh' Vs ls ls\<^sub>2" by(auto) from 2 FAccNone show ?case by(rule_tac x = ls\<^sub>2 in exI, auto elim!: FAccNone\<^sub>1) next case (FAccStatic e h l sh a h' l' sh' C fs F t D) have "PROP ?P e h l sh (addr a) h' l' sh' Vs ls" by fact with FAccStatic.prems obtain ls\<^sub>2 where 2: "?Post e h l sh (addr a) h' l' sh' Vs ls ls\<^sub>2" by(auto) from 2 FAccStatic show ?case by(rule_tac x = ls\<^sub>2 in exI, auto elim!: FAccStatic\<^sub>1) next case SFAcc then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (SFAccInit C F t D sh h l v' h' l' sh' sfs i v) have "PROP ?P (INIT D ([D],False) \ unit) h l sh (Val v') h' l' sh' Vs ls" by fact with SFAccInit.prems obtain ls\<^sub>2 where 1: "?Post (INIT D ([D],False) \ unit) h l sh (Val v') h' l' sh' Vs ls ls\<^sub>2" by(auto) from 1 SFAccInit show ?case by(rule_tac x = ls\<^sub>2 in exI, auto intro: SFAccInit\<^sub>1) next case (SFAccInitThrow C F t D sh h l a h' l' sh') have "PROP ?P (INIT D ([D],False) \ unit) h l sh (throw a) h' l' sh' Vs ls" by fact with SFAccInitThrow.prems obtain ls\<^sub>2 where 1: "?Post (INIT D ([D],False) \ unit) h l sh (throw a) h' l' sh' Vs ls ls\<^sub>2" by(auto) from 1 SFAccInitThrow show ?case by(rule_tac x = ls\<^sub>2 in exI, auto intro: SFAccInitThrow\<^sub>1) next case SFAccNone then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case SFAccNonStatic then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (FAss e\<^sub>1 h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs fs' F D h\<^sub>2') have "PROP ?P e\<^sub>1 h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with FAss.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAss.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAss show ?case by(auto intro!:FAss\<^sub>1) next case (FAssNull e\<^sub>1 h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 F D) have "PROP ?P e\<^sub>1 h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with FAssNull.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAssNull.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAssNull show ?case by(auto intro!:FAssNull\<^sub>1) next case FAssThrow1 thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (FAssThrow2 e\<^sub>1 h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 e h\<^sub>2 l\<^sub>2 sh\<^sub>2 F D) have "PROP ?P e\<^sub>1 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with FAssThrow2.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAssThrow2.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAssThrow2 show ?case by(auto intro!:FAssThrow\<^sub>2\<^sub>1) next case (FAssNone e\<^sub>1 h l sh a h' l' sh' e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F D) have "PROP ?P e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls" by fact with FAssNone.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAssNone.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAssNone show ?case by(auto intro!:FAssNone\<^sub>1) next case (FAssStatic e\<^sub>1 h l sh a h' l' sh' e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F t D) have "PROP ?P e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls" by fact with FAssStatic.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 FAssStatic.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 FAssStatic show ?case by(auto intro!:FAssStatic\<^sub>1) next case SFAss then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (SFAssInit e\<^sub>2 h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D v' h' l' sh' sfs i sfs' sh'') have "PROP ?P e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with SFAssInit.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) then have Init_size: "length Vs \ length ls\<^sub>1" using SFAssInit.prems(3) by linarith have "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h' l' sh' Vs ls\<^sub>1" by fact with 1 Init_size SFAssInit.prems obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h' l' sh' Vs ls\<^sub>1 ls\<^sub>2" by auto from 1 2 SFAssInit show ?case by(auto simp add: comp_def intro!: SFAssInit\<^sub>1 dest!:evals_final) next case (SFAssInitThrow e\<^sub>2 h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D a h\<^sub>2 l\<^sub>2 sh\<^sub>2) have "PROP ?P e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with SFAssInitThrow.prems obtain ls\<^sub>1 where 1: "?Post e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) then have Init_size: "length Vs \ length ls\<^sub>1" using SFAssInitThrow.prems(3) by linarith have "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 Init_size SFAssInitThrow.prems obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by auto from 1 2 SFAssInitThrow show ?case by(auto simp add: comp_def intro!: SFAssInitThrow\<^sub>1 dest!:evals_final) next case SFAssThrow then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (SFAssNone e\<^sub>2 h l sh v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C F D) have "PROP ?P e\<^sub>2 h l sh (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls" by fact with SFAssNone.prems obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h l sh (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls ls\<^sub>2" by(auto) from 2 SFAssNone show ?case by(rule_tac x = ls\<^sub>2 in exI, auto elim!: SFAssNone\<^sub>1) next case SFAssNonStatic then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (CallNull e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 es vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 M) have "PROP ?P e h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with CallNull.prems obtain ls\<^sub>1 where 1: "?Post e h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 CallNull.prems obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 CallNull show ?case by (auto simp add: comp_def elim!: CallNull\<^sub>1) next case CallObjThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (CallParamsThrow e h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 es vs ex es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 M) have "PROP ?P e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with CallParamsThrow.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 CallParamsThrow.prems obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 CallParamsThrow show ?case by (auto simp add: comp_def elim!: CallParamsThrow\<^sub>1 dest!:evals_final) next case (CallNone e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M) have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with CallNone.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 CallNone.prems obtain ls\<^sub>2 where 2: "?Posts ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 CallNone show ?case by (auto simp add: comp_def elim!: CallNone\<^sub>1 dest!:evals_final sees_method_compPD) next case (CallStatic e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body D) have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with CallStatic.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) let ?Vs = pns have mdecl: "P \ C sees M,Static: Ts\T = (pns, body) in D" by fact have mdecl\<^sub>1: "compP\<^sub>1 P \ C sees M,Static: Ts\T = (compE\<^sub>1 ?Vs body) in D" using sees_method_compP[OF mdecl, of "\b (pns,e). compE\<^sub>1 (case b of NonStatic \ this#pns | Static \ pns) e"] by(simp) have "PROP ?Ps ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 CallStatic.prems obtain ls\<^sub>2 where 2: "?Posts ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 mdecl\<^sub>1 CallStatic show ?case by (auto simp add: comp_def elim!: CallStatic\<^sub>1 dest!:evals_final) next case (Call e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 es vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body D l\<^sub>2' b' h\<^sub>3 l\<^sub>3 sh\<^sub>3) have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with Call.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?Ps es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 Call.prems obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" "size ls\<^sub>1 = size ls\<^sub>2" by(auto intro!:evals\<^sub>1_preserves_len) let ?Vs = "this#pns" let ?ls = "Addr a # vs @ replicate (max_vars body) undefined" have mdecl: "P \ C sees M,NonStatic: Ts\T = (pns, body) in D" by fact have fv_body: "fv body \ set ?Vs" and wf_size: "size Ts = size pns" using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def) have mdecl\<^sub>1: "compP\<^sub>1 P \ C sees M,NonStatic: Ts\T = (compE\<^sub>1 ?Vs body) in D" using sees_method_compP[OF mdecl, of "\b (pns,e). compE\<^sub>1 (case b of NonStatic \ this#pns | Static \ pns) e"] by(simp) have [simp]: "l\<^sub>2' = [this \ Addr a, pns [\] vs]" by fact have Call_size: "size vs = size pns" by fact have "PROP ?P body h\<^sub>2 l\<^sub>2' sh\<^sub>2 b' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls" by fact with 1 2 fv_body Call_size Call.prems obtain ls\<^sub>3 where 3: "?Post body h\<^sub>2 l\<^sub>2' sh\<^sub>2 b' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls ls\<^sub>3" by(auto) have hp: "h\<^sub>2 a = Some (C, fs)" by fact from 1 2 3 hp mdecl\<^sub>1 wf_size Call_size show ?case by(fastforce simp add: comp_def intro!: Call\<^sub>1 dest!:evals_final) next case (SCallParamsThrow es h l sh vs ex es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 C M) have "PROP ?Ps es h l sh (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls" by fact with SCallParamsThrow.prems obtain ls\<^sub>2 where 2: "?Posts es h l sh (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls ls\<^sub>2" by(auto) from 2 SCallParamsThrow show ?case by (fastforce simp add: comp_def elim!: SCallParamsThrow\<^sub>1 dest!:evals_final) next case (SCall ps h l sh vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C M Ts T pns body D sfs l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) have "PROP ?Ps ps h l sh (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls" by fact with SCall.prems obtain ls\<^sub>2 where 2: "?Posts ps h l sh (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls ls\<^sub>2" "size ls = size ls\<^sub>2" by(auto intro!:evals\<^sub>1_preserves_len) let ?Vs = "pns" let ?ls = "vs @ replicate (max_vars body) undefined" have mdecl: "P \ C sees M,Static: Ts\T = (pns, body) in D" by fact have fv_body: "fv body \ set ?Vs" and wf_size: "size Ts = size pns" using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def) have mdecl\<^sub>1: "compP\<^sub>1 P \ C sees M,Static: Ts\T = (compE\<^sub>1 ?Vs body) in D" using sees_method_compP[OF mdecl, of "\b (pns,e). compE\<^sub>1 (case b of NonStatic \ this#pns | Static \ pns) e"] by(simp) have [simp]: "l\<^sub>2' = [pns [\] vs]" by fact have SCall_size: "size vs = size pns" by fact have "PROP ?P body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls" by fact with 2 fv_body SCall_size SCall.prems obtain ls\<^sub>3 where 3: "?Post body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls ls\<^sub>3" by(auto) have shp: "sh\<^sub>2 D = \(sfs, Done)\ \ M = clinit \ sh\<^sub>2 D = \(sfs, Processing)\" by fact from 2 3 shp mdecl\<^sub>1 wf_size SCall_size show ?case by(fastforce simp add: comp_def intro!: SCall\<^sub>1 dest!:evals_final) next case (SCallNone ps h l sh vs h' l' sh' C M) have "PROP ?Ps ps h l sh (map Val vs) h' l' sh' Vs ls" by fact with SCallNone.prems obtain ls\<^sub>2 where 2: "?Posts ps h l sh (map Val vs) h' l' sh' Vs ls ls\<^sub>2" by(auto) from 2 SCallNone show ?case by(rule_tac x = ls\<^sub>2 in exI, auto simp add: comp_def elim!: SCallNone\<^sub>1 dest!:evals_final sees_method_compPD) next case (SCallNonStatic ps h l sh vs h' l' sh' C M Ts T pns body D) let ?Vs = "this#pns" have mdecl: "P \ C sees M,NonStatic: Ts\T = (pns, body) in D" by fact have mdecl\<^sub>1: "compP\<^sub>1 P \ C sees M,NonStatic: Ts\T = (compE\<^sub>1 ?Vs body) in D" using sees_method_compP[OF mdecl, of "\b (pns,e). compE\<^sub>1 (case b of NonStatic \ this#pns | Static \ pns) e"] by(simp) have "PROP ?Ps ps h l sh (map Val vs) h' l' sh' Vs ls" by fact with SCallNonStatic.prems obtain ls\<^sub>2 where 2: "?Posts ps h l sh (map Val vs) h' l' sh' Vs ls ls\<^sub>2" by(auto) from 2 mdecl\<^sub>1 SCallNonStatic show ?case by (auto simp add: comp_def elim!: SCallNonStatic\<^sub>1 dest!:evals_final) next case (SCallInitThrow ps h l sh vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D a h\<^sub>2 l\<^sub>2 sh\<^sub>2) have "PROP ?Ps ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with SCallInitThrow.prems obtain ls\<^sub>1 where 1: "?Posts ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1" by(auto intro!:evals\<^sub>1_preserves_len) then have Init_size: "length Vs \ length ls\<^sub>1" using SCallInitThrow.prems(3) by linarith have "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 Init_size SCallInitThrow.prems obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by auto let ?Vs = "pns" have mdecl: "P \ C sees M,Static: Ts\T = (pns, body) in D" by fact have mdecl\<^sub>1: "compP\<^sub>1 P \ C sees M,Static: Ts\T = (compE\<^sub>1 ?Vs body) in D" using sees_method_compP[OF mdecl, of "\b (pns,e). compE\<^sub>1 (case b of NonStatic \ this#pns | Static \ pns) e"] by(simp) from 1 2 mdecl\<^sub>1 SCallInitThrow show ?case by(auto simp add: comp_def intro!: SCallInitThrow\<^sub>1 dest!:evals_final) next case (SCallInit ps h l sh vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) have "PROP ?Ps ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact with SCallInit.prems obtain ls\<^sub>1 where 1: "?Posts ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1" by(auto intro!:evals\<^sub>1_preserves_len) then have Init_size: "length Vs \ length ls\<^sub>1" using SCallInit.prems(3) by linarith have "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact with 1 Init_size SCallInit.prems obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \ unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by auto let ?Vs = "pns" let ?ls = "vs @ replicate (max_vars body) undefined" have mdecl: "P \ C sees M,Static: Ts\T = (pns, body) in D" by fact have fv_body: "fv body \ set ?Vs" and wf_size: "size Ts = size pns" using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def) have mdecl\<^sub>1: "compP\<^sub>1 P \ C sees M,Static: Ts\T = (compE\<^sub>1 ?Vs body) in D" using sees_method_compP[OF mdecl, of "\b (pns,e). compE\<^sub>1 (case b of NonStatic \ this#pns | Static \ pns) e"] by(simp) have [simp]: "l\<^sub>2' = [pns [\] vs]" by fact have SCall_size: "size vs = size pns" by fact have nclinit: "M \ clinit" by fact have "PROP ?P body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls" by fact with 2 fv_body SCall_size SCallInit.prems obtain ls\<^sub>3 where 3: "?Post body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls ls\<^sub>3" by(auto) have shp: "\sfs. sh\<^sub>1 D = \(sfs, Done)\" by fact from 1 2 3 shp mdecl\<^sub>1 wf_size SCall_size nclinit show ?case by(auto simp add: comp_def intro!: SCallInit\<^sub>1 dest!:evals_final) next \ \ init cases \ case InitFinal then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (InitNone sh C C' Cs e h l e' h' l' sh') let ?sh1 = "sh(C \ (sblank P C, Prepared))" have "PROP ?P (INIT C' (C # Cs,False) \ e) h l ?sh1 e' h' l' sh' Vs ls" by fact with InitNone.prems obtain ls\<^sub>2 where 2: "?Post (INIT C' (C # Cs,False) \ e) h l ?sh1 e' h' l' sh' Vs ls ls\<^sub>2" by(auto) from 2 InitNone show ?case by (auto elim!: InitNone\<^sub>1) next case InitDone then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case InitProcessing then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case InitError then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case InitObject then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (InitNonObject sh C sfs D fs ms sh' C' Cs e h l e' h1 l1 sh1) let ?f = "(\b (pns,body). compE\<^sub>1 (case b of NonStatic \ this#pns | Static \ pns) body)" have cls: "class (compP ?f P) C = \(D,fs,map (compM ?f) ms)\" by(rule class_compP[OF InitNonObject.hyps(3)]) have "PROP ?P (INIT C' (D # C # Cs,False) \ e) h l sh' e' h1 l1 sh1 Vs ls" by fact with InitNonObject.prems obtain ls\<^sub>2 where 2: "?Post (INIT C' (D # C # Cs,False) \ e) h l sh' e' h1 l1 sh1 Vs ls ls\<^sub>2" by(auto) from 2 cls InitNonObject show ?case by (auto elim!: InitNonObject\<^sub>1) next case InitRInit then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) next case (RInit e h l sh v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1) have "PROP ?P e h l sh (Val v) h' l' sh' Vs ls" by fact with RInit.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v) h' l' sh' Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have "PROP ?P (INIT C' (Cs,True) \ e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1" by fact with 1 RInit.prems obtain ls\<^sub>2 where 2: "?Post (INIT C' (Cs,True) \ e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1 ls\<^sub>2" by(auto) from 1 2 RInit show ?case by (auto elim!: RInit\<^sub>1) next case (RInitInitFail e h l sh a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1) have "PROP ?P e h l sh (throw a) h' l' sh' Vs ls" by fact with RInitInitFail.prems obtain ls\<^sub>1 where 1: "?Post e h l sh (throw a) h' l' sh' Vs ls ls\<^sub>1" "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len) have fv: "fv (RI (D,throw a) ; Cs \ e') \ set Vs" using RInitInitFail.hyps(1) eval_final RInitInitFail.prems(1) subset_eq by fastforce have l': "l' \\<^sub>m [Vs [\] ls\<^sub>1]" by (simp add: "1"(1)) have "PROP ?P (RI (D,throw a) ; Cs \ e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1" by fact with 1 eval_final[OF RInitInitFail.hyps(1)] RInitInitFail.prems obtain ls\<^sub>2 where 2: "?Post (RI (D,throw a) ; Cs \ e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1 ls\<^sub>2" by fastforce from 1 2 RInitInitFail show ?case by(fastforce simp add: comp_def intro!: RInitInitFail\<^sub>1 dest!:eval_final) next case RInitFailFinal then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros) qed (*>*) subsection\Preservation of well-formedness\ text\ The compiler preserves well-formedness. Is less trivial than it may appear. We start with two simple properties: preservation of well-typedness \ lemma compE\<^sub>1_pres_wt: "\Vs Ts U. \ P,[Vs[\]Ts] \ e :: U; size Ts = size Vs \ \ compP f P,Ts \\<^sub>1 compE\<^sub>1 Vs e :: U" and "\Vs Ts Us. \ P,[Vs[\]Ts] \ es [::] Us; size Ts = size Vs \ \ compP f P,Ts \\<^sub>1 compEs\<^sub>1 Vs es [::] Us" (*<*) -apply(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) -apply clarsimp -apply(fastforce) -apply clarsimp -apply(fastforce split:bop.splits) -apply (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) -apply (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) -apply(fastforce) -apply(fastforce) -apply(fastforce) -apply(fastforce) -apply (fastforce dest!: sees_method_compP[where f = f]) -apply (fastforce dest!: sees_method_compP[where f = f]) -apply (fastforce simp:nth_append) -apply (fastforce) -apply (fastforce) -apply (fastforce) -apply (fastforce) -apply (fastforce simp:nth_append) -apply simp -apply simp -apply simp -apply (fastforce) -done +proof(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct) + case Var then show ?case + by (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) +next + case LAss then show ?case + by (fastforce simp:map_upds_apply_eq_Some split:if_split_asm) +next + case Call then show ?case + by (fastforce dest!: sees_method_compP[where f = f]) +next + case SCall then show ?case + by (fastforce dest!: sees_method_compP[where f = f]) +next + case Block then show ?case by (fastforce simp:nth_append) +next + case TryCatch then show ?case by (fastforce simp:nth_append) +qed fastforce+ (*>*) text\\noindent and the correct block numbering: \ lemma \: "\Vs n. size Vs = n \ \ (compE\<^sub>1 Vs e) n" and \s: "\Vs n. size Vs = n \ \s (compEs\<^sub>1 Vs es) n" (*<*)by (induct e and es rule: \.induct \s.induct) (force | simp,metis length_append_singleton)+(*>*) text\ The main complication is preservation of definite assignment @{term"\"}. \ lemma image_last_index: "A \ set(xs@[x]) \ last_index (xs @ [x]) ` A = (if x \ A then insert (size xs) (last_index xs ` (A-{x})) else last_index xs ` A)" (*<*) by(auto simp:image_def) (*>*) lemma A_compE\<^sub>1_None[simp]: "\Vs. \ e = None \ \ (compE\<^sub>1 Vs e) = None" and "\Vs. \s es = None \ \s (compEs\<^sub>1 Vs es) = None" (*<*)by(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct)(auto simp:hyperset_defs)(*>*) lemma A_compE\<^sub>1: "\A Vs. \ \ e = \A\; fv e \ set Vs \ \ \ (compE\<^sub>1 Vs e) = \last_index Vs ` A\" and "\A Vs. \ \s es = \A\; fvs es \ set Vs \ \ \s (compEs\<^sub>1 Vs es) = \last_index Vs ` A\" (*<*) proof(induct e and es rule: fv.induct fvs.induct) case (Block V' T e) hence "fv e \ set (Vs@[V'])" by fastforce moreover obtain B where "\ e = \B\" using Block.prems by(simp add: hyperset_defs) moreover from calculation have "B \ set (Vs@[V'])" by(auto dest!:A_fv) ultimately show ?case using Block by(auto simp add: hyperset_defs image_last_index last_index_size_conv) next case (TryCatch e\<^sub>1 C V' e\<^sub>2) hence fve\<^sub>2: "fv e\<^sub>2 \ set (Vs@[V'])" by auto show ?case proof (cases "\ e\<^sub>1") assume A\<^sub>1: "\ e\<^sub>1 = None" then obtain A\<^sub>2 where A\<^sub>2: "\ e\<^sub>2 = \A\<^sub>2\" using TryCatch by(simp add:hyperset_defs) hence "A\<^sub>2 \ set (Vs@[V'])" using TryCatch.prems A_fv[OF A\<^sub>2] by simp blast thus ?thesis using TryCatch fve\<^sub>2 A\<^sub>1 A\<^sub>2 by(auto simp add:hyperset_defs image_last_index last_index_size_conv) next fix A\<^sub>1 assume A\<^sub>1: "\ e\<^sub>1 = \A\<^sub>1\" show ?thesis proof (cases "\ e\<^sub>2") assume A\<^sub>2: "\ e\<^sub>2 = None" then show ?case using TryCatch A\<^sub>1 by(simp add:hyperset_defs) next fix A\<^sub>2 assume A\<^sub>2: "\ e\<^sub>2 = \A\<^sub>2\" have "A\<^sub>1 \ set Vs" using TryCatch.prems A_fv[OF A\<^sub>1] by simp blast moreover have "A\<^sub>2 \ set (Vs@[V'])" using TryCatch.prems A_fv[OF A\<^sub>2] by simp blast ultimately show ?thesis using TryCatch A\<^sub>1 A\<^sub>2 by (auto simp add: Diff_subset_conv last_index_size_conv subsetD hyperset_defs dest!: sym [of _ A]) qed qed next case (Cond e e\<^sub>1 e\<^sub>2) { assume "\ e = None \ \ e\<^sub>1 = None \ \ e\<^sub>2 = None" hence ?case using Cond by(auto simp add:hyperset_defs image_Un) } moreover { fix A A\<^sub>1 A\<^sub>2 assume "\ e = \A\" and A\<^sub>1: "\ e\<^sub>1 = \A\<^sub>1\" and A\<^sub>2: "\ e\<^sub>2 = \A\<^sub>2\" moreover have "A\<^sub>1 \ set Vs" using Cond.prems A_fv[OF A\<^sub>1] by simp blast moreover have "A\<^sub>2 \ set Vs" using Cond.prems A_fv[OF A\<^sub>2] by simp blast ultimately have ?case using Cond by(auto simp add:hyperset_defs image_Un inj_on_image_Int[OF inj_on_last_index]) } ultimately show ?case by fastforce qed (auto simp add:hyperset_defs) (*>*) lemma D_None[iff]: "\ (e::'a exp) None" and [iff]: "\s (es::'a exp list) None" (*<*)by(induct e and es rule: \.induct \s.induct)(simp_all)(*>*) lemma D_last_index_compE\<^sub>1: "\A Vs. \ A \ set Vs; fv e \ set Vs \ \ \ e \A\ \ \ (compE\<^sub>1 Vs e) \last_index Vs ` A\" and "\A Vs. \ A \ set Vs; fvs es \ set Vs \ \ \s es \A\ \ \s (compEs\<^sub>1 Vs es) \last_index Vs ` A\" (*<*) proof(induct e and es rule: \.induct \s.induct) case (BinOp e\<^sub>1 bop e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using BinOp by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] BinOp.prems by auto have "A \ A\<^sub>1 \ set Vs" using BinOp.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using BinOp Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (FAss e\<^sub>1 F D e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using FAss by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] FAss.prems by auto have "A \ A\<^sub>1 \ set Vs" using FAss.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using FAss Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (Call e\<^sub>1 M es) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using Call by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] Call.prems by auto have "A \ A\<^sub>1 \ set Vs" using Call.prems A_fv[OF Some] by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` (A \ A\<^sub>1)\" using Call Some by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (TryCatch e\<^sub>1 C V e\<^sub>2) have "\ A\{V} \ set(Vs@[V]); fv e\<^sub>2 \ set(Vs@[V]); \ e\<^sub>2 \A\{V}\\ \ \ (compE\<^sub>1 (Vs@[V]) e\<^sub>2) \last_index (Vs@[V]) ` (A\{V})\" by fact hence "\ (compE\<^sub>1 (Vs@[V]) e\<^sub>2) \last_index (Vs@[V]) ` (A\{V})\" using TryCatch.prems by(simp add:Diff_subset_conv) moreover have "last_index (Vs@[V]) ` A \ last_index Vs ` A \ {size Vs}" using TryCatch.prems by(auto simp add: image_last_index split:if_split_asm) ultimately show ?case using TryCatch by(auto simp:hyperset_defs elim!:D_mono') next case (Seq e\<^sub>1 e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using Seq by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] Seq.prems by auto have "A \ A\<^sub>1 \ set Vs" using Seq.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using Seq Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (Cond e e\<^sub>1 e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e) \last_index Vs ` A\" by simp show ?case proof (cases "\ e") case None thus ?thesis using Cond by simp next case (Some B) have indexB: "\ (compE\<^sub>1 Vs e) = \last_index Vs ` B\" using A_compE\<^sub>1[OF Some] Cond.prems by auto have "A \ B \ set Vs" using Cond.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` (A \ B)\" and "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ B)\" using Cond Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A \ last_index Vs ` B\" and "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` B\" by(simp add: image_Un)+ thus ?thesis using IH\<^sub>1 indexB by auto qed next case (While e\<^sub>1 e\<^sub>2) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using While by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] While.prems by auto have "A \ A\<^sub>1 \ set Vs" using While.prems A_fv[OF Some] by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` (A \ A\<^sub>1)\" using While Some by auto hence "\ (compE\<^sub>1 Vs e\<^sub>2) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed next case (Block V T e) have "\ A-{V} \ set(Vs@[V]); fv e \ set(Vs@[V]); \ e \A-{V}\ \ \ \ (compE\<^sub>1 (Vs@[V]) e) \last_index (Vs@[V]) ` (A-{V})\" by fact hence "\ (compE\<^sub>1 (Vs@[V]) e) \last_index (Vs@[V]) ` (A-{V})\" using Block.prems by(simp add:Diff_subset_conv) moreover have "size Vs \ last_index Vs ` A" using Block.prems by(auto simp add:image_def size_last_index_conv) ultimately show ?case using Block by(auto simp add: image_last_index Diff_subset_conv hyperset_defs elim!: D_mono') next case (Cons_exp e\<^sub>1 es) hence IH\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) \last_index Vs ` A\" by simp show ?case proof (cases "\ e\<^sub>1") case None thus ?thesis using Cons_exp by simp next case (Some A\<^sub>1) have indexA\<^sub>1: "\ (compE\<^sub>1 Vs e\<^sub>1) = \last_index Vs ` A\<^sub>1\" using A_compE\<^sub>1[OF Some] Cons_exp.prems by auto have "A \ A\<^sub>1 \ set Vs" using Cons_exp.prems A_fv[OF Some] by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` (A \ A\<^sub>1)\" using Cons_exp Some by auto hence "\s (compEs\<^sub>1 Vs es) \last_index Vs ` A \ last_index Vs ` A\<^sub>1\" by(simp add: image_Un) thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto qed qed (simp_all add:hyperset_defs) (*>*) lemma last_index_image_set: "distinct xs \ last_index xs ` set xs = {..*) lemma D_compE\<^sub>1: "\ \ e \set Vs\; fv e \ set Vs; distinct Vs \ \ \ (compE\<^sub>1 Vs e) \{.." (*<*)by(fastforce dest!: D_last_index_compE\<^sub>1[OF subset_refl] simp add:last_index_image_set)(*>*) lemma D_compE\<^sub>1': assumes "\ e \set(V#Vs)\" and "fv e \ set(V#Vs)" and "distinct(V#Vs)" shows "\ (compE\<^sub>1 (V#Vs) e) \{..length Vs}\" (*<*) proof - have "{..size Vs} = {..1) qed (*>*) lemma compP\<^sub>1_pres_wf: "wf_J_prog P \ wf_J\<^sub>1_prog (compP\<^sub>1 P)" (*<*) -apply simp -apply(rule wf_prog_compPI) - prefer 2 apply assumption -apply(case_tac m) -apply(simp add:wf_mdecl_def wf_J\<^sub>1_mdecl_def) -apply(clarify) apply(rename_tac C M b Ts T x1 x2 pns body) -apply(case_tac b) - apply clarsimp - apply(frule WT_fv) - apply(auto intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1 \)[1] -apply clarsimp -apply(frule WT_fv) -apply(fastforce intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1' \) -done +proof - + assume wf: "wf_J_prog P" + let ?f = "(\b (pns, body). + compE\<^sub>1 (case b of Static \ pns | NonStatic \ this # pns) body)" + let ?wf\<^sub>2 = "wf_J\<^sub>1_mdecl" + + { fix C M b Ts T m + assume cM: "P \ C sees M, b : Ts\T = m in C" + and wfm: "wf_mdecl wf_J_mdecl P C (M, b, Ts, T, m)" + obtain pns body where [simp]: "m = (pns, body)" by(cases m) simp + let ?E = "\b. case b of Static \ [pns [\] Ts] | NonStatic \ [pns [\] Ts, this \ Class C]" + obtain T' where WT: "P,?E b \ body :: T'" and subT: "P \ T' \ T" + using wfm by(cases b) (auto simp: wf_mdecl_def) + have fv: "fv body \ dom (?E b)" by(rule WT_fv[OF WT]) + have "wf_mdecl ?wf\<^sub>2 (compP ?f P) C (M, b, Ts, T, ?f b m)" + proof(cases b) + case Static then show ?thesis using cM wfm fv + by(auto simp:wf_mdecl_def wf_J\<^sub>1_mdecl_def + intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1 \) + next + case NonStatic then show ?thesis using cM wfm fv + by(clarsimp simp add:wf_mdecl_def wf_J\<^sub>1_mdecl_def) + (fastforce intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1' \) + qed + } + then show ?thesis by simp (rule wf_prog_compPI[OF _ wf]) +qed (*>*) end diff --git a/thys/JinjaDCI/Compiler/Correctness2.thy b/thys/JinjaDCI/Compiler/Correctness2.thy --- a/thys/JinjaDCI/Compiler/Correctness2.thy +++ b/thys/JinjaDCI/Compiler/Correctness2.thy @@ -1,3202 +1,3314 @@ (* Title: JinjaDCI/Compiler/Correctness2.thy Author: Tobias Nipkow, Susannah Mansky Copyright TUM 2003, UIUC 2019-20 Based on the Jinja theory Compiler/Correctness2.thy by Tobias Nipkow *) section \ Correctness of Stage 2 \ theory Correctness2 imports "HOL-Library.Sublist" Compiler2 J1WellForm "../J/EConform" begin (*<*)hide_const (open) Throw(*>*) subsection\ Instruction sequences \ text\ How to select individual instructions and subsequences of instructions from a program given the class, method and program counter. \ definition before :: "jvm_prog \ cname \ mname \ nat \ instr list \ bool" ("(_,_,_,_/ \ _)" [51,0,0,0,51] 50) where "P,C,M,pc \ is \ prefix is (drop pc (instrs_of P C M))" definition at :: "jvm_prog \ cname \ mname \ nat \ instr \ bool" ("(_,_,_,_/ \ _)" [51,0,0,0,51] 50) where "P,C,M,pc \ i \ (\is. drop pc (instrs_of P C M) = i#is)" lemma [simp]: "P,C,M,pc \ []" (*<*)by(simp add:before_def)(*>*) lemma [simp]: "P,C,M,pc \ (i#is) = (P,C,M,pc \ i \ P,C,M,pc + 1 \ is)" (*<*)by(fastforce simp add:before_def at_def prefix_def drop_Suc drop_tl)(*>*) (*<*) declare drop_drop[simp del] (*>*) lemma [simp]: "P,C,M,pc \ (is\<^sub>1 @ is\<^sub>2) = (P,C,M,pc \ is\<^sub>1 \ P,C,M,pc + size is\<^sub>1 \ is\<^sub>2)" (*<*) -apply(simp add:before_def prefix_def) -apply(subst add.commute) -apply(simp add: drop_drop[symmetric]) -apply fastforce -done +by(subst add.commute) + (fastforce simp add:before_def prefix_def drop_drop[symmetric]) (*>*) (*<*) declare drop_drop[simp] (*>*) lemma [simp]: "P,C,M,pc \ i \ instrs_of P C M ! pc = i" (*<*)by(clarsimp simp add:at_def strict_prefix_def nth_via_drop)(*>*) lemma beforeM: "P \ C sees M,b: Ts\T = body in D \ compP\<^sub>2 P,D,M,0 \ compE\<^sub>2 body @ [Return]" -(*<*) -apply(drule sees_method_idemp) -apply(simp add:before_def compP\<^sub>2_def compMb\<^sub>2_def) -done -(*>*) +(*<*)by(drule sees_method_idemp) (simp add:before_def compMb\<^sub>2_def)(*>*) text\ This lemma executes a single instruction by rewriting: \ lemma [simp]: "P,C,M,pc \ instr \ (P \ (None, h, (vs,ls,C,M,pc,ics) # frs, sh) -jvm\ \') = ((None, h, (vs,ls,C,M,pc,ics) # frs, sh) = \' \ (\\. exec(P,(None, h, (vs,ls,C,M,pc,ics) # frs, sh)) = Some \ \ P \ \ -jvm\ \'))" (*<*) -apply(simp only: exec_all_def) -apply(blast intro: converse_rtranclE converse_rtrancl_into_rtrancl) -done +by(simp only: exec_all_def) + (blast intro: converse_rtranclE converse_rtrancl_into_rtrancl) (*>*) subsection\ Exception tables \ definition pcs :: "ex_table \ nat set" where "pcs xt \ \(f,t,C,h,d) \ set xt. {f ..< t}" lemma pcs_subset: shows "(\pc d. pcs(compxE\<^sub>2 e pc d) \ {pc..2 e)})" and "(\pc d. pcs(compxEs\<^sub>2 es pc d) \ {pc..2 es)})" (*<*) -apply(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) -apply (simp_all add:pcs_def) -apply (fastforce split:bop.splits)+ -done +proof(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) + case Cast then show ?case by (fastforce simp:pcs_def) +next + case BinOp then show ?case by (fastforce simp:pcs_def split:bop.splits) +next + case LAss then show ?case by (fastforce simp: pcs_def) +next + case FAcc then show ?case by (fastforce simp: pcs_def) +next + case FAss then show ?case by (fastforce simp: pcs_def) +next + case SFAss then show ?case by (fastforce simp: pcs_def) +next + case Call then show ?case by (fastforce simp: pcs_def) +next + case SCall then show ?case by (fastforce simp: pcs_def) +next + case Seq then show ?case by (fastforce simp: pcs_def) +next + case Cond then show ?case by (fastforce simp: pcs_def) +next + case While then show ?case by (fastforce simp: pcs_def) +next + case throw then show ?case by (fastforce simp: pcs_def) +next + case TryCatch then show ?case by (fastforce simp: pcs_def) +next + case Cons_exp then show ?case by (fastforce simp: pcs_def) +qed (simp_all add:pcs_def) (*>*) lemma [simp]: "pcs [] = {}" (*<*)by(simp add:pcs_def)(*>*) lemma [simp]: "pcs (x#xt) = {fst x ..< fst(snd x)} \ pcs xt" (*<*)by(auto simp add: pcs_def)(*>*) lemma [simp]: "pcs(xt\<^sub>1 @ xt\<^sub>2) = pcs xt\<^sub>1 \ pcs xt\<^sub>2" (*<*)by(simp add:pcs_def)(*>*) lemma [simp]: "pc < pc\<^sub>0 \ pc\<^sub>0+size(compE\<^sub>2 e) \ pc \ pc \ pcs(compxE\<^sub>2 e pc\<^sub>0 d)" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc < pc\<^sub>0 \ pc\<^sub>0+size(compEs\<^sub>2 es) \ pc \ pc \ pcs(compxEs\<^sub>2 es pc\<^sub>0 d)" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc\<^sub>1 + size(compE\<^sub>2 e\<^sub>1) \ pc\<^sub>2 \ pcs(compxE\<^sub>2 e\<^sub>1 pc\<^sub>1 d\<^sub>1) \ pcs(compxE\<^sub>2 e\<^sub>2 pc\<^sub>2 d\<^sub>2) = {}" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc\<^sub>1 + size(compE\<^sub>2 e) \ pc\<^sub>2 \ pcs(compxE\<^sub>2 e pc\<^sub>1 d\<^sub>1) \ pcs(compxEs\<^sub>2 es pc\<^sub>2 d\<^sub>2) = {}" (*<*)using pcs_subset by fastforce(*>*) lemma [simp]: "pc \ pcs xt\<^sub>0 \ match_ex_table P C pc (xt\<^sub>0 @ xt\<^sub>1) = match_ex_table P C pc xt\<^sub>1" (*<*)by (induct xt\<^sub>0) (auto simp: matches_ex_entry_def)(*>*) lemma [simp]: "\ x \ set xt; pc \ pcs xt \ \ \ matches_ex_entry P D pc x" (*<*)by(auto simp:matches_ex_entry_def pcs_def)(*>*) lemma [simp]: assumes xe: "xe \ set(compxE\<^sub>2 e pc d)" and outside: "pc' < pc \ pc+size(compE\<^sub>2 e) \ pc'" shows "\ matches_ex_entry P C pc' xe" (*<*) proof assume "matches_ex_entry P C pc' xe" with xe have "pc' \ pcs(compxE\<^sub>2 e pc d)" by(force simp add:matches_ex_entry_def pcs_def) with outside show False by simp qed (*>*) lemma [simp]: assumes xe: "xe \ set(compxEs\<^sub>2 es pc d)" and outside: "pc' < pc \ pc+size(compEs\<^sub>2 es) \ pc'" shows "\ matches_ex_entry P C pc' xe" (*<*) proof assume "matches_ex_entry P C pc' xe" with xe have "pc' \ pcs(compxEs\<^sub>2 es pc d)" by(force simp add:matches_ex_entry_def pcs_def) with outside show False by simp qed (*>*) lemma match_ex_table_app[simp]: "\xte \ set xt\<^sub>1. \ matches_ex_entry P D pc xte \ match_ex_table P D pc (xt\<^sub>1 @ xt) = match_ex_table P D pc xt" (*<*)by(induct xt\<^sub>1) simp_all(*>*) lemma [simp]: "\x \ set xtab. \ matches_ex_entry P C pc x \ match_ex_table P C pc xtab = None" (*<*)using match_ex_table_app[where ?xt = "[]"] by fastforce(*>*) lemma match_ex_entry: "matches_ex_entry P C pc (start, end, catch_type, handler) = (start \ pc \ pc < end \ P \ C \\<^sup>* catch_type)" (*<*)by(simp add:matches_ex_entry_def)(*>*) definition caught :: "jvm_prog \ pc \ heap \ addr \ ex_table \ bool" where "caught P pc h a xt \ (\entry \ set xt. matches_ex_entry P (cname_of h a) pc entry)" definition beforex :: "jvm_prog \ cname \ mname \ ex_table \ nat set \ nat \ bool" ("(2_,/_,/_ \/ _ /'/ _,/_)" [51,0,0,0,0,51] 50) where "P,C,M \ xt / I,d \ (\xt\<^sub>0 xt\<^sub>1. ex_table_of P C M = xt\<^sub>0 @ xt @ xt\<^sub>1 \ pcs xt\<^sub>0 \ I = {} \ pcs xt \ I \ (\pc \ I. \C pc' d'. match_ex_table P C pc xt\<^sub>1 = \(pc',d')\ \ d' \ d))" definition dummyx :: "jvm_prog \ cname \ mname \ ex_table \ nat set \ nat \ bool" ("(2_,_,_ \/ _ '/_,_)" [51,0,0,0,0,51] 50) where "P,C,M \ xt/I,d \ P,C,M \ xt/I,d" +abbreviation +"beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1 + \ ex_table_of P C M = xt\<^sub>0 @ xt @ xt\<^sub>1 \ pcs xt\<^sub>0 \ I = {} + \ pcs xt \ I \ (\pc \ I. \C pc' d'. match_ex_table P C pc xt\<^sub>1 = \(pc',d')\ \ d' \ d)" + +lemma beforex_beforex\<^sub>0_eq: + "P,C,M \ xt / I,d \ \xt\<^sub>0 xt\<^sub>1. beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1" +using beforex_def by auto + lemma beforexD1: "P,C,M \ xt / I,d \ pcs xt \ I" (*<*)by(auto simp add:beforex_def)(*>*) lemma beforex_mono: "\ P,C,M \ xt/I,d'; d' \ d \ \ P,C,M \ xt/I,d" (*<*)by(fastforce simp:beforex_def)(*>*) lemma [simp]: "P,C,M \ xt/I,d \ P,C,M \ xt/I,Suc d" (*<*)by(fastforce intro:beforex_mono)(*>*) lemma beforex_append[simp]: "pcs xt\<^sub>1 \ pcs xt\<^sub>2 = {} \ P,C,M \ xt\<^sub>1 @ xt\<^sub>2/I,d = (P,C,M \ xt\<^sub>1/I-pcs xt\<^sub>2,d \ P,C,M \ xt\<^sub>2/I-pcs xt\<^sub>1,d \ P,C,M \ xt\<^sub>1@xt\<^sub>2/I,d)" -(*<*) -apply(rule iffI) - prefer 2 - apply(simp add:dummyx_def) -apply(auto simp add: beforex_def dummyx_def) - apply(rule_tac x = xt\<^sub>0 in exI) - apply auto -apply(rule_tac x = "xt\<^sub>0@xt\<^sub>1" in exI) -apply auto -done +(*<*)(is "?Q \ ?P = (?P1 \ ?P2 \ ?P3)" is "?Q \ ?P = ?P123") +proof - + assume pcs: ?Q + show ?thesis proof(rule iffI) + assume "?P123" then show ?P by(simp add:dummyx_def) + next + assume hyp: ?P + let ?xt = "xt\<^sub>1 @ xt\<^sub>2" + let ?beforex = "beforex\<^sub>0 P C M d" + obtain xt\<^sub>0 xt\<^sub>1' where beforex: "?beforex I ?xt xt\<^sub>0 xt\<^sub>1'" + using hyp by(clarsimp simp: beforex_def) + have "\xt\<^sub>0 xt\<^sub>1'. ?beforex (I - pcs xt\<^sub>2) xt\<^sub>1 xt\<^sub>0 xt\<^sub>1'" \ \?P1\ + using pcs beforex by(rule_tac x=xt\<^sub>0 in exI) auto + moreover have "\xt\<^sub>0 xt\<^sub>1'. ?beforex (I - pcs xt\<^sub>1) xt\<^sub>2 xt\<^sub>0 xt\<^sub>1'" \ \?P2\ + using pcs beforex by(rule_tac x="xt\<^sub>0@xt\<^sub>1" in exI) auto + moreover have ?P3 using hyp by(simp add: dummyx_def) + ultimately show ?P123 by (simp add: beforex_def) + qed +qed (*>*) lemma beforex_appendD1: - "\ P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d; - pcs xt\<^sub>1 \ J; J \ I; J \ pcs xt\<^sub>2 = {} \ - \ P,C,M \ xt\<^sub>1 / J,d" +assumes bx: "P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d" + and pcs: "pcs xt\<^sub>1 \ J" and JI: "J \ I" and Jpcs: "J \ pcs xt\<^sub>2 = {}" +shows "P,C,M \ xt\<^sub>1 / J,d" (*<*) -apply(auto simp:beforex_def) -apply(rule exI,rule exI,rule conjI, rule refl) -apply(rule conjI, blast) -apply(auto) -apply(subgoal_tac "pc \ pcs xt\<^sub>2") - prefer 2 apply blast -apply (auto split:if_split_asm) -done +proof - + let ?beforex = "beforex\<^sub>0 P C M d" + obtain xt\<^sub>0 xt\<^sub>1' where bx': "?beforex I (xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)]) xt\<^sub>0 xt\<^sub>1'" + using bx by(clarsimp simp:beforex_def) + let ?xt0 = xt\<^sub>0 and ?xt1 = "xt\<^sub>2 @ (f, t, D, h, d) # xt\<^sub>1'" + have "pcs xt\<^sub>0 \ J = {}" using bx' JI by blast + moreover { + fix pc C pc' d' assume pcJ: "pc\J" + then have "pc \ pcs xt\<^sub>2" using bx' Jpcs by blast + then have "match_ex_table P C pc (xt\<^sub>2 @ (f, t, D, h, d) # xt\<^sub>1') + = \(pc', d')\ \ d' \ d" + using bx' JI pcJ by (auto split:if_split_asm) + } + ultimately have "?beforex J xt\<^sub>1 ?xt0 ?xt1" using bx' pcs by simp + then show ?thesis using beforex_def by blast +qed (*>*) lemma beforex_appendD2: - "\ P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d; - pcs xt\<^sub>2 \ J; J \ I; J \ pcs xt\<^sub>1 = {} \ - \ P,C,M \ xt\<^sub>2 / J,d" +assumes bx: "P,C,M \ xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d" + and pcs: "pcs xt\<^sub>2 \ J" and JI: "J \ I" and Jpcs: "J \ pcs xt\<^sub>1 = {}" +shows "P,C,M \ xt\<^sub>2 / J,d" (*<*) -apply(auto simp:beforex_def) -apply(rule_tac x = "xt\<^sub>0 @ xt\<^sub>1" in exI) -apply fastforce -done +proof - + let ?beforex = "beforex\<^sub>0 P C M d" + obtain xt\<^sub>0 xt\<^sub>1' where bx': "?beforex I (xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)]) xt\<^sub>0 xt\<^sub>1'" + using bx by(clarsimp simp:beforex_def) + then have "\xt\<^sub>1''. beforex\<^sub>0 P C M d J xt\<^sub>2 (xt\<^sub>0 @ xt\<^sub>1) xt\<^sub>1''" + using assms by fastforce + then show ?thesis using beforex_def by blast +qed (*>*) lemma beforexM: "P \ C sees M,b: Ts\T = body in D \ compP\<^sub>2 P,D,M \ compxE\<^sub>2 body 0 0/{..2 body)},0" (*<*) -apply(drule sees_method_idemp) -apply(drule sees_method_compP[where f = compMb\<^sub>2]) -apply(simp add:beforex_def compP\<^sub>2_def compMb\<^sub>2_def) -apply(rule_tac x = "[]" in exI) -using pcs_subset apply fastforce -done +proof - + assume cM: "P \ C sees M,b: Ts\T = body in D" + let ?xt0 = "[]" + have "\xt1. beforex\<^sub>0 (compP\<^sub>2 P) D M 0 ({..2 body)}) (compxE\<^sub>2 body 0 0) ?xt0 xt1" + using sees_method_compP[where f = compMb\<^sub>2, OF sees_method_idemp[OF cM]] + pcs_subset by(fastforce simp add: compP\<^sub>2_def compMb\<^sub>2_def) + then show ?thesis using beforex_def by blast +qed (*>*) lemma match_ex_table_SomeD2: - "\ match_ex_table P D pc (ex_table_of P C M) = \(pc',d')\; - P,C,M \ xt/I,d; \x \ set xt. \ matches_ex_entry P D pc x; pc \ I \ - \ d' \ d" +assumes met: "match_ex_table P D pc (ex_table_of P C M) = \(pc',d')\" + and bx: "P,C,M \ xt/I,d" + and nmet: "\x \ set xt. \ matches_ex_entry P D pc x" and pcI: "pc \ I" +shows "d' \ d" (*<*) -apply(auto simp:beforex_def) -apply(subgoal_tac "pc \ pcs xt\<^sub>0") -apply auto -done +proof - + obtain xt\<^sub>0 xt\<^sub>1 where bx': "beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1" + using bx by(clarsimp simp:beforex_def) + then have "pc \ pcs xt\<^sub>0" using pcI by blast + then show ?thesis using bx' met nmet pcI by simp +qed (*>*) lemma match_ex_table_SomeD1: "\ match_ex_table P D pc (ex_table_of P C M) = \(pc',d')\; P,C,M \ xt / I,d; pc \ I; pc \ pcs xt \ \ d' \ d" (*<*)by(auto elim: match_ex_table_SomeD2)(*>*) subsection\ The correctness proof \ (*<*) declare nat_add_distrib[simp] caught_def[simp] declare fun_upd_apply[simp del] (*>*) definition handle :: "jvm_prog \ cname \ mname \ addr \ heap \ val list \ val list \ nat \ init_call_status \ frame list \ sheap \ jvm_state" where "handle P C M a h vs ls pc ics frs sh = find_handler P a h ((vs,ls,C,M,pc,ics) # frs) sh" lemma aux_isin[simp]: "\ B \ A; a \ B \ \ a \ A" (*<*)by blast(*>*) lemma handle_frs_tl_neq: "ics_of f \ No_ics \ (xp, h, f#frs, sh) \ handle P C M xa h' vs l pc ics frs sh'" by(simp add: handle_def find_handler_frs_tl_neq del: find_handler.simps) subsubsection "Correctness proof inductive hypothesis" \ \ frame definitions for use by correctness proof inductive hypothesis \ fun calling_to_called :: "frame \ frame" where "calling_to_called (stk,loc,D,M,pc,ics) = (stk,loc,D,M,pc,case ics of Calling C Cs \ Called (C#Cs))" fun calling_to_scalled :: "frame \ frame" where "calling_to_scalled (stk,loc,D,M,pc,ics) = (stk,loc,D,M,pc,case ics of Calling C Cs \ Called Cs)" fun calling_to_calling :: "frame \ cname \ frame" where "calling_to_calling (stk,loc,D,M,pc,ics) C' = (stk,loc,D,M,pc,case ics of Calling C Cs \ Calling C' (C#Cs))" fun calling_to_throwing :: "frame \ addr \ frame" where "calling_to_throwing (stk,loc,D,M,pc,ics) a = (stk,loc,D,M,pc,case ics of Calling C Cs \ Throwing (C#Cs) a)" fun calling_to_sthrowing :: "frame \ addr \ frame" where "calling_to_sthrowing (stk,loc,D,M,pc,ics) a = (stk,loc,D,M,pc,case ics of Calling C Cs \ Throwing Cs a)" \ \ pieces of the correctness proof's inductive hypothesis, which depend on the expression being compiled) \ fun Jcc_cond :: "J\<^sub>1_prog \ ty list \ cname \ mname \ val list \ pc \ init_call_status \ nat set \ heap \ sheap \ expr\<^sub>1 \ bool" where "Jcc_cond P E C M vs pc ics I h sh (INIT C\<^sub>0 (Cs,b) \ e') = ((\T. P,E,h,sh \\<^sub>1 INIT C\<^sub>0 (Cs,b) \ e' : T) \ unit = e' \ ics = No_ics)" | "Jcc_cond P E C M vs pc ics I h sh (RI(C',e\<^sub>0);Cs \ e') = (((e\<^sub>0 = C'\\<^sub>sclinit([]) \ (\T. P,E,h,sh \\<^sub>1 RI(C',e\<^sub>0);Cs \ e':T)) \ ((\a. e\<^sub>0 = Throw a) \ (\C \ set(C'#Cs). is_class P C))) \ unit = e' \ ics = No_ics)" | "Jcc_cond P E C M vs pc ics I h sh (C'\\<^sub>sM'(es)) = (let e = (C'\\<^sub>sM'(es)) in if M' = clinit \ es = [] then (\T. P,E,h,sh \\<^sub>1 e:T) \ (\Cs. ics = Called Cs) else (compP\<^sub>2 P,C,M,pc \ compE\<^sub>2 e \ compP\<^sub>2 P,C,M \ compxE\<^sub>2 e pc (size vs)/I,size vs \ {pc..2 e)} \ I \ \sub_RI e \ ics = No_ics) )" | "Jcc_cond P E C M vs pc ics I h sh e = (compP\<^sub>2 P,C,M,pc \ compE\<^sub>2 e \ compP\<^sub>2 P,C,M \ compxE\<^sub>2 e pc (size vs)/I,size vs \ {pc..2 e)} \ I \ \sub_RI e \ ics = No_ics)" fun Jcc_frames :: "jvm_prog \ cname \ mname \ val list \ val list \ pc \ init_call_status \ frame list \ expr\<^sub>1 \ frame list" where "Jcc_frames P C M vs ls pc ics frs (INIT C\<^sub>0 (C'#Cs,b) \ e') = (case b of False \ (vs,ls,C,M,pc,Calling C' Cs) # frs | True \ (vs,ls,C,M,pc,Called (C'#Cs)) # frs )" | "Jcc_frames P C M vs ls pc ics frs (INIT C\<^sub>0 (Nil,b) \ e') = (vs,ls,C,M,pc,Called [])#frs" | "Jcc_frames P C M vs ls pc ics frs (RI(C',e\<^sub>0);Cs \ e') = (case e\<^sub>0 of Throw a \ (vs,ls,C,M,pc,Throwing (C'#Cs) a) # frs | _ \ (vs,ls,C,M,pc,Called (C'#Cs)) # frs )" | "Jcc_frames P C M vs ls pc ics frs (C'\\<^sub>sM'(es)) = (if M' = clinit \ es = [] then create_init_frame P C'#(vs,ls,C,M,pc,ics)#frs else (vs,ls,C,M,pc,ics)#frs )" | "Jcc_frames P C M vs ls pc ics frs e = (vs,ls,C,M,pc,ics)#frs" fun Jcc_rhs :: "J\<^sub>1_prog \ ty list \ cname \ mname \ val list \ val list \ pc \ init_call_status \ frame list \ heap \ val list \ sheap \ val \ expr\<^sub>1 \ jvm_state" where "Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v (INIT C\<^sub>0 (Cs,b) \ e') = (None,h',(vs,ls,C,M,pc,Called [])#frs,sh')" | "Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v (RI(C',e\<^sub>0);Cs \ e') = (None,h',(vs,ls,C,M,pc,Called [])#frs,sh')" | "Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v (C'\\<^sub>sM'(es)) = (let e = (C'\\<^sub>sM'(es)) in if M' = clinit \ es = [] then (None,h',(vs,ls,C,M,pc,ics)#frs,sh'(C'\(fst(the(sh' C')),Done))) else (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 e),ics)#frs,sh') )" | "Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v e = (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 e),ics)#frs,sh')" fun Jcc_err :: "jvm_prog \ cname \ mname \ heap \ val list \ val list \ pc \ init_call_status \ frame list \ sheap \ nat set \ heap \ val list \ sheap \ addr \ expr\<^sub>1 \ bool" where "Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa (INIT C\<^sub>0 (Cs,b) \ e') = (\vs'. P \ (None,h,Jcc_frames P C M vs ls pc ics frs (INIT C\<^sub>0 (Cs,b) \ e'),sh) -jvm\ handle P C M xa h' (vs'@vs) ls pc ics frs sh')" | "Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa (RI(C',e\<^sub>0);Cs \ e') = (\vs'. P \ (None,h,Jcc_frames P C M vs ls pc ics frs (RI(C',e\<^sub>0);Cs \ e'),sh) -jvm\ handle P C M xa h' (vs'@vs) ls pc ics frs sh')" | "Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa (C'\\<^sub>sM'(es)) = (let e = (C'\\<^sub>sM'(es)) in if M' = clinit \ es = [] then case ics of Called Cs \ P \ (None,h,Jcc_frames P C M vs ls pc ics frs e,sh) -jvm\ (None,h',(vs,ls,C,M,pc,Throwing Cs xa)#frs,(sh'(C' \ (fst(the(sh' C')),Error)))) else (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compE\<^sub>2 e) \ \ caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (size vs)) \ (\vs'. P \ (None,h,Jcc_frames P C M vs ls pc ics frs e,sh) -jvm\ handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh')) )" | "Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa e = (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compE\<^sub>2 e) \ \ caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (size vs)) \ (\vs'. P \ (None,h,Jcc_frames P C M vs ls pc ics frs e,sh) -jvm\ handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh'))" fun Jcc_pieces :: "J\<^sub>1_prog \ ty list \ cname \ mname \ heap \ val list \ val list \ pc \ init_call_status \ frame list \ sheap \ nat set \ heap \ val list \ sheap \ val \ addr \ expr\<^sub>1 \ bool \ frame list \ jvm_state \ bool" where "Jcc_pieces P E C M h vs ls pc ics frs sh I h' ls' sh' v xa e = (Jcc_cond P E C M vs pc ics I h sh e, Jcc_frames (compP\<^sub>2 P) C M vs ls pc ics frs e, Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v e, Jcc_err (compP\<^sub>2 P) C M h vs ls pc ics frs sh I h' ls' sh' xa e)" \ \ @{text Jcc_pieces} lemmas \ lemma nsub_RI_Jcc_pieces: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and nsub: "\sub_RI e" shows "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa e = (let cond = P,C,M,pc \ compE\<^sub>2 e \ P,C,M \ compxE\<^sub>2 e pc (size vs)/I,size vs \ {pc..2 e)} \ I \ ics = No_ics; frs' = (vs,ls,C,M,pc,ics)#frs; rhs = (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 e),ics)#frs,sh'); err = (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compE\<^sub>2 e) \ \ caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (size vs)) \ (\vs'. P \ (None,h,frs',sh) -jvm\ handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh')) in (cond, frs',rhs, err) )" proof - have NC: "\C'. e \ C'\\<^sub>sclinit([])" using assms(2) proof(cases e) qed(simp_all) then show ?thesis using assms proof(cases e) case (SCall C M es) then have "M \ clinit" using nsub by simp then show ?thesis using SCall nsub proof(cases es) qed(simp_all) qed(simp_all) qed lemma Jcc_pieces_Cast: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e) = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e = (True, frs\<^sub>0, (xp',h',(v'#vs',ls',C\<^sub>0,M',pc' - 1,ics')#frs',sh'), (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compE\<^sub>2 e) \ \ caught P pc\<^sub>1 h\<^sub>1 xa (compxE\<^sub>2 e pc (size vs)) \ (\vs'. P \ (None,h\<^sub>0,frs\<^sub>0,sh\<^sub>0) -jvm\ handle P C M xa h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)))" proof - have pc: "{pc..2 e)} \ I" using assms by clarsimp show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] pc by clarsimp qed lemma Jcc_pieces_BinOp1: assumes "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e \bop\ e') = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "\err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 (I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs')))) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e = (True, frs\<^sub>0, (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C\<^sub>0,M',pc' - size (compE\<^sub>2 e') - 1,ics')#frs',sh\<^sub>1), err)" proof - have bef: "compP compMb\<^sub>2 P,C\<^sub>0,M' \ compxE\<^sub>2 e pc (length vs) / I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs'))),length vs" using assms by clarsimp have vs: "vs = vs'" using assms by simp show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] bef vs by clarsimp qed lemma Jcc_pieces_BinOp2: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e \bop\ e') = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "\err. Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (v\<^sub>1#vs) ls\<^sub>1 (pc + size (compE\<^sub>2 e)) ics frs sh\<^sub>1 (I - pcs (compxE\<^sub>2 e pc (length vs'))) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v' xa e' = (True, (v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs, (xp',h',(v'#v\<^sub>1#vs',ls',C\<^sub>0,M',pc' - 1,ics')#frs',sh'), (\pc\<^sub>1. pc + size (compE\<^sub>2 e) \ pc\<^sub>1 \ pc\<^sub>1 < pc + size (compE\<^sub>2 e) + length (compE\<^sub>2 e') \ \ caught P pc\<^sub>1 h\<^sub>2 xa (compxE\<^sub>2 e' (pc + size (compE\<^sub>2 e)) (Suc (length vs))) \ (\vs'. P \ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs,sh\<^sub>1) -jvm\ handle P C M xa h\<^sub>2 (vs'@v\<^sub>1#vs) ls\<^sub>2 pc\<^sub>1 ics frs sh\<^sub>2)))" proof - have bef: "compP compMb\<^sub>2 P\<^sub>1,C\<^sub>0,M' \ compxE\<^sub>2 e pc (length vs) / I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs'))),length vs" using assms by clarsimp have vs: "vs = vs'" using assms by simp show ?thesis using assms nsub_RI_Jcc_pieces[where e=e'] bef vs by clarsimp qed lemma Jcc_pieces_FAcc: assumes "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\F{D}) = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "\err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e = (True, frs\<^sub>0, (xp',h',(v'#vs',ls',C\<^sub>0,M',pc' - 1,ics')#frs',sh'), err)" proof - have pc: "{pc..2 e)} \ I" using assms by clarsimp then show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by clarsimp qed lemma Jcc_pieces_LAss: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (i:=e) = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e = (True, frs\<^sub>0, (xp',h',(v'#vs',ls',C\<^sub>0,M',pc' - 2,ics')#frs',sh'), (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compE\<^sub>2 e) \ \ caught P pc\<^sub>1 h\<^sub>1 xa (compxE\<^sub>2 e pc (size vs)) \ (\vs'. P \ (None,h\<^sub>0,frs\<^sub>0,sh\<^sub>0) -jvm\ handle P C M xa h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)))" proof - have pc: "{pc..2 e)} \ I" using assms by clarsimp show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] pc by clarsimp qed lemma Jcc_pieces_FAss1: assumes "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\F{D}:=e') = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "\err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 (I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs')))) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e = (True, frs\<^sub>0, (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C\<^sub>0,M',pc' - size (compE\<^sub>2 e') - 2,ics')#frs',sh\<^sub>1), err)" proof - show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by clarsimp qed lemma Jcc_pieces_FAss2: assumes "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\F{D}:=e') = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "Jcc_pieces P E C M h\<^sub>1 (v\<^sub>1#vs) ls\<^sub>1 (pc + size (compE\<^sub>2 e)) ics frs sh\<^sub>1 (I - pcs (compxE\<^sub>2 e pc (length vs'))) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v' xa e' = (True, (v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs, (xp',h',(v'#v\<^sub>1#vs',ls',C\<^sub>0,M',pc' - 2,ics')#frs',sh'), (\pc\<^sub>1. (pc + size (compE\<^sub>2 e)) \ pc\<^sub>1 \ pc\<^sub>1 < pc + size (compE\<^sub>2 e) + size(compE\<^sub>2 e') \ \ caught (compP\<^sub>2 P) pc\<^sub>1 h\<^sub>2 xa (compxE\<^sub>2 e' (pc + size (compE\<^sub>2 e)) (size (v\<^sub>1#vs))) \ (\vs'. (compP\<^sub>2 P) \ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs,sh\<^sub>1) -jvm\ handle (compP\<^sub>2 P) C M xa h\<^sub>2 (vs'@v\<^sub>1#vs) ls\<^sub>2 pc\<^sub>1 ics frs sh\<^sub>2)))" proof - show ?thesis using assms nsub_RI_Jcc_pieces[where e=e'] by clarsimp qed lemma Jcc_pieces_SFAss: assumes "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h' ls' sh' v xa (C'\\<^sub>sF{D}:=e) = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)" shows "\err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e = (True, frs\<^sub>0, (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C\<^sub>0,M',pc' - 2,ics')#frs',sh\<^sub>1), err)" proof - have pc: "{pc..2 e)} \ I" using assms by clarsimp show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] pc by clarsimp qed lemma Jcc_pieces_Call1: assumes "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>3 ls\<^sub>3 sh\<^sub>3 v xa (e\M\<^sub>0(es)) = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C',M',pc',ics')#frs',sh'), err)" shows "\err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs')))) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e = (True, frs\<^sub>0, (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C',M',pc' - size (compEs\<^sub>2 es) - 1,ics')#frs',sh\<^sub>1), err)" proof - show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by clarsimp qed lemma Jcc_pieces_clinit: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and cond: "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (C1\\<^sub>sclinit([]))" shows "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C1\\<^sub>sclinit([])) = (True, create_init_frame P C1 # (vs,ls,C,M,pc,ics)#frs, (None, h', (vs,ls,C,M,pc,ics)#frs, sh'(C1\(fst(the(sh' C1)),Done))), P \ (None,h,create_init_frame P C1 # (vs,ls,C,M,pc,ics)#frs,sh) -jvm\ (case ics of Called Cs \ (None,h',(vs,ls,C,M,pc,Throwing Cs xa)#frs,(sh'(C1 \ (fst(the(sh' C1)),Error))))))" using assms by(auto split: init_call_status.splits list.splits bool.splits) lemma Jcc_pieces_SCall_clinit_body: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa (C1\\<^sub>sclinit([])) = (True, frs', rhs', err')" and method: "P\<^sub>1 \ C1 sees clinit,Static: []\Void = body in D" shows "Jcc_pieces P\<^sub>1 [] D clinit h\<^sub>2 [] (replicate (max_vars body) undefined) 0 No_ics (tl frs') sh\<^sub>2 {..2 body)} h\<^sub>3 ls\<^sub>3 sh\<^sub>3 v xa body = (True, frs', (None,h\<^sub>3,([v],ls\<^sub>3,D,clinit,size(compE\<^sub>2 body), No_ics)#tl frs',sh\<^sub>3), \pc\<^sub>1. 0 \ pc\<^sub>1 \ pc\<^sub>1 < size(compE\<^sub>2 body) \ \ caught P pc\<^sub>1 h\<^sub>3 xa (compxE\<^sub>2 body 0 0) \ (\vs'. P \ (None,h\<^sub>2,frs',sh\<^sub>2) -jvm\ handle P D clinit xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>1 No_ics (tl frs') sh\<^sub>3))" proof - have M_in_D: "P\<^sub>1 \ D sees clinit,Static: []\Void = body in D" using method by(rule sees_method_idemp) hence M_code: "compP\<^sub>2 P\<^sub>1,D,clinit,0 \ compE\<^sub>2 body @ [Return]" and M_xtab: "compP\<^sub>2 P\<^sub>1,D,clinit \ compxE\<^sub>2 body 0 0/{..2 body)},0" by(rule beforeM, rule beforexM) have nsub: "\sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf method]) then show ?thesis using assms nsub_RI_Jcc_pieces M_code M_xtab by clarsimp qed lemma Jcc_pieces_Cons: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "P,C,M,pc \ compEs\<^sub>2 (e#es)" and "P,C,M \ compxEs\<^sub>2 (e#es) pc (size vs)/I,size vs" and "{pc..2 (e#es))} \ I" and "ics = No_ics" and "\sub_RIs (e#es)" shows "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs)))) h' ls' sh' v xa e = (True, (vs, ls, C, M, pc, ics) # frs, (None, h', (v#vs, ls', C, M, pc + length (compE\<^sub>2 e), ics) # frs, sh'), \pc\<^sub>1\pc. pc\<^sub>1 < pc + length (compE\<^sub>2 e) \ \ caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (length vs)) \ (\vs'. P \ (None, h, (vs, ls, C, M, pc, ics) # frs, sh) -jvm\ handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh'))" proof - show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by auto qed lemma Jcc_pieces_InitNone: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" shows "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs (sh(C\<^sub>0 \ (sblank P C\<^sub>0, Prepared))) I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), \vs'. P \ (None,h,frs',(sh(C\<^sub>0 \ (sblank P\<^sub>1 C\<^sub>0, Prepared)))) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh')" proof - have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \ e)" using assms by simp then obtain T where "P\<^sub>1,E,h,sh \\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \ unit : T" by fastforce then have "P\<^sub>1,E,h,sh(C\<^sub>0 \ (sblank P\<^sub>1 C\<^sub>0, Prepared)) \\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \ unit : T" by(auto simp: fun_upd_apply) then have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h (sh(C\<^sub>0 \ (sblank P\<^sub>1 C\<^sub>0, Prepared))) (INIT C' (C\<^sub>0 # Cs,False) \ unit))" by(simp only: exI) then show ?thesis using assms by clarsimp qed lemma Jcc_pieces_InitDP: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" shows "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (Cs,True) \ e) = (True, (calling_to_scalled (hd frs'))#(tl frs'), (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), \vs'. P \ (None,h,calling_to_scalled (hd frs')#(tl frs'),sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh')" proof - have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \ e)" using assms by simp then obtain T where "P\<^sub>1,E,h,sh \\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \ unit : T" by fastforce then have "P\<^sub>1,E,h,sh \\<^sub>1 INIT C' (Cs,True) \ unit : T" by (auto; metis list.sel(2) list.set_sel(2)) then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' (Cs,True) \ unit))" by(simp only: exI) show ?thesis using assms wtrt proof(cases Cs) case (Cons C1 Cs1) then show ?thesis using assms wtrt by(case_tac "method P C1 clinit") clarsimp qed(clarsimp) qed lemma Jcc_pieces_InitError: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" and err: "sh C\<^sub>0 = Some(sfs,Error)" shows "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (RI (C\<^sub>0, THROW NoClassDefFoundError);Cs \ e) = (True, (calling_to_throwing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs'), (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), \vs'. P \ (None,h, (calling_to_throwing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs'),sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh')" proof - show ?thesis using assms proof(cases Cs) case (Cons C1 Cs1) then show ?thesis using assms by(case_tac "method P C1 clinit", case_tac "method P C\<^sub>0 clinit") clarsimp qed(clarsimp) qed lemma Jcc_pieces_InitObj: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' (sh(C\<^sub>0 \ (sfs,Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" shows "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs (sh(C\<^sub>0 \ (sfs,Processing))) I h' l' sh'' v xa (INIT C' (C\<^sub>0 # Cs,True) \ e) = (True, calling_to_called (hd frs')#(tl frs'), (None, h', (vs, l, C, M, pc, Called []) # frs, sh''), \vs'. P \ (None,h,calling_to_called (hd frs')#(tl frs'),sh') -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'')" proof - have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \ e)" using assms by simp then obtain T where "P\<^sub>1,E,h,sh \\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \ unit : T" by fastforce then have "P\<^sub>1,E,h,sh(C\<^sub>0 \ (sfs,Processing)) \\<^sub>1 INIT C' (C\<^sub>0 # Cs,True) \ unit : T" using assms by clarsimp (auto simp: fun_upd_apply) then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h (sh(C\<^sub>0 \ (sfs,Processing))) (INIT C' (C\<^sub>0 # Cs,True) \ unit))" by(simp only: exI) show ?thesis using assms wtrt by clarsimp qed lemma Jcc_pieces_InitNonObj: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "is_class P\<^sub>1 D" and "D \ set (C\<^sub>0#Cs)" and "\C \ set (C\<^sub>0#Cs). P\<^sub>1 \ C \\<^sup>* D" and pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' (sh(C\<^sub>0 \ (sfs,Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" shows "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs (sh(C\<^sub>0 \ (sfs,Processing))) I h' l' sh'' v xa (INIT C' (D # C\<^sub>0 # Cs,False) \ e) = (True, calling_to_calling (hd frs') D#(tl frs'), (None, h', (vs, l, C, M, pc, Called []) # frs, sh''), \vs'. P \ (None,h,calling_to_calling (hd frs') D#(tl frs'),sh') -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'')" proof - have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \ e)" using assms by simp then obtain T where "P\<^sub>1,E,h,sh \\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \ unit : T" by fastforce then have "P\<^sub>1,E,h,sh(C\<^sub>0 \ (sfs,Processing)) \\<^sub>1 INIT C' (D # C\<^sub>0 # Cs,False) \ unit : T" using assms by clarsimp (auto simp: fun_upd_apply) then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h (sh(C\<^sub>0 \ (sfs,Processing))) (INIT C' (D # C\<^sub>0 # Cs,False) \ unit))" by(simp only: exI) show ?thesis using assms wtrt by clarsimp qed lemma Jcc_pieces_InitRInit: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,True) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" shows "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; Cs \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), \vs'. P \ (None,h,frs',sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh')" proof - have cond: "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,True) \ e)" using assms by simp then have clinit: "\T. P\<^sub>1,E,h,sh \\<^sub>1 C\<^sub>0\\<^sub>sclinit([]) : T" using wf by clarsimp (auto simp: is_class_def intro: wf\<^sub>1_types_clinit) then obtain T where cT: "P\<^sub>1,E,h,sh \\<^sub>1 C\<^sub>0\\<^sub>sclinit([]) : T" by blast obtain T where "P\<^sub>1,E,h,sh \\<^sub>1 INIT C' (C\<^sub>0 # Cs,True) \ unit : T" using cond by fastforce then have "P\<^sub>1,E,h,sh \\<^sub>1 RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; Cs \ unit : T" using assms by (auto intro: cT) then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; Cs \ unit))" by(simp only: exI) then show ?thesis using assms by simp qed lemma Jcc_pieces_RInit_clinit: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([]));Cs \ e) = (True, frs', (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)" shows "Jcc_pieces P\<^sub>1 E C M h vs l pc (Called Cs) (tl frs') sh I h' l' sh' v xa (C\<^sub>0\\<^sub>sclinit([])) = (True, create_init_frame P C\<^sub>0#(vs,l,C,M,pc,Called Cs)#tl frs', (None, h', (vs,l,C,M,pc,Called Cs)#tl frs', sh'(C\<^sub>0\(fst(the(sh' C\<^sub>0)),Done))), P \ (None,h,create_init_frame P C\<^sub>0#(vs,l,C,M,pc,Called Cs)#tl frs',sh) -jvm\ (None,h',(vs, l, C, M, pc, Throwing Cs xa) # tl frs',sh'(C\<^sub>0 \ (fst(the(sh' C\<^sub>0)),Error))))" proof - have cond: "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([]));Cs \ e)" using assms by simp then have wtrt: "\T. P\<^sub>1,E,h,sh \\<^sub>1 C\<^sub>0\\<^sub>sclinit([]) : T" using wf by clarsimp (auto simp: is_class_def intro: wf\<^sub>1_types_clinit) then show ?thesis using assms by clarsimp qed lemma Jcc_pieces_RInit_Init: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1" and proc: "\C' \ set Cs. \sfs. sh'' C' = \(sfs,Processing)\" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([]));Cs \ e) = (True, frs', (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)" shows "Jcc_pieces P\<^sub>1 E C M h' vs l pc ics frs sh'' I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (INIT (last (C\<^sub>0#Cs)) (Cs,True) \ e) = (True, (vs, l, C, M, pc, Called Cs) # frs, (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), \vs'. P \ (None,h',(vs, l, C, M, pc, Called Cs) # frs,sh'') -jvm\ handle P C M xa h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1)" proof - have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([]));Cs \ e)" using assms by simp then have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; Cs \ unit))" by simp then obtain T where riwt: "P\<^sub>1,E,h,sh \\<^sub>1 RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([]));Cs \ unit : T" by meson then have "P\<^sub>1,E,h',sh'' \\<^sub>1 INIT (last (C\<^sub>0#Cs)) (Cs,True) \ unit : T" using proc proof(cases Cs) qed(auto) then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h' sh'' (INIT (last (C\<^sub>0#Cs)) (Cs,True) \ unit))" by(simp only: exI) show ?thesis using assms wtrt proof(cases Cs) case (Cons C1 Cs1) then show ?thesis using assms wtrt by(case_tac "method P C1 clinit") clarsimp qed(clarsimp) qed lemma Jcc_pieces_RInit_RInit: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (C\<^sub>0,e);D#Cs \ e') = (True, frs', rhs, err)" and hd: "hd frs' = (vs1,l1,C1,M1,pc1,ics1)" shows "Jcc_pieces P\<^sub>1 E C M h' vs l pc ics frs sh'' I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (D,Throw xa) ; Cs \ e') = (True, (vs1, l1, C1, M1, pc1, Throwing (D#Cs) xa) # tl frs', (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), \vs'. P \ (None,h',(vs1, l1, C1, M1, pc1, Throwing (D#Cs) xa) # tl frs',sh'') -jvm\ handle P C M xa h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1)" using assms by(case_tac "method P D clinit", cases "e = C\<^sub>0\\<^sub>sclinit([])") clarsimp+ subsubsection "JVM stepping lemmas" lemma jvm_Invoke: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "P,C,M,pc \ Invoke M' (length Ts)" and ha: "h\<^sub>2 a = \(Ca, fs)\" and method: "P\<^sub>1 \ Ca sees M', NonStatic : Ts\T = body in D" and len: "length pvs = length Ts" and "ls\<^sub>2' = Addr a # pvs @ replicate (max_vars body) undefined" shows "P \ (None, h\<^sub>2, (rev pvs @ Addr a # vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2) -jvm\ (None, h\<^sub>2, ([], ls\<^sub>2', D, M', 0, No_ics) # (rev pvs @ Addr a # vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2)" proof - have cname: "cname_of h\<^sub>2 (the_Addr ((rev pvs @ Addr a # vs) ! length Ts)) = Ca" using ha method len by(auto simp: nth_append) have r: "(rev pvs @ Addr a # vs) ! (length Ts) = Addr a" using len by(auto simp: nth_append) have exm: "\Ts T m D b. P \ Ca sees M',b:Ts \ T = m in D" using sees_method_compP[OF method] by fastforce show ?thesis using assms cname r exm by simp qed lemma jvm_Invokestatic: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "P,C,M,pc \ Invokestatic C' M' (length Ts)" and sh: "sh\<^sub>2 D = Some(sfs,Done)" and method: "P\<^sub>1 \ C' sees M', Static : Ts\T = body in D" and len: "length pvs = length Ts" and "ls\<^sub>2' = pvs @ replicate (max_vars body) undefined" shows "P \ (None, h\<^sub>2, (rev pvs @ vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2) -jvm\ (None, h\<^sub>2, ([], ls\<^sub>2', D, M', 0, No_ics) # (rev pvs @ vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2)" proof - have exm: "\Ts T m D b. P \ C' sees M',b:Ts \ T = m in D" using sees_method_compP[OF method] by fastforce show ?thesis using assms exm by simp qed lemma jvm_Invokestatic_Called: assumes [simp]: "P \ compP\<^sub>2 P\<^sub>1" and "P,C,M,pc \ Invokestatic C' M' (length Ts)" and sh: "sh\<^sub>2 D = Some(sfs,i)" and method: "P\<^sub>1 \ C' sees M', Static : Ts\T = body in D" and len: "length pvs = length Ts" and "ls\<^sub>2' = pvs @ replicate (max_vars body) undefined" shows "P \ (None, h\<^sub>2, (rev pvs @ vs, ls\<^sub>2, C, M, pc, Called []) # frs, sh\<^sub>2) -jvm\ (None, h\<^sub>2, ([], ls\<^sub>2', D, M', 0, No_ics) # (rev pvs @ vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2)" proof - have exm: "\Ts T m D b. P \ C' sees M',b:Ts \ T = m in D" using sees_method_compP[OF method] by fastforce show ?thesis using assms exm by simp qed lemma jvm_Return_Init: "P,D,clinit,0 \ compE\<^sub>2 body @ [Return] \ P \ (None, h, (vs, ls, D, clinit, size(compE\<^sub>2 body), No_ics) # frs, sh) -jvm\ (None, h, frs, sh(D\(fst(the(sh D)),Done)))" -apply(simp add: exec_all_def1, rule r_into_rtrancl, rule exec_1I) -apply(cases frs, auto) -done +(is "?P \ P \ ?s1 -jvm\ ?s2") +proof - + assume ?P + then have "exec (P, ?s1) = \?s2\" by(cases frs) auto + then have "(?s1, ?s2) \ (exec_1 P)\<^sup>*" + by(rule exec_1I[THEN r_into_rtrancl]) + then show ?thesis by(simp add: exec_all_def1) +qed lemma jvm_InitNone: "\ ics_of f = Calling C Cs; sh C = None \ \ P \ (None,h,f#frs,sh) -jvm\ (None,h,f#frs,sh(C \ (sblank P C, Prepared)))" -apply(simp add: exec_all_def1, rule r_into_rtrancl, rule exec_1I) -apply(cases f) apply(rename_tac ics, case_tac ics, simp_all) -done +(is "\ ?P; ?Q \ \ P \ ?s1 -jvm\ ?s2") +proof - + assume assms: ?P ?Q + then obtain stk1 loc1 C1 M1 pc1 ics1 where "f = (stk1,loc1,C1,M1,pc1,ics1)" + by(cases f) simp + then have "exec (P, ?s1) = \?s2\" using assms + by(case_tac ics1) simp_all + then have "(?s1, ?s2) \ (exec_1 P)\<^sup>*" + by(rule exec_1I[THEN r_into_rtrancl]) + then show ?thesis by(simp add: exec_all_def1) +qed lemma jvm_InitDP: "\ ics_of f = Calling C Cs; sh C = \(sfs,i)\; i = Done \ i = Processing \ \ P \ (None,h,f#frs,sh) -jvm\ (None,h,(calling_to_scalled f)#frs,sh)" -apply(simp add: exec_all_def1, rule r_into_rtrancl, rule exec_1I) -apply(cases f) -apply(erule_tac P = "i = Done" in disjE) - apply simp_all -done +(is "\ ?P; ?Q; ?R \ \ P \ ?s1 -jvm\ ?s2") +proof - + assume assms: ?P ?Q ?R + then obtain stk1 loc1 C1 M1 pc1 ics1 where "f = (stk1,loc1,C1,M1,pc1,ics1)" + by(cases f) simp + then have "exec (P, ?s1) = \?s2\" using assms + by(case_tac i) simp_all + then have "(?s1, ?s2) \ (exec_1 P)\<^sup>*" + by(rule exec_1I[THEN r_into_rtrancl]) + then show ?thesis by(simp add: exec_all_def1) +qed lemma jvm_InitError: "sh C = \(sfs,Error)\ \ P \ (None,h,(vs,ls,C\<^sub>0,M,pc,Calling C Cs)#frs,sh) -jvm\ (None,h,(vs,ls,C\<^sub>0,M,pc,Throwing Cs (addr_of_sys_xcpt NoClassDefFoundError))#frs,sh)" by(clarsimp simp: exec_all_def1 intro!: r_into_rtrancl exec_1I) lemma exec_ErrorThrowing: "sh C = \(sfs,Error)\ \ exec (P, (None,h,calling_to_throwing (stk,loc,D,M,pc,Calling C Cs) a#frs,sh)) = Some (None,h,calling_to_sthrowing (stk,loc,D,M,pc,Calling C Cs) a #frs,sh)" by(clarsimp simp: exec_all_def1 fun_upd_idem_iff intro!: r_into_rtrancl exec_1I) lemma jvm_InitObj: "\ sh C = Some(sfs,Prepared); C = Object; sh' = sh(C \ (sfs,Processing)) \ \ P \ (None, h, (vs,ls,C\<^sub>0,M,pc,Calling C Cs)#frs, sh) -jvm\ (None, h, (vs,ls,C\<^sub>0,M,pc,Called (C#Cs))#frs,sh')" -apply(simp add: exec_all_def1, rule r_into_rtrancl, rule exec_1I) -apply(case_tac "method P C clinit", simp) -done +(is "\ ?P; ?Q; ?R \ \ P \ ?s1 -jvm\ ?s2") +proof - + assume assms: ?P ?Q ?R + then have "exec (P, ?s1) = \?s2\" + by(case_tac "method P C clinit") simp + then have "(?s1, ?s2) \ (exec_1 P)\<^sup>*" + by(rule exec_1I[THEN r_into_rtrancl]) + then show ?thesis by(simp add: exec_all_def1) +qed lemma jvm_InitNonObj: "\ sh C = Some(sfs,Prepared); C \ Object; class P C = Some (D,r); sh' = sh(C \ (sfs,Processing)) \ \ P \ (None, h, (vs,ls,C\<^sub>0,M,pc,Calling C Cs)#frs, sh) -jvm\ (None, h, (vs,ls,C\<^sub>0,M,pc,Calling D (C#Cs))#frs, sh')" -apply(simp add: exec_all_def1, rule r_into_rtrancl, rule exec_1I) -apply(case_tac "method P C clinit", simp) -done +(is "\ ?P; ?Q; ?R; ?S \ \ P \ ?s1 -jvm\ ?s2") +proof - + assume assms: ?P ?Q ?R ?S + then have "exec (P, ?s1) = \?s2\" + by(case_tac "method P C clinit") simp + then have "(?s1, ?s2) \ (exec_1 P)\<^sup>*" + by(rule exec_1I[THEN r_into_rtrancl]) + then show ?thesis by(simp add: exec_all_def1) +qed lemma jvm_RInit_throw: "P \ (None,h,(vs,l,C,M,pc,Throwing [] xa) # frs,sh) -jvm\ handle P C M xa h vs l pc No_ics frs sh" -apply(simp add: exec_all_def1, rule r_into_rtrancl, rule exec_1I) -apply(simp add: handle_def split: bool.splits) -done +(is "P \ ?s1 -jvm\ ?s2") +proof - + have "exec (P, ?s1) = \?s2\" + by(simp add: handle_def split: bool.splits) + then have "(?s1, ?s2) \ (exec_1 P)\<^sup>*" + by(rule exec_1I[THEN r_into_rtrancl]) + then show ?thesis by(simp add: exec_all_def1) +qed lemma jvm_RInit_throw': "P \ (None,h,(vs,l,C,M,pc,Throwing [C'] xa) # frs,sh) -jvm\ handle P C M xa h vs l pc No_ics frs (sh(C':=Some(fst(the(sh C')), Error)))" -apply(simp add: exec_all_def1) -apply(rule_tac y = "(None,h,(vs,l,C,M,pc,Throwing [] xa) # frs,sh(C':=Some(fst(the(sh C')), Error)))" in rtrancl_trans) - apply(rule r_into_rtrancl, rule exec_1I) - apply(simp add: handle_def) -apply(cut_tac jvm_RInit_throw) -apply(simp add: exec_all_def1) -done +(is "P \ ?s1 -jvm\ ?s2") +proof - + let ?sy = "(None,h,(vs,l,C,M,pc,Throwing [] xa) # frs,sh(C':=Some(fst(the(sh C')), Error)))" + have "exec (P, ?s1) = \?sy\" by simp + then have "(?s1, ?sy) \ (exec_1 P)\<^sup>*" + by(rule exec_1I[THEN r_into_rtrancl]) + also have "(?sy, ?s2) \ (exec_1 P)\<^sup>*" + using jvm_RInit_throw by(simp add: exec_all_def1) + ultimately show ?thesis by(simp add: exec_all_def1) +qed lemma jvm_Called: "P \ (None, h, (vs, l, C, M, pc, Called (C\<^sub>0 # Cs)) # frs, sh) -jvm\ (None, h, create_init_frame P C\<^sub>0 # (vs, l, C, M, pc, Called Cs) # frs, sh)" by(simp add: exec_all_def1 r_into_rtrancl exec_1I) lemma jvm_Throwing: "P \ (None, h, (vs, l, C, M, pc, Throwing (C\<^sub>0#Cs) xa') # frs, sh) -jvm\ (None, h, (vs, l, C, M, pc, Throwing Cs xa') # frs, sh(C\<^sub>0 \ (fst (the (sh C\<^sub>0)), Error)))" by(simp add: exec_all_def1 r_into_rtrancl exec_1I) subsubsection "Other lemmas for correctness proof" lemma assumes wf:"wf_prog wf_md P" and ex: "class P C = Some a" shows create_init_frame_wf_eq: "create_init_frame (compP\<^sub>2 P) C = (stk,loc,D,M,pc,ics) \ D=C" using wf_sees_clinit[OF wf ex] by(cases "method P C clinit", auto) lemma beforex_try: - "\ {pc..2(try e\<^sub>1 catch(Ci i) e\<^sub>2))} \ I; - P,C,M \ compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs) / I,size vs \ - \ P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs) / {pc..2 e\<^sub>1)},size vs" -apply(clarsimp simp:beforex_def split:if_split_asm) -apply(rename_tac xt\<^sub>0 xt\<^sub>1) apply(rule_tac x=xt\<^sub>0 in exI) -apply(auto simp: pcs_subset(1)) -using atLeastLessThan_iff by blast +assumes pcI: "{pc..2(try e\<^sub>1 catch(Ci i) e\<^sub>2))} \ I" + and bx: "P,C,M \ compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs) / I,size vs" +shows "P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs) / {pc..2 e\<^sub>1)},size vs" +proof - + obtain xt\<^sub>0 xt\<^sub>1 where + "beforex\<^sub>0 P C M (size vs) I (compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs)) xt\<^sub>0 xt\<^sub>1" + using bx by(clarsimp simp:beforex_def) + then have "\xt1. beforex\<^sub>0 P C M (size vs) {pc..2 e\<^sub>1)} + (compxE\<^sub>2 e\<^sub>1 pc (size vs)) xt\<^sub>0 xt1" + using pcI pcs_subset(1) atLeastLessThan_iff by simp blast + then show ?thesis using beforex_def by blast +qed \ \ Evaluation of initialization expressions \ (* --needs J1 and EConform; version for eval found in Equivalence *) lemma shows eval\<^sub>1_init_return: "P \\<^sub>1 \e,s\ \ \e',s'\ \ iconf (shp\<^sub>1 s) e \ (\Cs b. e = INIT C' (Cs,b) \ unit) \ (\C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs@[C'] \ unit) \ (\e\<^sub>0. e = RI(C',e\<^sub>0);Nil \ unit) \ (val_of e' = Some v \ (\sfs i. shp\<^sub>1 s' C' = \(sfs,i)\ \ (i = Done \ i = Processing))) \ (throw_of e' = Some a \ (\sfs i. shp\<^sub>1 s' C' = \(sfs,Error)\))" and "P \\<^sub>1 \es,s\ [\] \es',s'\ \ True" proof(induct rule: eval\<^sub>1_evals\<^sub>1.inducts) case (InitFinal\<^sub>1 e s e' s' C b) then show ?case by(auto simp: initPD_def dest: eval\<^sub>1_final_same) next case (InitDone\<^sub>1 sh C sfs C' Cs e h l e' s') then have "final e'" using eval\<^sub>1_final by simp then show ?case proof(rule finalE) fix v assume e': "e' = Val v" then show ?thesis using InitDone\<^sub>1 initPD_def proof(cases Cs) qed(auto) next fix a assume e': "e' = throw a" then show ?thesis using InitDone\<^sub>1 initPD_def proof(cases Cs) qed(auto) qed next case (InitProcessing\<^sub>1 sh C sfs C' Cs e h l e' s') then have "final e'" using eval\<^sub>1_final by simp then show ?case proof(rule finalE) fix v assume e': "e' = Val v" then show ?thesis using InitProcessing\<^sub>1 initPD_def proof(cases Cs) qed(auto) next fix a assume e': "e' = throw a" then show ?thesis using InitProcessing\<^sub>1 initPD_def proof(cases Cs) qed(auto) qed next case (InitError\<^sub>1 sh C sfs Cs e h l e' s' C') show ?case proof(cases Cs) case Nil then show ?thesis using InitError\<^sub>1 by simp next case (Cons C2 list) then have "final e'" using InitError\<^sub>1 eval\<^sub>1_final by simp then show ?thesis proof(rule finalE) fix v assume e': "e' = Val v" show ?thesis using InitError\<^sub>1.hyps(2) e' rinit\<^sub>1_throwE by blast next fix a assume e': "e' = throw a" then show ?thesis using Cons InitError\<^sub>1 cons_to_append[of list] by clarsimp qed qed next case (InitRInit\<^sub>1 C Cs h l sh e' s' C') show ?case proof(cases Cs) case Nil then show ?thesis using InitRInit\<^sub>1 by simp next case (Cons C' list) then show ?thesis using InitRInit\<^sub>1 Cons cons_to_append[of list] by clarsimp qed next case (RInit\<^sub>1 e s v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 s\<^sub>1) then have final: "final e\<^sub>1" using eval\<^sub>1_final by simp then show ?case proof(cases Cs) case Nil show ?thesis using final proof(rule finalE) fix v assume e': "e\<^sub>1 = Val v" show ?thesis using RInit\<^sub>1 Nil by(clarsimp, meson fun_upd_same initPD_def) next fix a assume e': "e\<^sub>1 = throw a" show ?thesis using RInit\<^sub>1 Nil by(clarsimp, meson fun_upd_same initPD_def) qed next case (Cons a list) show ?thesis using final proof(rule finalE) fix v assume e': "e\<^sub>1 = Val v" then show ?thesis using RInit\<^sub>1 Cons by(clarsimp, metis last.simps last_appendR list.distinct(1)) next fix a assume e': "e\<^sub>1 = throw a" then show ?thesis using RInit\<^sub>1 Cons by(clarsimp, metis last.simps last_appendR list.distinct(1)) qed qed next case (RInitInitFail\<^sub>1 e s a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 s\<^sub>1) then have final: "final e\<^sub>1" using eval\<^sub>1_final by simp then show ?case proof(rule finalE) fix v assume e': "e\<^sub>1 = Val v" then show ?thesis using RInitInitFail\<^sub>1 by(clarsimp, meson exp.distinct(101) rinit\<^sub>1_throwE) next fix a' assume e': "e\<^sub>1 = Throw a'" then have "iconf (sh'(C \ (sfs, Error))) a" using RInitInitFail\<^sub>1.hyps(1) eval\<^sub>1_final by fastforce then show ?thesis using RInitInitFail\<^sub>1 e' by(clarsimp, meson Cons_eq_append_conv list.inject) qed qed(auto simp: fun_upd_same) lemma init\<^sub>1_Val_PD: "P \\<^sub>1 \INIT C' (Cs,b) \ unit,s\ \ \Val v,s'\ \ iconf (shp\<^sub>1 s) (INIT C' (Cs,b) \ unit) \ \sfs i. shp\<^sub>1 s' C' = \(sfs,i)\ \ (i = Done \ i = Processing)" by(drule_tac v = v in eval\<^sub>1_init_return, simp+) lemma init\<^sub>1_throw_PD: "P \\<^sub>1 \INIT C' (Cs,b) \ unit,s\ \ \throw a,s'\ \ iconf (shp\<^sub>1 s) (INIT C' (Cs,b) \ unit) \ \sfs i. shp\<^sub>1 s' C' = \(sfs,Error)\" by(drule_tac a = a in eval\<^sub>1_init_return, simp+) -lemma rinit\<^sub>1_Val_PD: "P \\<^sub>1 \RI(C,e\<^sub>0);Cs \ unit,s\ \ \Val v,s'\ - \ iconf (shp\<^sub>1 s) (RI(C,e\<^sub>0);Cs \ unit) \ last(C#Cs) = C' - \ \sfs i. shp\<^sub>1 s' C' = \(sfs,i)\ \ (i = Done \ i = Processing)" -apply(drule_tac C' = C' and v = v in eval\<^sub>1_init_return, simp_all) -apply (metis append_butlast_last_id) -done +lemma rinit\<^sub>1_Val_PD: +assumes eval: "P \\<^sub>1 \RI(C,e\<^sub>0);Cs \ unit,s\ \ \Val v,s'\" + and iconf: "iconf (shp\<^sub>1 s) (RI(C,e\<^sub>0);Cs \ unit)" and last: "last(C#Cs) = C'" +shows "\sfs i. shp\<^sub>1 s' C' = \(sfs,i)\ \ (i = Done \ i = Processing)" +proof(cases Cs) + case Nil + then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] last by simp +next + case (Cons a list) + then have nNil: "Cs \ []" by simp + then have "\Cs'. Cs = Cs' @ [C']" using last append_butlast_last_id[OF nNil] + by(rule_tac x="butlast Cs" in exI) simp + then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] by simp +qed -lemma rinit\<^sub>1_throw_PD: "P \\<^sub>1 \RI(C,e\<^sub>0);Cs \ unit,s\ \ \throw a,s'\ - \ iconf (shp\<^sub>1 s) (RI(C,e\<^sub>0);Cs \ unit) \ last(C#Cs) = C' - \ \sfs i. shp\<^sub>1 s' C' = \(sfs,Error)\" -apply(drule_tac C' = C' and a = a in eval\<^sub>1_init_return, simp_all) -apply (metis append_butlast_last_id) -done +lemma rinit\<^sub>1_throw_PD: +assumes eval: "P \\<^sub>1 \RI(C,e\<^sub>0);Cs \ unit,s\ \ \throw a,s'\" + and iconf: "iconf (shp\<^sub>1 s) (RI(C,e\<^sub>0);Cs \ unit)" and last: "last(C#Cs) = C'" +shows "\sfs. shp\<^sub>1 s' C' = \(sfs,Error)\" +proof(cases Cs) + case Nil + then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] last by simp +next + case (Cons a list) + then have nNil: "Cs \ []" by simp + then have "\Cs'. Cs = Cs' @ [C']" using last append_butlast_last_id[OF nNil] + by(rule_tac x="butlast Cs" in exI) simp + then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] by simp +qed subsubsection "The proof" lemma fixes P\<^sub>1 defines [simp]: "P \ compP\<^sub>2 P\<^sub>1" assumes wf: "wf_J\<^sub>1_prog P\<^sub>1" shows Jcc: "P\<^sub>1 \\<^sub>1 \e,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\ \ \ef,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\ \ (\E C M pc ics v xa vs frs I. \ Jcc_cond P\<^sub>1 E C M vs pc ics I h\<^sub>0 sh\<^sub>0 e \ \ (ef = Val v \ P \ (None,h\<^sub>0,Jcc_frames P C M vs ls\<^sub>0 pc ics frs e,sh\<^sub>0) -jvm\ Jcc_rhs P\<^sub>1 E C M vs ls\<^sub>0 pc ics frs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v e) \ (ef = Throw xa \ Jcc_err P C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 xa e) )" (*<*) (is "_ \ (\E C M pc ics v xa vs frs I. PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 ef h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I)") (*>*) and "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\ [\] \fs,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\ \ (\C M pc ics ws xa es' vs frs I. \ P,C,M,pc \ compEs\<^sub>2 es; P,C,M \ compxEs\<^sub>2 es pc (size vs)/I,size vs; {pc..2 es)} \ I; ics = No_ics; \sub_RIs es \ \ (fs = map Val ws \ P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(rev ws @ vs,ls\<^sub>1,C,M,pc+size(compEs\<^sub>2 es),ics)#frs,sh\<^sub>1)) \ (fs = map Val ws @ Throw xa # es' \ (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compEs\<^sub>2 es) \ \ caught P pc\<^sub>1 h\<^sub>1 xa (compxEs\<^sub>2 es pc (size vs)) \ (\vs'. P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ handle P C M xa h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1))))" (*<*) (is "_ \ (\C M pc ics ws xa es' vs frs I. PROP ?Ps es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 fs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C M pc ics ws xa es' vs frs I)") proof (induct rule:eval\<^sub>1_evals\<^sub>1_inducts) case New\<^sub>1 thus ?case by auto next case (NewFail\<^sub>1 sh C' sfs h ls) let ?xa = "addr_of_sys_xcpt OutOfMemory" have "P \ (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\ handle P C M ?xa h vs ls pc ics frs sh" using NewFail\<^sub>1 by(clarsimp simp: handle_def) then show ?case by(auto intro!: exI[where x="[]"]) next case (NewInit\<^sub>1 sh C' h ls v' h' ls' sh' a FDTs h'') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (new C') = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (new C')),ics)#frs,sh'), err)" using NewInit\<^sub>1.prems(1) by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' ([C'],False) \ unit))" using has_fields_is_class[OF NewInit\<^sub>1.hyps(5)] by auto then obtain err' where pcs': "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v' xa (INIT C' ([C'],False) \ unit) = (True, (vs,ls,C,M,pc,Calling C' []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), err')" using NewInit\<^sub>1.prems(1) by auto have IH: "PROP ?P (INIT C' ([C'],False) \ unit) h ls sh (Val v') h' ls' sh' E C M pc ics v' xa vs frs I" by fact have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF NewInit\<^sub>1.hyps(2)]) obtain sfs i where sh': "sh' C' = Some(sfs,i)" using init\<^sub>1_Val_PD[OF NewInit\<^sub>1.hyps(2)] by clarsimp have "P \ (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\ (None,h,(vs,ls,C,M,pc,Calling C' [])#frs,sh)" proof(cases "sh C'") case None then show ?thesis using NewInit\<^sub>1.prems by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using NewInit\<^sub>1.hyps(1) NewInit\<^sub>1.prems Some by(cases ics; case_tac i) auto qed also have "P \ \ -jvm\ (None, h', (vs, ls, C, M, pc, Called []) # frs, sh')" using IH pcs' by auto also have "P \ \ -jvm\ (None, h'', (Addr a#vs, ls, C, M, Suc pc, ics) # frs, sh')" using NewInit\<^sub>1.hyps(1,2,4-6) NewInit\<^sub>1.prems sh' by(cases ics) auto finally show ?case using pcs ls by clarsimp next case (NewInitOOM\<^sub>1 sh C' h ls v' h' ls' sh') let ?xa = "addr_of_sys_xcpt OutOfMemory" obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (new C') = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (new C')),ics)#frs,sh'), err)" using NewInitOOM\<^sub>1.prems(1) by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' ([C'],False) \ unit))" using NewInitOOM\<^sub>1.hyps(5) by auto then obtain err' where pcs': "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v' xa (INIT C' ([C'],False) \ unit) = (True, (vs,ls,C,M,pc,Calling C' []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), err')" using NewInitOOM\<^sub>1.prems(1) by auto have IH: "PROP ?P (INIT C' ([C'],False) \ unit) h ls sh (Val v') h' ls' sh' E C M pc ics v' xa vs frs I" by fact have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF NewInitOOM\<^sub>1.hyps(2)]) have "iconf (shp\<^sub>1 (h, ls, sh)) (INIT C' ([C'],False) \ unit)" by simp then obtain sfs i where sh': "sh' C' = Some(sfs,i)" using init\<^sub>1_Val_PD[OF NewInitOOM\<^sub>1.hyps(2)] by clarsimp have "P \ (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\ (None,h,(vs,ls,C,M,pc,Calling C' [])#frs,sh)" proof(cases "sh C'") case None then show ?thesis using NewInitOOM\<^sub>1.prems by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using NewInitOOM\<^sub>1.hyps(1) NewInitOOM\<^sub>1.prems Some by(cases ics; case_tac i) auto qed also have "P \ \ -jvm\ (None, h', (vs, ls, C, M, pc, Called []) # frs, sh')" using IH pcs' by auto also have "P \ \ -jvm\ handle P C M ?xa h' vs ls pc ics frs sh'" using NewInitOOM\<^sub>1.hyps(1,2,4,5) NewInitOOM\<^sub>1.prems sh' by(auto simp: handle_def) finally show ?case using pcs ls by(simp, metis (no_types) append_Nil le_refl lessI) next case (NewInitThrow\<^sub>1 sh C' h ls a h' ls' sh') obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (new C') = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (new C')),ics)#frs,sh'), err)" using NewInitThrow\<^sub>1.prems(1) by clarsimp obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF NewInitThrow\<^sub>1.hyps(2)] by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' ([C'],False) \ unit))" using NewInitThrow\<^sub>1.hyps(4) by auto then obtain vs' where pcs': "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v a' (INIT C' ([C'],False) \ unit) = (True, (vs,ls,C,M,pc,Calling C' []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), P \ (None,h,(vs,ls,C,M,pc,Calling C' []) # frs,sh) -jvm\ handle P C M a' h' (vs'@vs) ls pc ics frs sh')" using NewInitThrow\<^sub>1.prems(1) by simp blast have IH: "PROP ?P (INIT C' ([C'],False) \ unit) h ls sh (throw a) h' ls' sh' E C M pc ics v a' vs frs I" by fact have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF NewInitThrow\<^sub>1.hyps(2)]) then have "P \ (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\ (None,h,(vs,ls,C,M,pc,Calling C' []) # frs,sh)" proof(cases "sh C'") case None then show ?thesis using NewInitThrow\<^sub>1.prems by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using NewInitThrow\<^sub>1.hyps(1) NewInitThrow\<^sub>1.prems Some by(cases ics; case_tac i) auto qed also have "P \ \ -jvm\ handle P C M a' h' (vs'@vs) ls pc ics frs sh'" using IH pcs' throw by auto finally show ?case using throw ls by auto next case (Cast\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs C') let ?pc = "pc + length(compE\<^sub>2 e)" obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)" using Cast\<^sub>1.prems(1) by auto have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact then have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using Jcc_pieces_Cast[OF assms(1) pcs, of "Addr a"] Cast\<^sub>1.prems pcs by auto also have "P \ \ -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)" using Cast\<^sub>1 by (auto simp add:cast_ok_def) finally show ?case by auto next case (CastNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C') let ?pc = "pc + length(compE\<^sub>2 e)" obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)" using CastNull\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs I" by fact then have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using Jcc_pieces_Cast[OF assms(1) pcs, of Null] CastNull\<^sub>1.prems pcs by auto also have "P \ \ -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)" using CastNull\<^sub>1 by (auto simp add:cast_ok_def) finally show ?case by auto next case (CastFail\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs C') let ?pc = "pc + length(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt ClassCast" obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)" using CastFail\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact then have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using Jcc_pieces_Cast[OF assms(1) pcs, of "Addr a"] CastFail\<^sub>1.prems pcs by auto also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using CastFail\<^sub>1 by (auto simp add:handle_def cast_ok_def) finally have exec: "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ \". show ?case (is "?N \ (?eq \ ?err)") proof show ?N by simp next { assume ?eq then have ?err using exec by (auto intro!: exI[where x="?pc"] exI[where x="[Addr a]"]) } thus "?eq \ ?err" by simp qed next case (CastThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C') obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)" using CastThrow\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact show ?case using IH Jcc_pieces_Cast[OF assms(1) pcs, of v] CastThrow\<^sub>1.prems pcs less_SucI by(simp, blast) next case Val\<^sub>1 thus ?case by auto next case Var\<^sub>1 thus ?case by auto next case (BinOp\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 v\<^sub>2 h\<^sub>2 ls\<^sub>2 sh\<^sub>2 bop w) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1 \bop\ e\<^sub>2) = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1 \bop\ e\<^sub>2)),ics)#frs,sh\<^sub>2), err)" using BinOp\<^sub>1.prems(1) by clarsimp let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val v\<^sub>1) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v\<^sub>1 xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v\<^sub>2) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v\<^sub>2 xa (v\<^sub>1#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using IH\<^sub>1 Jcc_pieces_BinOp1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v\<^sub>1] by simp also have "P \ \ -jvm\ (None,h\<^sub>2,(v\<^sub>2#v\<^sub>1#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using IH\<^sub>2 Jcc_pieces_BinOp2[OF assms(1) pcs, of h\<^sub>1 v\<^sub>1 ls\<^sub>1 sh\<^sub>1 v\<^sub>2] by (simp add: add.assoc) also have "P \ \ -jvm\ (None,h\<^sub>2,(w#vs,ls\<^sub>2,C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2)" using BinOp\<^sub>1 by(cases bop) auto finally show ?case using pcs by (auto split: bop.splits simp:add.assoc) next case (BinOpThrow\<^sub>1\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e h\<^sub>1 ls\<^sub>1 sh\<^sub>1 bop e\<^sub>2) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<^sub>1 \bop\ e\<^sub>2) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<^sub>1 \bop\ e\<^sub>2)),ics)#frs,sh\<^sub>1), err)" using BinOpThrow\<^sub>1\<^sub>1.prems(1) by clarsimp have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact show ?case using IH\<^sub>1 Jcc_pieces_BinOp1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v] BinOpThrow\<^sub>1\<^sub>1.prems nsub_RI_Jcc_pieces by auto next case (BinOpThrow\<^sub>2\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e h\<^sub>2 ls\<^sub>2 sh\<^sub>2 bop) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1 \bop\ e\<^sub>2) = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1 \bop\ e\<^sub>2)),ics)#frs,sh\<^sub>2), err)" using BinOpThrow\<^sub>2\<^sub>1.prems(1) by clarsimp let ?pc = "pc + length(compE\<^sub>2 e\<^sub>1)" have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val v\<^sub>1) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v\<^sub>1 xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc ics v xa (v\<^sub>1#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact let ?\\<^sub>1 = "(None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" have 1: "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ ?\\<^sub>1" using IH\<^sub>1 Jcc_pieces_BinOp1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v\<^sub>1] by simp have "(throw e = Val v \ P \ (None, h\<^sub>0, Jcc_frames P C M vs ls\<^sub>0 pc ics frs (e\<^sub>1 \bop\ e\<^sub>2), sh\<^sub>0) -jvm\ Jcc_rhs P\<^sub>1 E C M vs ls\<^sub>0 pc ics frs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v (e\<^sub>1 \bop\ e\<^sub>2)) \ (throw e = Throw xa \ (\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < pc + size(compE\<^sub>2 (e\<^sub>1 \bop\ e\<^sub>2)) \ \ caught P pc\<^sub>1 h\<^sub>2 xa (compxE\<^sub>2 (e\<^sub>1 \bop\ e\<^sub>2) pc (size vs)) \ (\vs'. P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>1 ics frs sh\<^sub>2)))" (is "?N \ (?eq \ (\pc\<^sub>2. ?H pc\<^sub>2))") proof show ?N by simp next { assume ?eq then obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc \ pc\<^sub>2 \ pc\<^sub>2 < ?pc + size(compE\<^sub>2 e\<^sub>2) \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc (size vs + 1))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@v\<^sub>1#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using IH\<^sub>2 Jcc_pieces_BinOp2[OF assms(1) pcs, of h\<^sub>1 v\<^sub>1 ls\<^sub>1 sh\<^sub>1 v] BinOpThrow\<^sub>2\<^sub>1.prems by clarsimp then have "?H pc\<^sub>2" using jvm_trans[OF 1 2] by(auto intro!: exI[where x="vs'@[v\<^sub>1]"]) hence "\pc\<^sub>2. ?H pc\<^sub>2" by iprover } thus "?eq \ (\pc\<^sub>2. ?H pc\<^sub>2)" by iprover qed then show ?case using pcs by simp blast next case (FAcc\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' fs F T D w) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\F{D}) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\F{D})),ics)#frs,sh\<^sub>1), err)" using FAcc\<^sub>1.prems(1) by clarsimp have "P\<^sub>1 \ D sees F,NonStatic:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAcc\<^sub>1.hyps(4)]]) then have field: "field P D F = (D,NonStatic,T)" by simp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact let ?pc = "pc + length(compE\<^sub>2 e)" have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_FAcc[OF pcs, of "Addr a"] pcs by simp also have "P \ \ -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)" using FAcc\<^sub>1 field by auto finally have "P \ (None, h\<^sub>0, frs', sh\<^sub>0) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)" by auto then show ?case using pcs by auto next case (FAccNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 F D) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\F{D}) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\F{D})),ics)#frs,sh\<^sub>1), err)" using FAccNull\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs I" by fact let ?pc = "pc + length(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt NullPointer" have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_FAcc[OF pcs, of Null] by simp also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using FAccNull\<^sub>1.prems by(fastforce simp:split_beta handle_def simp del: split_paired_Ex) finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[Null]"]) next case (FAccThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 F D) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\F{D}) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\F{D})),ics)#frs,sh\<^sub>1), err)" using FAccThrow\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact show ?case using IH Jcc_pieces_FAcc[OF pcs, of v] FAccThrow\<^sub>1.prems nsub_RI_Jcc_pieces less_Suc_eq by auto next case (FAccNone\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C fs F D) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\F{D}) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\F{D})),ics)#frs,sh\<^sub>1), err)" using FAccNone\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact let ?pc = "pc + length(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt NoSuchFieldError" have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_FAcc[OF pcs, of "Addr a"] by simp also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using FAccNone\<^sub>1 by(cases ics; clarsimp simp:split_beta handle_def simp del: split_paired_Ex) finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[Addr a]"]) next case (FAccStatic\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' fs F T D) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\F{D}) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\F{D})),ics)#frs,sh\<^sub>1), err)" using FAccStatic\<^sub>1.prems(1) by clarsimp have "P\<^sub>1 \ D sees F,Static:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAccStatic\<^sub>1.hyps(4)]]) then have field: "field P D F = (D,Static,T)" by simp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact let ?pc = "pc + length(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError" have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_FAcc[OF pcs, of "Addr a"] by simp also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using FAccStatic\<^sub>1 field by(fastforce simp:split_beta handle_def simp del: split_paired_Ex) finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[Addr a]"]) next case (SFAcc\<^sub>1 C' F t D sh sfs v' h ls) have has: "P\<^sub>1 \ D has F,Static:t in D" by(rule has_field_idemp[OF SFAcc\<^sub>1.hyps(1)]) have "P\<^sub>1 \ D sees F,Static:t in D" by(rule has_field_sees[OF has]) then have field: "field P D F = (D,Static,t)" by simp then have "P \ (None,h,Jcc_frames P C M vs ls pc ics frs (C'\\<^sub>sF{D}),sh) -jvm\ (None,h,(v'#vs,ls,C,M,Suc pc,ics)#frs,sh)" using SFAcc\<^sub>1 has by(cases ics) auto then show ?case by clarsimp next case (SFAccInit\<^sub>1 C' F t D sh h ls v' h' ls' sh' sfs i v'') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C'\\<^sub>sF{D}) = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D})),ics)#frs,sh'), err)" using SFAccInit\<^sub>1.prems(1) by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT D ([D],False) \ unit))" using has_field_is_class'[OF SFAccInit\<^sub>1.hyps(1)] by auto then obtain err' where pcs': "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v' xa (INIT D ([D],False) \ unit) = (True, (vs,ls,C,M,pc,Calling D []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), err')" using SFAccInit\<^sub>1.prems(1) by auto have IH: "PROP ?P (INIT D ([D],False) \ unit) h ls sh (Val v') h' ls' sh' E C M pc ics v' xa vs frs I" by fact have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF SFAccInit\<^sub>1.hyps(3)]) have has: "P\<^sub>1 \ D has F,Static:t in D" by(rule has_field_idemp[OF SFAccInit\<^sub>1.hyps(1)]) have "P\<^sub>1 \ D sees F,Static:t in D" by(rule has_field_sees[OF has]) then have field: "field P D F = (D,Static,t)" by simp have "P \ (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\ (None,h,(vs,ls,C,M,pc,Calling D [])#frs,sh)" proof(cases "sh D") case None then show ?thesis using SFAccInit\<^sub>1.hyps(1,2,5,6) SFAccInit\<^sub>1.prems field by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using SFAccInit\<^sub>1.hyps(1,2,5,6) SFAccInit\<^sub>1.prems field Some by(cases ics; case_tac i) auto qed also have "P \ ... -jvm\ (None, h', (vs, ls, C, M, pc, Called []) # frs, sh')" using IH pcs' by auto also have "P \ ... -jvm\ (None, h', (v''#vs, ls, C, M, Suc pc, ics) # frs, sh')" using SFAccInit\<^sub>1.hyps(1,2,5,6) SFAccInit\<^sub>1.prems has field by(cases ics) auto finally show ?case using pcs ls by clarsimp next case (SFAccInitThrow\<^sub>1 C' F t D sh h ls a h' ls' sh') obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C'\\<^sub>sF{D}) = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D})),ics)#frs,sh'), err)" using SFAccInitThrow\<^sub>1.prems(1) by clarsimp obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF SFAccInitThrow\<^sub>1.hyps(3)] by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT D ([D],False) \ unit))" using has_field_is_class'[OF SFAccInitThrow\<^sub>1.hyps(1)] by auto then obtain vs' where pcs': "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v a' (INIT D ([D],False) \ unit) = (True, (vs,ls,C,M,pc,Calling D []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), P \ (None,h,(vs,ls,C,M,pc,Calling D []) # frs,sh) -jvm\ handle P C M a' h' (vs'@vs) ls pc ics frs sh')" using SFAccInitThrow\<^sub>1.prems(1) by simp blast have IH: "PROP ?P (INIT D ([D],False) \ unit) h ls sh (throw a) h' ls' sh' E C M pc ics v a' vs frs I" by fact have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF SFAccInitThrow\<^sub>1.hyps(3)]) have has: "P\<^sub>1 \ D has F,Static:t in D" by(rule has_field_idemp[OF SFAccInitThrow\<^sub>1.hyps(1)]) have "P\<^sub>1 \ D sees F,Static:t in D" by(rule has_field_sees[OF has]) then have field: "field P D F = (D,Static,t)" by simp then have "P \ (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\ (None,h,(vs,ls,C,M,pc,Calling D []) # frs,sh)" proof(cases "sh D") case None then show ?thesis using SFAccInitThrow\<^sub>1.hyps(1,2) SFAccInitThrow\<^sub>1.prems field by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using SFAccInitThrow\<^sub>1.hyps(1,2) SFAccInitThrow\<^sub>1.prems field Some by(cases ics; case_tac i) auto qed also have "P \ \ -jvm\ handle P C M a' h' (vs'@vs) ls pc ics frs sh'" using IH pcs' throw by auto finally show ?case using throw ls by auto next case (SFAccNone\<^sub>1 C' F D h\<^sub>1 ls\<^sub>1 sh\<^sub>1) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\\<^sub>sF{D}) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D})),ics)#frs,sh\<^sub>1), err)" by clarsimp let ?xa = "addr_of_sys_xcpt NoSuchFieldError" have "P \ (None,h\<^sub>1,frs',sh\<^sub>1) -jvm\ handle P C M ?xa h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1" using SFAccNone\<^sub>1 pcs by(cases ics; clarsimp simp:split_beta handle_def simp del: split_paired_Ex) then show ?case using pcs by(auto intro!: exI[where x = pc] exI[where x="[]"]) next case (SFAccNonStatic\<^sub>1 C' F t D h\<^sub>1 ls\<^sub>1 sh\<^sub>1) let ?frs' = "(vs, ls\<^sub>1, C, M, pc, ics) # frs" let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError" have "P\<^sub>1 \ D sees F,NonStatic:t in D" by(rule has_field_sees[OF has_field_idemp[OF SFAccNonStatic\<^sub>1.hyps(1)]]) then have field: "field P D F = (D,NonStatic,t)" by simp have "P \ (None,h\<^sub>1,?frs',sh\<^sub>1) -jvm\ handle P C M ?xa h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1" using SFAccNonStatic\<^sub>1 proof(cases ics) case No_ics then show ?thesis using SFAccNonStatic\<^sub>1 field by (auto simp:split_beta handle_def simp del: split_paired_Ex) qed(simp_all) then show ?case by (auto intro!: exI[where x = pc] exI[where x="[]"]) next case (LAss\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 i ls\<^sub>2) let ?pc = "pc + length(compE\<^sub>2 e)" obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (i:=e) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (i:=e)),ics)#frs,sh\<^sub>1), err)" using LAss\<^sub>1.prems(1) by auto have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact then have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using Jcc_pieces_LAss[OF assms(1) pcs, of w] LAss\<^sub>1.prems pcs by auto also have "P \ \ -jvm\ (None,h\<^sub>1,(Unit#vs,ls\<^sub>2,C,M,?pc+2,ics)#frs,sh\<^sub>1)" using LAss\<^sub>1 by (auto simp add:cast_ok_def) finally show ?case by auto next case (LAssThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 i) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (i:=e) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (i:=e)),ics)#frs,sh\<^sub>1), err)" using LAssThrow\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact show ?case using IH Jcc_pieces_LAss[OF assms(1) pcs, of v] LAssThrow\<^sub>1.prems pcs less_SucI by(simp, blast) next case (FAss\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs F T D fs' h\<^sub>2') obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\F{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)" using FAss\<^sub>1.prems(1) by clarsimp have "P\<^sub>1 \ D sees F,NonStatic:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAss\<^sub>1.hyps(6)]]) then have field: "field P D F = (D,NonStatic,T)" by simp let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 "Addr a"] by simp also have "P \ \ -jvm\ (None,h\<^sub>2,(w#Addr a#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 "Addr a" ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc) also have "P \ \ -jvm\ (None,h\<^sub>2',(Unit#vs,ls\<^sub>2,C,M,?pc\<^sub>2+2,ics)#frs,sh\<^sub>2)" using FAss\<^sub>1 field by auto finally show ?case using pcs by (auto simp:add.assoc) next case (FAssNull\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 F D) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\F{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)" using FAssNull\<^sub>1.prems(1) by clarsimp let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" let ?xa = "addr_of_sys_xcpt NullPointer" have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Null#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 Null] by simp also have "P \ \ -jvm\ (None,h\<^sub>2,(w#Null#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 Null ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc) also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (w#Null#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using FAssNull\<^sub>1 by(fastforce simp:split_beta handle_def simp del: split_paired_Ex) finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="w#[Null]"]) next case (FAssThrow\<^sub>2\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 F D) let ?frs' = "(vs, ls\<^sub>0, C, M, pc, ics) # frs" obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\F{D} := e\<^sub>2) = (True, ?frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)" using FAssThrow\<^sub>2\<^sub>1.prems(1) by clarsimp let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?\\<^sub>1 = "(None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v xa (w#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have 1: "P \ (None,h\<^sub>0,?frs',sh\<^sub>0) -jvm\ ?\\<^sub>1" using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 w] by simp show ?case (is "?N \ (?eq \ ?err)") proof show ?N by simp next { assume ?eq moreover have "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v xa (w#vs) frs (I - pcs (compxE\<^sub>2 e\<^sub>1 pc (length vs)))" by fact ultimately obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc\<^sub>1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1 + size(compE\<^sub>2 e\<^sub>2) \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1 (size vs + 1))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@w#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using FAssThrow\<^sub>2\<^sub>1.prems Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 w ls\<^sub>1 sh\<^sub>1] by auto have ?err using Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 w ls\<^sub>1 sh\<^sub>1] pc\<^sub>2 jvm_trans[OF 1 2] by(auto intro!: exI[where x=pc\<^sub>2] exI[where x="vs'@[w]"]) } thus "?eq \ ?err" by simp qed next case (FAssThrow\<^sub>1\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 F D e\<^sub>2) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<^sub>1\F{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)" using FAssThrow\<^sub>1\<^sub>1.prems(1) by clarsimp have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact show ?case using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v] FAssThrow\<^sub>1\<^sub>1.prems nsub_RI_Jcc_pieces by auto next case (FAssNone\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs F D) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\F{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)" using FAssNone\<^sub>1.prems(1) by clarsimp let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" let ?xa = "addr_of_sys_xcpt NoSuchFieldError" have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 "Addr a"] by simp also have "P \ \ -jvm\ (None,h\<^sub>2,(w#Addr a#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 "Addr a" ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc) also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (w#Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using FAssNone\<^sub>1 by(fastforce simp:split_beta handle_def simp del: split_paired_Ex) finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="w#[Addr a]"]) next case (FAssStatic\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs F T D) obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\F{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)" using FAssStatic\<^sub>1.prems(1) by clarsimp have "P\<^sub>1 \ D sees F,Static:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAssStatic\<^sub>1.hyps(6)]]) then have field: "field P D F = (D,Static,T)" by simp let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)" let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError" have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 "Addr a"] by simp also have "P \ \ -jvm\ (None,h\<^sub>2,(w#Addr a#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 "Addr a" ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc) also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (w#Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using FAssStatic\<^sub>1 field by(fastforce simp:split_beta handle_def simp del: split_paired_Ex) finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="w#[Addr a]"]) next case (SFAss\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F T D sfs sfs' sh\<^sub>1') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\\<^sub>sF{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)" using SFAss\<^sub>1.prems(1) by clarsimp have "P\<^sub>1 \ D sees F,Static:T in D" by(rule has_field_sees[OF has_field_idemp[OF SFAss\<^sub>1.hyps(3)]]) then have field: "field P D F = (D,Static,T)" by simp have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)" have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp also have "P \ \ -jvm\ (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1')" using SFAss\<^sub>1.hyps(3-6) SFAss\<^sub>1.prems(1) field by auto also have "P \ ... -jvm\ (None,h\<^sub>1,(Unit#vs,ls\<^sub>1,C,M,?pc+2,ics)#frs,sh\<^sub>1')" using SFAss\<^sub>1 by auto finally show ?case using pcs by auto next case (SFAssInit\<^sub>1 e\<^sub>2 h ls sh w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F t D v' h' ls' sh' sfs i sfs' sh'') let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)" obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh'' v xa (C'\\<^sub>sF{D}:=e\<^sub>2) = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D}:=e\<^sub>2)),ics)#frs,sh''), err)" using SFAssInit\<^sub>1.prems(1) by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \ unit))" using has_field_is_class'[OF SFAssInit\<^sub>1.hyps(3)] by auto then obtain err' where pcs': "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1 I h' ls' sh' v' xa (INIT D ([D],False) \ unit) = (True, (w#vs,ls\<^sub>1,C,M,?pc,Calling D []) # frs, (None,h',(w#vs,ls\<^sub>1,C,M,?pc,Called [])#frs,sh'), err')" using SFAssInit\<^sub>1.prems(1) by simp have ls: "ls\<^sub>1 = ls'" by(rule init\<^sub>1_same_loc[OF SFAssInit\<^sub>1.hyps(5)]) have has: "P\<^sub>1 \ D has F,Static:t in D" by(rule has_field_idemp[OF SFAssInit\<^sub>1.hyps(3)]) have "P\<^sub>1 \ D sees F,Static:t in D" by(rule has_field_sees[OF has]) then have field: "field P D F = (D,Static,t)" by simp have IH: "PROP ?P e\<^sub>2 h ls sh (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact have IHI: "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v') h' ls' sh' E C M ?pc ics v' xa (w#vs) frs I" by fact have "P \ (None,h,frs',sh) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_SFAss[OF pcs, where v'=w] by simp also have "P \ \ -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,Calling D [])#frs,sh\<^sub>1)" proof(cases "sh\<^sub>1 D") case None then show ?thesis using None SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems field by(cases ics, auto) next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems field Some by(cases ics; case_tac i) auto qed also have "P \ ... -jvm\ (None, h', (w#vs, ls\<^sub>1, C, M, ?pc, Called []) # frs, sh')" using IHI pcs' by clarsimp also have "P \ ... -jvm\ (None, h', (vs, ls\<^sub>1, C, M, ?pc + 1, ics) # frs, sh'')" using SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems has field by(cases ics) auto also have "P \ ... -jvm\ (None, h', (Unit#vs, ls\<^sub>1, C, M, ?pc + 2, ics) # frs, sh'')" using SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems has field by(cases ics) auto finally show ?case using pcs ls by simp blast next case (SFAssInitThrow\<^sub>1 e\<^sub>2 h ls sh w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F t D a h' ls' sh') let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)" obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C'\\<^sub>sF{D}:=e\<^sub>2) = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D}:=e\<^sub>2)),ics)#frs,sh'), err)" using SFAssInitThrow\<^sub>1.prems(1) by clarsimp obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF SFAssInitThrow\<^sub>1.hyps(5)] by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \ unit))" using has_field_is_class'[OF SFAssInitThrow\<^sub>1.hyps(3)] by auto then obtain vs' where pcs': "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1 I h' ls' sh' v a' (INIT D ([D],False) \ unit) = (True, (w#vs,ls\<^sub>1,C,M,?pc,Calling D []) # frs, (None,h',(w#vs,ls\<^sub>1,C,M,?pc,Called [])#frs,sh'), P \ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,Calling D []) # frs,sh\<^sub>1) -jvm\ handle P C M a' h' (vs'@w#vs) ls\<^sub>1 ?pc ics frs sh')" using SFAssInitThrow\<^sub>1.prems(1) by simp blast have ls: "ls\<^sub>1 = ls'" by(rule init\<^sub>1_same_loc[OF SFAssInitThrow\<^sub>1.hyps(5)]) have has: "P\<^sub>1 \ D has F,Static:t in D" by(rule has_field_idemp[OF SFAssInitThrow\<^sub>1.hyps(3)]) have "P\<^sub>1 \ D sees F,Static:t in D" by(rule has_field_sees[OF has]) then have field: "field P D F = (D,Static,t)" by simp have IH: "PROP ?P e\<^sub>2 h ls sh (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact have IHI: "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw a) h' ls' sh' E C M ?pc ics v a' (w#vs) frs I" by fact have "P \ (None,h,(vs, ls, C, M, pc, ics) # frs,sh) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp blast also have "P \ \ -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,Calling D [])#frs,sh\<^sub>1)" proof(cases "sh\<^sub>1 D") case None then show ?thesis using SFAssInitThrow\<^sub>1.hyps(1,3,4,5) SFAssInitThrow\<^sub>1.prems field by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using SFAssInitThrow\<^sub>1.hyps(1,3,4,5) SFAssInitThrow\<^sub>1.prems field Some by(cases ics; case_tac i) auto qed also have "P \ ... -jvm\ handle P C M a' h' (vs'@w#vs) ls\<^sub>1 ?pc ics frs sh'" using IHI pcs' throw by auto finally show ?case using throw ls by(auto intro!: exI[where x = ?pc] exI[where x="vs'@[w]"]) next case (SFAssThrow\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F D) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\\<^sub>sF{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)" using SFAssThrow\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact show ?case using IH Jcc_pieces_SFAss[OF pcs, where v'=v] SFAssThrow\<^sub>1.prems nsub_RI_Jcc_pieces less_Suc_eq by auto next case (SFAssNone\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F D) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\\<^sub>sF{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)" using SFAssNone\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)" let ?xa = "addr_of_sys_xcpt NoSuchFieldError" have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using SFAssNone\<^sub>1 by(cases ics; clarsimp simp add: handle_def) finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[w]"]) next case (SFAssNonStatic\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F T D) then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\\<^sub>sF{D} := e\<^sub>2) = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)" using SFAssNonStatic\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)" let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError" have "P\<^sub>1 \ D sees F,NonStatic:T in D" by(rule has_field_sees[OF has_field_idemp[OF SFAssNonStatic\<^sub>1.hyps(3)]]) then have field: "field P D F = (D,NonStatic,T)" by simp have "P \ (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using SFAssNonStatic\<^sub>1 proof(cases ics) case No_ics then show ?thesis using SFAssNonStatic\<^sub>1 field by (auto simp:split_beta handle_def simp del: split_paired_Ex) qed(simp_all) finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[w]"]) next case (Call\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 Ca fs M' Ts T body D ls\<^sub>2' f h\<^sub>3 ls\<^sub>3 sh\<^sub>3) let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs" let ?\\<^sub>0 = "(None,h\<^sub>0,?frs\<^sub>0,sh\<^sub>0)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>1 = "(None,h\<^sub>1,(Addr a#vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" let ?frs\<^sub>2 = "(rev pvs @ Addr a # vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" let ?frs\<^sub>2' = "([], ls\<^sub>2', D,M',0,No_ics) # ?frs\<^sub>2" let ?\\<^sub>2' = "(None, h\<^sub>2, ?frs\<^sub>2', sh\<^sub>2)" have nclinit: "M' \ clinit" using wf_sees_clinit1[OF wf] visible_method_exists[OF Call\<^sub>1.hyps(6)] sees_method_idemp[OF Call\<^sub>1.hyps(6)] by fastforce have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) have invoke: "P,C,M,?pc\<^sub>2 \ Invoke M' (length Ts)" using Call\<^sub>1.hyps(7) Call\<^sub>1.prems(1) by clarsimp have nsub: "\ sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf Call\<^sub>1.hyps(6)]) obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa (e\M'(es)) = (True, ?frs\<^sub>0, (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>3), err)" using Call\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa (map Val pvs) (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact have "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp also have "P \ \ -jvm\ ?\\<^sub>2" using IH_es Call\<^sub>1.prems by fastforce also have "P \ \ -jvm\ ?\\<^sub>2'" using jvm_Invoke[OF assms(1) invoke _ Call\<^sub>1.hyps(6-8)] Call\<^sub>1.hyps(5) Call\<^sub>1.prems(1) by simp finally have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>2'". have "P\<^sub>1 \ Ca sees M',NonStatic: Ts\T = body in D" by fact then have M'_in_D: "P\<^sub>1 \ D sees M',NonStatic: Ts\T = body in D" by(rule sees_method_idemp) have M'_code: "compP\<^sub>2 P\<^sub>1,D,M',0 \ compE\<^sub>2 body @ [Return]" using beforeM M'_in_D by simp have M'_xtab: "compP\<^sub>2 P\<^sub>1,D,M' \ compxE\<^sub>2 body 0 0/{..2 body)},0" using M'_in_D by(rule beforexM) have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 f h\<^sub>3 ls\<^sub>3 sh\<^sub>3 (Class D # Ts) D M' 0 No_ics v xa [] ?frs\<^sub>2 ({..2 body)})" by fact have cond: "Jcc_cond P\<^sub>1 (Class D # Ts) D M' [] 0 No_ics {..2 body)} h\<^sub>2 sh\<^sub>2 body" using nsub_RI_Jcc_pieces[OF assms(1) nsub] M'_code M'_xtab by clarsimp show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>2' -jvm\ (None,h\<^sub>3,([v],ls\<^sub>3,D,M',size(compE\<^sub>2 body),No_ics)#?frs\<^sub>2,sh\<^sub>3)" using val IH_body Call\<^sub>1.prems M'_code cond nsub_RI_Jcc_pieces nsub by auto also have "P \ \ -jvm\ (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>3)" using Call\<^sub>1.hyps(7) M'_code M'_in_D nclinit by(cases T, auto) finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw with IH_body obtain pc\<^sub>2 vs' where pc\<^sub>2: "0 \ pc\<^sub>2 \ pc\<^sub>2 < size(compE\<^sub>2 body) \ \ caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and 2: "P \ ?\\<^sub>2' -jvm\ handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3" using Call\<^sub>1.prems M'_code M'_xtab cond nsub_RI_Jcc_pieces nsub by (auto simp del:split_paired_Ex) have "handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3 = handle P C M xa h\<^sub>3 (rev pvs @ Addr a # vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>3" using pc\<^sub>2 M'_in_D nclinit by(auto simp add:handle_def) then show "?err" using pc\<^sub>2 jvm_trans[OF 1 2] by(auto intro!:exI[where x="?pc\<^sub>2"] exI[where x="rev pvs@[Addr a]"]) qed qed next case (CallParamsThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 pvs ex es'' M') let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc,ics)#frs,sh\<^sub>0)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>1 = "(None,h\<^sub>1,(w # vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\M'(es)) = (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)" using CallParamsThrow\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp have Isubs: "{?pc\<^sub>1..2} \ I - pcs (compxE\<^sub>2 e pc (length vs))" using CallParamsThrow\<^sub>1.prems by clarsimp show ?case (is "?N \ (?eq \ ?err)") proof show ?N by simp next { assume ?eq moreover have "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa es'' (w#vs) frs (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact ultimately obtain vs' where "\pc\<^sub>2. (?pc\<^sub>1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1 + size(compEs\<^sub>2 es) \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es ?pc\<^sub>1 (size vs + 1))) \ P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@w#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" (is "\pc\<^sub>2. ?PC pc\<^sub>2 \ ?Exec pc\<^sub>2") using CallParamsThrow\<^sub>1 Isubs by auto then obtain pc\<^sub>2 where pc\<^sub>2: "?PC pc\<^sub>2" and 2: "?Exec pc\<^sub>2" by iprover then have "?err" using pc\<^sub>2 jvm_trans[OF 1 2] by(auto intro!: exI[where x="pc\<^sub>2"] exI[where x="vs'@[w]"]) } thus "?eq \ ?err" by simp qed next case (CallNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 M') have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" let ?xa = "addr_of_sys_xcpt NullPointer" obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\M'(es)) = (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)" using CallNull\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa (map Val pvs) (Null#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact have Isubs: "{pc + length (compE\<^sub>2 e)..2 e) + length (compEs\<^sub>2 es)} \ I - pcs (compxE\<^sub>2 e pc (length vs))" using CallNull\<^sub>1.prems by clarsimp have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using Jcc_pieces_Call1[OF pcs] IH by clarsimp also have "P \ \ -jvm\ (None,h\<^sub>2,(rev pvs@Null#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using CallNull\<^sub>1 IH_es Isubs by auto also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (rev pvs@Null#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using CallNull\<^sub>1.prems by(auto simp:split_beta handle_def nth_append simp del: split_paired_Ex) finally show ?case by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="rev pvs@[Null]"]) next case (CallObjThrow\<^sub>1 e h ls sh e' h' ls' sh' M' es) obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (e\M'(es)) = (True, (vs, ls, C,M,pc,ics)#frs, (None, h', (v#vs, ls', C,M,pc+size(compE\<^sub>2 (e\M'(es))),ics)#frs,sh'), err)" using CallObjThrow\<^sub>1.prems(1) by clarsimp obtain a' where throw: "throw e' = Throw a'" using eval\<^sub>1_final[OF CallObjThrow\<^sub>1.hyps(1)] by clarsimp have IH: "PROP ?P e h ls sh (throw e') h' ls' sh' E C M pc ics v a' vs frs (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact show ?case using IH Jcc_pieces_Call1[OF pcs] throw CallObjThrow\<^sub>1.prems nsub_RI_Jcc_pieces by auto next case (CallNone\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs M') let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs" let ?\\<^sub>0 = "(None,h\<^sub>0,?frs\<^sub>0,sh\<^sub>0)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>1 = "(None,h\<^sub>1,(Addr a#vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" let ?frs\<^sub>2 = "(rev pvs @ Addr a # vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" let ?xa = "addr_of_sys_xcpt NoSuchMethodError" have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) have aux: "(rev pvs @ Addr a # vs) ! length pvs = Addr a" by (metis length_rev nth_append_length) have nmeth: "\(\b Ts T body D. P \ C' sees M', b : Ts\T = body in D)" using sees_method_compPD CallNone\<^sub>1.hyps(6) by fastforce obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\M'(es)) = (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)" using CallNone\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa (map Val pvs) (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact have "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp also have "P \ \ -jvm\ ?\\<^sub>2" using IH_es CallNone\<^sub>1.prems by fastforce also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (rev pvs@Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using CallNone\<^sub>1.hyps(5) CallNone\<^sub>1.prems aux nmeth by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2", auto simp: handle_def) finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="rev pvs@[Addr a]"]) next case (CallStatic\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs M' Ts T body D) let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs" let ?\\<^sub>0 = "(None,h\<^sub>0,?frs\<^sub>0,sh\<^sub>0)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>1 = "(None,h\<^sub>1,(Addr a#vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" let ?frs\<^sub>2 = "(rev pvs @ Addr a # vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError" have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) have aux: "(rev pvs @ Addr a # vs) ! length pvs = Addr a" by (metis length_rev nth_append_length) obtain body' where method: "P \ C' sees M', Static : Ts\T = body' in D" by (metis CallStatic\<^sub>1.hyps(6) P_def compP\<^sub>2_def sees_method_compP) obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\M'(es)) = (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)" using CallStatic\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa (map Val pvs) (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact have "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp also have "P \ \ -jvm\ ?\\<^sub>2" using IH_es CallStatic\<^sub>1.prems by fastforce also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (rev pvs@Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using CallStatic\<^sub>1.hyps(5) CallStatic\<^sub>1.prems aux method by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2") (auto simp: handle_def; meson frames_of.cases) finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="rev pvs@[Addr a]"]) next case (SCallParamsThrow\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 pvs ex es'' C' M') show ?case proof(cases "M' = clinit \ es = []") case clinit: True then show ?thesis using SCallParamsThrow\<^sub>1.hyps(1,3) evals\<^sub>1_cases(1) by fastforce next case nclinit: False let ?\\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)" have Isubs: "{pc..2 es)} \ I" using SCallParamsThrow\<^sub>1.prems nclinit by clarsimp show ?thesis (is "?N \ (?eq \ ?err)") proof show ?N by simp next { assume ?eq moreover have "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa es'' vs frs I" by fact ultimately have "\pc\<^sub>2. (pc \ pc\<^sub>2 \ pc\<^sub>2 < pc + size(compEs\<^sub>2 es) \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es pc (size vs))) \ (\vs'. P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2)" (is "\pc\<^sub>2. ?PC pc\<^sub>2 \ ?Exec pc\<^sub>2") using SCallParamsThrow\<^sub>1 Isubs nclinit by auto then obtain pc\<^sub>2 where pc\<^sub>2: "?PC pc\<^sub>2" and 2: "?Exec pc\<^sub>2" by iprover then have "?err" using pc\<^sub>2 2 by(auto intro: exI[where x="pc\<^sub>2"]) } thus "?eq \ ?err" by iprover qed qed next case (SCallNone\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' M') show ?case proof(cases "M' = clinit \ es = []") case clinit: True then show ?thesis using SCallNone\<^sub>1.hyps(3) SCallNone\<^sub>1.prems by auto next case nclinit: False let ?\\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)" let ?frs\<^sub>2 = "(rev pvs @ vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" let ?xa = "addr_of_sys_xcpt NoSuchMethodError" have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) have nmeth: "\(\b Ts T body D. P \ C' sees M', b : Ts\T = body in D)" using sees_method_compPD SCallNone\<^sub>1.hyps(3) by fastforce have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa (map Val pvs) vs frs I" by fact have "P \ ?\\<^sub>1 -jvm\ ?\\<^sub>2" using IH_es SCallNone\<^sub>1.prems nclinit by auto fastforce+ also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (rev pvs@vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using SCallNone\<^sub>1.prems nmeth nclinit by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2", auto simp: handle_def) finally show ?thesis using nclinit by (auto intro: exI[where x = ?pc\<^sub>2]) qed next case (SCallNonStatic\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' M' Ts T body D) show ?case proof(cases "M' = clinit \ es = []") case clinit: True then show ?thesis using SCallNonStatic\<^sub>1.hyps(3) SCallNonStatic\<^sub>1.prems sees_method_fun by fastforce next case nclinit: False let ?\\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)" let ?frs\<^sub>2 = "(rev pvs @ vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError" have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) obtain body' where method: "P \ C' sees M', NonStatic : Ts\T = body' in D" by (metis SCallNonStatic\<^sub>1.hyps(3) P_def compP\<^sub>2_def sees_method_compP) have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa (map Val pvs) vs frs I" by fact have "P \ ?\\<^sub>1 -jvm\ ?\\<^sub>2" using IH_es SCallNonStatic\<^sub>1.prems nclinit by auto fastforce+ also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>2 (rev pvs@vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2" using SCallNonStatic\<^sub>1.prems method nclinit by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2") (auto simp: handle_def; meson frames_of.cases) finally show ?thesis using nclinit by (auto intro: exI[where x = ?pc\<^sub>2]) qed next case (SCallInitThrow\<^sub>1 es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 pvs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' M' Ts T body D a h\<^sub>2 ls\<^sub>2 sh\<^sub>2) show ?case proof(cases "M' = clinit \ es = []") case clinit: True then show ?thesis using SCallInitThrow\<^sub>1 by simp next case nclinit: False let ?\\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc,ics)#frs,sh\<^sub>0)" let ?pc\<^sub>1 = "pc + length(compEs\<^sub>2 es)" let ?frs\<^sub>1 = "(rev pvs @ vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs" let ?\\<^sub>1 = "(None,h\<^sub>1,?frs\<^sub>1,sh\<^sub>1)" let ?frs\<^sub>1' = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Calling D [])#frs" let ?\\<^sub>1' = "(None,h\<^sub>1,?frs\<^sub>1',sh\<^sub>1)" let ?frs\<^sub>2 = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Called [])#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" have ls: "ls\<^sub>1 = ls\<^sub>2" by(rule init\<^sub>1_same_loc[OF SCallInitThrow\<^sub>1.hyps(6)]) have method: "\m'. P \ C' sees M',Static:Ts\T = m' in D" using SCallInitThrow\<^sub>1.hyps(3) by (metis P_def compP\<^sub>2_def sees_method_compP) obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF SCallInitThrow\<^sub>1.hyps(6)] by clarsimp have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \ unit))" using sees_method_is_class'[OF SCallInitThrow\<^sub>1.hyps(3)] by auto then obtain err' where pcs': "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (rev pvs@vs) ls\<^sub>1 ?pc\<^sub>1 ics frs sh\<^sub>1 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (INIT D ([D],False) \ unit) = (True, ?frs\<^sub>1', (None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2), err')" using SCallInitThrow\<^sub>1.prems(1) nclinit by auto have IHI: "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v a' (rev pvs@vs) frs I" by fact have IH_es: "PROP ?Ps es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (map Val pvs) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C M pc ics pvs xa (map Val pvs) vs frs I" by fact have "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using IH_es SCallInitThrow\<^sub>1.prems nclinit by auto fastforce+ also have "P \ \ -jvm\ ?\\<^sub>1'" proof(cases "sh\<^sub>1 D") case None then show ?thesis using SCallInitThrow\<^sub>1.hyps(1,3-6) SCallInitThrow\<^sub>1.prems method by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using SCallInitThrow\<^sub>1.hyps(1,3-6) SCallInitThrow\<^sub>1.prems method Some by(cases ics; case_tac i, auto) qed also obtain vs' where "P \ \ -jvm\ handle P C M a' h\<^sub>2 (vs'@rev pvs@vs) ls\<^sub>1 ?pc\<^sub>1 ics frs sh\<^sub>2" using IHI pcs' throw by auto finally show ?thesis using nclinit throw ls by(auto intro!: exI[where x="?pc\<^sub>1"] exI[where x="vs'@rev pvs"]) qed next case (SCallInit\<^sub>1 es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 pvs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' M' Ts T body D v' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 ls\<^sub>2' e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3) show ?case proof(cases "M' = clinit \ es = []") case clinit: True then show ?thesis using SCallInit\<^sub>1 by simp next case nclinit: False let ?\\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc,ics)#frs,sh\<^sub>0)" let ?pc\<^sub>1 = "pc + length(compEs\<^sub>2 es)" let ?frs\<^sub>1 = "(rev pvs @ vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs" let ?\\<^sub>1 = "(None,h\<^sub>1,?frs\<^sub>1,sh\<^sub>1)" let ?frs\<^sub>1' = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Calling D [])#frs" let ?\\<^sub>1' = "(None,h\<^sub>1,?frs\<^sub>1',sh\<^sub>1)" let ?frs\<^sub>2 = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Called [])#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" let ?frs\<^sub>2' = "([], ls\<^sub>2', D,M',0,No_ics) # ?frs\<^sub>1" let ?\\<^sub>2' = "(None, h\<^sub>2, ?frs\<^sub>2', sh\<^sub>2)" have nclinit': "M' \ clinit" by fact have ics: "ics = No_ics" using SCallInit\<^sub>1.hyps(5) SCallInit\<^sub>1.prems by simp have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>0, ls\<^sub>0, sh\<^sub>0)\ [\] \map Val pvs,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) have invoke: "P,C,M,?pc\<^sub>1 \ Invokestatic C' M' (length Ts)" using SCallInit\<^sub>1.hyps(8) SCallInit\<^sub>1.prems nclinit by(auto simp: add.assoc) have nsub: "\ sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf SCallInit\<^sub>1.hyps(3)]) have ls: "ls\<^sub>1 = ls\<^sub>2" by(rule init\<^sub>1_same_loc[OF SCallInit\<^sub>1.hyps(6)]) obtain sfs i where sh\<^sub>2: "sh\<^sub>2 D = Some(sfs,i)" using init\<^sub>1_Val_PD[OF SCallInit\<^sub>1.hyps(6)] by clarsimp have method: "\m'. P \ C' sees M',Static:Ts\T = m' in D" using SCallInit\<^sub>1.hyps(3) by (metis P_def compP\<^sub>2_def sees_method_compP) have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \ unit))" using sees_method_is_class'[OF SCallInit\<^sub>1.hyps(3)] by auto then obtain err' where pcs': "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (rev pvs@vs) ls\<^sub>1 ?pc\<^sub>1 ics frs sh\<^sub>1 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v' xa (INIT D ([D],False) \ unit) = (True, ?frs\<^sub>1', (None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2), err')" using SCallInit\<^sub>1.prems(1) nclinit by auto have IHI: "PROP ?P (INIT D ([D],False) \ unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v') h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v' xa (rev pvs@vs) frs I" by fact have IH_es: "PROP ?Ps es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (map Val pvs) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C M pc ics pvs xa (map Val pvs) vs frs I" by fact have "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using IH_es SCallInit\<^sub>1.prems nclinit by auto fastforce+ also have "P \ \ -jvm\ ?\\<^sub>1'" proof(cases "sh\<^sub>1 D") case None then show ?thesis using SCallInit\<^sub>1.hyps(1,3-6,8-10) SCallInit\<^sub>1.prems method by(cases ics) auto next case (Some a) then obtain sfs i where "a = (sfs,i)" by(cases a) then show ?thesis using SCallInit\<^sub>1.hyps(1,3-6,8-10) SCallInit\<^sub>1.prems method Some by(cases ics; case_tac i, auto) qed also have "P \ \ -jvm\ ?\\<^sub>2" using IHI pcs' by auto also have "P \ \ -jvm\ ?\\<^sub>2'" using jvm_Invokestatic_Called[OF assms(1) invoke _ SCallInit\<^sub>1.hyps(3,8,9)] sh\<^sub>2 ics by auto finally have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>2'". have "P\<^sub>1 \ C' sees M',Static: Ts\T = body in D" by fact then have M'_in_D: "P\<^sub>1 \ D sees M',Static: Ts\T = body in D" by(rule sees_method_idemp) have M'_code: "compP\<^sub>2 P\<^sub>1,D,M',0 \ compE\<^sub>2 body @ [Return]" using beforeM M'_in_D by simp have M'_xtab: "compP\<^sub>2 P\<^sub>1,D,M' \ compxE\<^sub>2 body 0 0/{..2 body)},0" using M'_in_D by(rule beforexM) have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3 (Class D # Ts) D M' 0 No_ics v xa [] ?frs\<^sub>1 ({..2 body)})" by fact have cond: "Jcc_cond P\<^sub>1 (Class D # Ts) D M' [] 0 No_ics {..2 body)} h\<^sub>2 sh\<^sub>2 body" using nsub_RI_Jcc_pieces[OF assms(1) nsub] M'_code M'_xtab by clarsimp show ?thesis (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>2' -jvm\ (None,h\<^sub>3,([v],ls\<^sub>3,D,M',size(compE\<^sub>2 body),No_ics)#?frs\<^sub>1,sh\<^sub>3)" using val IH_body SCallInit\<^sub>1.prems M'_code cond nsub_RI_Jcc_pieces nsub by auto also have "P \ \ -jvm\ (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>3)" using SCallInit\<^sub>1.hyps(8) M'_code M'_in_D ls nclinit' by(cases T, auto) finally show ?trans using nclinit by(auto simp:add.assoc) qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw with IH_body obtain pc\<^sub>2 vs' where pc\<^sub>2: "0 \ pc\<^sub>2 \ pc\<^sub>2 < size(compE\<^sub>2 body) \ \ caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and 2: "P \ ?\\<^sub>2' -jvm\ handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>1 sh\<^sub>3" using SCallInit\<^sub>1.prems M'_code M'_xtab cond nsub_RI_Jcc_pieces nsub by (auto simp del:split_paired_Ex) have "handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>1 sh\<^sub>3 = handle P C M xa h\<^sub>3 (rev pvs @ vs) ls\<^sub>2 ?pc\<^sub>1 ics frs sh\<^sub>3" using pc\<^sub>2 M'_in_D ls nclinit' by(auto simp add:handle_def) then show "?err" using pc\<^sub>2 jvm_trans[OF 1 2] nclinit by(auto intro!:exI[where x="?pc\<^sub>1"] exI[where x="rev pvs"]) qed qed qed next case (SCall\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' M' Ts T body D sfs ls\<^sub>2' e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3) show ?case proof(cases "M' = clinit \ es = []") case clinit: True then have s1: "pvs = []" "h\<^sub>1 = h\<^sub>2" "ls\<^sub>1 = ls\<^sub>2" "sh\<^sub>1 = sh\<^sub>2" using SCall\<^sub>1.hyps(1) evals\<^sub>1_cases(1) by blast+ then have ls\<^sub>2': "ls\<^sub>2' = replicate (max_vars body) undefined" using SCall\<^sub>1.hyps(6) clinit by simp let ?frs = "create_init_frame P C' # (vs, ls\<^sub>1, C,M,pc,ics)#frs" let ?\\<^sub>1 = "(None,h\<^sub>1,?frs,sh\<^sub>1)" have method: "P\<^sub>1 \ C' sees clinit,Static: []\Void = body in C'" using SCall\<^sub>1.hyps(3) clinit s1(1) wf_sees_clinit[OF wf] by (metis is_class_def option.collapse sees_method_fun sees_method_is_class) then have M_code: "compP\<^sub>2 P\<^sub>1,C',clinit,0 \ compE\<^sub>2 body @ [Return]" by(rule beforeM) have pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1 I h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa (C'\\<^sub>sclinit([])) = (True, ?frs, (None, h\<^sub>3, tl ?frs, sh\<^sub>3(C'\(fst(the(sh\<^sub>3 C')),Done))), P \ (None, h\<^sub>1, ?frs, sh\<^sub>1) -jvm\ (case ics of Called Cs \ (None, h\<^sub>3, (vs, ls\<^sub>1, C, M, pc, Throwing Cs xa) # frs, sh\<^sub>3(C' \ (fst (the (sh\<^sub>3 C')), Error)))))" using Jcc_pieces_clinit[OF assms(1),of E C M vs pc ics I h\<^sub>1 sh\<^sub>1 C' ls\<^sub>1 frs h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa] SCall\<^sub>1.prems(1) clinit s1(1) by clarsimp have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3 [] C' clinit 0 No_ics v xa [] (tl ?frs) ({..2 body)})" by fact show ?thesis (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val then have "P \ ?\\<^sub>1 -jvm\ (None, h\<^sub>3, ([v], ls\<^sub>3, C', clinit, size(compE\<^sub>2 body), No_ics) # tl ?frs,sh\<^sub>3)" using IH_body Jcc_pieces_SCall_clinit_body[OF assms(1) wf pcs method] s1 ls\<^sub>2' by clarsimp also have "P \ \ -jvm\ (None, h\<^sub>3, tl ?frs, sh\<^sub>3(C'\(fst(the(sh\<^sub>3 C')),Done)))" using jvm_Return_Init[OF M_code] by simp finally show ?trans using pcs s1 clinit by simp qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw with IH_body obtain pc\<^sub>2 vs2 where pc\<^sub>2: "0 \ pc\<^sub>2 \ pc\<^sub>2 < size(compE\<^sub>2 body) \ \ caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C' clinit xa h\<^sub>3 vs2 ls\<^sub>3 pc\<^sub>2 No_ics (tl ?frs) sh\<^sub>3" using SCall\<^sub>1.prems Jcc_pieces_SCall_clinit_body[OF assms(1) wf pcs method] s1 ls\<^sub>2' by clarsimp show ?err using SCall\<^sub>1.prems(1) clinit proof(cases ics) case (Called Cs) note 2 also have "handle P C' clinit xa h\<^sub>3 vs2 ls\<^sub>3 pc\<^sub>2 No_ics (tl ?frs) sh\<^sub>3 = (None, h\<^sub>3, (vs, ls\<^sub>1, C, M, pc, Throwing (C'#Cs) xa) # frs, sh\<^sub>3)" using Called pc\<^sub>2 method by(simp add: handle_def) also have "P \ \ -jvm\ (None, h\<^sub>3, (vs, ls\<^sub>1, C, M, pc, Throwing Cs xa) # frs, sh\<^sub>3(C' \ (fst (the (sh\<^sub>3 C')), Error)))" using Called jvm_Throwing by simp finally show ?thesis using pcs clinit Called by(clarsimp intro!: exI[where x="[]"]) qed(auto) qed qed next case nclinit: False let ?\\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)" let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)" let ?frs\<^sub>2 = "(rev pvs @ vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs" let ?\\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)" let ?frs\<^sub>2' = "([], ls\<^sub>2', D,M',0,No_ics) # ?frs\<^sub>2" let ?\\<^sub>2' = "(None, h\<^sub>2, ?frs\<^sub>2', sh\<^sub>2)" have nclinit': "M' \ clinit" using wf_sees_clinit1[OF wf] visible_method_exists[OF SCall\<^sub>1.hyps(3)] sees_method_idemp[OF SCall\<^sub>1.hyps(3)] nclinit SCall\<^sub>1.hyps(5) evals\<^sub>1_preserves_elen[OF SCall\<^sub>1.hyps(1)] by fastforce have "P\<^sub>1 \\<^sub>1 \es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\ [\] \map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\" by fact hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen) have invoke: "P,C,M,?pc\<^sub>2 \ Invokestatic C' M' (length Ts)" using SCall\<^sub>1.hyps(5) SCall\<^sub>1.prems nclinit by(auto simp: add.assoc) have nsub: "\ sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf SCall\<^sub>1.hyps(3)]) have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa (map Val pvs) vs frs I" by fact have "P \ ?\\<^sub>1 -jvm\ ?\\<^sub>2" using IH_es SCall\<^sub>1.prems nclinit by auto fastforce+ also have "P \ \ -jvm\ ?\\<^sub>2'" using jvm_Invokestatic[OF assms(1) invoke _ SCall\<^sub>1.hyps(3,5,6)] SCall\<^sub>1.hyps(4) SCall\<^sub>1.prems nclinit by auto finally have 1: "P \ ?\\<^sub>1 -jvm\ ?\\<^sub>2'". have "P\<^sub>1 \ C' sees M',Static: Ts\T = body in D" by fact then have M'_in_D: "P\<^sub>1 \ D sees M',Static: Ts\T = body in D" by(rule sees_method_idemp) have M'_code: "compP\<^sub>2 P\<^sub>1,D,M',0 \ compE\<^sub>2 body @ [Return]" using beforeM M'_in_D by simp have M'_xtab: "compP\<^sub>2 P\<^sub>1,D,M' \ compxE\<^sub>2 body 0 0/{..2 body)},0" using M'_in_D by(rule beforexM) have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3 (Class D # Ts) D M' 0 No_ics v xa [] ?frs\<^sub>2 ({..2 body)})" by fact have cond: "Jcc_cond P\<^sub>1 (Class D # Ts) D M' [] 0 No_ics {..2 body)} h\<^sub>2 sh\<^sub>2 body" using nsub_RI_Jcc_pieces[OF assms(1) nsub] M'_code M'_xtab by clarsimp show ?thesis (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>2' -jvm\ (None,h\<^sub>3,([v],ls\<^sub>3,D,M',size(compE\<^sub>2 body),No_ics)#?frs\<^sub>2,sh\<^sub>3)" using val IH_body SCall\<^sub>1.prems M'_code cond nsub_RI_Jcc_pieces nsub by auto also have "P \ \ -jvm\ (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>3)" using SCall\<^sub>1.hyps(5) M'_code M'_in_D nclinit' by(cases T, auto) finally show ?trans using nclinit by(auto simp:add.assoc) qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw with IH_body obtain pc\<^sub>2 vs' where pc\<^sub>2: "0 \ pc\<^sub>2 \ pc\<^sub>2 < size(compE\<^sub>2 body) \ \ caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and 2: "P \ ?\\<^sub>2' -jvm\ handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3" using SCall\<^sub>1.prems M'_code M'_xtab cond nsub_RI_Jcc_pieces nsub by (auto simp del:split_paired_Ex) have "handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3 = handle P C M xa h\<^sub>3 (rev pvs @ vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>3" using pc\<^sub>2 M'_in_D nclinit' by(auto simp add:handle_def) then show "?err" using pc\<^sub>2 jvm_trans[OF 1 2] nclinit by(auto intro:exI[where x="?pc\<^sub>2"]) qed qed qed next case Block\<^sub>1 then show ?case using nsub_RI_Jcc_pieces by auto next case (Seq\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>1)" let ?I = "I - pcs (compxE\<^sub>2 e\<^sub>2 (Suc ?pc\<^sub>1) (length vs))" have Isub: "{pc..2 e\<^sub>1)} \ ?I" using Seq\<^sub>1.prems by clarsimp have IH: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs ?I" by fact have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using Seq\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by auto also have "P \ \ -jvm\ ?\\<^sub>1" using Seq\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". let ?pc\<^sub>2 = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>2)" let ?I' = "I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs))" have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (?pc\<^sub>1+1) ics v xa vs frs ?I'" by fact have Isub2: "{Suc (pc + length (compE\<^sub>2 e\<^sub>1))..2 e\<^sub>1) + length (compE\<^sub>2 e\<^sub>2))} \ ?I'" using Seq\<^sub>1.prems by clarsimp show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note eval\<^sub>1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using val Seq\<^sub>1.prems nsub_RI_Jcc_pieces IH\<^sub>2 Isub2 by auto finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw then obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc\<^sub>1+1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2 \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 (?pc\<^sub>1+1) (size vs))" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using IH\<^sub>2 Seq\<^sub>1.prems nsub_RI_Jcc_pieces Isub2 by auto show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2]) qed qed next case (SeqThrow\<^sub>1 e\<^sub>0 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>1) let ?I = "I - pcs (compxE\<^sub>2 e\<^sub>1 (Suc (pc + length (compE\<^sub>2 e\<^sub>0))) (length vs))" obtain a' where throw: "throw e = Throw a'" using eval\<^sub>1_final[OF SeqThrow\<^sub>1.hyps(1)] by clarsimp have Isub: "{pc..2 e\<^sub>0)} \ ?I" using SeqThrow\<^sub>1.prems by clarsimp have "PROP ?P e\<^sub>0 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v a' vs frs ?I" by fact then show ?case using SeqThrow\<^sub>1.prems throw nsub_RI_Jcc_pieces Isub by auto next case (CondT\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>1 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>1)" let ?d = "size vs" let ?xt\<^sub>1 = "compxE\<^sub>2 e\<^sub>1 (pc+size(compE\<^sub>2 e)+1) ?d" let ?xt\<^sub>2 = "compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e)+size(compE\<^sub>2 e\<^sub>1)+2) ?d" let ?I = "I - (pcs ?xt\<^sub>1 \ pcs ?xt\<^sub>2)" have Isub: "{pc..2 e)} \ ?I" using CondT\<^sub>1.prems by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 true h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool True) xa vs frs ?I" by fact have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool(True)#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using CondT\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by(auto simp: Int_Un_distrib) also have "P \ \ -jvm\ ?\\<^sub>1" using CondT\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". let ?pc\<^sub>1' = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>2' = "?pc\<^sub>1' + 1 + length(compE\<^sub>2 e\<^sub>2)" let ?I' = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>2 (?pc\<^sub>1'+1) ?d)" have IH2: "PROP ?P e\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (?pc\<^sub>1+1) ics v xa vs frs ?I'" by fact show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note eval\<^sub>1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>1',ics)#frs,sh\<^sub>2)" using val CondT\<^sub>1.prems nsub_RI_Jcc_pieces IH2 by(fastforce simp:Int_Un_distrib) also have "P \ \ -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2',ics)#frs,sh\<^sub>2)" using CondT\<^sub>1 nsub_RI_Jcc_pieces by(auto simp:add.assoc) finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw moreover note IH2 ultimately obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc\<^sub>1+1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1' \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) (size vs))" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using CondT\<^sub>1.prems nsub_RI_Jcc_pieces by (fastforce simp:Int_Un_distrib) show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2]) qed qed next case (CondF\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>1) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?pc\<^sub>2 = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>1)+ 1" let ?pc\<^sub>2' = "?pc\<^sub>2 + length(compE\<^sub>2 e\<^sub>2)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>1)" let ?d = "size vs" let ?xt\<^sub>1 = "compxE\<^sub>2 e\<^sub>1 (pc+size(compE\<^sub>2 e)+1) ?d" let ?xt\<^sub>2 = "compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e)+size(compE\<^sub>2 e\<^sub>1)+2) ?d" let ?I = "I - (pcs ?xt\<^sub>1 \ pcs ?xt\<^sub>2)" let ?I' = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) ?d)" have pcs: "pcs(compxE\<^sub>2 e pc ?d) \ pcs(?xt\<^sub>1 @ ?xt\<^sub>2) = {}" using CondF\<^sub>1.prems by (simp add:Int_Un_distrib) have Isub: "{pc..2 e)} \ ?I" using CondF\<^sub>1.prems by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 false h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool False) xa vs frs ?I" by fact have IH2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>2 ics v xa vs frs ?I'" by fact have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool(False)#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using CondF\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub pcs by auto also have "P \ \ -jvm\ ?\\<^sub>1" using CondF\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note eval\<^sub>1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2',ics)#frs,sh\<^sub>2)" using val CondF\<^sub>1.prems nsub_RI_Jcc_pieces IH2 by(fastforce simp:Int_Un_distrib) finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ ?err") proof let ?I' = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) ?d)" assume throw: ?throw then obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc\<^sub>2 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2' \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>2 ?d)" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using CondF\<^sub>1.prems nsub_RI_Jcc_pieces IH2 by(fastforce simp:Int_Un_distrib) show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2]) qed qed next case (CondThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 f h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>1 e\<^sub>2) let ?d = "size vs" let ?xt\<^sub>1 = "compxE\<^sub>2 e\<^sub>1 (pc+size(compE\<^sub>2 e)+1) ?d" let ?xt\<^sub>2 = "compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e)+size(compE\<^sub>2 e\<^sub>1)+2) ?d" let ?I = "I - (pcs ?xt\<^sub>1 \ pcs ?xt\<^sub>2)" have Isub: "{pc..2 e)} \ ?I" using CondThrow\<^sub>1.prems by clarsimp have "pcs(compxE\<^sub>2 e pc ?d) \ pcs(?xt\<^sub>1 @ ?xt\<^sub>2) = {}" using CondThrow\<^sub>1.prems by (simp add:Int_Un_distrib) moreover have "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw f) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs ?I" by fact ultimately show ?case using CondThrow\<^sub>1.prems nsub_RI_Jcc_pieces Isub by auto next case (WhileF\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 c) let ?pc = "pc + length(compE\<^sub>2 e)" let ?pc' = "?pc + length(compE\<^sub>2 c) + 3" have Isub: "{pc..2 e)} \ I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))" using WhileF\<^sub>1.prems by clarsimp have Isub2: "{Suc (pc + length (compE\<^sub>2 e))..2 e) + length (compE\<^sub>2 c))} \ I - pcs (compxE\<^sub>2 e pc (length vs))" using WhileF\<^sub>1.prems by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 false h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool False) xa vs frs (I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs)))" by fact have "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(Bool False#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using WhileF\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by auto also have "P \ \ -jvm\ (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc',ics)#frs,sh\<^sub>1)" using WhileF\<^sub>1 by (auto simp:add.assoc) also have "P \ \ -jvm\ (None,h\<^sub>1,(Unit#vs,ls\<^sub>1,C,M,?pc'+1,ics)#frs,sh\<^sub>1)" using WhileF\<^sub>1.prems by (auto simp:eval_nat_numeral) finally show ?case by (simp add:add.assoc eval_nat_numeral) next case (WhileT\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 c v\<^sub>1 h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>3 h\<^sub>3 ls\<^sub>3 sh\<^sub>3) let ?pc = "pc + length(compE\<^sub>2 e)" let ?pc' = "?pc + length(compE\<^sub>2 c) + 1" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?\\<^sub>2 = "(None,h\<^sub>2,(vs,ls\<^sub>2,C,M,pc,ics)#frs,sh\<^sub>2)" have Isub: "{pc..2 e)} \ I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))" using WhileT\<^sub>1.prems by clarsimp have Isub2: "{Suc (pc + length (compE\<^sub>2 e))..2 e) + length (compE\<^sub>2 c))} \ I - pcs (compxE\<^sub>2 e pc (length vs))" using WhileT\<^sub>1.prems by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 true h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool True) xa vs frs (I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs)))" by fact have IH2: "PROP ?P c h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v\<^sub>1) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (Suc ?pc) ics v\<^sub>1 xa vs frs (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool True#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)" using WhileT\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by auto also have "P \ \ -jvm\ (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)" using WhileT\<^sub>1.prems by auto also have "P \ \ -jvm\ (None,h\<^sub>2,(v\<^sub>1#vs,ls\<^sub>2,C,M,?pc',ics)#frs,sh\<^sub>2)" using WhileT\<^sub>1.prems nsub_RI_Jcc_pieces IH2 Isub2 by auto also have "P \ \ -jvm\ ?\\<^sub>2" using WhileT\<^sub>1.prems by auto finally have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>2". show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>2 -jvm\ (None,h\<^sub>3,(v#vs,ls\<^sub>3,C,M,?pc'+3,ics)#frs,sh\<^sub>3)" using val WhileT\<^sub>1 by (auto simp add:add.assoc eval_nat_numeral) finally show ?trans by(simp add:add.assoc eval_nat_numeral) qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw moreover have "PROP ?P (while (e) c) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>3 h\<^sub>3 ls\<^sub>3 sh\<^sub>3 E C M pc ics v xa vs frs I" by fact ultimately obtain pc\<^sub>2 vs' where pc\<^sub>2: "pc \ pc\<^sub>2 \ pc\<^sub>2 < ?pc'+3 \ \ caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 (while (e) c) pc (size vs))" and 2: "P \ ?\\<^sub>2 -jvm\ handle P C M xa h\<^sub>3 (vs'@vs) ls\<^sub>3 pc\<^sub>2 ics frs sh\<^sub>3" using WhileT\<^sub>1.prems by (auto simp:add.assoc eval_nat_numeral) show "?err" using pc\<^sub>2 jvm_trans[OF 1 2] by(auto intro: exI[where x=pc\<^sub>2]) qed qed next case (WhileCondThrow\<^sub>1 e h ls sh e' h' ls' sh' c) let ?I = "I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))" obtain a' where throw: "throw e' = Throw a'" using eval\<^sub>1_final[OF WhileCondThrow\<^sub>1.hyps(1)] by clarsimp have Isub: "{pc..2 e)} \ ?I" using WhileCondThrow\<^sub>1.prems by clarsimp have "PROP ?P e h ls sh (throw e') h' ls' sh' E C M pc ics v a' vs frs ?I" by fact then show ?case using WhileCondThrow\<^sub>1.prems throw nsub_RI_Jcc_pieces Isub by auto next case (WhileBodyThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 c e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>1)" let ?I = "I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))" have Isub: "{pc..2 e)} \ ?I" using WhileBodyThrow\<^sub>1.prems by clarsimp have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 true h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool True) xa vs frs ?I" by fact then have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,(Bool(True)#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using WhileBodyThrow\<^sub>1.prems nsub_RI_Jcc_pieces Isub by auto also have "P \ \ -jvm\ ?\\<^sub>1" using WhileBodyThrow\<^sub>1 by auto finally have eval\<^sub>1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1". let ?pc\<^sub>1' = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 c)" show ?case (is "?Norm \ ?Err") proof show ?Norm by simp next show ?Err (is "?throw \ ?err") proof assume throw: ?throw moreover have "PROP ?P c h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (?pc\<^sub>1+1) ics v xa vs frs (I - pcs (compxE\<^sub>2 e pc (size vs)))" by fact ultimately obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc\<^sub>1+1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>1' \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 c (?pc\<^sub>1+1) (size vs))" and eval\<^sub>2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using WhileBodyThrow\<^sub>1.prems nsub_RI_Jcc_pieces by (fastforce simp:Int_Un_distrib) show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2]) qed qed next case (Throw\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1) let ?pc = "pc + size(compE\<^sub>2 e)" have Isub: "{pc..2 e)} \ I" using Throw\<^sub>1.prems by clarsimp show ?case (is "?Norm \ ?Err") proof show ?Norm by simp next show ?Err (is "?throw \ ?err") proof assume throw:?throw have "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) a vs frs I" by fact then have "P \ (None, h\<^sub>0, (vs, ls\<^sub>0, C, M, pc, ics) # frs, sh\<^sub>0) -jvm\ (None, h\<^sub>1, (Addr xa#vs, ls\<^sub>1, C, M, ?pc, ics) # frs, sh\<^sub>1)" using Throw\<^sub>1 nsub_RI_Jcc_pieces Isub throw by auto also have "P \ \ -jvm\ handle P C M xa h\<^sub>1 (Addr xa#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using Throw\<^sub>1.prems by(auto simp add:handle_def) finally show "?err" by(auto intro!: exI[where x="?pc"] exI[where x="[Addr xa]"]) qed qed next case (ThrowNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1) let ?pc = "pc + size(compE\<^sub>2 e)" let ?xa = "addr_of_sys_xcpt NullPointer" have Isub: "{pc..2 e)} \ I" using ThrowNull\<^sub>1.prems by clarsimp show ?case (is "?Norm \ ?Err") proof show ?Norm by simp next show ?Err (is "?throw \ ?err") proof assume throw: ?throw have "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs I" by fact then have "P \ (None, h\<^sub>0, (vs, ls\<^sub>0, C, M, pc, ics) # frs, sh\<^sub>0) -jvm\ (None, h\<^sub>1, (Null#vs, ls\<^sub>1, C, M, ?pc, ics) # frs, sh\<^sub>1)" using ThrowNull\<^sub>1.prems nsub_RI_Jcc_pieces Isub by auto also have "P \ \ -jvm\ handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1" using ThrowNull\<^sub>1.prems by(auto simp add:handle_def) finally show "?err" using throw by(auto intro!: exI[where x="?pc"] exI[where x="[Null]"]) qed qed next case (ThrowThrow\<^sub>1 e h ls sh e' h' ls' sh') obtain a' where throw: "throw e' = Throw a'" using eval\<^sub>1_final[OF ThrowThrow\<^sub>1.hyps(1)] by clarsimp have Isub: "{pc..2 e)} \ I" using ThrowThrow\<^sub>1.prems by clarsimp have "PROP ?P e h ls sh (throw e') h' ls' sh' E C M pc ics v a' vs frs I" by fact then show ?case using ThrowThrow\<^sub>1.prems throw nsub_RI_Jcc_pieces Isub by auto next case (Try\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 Ci i e\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>1' = "?pc\<^sub>1 + 2 + length(compE\<^sub>2 e\<^sub>2)" have "{pc..2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \ I" using Try\<^sub>1.prems by simp also have "P,C,M \ compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs) / I,size vs" using Try\<^sub>1.prems by simp ultimately have "P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs) / {pc..2 e\<^sub>1)},size vs" by(rule beforex_try) hence "P \ (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" using Try\<^sub>1 nsub_RI_Jcc_pieces by auto blast also have "P \ \ -jvm\ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1',ics)#frs,sh\<^sub>1)" using Try\<^sub>1.prems by auto finally show ?case by (auto simp:add.assoc) next case (TryCatch\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs Ci i e\<^sub>2 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2) let ?e = "try e\<^sub>1 catch(Ci i) e\<^sub>2" let ?xt = "compxE\<^sub>2 ?e pc (size vs)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?ls\<^sub>1 = "ls\<^sub>1[i := Addr a]" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?pc\<^sub>1' = "?pc\<^sub>1 + 2" let ?\\<^sub>1 = "(None,h\<^sub>1,(vs,?ls\<^sub>1,C,M, ?pc\<^sub>1',ics) # frs,sh\<^sub>1)" have I: "{pc..2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \ I" and beforex: "P,C,M \ ?xt/I,size vs" using TryCatch\<^sub>1.prems by simp+ have "P \ ?\\<^sub>0 -jvm\ (None,h\<^sub>1,((Addr a)#vs,ls\<^sub>1,C,M, ?pc\<^sub>1+1,ics) # frs,sh\<^sub>1)" proof - have ics: "ics = No_ics" using TryCatch\<^sub>1.prems by auto have "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Throw a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v a vs frs {pc..2 e\<^sub>1)}" by fact moreover have "P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs)/{pc..1},size vs" using beforex I pcs_subset by(force elim!: beforex_appendD1) ultimately have "\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < ?pc\<^sub>1 \ \ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs)) \ (\vs'. P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)" using TryCatch\<^sub>1.prems nsub_RI_Jcc_pieces by auto then obtain pc\<^sub>1 vs' where pc\<^sub>1_in_e\<^sub>1: "pc \ pc\<^sub>1" "pc\<^sub>1 < ?pc\<^sub>1" and pc\<^sub>1_not_caught: "\ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs))" and 0: "P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1" by iprover from beforex obtain xt\<^sub>0 xt\<^sub>1 where ex_tab: "ex_table_of P C M = xt\<^sub>0 @ ?xt @ xt\<^sub>1" and disj: "pcs xt\<^sub>0 \ I = {}" by(auto simp:beforex_def) have hp: "h\<^sub>1 a = Some (D, fs)" "P\<^sub>1 \ D \\<^sup>* Ci" by fact+ have "pc\<^sub>1 \ pcs xt\<^sub>0" using pc\<^sub>1_in_e\<^sub>1 I disj by auto with pc\<^sub>1_in_e\<^sub>1 pc\<^sub>1_not_caught hp show ?thesis using ex_tab 0 ics by(simp add:handle_def matches_ex_entry_def) qed also have "P \ \ -jvm\ ?\\<^sub>1" using TryCatch\<^sub>1 by auto finally have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" . let ?pc\<^sub>2 = "?pc\<^sub>1' + length(compE\<^sub>2 e\<^sub>2)" let ?I\<^sub>2 = "{?pc\<^sub>1' ..< ?pc\<^sub>2}" have "P,C,M \ compxE\<^sub>2 ?e pc (size vs) / I,size vs" by fact hence beforex\<^sub>2: "P,C,M \ compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1' (size vs) / ?I\<^sub>2, size vs" using I pcs_subset[of _ ?pc\<^sub>1'] by(auto elim!:beforex_appendD2) have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ?ls\<^sub>1 sh\<^sub>1 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1' ics v xa vs frs ?I\<^sub>2" by fact show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using val beforex\<^sub>2 IH\<^sub>2 TryCatch\<^sub>1.prems nsub_RI_Jcc_pieces by auto finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ ?err") proof assume throw: ?throw then obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc\<^sub>1+2 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2 \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1' (size vs))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using IH\<^sub>2 beforex\<^sub>2 TryCatch\<^sub>1.prems nsub_RI_Jcc_pieces by auto show ?err using pc\<^sub>2 jvm_trans[OF 1 2] by (simp add:match_ex_entry) (auto intro: exI[where x=pc\<^sub>2]) qed qed next case (TryThrow\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs Ci i e\<^sub>2) let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)" let ?e = "try e\<^sub>1 catch(Ci i) e\<^sub>2" let ?xt = "compxE\<^sub>2 ?e pc (size vs)" have I: "{pc..2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \ I" and beforex: "P,C,M \ ?xt/I,size vs" using TryThrow\<^sub>1.prems by simp+ have "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Throw a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v a vs frs {pc..2 e\<^sub>1)}" by fact moreover have "P,C,M \ compxE\<^sub>2 e\<^sub>1 pc (size vs)/{pc..1},size vs" using beforex I pcs_subset by(force elim!: beforex_appendD1) ultimately have "\pc\<^sub>1. pc \ pc\<^sub>1 \ pc\<^sub>1 < ?pc\<^sub>1 \ \ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs)) \ (\vs'. P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)" using TryThrow\<^sub>1.prems nsub_RI_Jcc_pieces by auto then obtain pc\<^sub>1 vs' where pc\<^sub>1_in_e\<^sub>1: "pc \ pc\<^sub>1" "pc\<^sub>1 < ?pc\<^sub>1" and pc\<^sub>1_not_caught: "\ caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs))" and 0: "P \ ?\\<^sub>0 -jvm\ handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1" by iprover show ?case (is "?N \ (?eq \ ?err)") proof show ?N by simp next { assume ?eq with TryThrow\<^sub>1 pc\<^sub>1_in_e\<^sub>1 pc\<^sub>1_not_caught 0 have "?err" by (simp add:match_ex_entry) auto } thus "?eq \ ?err" by iprover qed next case Nil\<^sub>1 thus ?case by simp next case (Cons\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es fs h\<^sub>2 ls\<^sub>2 sh\<^sub>2) let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)" let ?\\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)" let ?\\<^sub>1 = "(None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)" have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val v) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 [] C M pc ics v xa vs frs (I - pcs (compxEs\<^sub>2 es ?pc\<^sub>1 (Suc (length vs))))" by fact then have 1: "P \ ?\\<^sub>0 -jvm\ ?\\<^sub>1" using Jcc_pieces_Cons[OF _ Cons\<^sub>1.prems(1-5)] by auto let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)" have IHs: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 fs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics (tl ws) xa es' (v#vs) frs (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact show ?case (is "?Norm \ ?Err") proof show ?Norm (is "?val \ ?trans") proof assume val: ?val note 1 also have "P \ ?\\<^sub>1 -jvm\ (None,h\<^sub>2,(rev(ws) @ vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)" using val IHs Cons\<^sub>1.prems by fastforce finally show ?trans by(simp add:add.assoc) qed next show ?Err (is "?throw \ (\pc\<^sub>2. ?H pc\<^sub>2)") proof assume throw: ?throw then obtain pc\<^sub>2 vs' where pc\<^sub>2: "?pc\<^sub>1 \ pc\<^sub>2 \ pc\<^sub>2 < ?pc\<^sub>2 \ \ caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es ?pc\<^sub>1 (size vs + 1))" and 2: "P \ ?\\<^sub>1 -jvm\ handle P C M xa h\<^sub>2 (vs'@v#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2" using IHs Cons\<^sub>1.prems by(fastforce simp:Cons_eq_append_conv neq_Nil_conv) have "?H pc\<^sub>2" using Cons\<^sub>1.prems pc\<^sub>2 jvm_trans[OF 1 2] by(auto intro!: exI[where x="vs'@[v]"]) thus "\pc\<^sub>2. ?H pc\<^sub>2" by iprover qed qed next case (ConsThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es) then show ?case using Jcc_pieces_Cons[OF _ ConsThrow\<^sub>1.prems(1-5)] by (fastforce simp:Cons_eq_append_conv) next case InitFinal\<^sub>1 then show ?case using eval\<^sub>1_final_same[OF InitFinal\<^sub>1.hyps(1)] by clarsimp next case (InitNone\<^sub>1 sh C\<^sub>0 C' Cs e h l e' h' l' sh') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" using InitNone\<^sub>1.prems(1) by clarsimp let ?sh = "(sh(C\<^sub>0 \ (sblank P\<^sub>1 C\<^sub>0, Prepared)))" obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs" and frs\<^sub>1: "frs' \ Nil" using pcs by clarsimp then have 1: "P \ (None,h,frs',sh) -jvm\ (None,h,frs',?sh)" using InitNone\<^sub>1 jvm_InitNone[where P = P] by(cases frs', simp+) show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 note 1 also have "P \ (None,h,frs',?sh) -jvm\ (None,h',(vs,l,C,M,pc,Called [])#frs,sh')" using InitNone\<^sub>1.hyps(3)[of E] Jcc_pieces_InitNone[OF assms(1) pcs] InitNone\<^sub>1.prems val by clarsimp finally have ?jvm1 using pcs by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 note 1 also obtain vs' where "P \ (None,h,frs',?sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'" using InitNone\<^sub>1.hyps(3)[of E] Jcc_pieces_InitNone[OF assms(1) pcs] throw by clarsimp presburger finally have ?err using pcs by auto } thus "?e2 \ ?err" by simp qed next case (InitDone\<^sub>1 sh C\<^sub>0 sfs C' Cs e h l e' h' l' sh') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" using InitDone\<^sub>1.prems(1) by clarsimp let ?frs' = "(calling_to_scalled (hd frs'))#(tl frs')" have IH: "PROP ?P (INIT C' (Cs,True) \ e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I" by fact obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs" and frs\<^sub>1: "frs' \ Nil" using pcs by clarsimp then have 1: "P \ (None,h,frs',sh) -jvm\ (None,h,?frs',sh)" using InitDone\<^sub>1 jvm_InitDP[where P = P] by(cases frs', simp+) show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 note 1 also have "P \ (None,h,?frs',sh) -jvm\ (None,h',(vs,l,C,M,pc,Called [])#frs,sh')" using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitDone\<^sub>1.prems val by clarsimp finally have ?jvm1 using pcs by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 note 1 also obtain vs' where "P \ (None,h,?frs',sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'" using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitDone\<^sub>1.prems throw by clarsimp finally have ?err using pcs by auto } thus "?e2 \ ?err" by simp qed next case (InitProcessing\<^sub>1 sh C\<^sub>0 sfs C' Cs e h l e' h' l' sh') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" using InitProcessing\<^sub>1.prems(1) by clarsimp let ?frs' = "(calling_to_scalled (hd frs'))#(tl frs')" have IH: "PROP ?P (INIT C' (Cs,True) \ e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I" by fact obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs" and frs\<^sub>1: "frs' \ Nil" using pcs by clarsimp then have 1: "P \ (None,h,frs',sh) -jvm\ (None,h,?frs',sh)" using InitProcessing\<^sub>1 jvm_InitDP[where P = P] by(cases frs', simp+) show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 note 1 also have "P \ (None,h,?frs',sh) -jvm\ (None,h',(vs,l,C,M,pc,Called [])#frs,sh')" using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitProcessing\<^sub>1.prems val by clarsimp finally have ?jvm1 using pcs by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 note 1 also obtain vs' where "P \ (None,h,?frs',sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'" using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitProcessing\<^sub>1.prems throw by clarsimp finally have ?err using pcs by auto } thus "?e2 \ ?err" by simp qed next case (InitError\<^sub>1 sh C\<^sub>0 sfs Cs e h l e' h' l' sh' C') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" using InitError\<^sub>1.prems(1) by clarsimp let ?e\<^sub>0 = "THROW NoClassDefFoundError" let ?frs' = "(calling_to_sthrowing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs')" have IH: "PROP ?P (RI (C\<^sub>0,?e\<^sub>0) ; Cs \ e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I" by fact obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs" and frs\<^sub>1: "frs' \ Nil" and tl: "tl frs' = frs" using pcs by clarsimp then have 1: "P \ (None,h,frs',sh) -jvm\ (None,h,?frs',sh)" proof(cases frs') case (Cons a list) obtain vs' l' C' M' pc' ics' where a: "a = (vs',l',C',M',pc',ics')" by(cases a) then have "ics' = Calling C\<^sub>0 Cs" using Cons ics by simp then show ?thesis using Cons a IH InitError\<^sub>1.prems jvm_InitError[where P = P] InitError\<^sub>1.hyps(1) by simp qed(simp) show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 then have False using val rinit\<^sub>1_throw[OF InitError\<^sub>1.hyps(2)] by blast then have ?jvm1 using pcs by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 let ?frs = "(calling_to_throwing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs')" have exec: "exec (P, (None,h,?frs,sh)) = Some (None,h,?frs',sh)" using exec_ErrorThrowing[where sh=sh, OF InitError\<^sub>1.hyps(1)] ics by(cases "hd frs'", simp) obtain vs' where 2: "P \ (None,h,?frs,sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'" using IH Jcc_pieces_InitError[OF assms(1) pcs InitError\<^sub>1.hyps(1)] throw by clarsimp have neq: "(None, h, ?frs, sh) \ handle P C M xa h' (vs' @ vs) l pc ics frs sh'" using tl ics by(cases "hd frs'", simp add: handle_frs_tl_neq) note 1 also have "P \ (None,h,?frs',sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'" using exec_1_exec_all_conf[OF exec 2] neq by simp finally have ?err using pcs by auto } thus "?e2 \ ?err" by simp qed next case (InitObject\<^sub>1 sh C\<^sub>0 sfs sh' C' Cs e h l e' h' l' sh'') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' (sh(C\<^sub>0 \ (sfs, Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" using InitObject\<^sub>1.prems(1) by clarsimp let ?frs' = "(calling_to_called (hd frs'))#(tl frs')" have IH: "PROP ?P (INIT C' (C\<^sub>0#Cs,True) \ e) h l sh' e' h' l' sh'' E C M pc ics v xa vs frs I" by fact obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs" and frs\<^sub>1: "frs' \ Nil" using pcs by clarsimp then have 1: "P \ (None,h,frs',sh) -jvm\ (None,h,?frs',sh')" proof(cases frs') case (Cons a list) obtain vs' l' C' M' pc' ics' where a: "a = (vs',l',C',M',pc',ics')" by(cases a) then have "ics' = Calling C\<^sub>0 Cs" using Cons ics by simp then show ?thesis using Cons Nil a IH InitObject\<^sub>1 jvm_InitObj[where P = P] by simp qed(simp) show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 note 1 also have "P \ (None,h,?frs',sh') -jvm\ (None,h',(vs,l,C,M,pc,Called [])#frs,sh'')" using IH Jcc_pieces_InitObj[OF assms(1) pcs] InitObject\<^sub>1 val by simp finally have ?jvm1 using pcs by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 note 1 also obtain vs' where "P \ (None,h,?frs',sh') -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh''" using IH Jcc_pieces_InitObj[OF assms(1) pcs] InitObject\<^sub>1 throw by clarsimp finally have ?err using pcs by auto } thus "?e2 \ ?err" by simp qed next case (InitNonObject\<^sub>1 sh C\<^sub>0 sfs D a b sh' C' Cs e h l e' h' l' sh'') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' (sh(C\<^sub>0 \ (sfs,Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" using InitNonObject\<^sub>1.prems(1) by clarsimp let ?frs' = "(calling_to_calling (hd frs') D)#(tl frs')" have cls1: "is_class P\<^sub>1 D" using InitNonObject\<^sub>1.hyps(2,3) class_wf wf wf_cdecl_supD by blast have cls_aux: "distinct (C\<^sub>0#Cs) \ supercls_lst P\<^sub>1 (C\<^sub>0#Cs)" using InitNonObject\<^sub>1.prems(1) by auto then have cls2: "D \ set (C\<^sub>0 # Cs)" proof - have "distinct (D # C\<^sub>0 # Cs)" using InitNonObject\<^sub>1.hyps(2,3) cls_aux wf wf_supercls_distinct_app by blast then show "D \ set (C\<^sub>0 # Cs)" by (metis distinct.simps(2)) qed have cls3: "\C\set (C\<^sub>0 # Cs). P\<^sub>1 \ C \\<^sup>* D" using InitNonObject\<^sub>1.hyps(2,3) cls_aux by (metis r_into_rtrancl rtrancl_into_rtrancl set_ConsD subcls1.subcls1I supercls_lst.simps(1)) have IH: "PROP ?P (INIT C' (D # C\<^sub>0 # Cs,False) \ e) h l sh' e' h' l' sh'' E C M pc ics v xa vs frs I" by fact obtain r where cls: "class P C\<^sub>0 = \(D, r)\" using InitNonObject\<^sub>1.hyps(3) by (metis assms class_compP compP\<^sub>2_def) obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs" and frs\<^sub>1: "frs' \ Nil" using pcs by clarsimp then have 1: "P \ (None,h,frs',sh) -jvm\ (None,h,?frs',sh')" proof(cases frs') case (Cons a list) obtain vs' l' C' M' pc' ics' where a: "a = (vs',l',C',M',pc',ics')" by(cases a) then have "ics' = Calling C\<^sub>0 Cs" using Cons ics by simp then show ?thesis using Cons a IH InitNonObject\<^sub>1 jvm_InitNonObj[OF _ _ cls] by simp qed(simp) show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 note 1 also have "P \ (None,h,?frs',sh') -jvm\ (None,h',(vs,l,C,M,pc,Called [])#frs,sh'')" using IH Jcc_pieces_InitNonObj[OF assms(1) cls1 cls2 cls3 pcs] InitNonObject\<^sub>1 val by simp finally have ?jvm1 using pcs by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 note 1 also obtain vs' where "P \ (None,h,?frs',sh') -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh''" using IH Jcc_pieces_InitNonObj[OF assms(1) cls1 cls2 cls3 pcs] InitNonObject\<^sub>1 throw by clarsimp finally have ?err using pcs by auto } thus "?e2 \ ?err" by simp qed next case (InitRInit\<^sub>1 C\<^sub>0 Cs e h l sh e' h' l' sh' C') then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,True) \ e) = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)" using InitRInit\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; Cs \ e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I" by fact show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 have "P \ (None,h,frs',sh) -jvm\ (None,h',(vs,l,C,M,pc,Called [])#frs,sh')" using IH Jcc_pieces_InitRInit[OF assms(1,2) pcs] InitRInit\<^sub>1.prems val by simp then have ?jvm1 using pcs by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 obtain vs' where "P \ (None,h,frs',sh) -jvm\ handle P C M xa h' (vs'@vs) l pc ics frs sh'" using IH Jcc_pieces_InitRInit[OF assms(1,2) pcs] InitRInit\<^sub>1 throw by clarsimp then have ?err using pcs by auto } thus "?e2 \ ?err" by simp qed next case (RInit\<^sub>1 e h l sh v1 h' l' sh' C\<^sub>0 sfs i sh'' C' Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1) let ?frs = "(vs,l,C,M,pc,Called (C\<^sub>0#Cs)) # frs" let ?frs' = "(vs,l,C,M,pc,Called Cs) # frs" have clinit: "e = C\<^sub>0\\<^sub>sclinit([])" using RInit\<^sub>1 by (metis Jcc_cond.simps(2) eval\<^sub>1_final_same exp.distinct(101) final_def) then obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; Cs \ e') = (True, ?frs, (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)" using RInit\<^sub>1.prems(1) by simp have shC: "\C'\set Cs. \sfs. sh C' = \(sfs, Processing)\" using RInit\<^sub>1.prems(1) clinit by clarsimp then have shC'': "\C'\set Cs. \sfs. sh'' C' = \(sfs, Processing)\" using clinit\<^sub>1_proc_pres[OF wf] RInit\<^sub>1.hyps(1) clinit RInit\<^sub>1.hyps(4) RInit\<^sub>1.prems(1) by (auto simp: fun_upd_apply) have loc: "l = l'" using clinit\<^sub>1_loc_pres RInit\<^sub>1.hyps(1) clinit by simp have IH: "PROP ?P e h l sh (Val v1) h' l' sh' E C M pc (Called Cs) v1 xa vs (tl ?frs') I" by fact then have IH': "PROP ?P (C\<^sub>0\\<^sub>sclinit([])) h l sh (Val v1) h' l' sh' E C M pc (Called Cs) v1 xa vs (tl ?frs') I" using clinit by simp have IH2: "PROP ?P (INIT C' (Cs,True) \ e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact have "P \ (None,h,?frs,sh) -jvm\ (None,h,create_init_frame P C\<^sub>0 # ?frs',sh)" by(rule jvm_Called) also have "P \ \ -jvm\ (None,h',?frs',sh'')" using IH' Jcc_pieces_RInit_clinit[OF assms(1-2) pcs,of h' l' sh'] RInit\<^sub>1.hyps(3,4) by simp finally have jvm1: "P \ (None,h,?frs,sh) -jvm\ (None,h',?frs',sh'')" . show ?case (is "(?e1 \ ?jvm1) \ (?e2 \ ?err)") proof(rule conjI) { assume val: ?e1 note jvm1 also have "P \ (None,h',?frs',sh'') -jvm\ (None,h\<^sub>1,(vs,l,C,M,pc,Called [])#frs,sh\<^sub>1)" using IH2 Jcc_pieces_RInit_Init[OF assms(1-2) shC'' pcs,of h'] RInit\<^sub>1.hyps(5) loc val by auto finally have ?jvm1 using pcs clinit by simp } thus "?e1 \ ?jvm1" by simp next { assume throw: ?e2 note jvm1 also obtain vs' where "P \ (None,h',?frs',sh'') -jvm\ handle P C M xa h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1" using IH2 Jcc_pieces_RInit_Init[OF assms(1-2) shC'' pcs,of h'] RInit\<^sub>1.hyps(5) loc throw by auto finally have ?err using pcs clinit by auto } thus "?e2 \ ?err" by simp qed next case (RInitInitFail\<^sub>1 e h l sh a h' l' sh' C\<^sub>0 sfs i sh'' D Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1) let ?frs = "(vs,l,C,M,pc,Called (C\<^sub>0#D#Cs)) # frs" let ?frs' = "(vs,l,C,M,pc,Called (D#Cs)) # frs" let "?frsT" = "\xa1. (vs,l,C,M,pc,Throwing (C\<^sub>0#D#Cs) xa1) # frs" let "?frsT'" = "\xa1. (vs,l,C,M,pc,Throwing (D#Cs) xa1) # frs" obtain xa' where xa': "throw a = Throw xa'" by (metis RInitInitFail\<^sub>1.hyps(1) eval\<^sub>1_final exp.distinct(101) final_def) have e\<^sub>1: "e\<^sub>1 = Throw xa'" using xa' rinit\<^sub>1_throw RInitInitFail\<^sub>1.hyps(5) by simp show ?case proof(cases "e = C\<^sub>0\\<^sub>sclinit([])") case clinit: True then obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa' (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; D # Cs \ e') = (True, ?frs, (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)" using RInitInitFail\<^sub>1.prems(1) by simp have loc: "l = l'" using clinit\<^sub>1_loc_pres RInitInitFail\<^sub>1.hyps(1) clinit by simp have IH: "PROP ?P e h l sh (throw a) h' l' sh' E C M pc (Called (D#Cs)) v xa' vs frs I" by fact then have IH': "PROP ?P (C\<^sub>0\\<^sub>sclinit([])) h l sh (Throw xa') h' l' sh' E C M pc (Called (D#Cs)) v xa' vs frs I" using clinit xa' by simp have IH2: "PROP ?P (RI (D,throw a) ; Cs \ e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 E C M pc ics v xa' vs frs I" by fact have "P \ (None,h,?frs,sh) -jvm\ (None,h,create_init_frame P C\<^sub>0 # ?frs',sh)" by(rule jvm_Called) also have "P \ \ -jvm\ (None,h',(vs, l, C, M, pc, Throwing (D#Cs) xa') # frs,sh'')" using IH' Jcc_pieces_RInit_clinit[OF assms(1-2) pcs,of h' l' sh'] RInitInitFail\<^sub>1.hyps(3,4) by simp also obtain vs'' where "P \ \ -jvm\ handle P C M xa' h\<^sub>1 (vs''@vs) l pc ics frs sh\<^sub>1" using IH2 pcs Jcc_pieces_RInit_RInit[OF assms(1) pcs] RInitInitFail\<^sub>1.hyps(3,4) xa' loc e\<^sub>1 xa' by clarsimp finally show ?thesis using pcs e\<^sub>1 clinit by auto next case throw: False then have eT: "e = Throw xa'" "h = h'" "l = l'" "sh = sh'" using xa' RInitInitFail\<^sub>1.prems(1) eval\<^sub>1_final_same[OF RInitInitFail\<^sub>1.hyps(1)] by clarsimp+ obtain a' where "class P\<^sub>1 C\<^sub>0 = \a'\" using RInitInitFail\<^sub>1.prems by(auto simp: is_class_def) then obtain stk' loc' M' pc' ics' where "create_init_frame P C\<^sub>0 = (stk',loc',C\<^sub>0,M',pc',ics')" using create_init_frame_wf_eq[OF wf] by(cases "create_init_frame P C\<^sub>0", simp) then obtain rhs err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh'' v xa' (RI (C\<^sub>0,e) ; D#Cs \ e') = (True, ?frsT xa', rhs, err)" using RInitInitFail\<^sub>1.prems(1) eT by clarsimp have IH2: "PROP ?P (RI (D,throw a) ; Cs \ e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 E C M pc ics v xa' vs frs I" by fact have "P \ (None,h,?frsT xa',sh') -jvm\ (None,h,?frsT' xa',sh'(C\<^sub>0 \ (fst (the (sh' C\<^sub>0)), Error)))" by(rule jvm_Throwing) also obtain vs' where "P \ ... -jvm\ handle P C M xa' h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1" using IH2 Jcc_pieces_RInit_RInit[OF assms(1) pcs] RInitInitFail\<^sub>1.hyps(3,4) eT e\<^sub>1 xa' by clarsimp finally show ?thesis using pcs e\<^sub>1 throw eT by auto qed next case (RInitFailFinal\<^sub>1 e h l sh a h' l' sh' C\<^sub>0 sfs i sh'' e'') let ?frs = "(vs,l,C,M,pc,Called [C\<^sub>0]) # frs" let ?frs' = "(vs,l,C,M,pc,Called []) # frs" let "?frsT" = "\xa1. (vs,l,C,M,pc,Throwing [C\<^sub>0] xa1) # frs" let "?frsT'" = "\xa1. (vs,l,C,M,pc,Throwing [] xa1) # frs" obtain xa' where xa': "throw a = Throw xa'" by (metis RInitFailFinal\<^sub>1.hyps(1) eval\<^sub>1_final exp.distinct(101) final_def) show ?case proof(cases "e = C\<^sub>0\\<^sub>sclinit([])") case clinit: True then obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh'' v xa' (RI (C\<^sub>0,C\<^sub>0\\<^sub>sclinit([])) ; [] \ unit) = (True, ?frs, (None, h', ?frs', sh''), err)" using RInitFailFinal\<^sub>1.prems(1) by clarsimp have IH: "PROP ?P e h l sh (throw a) h' l' sh' E C M pc (Called []) v xa' vs frs I" by fact then have IH': "PROP ?P (C\<^sub>0\\<^sub>sclinit([])) h l sh (throw a) h' l' sh' E C M pc (Called []) v xa' vs frs I" using clinit by simp have "P \ (None,h,?frs,sh) -jvm\ (None,h,create_init_frame P C\<^sub>0 # ?frs',sh)" by(rule jvm_Called) also have "P \ \ -jvm\ (None,h',?frsT' xa',sh'')" using IH' Jcc_pieces_RInit_clinit[OF assms(1-2) pcs,of h' l' sh'] xa' RInitFailFinal\<^sub>1.hyps(3,4) by simp also have "P \ \ -jvm\ handle (compP compMb\<^sub>2 P\<^sub>1) C M xa' h' vs l pc No_ics frs sh''" using RInitFailFinal\<^sub>1.hyps(3,4) jvm_RInit_throw[where h=h' and sh=sh''] by simp finally show ?thesis using xa' pcs clinit by(clarsimp intro!: exI[where x="[]"]) next case throw: False then have eT: "e = Throw xa'" "h = h'" "sh = sh'" using xa' RInitFailFinal\<^sub>1.prems(1) eval\<^sub>1_final_same[OF RInitFailFinal\<^sub>1.hyps(1)] by clarsimp+ obtain a where "class P\<^sub>1 C\<^sub>0 = \a\" using RInitFailFinal\<^sub>1.prems by(auto simp: is_class_def) then obtain stk' loc' M' pc' ics' where "create_init_frame P C\<^sub>0 = (stk',loc',C\<^sub>0,M',pc',ics')" using create_init_frame_wf_eq[OF wf] by(cases "create_init_frame P C\<^sub>0", simp) then obtain rhs err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh'' v xa' (RI (C\<^sub>0,e) ; [] \ unit) = (True, ?frsT xa', rhs, err)" using RInitFailFinal\<^sub>1.prems(1) eT by clarsimp have "P \ (None,h,?frsT xa',sh) -jvm\ (None,h,?frsT' xa',sh(C\<^sub>0 \ (fst (the (sh C\<^sub>0)), Error)))" by(rule jvm_Throwing) also have "P \ \ -jvm\ handle P C M xa' h' vs l pc No_ics frs sh''" using RInitFailFinal\<^sub>1.hyps(3,4) jvm_RInit_throw[where h=h and sh=sh''] eT by simp finally show ?thesis using pcs xa' by(clarsimp intro!: exI[where x="[]"]) qed qed (*>*) (*FIXME move! *) lemma atLeast0AtMost[simp]: "{0::nat..n} = {..n}" by auto lemma atLeast0LessThan[simp]: "{0::nat.. addr option" where "exception (Throw a) = Some a" | "exception e = None" lemma comp\<^sub>2_correct: assumes wf: "wf_J\<^sub>1_prog P\<^sub>1" and "method": "P\<^sub>1 \ C sees M,b:Ts\T = body in C" and eval: "P\<^sub>1 \\<^sub>1 \body,(h,ls,sh)\ \ \e',(h',ls',sh')\" and nclinit: "M \ clinit" shows "compP\<^sub>2 P\<^sub>1 \ (None,h,[([],ls,C,M,0,No_ics)],sh) -jvm\ (exception e',h',[],sh')" (*<*) (is "_ \ ?\\<^sub>0 -jvm\ ?\\<^sub>1") proof - let ?P = "compP\<^sub>2 P\<^sub>1" let ?E = "case b of Static \ Ts | NonStatic \ Class C#Ts" have nsub: "\sub_RI body" using sees_wf\<^sub>1_nsub_RI[OF wf method] by simp have code: "?P,C,M,0 \ compE\<^sub>2 body" using beforeM[OF "method"] by auto have xtab: "?P,C,M \ compxE\<^sub>2 body 0 (size[])/{..2 body)},size[]" using beforexM[OF "method"] by auto have cond: "Jcc_cond P\<^sub>1 ?E C M [] 0 No_ics {..2 body)} h sh body" using nsub_RI_Jcc_pieces nsub code xtab by auto \ \Distinguish if e' is a value or an exception\ { fix v assume [simp]: "e' = Val v" have "?P \ ?\\<^sub>0 -jvm\ (None,h',[([v],ls',C,M,size(compE\<^sub>2 body),No_ics)],sh')" using Jcc[OF wf eval cond] nsub_RI_Jcc_pieces[OF _ nsub] by auto also have "?P \ \ -jvm\ ?\\<^sub>1" using beforeM[OF "method"] nclinit by auto finally have ?thesis . } moreover { fix a assume [simp]: "e' = Throw a" obtain pc vs' where pc: "0 \ pc \ pc < size(compE\<^sub>2 body) \ \ caught ?P pc h' a (compxE\<^sub>2 body 0 0)" and 1: "?P \ ?\\<^sub>0 -jvm\ handle ?P C M a h' vs' ls' pc No_ics [] sh'" using Jcc[OF wf eval cond] nsub_RI_Jcc_pieces[OF _ nsub] by auto meson from pc have "handle ?P C M a h' vs' ls' pc No_ics [] sh' = ?\\<^sub>1" using xtab "method" nclinit by(auto simp:handle_def compMb\<^sub>2_def) with 1 have ?thesis by simp } ultimately show ?thesis using eval\<^sub>1_final[OF eval] by(auto simp:final_def) qed (*>*) end diff --git a/thys/JinjaDCI/Compiler/J1.thy b/thys/JinjaDCI/Compiler/J1.thy --- a/thys/JinjaDCI/Compiler/J1.thy +++ b/thys/JinjaDCI/Compiler/J1.thy @@ -1,535 +1,548 @@ (* Title: JinjaDCI/Compiler/J1.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory Compiler/J1.thy by Tobias Nipkow *) chapter \ Compilation \label{cha:comp} \ section \ An Intermediate Language \ theory J1 imports "../J/BigStep" begin type_synonym expr\<^sub>1 = "nat exp" type_synonym J\<^sub>1_prog = "expr\<^sub>1 prog" type_synonym state\<^sub>1 = "heap \ (val list) \ sheap" definition hp\<^sub>1 :: "state\<^sub>1 \ heap" where "hp\<^sub>1 \ fst" definition lcl\<^sub>1 :: "state\<^sub>1 \ val list" where "lcl\<^sub>1 \ fst \ snd" definition shp\<^sub>1 :: "state\<^sub>1 \ sheap" where "shp\<^sub>1 \ snd \ snd" (*<*) declare hp\<^sub>1_def[simp] lcl\<^sub>1_def[simp] shp\<^sub>1_def[simp] (*>*) primrec max_vars :: "'a exp \ nat" and max_varss :: "'a exp list \ nat" where "max_vars(new C) = 0" | "max_vars(Cast C e) = max_vars e" | "max_vars(Val v) = 0" | "max_vars(e\<^sub>1 \bop\ e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)" | "max_vars(Var V) = 0" | "max_vars(V:=e) = max_vars e" | "max_vars(e\F{D}) = max_vars e" | "max_vars(C\\<^sub>sF{D}) = 0" | "max_vars(FAss e\<^sub>1 F D e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)" | "max_vars(SFAss C F D e\<^sub>2) = max_vars e\<^sub>2" | "max_vars(e\M(es)) = max (max_vars e) (max_varss es)" | "max_vars(C\\<^sub>sM(es)) = max_varss es" | "max_vars({V:T; e}) = max_vars e + 1" | "max_vars(e\<^sub>1;;e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)" | "max_vars(if (e) e\<^sub>1 else e\<^sub>2) = max (max_vars e) (max (max_vars e\<^sub>1) (max_vars e\<^sub>2))" | "max_vars(while (b) e) = max (max_vars b) (max_vars e)" | "max_vars(throw e) = max_vars e" | "max_vars(try e\<^sub>1 catch(C V) e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2 + 1)" | "max_vars(INIT C (Cs,b) \ e) = max_vars e" | "max_vars(RI(C,e);Cs \ e') = max (max_vars e) (max_vars e')" | "max_varss [] = 0" | "max_varss (e#es) = max (max_vars e) (max_varss es)" inductive eval\<^sub>1 :: "J\<^sub>1_prog \ expr\<^sub>1 \ state\<^sub>1 \ expr\<^sub>1 \ state\<^sub>1 \ bool" ("_ \\<^sub>1 ((1\_,/_\) \/ (1\_,/_\))" [51,0,0,0,0] 81) and evals\<^sub>1 :: "J\<^sub>1_prog \ expr\<^sub>1 list \ state\<^sub>1 \ expr\<^sub>1 list \ state\<^sub>1 \ bool" ("_ \\<^sub>1 ((1\_,/_\) [\]/ (1\_,/_\))" [51,0,0,0,0] 81) for P :: J\<^sub>1_prog where New\<^sub>1: "\ sh C = Some (sfs, Done); new_Addr h = Some a; P \ C has_fields FDTs; h' = h(a\blank P C) \ \ P \\<^sub>1 \new C,(h,l,sh)\ \ \addr a,(h',l,sh)\" | NewFail\<^sub>1: "\ sh C = Some (sfs, Done); new_Addr h = None \ \ P \\<^sub>1 \new C, (h,l,sh)\ \ \THROW OutOfMemory,(h,l,sh)\" | NewInit\<^sub>1: "\ \sfs. sh C = Some (sfs, Done); P \\<^sub>1 \INIT C ([C],False) \ unit,(h,l,sh)\ \ \Val v',(h',l',sh')\; new_Addr h' = Some a; P \ C has_fields FDTs; h'' = h'(a\blank P C) \ \ P \\<^sub>1 \new C,(h,l,sh)\ \ \addr a,(h'',l',sh')\" | NewInitOOM\<^sub>1: "\ \sfs. sh C = Some (sfs, Done); P \\<^sub>1 \INIT C ([C],False) \ unit,(h,l,sh)\ \ \Val v',(h',l',sh')\; new_Addr h' = None; is_class P C \ \ P \\<^sub>1 \new C,(h,l,sh)\ \ \THROW OutOfMemory,(h',l',sh')\" | NewInitThrow\<^sub>1: "\ \sfs. sh C = Some (sfs, Done); P \\<^sub>1 \INIT C ([C],False) \ unit,(h,l,sh)\ \ \throw a,s'\; is_class P C \ \ P \\<^sub>1 \new C,(h,l,sh)\ \ \throw a,s'\" | Cast\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,l,sh)\; h a = Some(D,fs); P \ D \\<^sup>* C \ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \addr a,(h,l,sh)\" | CastNull\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \null,s\<^sub>1\" | CastFail\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,l,sh)\; h a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \THROW ClassCast,(h,l,sh)\" | CastThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \Cast C e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | Val\<^sub>1: "P \\<^sub>1 \Val v,s\ \ \Val v,s\" | BinOp\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v\<^sub>2,s\<^sub>2\; binop(bop,v\<^sub>1,v\<^sub>2) = Some v \ \ P \\<^sub>1 \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\ \ \Val v,s\<^sub>2\" | BinOpThrow\<^sub>1\<^sub>1: "P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \throw e,s\<^sub>1\ \ P \\<^sub>1 \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0\ \ \throw e,s\<^sub>1\" | BinOpThrow\<^sub>2\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \throw e,s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\ \ \throw e,s\<^sub>2\" | Var\<^sub>1: "\ ls!i = v; i < size ls \ \ P \\<^sub>1 \Var i,(h,ls,sh)\ \ \Val v,(h,ls,sh)\" | LAss\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \Val v,(h,ls,sh)\; i < size ls; ls' = ls[i := v] \ \ P \\<^sub>1 \i:= e,s\<^sub>0\ \ \unit,(h,ls',sh)\" | LAssThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \i:= e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAcc\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,ls,sh)\; h a = Some(C,fs); P \ C has F,NonStatic:t in D; fs(F,D) = Some v \ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \Val v,(h,ls,sh)\" | FAccNull\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \THROW NullPointer,s\<^sub>1\" | FAccThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAccNone\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,ls,sh)\; h a = Some(C,fs); \(\b t. P \ C has F,b:t in D) \ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \THROW NoSuchFieldError,(h,ls,sh)\" | FAccStatic\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,(h,ls,sh)\; h a = Some(C,fs); P \ C has F,Static:t in D \ \ P \\<^sub>1 \e\F{D},s\<^sub>0\ \ \THROW IncompatibleClassChangeError,(h,ls,sh)\" | SFAcc\<^sub>1: "\ P \ C has F,Static:t in D; sh D = Some (sfs,Done); sfs F = Some v \ \ P \\<^sub>1 \C\\<^sub>sF{D},(h,ls,sh)\ \ \Val v,(h,ls,sh)\" | SFAccInit\<^sub>1: "\ P \ C has F,Static:t in D; \sfs. sh D = Some (sfs,Done); P \\<^sub>1 \INIT D ([D],False) \ unit,(h,ls,sh)\ \ \Val v',(h',ls',sh')\; sh' D = Some (sfs,i); sfs F = Some v \ \ P \\<^sub>1 \C\\<^sub>sF{D},(h,ls,sh)\ \ \Val v,(h',ls',sh')\" | SFAccInitThrow\<^sub>1: "\ P \ C has F,Static:t in D; \sfs. sh D = Some (sfs,Done); P \\<^sub>1 \INIT D ([D],False) \ unit,(h,ls,sh)\ \ \throw a,s'\ \ \ P \\<^sub>1 \C\\<^sub>sF{D},(h,ls,sh)\ \ \throw a,s'\" | SFAccNone\<^sub>1: "\ \(\b t. P \ C has F,b:t in D) \ \ P \\<^sub>1 \C\\<^sub>sF{D},s\ \ \THROW NoSuchFieldError,s\" | SFAccNonStatic\<^sub>1: "\ P \ C has F,NonStatic:t in D \ \ P \\<^sub>1 \C\\<^sub>sF{D},s\ \ \THROW IncompatibleClassChangeError,s\" | FAss\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\; h\<^sub>2 a = Some(C,fs); P \ C has F,NonStatic:t in D; fs' = fs((F,D)\v); h\<^sub>2' = h\<^sub>2(a\(C,fs')) \ \ P \\<^sub>1 \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \unit,(h\<^sub>2',l\<^sub>2,sh\<^sub>2)\" | FAssNull\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \null,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v,s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \THROW NullPointer,s\<^sub>2\" | FAssThrow\<^sub>1\<^sub>1: "P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \throw e',s\<^sub>1\" | FAssThrow\<^sub>2\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \throw e',s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \throw e',s\<^sub>2\" | FAssNone\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\; h\<^sub>2 a = Some(C,fs); \(\b t. P \ C has F,b:t in D) \ \ P \\<^sub>1 \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \THROW NoSuchFieldError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\" | FAssStatic\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\; h\<^sub>2 a = Some(C,fs); P \ C has F,Static:t in D \ \ P \\<^sub>1 \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0\ \ \THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\" | SFAss\<^sub>1: "\ P \\<^sub>1 \e\<^sub>2,s\<^sub>0\ \ \Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\; P \ C has F,Static:t in D; sh\<^sub>1 D = Some(sfs,Done); sfs' = sfs(F\v); sh\<^sub>1' = sh\<^sub>1(D\(sfs',Done)) \ \ P \\<^sub>1 \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\ \ \unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1')\" | SFAssInit\<^sub>1: "\ P \\<^sub>1 \e\<^sub>2,s\<^sub>0\ \ \Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\; P \ C has F,Static:t in D; \sfs. sh\<^sub>1 D = Some(sfs,Done); P \\<^sub>1 \INIT D ([D],False) \ unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\ \ \Val v',(h',l',sh')\; sh' D = Some(sfs,i); sfs' = sfs(F\v); sh'' = sh'(D\(sfs',i)) \ \ P \\<^sub>1 \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\ \ \unit,(h',l',sh'')\" | SFAssInitThrow\<^sub>1: "\ P \\<^sub>1 \e\<^sub>2,s\<^sub>0\ \ \Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\; P \ C has F,Static:t in D; \sfs. sh\<^sub>1 D = Some(sfs,Done); P \\<^sub>1 \INIT D ([D],False) \ unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\ \ \throw a,s'\ \ \ P \\<^sub>1 \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\ \ \throw a,s'\" | SFAssThrow\<^sub>1: "P \\<^sub>1 \e\<^sub>2,s\<^sub>0\ \ \throw e',s\<^sub>2\ \ P \\<^sub>1 \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\ \ \throw e',s\<^sub>2\" | SFAssNone\<^sub>1: "\ P \\<^sub>1 \e\<^sub>2,s\<^sub>0\ \ \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\; \(\b t. P \ C has F,b:t in D) \ \ P \\<^sub>1 \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\ \ \THROW NoSuchFieldError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\" | SFAssNonStatic\<^sub>1: "\ P \\<^sub>1 \e\<^sub>2,s\<^sub>0\ \ \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\; P \ C has F,NonStatic:t in D \ \ P \\<^sub>1 \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\ \ \THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\" | CallObjThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \throw e',s\<^sub>1\" | CallNull\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \map Val vs,s\<^sub>2\ \ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \THROW NullPointer,s\<^sub>2\" | Call\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\; h\<^sub>2 a = Some(C,fs); P \ C sees M,NonStatic:Ts\T = body in D; size vs = size Ts; ls\<^sub>2' = (Addr a) # vs @ replicate (max_vars body) undefined; P \\<^sub>1 \body,(h\<^sub>2,ls\<^sub>2',sh\<^sub>2)\ \ \e',(h\<^sub>3,ls\<^sub>3,sh\<^sub>3)\ \ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \e',(h\<^sub>3,ls\<^sub>2,sh\<^sub>3)\" | CallParamsThrow\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \es',s\<^sub>2\; es' = map Val vs @ throw ex # es\<^sub>2 \ \ P \\<^sub>1 \e\M(es),s\<^sub>0\ \ \throw ex,s\<^sub>2\" | CallNone\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \ps,s\<^sub>1\ [\] \map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\; h\<^sub>2 a = Some(C,fs); \(\b Ts T body D. P \ C sees M,b:Ts\T = body in D) \ \ P \\<^sub>1 \e\M(ps),s\<^sub>0\ \ \THROW NoSuchMethodError,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\" | CallStatic\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,s\<^sub>1\; P \\<^sub>1 \ps,s\<^sub>1\ [\] \map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\; h\<^sub>2 a = Some(C,fs); P \ C sees M,Static:Ts\T = body in D \ \ P \\<^sub>1 \e\M(ps),s\<^sub>0\ \ \THROW IncompatibleClassChangeError,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\" | SCallParamsThrow\<^sub>1: "\ P \\<^sub>1 \es,s\<^sub>0\ [\] \es',s\<^sub>2\; es' = map Val vs @ throw ex # es\<^sub>2 \ \ P \\<^sub>1 \C\\<^sub>sM(es),s\<^sub>0\ \ \throw ex,s\<^sub>2\" | SCallNone\<^sub>1: "\ P \\<^sub>1 \ps,s\<^sub>0\ [\] \map Val vs,s\<^sub>2\; \(\b Ts T body D. P \ C sees M,b:Ts\T = body in D) \ \ P \\<^sub>1 \C\\<^sub>sM(ps),s\<^sub>0\ \ \THROW NoSuchMethodError,s\<^sub>2\" | SCallNonStatic\<^sub>1: "\ P \\<^sub>1 \ps,s\<^sub>0\ [\] \map Val vs,s\<^sub>2\; P \ C sees M,NonStatic:Ts\T = body in D \ \ P \\<^sub>1 \C\\<^sub>sM(ps),s\<^sub>0\ \ \THROW IncompatibleClassChangeError,s\<^sub>2\" | SCallInitThrow\<^sub>1: "\ P \\<^sub>1 \ps,s\<^sub>0\ [\] \map Val vs,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\; P \ C sees M,Static:Ts\T = body in D; \sfs. sh\<^sub>1 D = Some(sfs,Done); M \ clinit; P \\<^sub>1 \INIT D ([D],False) \ unit,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\ \ \throw a,s'\ \ \ P \\<^sub>1 \C\\<^sub>sM(ps),s\<^sub>0\ \ \throw a,s'\" | SCallInit\<^sub>1: "\ P \\<^sub>1 \ps,s\<^sub>0\ [\] \map Val vs,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\; P \ C sees M,Static:Ts\T = body in D; \sfs. sh\<^sub>1 D = Some(sfs,Done); M \ clinit; P \\<^sub>1 \INIT D ([D],False) \ unit,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\ \ \Val v',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\; size vs = size Ts; ls\<^sub>2' = vs @ replicate (max_vars body) undefined; P \\<^sub>1 \body,(h\<^sub>2,ls\<^sub>2',sh\<^sub>2)\ \ \e',(h\<^sub>3,ls\<^sub>3,sh\<^sub>3)\ \ \ P \\<^sub>1 \C\\<^sub>sM(ps),s\<^sub>0\ \ \e',(h\<^sub>3,ls\<^sub>2,sh\<^sub>3)\" | SCall\<^sub>1: "\ P \\<^sub>1 \ps,s\<^sub>0\ [\] \map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\; P \ C sees M,Static:Ts\T = body in D; sh\<^sub>2 D = Some(sfs,Done) \ (M = clinit \ sh\<^sub>2 D = \(sfs, Processing)\); size vs = size Ts; ls\<^sub>2' = vs @ replicate (max_vars body) undefined; P \\<^sub>1 \body,(h\<^sub>2,ls\<^sub>2',sh\<^sub>2)\ \ \e',(h\<^sub>3,ls\<^sub>3,sh\<^sub>3)\ \ \ P \\<^sub>1 \C\\<^sub>sM(ps),s\<^sub>0\ \ \e',(h\<^sub>3,ls\<^sub>2,sh\<^sub>3)\" | Block\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \e',s\<^sub>1\ \ P \\<^sub>1 \Block i T e,s\<^sub>0\ \ \e',s\<^sub>1\" | Seq\<^sub>1: "\ P \\<^sub>1 \e\<^sub>0,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \e\<^sub>1,s\<^sub>1\ \ \e\<^sub>2,s\<^sub>2\ \ \ P \\<^sub>1 \e\<^sub>0;;e\<^sub>1,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" | SeqThrow\<^sub>1: "P \\<^sub>1 \e\<^sub>0,s\<^sub>0\ \ \throw e,s\<^sub>1\ \ P \\<^sub>1 \e\<^sub>0;;e\<^sub>1,s\<^sub>0\ \ \throw e,s\<^sub>1\" | CondT\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \\<^sub>1 \e\<^sub>1,s\<^sub>1\ \ \e',s\<^sub>2\ \ \ P \\<^sub>1 \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \ \e',s\<^sub>2\" | CondF\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \false,s\<^sub>1\; P \\<^sub>1 \e\<^sub>2,s\<^sub>1\ \ \e',s\<^sub>2\ \ \ P \\<^sub>1 \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \ \e',s\<^sub>2\" | CondThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\ \ \throw e',s\<^sub>1\" | WhileF\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \false,s\<^sub>1\ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \unit,s\<^sub>1\" | WhileT\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \\<^sub>1 \c,s\<^sub>1\ \ \Val v\<^sub>1,s\<^sub>2\; P \\<^sub>1 \while (e) c,s\<^sub>2\ \ \e\<^sub>3,s\<^sub>3\ \ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \e\<^sub>3,s\<^sub>3\" | WhileCondThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \throw e',s\<^sub>1\" | WhileBodyThrow\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \true,s\<^sub>1\; P \\<^sub>1 \c,s\<^sub>1\ \ \throw e',s\<^sub>2\\ \ P \\<^sub>1 \while (e) c,s\<^sub>0\ \ \throw e',s\<^sub>2\" | Throw\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \addr a,s\<^sub>1\ \ P \\<^sub>1 \throw e,s\<^sub>0\ \ \Throw a,s\<^sub>1\" | ThrowNull\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \null,s\<^sub>1\ \ P \\<^sub>1 \throw e,s\<^sub>0\ \ \THROW NullPointer,s\<^sub>1\" | ThrowThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \throw e,s\<^sub>0\ \ \throw e',s\<^sub>1\" | Try\<^sub>1: "P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\ \ P \\<^sub>1 \try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\ \ \Val v\<^sub>1,s\<^sub>1\" | TryCatch\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\; h\<^sub>1 a = Some(D,fs); P \ D \\<^sup>* C; i < length ls\<^sub>1; P \\<^sub>1 \e\<^sub>2,(h\<^sub>1,ls\<^sub>1[i:=Addr a],sh\<^sub>1)\ \ \e\<^sub>2',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\ \ \ P \\<^sub>1 \try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\" | TryThrow\<^sub>1: "\ P \\<^sub>1 \e\<^sub>1,s\<^sub>0\ \ \Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\; h\<^sub>1 a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \\<^sub>1 \try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\ \ \Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\" | Nil\<^sub>1: "P \\<^sub>1 \[],s\ [\] \[],s\" | Cons\<^sub>1: "\ P \\<^sub>1 \e,s\<^sub>0\ \ \Val v,s\<^sub>1\; P \\<^sub>1 \es,s\<^sub>1\ [\] \es',s\<^sub>2\ \ \ P \\<^sub>1 \e#es,s\<^sub>0\ [\] \Val v # es',s\<^sub>2\" | ConsThrow\<^sub>1: "P \\<^sub>1 \e,s\<^sub>0\ \ \throw e',s\<^sub>1\ \ P \\<^sub>1 \e#es,s\<^sub>0\ [\] \throw e' # es, s\<^sub>1\" \ \ init rules \ | InitFinal\<^sub>1: "P \\<^sub>1 \e,s\ \ \e',s'\ \ P \\<^sub>1 \INIT C (Nil,b) \ e,s\ \ \e',s'\" | InitNone\<^sub>1: "\ sh C = None; P \\<^sub>1 \INIT C' (C#Cs,False) \ e,(h,l,sh(C \ (sblank P C, Prepared)))\ \ \e',s'\ \ \ P \\<^sub>1 \INIT C' (C#Cs,False) \ e,(h,l,sh)\ \ \e',s'\" | InitDone\<^sub>1: "\ sh C = Some(sfs,Done); P \\<^sub>1 \INIT C' (Cs,True) \ e,(h,l,sh)\ \ \e',s'\ \ \ P \\<^sub>1 \INIT C' (C#Cs,False) \ e,(h,l,sh)\ \ \e',s'\" | InitProcessing\<^sub>1: "\ sh C = Some(sfs,Processing); P \\<^sub>1 \INIT C' (Cs,True) \ e,(h,l,sh)\ \ \e',s'\ \ \ P \\<^sub>1 \INIT C' (C#Cs,False) \ e,(h,l,sh)\ \ \e',s'\" | InitError\<^sub>1: "\ sh C = Some(sfs,Error); P \\<^sub>1 \RI (C, THROW NoClassDefFoundError);Cs \ e,(h,l,sh)\ \ \e',s'\ \ \ P \\<^sub>1 \INIT C' (C#Cs,False) \ e,(h,l,sh)\ \ \e',s'\" | InitObject\<^sub>1: "\ sh C = Some(sfs,Prepared); C = Object; sh' = sh(C \ (sfs,Processing)); P \\<^sub>1 \INIT C' (C#Cs,True) \ e,(h,l,sh')\ \ \e',s'\ \ \ P \\<^sub>1 \INIT C' (C#Cs,False) \ e,(h,l,sh)\ \ \e',s'\" | InitNonObject\<^sub>1: "\ sh C = Some(sfs,Prepared); C \ Object; class P C = Some (D,r); sh' = sh(C \ (sfs,Processing)); P \\<^sub>1 \INIT C' (D#C#Cs,False) \ e,(h,l,sh')\ \ \e',s'\ \ \ P \\<^sub>1 \INIT C' (C#Cs,False) \ e,(h,l,sh)\ \ \e',s'\" | InitRInit\<^sub>1: "P \\<^sub>1 \RI (C,C\\<^sub>sclinit([]));Cs \ e,(h,l,sh)\ \ \e',s'\ \ P \\<^sub>1 \INIT C' (C#Cs,True) \ e,(h,l,sh)\ \ \e',s'\" | RInit\<^sub>1: "\ P \\<^sub>1 \e,s\ \ \Val v, (h',l',sh')\; sh' C = Some(sfs, i); sh'' = sh'(C \ (sfs, Done)); C' = last(C#Cs); P \\<^sub>1 \INIT C' (Cs,True) \ e', (h',l',sh'')\ \ \e\<^sub>1,s\<^sub>1\ \ \ P \\<^sub>1 \RI (C,e);Cs \ e',s\ \ \e\<^sub>1,s\<^sub>1\" | RInitInitFail\<^sub>1: "\ P \\<^sub>1 \e,s\ \ \throw a, (h',l',sh')\; sh' C = Some(sfs, i); sh'' = sh'(C \ (sfs, Error)); P \\<^sub>1 \RI (D,throw a);Cs \ e', (h',l',sh'')\ \ \e\<^sub>1,s\<^sub>1\ \ \ P \\<^sub>1 \RI (C,e);D#Cs \ e',s\ \ \e\<^sub>1,s\<^sub>1\" | RInitFailFinal\<^sub>1: "\ P \\<^sub>1 \e,s\ \ \throw a, (h',l',sh')\; sh' C = Some(sfs, i); sh'' = sh'(C \ (sfs, Error)) \ \ P \\<^sub>1 \RI (C,e);Nil \ e',s\ \ \throw a, (h',l',sh'')\" (*<*) lemmas eval\<^sub>1_evals\<^sub>1_induct = eval\<^sub>1_evals\<^sub>1.induct [split_format (complete)] and eval\<^sub>1_evals\<^sub>1_inducts = eval\<^sub>1_evals\<^sub>1.inducts [split_format (complete)] (*>*) inductive_cases eval\<^sub>1_cases [cases set]: "P \\<^sub>1 \new C,s\ \ \e',s'\" "P \\<^sub>1 \Cast C e,s\ \ \e',s'\" "P \\<^sub>1 \Val v,s\ \ \e',s'\" "P \\<^sub>1 \e\<^sub>1 \bop\ e\<^sub>2,s\ \ \e',s'\" "P \\<^sub>1 \Var v,s\ \ \e',s'\" "P \\<^sub>1 \V:=e,s\ \ \e',s'\" "P \\<^sub>1 \e\F{D},s\ \ \e',s'\" "P \\<^sub>1 \C\\<^sub>sF{D},s\ \ \e',s'\" "P \\<^sub>1 \e\<^sub>1\F{D}:=e\<^sub>2,s\ \ \e',s'\" "P \\<^sub>1 \C\\<^sub>sF{D}:=e\<^sub>2,s\ \ \e',s'\" "P \\<^sub>1 \e\M(es),s\ \ \e',s'\" "P \\<^sub>1 \C\\<^sub>sM(es),s\ \ \e',s'\" "P \\<^sub>1 \{V:T;e\<^sub>1},s\ \ \e',s'\" "P \\<^sub>1 \e\<^sub>1;;e\<^sub>2,s\ \ \e',s'\" "P \\<^sub>1 \if (e) e\<^sub>1 else e\<^sub>2,s\ \ \e',s'\" "P \\<^sub>1 \while (b) c,s\ \ \e',s'\" "P \\<^sub>1 \throw e,s\ \ \e',s'\" "P \\<^sub>1 \try e\<^sub>1 catch(C V) e\<^sub>2,s\ \ \e',s'\" "P \\<^sub>1 \INIT C (Cs,b) \ e,s\ \ \e',s'\" "P \\<^sub>1 \RI (C,e);Cs \ e\<^sub>0,s\ \ \e',s'\" inductive_cases evals\<^sub>1_cases [cases set]: "P \\<^sub>1 \[],s\ [\] \e',s'\" "P \\<^sub>1 \e#es,s\ [\] \e',s'\" (*>*) lemma eval\<^sub>1_final: "P \\<^sub>1 \e,s\ \ \e',s'\ \ final e'" and evals\<^sub>1_final: "P \\<^sub>1 \es,s\ [\] \es',s'\ \ finals es'" (*<*)by(induct rule:eval\<^sub>1_evals\<^sub>1.inducts, simp_all)(*>*) -lemma eval\<^sub>1_final_same: "\ P \\<^sub>1 \e,s\ \ \e',s'\; final e \ \ e = e' \ s = s'" +lemma eval\<^sub>1_final_same: +assumes eval: "P \\<^sub>1 \e,s\ \ \e',s'\" shows "final e \ e = e' \ s = s'" (*<*) -apply(erule finalE) - using eval\<^sub>1_cases(3) apply blast -by (metis eval\<^sub>1_cases(3,17) exp.distinct(101) exp.inject(3) val.distinct(13)) +proof(erule finalE) + fix v assume Val: "e = Val v" + then show ?thesis using eval eval\<^sub>1_cases(3) by blast +next + fix a assume "e = Throw a" + then show ?thesis using eval + by (metis eval\<^sub>1_cases(3,17) exp.distinct(101) exp.inject(3) val.distinct(13)) +qed (*>*) subsection "Property preservation" lemma eval\<^sub>1_preserves_len: "P \\<^sub>1 \e\<^sub>0,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\ \ \e\<^sub>1,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\ \ length ls\<^sub>0 = length ls\<^sub>1" and evals\<^sub>1_preserves_len: "P \\<^sub>1 \es\<^sub>0,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\ [\] \es\<^sub>1,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\ \ length ls\<^sub>0 = length ls\<^sub>1" (*<*)by (induct rule:eval\<^sub>1_evals\<^sub>1_inducts, simp_all)(*>*) lemma evals\<^sub>1_preserves_elen: "\es' s s'. P \\<^sub>1 \es,s\ [\] \es',s'\ \ length es = length es'" -(*<*) -apply(induct es type:list) -apply (auto elim:evals\<^sub>1.cases) -done -(*>*) +(*<*)by(induct es type:list) (auto elim:evals\<^sub>1.cases)(*>*) lemma clinit\<^sub>1_loc_pres: "P \\<^sub>1 \C\<^sub>0\\<^sub>sclinit([]),(h,l,sh)\ \ \e',(h',l',sh')\ \ l = l'" by(auto elim!: eval\<^sub>1_cases(12) evals\<^sub>1_cases(1)) lemma shows init\<^sub>1_ri\<^sub>1_same_loc: "P \\<^sub>1 \e,(h,l,sh)\ \ \e',(h',l',sh')\ \ (\C Cs b M a. e = INIT C (Cs,b) \ unit \ e = C\\<^sub>sM([]) \ e = RI (C,Throw a) ; Cs \ unit \ e = RI (C,C\\<^sub>sclinit([])) ; Cs \ unit \ l = l')" and "P \\<^sub>1 \es,(h,l,sh)\ [\] \es',(h',l',sh')\ \ True" proof(induct rule: eval\<^sub>1_evals\<^sub>1_inducts) case (RInitInitFail\<^sub>1 e h l sh a') then show ?case using eval\<^sub>1_final[of _ _ _ "throw a'"] by(fastforce dest: eval\<^sub>1_final_same[of _ "Throw a"]) next case RInitFailFinal\<^sub>1 then show ?case by(auto dest: eval\<^sub>1_final_same) qed(auto dest: evals\<^sub>1_cases eval\<^sub>1_cases(17) eval\<^sub>1_final_same) lemma init\<^sub>1_same_loc: "P \\<^sub>1 \INIT C (Cs,b) \ unit,(h,l,sh)\ \ \e',(h',l',sh')\ \ l = l'" by(simp add: init\<^sub>1_ri\<^sub>1_same_loc) theorem eval\<^sub>1_hext: "P \\<^sub>1 \e,(h,l,sh)\ \ \e',(h',l',sh')\ \ h \ h'" and evals\<^sub>1_hext: "P \\<^sub>1 \es,(h,l,sh)\ [\] \es',(h',l',sh')\ \ h \ h'" (*<*) proof (induct rule: eval\<^sub>1_evals\<^sub>1_inducts) case New\<^sub>1 thus ?case by(fastforce intro!: hext_new intro:LeastI simp:new_Addr_def split:if_split_asm simp del:fun_upd_apply) next case NewInit\<^sub>1 thus ?case by (meson hext_new hext_trans new_Addr_SomeD) next case FAss\<^sub>1 thus ?case by(auto simp:sym[THEN hext_upd_obj] simp del:fun_upd_apply elim!: hext_trans) qed (auto elim!: hext_trans) (*>*) subsection "Initialization" lemma rinit\<^sub>1_throw: "P\<^sub>1 \\<^sub>1 \RI (D,Throw xa) ; Cs \ e,(h, l, sh)\ \ \e',(h', l', sh')\ \ e' = Throw xa" -apply(induct Cs arbitrary: D h l sh h' l' sh') - apply(drule eval\<^sub>1_cases(20), auto elim: eval\<^sub>1_cases) -apply(drule eval\<^sub>1_cases(20), auto elim: eval\<^sub>1_cases dest: eval\<^sub>1_final_same simp: final_def) -done +proof(induct Cs arbitrary: D h l sh h' l' sh') + case Nil then show ?case + proof(rule eval\<^sub>1_cases(20)) qed(auto elim: eval\<^sub>1_cases) +next + case (Cons C Cs) show ?case using Cons.prems + proof(induct rule: eval\<^sub>1_cases(20)) + case RInit\<^sub>1: 1 + then show ?case using Cons.hyps by(auto elim: eval\<^sub>1_cases) + next + case RInitInitFail\<^sub>1: 2 + then show ?case using Cons.hyps eval\<^sub>1_final_same final_def by blast + next + case RInitFailFinal\<^sub>1: 3 then show ?case by simp + qed +qed lemma rinit\<^sub>1_throwE: "P \\<^sub>1 \RI (C,throw e) ; Cs \ e\<^sub>0,s\ \ \e',s'\ \ \a s\<^sub>t. e' = throw a \ P \\<^sub>1 \throw e,s\ \ \throw a,s\<^sub>t\" proof(induct Cs arbitrary: C e s) case Nil then show ?case proof(rule eval\<^sub>1_cases(20)) \ \ RI \ fix v h' l' sh' assume "P \\<^sub>1 \throw e,s\ \ \Val v,(h', l', sh')\" then show ?case using eval\<^sub>1_cases(17) by blast qed(auto) next case (Cons C' Cs') show ?case using Cons.prems(1) proof(rule eval\<^sub>1_cases(20)) \ \ RI \ fix v h' l' sh' assume "P \\<^sub>1 \throw e,s\ \ \Val v,(h', l', sh')\" then show ?case using eval\<^sub>1_cases(17) by blast next fix a h' l' sh' sfs i D Cs'' assume e''step: "P \\<^sub>1 \throw e,s\ \ \throw a,(h', l', sh')\" and shC: "sh' C = \(sfs, i)\" and riD: "P \\<^sub>1 \RI (D,throw a) ; Cs'' \ e\<^sub>0,(h', l', sh'(C \ (sfs, Error)))\ \ \e',s'\" and "C' # Cs' = D # Cs''" then show ?thesis using Cons.hyps eval\<^sub>1_final eval\<^sub>1_final_same by blast qed(simp) qed end diff --git a/thys/JinjaDCI/Compiler/J1WellForm.thy b/thys/JinjaDCI/Compiler/J1WellForm.thy --- a/thys/JinjaDCI/Compiler/J1WellForm.thy +++ b/thys/JinjaDCI/Compiler/J1WellForm.thy @@ -1,531 +1,519 @@ (* Title: JinjaDCI/Compiler/J1WellForm.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory Compiler/J1WellForm.thy by Tobias Nipkow *) section \ Well-Formedness of Intermediate Language \ theory J1WellForm imports "../J/JWellForm" J1 begin subsection "Well-Typedness" type_synonym env\<^sub>1 = "ty list" \ \type environment indexed by variable number\ inductive WT\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1, expr\<^sub>1 , ty ] \ bool" ("(_,_ \\<^sub>1/ _ :: _)" [51,51,51]50) and WTs\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1, expr\<^sub>1 list, ty list] \ bool" ("(_,_ \\<^sub>1/ _ [::] _)" [51,51,51]50) for P :: J\<^sub>1_prog where WTNew\<^sub>1: "is_class P C \ P,E \\<^sub>1 new C :: Class C" | WTCast\<^sub>1: "\ P,E \\<^sub>1 e :: Class D; is_class P C; P \ C \\<^sup>* D \ P \ D \\<^sup>* C \ \ P,E \\<^sub>1 Cast C e :: Class C" | WTVal\<^sub>1: "typeof v = Some T \ P,E \\<^sub>1 Val v :: T" | WTVar\<^sub>1: "\ E!i = T; i < size E \ \ P,E \\<^sub>1 Var i :: T" | WTBinOp\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1 :: T\<^sub>1; P,E \\<^sub>1 e\<^sub>2 :: T\<^sub>2; case bop of Eq \ (P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1) \ T = Boolean | Add \ T\<^sub>1 = Integer \ T\<^sub>2 = Integer \ T = Integer \ \ P,E \\<^sub>1 e\<^sub>1 \bop\ e\<^sub>2 :: T" | WTLAss\<^sub>1: "\ E!i = T; i < size E; P,E \\<^sub>1 e :: T'; P \ T' \ T \ \ P,E \\<^sub>1 i:=e :: Void" | WTFAcc\<^sub>1: "\ P,E \\<^sub>1 e :: Class C; P \ C sees F,NonStatic:T in D \ \ P,E \\<^sub>1 e\F{D} :: T" | WTSFAcc\<^sub>1: "\ P \ C sees F,Static:T in D \ \ P,E \\<^sub>1 C\\<^sub>sF{D} :: T" | WTFAss\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1 :: Class C; P \ C sees F,NonStatic:T in D; P,E \\<^sub>1 e\<^sub>2 :: T'; P \ T' \ T \ \ P,E \\<^sub>1 e\<^sub>1\F{D} := e\<^sub>2 :: Void" | WTSFAss\<^sub>1: "\ P \ C sees F,Static:T in D; P,E \\<^sub>1 e\<^sub>2 :: T'; P \ T' \ T \ \ P,E \\<^sub>1 C\\<^sub>sF{D}:=e\<^sub>2 :: Void" | WTCall\<^sub>1: "\ P,E \\<^sub>1 e :: Class C; P \ C sees M,NonStatic:Ts' \ T = m in D; P,E \\<^sub>1 es [::] Ts; P \ Ts [\] Ts' \ \ P,E \\<^sub>1 e\M(es) :: T" | WTSCall\<^sub>1: "\ P \ C sees M,Static:Ts \ T = m in D; P,E \\<^sub>1 es [::] Ts'; P \ Ts' [\] Ts; M \ clinit \ \ P,E \\<^sub>1 C\\<^sub>sM(es) :: T" | WTBlock\<^sub>1: "\ is_type P T; P,E@[T] \\<^sub>1 e::T' \ \ P,E \\<^sub>1 {i:T; e} :: T'" | WTSeq\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1::T\<^sub>1; P,E \\<^sub>1 e\<^sub>2::T\<^sub>2 \ \ P,E \\<^sub>1 e\<^sub>1;;e\<^sub>2 :: T\<^sub>2" | WTCond\<^sub>1: "\ P,E \\<^sub>1 e :: Boolean; P,E \\<^sub>1 e\<^sub>1::T\<^sub>1; P,E \\<^sub>1 e\<^sub>2::T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1; P \ T\<^sub>1 \ T\<^sub>2 \ T = T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ T = T\<^sub>1 \ \ P,E \\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 :: T" | WTWhile\<^sub>1: "\ P,E \\<^sub>1 e :: Boolean; P,E \\<^sub>1 c::T \ \ P,E \\<^sub>1 while (e) c :: Void" | WTThrow\<^sub>1: "P,E \\<^sub>1 e :: Class C \ P,E \\<^sub>1 throw e :: Void" | WTTry\<^sub>1: "\ P,E \\<^sub>1 e\<^sub>1 :: T; P,E@[Class C] \\<^sub>1 e\<^sub>2 :: T; is_class P C \ \ P,E \\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 :: T" | WTNil\<^sub>1: "P,E \\<^sub>1 [] [::] []" | WTCons\<^sub>1: "\ P,E \\<^sub>1 e :: T; P,E \\<^sub>1 es [::] Ts \ \ P,E \\<^sub>1 e#es [::] T#Ts" (*<*) declare WT\<^sub>1_WTs\<^sub>1.intros[intro!] declare WTNil\<^sub>1[iff] lemmas WT\<^sub>1_WTs\<^sub>1_induct = WT\<^sub>1_WTs\<^sub>1.induct [split_format (complete)] and WT\<^sub>1_WTs\<^sub>1_inducts = WT\<^sub>1_WTs\<^sub>1.inducts [split_format (complete)] inductive_cases eee[elim!]: "P,E \\<^sub>1 Val v :: T" "P,E \\<^sub>1 Var i :: T" "P,E \\<^sub>1 Cast D e :: T" "P,E \\<^sub>1 i:=e :: T" "P,E \\<^sub>1 {i:U; e} :: T" "P,E \\<^sub>1 e\<^sub>1;;e\<^sub>2 :: T" "P,E \\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 :: T" "P,E \\<^sub>1 while (e) c :: T" "P,E \\<^sub>1 throw e :: T" "P,E \\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 :: T" "P,E \\<^sub>1 e\F{D} :: T" "P,E \\<^sub>1 C\\<^sub>sF{D} :: T" "P,E \\<^sub>1 e\<^sub>1\F{D}:=e\<^sub>2 :: T" "P,E \\<^sub>1 C\\<^sub>sF{D}:=e\<^sub>2 :: T" "P,E \\<^sub>1 e\<^sub>1 \bop\ e\<^sub>2 :: T" "P,E \\<^sub>1 new C :: T" "P,E \\<^sub>1 e\M(es) :: T" "P,E \\<^sub>1 C\\<^sub>sM(es) :: T" "P,E \\<^sub>1 [] [::] Ts" "P,E \\<^sub>1 e#es [::] Ts" (*>*) lemma init_nWT\<^sub>1 [simp]:"\P,E \\<^sub>1 INIT C (Cs,b) \ e :: T" by(auto elim:WT\<^sub>1.cases) lemma rinit_nWT\<^sub>1 [simp]:"\P,E \\<^sub>1 RI(C,e);Cs \ e' :: T" by(auto elim:WT\<^sub>1.cases) lemma WTs\<^sub>1_same_size: "\Ts. P,E \\<^sub>1 es [::] Ts \ size es = size Ts" (*<*)by (induct es type:list) auto(*>*) lemma WT\<^sub>1_unique: "P,E \\<^sub>1 e :: T\<^sub>1 \ (\T\<^sub>2. P,E \\<^sub>1 e :: T\<^sub>2 \ T\<^sub>1 = T\<^sub>2)" and WTs\<^sub>1_unique: "P,E \\<^sub>1 es [::] Ts\<^sub>1 \ (\Ts\<^sub>2. P,E \\<^sub>1 es [::] Ts\<^sub>2 \ Ts\<^sub>1 = Ts\<^sub>2)" (*<*) -apply(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) -apply blast -apply blast -apply clarsimp -apply blast -apply clarsimp -apply(case_tac bop) -apply clarsimp -apply clarsimp -apply blast -apply (blast dest:sees_field_idemp sees_field_fun) -apply (blast dest:sees_field_fun) -apply blast -apply (blast dest:sees_field_fun) -apply (blast dest:sees_method_idemp sees_method_fun) -apply (blast dest:sees_method_fun) -apply blast -apply blast -apply blast -apply blast -apply clarify -apply blast -apply blast -apply blast -done +proof(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) + case WTVal\<^sub>1 then show ?case by clarsimp +next + case (WTBinOp\<^sub>1 E e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 bop T) + then show ?case by(case_tac bop) force+ +next + case WTFAcc\<^sub>1 then show ?case + by (blast dest:sees_field_idemp sees_field_fun) +next + case WTSFAcc\<^sub>1 then show ?case by (blast dest:sees_field_fun) +next + case WTSFAss\<^sub>1 then show ?case by (blast dest:sees_field_fun) +next + case WTCall\<^sub>1 then show ?case + by (blast dest:sees_method_idemp sees_method_fun) +next + case WTSCall\<^sub>1 then show ?case by (blast dest:sees_method_fun) +qed blast+ (*>*) lemma assumes wf: "wf_prog p P" shows WT\<^sub>1_is_type: "P,E \\<^sub>1 e :: T \ set E \ types P \ is_type P T" and "P,E \\<^sub>1 es [::] Ts \ True" (*<*) -apply(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) -apply simp -apply simp -apply (simp add:typeof_lit_is_type) -apply (blast intro:nth_mem) -apply(simp split:bop.splits) -apply simp -apply (simp add:sees_field_is_type[OF _ wf]) -apply (simp add:sees_field_is_type[OF _ wf]) -apply simp -apply simp -apply(fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def) -apply(fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def) -apply simp -apply simp -apply blast -apply simp -apply simp -apply simp -apply simp -apply simp -done +proof(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts) + case WTVal\<^sub>1 then show ?case by (simp add:typeof_lit_is_type) +next + case WTVar\<^sub>1 then show ?case by (blast intro:nth_mem) +next + case WTBinOp\<^sub>1 then show ?case by (simp split:bop.splits) +next + case WTFAcc\<^sub>1 then show ?case + by (simp add:sees_field_is_type[OF _ wf]) +next + case WTSFAcc\<^sub>1 then show ?case + by (simp add:sees_field_is_type[OF _ wf]) +next + case WTCall\<^sub>1 then show ?case + by (fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def) +next + case WTSCall\<^sub>1 then show ?case + by (fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def) +next + case WTCond\<^sub>1 then show ?case by blast +qed simp+ (*>*) lemma WT\<^sub>1_nsub_RI: "P,E \\<^sub>1 e :: T \ \sub_RI e" and WTs\<^sub>1_nsub_RIs: "P,E \\<^sub>1 es [::] Ts \ \sub_RIs es" proof(induct rule: WT\<^sub>1_WTs\<^sub>1.inducts) qed(simp_all) subsection\ Runtime Well-Typedness \ inductive WTrt\<^sub>1 :: "J\<^sub>1_prog \ heap \ sheap \ env\<^sub>1 \ expr\<^sub>1 \ ty \ bool" and WTrts\<^sub>1 :: "J\<^sub>1_prog \ heap \ sheap \ env\<^sub>1 \ expr\<^sub>1 list \ ty list \ bool" and WTrt2\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1,heap,sheap,expr\<^sub>1,ty] \ bool" ("_,_,_,_ \\<^sub>1 _ : _" [51,51,51,51]50) and WTrts2\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1,heap,sheap,expr\<^sub>1 list, ty list] \ bool" ("_,_,_,_ \\<^sub>1 _ [:] _" [51,51,51,51]50) for P :: J\<^sub>1_prog and h :: heap and sh :: sheap where "P,E,h,sh \\<^sub>1 e : T \ WTrt\<^sub>1 P h sh E e T" | "P,E,h,sh \\<^sub>1 es[:]Ts \ WTrts\<^sub>1 P h sh E es Ts" | WTrtNew\<^sub>1: "is_class P C \ P,E,h,sh \\<^sub>1 new C : Class C" | WTrtCast\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : T; is_refT T; is_class P C \ \ P,E,h,sh \\<^sub>1 Cast C e : Class C" | WTrtVal\<^sub>1: "typeof\<^bsub>h\<^esub> v = Some T \ P,E,h,sh \\<^sub>1 Val v : T" | WTrtVar\<^sub>1: "\ E!i = T; i < size E \ \ P,E,h,sh \\<^sub>1 Var i : T" | WTrtBinOpEq\<^sub>1: "\ P,E,h,sh \\<^sub>1 e\<^sub>1 : T\<^sub>1; P,E,h,sh \\<^sub>1 e\<^sub>2 : T\<^sub>2 \ \ P,E,h,sh \\<^sub>1 e\<^sub>1 \Eq\ e\<^sub>2 : Boolean" | WTrtBinOpAdd\<^sub>1: "\ P,E,h,sh \\<^sub>1 e\<^sub>1 : Integer; P,E,h,sh \\<^sub>1 e\<^sub>2 : Integer \ \ P,E,h,sh \\<^sub>1 e\<^sub>1 \Add\ e\<^sub>2 : Integer" | WTrtLAss\<^sub>1: "\ E!i = T; i < size E; P,E,h,sh \\<^sub>1 e : T'; P \ T' \ T \ \ P,E,h,sh \\<^sub>1 i:=e : Void" | WTrtFAcc\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : Class C; P \ C has F,NonStatic:T in D \ \ P,E,h,sh \\<^sub>1 e\F{D} : T" | WTrtFAccNT\<^sub>1: "P,E,h,sh \\<^sub>1 e : NT \ P,E,h,sh \\<^sub>1 e\F{D} : T" | WTrtSFAcc\<^sub>1: "\ P \ C has F,Static:T in D \ \ P,E,h,sh \\<^sub>1 C\\<^sub>sF{D} : T" | WTrtFAss\<^sub>1: "\ P,E,h,sh \\<^sub>1 e\<^sub>1 : Class C; P \ C has F,NonStatic:T in D; P,E,h,sh \\<^sub>1 e\<^sub>2 : T\<^sub>2; P \ T\<^sub>2 \ T \ \ P,E,h,sh \\<^sub>1 e\<^sub>1\F{D}:=e\<^sub>2 : Void" | WTrtFAssNT\<^sub>1: "\ P,E,h,sh \\<^sub>1 e\<^sub>1:NT; P,E,h,sh \\<^sub>1 e\<^sub>2 : T\<^sub>2 \ \ P,E,h,sh \\<^sub>1 e\<^sub>1\F{D}:=e\<^sub>2 : Void" | WTrtSFAss\<^sub>1: "\ P,E,h,sh \\<^sub>1 e\<^sub>2 : T\<^sub>2; P \ C has F,Static:T in D; P \ T\<^sub>2 \ T \ \ P,E,h,sh \\<^sub>1 C\\<^sub>sF{D}:=e\<^sub>2 : Void" | WTrtCall\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : Class C; P \ C sees M,NonStatic:Ts \ T = m in D; P,E,h,sh \\<^sub>1 es [:] Ts'; P \ Ts' [\] Ts \ \ P,E,h,sh \\<^sub>1 e\M(es) : T" | WTrtCallNT\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : NT; P,E,h,sh \\<^sub>1 es [:] Ts \ \ P,E,h,sh \\<^sub>1 e\M(es) : T" | WTrtSCall\<^sub>1: "\ P \ C sees M,Static:Ts \ T = m in D; P,E,h,sh \\<^sub>1 es [:] Ts'; P \ Ts' [\] Ts; M = clinit \ sh D = \(sfs,Processing)\ \ es = map Val vs \ \ P,E,h,sh \\<^sub>1 C\\<^sub>sM(es) : T" | WTrtBlock\<^sub>1: "P,E@[T],h,sh \\<^sub>1 e : T' \ P,E,h,sh \\<^sub>1 {i:T; e} : T'" | WTrtSeq\<^sub>1: "\ P,E,h,sh \\<^sub>1 e\<^sub>1:T\<^sub>1; P,E,h,sh \\<^sub>1 e\<^sub>2:T\<^sub>2 \ \ P,E,h,sh \\<^sub>1 e\<^sub>1;;e\<^sub>2 : T\<^sub>2" | WTrtCond\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : Boolean; P,E,h,sh \\<^sub>1 e\<^sub>1:T\<^sub>1; P,E,h,sh \\<^sub>1 e\<^sub>2:T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1; P \ T\<^sub>1 \ T\<^sub>2 \ T = T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ T = T\<^sub>1 \ \ P,E,h,sh \\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 : T" | WTrtWhile\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : Boolean; P,E,h,sh \\<^sub>1 c:T \ \ P,E,h,sh \\<^sub>1 while(e) c : Void" | WTrtThrow\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : T\<^sub>r; is_refT T\<^sub>r \ \ P,E,h,sh \\<^sub>1 throw e : T" | WTrtTry\<^sub>1: "\ P,E,h,sh \\<^sub>1 e\<^sub>1 : T\<^sub>1; P,E@[Class C],h,sh \\<^sub>1 e\<^sub>2 : T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ \ P,E,h,sh \\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 : T\<^sub>2" | WTrtInit\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : T; \C' \ set (C#Cs). is_class P C'; \sub_RI e; \C' \ set (tl Cs). \sfs. sh C' = \(sfs,Processing)\; b \ (\C' \ set Cs. \sfs. sh C' = \(sfs,Processing)\); distinct Cs; supercls_lst P Cs \ \ P,E,h,sh \\<^sub>1 INIT C (Cs, b) \ e : T" | WTrtRI\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : T; P,E,h,sh \\<^sub>1 e' : T'; \C' \ set (C#Cs). is_class P C'; \sub_RI e'; \C' \ set (C#Cs). not_init C' e; \C' \ set Cs. \sfs. sh C' = \(sfs,Processing)\; \sfs. sh C = \(sfs, Processing)\ \ (sh C = \(sfs, Error)\ \ e = THROW NoClassDefFoundError); distinct (C#Cs); supercls_lst P (C#Cs) \ \ P,E,h,sh \\<^sub>1 RI(C, e);Cs \ e' : T'" \ \well-typed expression lists\ | WTrtNil\<^sub>1: "P,E,h,sh \\<^sub>1 [] [:] []" | WTrtCons\<^sub>1: "\ P,E,h,sh \\<^sub>1 e : T; P,E,h,sh \\<^sub>1 es [:] Ts \ \ P,E,h,sh \\<^sub>1 e#es [:] T#Ts" (*<*) declare WTrt\<^sub>1_WTrts\<^sub>1.intros[intro!] WTrtNil\<^sub>1[iff] declare WTrtFAcc\<^sub>1[rule del] WTrtFAccNT\<^sub>1[rule del] WTrtSFAcc\<^sub>1[rule del] WTrtFAss\<^sub>1[rule del] WTrtFAssNT\<^sub>1[rule del] WTrtSFAss\<^sub>1[rule del] WTrtCall\<^sub>1[rule del] WTrtCallNT\<^sub>1[rule del] WTrtSCall\<^sub>1[rule del] lemmas WTrt\<^sub>1_induct = WTrt\<^sub>1_WTrts\<^sub>1.induct [split_format (complete)] and WTrt\<^sub>1_inducts = WTrt\<^sub>1_WTrts\<^sub>1.inducts [split_format (complete)] (*>*) (*<*) inductive_cases WTrt\<^sub>1_elim_cases[elim!]: "P,E,h,sh \\<^sub>1 Val v : T" "P,E,h,sh \\<^sub>1 Var i : T" "P,E,h,sh \\<^sub>1 v :=e : T" "P,E,h,sh \\<^sub>1 {i:U; e} : T" "P,E,h,sh \\<^sub>1 e\<^sub>1;;e\<^sub>2 : T\<^sub>2" "P,E,h,sh \\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 : T" "P,E,h,sh \\<^sub>1 while(e) c : T" "P,E,h,sh \\<^sub>1 throw e : T" "P,E,h,sh \\<^sub>1 try e\<^sub>1 catch(C V) e\<^sub>2 : T" "P,E,h,sh \\<^sub>1 Cast D e : T" "P,E,h,sh \\<^sub>1 e\F{D} : T" "P,E,h,sh \\<^sub>1 C\\<^sub>sF{D} : T" "P,E,h,sh \\<^sub>1 e\F{D} := v : T" "P,E,h,sh \\<^sub>1 C\\<^sub>sF{D} := v : T" "P,E,h,sh \\<^sub>1 e\<^sub>1 \bop\ e\<^sub>2 : T" "P,E,h,sh \\<^sub>1 new C : T" "P,E,h,sh \\<^sub>1 e\M{D}(es) : T" "P,E,h,sh \\<^sub>1 C\\<^sub>sM{D}(es) : T" "P,E,h,sh \\<^sub>1 INIT C (Cs,b) \ e : T" "P,E,h,sh \\<^sub>1 RI(C,e);Cs \ e' : T" "P,E,h,sh \\<^sub>1 [] [:] Ts" "P,E,h,sh \\<^sub>1 e#es [:] Ts" (*>*) lemma WT\<^sub>1_implies_WTrt\<^sub>1: "P,E \\<^sub>1 e :: T \ P,E,h,sh \\<^sub>1 e : T" and WTs\<^sub>1_implies_WTrts\<^sub>1: "P,E \\<^sub>1 es [::] Ts \ P,E,h,sh \\<^sub>1 es [:] Ts" (*<*) -apply(induct rule: WT\<^sub>1_WTs\<^sub>1_inducts) -apply fast -apply (fast) -apply(fastforce dest:typeof_lit_typeof) -apply(fast) -apply(rename_tac E e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 bop T) apply(case_tac bop) - apply(fastforce) - apply(fastforce) -apply(fastforce) -apply(meson WTrtFAcc\<^sub>1 has_visible_field) -apply(meson WTrtSFAcc\<^sub>1 has_visible_field) -apply(meson WTrtFAss\<^sub>1 has_visible_field) -apply(meson WTrtSFAss\<^sub>1 has_visible_field) -apply(fastforce simp: WTrtCall\<^sub>1) -apply(fastforce simp: WTrtSCall\<^sub>1) -apply(fastforce) -apply(fastforce) -apply(fastforce simp: WTrtCond\<^sub>1) -apply(fastforce) -apply(fastforce) -apply(fastforce) -apply(simp) -apply(fast) -done +proof(induct rule: WT\<^sub>1_WTs\<^sub>1_inducts) + case WTVal\<^sub>1 then show ?case by (fastforce dest:typeof_lit_typeof) +next + case (WTBinOp\<^sub>1 E e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 bop T) + then show ?case by (case_tac bop) fastforce+ +next + case WTFAcc\<^sub>1 then show ?case + by (fastforce simp: WTrtFAcc\<^sub>1 has_visible_field) +next + case WTSFAcc\<^sub>1 then show ?case + by (fastforce simp: WTrtSFAcc\<^sub>1 has_visible_field) +next + case WTFAss\<^sub>1 then show ?case by (meson WTrtFAss\<^sub>1 has_visible_field) +next + case WTSFAss\<^sub>1 then show ?case by (meson WTrtSFAss\<^sub>1 has_visible_field) +next + case WTCall\<^sub>1 then show ?case by (fastforce simp: WTrtCall\<^sub>1) +next + case WTSCall\<^sub>1 then show ?case by (fastforce simp: WTrtSCall\<^sub>1) +qed fastforce+ (*>*) subsection\ Well-formedness\ \ \Indices in blocks increase by 1\ primrec \ :: "expr\<^sub>1 \ nat \ bool" and \s :: "expr\<^sub>1 list \ nat \ bool" where "\ (new C) i = True" | "\ (Cast C e) i = \ e i" | "\ (Val v) i = True" | "\ (e\<^sub>1 \bop\ e\<^sub>2) i = (\ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (Var j) i = True" | "\ (e\F{D}) i = \ e i" | "\ (C\\<^sub>sF{D}) i = True" | "\ (j:=e) i = \ e i" | "\ (e\<^sub>1\F{D} := e\<^sub>2) i = (\ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (C\\<^sub>sF{D} := e\<^sub>2) i = \ e\<^sub>2 i" | "\ (e\M(es)) i = (\ e i \ \s es i)" | "\ (C\\<^sub>sM(es)) i = \s es i" | "\ ({j:T ; e}) i = (i = j \ \ e (i+1))" | "\ (e\<^sub>1;;e\<^sub>2) i = (\ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (if (e) e\<^sub>1 else e\<^sub>2) i = (\ e i \ \ e\<^sub>1 i \ \ e\<^sub>2 i)" | "\ (throw e) i = \ e i" | "\ (while (e) c) i = (\ e i \ \ c i)" | "\ (try e\<^sub>1 catch(C j) e\<^sub>2) i = (\ e\<^sub>1 i \ i=j \ \ e\<^sub>2 (i+1))" | "\ (INIT C (Cs,b) \ e) i = \ e i" | "\ (RI(C,e);Cs \ e') i = (\ e i \ \ e' i)" | "\s [] i = True" | "\s (e#es) i = (\ e i \ \s es i)" definition wf_J\<^sub>1_mdecl :: "J\<^sub>1_prog \ cname \ expr\<^sub>1 mdecl \ bool" where "wf_J\<^sub>1_mdecl P C \ \(M,b,Ts,T,body). \sub_RI body \ (case b of NonStatic \ (\T'. P,Class C#Ts \\<^sub>1 body :: T' \ P \ T' \ T) \ \ body \{..size Ts}\ \ \ body (size Ts + 1) | Static \ (\T'. P,Ts \\<^sub>1 body :: T' \ P \ T' \ T) \ \ body \{.. \ \ body (size Ts))" lemma wf_J\<^sub>1_mdecl_NonStatic[simp]: "wf_J\<^sub>1_mdecl P C (M,NonStatic,Ts,T,body) \ (\sub_RI body \ (\T'. P,Class C#Ts \\<^sub>1 body :: T' \ P \ T' \ T) \ \ body \{..size Ts}\ \ \ body (size Ts + 1))" (*<*)by (simp add:wf_J\<^sub>1_mdecl_def)(*>*) lemma wf_J\<^sub>1_mdecl_Static[simp]: "wf_J\<^sub>1_mdecl P C (M,Static,Ts,T,body) \ (\sub_RI body \ (\T'. P,Ts \\<^sub>1 body :: T' \ P \ T' \ T) \ \ body \{.. \ \ body (size Ts))" (*<*)by (simp add:wf_J\<^sub>1_mdecl_def)(*>*) abbreviation "wf_J\<^sub>1_prog == wf_prog wf_J\<^sub>1_mdecl" lemma sees_wf\<^sub>1_nsub_RI: - "\ wf_J\<^sub>1_prog P; P \ C sees M,b : Ts\T = body in D \ \ \sub_RI body" -apply(drule sees_wf_mdecl, simp) -apply(unfold wf_J\<^sub>1_mdecl_def wf_mdecl_def, simp) -done +assumes wf: "wf_J\<^sub>1_prog P" and cM: "P \ C sees M,b : Ts\T = body in D" +shows "\sub_RI body" +using sees_wf_mdecl[OF wf cM] by(simp add: wf_J\<^sub>1_mdecl_def wf_mdecl_def) lemma wf\<^sub>1_types_clinit: assumes wf:"wf_prog wf_md P" and ex: "class P C = Some a" and proc: "sh C = \(sfs, Processing)\" shows "P,E,h,sh \\<^sub>1 C\\<^sub>sclinit([]) : Void" proof - from ex obtain D fs ms where "a = (D,fs,ms)" by(cases a) then have sP: "(C, D, fs, ms) \ set P" using ex map_of_SomeD[of P C a] by(simp add: class_def) then have "wf_clinit ms" using assms by(unfold wf_prog_def wf_cdecl_def, auto) then obtain m where sm: "(clinit, Static, [], Void, m) \ set ms" by(unfold wf_clinit_def) auto then have "P \ C sees clinit,Static:[] \ Void = m in C" using mdecl_visible[OF wf sP sm] by simp then show ?thesis using WTrtSCall\<^sub>1 proc by blast qed lemma assumes wf: "wf_J\<^sub>1_prog P" shows eval\<^sub>1_proc_pres: "P \\<^sub>1 \e,(h,l,sh)\ \ \e',(h',l',sh')\ \ not_init C e \ \sfs. sh C = \(sfs, Processing)\ \ \sfs'. sh' C = \(sfs', Processing)\" and evals\<^sub>1_proc_pres: "P \\<^sub>1 \es,(h,l,sh)\ [\] \es',(h',l',sh')\ \ not_inits C es \ \sfs. sh C = \(sfs, Processing)\ \ \sfs'. sh' C = \(sfs', Processing)\" (*<*) proof(induct rule:eval\<^sub>1_evals\<^sub>1_inducts) case Call\<^sub>1 then show ?case using sees_wf\<^sub>1_nsub_RI[OF wf Call\<^sub>1.hyps(6)] nsub_RI_not_init by auto next case (SCallInit\<^sub>1 ps h l sh vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C' M Ts T body D v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) then show ?case using SCallInit\<^sub>1 sees_wf\<^sub>1_nsub_RI[OF wf SCallInit\<^sub>1.hyps(3)] nsub_RI_not_init[of body] by auto next case SCall\<^sub>1 then show ?case using sees_wf\<^sub>1_nsub_RI[OF wf SCall\<^sub>1.hyps(3)] nsub_RI_not_init by auto next case (InitNone\<^sub>1 sh C1 C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto next case (InitDone\<^sub>1 sh C sfs C' Cs h l e' a a b) then show ?case by(cases Cs, auto) next case (InitProcessing\<^sub>1 sh C sfs C' Cs h l e' a a b) then show ?case by(cases Cs, auto) next case (InitError\<^sub>1 sh C1 sfs Cs h l e' a a b C') then show ?case by(cases "C = C1") auto next case (InitObject\<^sub>1 sh C1 sfs sh' C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto next case (InitNonObject\<^sub>1 sh C1 sfs D a b sh' C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto next case (RInit\<^sub>1 e a a b v h' l' sh' C sfs i sh'' C' Cs e\<^sub>1 a a b) then show ?case by(cases Cs, auto) next case (RInitInitFail\<^sub>1 e h l sh a h' l' sh' C1 sfs i sh'' D Cs e\<^sub>1 h1 l1 sh1) then show ?case using eval\<^sub>1_final by fastforce qed(auto) +(*>*) lemma clinit\<^sub>1_proc_pres: "\ wf_J\<^sub>1_prog P; P \\<^sub>1 \C\<^sub>0\\<^sub>sclinit([]),(h,l,sh)\ \ \e',(h',l',sh')\; sh C' = \(sfs,Processing)\ \ \ \sfs. sh' C' = \(sfs,Processing)\" by(auto dest: eval\<^sub>1_proc_pres) end diff --git a/thys/JinjaDCI/Compiler/PCompiler.thy b/thys/JinjaDCI/Compiler/PCompiler.thy --- a/thys/JinjaDCI/Compiler/PCompiler.thy +++ b/thys/JinjaDCI/Compiler/PCompiler.thy @@ -1,314 +1,393 @@ (* Title: JinjaDCI/Compiler/PCompiler.thy Author: Tobias Nipkow, Susannah Mansky Copyright TUM 2003, UIUC 2019-20 Based on the Jinja theory Common/PCompiler.thy by Tobias Nipkow *) section \ Program Compilation \ theory PCompiler imports "../Common/WellForm" begin definition compM :: "(staticb \ 'a \ 'b) \ 'a mdecl \ 'b mdecl" where "compM f \ \(M, b, Ts, T, m). (M, b, Ts, T, f b m)" definition compC :: "(staticb \ 'a \ 'b) \ 'a cdecl \ 'b cdecl" where "compC f \ \(C,D,Fdecls,Mdecls). (C,D,Fdecls, map (compM f) Mdecls)" definition compP :: "(staticb \ 'a \ 'b) \ 'a prog \ 'b prog" where "compP f \ map (compC f)" text\ Compilation preserves the program structure. Therefore lookup functions either commute with compilation (like method lookup) or are preserved by it (like the subclass relation). \ lemma map_of_map4: "map_of (map (\(x,a,b,c).(x,a,b,f c)) ts) = map_option (\(a,b,c).(a,b,f c)) \ (map_of ts)" (*<*) -apply(induct ts) - apply simp -apply(rule ext) -apply fastforce -done +proof(induct ts) + case Nil then show ?case by simp +qed fastforce (*>*) lemma map_of_map245: "map_of (map (\(x,a,b,c,d).(x,a,b,c,f a c d)) ts) = map_option (\(a,b,c,d).(a,b,c,f a c d)) \ (map_of ts)" (*<*) -apply(induct ts) - apply simp -apply(rule ext) -apply fastforce -done +proof(induct ts) + case Nil then show ?case by simp +qed fastforce (*>*) lemma class_compP: "class P C = Some (D, fs, ms) \ class (compP f P) C = Some (D, fs, map (compM f) ms)" (*<*)by(simp add:class_def compP_def compC_def map_of_map4)(*>*) lemma class_compPD: "class (compP f P) C = Some (D, fs, cms) \ \ms. class P C = Some(D,fs,ms) \ cms = map (compM f) ms" (*<*)by(clarsimp simp add:class_def compP_def compC_def map_of_map4)(*>*) lemma [simp]: "is_class (compP f P) C = is_class P C" (*<*)by(auto simp:is_class_def dest: class_compP class_compPD)(*>*) lemma [simp]: "class (compP f P) C = map_option (\c. snd(compC f (C,c))) (class P C)" (*<*) -apply(simp add:compP_def compC_def class_def map_of_map4) -apply(simp add:split_def) -done +by(simp add:compP_def compC_def class_def map_of_map4) + (simp add:split_def) (*>*) lemma sees_methods_compP: "P \ C sees_methods Mm \ compP f P \ C sees_methods (map_option (\((b,Ts,T,m),D). ((b,Ts,T,f b m),D)) \ Mm)" -(*<*) -apply(erule Methods.induct) - apply(rule sees_methods_Object) - apply(erule class_compP) - apply(rule ext) - apply(simp add:compM_def map_of_map245 option.map_comp) - apply(case_tac "map_of ms x") - apply simp - apply fastforce -apply(rule sees_methods_rec) - apply(erule class_compP) - apply assumption - apply assumption -apply(rule ext) -apply(simp add:map_add_def compM_def map_of_map245 option.map_comp split:option.split) -done +(*<*)(is "?P \ compP f P \ C sees_methods (?map Mm)") +proof(induct rule: Methods.induct) + case Object: (sees_methods_Object D fs ms Mm) + let ?Mm1 = "\x. map_option ((\m. (m, Object)) \ (\(b, Ts, T, m). (b, Ts, T, f b m))) (map_of ms x)" + let ?Mm2 = "\x. map_option (case_prod (\(b, Ts, T, m). + Pair (b, Ts, T, f b m)) \ (\m. (m, Object))) (map_of ms x)" + have Mm_eq: "\x. ?Mm1 x = ?Mm2 x" + proof - + fix x show "?Mm1 x = ?Mm2 x" + proof(cases "map_of ms x") + case None then show ?thesis by simp + qed fastforce + qed + + have Mm: "Mm = map_option (\m. (m, Object)) \ map_of ms" by fact + let ?Mm = "map_option (\m. (m, Object)) \ map_of (map (compM f) ms)" + let ?Mm' = "?map Mm" + have "?Mm' = ?Mm" + by(rule ext) (simp add:Mm Mm_eq compM_def map_of_map245 option.map_comp) + then show ?case by(rule sees_methods_Object[OF class_compP[OF Object(1)]]) +next + case rec: (sees_methods_rec C D fs ms Mm Mm') + have Mm': "Mm' = Mm ++ (map_option (\m. (m, C)) \ map_of ms)" by fact + let ?Mm' = "?map Mm'" + let ?Mm'' = "(?map Mm) ++ (map_option (\m. (m, C)) \ map_of (map (compM f) ms))" + have "?Mm' = ?Mm''" + by(rule ext) (simp add:Mm' map_add_def compM_def map_of_map245) + moreover have "compP f P \ C sees_methods ?Mm''" + using sees_methods_rec[OF class_compP[OF rec(1)] rec(2,4)] by fast + ultimately show "compP f P \ C sees_methods ?Mm'" by simp +qed (*>*) lemma sees_method_compP: "P \ C sees M,b: Ts\T = m in D \ compP f P \ C sees M,b: Ts\T = (f b m) in D" (*<*)by(fastforce elim:sees_methods_compP simp add:Method_def)(*>*) lemma [simp]: "P \ C sees M,b: Ts\T = m in D \ method (compP f P) C M = (D,b,Ts,T,f b m)" (*<*) -apply(drule sees_method_compP) -apply(simp add:method_def) -apply(rule the_equality) - apply simp -apply(fastforce dest:sees_method_fun) -done +proof - + let ?P = "\(D, b, Ts, T, m). compP f P \ C sees M, b : Ts\T = m in D" + let ?a = "(D, b, Ts, T, f b m)" + assume cM: "P \ C sees M,b: Ts\T = m in D" + have compP_cM: "?P ?a" using sees_method_compP[OF cM] by simp + moreover { + fix x assume "?P x" then have "x = ?a" + using compP_cM by(fastforce dest:sees_method_fun) + } + ultimately have "(THE x. ?P x) = ?a" by(rule the_equality) + then show ?thesis by(simp add:method_def) +qed (*>*) lemma sees_methods_compPD: "\ cP \ C sees_methods Mm'; cP = compP f P \ \ \Mm. P \ C sees_methods Mm \ Mm' = (map_option (\((b,Ts,T,m),D). ((b,Ts,T,f b m),D)) \ Mm)" -(*<*) -apply(erule Methods.induct) - apply(clarsimp simp:compC_def) - apply(rule exI) - apply(rule conjI, erule sees_methods_Object) - apply(rule refl) - apply(rule ext) - apply(simp add:compM_def map_of_map245 option.map_comp) - apply(case_tac "map_of b x") - apply simp - apply fastforce -apply(clarsimp simp:compC_def) -apply(rule exI, rule conjI) - apply(erule (2) sees_methods_rec) - apply(rule refl) -apply(rule ext) -apply(simp add:map_add_def compM_def map_of_map245 option.map_comp split:option.split) -done +(*<*)(is "\ ?P; ?Q \ \ \Mm. P \ C sees_methods Mm \ Mm' = (?map Mm)") +proof(induct rule: Methods.induct) + case Object: (sees_methods_Object D fs ms Mm) + then obtain ms' where P_Obj: "class P Object = \(D, fs, ms')\" + and ms: "ms = map (compM f) ms'" by(clarsimp simp:compC_def) + + let ?Mm1 = "\x. map_option ((\m. (m, Object)) \ (\(b, Ts, T, m). (b, Ts, T, f b m))) (map_of ms' x)" + let ?Mm2 = "\x. map_option (case_prod (\(b, Ts, T, m). Pair (b, Ts, T, f b m)) \ (\m. (m, Object))) + (map_of ms' x)" + have Mm_eq: "\x. ?Mm1 x = ?Mm2 x" + proof - + fix x show "?Mm1 x = ?Mm2 x" + proof(cases "map_of ms' x") + case None then show ?thesis by simp + qed fastforce + qed + + let ?Mm = "map_option (\m. (m,Object)) \ map_of ms'" + let ?Mm' = "?map ?Mm" + have Mm: "Mm = map_option (\m. (m, Object)) \ map_of ms" by fact + have "P \ Object sees_methods ?Mm" + using sees_methods_Object[OF P_Obj] by simp + moreover have "Mm = ?Mm'" + by(rule ext) (simp add:Mm_eq Mm ms compM_def map_of_map245 option.map_comp) + ultimately show ?case by fast +next + case rec: (sees_methods_rec C D fs ms Mm Mm') + then obtain ms' Mm\<^sub>D where P_D: "class P C = \(D, fs, ms')\" + and ms: "ms = map (compM f) ms'" and C_nObj: "C \ Object" + and Mm\<^sub>D: "P \ D sees_methods Mm\<^sub>D" + and Mm: "Mm = (\a. map_option (case_prod (\(b, Ts, T, m). Pair (b, Ts, T, f b m))) (Mm\<^sub>D a))" + by(clarsimp simp:compC_def) + + let ?Mm = "Mm\<^sub>D ++ (map_option (\m. (m, C)) \ map_of ms')" + let ?Mm1 = "Mm ++ (map_option (\m. (m, C)) \ map_of ms)" + let ?Mm2 = "Mm ++ (map_option (\m. (m, C)) \ map_of (map (compM f) ms'))" + let ?Mm3 = "?map ?Mm" + have "Mm' = ?Mm1" by fact + also have "\ = ?Mm2" using ms by simp + also have "\ = ?Mm3" + by(rule ext)(simp add:Mm map_add_def compM_def map_of_map245) + moreover have "P \ C sees_methods ?Mm" + using sees_methods_rec[OF P_D C_nObj Mm\<^sub>D] by simp + ultimately show ?case by fast +qed (*>*) lemma sees_method_compPD: "compP f P \ C sees M,b: Ts\T = fm in D \ \m. P \ C sees M,b: Ts\T = m in D \ f b m = fm" (*<*) -apply(simp add:Method_def) -apply clarify -apply(drule sees_methods_compPD[OF _ refl]) -apply clarsimp -apply blast -done +proof - + assume "compP f P \ C sees M,b: Ts\T = fm in D" + then obtain Mm where Mm: "compP f P \ C sees_methods Mm" + and MmM: "Mm M = \((b, Ts, T, fm), D)\" + by(clarsimp simp:Method_def) + show ?thesis using sees_methods_compPD[OF Mm refl] MmM + by(fastforce simp: Method_def) +qed (*>*) lemma [simp]: "subcls1(compP f P) = subcls1 P" (*<*) by(fastforce simp add: is_class_def compC_def intro:subcls1I order_antisym dest:subcls1D) (*>*) lemma compP_widen[simp]: "(compP f P \ T \ T') = (P \ T \ T')" (*<*)by(cases T')(simp_all add:widen_Class)(*>*) lemma [simp]: "(compP f P \ Ts [\] Ts') = (P \ Ts [\] Ts')" (*<*) -apply(induct Ts) - apply simp -apply(cases Ts') - apply(auto simp:fun_of_def) -done +proof(induct Ts) + case (Cons a Ts) + then show ?case by(cases Ts')(auto simp:fun_of_def) +qed simp (*>*) lemma [simp]: "is_type (compP f P) T = is_type P T" (*<*)by(cases T) simp_all(*>*) lemma [simp]: "(compP (f::staticb\'a\'b) P \ C has_fields FDTs) = (P \ C has_fields FDTs)" (*<*) (is "?A = ?B") proof { fix cP::"'b prog" assume "cP \ C has_fields FDTs" hence "cP = compP f P \ P \ C has_fields FDTs" proof induct case has_fields_Object thus ?case by(fast intro:Fields.has_fields_Object dest:class_compPD) next case has_fields_rec thus ?case by(fast intro:Fields.has_fields_rec dest:class_compPD) qed } note lem = this assume ?A with lem show ?B by blast next assume ?B thus ?A proof induct case has_fields_Object thus ?case by(fast intro:Fields.has_fields_Object class_compP) next case has_fields_rec thus ?case by(fast intro:Fields.has_fields_rec class_compP) qed qed (*>*) lemma fields_compP [simp]: "fields (compP f P) C = fields P C" (*<*)by(simp add:fields_def)(*>*) lemma ifields_compP [simp]: "ifields (compP f P) C = ifields P C" (*<*)by(simp add:ifields_def)(*>*) lemma blank_compP [simp]: "blank (compP f P) C = blank P C" (*<*)by(simp add:blank_def)(*>*) lemma isfields_compP [simp]: "isfields (compP f P) C = isfields P C" (*<*)by(simp add:isfields_def)(*>*) lemma sblank_compP [simp]: "sblank (compP f P) C = sblank P C" (*<*)by(simp add:sblank_def)(*>*) lemma sees_fields_compP [simp]: "(compP f P \ C sees F,b:T in D) = (P \ C sees F,b:T in D)" (*<*)by(simp add:sees_field_def)(*>*) lemma has_field_compP [simp]: "(compP f P \ C has F,b:T in D) = (P \ C has F,b:T in D)" (*<*)by(simp add:has_field_def)(*>*) lemma field_compP [simp]: "field (compP f P) F D = field P F D" (*<*)by(simp add:field_def)(*>*) subsection\Invariance of @{term wf_prog} under compilation \ lemma [iff]: "distinct_fst (compP f P) = distinct_fst P" (*<*) -apply(simp add:distinct_fst_def compP_def compC_def) -apply(induct P) -apply (auto simp:image_iff) -done +by (induct P) + (auto simp:distinct_fst_def compP_def compC_def image_iff) (*>*) lemma [iff]: "distinct_fst (map (compM f) ms) = distinct_fst ms" (*<*) -apply(simp add:distinct_fst_def compM_def) -apply(induct ms) -apply (auto simp:image_iff) -done +by (induct ms) + (auto simp:distinct_fst_def compM_def image_iff) (*>*) lemma [iff]: "wf_syscls (compP f P) = wf_syscls P" (*<*)by(simp add:wf_syscls_def compP_def compC_def image_def Bex_def)(*>*) lemma [iff]: "wf_fdecl (compP f P) = wf_fdecl P" (*<*)by(simp add:wf_fdecl_def)(*>*) lemma wf_clinit_compM [iff]: "wf_clinit (map (compM f) ms) = wf_clinit ms" (*<*) -apply(simp add: wf_clinit_def compM_def) -apply(rule iffI) - apply clarsimp apply(rename_tac m) - apply(rule_tac x = m in exI, simp+) -apply clarsimp apply(rename_tac m) -apply(rule_tac x = "f Static m" in exI, simp add: rev_image_eqI) -done +proof(rule iffI) + assume "wf_clinit (map (compM f) ms)" + then obtain m where "(clinit, Static, [], Void, m) \ set ms" + by(clarsimp simp: wf_clinit_def compM_def) + then show "wf_clinit ms" by(fastforce simp: wf_clinit_def) +next + assume "wf_clinit ms" + then obtain m where "(clinit, Static, [], Void, m) \ set ms" + by(clarsimp simp: wf_clinit_def compM_def) + then have "\m. (clinit, Static, [], Void, m) + \ (\x. case x of (M, b, Ts, T, m) \ (M, b, Ts, T, f b m)) ` set ms" + by(rule_tac x = "f Static m" in exI) (simp add: rev_image_eqI) + then show "wf_clinit (map (compM f) ms)" + by(simp add: wf_clinit_def compM_def) +qed (*>*) lemma set_compP: "((C,D,fs,ms') \ set(compP f P)) = (\ms. (C,D,fs,ms) \ set P \ ms' = map (compM f) ms)" (*<*)by(fastforce simp add:compP_def compC_def image_iff Bex_def)(*>*) lemma wf_cdecl_compPI: "\ \C M b Ts T m. \ wf_mdecl wf\<^sub>1 P C (M,b,Ts,T,m); P \ C sees M,b:Ts\T = m in C \ \ wf_mdecl wf\<^sub>2 (compP f P) C (M,b,Ts,T, f b m); \x\set P. wf_cdecl wf\<^sub>1 P x; x \ set (compP f P); wf_prog p P \ \ wf_cdecl wf\<^sub>2 (compP f P) x" (*<*) -apply(clarsimp simp add:wf_cdecl_def Ball_def set_compP) -apply(rename_tac C D fs ms) -apply(rule conjI) - apply (clarsimp simp:compM_def) - apply (drule (2) mdecl_visible) - apply simp -apply(clarify) -apply(drule sees_method_compPD[where f = f]) -apply clarsimp -apply(fastforce simp:image_iff compM_def) -done +proof - + assume + wfm: "\C M b Ts T m. \ wf_mdecl wf\<^sub>1 P C (M,b,Ts,T,m); P \ C sees M,b:Ts\T = m in C \ + \ wf_mdecl wf\<^sub>2 (compP f P) C (M,b,Ts,T, f b m)" + and wfc: "\x\set P. wf_cdecl wf\<^sub>1 P x" + and compP: "x \ set (compP f P)" and wf: "wf_prog p P" + obtain C D fs ms where x: "x = (C, D, fs, map (compM f) ms)" + and x_set: "(C, D, fs, ms) \ set P" + using compP by(case_tac x) (clarsimp simp: set_compP) + have wfc': "wf_cdecl wf\<^sub>1 P (C, D, fs, ms)" using wfc x_set by fast + let ?P = "compP f P" and ?ms = "compM f ` set ms" + { fix M b Ts T m + assume M: "(M,b,Ts,T,m) \ set ms" + then have "wf_mdecl wf\<^sub>1 P C (M, b, Ts, T, m)" using wfc' + by(simp add:wf_cdecl_def) + moreover have cM: "P \ C sees M, b : Ts\T = m in C" using M + by(rule mdecl_visible[OF wf x_set]) + ultimately have "wf_mdecl wf\<^sub>2 (compP f P) C (M, b, Ts, T, f b m)" + by(rule wfm) + } + then have "\m \ ?ms. wf_mdecl wf\<^sub>2 ?P C m" + by (clarsimp simp:compM_def) + moreover have "C \ Object \ + (\(M,b,Ts,T,m)\?ms. + \D' b' Ts' T' m'. ?P \ D sees M,b':Ts' \ T' = m' in D' \ + b = b' \ P \ Ts' [\] Ts \ P \ T \ T')" + proof - + { fix M b Ts T m D' b' Ts' T' m' + assume "C \ Object" and "(M,b,Ts,T,m)\?ms" + and dM: "?P \ D sees M,b':Ts' \ T' = m' in D'" + then have "b = b' \ P \ Ts' [\] Ts \ P \ T \ T'" + using wfc' sees_method_compPD[OF dM] + by(fastforce simp:wf_cdecl_def image_iff compM_def) + } + then show ?thesis by fast + qed + moreover have "(\f\set fs. wf_fdecl P f) \ distinct_fst fs + \ distinct_fst ms \ wf_clinit ms + \ (C \ Object \ is_class P D \ \ P \ D \\<^sup>* C)" using wfc' + by(simp add: wf_cdecl_def) + ultimately show ?thesis using x by(simp add:wf_cdecl_def) +qed (*>*) lemma wf_prog_compPI: assumes lift: "\C M b Ts T m. \ P \ C sees M,b:Ts\T = m in C; wf_mdecl wf\<^sub>1 P C (M,b,Ts,T,m) \ \ wf_mdecl wf\<^sub>2 (compP f P) C (M,b,Ts,T, f b m)" and wf: "wf_prog wf\<^sub>1 P" shows "wf_prog wf\<^sub>2 (compP f P)" (*<*) using wf by (simp add:wf_prog_def) (blast intro:wf_cdecl_compPI lift wf) (*>*) end diff --git a/thys/JinjaDCI/Compiler/TypeComp.thy b/thys/JinjaDCI/Compiler/TypeComp.thy --- a/thys/JinjaDCI/Compiler/TypeComp.thy +++ b/thys/JinjaDCI/Compiler/TypeComp.thy @@ -1,1471 +1,1619 @@ (* Title: JinjaDCI/Compiler/TypeComp.thy Author: Tobias Nipkow, Susannah Mansky Copyright TUM 2003, UIUC 2019-20 Based on the Jinja theory Compiler/TypeComp.thy by Tobias Nipkow *) section \ Preservation of Well-Typedness \ theory TypeComp imports Compiler "../BV/BVSpec" begin (*<*) declare nth_append[simp] (*>*) lemma max_stack1: "P,E \\<^sub>1 e :: T \ 1 \ max_stack e" (*<*)using max_stack1'[OF WT\<^sub>1_nsub_RI] by simp(*>*) locale TC0 = fixes P :: "J\<^sub>1_prog" and mxl :: nat begin definition "ty E e = (THE T. P,E \\<^sub>1 e :: T)" definition "ty\<^sub>l E A' = map (\i. if i \ A' \ i < size E then OK(E!i) else Err) [0..i' ST E A = (case A of None \ None | \A'\ \ Some(ST, ty\<^sub>l E A'))" definition "after E A ST e = ty\<^sub>i' (ty E e # ST) E (A \ \ e)" end lemma (in TC0) ty_def2 [simp]: "P,E \\<^sub>1 e :: T \ ty E e = T" -(*<*) -apply (unfold ty_def) -apply(blast intro: the_equality WT\<^sub>1_unique) -done -(*>*) +(*<*)by(unfold ty_def) (blast intro: the_equality WT\<^sub>1_unique)(*>*) lemma (in TC0) [simp]: "ty\<^sub>i' ST E None = None" (*<*)by (simp add: ty\<^sub>i'_def)(*>*) lemma (in TC0) ty\<^sub>l_app_diff[simp]: "ty\<^sub>l (E@[T]) (A - {size E}) = ty\<^sub>l E A" (*<*)by(auto simp add:ty\<^sub>l_def hyperset_defs)(*>*) lemma (in TC0) ty\<^sub>i'_app_diff[simp]: "ty\<^sub>i' ST (E @ [T]) (A \ size E) = ty\<^sub>i' ST E A" (*<*)by(auto simp add:ty\<^sub>i'_def hyperset_defs)(*>*) lemma (in TC0) ty\<^sub>l_antimono: "A \ A' \ P \ ty\<^sub>l E A' [\\<^sub>\] ty\<^sub>l E A" (*<*)by(auto simp:ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>i'_antimono: "A \ A' \ P \ ty\<^sub>i' ST E \A'\ \' ty\<^sub>i' ST E \A\" (*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>l_env_antimono: "P \ ty\<^sub>l (E@[T]) A [\\<^sub>\] ty\<^sub>l E A" (*<*)by(auto simp:ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>i'_env_antimono: "P \ ty\<^sub>i' ST (E@[T]) A \' ty\<^sub>i' ST E A" (*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>i'_incr: "P \ ty\<^sub>i' ST (E @ [T]) \insert (size E) A\ \' ty\<^sub>i' ST E \A\" (*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>l_incr: "P \ ty\<^sub>l (E @ [T]) (insert (size E) A) [\\<^sub>\] ty\<^sub>l E A" (*<*)by(auto simp: hyperset_defs ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>l_in_types: "set E \ types P \ ty\<^sub>l E A \ list mxl (err (types P))" (*<*)by(auto simp add:ty\<^sub>l_def intro!:listI dest!: nth_mem)(*>*) locale TC1 = TC0 begin primrec compT :: "ty list \ nat hyperset \ ty list \ expr\<^sub>1 \ ty\<^sub>i' list" and compTs :: "ty list \ nat hyperset \ ty list \ expr\<^sub>1 list \ ty\<^sub>i' list" where "compT E A ST (new C) = []" | "compT E A ST (Cast C e) = compT E A ST e @ [after E A ST e]" | "compT E A ST (Val v) = []" | "compT E A ST (e\<^sub>1 \bop\ e\<^sub>2) = (let ST\<^sub>1 = ty E e\<^sub>1#ST; A\<^sub>1 = A \ \ e\<^sub>1 in compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @ compT E A\<^sub>1 ST\<^sub>1 e\<^sub>2 @ [after E A\<^sub>1 ST\<^sub>1 e\<^sub>2])" | "compT E A ST (Var i) = []" | "compT E A ST (i := e) = compT E A ST e @ [after E A ST e, ty\<^sub>i' ST E (A \ \ e \ \{i}\)]" | "compT E A ST (e\F{D}) = compT E A ST e @ [after E A ST e]" | "compT E A ST (C\\<^sub>sF{D}) = []" | "compT E A ST (e\<^sub>1\F{D} := e\<^sub>2) = (let ST\<^sub>1 = ty E e\<^sub>1#ST; A\<^sub>1 = A \ \ e\<^sub>1; A\<^sub>2 = A\<^sub>1 \ \ e\<^sub>2 in compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @ compT E A\<^sub>1 ST\<^sub>1 e\<^sub>2 @ [after E A\<^sub>1 ST\<^sub>1 e\<^sub>2] @ [ty\<^sub>i' ST E A\<^sub>2])" | "compT E A ST (C\\<^sub>sF{D} := e\<^sub>2) = compT E A ST e\<^sub>2 @ [after E A ST e\<^sub>2] @ [ty\<^sub>i' ST E (A \ \ e\<^sub>2)]" | "compT E A ST {i:T; e} = compT (E@[T]) (A\i) ST e" | "compT E A ST (e\<^sub>1;;e\<^sub>2) = (let A\<^sub>1 = A \ \ e\<^sub>1 in compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1, ty\<^sub>i' ST E A\<^sub>1] @ compT E A\<^sub>1 ST e\<^sub>2)" | "compT E A ST (if (e) e\<^sub>1 else e\<^sub>2) = (let A\<^sub>0 = A \ \ e; \ = ty\<^sub>i' ST E A\<^sub>0 in compT E A ST e @ [after E A ST e, \] @ compT E A\<^sub>0 ST e\<^sub>1 @ [after E A\<^sub>0 ST e\<^sub>1, \] @ compT E A\<^sub>0 ST e\<^sub>2)" | "compT E A ST (while (e) c) = (let A\<^sub>0 = A \ \ e; A\<^sub>1 = A\<^sub>0 \ \ c; \ = ty\<^sub>i' ST E A\<^sub>0 in compT E A ST e @ [after E A ST e, \] @ compT E A\<^sub>0 ST c @ [after E A\<^sub>0 ST c, ty\<^sub>i' ST E A\<^sub>1, ty\<^sub>i' ST E A\<^sub>0])" | "compT E A ST (throw e) = compT E A ST e @ [after E A ST e]" | "compT E A ST (e\M(es)) = compT E A ST e @ [after E A ST e] @ compTs E (A \ \ e) (ty E e # ST) es" | "compT E A ST (C\\<^sub>sM(es)) = compTs E A ST es" | "compT E A ST (try e\<^sub>1 catch(C i) e\<^sub>2) = compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @ [ty\<^sub>i' (Class C#ST) E A, ty\<^sub>i' ST (E@[Class C]) (A \ \{i}\)] @ compT (E@[Class C]) (A \ \{i}\) ST e\<^sub>2" | "compT E A ST (INIT C (Cs,b) \ e) = []" | "compT E A ST (RI(C,e');Cs \ e) = []" | "compTs E A ST [] = []" | "compTs E A ST (e#es) = compT E A ST e @ [after E A ST e] @ compTs E (A \ (\ e)) (ty E e # ST) es" definition compT\<^sub>a :: "ty list \ nat hyperset \ ty list \ expr\<^sub>1 \ ty\<^sub>i' list" where "compT\<^sub>a E A ST e = compT E A ST e @ [after E A ST e]" end lemma compE\<^sub>2_not_Nil[simp]: "P,E \\<^sub>1 e :: T \ compE\<^sub>2 e \ []" (*<*)by(simp add: compE\<^sub>2_not_Nil' WT\<^sub>1_nsub_RI)(*>*) lemma (in TC1) compT_sizes': shows "\E A ST. \sub_RI e \ size(compT E A ST e) = size(compE\<^sub>2 e) - 1" and "\E A ST. \sub_RIs es \ size(compTs E A ST es) = size(compEs\<^sub>2 es)" (*<*) -apply(induct e and es rule: compE\<^sub>2.induct compEs\<^sub>2.induct) -apply(auto split:bop.splits nat_diff_split simp: compE\<^sub>2_not_Nil') -done +by(induct e and es rule: compE\<^sub>2.induct compEs\<^sub>2.induct) + (auto split:bop.splits nat_diff_split simp: compE\<^sub>2_not_Nil') (*>*) lemma (in TC1) compT_sizes[simp]: shows "\E A ST. P,E \\<^sub>1 e :: T \ size(compT E A ST e) = size(compE\<^sub>2 e) - 1" and "\E A ST. P,E \\<^sub>1 es [::] Ts \ size(compTs E A ST es) = size(compEs\<^sub>2 es)" (*<*)using compT_sizes' WT\<^sub>1_nsub_RI WTs\<^sub>1_nsub_RIs by auto(*>*) lemma (in TC1) [simp]: "\ST E. \\\ \ set (compT E None ST e)" and [simp]: "\ST E. \\\ \ set (compTs E None ST es)" (*<*)by(induct e and es rule: compT.induct compTs.induct) (simp_all add:after_def)(*>*) lemma (in TC0) pair_eq_ty\<^sub>i'_conv: "(\(ST, LT)\ = ty\<^sub>i' ST\<^sub>0 E A) = (case A of None \ False | Some A \ (ST = ST\<^sub>0 \ LT = ty\<^sub>l E A))" (*<*)by(simp add:ty\<^sub>i'_def)(*>*) lemma (in TC0) pair_conv_ty\<^sub>i': "\(ST, ty\<^sub>l E A)\ = ty\<^sub>i' ST E \A\" (*<*)by(simp add:ty\<^sub>i'_def)(*>*) (*<*) declare (in TC0) ty\<^sub>i'_antimono [intro!] after_def[simp] pair_conv_ty\<^sub>i'[simp] pair_eq_ty\<^sub>i'_conv[simp] (*>*) lemma (in TC1) compT_LT_prefix: "\E A ST\<^sub>0. \ \(ST,LT)\ \ set(compT E A ST\<^sub>0 e); \ e (size E) \ \ P \ \(ST,LT)\ \' ty\<^sub>i' ST E A" and "\E A ST\<^sub>0. \ \(ST,LT)\ \ set(compTs E A ST\<^sub>0 es); \s es (size E) \ \ P \ \(ST,LT)\ \' ty\<^sub>i' ST E A" (*<*) proof(induct e and es rule: compT.induct compTs.induct) case FAss thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case BinOp thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans split:bop.splits) next case Seq thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case While thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case Cond thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case Block thus ?case by(force simp add:hyperset_defs ty\<^sub>i'_def simp del:pair_conv_ty\<^sub>i' elim!:sup_state_opt_trans) next case Call thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case Cons_exp thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case TryCatch thus ?case by(fastforce simp:hyperset_defs intro!: ty\<^sub>i'_incr elim!:sup_state_opt_trans) qed (auto simp:hyperset_defs) declare (in TC0) ty\<^sub>i'_antimono [rule del] after_def[simp del] pair_conv_ty\<^sub>i'[simp del] pair_eq_ty\<^sub>i'_conv[simp del] (*>*) lemma [iff]: "OK None \ states P mxs mxl" (*<*)by(simp add: JVM_states_unfold)(*>*) lemma (in TC0) after_in_states: - "\ wf_prog p P; P,E \\<^sub>1 e :: T; set E \ types P; set ST \ types P; - size ST + max_stack e \ mxs \ - \ OK (after E A ST e) \ states P mxs mxl" +assumes wf: "wf_prog p P" and wt: "P,E \\<^sub>1 e :: T" + and Etypes: "set E \ types P" and STtypes: "set ST \ types P" + and stack: "size ST + max_stack e \ mxs" +shows "OK (after E A ST e) \ states P mxs mxl" (*<*) -apply(subgoal_tac "size ST + 1 \ mxs") - apply(simp add: after_def ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types) - apply(blast intro!:listI WT\<^sub>1_is_type) -using max_stack1[where e=e] apply fastforce -done +proof - + have "size ST + 1 \ mxs" using max_stack1[where e=e] wt stack + by fastforce + then show ?thesis using assms + by(simp add: after_def ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types) + (blast intro!:listI WT\<^sub>1_is_type) +qed (*>*) lemma (in TC0) OK_ty\<^sub>i'_in_statesI[simp]: "\ set E \ types P; set ST \ types P; size ST \ mxs \ \ OK (ty\<^sub>i' ST E A) \ states P mxs mxl" (*<*) -apply(simp add:ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types) -apply(blast intro!:listI) -done +by(simp add:ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types) + (blast intro!:listI) (*>*) lemma is_class_type_aux: "is_class P C \ is_type P (Class C)" (*<*)by(simp)(*>*) (*<*) declare is_type_simps[simp del] subsetI[rule del] (*>*) theorem (in TC1) compT_states: assumes wf: "wf_prog p P" shows "\E T A ST. \ P,E \\<^sub>1 e :: T; set E \ types P; set ST \ types P; size ST + max_stack e \ mxs; size E + max_vars e \ mxl \ \ OK ` set(compT E A ST e) \ states P mxs mxl" (*<*)(is "\E T A ST. PROP ?P e E T A ST")(*>*) and "\E Ts A ST. \ P,E \\<^sub>1 es[::]Ts; set E \ types P; set ST \ types P; size ST + max_stacks es \ mxs; size E + max_varss es \ mxl \ \ OK ` set(compTs E A ST es) \ states P mxs mxl" (*<*)(is "\E Ts A ST. PROP ?Ps es E Ts A ST") proof(induct e and es rule: compT.induct compTs.induct) case new thus ?case by(simp) next case (Cast C e) thus ?case by (auto simp:after_in_states[OF wf]) next case Val thus ?case by(simp) next case Var thus ?case by(simp) next case LAss thus ?case by(auto simp:after_in_states[OF wf]) next case FAcc thus ?case by(auto simp:after_in_states[OF wf]) next case SFAcc thus ?case by(auto simp:after_in_states[OF wf]) next case FAss thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case SFAss thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case Seq thus ?case by(auto simp:image_Un after_in_states[OF wf]) next case BinOp thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case Cond thus ?case by(force simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case While thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case Block thus ?case by(auto) next case (TryCatch e\<^sub>1 C i e\<^sub>2) moreover have "size ST + 1 \ mxs" using TryCatch.prems max_stack1[where e=e\<^sub>1] by fastforce ultimately show ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf] is_class_type_aux) next case Nil_exp thus ?case by simp next case Cons_exp thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case throw thus ?case by(auto simp: WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case Call thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case SCall thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case INIT thus ?case by simp next case RI thus ?case by simp qed declare is_type_simps[simp] subsetI[intro!] (*>*) definition shift :: "nat \ ex_table \ ex_table" where "shift n xt \ map (\(from,to,C,handler,depth). (from+n,to+n,C,handler+n,depth)) xt" lemma [simp]: "shift 0 xt = xt" (*<*)by(induct xt)(auto simp:shift_def)(*>*) lemma [simp]: "shift n [] = []" (*<*)by(simp add:shift_def)(*>*) lemma [simp]: "shift n (xt\<^sub>1 @ xt\<^sub>2) = shift n xt\<^sub>1 @ shift n xt\<^sub>2" (*<*)by(simp add:shift_def)(*>*) lemma [simp]: "shift m (shift n xt) = shift (m+n) xt" (*<*)by(induct xt)(auto simp:shift_def)(*>*) lemma [simp]: "pcs (shift n xt) = {pc+n|pc. pc \ pcs xt}" (*<*) -apply(auto simp:shift_def pcs_def) -apply(rule_tac x = "x-n" in exI) -apply (force split:nat_diff_split) -done +proof - + { fix x f t C h d + assume "(f,t,C,h,d) \ set xt" and "f + n \ x" + and "x < t + n" + then have "\pc. x = pc + n \ (\x\set xt. pc \ (case x of (f, t, C, h, d) \ {f..*) lemma shift_compxE\<^sub>2: shows "\pc pc' d. shift pc (compxE\<^sub>2 e pc' d) = compxE\<^sub>2 e (pc' + pc) d" and "\pc pc' d. shift pc (compxEs\<^sub>2 es pc' d) = compxEs\<^sub>2 es (pc' + pc) d" (*<*) -apply(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) -apply(auto simp:shift_def ac_simps) -done +by(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) + (auto simp:shift_def ac_simps) (*>*) lemma compxE\<^sub>2_size_convs[simp]: shows "n \ 0 \ compxE\<^sub>2 e n d = shift n (compxE\<^sub>2 e 0 d)" and "n \ 0 \ compxEs\<^sub>2 es n d = shift n (compxEs\<^sub>2 es 0 d)" (*<*)by(simp_all add:shift_compxE\<^sub>2)(*>*) locale TC2 = TC1 + fixes T\<^sub>r :: ty and mxs :: pc begin definition wt_instrs :: "instr list \ ex_table \ ty\<^sub>i' list \ bool" ("(\ _, _ /[::]/ _)" [0,0,51] 50) where "\ is,xt [::] \s \ size is < size \s \ pcs xt \ {0.. (\pc< size is. P,T\<^sub>r,mxs,size \s,xt \ is!pc,pc :: \s)" end notation TC2.wt_instrs ("(_,_,_ \/ _, _ /[::]/ _)" [50,50,50,50,50,51] 50) (*<*) lemmas (in TC2) wt_defs = wt_instrs_def wt_instr_def app_def eff_def norm_eff_def (*>*) lemma (in TC2) [simp]: "\s \ [] \ \ [],[] [::] \s" (*<*) by (simp add: wt_defs) (*>*) lemma [simp]: "eff i P pc et None = []" (*<*)by (simp add: Effect.eff_def)(*>*) (*<*) declare split_comp_eq[simp del] (*>*) lemma wt_instr_appR: "\ P,T,m,mpc,xt \ is!pc,pc :: \s; pc < size is; size is < size \s; mpc \ size \s; mpc \ mpc' \ \ P,T,m,mpc',xt \ is!pc,pc :: \s@\s'" (*<*)by (fastforce simp:wt_instr_def app_def)(*>*) lemma relevant_entries_shift [simp]: "relevant_entries P i (pc+n) (shift n xt) = shift n (relevant_entries P i pc xt)" (*<*) - apply (induct xt) - apply (unfold relevant_entries_def shift_def) - apply simp - apply (auto simp add: is_relevant_entry_def) - done +proof(induct xt) + case Nil + then show ?case by (simp add: relevant_entries_def shift_def) +next + case (Cons a xt) + then show ?case + by (auto simp add: relevant_entries_def shift_def is_relevant_entry_def) +qed (*>*) lemma [simp]: "xcpt_eff i P (pc+n) \ (shift n xt) = map (\(pc,\). (pc + n, \)) (xcpt_eff i P pc \ xt)" (*<*) -apply(simp add: xcpt_eff_def) -apply(cases \) -apply(auto simp add: shift_def) -done +proof - + obtain ST LT where "\ = (ST, LT)" by(cases \) simp + then show ?thesis by(simp add: xcpt_eff_def) (auto simp add: shift_def) +qed (*>*) lemma [simp]: "app\<^sub>i (i, P, pc, m, T, \) \ eff i P (pc+n) (shift n xt) (Some \) = map (\(pc,\). (pc+n,\)) (eff i P pc xt (Some \))" -(*<*) -apply(simp add:eff_def norm_eff_def) -apply(cases "i",auto) -done -(*>*) +(*<*)by(cases "i") (auto simp add:eff_def norm_eff_def)(*>*) lemma [simp]: "xcpt_app i P (pc+n) mxs (shift n xt) \ = xcpt_app i P pc mxs xt \" (*<*)by (simp add: xcpt_app_def) (auto simp add: shift_def)(*>*) lemma wt_instr_appL: - "\ P,T,m,mpc,xt \ i,pc :: \s; pc < size \s; mpc \ size \s \ - \ P,T,m,mpc + size \s',shift (size \s') xt \ i,pc+size \s' :: \s'@\s" +assumes "P,T,m,mpc,xt \ i,pc :: \s" and "pc < size \s" and "mpc \ size \s" +shows "P,T,m,mpc + size \s',shift (size \s') xt \ i,pc+size \s' :: \s'@\s" (*<*) -apply(auto simp:wt_instr_def app_def) -prefer 2 apply(fast) -prefer 2 apply(fast) -apply(cases "i",auto) -done +proof - + let ?t = "(\s'@\s)!(pc+size \s')" + show ?thesis + proof(cases ?t) + case (Some \) + obtain ST LT where [simp]: "\ = (ST, LT)" by(cases \) simp + have "app\<^sub>i (i, P, pc + length \s', m, T, \)" + using Some assms by(cases "i") (auto simp:wt_instr_def app_def) + moreover { + fix pc' \' assume "(pc',\') \ set (eff i P pc xt ?t)" + then have "P \ \' \' \s!pc'" and "pc' < mpc" + using Some assms by(auto simp:wt_instr_def app_def) + } + ultimately show ?thesis using Some assms + by(fastforce simp:wt_instr_def app_def) + qed (auto simp:wt_instr_def app_def) +qed (*>*) lemma wt_instr_Cons: - "\ P,T,m,mpc - 1,[] \ i,pc - 1 :: \s; - 0 < pc; 0 < mpc; pc < size \s + 1; mpc \ size \s + 1 \ - \ P,T,m,mpc,[] \ i,pc :: \#\s" +assumes wti: "P,T,m,mpc - 1,[] \ i,pc - 1 :: \s" + and pcl: "0 < pc" and mpcl: "0 < mpc" + and pcu: "pc < size \s + 1" and mpcu: "mpc \ size \s + 1" +shows "P,T,m,mpc,[] \ i,pc :: \#\s" (*<*) -apply(drule wt_instr_appL[where \s' = "[\]"]) -apply arith -apply arith -apply (simp split:nat_diff_split_asm) -done +proof - + have "pc - 1 < length \s" using pcl pcu by arith + moreover have "mpc - 1 \ length \s" using mpcl mpcu by arith + ultimately have + "P,T,m,mpc - 1 + length [\],shift (length [\]) [] \ i,pc - 1 + length [\] :: [\] @ \s" + by(rule wt_instr_appL[where \s' = "[\]", OF wti]) + then show ?thesis using pcl mpcl by (simp split:nat_diff_split_asm) +qed (*>*) lemma wt_instr_append: - "\ P,T,m,mpc - size \s',[] \ i,pc - size \s' :: \s; - size \s' \ pc; size \s' \ mpc; pc < size \s + size \s'; mpc \ size \s + size \s' \ - \ P,T,m,mpc,[] \ i,pc :: \s'@\s" +assumes wti: "P,T,m,mpc - size \s',[] \ i,pc - size \s' :: \s" + and pcl: "size \s' \ pc" and mpcl: "size \s' \ mpc" + and pcu: "pc < size \s + size \s'" and mpcu: "mpc \ size \s + size \s'" +shows "P,T,m,mpc,[] \ i,pc :: \s'@\s" (*<*) -apply(drule wt_instr_appL[where \s' = \s']) -apply arith -apply arith -apply (simp split:nat_diff_split_asm) -done +proof - + have "pc - length \s' < length \s" using pcl pcu by arith + moreover have "mpc - length \s' \ length \s" using mpcl mpcu by arith +thm wt_instr_appL[where \s' = "\s'", OF wti] + ultimately have "P,T,m,mpc - length \s' + length \s',shift (length \s') [] + \ i,pc - length \s' + length \s' :: \s' @ \s" + by(rule wt_instr_appL[where \s' = "\s'", OF wti]) + then show ?thesis using pcl mpcl by (simp split:nat_diff_split_asm) +qed (*>*) lemma xcpt_app_pcs: "pc \ pcs xt \ xcpt_app i P pc mxs xt \" (*<*) by (auto simp add: xcpt_app_def relevant_entries_def is_relevant_entry_def pcs_def) (*>*) lemma xcpt_eff_pcs: "pc \ pcs xt \ xcpt_eff i P pc \ xt = []" (*<*) by (cases \) (auto simp add: is_relevant_entry_def xcpt_eff_def relevant_entries_def pcs_def intro!: filter_False) (*>*) lemma pcs_shift: "pc < n \ pc \ pcs (shift n xt)" (*<*)by (auto simp add: shift_def pcs_def)(*>*) lemma wt_instr_appRx: "\ P,T,m,mpc,xt \ is!pc,pc :: \s; pc < size is; size is < size \s; mpc \ size \s \ \ P,T,m,mpc,xt @ shift (size is) xt' \ is!pc,pc :: \s" (*<*)by (auto simp:wt_instr_def eff_def app_def xcpt_app_pcs xcpt_eff_pcs)(*>*) lemma wt_instr_appLx: "\ P,T,m,mpc,xt \ i,pc :: \s; pc \ pcs xt' \ \ P,T,m,mpc,xt'@xt \ i,pc :: \s" (*<*)by (auto simp:wt_instr_def app_def eff_def xcpt_app_pcs xcpt_eff_pcs)(*>*) lemma (in TC2) wt_instrs_extR: "\ is,xt [::] \s \ \ is,xt [::] \s @ \s'" (*<*)by(auto simp add:wt_instrs_def wt_instr_appR)(*>*) lemma (in TC2) wt_instrs_ext: - "\ \ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@\s\<^sub>2; \ is\<^sub>2,xt\<^sub>2 [::] \s\<^sub>2; size \s\<^sub>1 = size is\<^sub>1 \ - \ \ is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" +assumes wt\<^sub>1: "\ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@\s\<^sub>2" and wt\<^sub>2: "\ is\<^sub>2,xt\<^sub>2 [::] \s\<^sub>2" + and \s_size: "size \s\<^sub>1 = size is\<^sub>1" + shows "\ is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" (*<*) -apply(clarsimp simp:wt_instrs_def) -apply(rule conjI, fastforce) -apply(rule conjI, fastforce) -apply clarsimp -apply(rule conjI, fastforce simp:wt_instr_appRx) -apply clarsimp -apply(erule_tac x = "pc - size is\<^sub>1" in allE)+ -apply(thin_tac "P \ Q" for P Q) -apply(erule impE, arith) -apply(drule_tac \s' = "\s\<^sub>1" in wt_instr_appL) - apply arith - apply simp -apply(fastforce simp add:add.commute intro!: wt_instr_appLx) -done +proof - + let ?is = "is\<^sub>1@is\<^sub>2" and ?xt = "xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2" + and ?\s = "\s\<^sub>1@\s\<^sub>2" + have "size ?is < size ?\s" using wt\<^sub>2 \s_size by(fastforce simp:wt_instrs_def) + moreover have "pcs ?xt \ {0..1 wt\<^sub>2 + by(fastforce simp:wt_instrs_def) + moreover { + fix pc assume pc: "pcr,mxs,size ?\s,?xt \ ?is!pc,pc :: ?\s" + proof(cases "pc < length is\<^sub>1") + case True then show ?thesis using wt\<^sub>1 pc + by(fastforce simp: wt_instrs_def wt_instr_appRx) + next + case False + then have "pc - length is\<^sub>1 < length is\<^sub>2" using pc by fastforce + then have "P,T\<^sub>r,mxs,length \s\<^sub>2,xt\<^sub>2 \ is\<^sub>2 ! (pc - length is\<^sub>1),pc - length is\<^sub>1 :: \s\<^sub>2" + using wt\<^sub>2 by(clarsimp simp: wt_instrs_def) + moreover have "pc - length is\<^sub>1 < length \s\<^sub>2" using pc wt\<^sub>2 + by(clarsimp simp: wt_instrs_def) arith + moreover have "length \s\<^sub>2 \ length \s\<^sub>2" by simp + moreover have "pc - length is\<^sub>1 + length \s\<^sub>1 \ pcs xt\<^sub>1" using wt\<^sub>1 \s_size + by(fastforce simp: wt_instrs_def) + ultimately have "P,T\<^sub>r,mxs,length \s\<^sub>2 + length \s\<^sub>1,xt\<^sub>1 @ shift (length \s\<^sub>1) xt\<^sub>2 + \ is\<^sub>2 ! (pc - length is\<^sub>1),pc - length is\<^sub>1 + length \s\<^sub>1 :: \s\<^sub>1 @ \s\<^sub>2" + by(rule wt_instr_appLx[OF wt_instr_appL[where \s' = "\s\<^sub>1"]]) + then show ?thesis using False \s_size by(simp add:add.commute) + qed + } + ultimately show ?thesis by(clarsimp simp:wt_instrs_def) +qed (*>*) corollary (in TC2) wt_instrs_ext2: "\ \ is\<^sub>2,xt\<^sub>2 [::] \s\<^sub>2; \ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@\s\<^sub>2; size \s\<^sub>1 = size is\<^sub>1 \ \ \ is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" (*<*)by(rule wt_instrs_ext)(*>*) corollary (in TC2) wt_instrs_ext_prefix [trans]: "\ \ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@\s\<^sub>2; \ is\<^sub>2,xt\<^sub>2 [::] \s\<^sub>3; size \s\<^sub>1 = size is\<^sub>1; prefix \s\<^sub>3 \s\<^sub>2 \ \ \ is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" (*<*)by(bestsimp simp:prefix_def elim: wt_instrs_ext dest:wt_instrs_extR)(*>*) corollary (in TC2) wt_instrs_app: assumes is\<^sub>1: "\ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@[\]" assumes is\<^sub>2: "\ is\<^sub>2,xt\<^sub>2 [::] \#\s\<^sub>2" assumes s: "size \s\<^sub>1 = size is\<^sub>1" shows "\ is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\#\s\<^sub>2" (*<*) proof - from is\<^sub>1 have "\ is\<^sub>1,xt\<^sub>1 [::] (\s\<^sub>1@[\])@\s\<^sub>2" by (rule wt_instrs_extR) hence "\ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@\#\s\<^sub>2" by simp from this is\<^sub>2 s show ?thesis by (rule wt_instrs_ext) qed (*>*) corollary (in TC2) wt_instrs_app_last[trans]: - "\ \ is\<^sub>2,xt\<^sub>2 [::] \#\s\<^sub>2; \ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1; - last \s\<^sub>1 = \; size \s\<^sub>1 = size is\<^sub>1+1 \ - \ \ is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" +assumes "\ is\<^sub>2,xt\<^sub>2 [::] \#\s\<^sub>2" "\ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1" + "last \s\<^sub>1 = \" "size \s\<^sub>1 = size is\<^sub>1+1" +shows "\ is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" (*<*) -apply(cases \s\<^sub>1 rule:rev_cases) - apply simp -apply(simp add:wt_instrs_app) -done +using assms proof(cases \s\<^sub>1 rule:rev_cases) + case (snoc ys y) + then show ?thesis using assms by(simp add:wt_instrs_app) +qed simp (*>*) corollary (in TC2) wt_instrs_append_last[trans]: - "\ \ is,xt [::] \s; P,T\<^sub>r,mxs,mpc,[] \ i,pc :: \s; - pc = size is; mpc = size \s; size is + 1 < size \s \ - \ \ is@[i],xt [::] \s" +assumes wtis: "\ is,xt [::] \s" and wti: "P,T\<^sub>r,mxs,mpc,[] \ i,pc :: \s" + and pc: "pc = size is" and mpc: "mpc = size \s" and is_\s: "size is + 1 < size \s" +shows "\ is@[i],xt [::] \s" (*<*) -apply(clarsimp simp add:wt_instrs_def) -apply(rule conjI, fastforce) -apply(fastforce intro!:wt_instr_appLx[where xt = "[]",simplified] - dest!:less_antisym) -done +proof - + have pc_xt: "pc \ pcs xt" using wtis pc by (fastforce simp:wt_instrs_def) + have "pcs xt \ {.. pc' < length is" "pc' < Suc (length is)" + have "P,T\<^sub>r,mxs,length \s,xt \ i,pc' :: \s" + using wt_instr_appLx[where xt = "[]",simplified,OF wti pc_xt] + less_antisym[OF pc'] pc mpc by simp + } + ultimately show ?thesis using wtis is_\s by(clarsimp simp add:wt_instrs_def) +qed (*>*) corollary (in TC2) wt_instrs_app2: "\ \ is\<^sub>2,xt\<^sub>2 [::] \'#\s\<^sub>2; \ is\<^sub>1,xt\<^sub>1 [::] \#\s\<^sub>1@[\']; xt' = xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2; size \s\<^sub>1+1 = size is\<^sub>1 \ \ \ is\<^sub>1@is\<^sub>2,xt' [::] \#\s\<^sub>1@\'#\s\<^sub>2" (*<*)using wt_instrs_app[where ?\s\<^sub>1.0 = "\ # \s\<^sub>1"] by simp (*>*) corollary (in TC2) wt_instrs_app2_simp[trans,simp]: "\ \ is\<^sub>2,xt\<^sub>2 [::] \'#\s\<^sub>2; \ is\<^sub>1,xt\<^sub>1 [::] \#\s\<^sub>1@[\']; size \s\<^sub>1+1 = size is\<^sub>1 \ \ \ is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \#\s\<^sub>1@\'#\s\<^sub>2" (*<*)using wt_instrs_app[where ?\s\<^sub>1.0 = "\ # \s\<^sub>1"] by simp(*>*) corollary (in TC2) wt_instrs_Cons[simp]: "\ \s \ []; \ [i],[] [::] [\,\']; \ is,xt [::] \'#\s \ \ \ i#is,shift 1 xt [::] \#\'#\s" (*<*) using wt_instrs_app2[where ?is\<^sub>1.0 = "[i]" and ?\s\<^sub>1.0 = "[]" and ?is\<^sub>2.0 = "is" and ?xt\<^sub>1.0 = "[]"] by simp corollary (in TC2) wt_instrs_Cons2[trans]: assumes \s: "\ is,xt [::] \s" assumes i: "P,T\<^sub>r,mxs,mpc,[] \ i,0 :: \#\s" assumes mpc: "mpc = size \s + 1" shows "\ i#is,shift 1 xt [::] \#\s" (*<*) proof - from \s have "\s \ []" by (auto simp: wt_instrs_def) with mpc i have "\ [i],[] [::] [\]@\s" by (simp add: wt_instrs_def) with \s show ?thesis by (fastforce dest: wt_instrs_ext) qed (*>*) lemma (in TC2) wt_instrs_last_incr[trans]: - "\ \ is,xt [::] \s@[\]; P \ \ \' \' \ \ \ is,xt [::] \s@[\']" +assumes wtis: "\ is,xt [::] \s@[\]" and ss: "P \ \ \' \'" +shows "\ is,xt [::] \s@[\']" (*<*) -apply(clarsimp simp add:wt_instrs_def wt_instr_def) -apply(rule conjI) -apply(fastforce) -apply(clarsimp) -apply(rename_tac pc' tau') -apply(erule allE, erule (1) impE) -apply(clarsimp) -apply(drule (1) bspec) -apply(clarsimp) -apply(subgoal_tac "pc' = size \s") -prefer 2 -apply(clarsimp simp:app_def) -apply(drule (1) bspec) -apply(clarsimp) -apply(auto elim!:sup_state_opt_trans) -done +proof - + let ?\s = "\s@[\]" and ?\s' = "\s@[\']" + { fix pc assume pc: "pc< size is" + let ?i = "is!pc" + have app_pc: "app (is ! pc) P mxs T\<^sub>r pc (length ?\s) xt (\s ! pc)" + using wtis pc by(clarsimp simp add:wt_instrs_def wt_instr_def) + then have Apc\': "\pc' \'. (pc',\') \ set (eff ?i P pc xt (?\s!pc)) + \ pc' < length ?\s" + using wtis pc by(fastforce simp add:wt_instrs_def app_def) + have Aepc\': "\pc' \'. (pc',\') \ set (eff ?i P pc xt (?\s!pc)) + \ P \ \' \' ?\s!pc'" + using wtis pc by(fastforce simp add:wt_instrs_def wt_instr_def) + { fix pc1 \1 assume pc\1: "(pc1,\1) \ set (eff ?i P pc xt (?\s'!pc))" + then have epc\': "(pc1,\1) \ set (eff ?i P pc xt (?\s!pc))" + using wtis pc by(simp add:wt_instrs_def) + have "P \ \1 \' ?\s'!pc1" + proof(cases "pc1 < length \s") + case True + then show ?thesis using wtis pc pc\1 + by(fastforce simp add:wt_instrs_def wt_instr_def) + next + case False + then have "pc1 < length ?\s" using Apc\'[OF epc\'] by simp + then have [simp]: "pc1 = size \s" using False by clarsimp + have "P \ \1 \' \" using Aepc\'[OF epc\'] by simp + then have "P \ \1 \' \'" by(rule sup_state_opt_trans[OF _ ss]) + then show ?thesis by simp + qed + } + then have "P,T\<^sub>r,mxs,size ?\s',xt \ is!pc,pc :: ?\s'" using wtis pc + by(clarsimp simp add:wt_instrs_def wt_instr_def) + } + then show ?thesis using wtis by(simp add:wt_instrs_def) +qed (*>*) lemma [iff]: "xcpt_app i P pc mxs [] \" (*<*)by (simp add: xcpt_app_def relevant_entries_def)(*>*) lemma [simp]: "xcpt_eff i P pc \ [] = []" (*<*)by (simp add: xcpt_eff_def relevant_entries_def)(*>*) lemma (in TC2) wt_New: "\ is_class P C; size ST < mxs \ \ \ [New C],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (Class C#ST) E A]" (*<*)by(simp add:wt_defs ty\<^sub>i'_def)(*>*) lemma (in TC2) wt_Cast: "is_class P C \ \ [Checkcast C],[] [::] [ty\<^sub>i' (Class D # ST) E A, ty\<^sub>i' (Class C # ST) E A]" (*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_Push: "\ size ST < mxs; typeof v = Some T \ \ \ [Push v],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (T#ST) E A]" (*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_Pop: "\ [Pop],[] [::] (ty\<^sub>i' (T#ST) E A # ty\<^sub>i' ST E A # \s)" (*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_CmpEq: "\ P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1\ \ \ [CmpEq],[] [::] [ty\<^sub>i' (T\<^sub>2 # T\<^sub>1 # ST) E A, ty\<^sub>i' (Boolean # ST) E A]" (*<*) by(auto simp:ty\<^sub>i'_def wt_defs elim!: refTE not_refTE) (*>*) lemma (in TC2) wt_IAdd: "\ [IAdd],[] [::] [ty\<^sub>i' (Integer#Integer#ST) E A, ty\<^sub>i' (Integer#ST) E A]" (*<*)by(simp add:ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_Load: "\ size ST < mxs; size E \ mxl; i \\ A; i < size E \ \ \ [Load i],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (E!i # ST) E A]" (*<*)by(auto simp add:ty\<^sub>i'_def wt_defs ty\<^sub>l_def hyperset_defs)(*>*) lemma (in TC2) wt_Store: "\ P \ T \ E!i; i < size E; size E \ mxl \ \ \ [Store i],[] [::] [ty\<^sub>i' (T#ST) E A, ty\<^sub>i' ST E (\{i}\ \ A)]" (*<*) by(auto simp:hyperset_defs nth_list_update ty\<^sub>i'_def wt_defs ty\<^sub>l_def intro:list_all2_all_nthI) (*>*) lemma (in TC2) wt_Get: "\ P \ C sees F,NonStatic:T in D \ \ \ [Getfield F D],[] [::] [ty\<^sub>i' (Class C # ST) E A, ty\<^sub>i' (T # ST) E A]" (*<*)by(auto simp: ty\<^sub>i'_def wt_defs dest: sees_field_idemp sees_field_decl_above)(*>*) lemma (in TC2) wt_GetS: "\ size ST < mxs; P \ C sees F,Static:T in D \ \ \ [Getstatic C F D],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (T # ST) E A]" (*<*)by(auto simp: ty\<^sub>i'_def wt_defs dest: sees_field_idemp sees_field_decl_above)(*>*) lemma (in TC2) wt_Put: "\ P \ C sees F,NonStatic:T in D; P \ T' \ T \ \ \ [Putfield F D],[] [::] [ty\<^sub>i' (T' # Class C # ST) E A, ty\<^sub>i' ST E A]" (*<*)by(auto intro: sees_field_idemp sees_field_decl_above simp: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_PutS: "\ P \ C sees F,Static:T in D; P \ T' \ T \ \ \ [Putstatic C F D],[] [::] [ty\<^sub>i' (T' # ST) E A, ty\<^sub>i' ST E A]" (*<*)by(auto intro: sees_field_idemp sees_field_decl_above simp: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_Throw: "\ [Throw],[] [::] [ty\<^sub>i' (Class C # ST) E A, \']" (*<*)by(auto simp: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_IfFalse: "\ 2 \ i; nat i < size \s + 2; P \ ty\<^sub>i' ST E A \' \s ! nat(i - 2) \ \ \ [IfFalse i],[] [::] ty\<^sub>i' (Boolean # ST) E A # ty\<^sub>i' ST E A # \s" (*<*) by(simp add: ty\<^sub>i'_def wt_defs eval_nat_numeral nat_diff_distrib) (*>*) lemma wt_Goto: "\ 0 \ int pc + i; nat (int pc + i) < size \s; size \s \ mpc; P \ \s!pc \' \s ! nat (int pc + i) \ \ P,T,mxs,mpc,[] \ Goto i,pc :: \s" (*<*)by(clarsimp simp add: TC2.wt_defs)(*>*) lemma (in TC2) wt_Invoke: "\ size es = size Ts'; P \ C sees M,NonStatic: Ts\T = m in D; P \ Ts' [\] Ts \ \ \ [Invoke M (size es)],[] [::] [ty\<^sub>i' (rev Ts' @ Class C # ST) E A, ty\<^sub>i' (T#ST) E A]" (*<*)by(fastforce simp add: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_Invokestatic: "\ size ST < mxs; size es = size Ts'; M \ clinit; P \ C sees M,Static: Ts\T = m in D; P \ Ts' [\] Ts \ \ \ [Invokestatic C M (size es)],[] [::] [ty\<^sub>i' (rev Ts' @ ST) E A, ty\<^sub>i' (T#ST) E A]" (*<*)by(fastforce simp add: ty\<^sub>i'_def wt_defs)(*>*) corollary (in TC2) wt_instrs_app3[simp]: "\ \ is\<^sub>2,[] [::] (\' # \s\<^sub>2); \ is\<^sub>1,xt\<^sub>1 [::] \ # \s\<^sub>1 @ [\']; size \s\<^sub>1+1 = size is\<^sub>1\ \ \ (is\<^sub>1 @ is\<^sub>2),xt\<^sub>1 [::] \ # \s\<^sub>1 @ \' # \s\<^sub>2" (*<*)using wt_instrs_app2[where ?xt\<^sub>2.0 = "[]"] by (simp add:shift_def)(*>*) corollary (in TC2) wt_instrs_Cons3[simp]: "\ \s \ []; \ [i],[] [::] [\,\']; \ is,[] [::] \'#\s \ \ \ (i # is),[] [::] \ # \' # \s" (*<*) using wt_instrs_Cons[where ?xt = "[]"] by (simp add:shift_def) (*<*) declare nth_append[simp del] declare [[simproc del: list_to_set_comprehension]] (*>*) lemma (in TC2) wt_instrs_xapp[trans]: - "\ \ is\<^sub>1 @ is\<^sub>2, xt [::] \s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \s\<^sub>2; - \\ \ set \s\<^sub>1. \ST' LT'. \ = Some(ST',LT') \ - size ST \ size ST' \ P \ Some (drop (size ST' - size ST) ST',LT') \' ty\<^sub>i' ST E A; - size is\<^sub>1 = size \s\<^sub>1; is_class P C; size ST < mxs \ \ - \ is\<^sub>1 @ is\<^sub>2, xt @ [(0,size is\<^sub>1 - 1,C,size is\<^sub>1,size ST)] [::] \s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \s\<^sub>2" -(*<*) -apply(simp add:wt_instrs_def) -apply(rule conjI) - apply(clarsimp) - apply arith -apply clarsimp -apply(erule allE, erule (1) impE) -apply(clarsimp simp add: wt_instr_def app_def eff_def) -apply(rule conjI) - apply (thin_tac "\x\ A \ B. P x" for A B P) - apply (thin_tac "\x\ A \ B. P x" for A B P) - apply (clarsimp simp add: xcpt_app_def relevant_entries_def) - apply (simp add: nth_append is_relevant_entry_def split!: if_splits) - apply (drule_tac x="\s\<^sub>1!pc" in bspec) - apply (blast intro: nth_mem) - apply fastforce -apply (rule conjI) - apply clarsimp - apply (erule disjE, blast) - apply (erule disjE, blast) - apply (clarsimp simp add: xcpt_eff_def relevant_entries_def split: if_split_asm) -apply clarsimp -apply (erule disjE, blast) -apply (erule disjE, blast) -apply (clarsimp simp add: xcpt_eff_def relevant_entries_def split: if_split_asm) -apply (simp add: nth_append is_relevant_entry_def split: if_split_asm) - apply (drule_tac x = "\s\<^sub>1!pc" in bspec) - apply (blast intro: nth_mem) - apply (fastforce simp add: ty\<^sub>i'_def) -done +assumes wtis: "\ is\<^sub>1 @ is\<^sub>2, xt [::] \s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \s\<^sub>2" + and P_\s\<^sub>1: "\\ \ set \s\<^sub>1. \ST' LT'. \ = Some(ST',LT') \ + size ST \ size ST' \ P \ Some (drop (size ST' - size ST) ST',LT') \' ty\<^sub>i' ST E A" + and is_\s: "size is\<^sub>1 = size \s\<^sub>1" and PC: "is_class P C" and ST_mxs: "size ST < mxs" +shows "\ is\<^sub>1 @ is\<^sub>2, xt @ [(0,size is\<^sub>1 - 1,C,size is\<^sub>1,size ST)] [::] \s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \s\<^sub>2" +(*<*)(is "\ ?is, xt@[?xte] [::] ?\s" is "\ ?is, ?xt' [::] ?\s") +proof - + have P_\s\<^sub>1': "\\. \ \ set \s\<^sub>1 \ (\ST' LT'. \ = Some(ST',LT') \ + size ST \ size ST' \ P \ Some (drop (size ST' - size ST) ST',LT') \' ty\<^sub>i' ST E A)" + using P_\s\<^sub>1 by fast + have "size ?is < size ?\s" and "pcs ?xt' \ {0..pc. pc< size ?is \ P,T\<^sub>r,mxs,size ?\s,xt \ ?is!pc,pc :: ?\s" + using wtis by(simp_all add:wt_instrs_def) + moreover { + fix pc let ?mpc = "size ?\s" and ?i = "?is!pc" and ?t = "?\s!pc" + assume "pc< size ?is" + then have wti: "P,T\<^sub>r,mxs,?mpc,xt \ ?i,pc :: ?\s" by(rule P_pc) + then have app: "app ?i P mxs T\<^sub>r pc ?mpc xt ?t" + and eff_ss: "\pc' \'. (pc',\') \ set (eff ?i P pc xt (?\s!pc)) + \ P \ \' \' ?\s!pc'" + by(fastforce simp add: wt_instr_def)+ + have "app ?i P mxs T\<^sub>r pc ?mpc ?xt' ?t + \ (\(pc',\') \ set (eff ?i P pc ?xt' ?t). P \ \' \' ?\s!pc')" + proof (cases ?t) + case (Some \) + obtain ST' LT' where \: "\ = (ST', LT')" by(cases \) simp + have app\<^sub>i: "app\<^sub>i (?i,P,pc,mxs,T\<^sub>r,\)" and xcpt_app: "xcpt_app ?i P pc mxs xt \" + and eff_pc: "\pc' \'. (pc',\') \ set (eff ?i P pc xt ?t) \ pc' < ?mpc" + using Some app by(fastforce simp add: app_def)+ + have "xcpt_app ?i P pc mxs ?xt' \" + proof(cases "pc < length \s\<^sub>1 - 1") + case False then show ?thesis using Some \ is_\s xcpt_app + by (clarsimp simp: xcpt_app_def relevant_entries_def + is_relevant_entry_def) + next + case True + then have True': "pc < length \s\<^sub>1" by simp + then have "\s\<^sub>1 ! pc = ?t" by(fastforce simp: nth_append) + moreover have \s\<^sub>1_pc: "\s\<^sub>1 ! pc \ set \s\<^sub>1" by(rule nth_mem[OF True']) + ultimately show ?thesis + using Some \ PC ST_mxs xcpt_app P_\s\<^sub>1'[OF \s\<^sub>1_pc] + by (simp add: xcpt_app_def relevant_entries_def) + qed + moreover { + fix pc' \' assume efft: "(pc',\') \ set (eff ?i P pc ?xt' ?t)" + have "pc' < ?mpc \ P \ \' \' ?\s!pc'" (is "?P1 \ ?P2") + proof(cases "(pc',\') \ set (eff ?i P pc xt ?t)") + case True + have ?P1 using True by(rule eff_pc) + moreover have ?P2 using True by(rule eff_ss) + ultimately show ?thesis by simp + next + case False + then have xte: "(pc',\') \ set (xcpt_eff ?i P pc \ [?xte])" + using efft Some by(clarsimp simp: eff_def) + then have ?P1 using Some \ is_\s + by(clarsimp simp: xcpt_eff_def relevant_entries_def split: if_split_asm) + moreover have ?P2 + proof(cases "pc < length \s\<^sub>1 - 1") + case False': False + then show ?thesis using False Some \ xte is_\s + by(simp add: xcpt_eff_def relevant_entries_def is_relevant_entry_def) + next + case True + then have True': "pc < length \s\<^sub>1" by simp + have \s\<^sub>1_pc: "\s\<^sub>1 ! pc \ set \s\<^sub>1" by(rule nth_mem[OF True']) + have "P \ \(Class C # drop (length ST' - length ST) ST', LT')\ + \' ty\<^sub>i' (Class C # ST) E A" + using True' Some \ P_\s\<^sub>1'[OF \s\<^sub>1_pc] + by (fastforce simp: nth_append ty\<^sub>i'_def) + then show ?thesis using \ xte is_\s + by(simp add: xcpt_eff_def relevant_entries_def split: if_split_asm) + qed + ultimately show ?thesis by simp + qed + } + ultimately show ?thesis using Some app\<^sub>i by(fastforce simp add: app_def) + qed simp + then have "P,T\<^sub>r,mxs,size ?\s,?xt' \ ?is!pc,pc :: ?\s" + by(simp add: wt_instr_def) + } + ultimately show ?thesis by(simp add:wt_instrs_def) +qed declare [[simproc add: list_to_set_comprehension]] declare nth_append[simp] (*>*) lemma drop_Cons_Suc: "\xs. drop n xs = y#ys \ drop (Suc n) xs = ys" - apply (induct n) - apply simp - apply (simp add: drop_Suc) - done +proof(induct n) + case (Suc n) then show ?case by(simp add: drop_Suc) +qed simp lemma drop_mess: - "\Suc (length xs\<^sub>0) \ length xs; drop (length xs - Suc (length xs\<^sub>0)) xs = x # xs\<^sub>0\ - \ drop (length xs - length xs\<^sub>0) xs = xs\<^sub>0" -apply (cases xs) - apply simp -apply (simp add: Suc_diff_le) -apply (case_tac "length list - length xs\<^sub>0") - apply simp -apply (simp add: drop_Cons_Suc) -done +assumes "Suc (length xs\<^sub>0) \ length xs" + and "drop (length xs - Suc (length xs\<^sub>0)) xs = x # xs\<^sub>0" +shows "drop (length xs - length xs\<^sub>0) xs = xs\<^sub>0" +using assms proof(cases xs) + case (Cons a list) then show ?thesis using assms + proof(cases "length list - length xs\<^sub>0") + case Suc then show ?thesis using Cons assms + by (simp add: Suc_diff_le drop_Cons_Suc) + qed simp +qed simp (*<*) declare (in TC0) after_def[simp] pair_eq_ty\<^sub>i'_conv[simp] (*>*) lemma (in TC1) compT_ST_prefix: "\E A ST\<^sub>0. \(ST,LT)\ \ set(compT E A ST\<^sub>0 e) \ size ST\<^sub>0 \ size ST \ drop (size ST - size ST\<^sub>0) ST = ST\<^sub>0" and "\E A ST\<^sub>0. \(ST,LT)\ \ set(compTs E A ST\<^sub>0 es) \ size ST\<^sub>0 \ size ST \ drop (size ST - size ST\<^sub>0) ST = ST\<^sub>0" (*<*) proof(induct e and es rule: compT.induct compTs.induct) case (FAss e\<^sub>1 F D e\<^sub>2) moreover { let ?ST\<^sub>0 = "ty E e\<^sub>1 # ST\<^sub>0" fix A assume "\(ST, LT)\ \ set (compT E A ?ST\<^sub>0 e\<^sub>2)" with FAss have "length ?ST\<^sub>0 \ length ST \ drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast hence ?case by (clarsimp simp add: drop_mess) } ultimately show ?case by auto next case TryCatch thus ?case by auto next case Block thus ?case by auto next case Seq thus ?case by auto next case While thus ?case by auto next case Cond thus ?case by auto next case (Call e M es) moreover { let ?ST\<^sub>0 = "ty E e # ST\<^sub>0" fix A assume "\(ST, LT)\ \ set (compTs E A ?ST\<^sub>0 es)" with Call have "length ?ST\<^sub>0 \ length ST \ drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast hence ?case by (clarsimp simp add: drop_mess) } ultimately show ?case by auto next case (Cons_exp e es) moreover { let ?ST\<^sub>0 = "ty E e # ST\<^sub>0" fix A assume "\(ST, LT)\ \ set (compTs E A ?ST\<^sub>0 es)" with Cons_exp have "length ?ST\<^sub>0 \ length ST \ drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast hence ?case by (clarsimp simp add: drop_mess) } ultimately show ?case by auto next case (BinOp e\<^sub>1 bop e\<^sub>2) moreover { let ?ST\<^sub>0 = "ty E e\<^sub>1 # ST\<^sub>0" fix A assume "\(ST, LT)\ \ set (compT E A ?ST\<^sub>0 e\<^sub>2)" with BinOp have "length ?ST\<^sub>0 \ length ST \ drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast hence ?case by (clarsimp simp add: drop_mess) } ultimately show ?case by auto next case new thus ?case by auto next case Val thus ?case by auto next case Cast thus ?case by auto next case Var thus ?case by auto next case LAss thus ?case by auto next case throw thus ?case by auto next case FAcc thus ?case by auto next case SFAcc thus ?case by auto next case SFAss thus ?case by auto next case (SCall C M es) thus ?case by auto next case INIT thus ?case by auto next case RI thus ?case by auto next case Nil_exp thus ?case by auto qed declare (in TC0) after_def[simp del] pair_eq_ty\<^sub>i'_conv[simp del] (*>*) (* FIXME *) lemma fun_of_simp [simp]: "fun_of S x y = ((x,y) \ S)" (*<*) by (simp add: fun_of_def)(*>*) theorem (in TC2) compT_wt_instrs: "\E T A ST. \ P,E \\<^sub>1 e :: T; \ e A; \ e (size E); size ST + max_stack e \ mxs; size E + max_vars e \ mxl \ \ \ compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ty\<^sub>i' ST E A # compT E A ST e @ [after E A ST e]" (*<*)(is "\E T A ST. PROP ?P e E T A ST")(*>*) and "\E Ts A ST. \ P,E \\<^sub>1 es[::]Ts; \s es A; \s es (size E); size ST + max_stacks es \ mxs; size E + max_varss es \ mxl \ \ let \s = ty\<^sub>i' ST E A # compTs E A ST es in \ compEs\<^sub>2 es,compxEs\<^sub>2 es 0 (size ST) [::] \s \ last \s = ty\<^sub>i' (rev Ts @ ST) E (A \ \s es)" (*<*) (is "\E Ts A ST. PROP ?Ps es E Ts A ST") proof(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) case (TryCatch e\<^sub>1 C i e\<^sub>2) hence [simp]: "i = size E" by simp have wt\<^sub>1: "P,E \\<^sub>1 e\<^sub>1 :: T" and wt\<^sub>2: "P,E@[Class C] \\<^sub>1 e\<^sub>2 :: T" and "class": "is_class P C" using TryCatch by auto let ?A\<^sub>1 = "A \ \ e\<^sub>1" let ?A\<^sub>i = "A \ \{i}\" let ?E\<^sub>i = "E @ [Class C]" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>1 = "compT E A ST e\<^sub>1" let ?\\<^sub>1 = "ty\<^sub>i' (T#ST) E ?A\<^sub>1" let ?\\<^sub>2 = "ty\<^sub>i' (Class C#ST) E A" let ?\\<^sub>3 = "ty\<^sub>i' ST ?E\<^sub>i ?A\<^sub>i" let ?\s\<^sub>2 = "compT ?E\<^sub>i ?A\<^sub>i ST e\<^sub>2" let ?\\<^sub>2' = "ty\<^sub>i' (T#ST) ?E\<^sub>i (?A\<^sub>i \ \ e\<^sub>2)" let ?\' = "ty\<^sub>i' (T#ST) E (A \ \ e\<^sub>1 \ (\ e\<^sub>2 \ i))" let ?go = "Goto (int(size(compE\<^sub>2 e\<^sub>2)) + 2)" have "PROP ?P e\<^sub>2 ?E\<^sub>i T ?A\<^sub>i ST" by fact hence "\ compE\<^sub>2 e\<^sub>2,compxE\<^sub>2 e\<^sub>2 0 (size ST) [::] (?\\<^sub>3 # ?\s\<^sub>2) @ [?\\<^sub>2']" using TryCatch.prems by(auto simp:after_def) also have "?A\<^sub>i \ \ e\<^sub>2 = (A \ \ e\<^sub>2) \ \{size E}\" by(fastforce simp:hyperset_defs) also have "P \ ty\<^sub>i' (T#ST) ?E\<^sub>i \ \' ty\<^sub>i' (T#ST) E (A \ \ e\<^sub>2)" by(simp add:hyperset_defs ty\<^sub>l_incr ty\<^sub>i'_def) also have "P \ \ \' ty\<^sub>i' (T#ST) E (A \ \ e\<^sub>1 \ (\ e\<^sub>2 \ i))" by(auto intro!: ty\<^sub>l_antimono simp:hyperset_defs ty\<^sub>i'_def) also have "(?\\<^sub>3 # ?\s\<^sub>2) @ [?\'] = ?\\<^sub>3 # ?\s\<^sub>2 @ [?\']" by simp also have "\ [Store i],[] [::] ?\\<^sub>2 # [] @ [?\\<^sub>3]" using TryCatch.prems by(auto simp:nth_list_update wt_defs ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth hyperset_defs) also have "[] @ (?\\<^sub>3 # ?\s\<^sub>2 @ [?\']) = (?\\<^sub>3 # ?\s\<^sub>2 @ [?\'])" by simp also have "P,T\<^sub>r,mxs,size(compE\<^sub>2 e\<^sub>2)+3,[] \ ?go,0 :: ?\\<^sub>1#?\\<^sub>2#?\\<^sub>3#?\s\<^sub>2 @ [?\']" using wt\<^sub>2 by (auto simp: hyperset_defs ty\<^sub>i'_def wt_defs nth_Cons nat_add_distrib fun_of_def intro: ty\<^sub>l_antimono list_all2_refl split:nat.split) also have "\ compE\<^sub>2 e\<^sub>1,compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\ # ?\s\<^sub>1 @ [?\\<^sub>1]" using TryCatch by(auto simp:after_def) also have "?\ # ?\s\<^sub>1 @ ?\\<^sub>1 # ?\\<^sub>2 # ?\\<^sub>3 # ?\s\<^sub>2 @ [?\'] = (?\ # ?\s\<^sub>1 @ [?\\<^sub>1]) @ ?\\<^sub>2 # ?\\<^sub>3 # ?\s\<^sub>2 @ [?\']" by simp also have "compE\<^sub>2 e\<^sub>1 @ ?go # [Store i] @ compE\<^sub>2 e\<^sub>2 = (compE\<^sub>2 e\<^sub>1 @ [?go]) @ (Store i # compE\<^sub>2 e\<^sub>2)" by simp also let "?Q \" = "\ST' LT'. \ = \(ST', LT')\ \ size ST \ size ST' \ P \ Some (drop (size ST' - size ST) ST',LT') \' ty\<^sub>i' ST E A" { have "?Q (ty\<^sub>i' ST E A)" by (clarsimp simp add: ty\<^sub>i'_def) moreover have "?Q (ty\<^sub>i' (T # ST) E ?A\<^sub>1)" by (fastforce simp add: ty\<^sub>i'_def hyperset_defs intro!: ty\<^sub>l_antimono) moreover have "\\. \ \ set (compT E A ST e\<^sub>1) \ ?Q \" using TryCatch.prems by clarsimp (frule compT_ST_prefix, fastforce dest!: compT_LT_prefix simp add: ty\<^sub>i'_def) ultimately have "\\\set (ty\<^sub>i' ST E A # compT E A ST e\<^sub>1 @ [ty\<^sub>i' (T # ST) E ?A\<^sub>1]). ?Q \" by auto } also from TryCatch.prems max_stack1[OF wt\<^sub>1] have "size ST + 1 \ mxs" by auto ultimately show ?case using wt\<^sub>1 wt\<^sub>2 TryCatch.prems "class" by (simp add:after_def) next case new thus ?case by(auto simp add:after_def wt_New) next case (BinOp e\<^sub>1 bop e\<^sub>2) let ?op = "case bop of Eq \ [CmpEq] | Add \ [IAdd]" have T: "P,E \\<^sub>1 e\<^sub>1 \bop\ e\<^sub>2 :: T" by fact then obtain T\<^sub>1 T\<^sub>2 where T\<^sub>1: "P,E \\<^sub>1 e\<^sub>1 :: T\<^sub>1" and T\<^sub>2: "P,E \\<^sub>1 e\<^sub>2 :: T\<^sub>2" and bopT: "case bop of Eq \ (P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1) \ T = Boolean | Add \ T\<^sub>1 = Integer \ T\<^sub>2 = Integer \ T = Integer" by auto let ?A\<^sub>1 = "A \ \ e\<^sub>1" let ?A\<^sub>2 = "?A\<^sub>1 \ \ e\<^sub>2" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>1 = "compT E A ST e\<^sub>1" let ?\\<^sub>1 = "ty\<^sub>i' (T\<^sub>1#ST) E ?A\<^sub>1" let ?\s\<^sub>2 = "compT E ?A\<^sub>1 (T\<^sub>1#ST) e\<^sub>2" let ?\\<^sub>2 = "ty\<^sub>i' (T\<^sub>2#T\<^sub>1#ST) E ?A\<^sub>2" let ?\' = "ty\<^sub>i' (T#ST) E ?A\<^sub>2" from bopT have "\ ?op,[] [::] [?\\<^sub>2,?\']" by (cases bop) (auto simp add: wt_CmpEq wt_IAdd) also have "PROP ?P e\<^sub>2 E T\<^sub>2 ?A\<^sub>1 (T\<^sub>1#ST)" by fact with BinOp.prems T\<^sub>2 have "\ compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size (T\<^sub>1#ST)) [::] ?\\<^sub>1#?\s\<^sub>2@[?\\<^sub>2]" by (auto simp: after_def) also from BinOp T\<^sub>1 have "\ compE\<^sub>2 e\<^sub>1, compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\#?\s\<^sub>1@[?\\<^sub>1]" by (auto simp: after_def) finally show ?case using T T\<^sub>1 T\<^sub>2 by (simp add: after_def hyperUn_assoc) next case (Cons_exp e es) have "P,E \\<^sub>1 e # es [::] Ts" by fact then obtain T\<^sub>e Ts' where T\<^sub>e: "P,E \\<^sub>1 e :: T\<^sub>e" and Ts': "P,E \\<^sub>1 es [::] Ts'" and Ts: "Ts = T\<^sub>e#Ts'" by auto let ?A\<^sub>e = "A \ \ e" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>e = "compT E A ST e" let ?\\<^sub>e = "ty\<^sub>i' (T\<^sub>e#ST) E ?A\<^sub>e" let ?\s' = "compTs E ?A\<^sub>e (T\<^sub>e#ST) es" let ?\s = "?\ # ?\s\<^sub>e @ (?\\<^sub>e # ?\s')" have Ps: "PROP ?Ps es E Ts' ?A\<^sub>e (T\<^sub>e#ST)" by fact with Cons_exp.prems T\<^sub>e Ts' have "\ compEs\<^sub>2 es, compxEs\<^sub>2 es 0 (size (T\<^sub>e#ST)) [::] ?\\<^sub>e#?\s'" by (simp add: after_def) also from Cons_exp T\<^sub>e have "\ compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ?\#?\s\<^sub>e@[?\\<^sub>e]" by (auto simp: after_def) moreover from Ps Cons_exp.prems T\<^sub>e Ts' Ts have "last ?\s = ty\<^sub>i' (rev Ts@ST) E (?A\<^sub>e \ \s es)" by simp ultimately show ?case using T\<^sub>e by (simp add: after_def hyperUn_assoc) next case (FAss e\<^sub>1 F D e\<^sub>2) hence Void: "P,E \\<^sub>1 e\<^sub>1\F{D} := e\<^sub>2 :: Void" by auto then obtain C T T' where C: "P,E \\<^sub>1 e\<^sub>1 :: Class C" and sees: "P \ C sees F,NonStatic:T in D" and T': "P,E \\<^sub>1 e\<^sub>2 :: T'" and T'_T: "P \ T' \ T" by auto let ?A\<^sub>1 = "A \ \ e\<^sub>1" let ?A\<^sub>2 = "?A\<^sub>1 \ \ e\<^sub>2" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>1 = "compT E A ST e\<^sub>1" let ?\\<^sub>1 = "ty\<^sub>i' (Class C#ST) E ?A\<^sub>1" let ?\s\<^sub>2 = "compT E ?A\<^sub>1 (Class C#ST) e\<^sub>2" let ?\\<^sub>2 = "ty\<^sub>i' (T'#Class C#ST) E ?A\<^sub>2" let ?\\<^sub>3 = "ty\<^sub>i' ST E ?A\<^sub>2" let ?\' = "ty\<^sub>i' (Void#ST) E ?A\<^sub>2" from FAss.prems sees T'_T have "\ [Putfield F D,Push Unit],[] [::] [?\\<^sub>2,?\\<^sub>3,?\']" by (fastforce simp add: wt_Push wt_Put) also have "PROP ?P e\<^sub>2 E T' ?A\<^sub>1 (Class C#ST)" by fact with FAss.prems T' have "\ compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size ST+1) [::] ?\\<^sub>1#?\s\<^sub>2@[?\\<^sub>2]" by (auto simp add: after_def hyperUn_assoc) also from FAss C have "\ compE\<^sub>2 e\<^sub>1, compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\#?\s\<^sub>1@[?\\<^sub>1]" by (auto simp add: after_def) finally show ?case using Void C T' by (simp add: after_def hyperUn_assoc) next case (SFAss C F D e\<^sub>2) hence Void: "P,E \\<^sub>1 C\\<^sub>sF{D} := e\<^sub>2 :: Void" by auto then obtain T T' where sees: "P \ C sees F,Static:T in D" and T': "P,E \\<^sub>1 e\<^sub>2 :: T'" and T'_T: "P \ T' \ T" by auto let ?A\<^sub>2 = "A \ \ e\<^sub>2" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>2 = "compT E A ST e\<^sub>2" let ?\\<^sub>2 = "ty\<^sub>i' (T'#ST) E ?A\<^sub>2" let ?\\<^sub>3 = "ty\<^sub>i' ST E ?A\<^sub>2" let ?\' = "ty\<^sub>i' (Void#ST) E ?A\<^sub>2" from SFAss.prems sees T'_T max_stack1[OF T'] have "\ [Putstatic C F D,Push Unit],[] [::] [?\\<^sub>2,?\\<^sub>3,?\']" by (fastforce simp add: wt_Push wt_PutS) also have "PROP ?P e\<^sub>2 E T' A ST" by fact with SFAss.prems T' have "\ compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size ST) [::] ?\#?\s\<^sub>2@[?\\<^sub>2]" by (auto simp add: after_def hyperUn_assoc) finally show ?case using Void T' by (simp add: after_def hyperUn_assoc) next case Val thus ?case by(auto simp:after_def wt_Push) next case Cast thus ?case by (auto simp:after_def wt_Cast) next case (Block i T\<^sub>i e) let ?\s = "ty\<^sub>i' ST E A # compT (E @ [T\<^sub>i]) (A\i) ST e" have IH: "PROP ?P e (E@[T\<^sub>i]) T (A\i) ST" by fact hence "\ compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ?\s @ [ty\<^sub>i' (T#ST) (E@[T\<^sub>i]) (A\(size E) \ \ e)]" using Block.prems by (auto simp add: after_def) also have "P \ ty\<^sub>i' (T # ST) (E@[T\<^sub>i]) (A \ size E \ \ e) \' ty\<^sub>i' (T # ST) (E@[T\<^sub>i]) ((A \ \ e) \ size E)" by(auto simp add:hyperset_defs intro: ty\<^sub>i'_antimono) also have "\ = ty\<^sub>i' (T # ST) E (A \ \ e)" by simp also have "P \ \ \' ty\<^sub>i' (T # ST) E (A \ (\ e \ i))" by(auto simp add:hyperset_defs intro: ty\<^sub>i'_antimono) finally show ?case using Block.prems by(simp add: after_def) next case Var thus ?case by(auto simp:after_def wt_Load) next case FAcc thus ?case by(auto simp:after_def wt_Get) next case SFAcc thus ?case by(auto simp: after_def wt_GetS) next case (LAss i e) then obtain T' where wt: "P,E \\<^sub>1 e :: T'" by auto show ?case using max_stack1[OF wt] LAss by(auto simp: hyper_insert_comm after_def wt_Store wt_Push) next case Nil_exp thus ?case by auto next case throw thus ?case by(auto simp add: after_def wt_Throw) next case (While e c) obtain Tc where wte: "P,E \\<^sub>1 e :: Boolean" and wtc: "P,E \\<^sub>1 c :: Tc" and [simp]: "T = Void" using While by auto have [simp]: "ty E (while (e) c) = Void" using While by simp let ?A\<^sub>0 = "A \ \ e" let ?A\<^sub>1 = "?A\<^sub>0 \ \ c" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>e = "compT E A ST e" let ?\\<^sub>e = "ty\<^sub>i' (Boolean#ST) E ?A\<^sub>0" let ?\\<^sub>1 = "ty\<^sub>i' ST E ?A\<^sub>0" let ?\s\<^sub>c = "compT E ?A\<^sub>0 ST c" let ?\\<^sub>c = "ty\<^sub>i' (Tc#ST) E ?A\<^sub>1" let ?\\<^sub>2 = "ty\<^sub>i' ST E ?A\<^sub>1" let ?\' = "ty\<^sub>i' (Void#ST) E ?A\<^sub>0" let ?\s = "(?\ # ?\s\<^sub>e @ [?\\<^sub>e]) @ ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c, ?\\<^sub>2, ?\\<^sub>1, ?\']" have "\ [],[] [::] [] @ ?\s" by(simp add:wt_instrs_def) also have "PROP ?P e E Boolean A ST" by fact hence "\ compE\<^sub>2 e,compxE\<^sub>2 e 0 (size ST) [::] ?\ # ?\s\<^sub>e @ [?\\<^sub>e]" using While.prems by (auto simp:after_def) also have "[] @ ?\s = (?\ # ?\s\<^sub>e) @ ?\\<^sub>e # ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c,?\\<^sub>2,?\\<^sub>1,?\']" by simp also let ?n\<^sub>e = "size(compE\<^sub>2 e)" let ?n\<^sub>c = "size(compE\<^sub>2 c)" let ?if = "IfFalse (int ?n\<^sub>c + 3)" have "\ [?if],[] [::] ?\\<^sub>e # ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c, ?\\<^sub>2, ?\\<^sub>1, ?\']" using wtc by(simp add: wt_instr_Cons wt_instr_append wt_IfFalse nat_add_distrib split: nat_diff_split) also have "(?\ # ?\s\<^sub>e) @ (?\\<^sub>e # ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c, ?\\<^sub>2, ?\\<^sub>1, ?\']) = ?\s" by simp also have "PROP ?P c E Tc ?A\<^sub>0 ST" by fact hence "\ compE\<^sub>2 c,compxE\<^sub>2 c 0 (size ST) [::] ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c]" using While.prems wtc by (auto simp:after_def) also have "?\s = (?\ # ?\s\<^sub>e @ [?\\<^sub>e,?\\<^sub>1] @ ?\s\<^sub>c) @ [?\\<^sub>c,?\\<^sub>2,?\\<^sub>1,?\']" by simp also have "\ [Pop],[] [::] [?\\<^sub>c, ?\\<^sub>2]" by(simp add:wt_Pop) also have "(?\ # ?\s\<^sub>e @ [?\\<^sub>e,?\\<^sub>1] @ ?\s\<^sub>c) @ [?\\<^sub>c,?\\<^sub>2,?\\<^sub>1,?\'] = ?\s" by simp also let ?go = "Goto (-int(?n\<^sub>c+?n\<^sub>e+2))" have "P \ ?\\<^sub>2 \' ?\" by(fastforce intro: ty\<^sub>i'_antimono simp: hyperset_defs) hence "P,T\<^sub>r,mxs,size ?\s,[] \ ?go,?n\<^sub>e+?n\<^sub>c+2 :: ?\s" using wte wtc by(simp add: wt_Goto split: nat_diff_split) also have "?\s = (?\ # ?\s\<^sub>e @ [?\\<^sub>e,?\\<^sub>1] @ ?\s\<^sub>c @ [?\\<^sub>c, ?\\<^sub>2]) @ [?\\<^sub>1, ?\']" by simp also have "\ [Push Unit],[] [::] [?\\<^sub>1,?\']" using While.prems max_stack1[OF wtc] by(auto simp add:wt_Push) finally show ?case using wtc wte by (simp add:after_def) next case (Cond e e\<^sub>1 e\<^sub>2) obtain T\<^sub>1 T\<^sub>2 where wte: "P,E \\<^sub>1 e :: Boolean" and wt\<^sub>1: "P,E \\<^sub>1 e\<^sub>1 :: T\<^sub>1" and wt\<^sub>2: "P,E \\<^sub>1 e\<^sub>2 :: T\<^sub>2" and sub\<^sub>1: "P \ T\<^sub>1 \ T" and sub\<^sub>2: "P \ T\<^sub>2 \ T" using Cond by auto have [simp]: "ty E (if (e) e\<^sub>1 else e\<^sub>2) = T" using Cond by simp let ?A\<^sub>0 = "A \ \ e" let ?A\<^sub>2 = "?A\<^sub>0 \ \ e\<^sub>2" let ?A\<^sub>1 = "?A\<^sub>0 \ \ e\<^sub>1" let ?A' = "?A\<^sub>0 \ \ e\<^sub>1 \ \ e\<^sub>2" let ?\\<^sub>2 = "ty\<^sub>i' ST E ?A\<^sub>0" let ?\' = "ty\<^sub>i' (T#ST) E ?A'" let ?\s\<^sub>2 = "compT E ?A\<^sub>0 ST e\<^sub>2" have "PROP ?P e\<^sub>2 E T\<^sub>2 ?A\<^sub>0 ST" by fact hence "\ compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size ST) [::] (?\\<^sub>2#?\s\<^sub>2) @ [ty\<^sub>i' (T\<^sub>2#ST) E ?A\<^sub>2]" using Cond.prems wt\<^sub>2 by(auto simp add:after_def) also have "P \ ty\<^sub>i' (T\<^sub>2#ST) E ?A\<^sub>2 \' ?\'" using sub\<^sub>2 by(auto simp add: hyperset_defs ty\<^sub>i'_def intro!: ty\<^sub>l_antimono) also let ?\\<^sub>3 = "ty\<^sub>i' (T\<^sub>1 # ST) E ?A\<^sub>1" let ?g\<^sub>2 = "Goto(int (size (compE\<^sub>2 e\<^sub>2) + 1))" from sub\<^sub>1 have "P,T\<^sub>r,mxs,size(compE\<^sub>2 e\<^sub>2)+2,[] \ ?g\<^sub>2,0 :: ?\\<^sub>3#(?\\<^sub>2#?\s\<^sub>2)@[?\']" using wt\<^sub>2 by(auto simp: hyperset_defs wt_defs nth_Cons ty\<^sub>i'_def split:nat.split intro!: ty\<^sub>l_antimono) also let ?\s\<^sub>1 = "compT E ?A\<^sub>0 ST e\<^sub>1" have "PROP ?P e\<^sub>1 E T\<^sub>1 ?A\<^sub>0 ST" by fact hence "\ compE\<^sub>2 e\<^sub>1,compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\\<^sub>2 # ?\s\<^sub>1 @ [?\\<^sub>3]" using Cond.prems wt\<^sub>1 by(auto simp add:after_def) also let ?\s\<^sub>1\<^sub>2 = "?\\<^sub>2 # ?\s\<^sub>1 @ ?\\<^sub>3 # (?\\<^sub>2 # ?\s\<^sub>2) @ [?\']" let ?\\<^sub>1 = "ty\<^sub>i' (Boolean#ST) E ?A\<^sub>0" let ?g\<^sub>1 = "IfFalse(int (size (compE\<^sub>2 e\<^sub>1) + 2))" let ?code = "compE\<^sub>2 e\<^sub>1 @ ?g\<^sub>2 # compE\<^sub>2 e\<^sub>2" have "\ [?g\<^sub>1],[] [::] [?\\<^sub>1] @ ?\s\<^sub>1\<^sub>2" using wt\<^sub>1 by(simp add: wt_IfFalse nat_add_distrib split:nat_diff_split) also (wt_instrs_ext2) have "[?\\<^sub>1] @ ?\s\<^sub>1\<^sub>2 = ?\\<^sub>1 # ?\s\<^sub>1\<^sub>2" by simp also let ?\ = "ty\<^sub>i' ST E A" have "PROP ?P e E Boolean A ST" by fact hence "\ compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ?\ # compT E A ST e @ [?\\<^sub>1]" using Cond.prems wte by(auto simp add:after_def) finally show ?case using wte wt\<^sub>1 wt\<^sub>2 by(simp add:after_def hyperUn_assoc) next case (Call e M es) obtain C D Ts m Ts' where C: "P,E \\<^sub>1 e :: Class C" and "method": "P \ C sees M,NonStatic:Ts \ T = m in D" and wtes: "P,E \\<^sub>1 es [::] Ts'" and subs: "P \ Ts' [\] Ts" using Call.prems by auto from wtes have same_size: "size es = size Ts'" by(rule WTs\<^sub>1_same_size) let ?A\<^sub>0 = "A \ \ e" let ?A\<^sub>1 = "?A\<^sub>0 \ \s es" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>e = "compT E A ST e" let ?\\<^sub>e = "ty\<^sub>i' (Class C # ST) E ?A\<^sub>0" let ?\s\<^sub>e\<^sub>s = "compTs E ?A\<^sub>0 (Class C # ST) es" let ?\\<^sub>1 = "ty\<^sub>i' (rev Ts' @ Class C # ST) E ?A\<^sub>1" let ?\' = "ty\<^sub>i' (T # ST) E ?A\<^sub>1" have "\ [Invoke M (size es)],[] [::] [?\\<^sub>1,?\']" by(rule wt_Invoke[OF same_size "method" subs]) also have "PROP ?Ps es E Ts' ?A\<^sub>0 (Class C # ST)" by fact hence "\ compEs\<^sub>2 es,compxEs\<^sub>2 es 0 (size ST+1) [::] ?\\<^sub>e # ?\s\<^sub>e\<^sub>s" "last (?\\<^sub>e # ?\s\<^sub>e\<^sub>s) = ?\\<^sub>1" using Call.prems wtes by(auto simp add:after_def) also have "(?\\<^sub>e # ?\s\<^sub>e\<^sub>s) @ [?\'] = ?\\<^sub>e # ?\s\<^sub>e\<^sub>s @ [?\']" by simp also have "\ compE\<^sub>2 e,compxE\<^sub>2 e 0 (size ST) [::] ?\ # ?\s\<^sub>e @ [?\\<^sub>e]" using Call C by(auto simp add:after_def) finally show ?case using Call.prems C wtes by(simp add:after_def hyperUn_assoc) next case (SCall C M es) obtain D Ts m Ts' where "method": "P \ C sees M,Static:Ts \ T = m in D" and wtes: "P,E \\<^sub>1 es [::] Ts'" and subs: "P \ Ts' [\] Ts" using SCall.prems by auto from SCall.prems(1) have nclinit: "M \ clinit" by auto from wtes have same_size: "size es = size Ts'" by(rule WTs\<^sub>1_same_size) have mxs: "length ST < mxs" using WT\<^sub>1_nsub_RI[OF SCall.prems(1)] SCall.prems(4) by simp let ?A\<^sub>1 = "A \ \s es" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>e\<^sub>s = "compTs E A ST es" let ?\\<^sub>1 = "ty\<^sub>i' (rev Ts' @ ST) E ?A\<^sub>1" let ?\' = "ty\<^sub>i' (T # ST) E ?A\<^sub>1" have "\ [Invokestatic C M (size es)],[] [::] [?\\<^sub>1,?\']" by(rule wt_Invokestatic[OF mxs same_size nclinit "method" subs]) also have "PROP ?Ps es E Ts' A ST" by fact hence "\ compEs\<^sub>2 es,compxEs\<^sub>2 es 0 (size ST) [::] ?\ # ?\s\<^sub>e\<^sub>s" "last (?\ # ?\s\<^sub>e\<^sub>s) = ?\\<^sub>1" using SCall.prems wtes by(auto simp add:after_def) also have "(?\ # ?\s\<^sub>e\<^sub>s) @ [?\'] = ?\ # ?\s\<^sub>e\<^sub>s @ [?\']" by simp finally show ?case using SCall.prems wtes by(simp add:after_def hyperUn_assoc) next case Seq thus ?case by(auto simp:after_def) (fastforce simp:wt_Push wt_Pop hyperUn_assoc intro:wt_instrs_app2 wt_instrs_Cons) next case (INIT C Cs b e) have "P,E \\<^sub>1 INIT C (Cs,b) \ e :: T" by fact thus ?case using WT\<^sub>1_nsub_RI by simp next case (RI C e' Cs e) have "P,E \\<^sub>1 RI (C,e') ; Cs \ e :: T" by fact thus ?case using WT\<^sub>1_nsub_RI by simp qed (*>*) lemma [simp]: "types (compP f P) = types P" (*<*)by auto(*>*) lemma [simp]: "states (compP f P) mxs mxl = states P mxs mxl" (*<*)by (simp add: JVM_states_unfold)(*>*) lemma [simp]: "app\<^sub>i (i, compP f P, pc, mpc, T, \) = app\<^sub>i (i, P, pc, mpc, T, \)" -(*<*) - apply (cases \) - apply (cases i) - apply auto -\ \ Invoke \ - apply (fastforce dest!: sees_method_compPD) - apply (force dest: sees_method_compP) -\ \ Invokestatic \ - apply (force dest!: sees_method_compPD) - apply (force dest: sees_method_compP) - done +(*<*)(is "?A = ?B") +proof - + obtain ST LT where \: "\ = (ST, LT)" by(cases \) simp + then show ?thesis proof(cases i) + case Invoke show ?thesis + proof(rule iffI) + assume ?A then show ?B using Invoke \ + by auto (fastforce dest!: sees_method_compPD) + next + assume ?B then show ?A using Invoke \ + by auto (force dest: sees_method_compP) + qed + next + case (Invokestatic x111 x112 x113) show ?thesis + proof(rule iffI) + assume ?A then show ?B using Invokestatic \ + by auto (fastforce dest!: sees_method_compPD) + next + assume ?B then show ?A using Invokestatic \ + by auto (force dest: sees_method_compP) + qed + qed auto +qed (*>*) lemma [simp]: "is_relevant_entry (compP f P) i = is_relevant_entry P i" (*<*) - apply (rule ext)+ - apply (unfold is_relevant_entry_def) - apply (cases i) - apply auto - done +proof - + { fix pc e + have "is_relevant_entry (compP f P) i pc e = is_relevant_entry P i pc e" + by(cases i) (auto simp: is_relevant_entry_def) + } + then show ?thesis by fast +qed (*>*) lemma [simp]: "relevant_entries (compP f P) i pc xt = relevant_entries P i pc xt" (*<*) by (simp add: relevant_entries_def)(*>*) lemma [simp]: "app i (compP f P) mpc T pc mxl xt \ = app i P mpc T pc mxl xt \" (*<*) - apply (simp add: app_def xcpt_app_def eff_def xcpt_eff_def norm_eff_def) - apply (fastforce simp add: image_def) - done +by (simp add: app_def xcpt_app_def eff_def xcpt_eff_def norm_eff_def) + (fastforce simp add: image_def) (*>*) -lemma [simp]: "app i P mpc T pc mxl xt \ \ eff i (compP f P) pc xt \ = eff i P pc xt \" +lemma [simp]: +assumes "app i P mpc T pc mxl xt \" +shows "eff i (compP f P) pc xt \ = eff i P pc xt \" (*<*) - apply (clarsimp simp add: eff_def norm_eff_def xcpt_eff_def app_def) - apply (cases i) - apply auto - done +using assms +proof(clarsimp simp: eff_def norm_eff_def xcpt_eff_def app_def, cases i) +qed auto (*>*) lemma [simp]: "subtype (compP f P) = subtype P" -(*<*) - apply (rule ext)+ - apply (simp) - done -(*>*) +(*<*)by (rule ext)+ simp(*>*) lemma [simp]: "compP f P \ \ \' \' = P \ \ \' \'" (*<*) by (simp add: sup_state_opt_def sup_state_def sup_ty_opt_def)(*>*) lemma [simp]: "compP f P,T,mpc,mxl,xt \ i,pc :: \s = P,T,mpc,mxl,xt \ i,pc :: \s" (*<*)by (simp add: wt_instr_def cong: conj_cong)(*>*) declare TC1.compT_sizes[simp] TC0.ty_def2[simp] context TC2 begin lemma compT_method_NonStatic: fixes e and A and C and Ts and mxl\<^sub>0 defines [simp]: "E \ Class C # Ts" and [simp]: "A \ \{..size Ts}\" and [simp]: "A' \ A \ \ e" and [simp]: "mxl\<^sub>0 \ max_vars e" assumes mxs: "max_stack e = mxs" and mxl: "Suc (length Ts + max_vars e) = mxl" - assumes assm: "wf_prog p P" "P,E \\<^sub>1 e :: T" "\ e A" "\ e (size E)" - "set E \ types P" "P \ T \ T\<^sub>r" + assumes wf: "wf_prog p P" and wte: "P,E \\<^sub>1 e :: T" and \: "\ e A" + and \: "\ e (size E)" and E_P: "set E \ types P" and wid: "P \ T \ T\<^sub>r" shows "wt_method (compP\<^sub>2 P) C NonStatic Ts T\<^sub>r mxs mxl\<^sub>0 (compE\<^sub>2 e @ [Return]) (compxE\<^sub>2 e 0 0) (ty\<^sub>i' [] E A # compT\<^sub>a E A [] e)" -(*<*) -using assms apply (simp add: wt_method_def compT\<^sub>a_def after_def mxl) -apply (rule conjI) - apply (simp add: check_types_def OK_ty\<^sub>i'_in_statesI) - apply (rule conjI) - apply (drule (1) WT\<^sub>1_is_type) - apply simp - apply (insert max_stack1 [where e=e]) - apply (rule OK_ty\<^sub>i'_in_statesI) apply (simp_all add: mxs)[3] - apply (erule compT_states(1)) - apply assumption - apply (simp_all add: mxs mxl)[4] -apply (rule conjI) - apply (auto simp add: wt_start_def ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth - nth_Cons mxl split: nat.split dest: less_antisym)[1] -apply (frule (1) TC2.compT_wt_instrs [of P _ _ _ _ "[]" "max_stack e" "Suc (length Ts + max_vars e)" T\<^sub>r]) - apply simp_all -apply (clarsimp simp: after_def) -apply hypsubst_thin -apply (rule conjI) - apply (clarsimp simp: wt_instrs_def after_def mxl mxs) -apply clarsimp -apply (drule (1) less_antisym) -apply (clarsimp simp: wt_defs xcpt_app_pcs xcpt_eff_pcs ty\<^sub>i'_def) -done +(*<*)(is "wt_method ?P C ?b Ts T\<^sub>r mxs mxl\<^sub>0 ?is ?xt ?\s") +proof - + let ?n = "length E + mxl\<^sub>0" + have wt_compE: "P,T\<^sub>r,mxs \ compE\<^sub>2 e, compxE\<^sub>2 e 0 (length []) [::] + TC0.ty\<^sub>i' ?n [] E A # TC1.compT P ?n E A [] e @[TC0.after P ?n E A [] e]" + using mxs TC2.compT_wt_instrs [OF wte \ \, of "[]" mxs ?n T\<^sub>r] by simp + + have "OK (ty\<^sub>i' [T] E A') \ states P mxs mxl" + using mxs WT\<^sub>1_is_type[OF wf wte E_P] max_stack1[OF wte] OK_ty\<^sub>i'_in_statesI[OF E_P] + by simp + moreover have "OK ` set (compT E A [] e) \ states P mxs mxl" + using mxs mxl wid compT_states(1)[OF wf wte E_P] by simp + ultimately have "check_types ?P mxs ?n (map OK ?\s)" + using mxl wte E_P by (simp add: compT\<^sub>a_def after_def check_types_def) + moreover have "wt_start ?P C ?b Ts mxl\<^sub>0 ?\s" using mxl + by (auto simp: wt_start_def ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth nth_Cons + split: nat.split) + moreover { + fix pc assume pc: "pc < size ?is" + then have "?P,T\<^sub>r,mxs,size ?is,?xt \ ?is!pc,pc :: ?\s" + proof(cases "pc < length (compE\<^sub>2 e)") + case True + then show ?thesis using mxs wte wt_compE + by (clarsimp simp: compT\<^sub>a_def mxl after_def wt_instrs_def) + next + case False + have "length (compE\<^sub>2 e) = pc" using less_antisym[OF False] pc by simp + then show ?thesis using mxl wte E_P wid + by (clarsimp simp: compT\<^sub>a_def after_def wt_defs xcpt_app_pcs xcpt_eff_pcs ty\<^sub>i'_def) + qed + } + moreover have "0 < size ?is" and "size ?\s = size ?is" + using wte by (simp_all add: compT\<^sub>a_def) + ultimately show ?thesis by(simp add: wt_method_def) +qed (*>*) lemma compT_method_Static: fixes e and A and C and Ts and mxl\<^sub>0 defines [simp]: "E \ Ts" and [simp]: "A \ \{.." and [simp]: "A' \ A \ \ e" and [simp]: "mxl\<^sub>0 \ max_vars e" assumes mxs: "max_stack e = mxs" and mxl: "length Ts + max_vars e = mxl" - assumes assm: "wf_prog p P" "P,E \\<^sub>1 e :: T" "\ e A" "\ e (size E)" - "set E \ types P" "P \ T \ T\<^sub>r" + assumes wf: "wf_prog p P" and wte: "P,E \\<^sub>1 e :: T" and \: "\ e A" + and \: "\ e (size E)" and E_P: "set E \ types P" and wid: "P \ T \ T\<^sub>r" shows "wt_method (compP\<^sub>2 P) C Static Ts T\<^sub>r mxs mxl\<^sub>0 (compE\<^sub>2 e @ [Return]) (compxE\<^sub>2 e 0 0) (ty\<^sub>i' [] E A # compT\<^sub>a E A [] e)" -(*<*) -using assms apply (simp add: wt_method_def compT\<^sub>a_def after_def mxl) -apply (rule conjI) - apply (simp add: check_types_def OK_ty\<^sub>i'_in_statesI) - apply (rule conjI) - apply (drule (1) WT\<^sub>1_is_type) - apply simp - apply (insert max_stack1 [where e=e]) - apply (rule OK_ty\<^sub>i'_in_statesI) apply (simp_all add: mxs)[3] - apply (erule compT_states(1)) - apply assumption - apply (simp_all add: mxs mxl)[4] -apply (rule conjI) - apply (auto simp add: wt_start_def ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth - nth_Cons mxl split: nat.split dest: less_antisym)[1] -apply (frule (1) TC2.compT_wt_instrs [of P _ _ _ _ "[]" "max_stack e" "length Ts + max_vars e" T\<^sub>r]) - apply simp_all -apply (clarsimp simp: after_def) -apply hypsubst_thin -apply (rule conjI) - apply (clarsimp simp: wt_instrs_def after_def mxl mxs) -apply clarsimp -apply (drule (1) less_antisym) -apply (clarsimp simp: wt_defs xcpt_app_pcs xcpt_eff_pcs ty\<^sub>i'_def) -done +(*<*)(is "wt_method ?P C ?b Ts T\<^sub>r mxs mxl\<^sub>0 ?is ?xt ?\s") +proof - + let ?n = "length E + mxl\<^sub>0" + have wt_compE: "P,T\<^sub>r,mxs \ compE\<^sub>2 e, compxE\<^sub>2 e 0 (length []) [::] + TC0.ty\<^sub>i' ?n [] E A # TC1.compT P ?n E A [] e @[TC0.after P ?n E A [] e]" + using mxs TC2.compT_wt_instrs [OF wte \ \, of "[]" mxs ?n T\<^sub>r] by simp + + have "OK (ty\<^sub>i' [T] E A') \ states P mxs mxl" + using mxs WT\<^sub>1_is_type[OF wf wte E_P] max_stack1[OF wte] OK_ty\<^sub>i'_in_statesI[OF E_P] + by simp + moreover have "OK ` set (compT E A [] e) \ states P mxs mxl" + using mxs mxl wid compT_states(1)[OF wf wte E_P] by simp + ultimately have "check_types ?P mxs ?n (map OK ?\s)" + using mxl wte E_P by (simp add: compT\<^sub>a_def after_def check_types_def) + moreover have "wt_start ?P C ?b Ts mxl\<^sub>0 ?\s" using mxl + by (auto simp: wt_start_def ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth nth_Cons + split: nat.split) + moreover { + fix pc assume pc: "pc < size ?is" + then have "?P,T\<^sub>r,mxs,size ?is,?xt \ ?is!pc,pc :: ?\s" + proof(cases "pc < length (compE\<^sub>2 e)") + case True + then show ?thesis using mxs wte wt_compE + by (clarsimp simp: compT\<^sub>a_def mxl after_def wt_instrs_def) + next + case False + have "length (compE\<^sub>2 e) = pc" using less_antisym[OF False] pc by simp + then show ?thesis using mxl wte E_P wid + by (clarsimp simp: compT\<^sub>a_def after_def wt_defs xcpt_app_pcs xcpt_eff_pcs ty\<^sub>i'_def) + qed + } + moreover have "0 < size ?is" and "size ?\s = size ?is" + using wte by (simp_all add: compT\<^sub>a_def) + ultimately show ?thesis by(simp add: wt_method_def) +qed (*>*) end definition compTP :: "J\<^sub>1_prog \ ty\<^sub>P" where "compTP P C M = ( let (D,b,Ts,T,e) = method P C M; E = case b of Static \ Ts | NonStatic \ Class C # Ts; A = case b of Static \ \{.. | NonStatic \ \{..size Ts}\; mxl = (case b of Static \ 0 | NonStatic \ 1) + size Ts + max_vars e in (TC0.ty\<^sub>i' mxl [] E A # TC1.compT\<^sub>a P mxl E A [] e))" theorem wt_compP\<^sub>2: - "wf_J\<^sub>1_prog P \ wf_jvm_prog (compP\<^sub>2 P)" +assumes wf: "wf_J\<^sub>1_prog P" shows "wf_jvm_prog (compP\<^sub>2 P)" (*<*) - apply (simp add: wf_jvm_prog_def wf_jvm_prog_phi_def) - apply(rule_tac x = "compTP P" in exI) - apply (rule wf_prog_compPI) - prefer 2 apply assumption - apply (simp add: compTP_def) apply(rename_tac C M b Ts T m) - apply(case_tac b) -\ \ Static \ - apply (clarsimp simp add: wf_mdecl_def) - apply (rule TC2.compT_method_Static [simplified]) - apply (rule refl) - apply (rule refl) - apply assumption - apply assumption - apply assumption - apply assumption - apply (drule (1) sees_wf_mdecl) - apply (simp add: wf_mdecl_def) - apply (blast intro: sees_method_is_class) - apply assumption -\ \ NonStatic \ - apply (clarsimp simp add: wf_mdecl_def) - apply (rule TC2.compT_method_NonStatic [simplified]) - apply (rule refl) - apply (rule refl) - apply assumption - apply assumption - apply assumption - apply assumption - apply (drule (1) sees_wf_mdecl) - apply (simp add: wf_mdecl_def) - apply (blast intro: sees_method_is_class) - apply assumption - done +proof - + let ?\ = "compTP P" and ?f = compMb\<^sub>2 + let ?wf\<^sub>2 = "\P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). + wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (?\ C M)" + and ?P = "compP ?f P" + { fix C M b Ts T m let ?md = "(M,b,Ts,T,m)::expr\<^sub>1 mdecl" + assume cM: "P \ C sees M, b : Ts\T = m in C" + and wfm: "wf_mdecl wf_J\<^sub>1_mdecl P C ?md" + then have Ts_types: "\T'\set Ts. is_type P T'" + and T_type: "is_type P T" and wfm\<^sub>1: "wf_J\<^sub>1_mdecl P C ?md" + by(simp_all add: wf_mdecl_def) + have Ts_P: "set Ts \ types P" using sees_wf_mdecl[OF wf cM] + by (clarsimp simp: wf_mdecl_def) + then have CTs_P: "is_class P C \ set Ts \ types P" + using sees_method_is_class[OF cM] by simp + + have "?wf\<^sub>2 ?P C (M,b,Ts,T,?f b m)" + proof(cases b) + case Static + then obtain T' where wte: "P,Ts \\<^sub>1 m :: T'" and wid: "P \ T' \ T" + and \: "\ m \{.." and \: "\ m (size Ts)" + using wfm\<^sub>1 by(auto simp: wf_mdecl_def) + show ?thesis using Static cM TC2.compT_method_Static [OF _ _ wf wte \ \ Ts_P wid] + by(simp add: compTP_def) + next + case NonStatic + then obtain T' where wte: "P,Class C#Ts \\<^sub>1 m :: T'" and wid: "P \ T' \ T" + and \: "\ m \{..size Ts}\" and \: "\ m (Suc (size Ts))" + using wfm\<^sub>1 by(auto simp: wf_mdecl_def) + have Ts_P: "set Ts \ types P" using sees_wf_mdecl[OF wf cM] + by (fastforce simp: wf_mdecl_def intro: sees_method_is_class) + show ?thesis using NonStatic cM + TC2.compT_method_NonStatic [simplified, OF _ _ wf wte \ \ CTs_P wid] + by(simp add: compTP_def) + qed + then have "wf_mdecl ?wf\<^sub>2 ?P C (M, b, Ts, T, ?f b m)" + using Ts_types T_type by(simp add: wf_mdecl_def) + } + then have "wf_prog ?wf\<^sub>2 (compP ?f P)" by(rule wf_prog_compPI[OF _ wf]) + then show ?thesis by (simp add: wf_jvm_prog_def wf_jvm_prog_phi_def) fast +qed (*>*) theorem wt_J2JVM: "wf_J_prog P \ wf_jvm_prog (J2JVM P)" (*<*) -apply(simp only:o_def J2JVM_def) -apply(blast intro:wt_compP\<^sub>2 compP\<^sub>1_pres_wf) -done +by(simp only:o_def J2JVM_def) + (blast intro:wt_compP\<^sub>2 compP\<^sub>1_pres_wf) end diff --git a/thys/JinjaDCI/J/DefAss.thy b/thys/JinjaDCI/J/DefAss.thy --- a/thys/JinjaDCI/J/DefAss.thy +++ b/thys/JinjaDCI/J/DefAss.thy @@ -1,189 +1,192 @@ (* Title: JinjaDCI/J/DefAss.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory J/DefAss.thy by Tobias Nipkow *) section \ Definite assignment \ theory DefAss imports BigStep begin subsection "Hypersets" type_synonym 'a hyperset = "'a set option" definition hyperUn :: "'a hyperset \ 'a hyperset \ 'a hyperset" (infixl "\" 65) where "A \ B \ case A of None \ None | \A\ \ (case B of None \ None | \B\ \ \A \ B\)" definition hyperInt :: "'a hyperset \ 'a hyperset \ 'a hyperset" (infixl "\" 70) where "A \ B \ case A of None \ B | \A\ \ (case B of None \ \A\ | \B\ \ \A \ B\)" definition hyperDiff1 :: "'a hyperset \ 'a \ 'a hyperset" (infixl "\" 65) where "A \ a \ case A of None \ None | \A\ \ \A - {a}\" definition hyper_isin :: "'a \ 'a hyperset \ bool" (infix "\\" 50) where "a \\ A \ case A of None \ True | \A\ \ a \ A" definition hyper_subset :: "'a hyperset \ 'a hyperset \ bool" (infix "\" 50) where "A \ B \ case B of None \ True | \B\ \ (case A of None \ False | \A\ \ A \ B)" lemmas hyperset_defs = hyperUn_def hyperInt_def hyperDiff1_def hyper_isin_def hyper_subset_def lemma [simp]: "\{}\ \ A = A \ A \ \{}\ = A" (*<*)by(simp add:hyperset_defs)(*>*) lemma [simp]: "\A\ \ \B\ = \A \ B\ \ \A\ \ a = \A - {a}\" (*<*)by(simp add:hyperset_defs)(*>*) lemma [simp]: "None \ A = None \ A \ None = None" (*<*)by(simp add:hyperset_defs)(*>*) lemma [simp]: "a \\ None \ None \ a = None" (*<*)by(simp add:hyperset_defs)(*>*) lemma hyper_isin_union: "x \\ \A\ \ x \\ \A\ \ B" by(case_tac B, auto simp: hyper_isin_def) lemma hyperUn_assoc: "(A \ B) \ C = A \ (B \ C)" (*<*)by(simp add:hyperset_defs Un_assoc)(*>*) lemma hyper_insert_comm: "A \ \{a}\ = \{a}\ \ A \ A \ (\{a}\ \ B) = \{a}\ \ (A \ B)" (*<*)by(simp add:hyperset_defs)(*>*) lemma hyper_comm: "A \ B = B \ A \ A \ B \ C = B \ A \ C" -apply(case_tac A, simp, case_tac B, simp) -apply(case_tac C, simp add: Un_commute) -apply(simp add: Un_left_commute Un_commute) -done +(*<*) +proof(cases A) + case SomeA: Some then show ?thesis + proof(cases B) + case SomeB: Some with SomeA show ?thesis + proof(cases C) + case SomeC: Some with SomeA SomeB show ?thesis + by(simp add: Un_left_commute Un_commute) + qed (simp add: Un_commute) + qed simp +qed simp +(*>*) subsection "Definite assignment" primrec \ :: "'a exp \ 'a hyperset" and \s :: "'a exp list \ 'a hyperset" where "\ (new C) = \{}\" | "\ (Cast C e) = \ e" | "\ (Val v) = \{}\" | "\ (e\<^sub>1 \bop\ e\<^sub>2) = \ e\<^sub>1 \ \ e\<^sub>2" | "\ (Var V) = \{}\" | "\ (LAss V e) = \{V}\ \ \ e" | "\ (e\F{D}) = \ e" | "\ (C\\<^sub>sF{D}) = \{}\" | "\ (e\<^sub>1\F{D}:=e\<^sub>2) = \ e\<^sub>1 \ \ e\<^sub>2" | "\ (C\\<^sub>sF{D}:=e\<^sub>2) = \ e\<^sub>2" | "\ (e\M(es)) = \ e \ \s es" | "\ (C\\<^sub>sM(es)) = \s es" | "\ ({V:T; e}) = \ e \ V" | "\ (e\<^sub>1;;e\<^sub>2) = \ e\<^sub>1 \ \ e\<^sub>2" | "\ (if (e) e\<^sub>1 else e\<^sub>2) = \ e \ (\ e\<^sub>1 \ \ e\<^sub>2)" | "\ (while (b) e) = \ b" | "\ (throw e) = None" | "\ (try e\<^sub>1 catch(C V) e\<^sub>2) = \ e\<^sub>1 \ (\ e\<^sub>2 \ V)" | "\ (INIT C (Cs,b) \ e) = \{}\" | "\ (RI (C,e);Cs \ e') = \ e" | "\s ([]) = \{}\" | "\s (e#es) = \ e \ \s es" primrec \ :: "'a exp \ 'a hyperset \ bool" and \s :: "'a exp list \ 'a hyperset \ bool" where "\ (new C) A = True" | "\ (Cast C e) A = \ e A" | "\ (Val v) A = True" | "\ (e\<^sub>1 \bop\ e\<^sub>2) A = (\ e\<^sub>1 A \ \ e\<^sub>2 (A \ \ e\<^sub>1))" | "\ (Var V) A = (V \\ A)" | "\ (LAss V e) A = \ e A" | "\ (e\F{D}) A = \ e A" | "\ (C\\<^sub>sF{D}) A = True" | "\ (e\<^sub>1\F{D}:=e\<^sub>2) A = (\ e\<^sub>1 A \ \ e\<^sub>2 (A \ \ e\<^sub>1))" | "\ (C\\<^sub>sF{D}:=e\<^sub>2) A = \ e\<^sub>2 A" | "\ (e\M(es)) A = (\ e A \ \s es (A \ \ e))" | "\ (C\\<^sub>sM(es)) A = \s es A" | "\ ({V:T; e}) A = \ e (A \ V)" | "\ (e\<^sub>1;;e\<^sub>2) A = (\ e\<^sub>1 A \ \ e\<^sub>2 (A \ \ e\<^sub>1))" | "\ (if (e) e\<^sub>1 else e\<^sub>2) A = (\ e A \ \ e\<^sub>1 (A \ \ e) \ \ e\<^sub>2 (A \ \ e))" | "\ (while (e) c) A = (\ e A \ \ c (A \ \ e))" | "\ (throw e) A = \ e A" | "\ (try e\<^sub>1 catch(C V) e\<^sub>2) A = (\ e\<^sub>1 A \ \ e\<^sub>2 (A \ \{V}\))" | "\ (INIT C (Cs,b) \ e) A = \ e A" | "\ (RI (C,e);Cs \ e') A = (\ e A \ \ e' A)" | "\s ([]) A = True" | "\s (e#es) A = (\ e A \ \s es (A \ \ e))" lemma As_map_Val[simp]: "\s (map Val vs) = \{}\" (*<*)by (induct vs) simp_all(*>*) lemma D_append[iff]: "\A. \s (es @ es') A = (\s es A \ \s es' (A \ \s es))" (*<*)by (induct es type:list) (auto simp:hyperUn_assoc)(*>*) lemma A_fv: "\A. \ e = \A\ \ A \ fv e" and "\A. \s es = \A\ \ A \ fvs es" (*<*) -apply(induct e and es rule: \.induct \s.induct) -apply (simp_all add:hyperset_defs) -apply blast+ -done +by (induct e and es rule: \.induct \s.induct) + (fastforce simp add:hyperset_defs)+ (*>*) lemma sqUn_lem: "A \ A' \ A \ B \ A' \ B" (*<*)by(simp add:hyperset_defs) blast(*>*) lemma diff_lem: "A \ A' \ A \ b \ A' \ b" (*<*)by(simp add:hyperset_defs) blast(*>*) (* This order of the premises avoids looping of the simplifier *) lemma D_mono: "\A A'. A \ A' \ \ e A \ \ (e::'a exp) A'" and Ds_mono: "\A A'. A \ A' \ \s es A \ \s (es::'a exp list) A'" (*<*) -apply(induct e and es rule: \.induct \s.induct) -apply simp -apply simp -apply simp -apply simp apply (iprover dest:sqUn_lem) -apply (fastforce simp add:hyperset_defs) -apply simp -apply simp -apply simp -apply simp apply (iprover dest:sqUn_lem) -apply simp -apply simp apply (iprover dest:sqUn_lem) -apply simp -apply simp apply (iprover dest:diff_lem) -apply simp apply (iprover dest:sqUn_lem) -apply simp apply (iprover dest:sqUn_lem) -apply simp apply (iprover dest:sqUn_lem) -apply simp -apply simp apply (iprover dest:sqUn_lem) -apply simp -apply simp -apply simp -apply simp apply (iprover dest:sqUn_lem) -done +proof(induct e and es rule: \.induct \s.induct) + case BinOp then show ?case by simp (iprover dest:sqUn_lem) +next + case Var then show ?case by (fastforce simp add:hyperset_defs) +next + case FAss then show ?case by simp (iprover dest:sqUn_lem) +next + case Call then show ?case by simp (iprover dest:sqUn_lem) +next + case Block then show ?case by simp (iprover dest:diff_lem) +next + case Seq then show ?case by simp (iprover dest:sqUn_lem) +next + case Cond then show ?case by simp (iprover dest:sqUn_lem) +next + case While then show ?case by simp (iprover dest:sqUn_lem) +next + case TryCatch then show ?case by simp (iprover dest:sqUn_lem) +next + case Cons_exp then show ?case by simp (iprover dest:sqUn_lem) +qed simp_all (*>*) (* And this is the order of premises preferred during application: *) lemma D_mono': "\ e A \ A \ A' \ \ e A'" and Ds_mono': "\s es A \ A \ A' \ \s es A'" (*<*)by(blast intro:D_mono, blast intro:Ds_mono)(*>*) lemma Ds_Vals: "\s (map Val vs) A" by(induct vs, auto) 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,350 +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\" -apply(induct e, auto) apply(rename_tac x1 x2 x3 x4) -apply(case_tac x4, auto) -done +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) + 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) +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) + 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) + 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" -apply(cut_tac e = e and R = "\e. \sub_RI e \ iconf sh e" in subexp_induct) - apply(rename_tac ea) apply(case_tac ea, simp_all) -apply(clarsimp simp: nsub_RI_iconf_aux) -done - -lemma nsub_RIs_iconfs: "\sub_RIs es \ iconfs sh es" -apply(cut_tac es = es and R = "\e. \sub_RI e \ iconf sh e" - and Rs = "\es. \sub_RIs es \ iconfs sh es" in subexps_induct) - apply(rename_tac esa) apply(case_tac esa, simp_all) - apply(clarsimp simp: nsub_RI_iconf_aux)+ -done + 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" + by(rule subexp_induct[where ?Rs = ?Rs]; clarsimp simp: nsub_RI_iconf_aux) + 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) + 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) + 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) \" -apply(unfold bconf_def, cases b, auto) -apply(drule val_of_spec, simp) -done + 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) \)" -apply(unfold bconf_def, cases b, auto) -apply(drule val_of_spec, simp) -done + 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) \" -apply(unfold bconf_def, cases b, auto) -apply(drule val_of_spec, simp) -done + 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) \" -apply(unfold bconf_def, cases b, auto) -apply(drule val_of_spec, simp) -done + 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) \)" -apply(unfold bconf_def, cases b, auto) -apply(drule val_of_spec, simp) -done + 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(unfold bconf_def, cases b, auto) + 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) + 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) + 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) + 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(unfold bconf_def, cases b, auto simp: assigned_def) + 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(unfold bconf_def, cases b, auto) + 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) \))" (* \ P,sh \\<^sub>b (e1,b) \"*) -by(unfold bconf_def, cases b, auto dest: val_of_spec lass_val_of_spec) + | 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') \" -apply(unfold bconf_def, cases "bool_of b") apply auto[1] -apply(rename_tac a) apply(case_tac a) - apply(simp, drule bool_of_specT) apply auto[1] -apply(simp, drule bool_of_specF) apply auto[1] -done +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(unfold bconf_def, cases b, auto) + by(cases b) (auto simp: bconf_def) lemma bconf_Throw[iff]: "P,sh \\<^sub>b (throw e,b) \ \ P,sh \\<^sub>b (e,b) \" -apply(unfold bconf_def, cases b, auto) -apply(drule val_of_spec, simp) -done + 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) \" -apply(unfold bconf_def, cases b, auto) -apply(drule val_of_spec, simp) -done + 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(unfold bconf_def, cases b, auto) + 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) \" -apply(unfold bconf_def, cases b, auto) -apply(drule val_of_spec, simp) -done + 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) + 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,4191 @@ (* 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(rule InitNoneRed[THEN converse_rtrancl_into_rtrancl])(*>*) 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(rule RedInitDone[THEN converse_rtrancl_into_rtrancl])(*>*) 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(rule RedInitProcessing[THEN converse_rtrancl_into_rtrancl])(*>*) 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(rule RedInitError[THEN converse_rtrancl_into_rtrancl])(*>*) 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(rule InitObjectRed[THEN converse_rtrancl_into_rtrancl])(*>*) 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(rule InitNonObjectSuperRed[THEN converse_rtrancl_into_rtrancl])(*>*) 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(rule RedInitRInit[THEN converse_rtrancl_into_rtrancl])(*>*) 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 + by(rule rtrancl_into_rtrancl[OF step(3) RInitRed[OF step(2)]]) +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'\" + \ 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 +by(rule rtrancl_trans[OF RInitReds + RedRInit[THEN converse_rtrancl_into_rtrancl]]) (*>*) + 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\" + "\ 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\" (*<*) -apply(rule rtrancl_trans) - apply(erule RInitReds) -apply(rule converse_rtrancl_into_rtrancl) - apply(erule (1) RInitInitThrow) -apply assumption -done +by(rule rtrancl_trans[OF RInitReds + RInitInitThrow[THEN converse_rtrancl_into_rtrancl]]) (*>*) 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(rule RInitReds[THEN rtrancl_into_rtrancl, OF _ 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(rule NewInitDoneRed[THEN converse_rtrancl_into_rtrancl, + OF _ 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(rule NewInitDoneRed[THEN converse_rtrancl_into_rtrancl, + OF _ 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 +assumes nDone: "\sfs. shp s C = Some (sfs, Done)" + and INIT_steps: "P \ \INIT C ([C],False) \ unit,s,False\ \* \Val v',(h',l',sh'),b'\" + and Addr: "new_Addr h' = Some a" and has: "P \ C has_fields FDTs" + and h'': "h'' = h'(a\blank P C)" and cls_C: "is_class P C" +shows "P \ \new C,s,False\ \* \addr a,(h'',l',sh'),False\" +(*<*)(is "(?a, ?c) \ (red P)\<^sup>*") +proof - + let ?b = "(INIT C ([C],False) \ new C,s,False)" + let ?b' = "(new C,(h', l', sh'),icheck P C (new C::expr))" + have b'c: "(?b', ?c) \ (red P)\<^sup>*" + using RedNew[OF Addr has h'', THEN r_into_rtrancl] by simp + obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp + have "(?a, ?b) \ (red P)\<^sup>*" + using NewInitRed[OF _ cls_C] nDone by fastforce + also have "(?b, ?c) \ (red P)\<^sup>*" + by(rule InitSeqReds[OF INIT_steps b'c]) simp + ultimately show ?thesis by simp +qed (*>*) 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 +assumes nDone: "\sfs. shp s C = Some (sfs, Done)" + and INIT_steps: "P \ \INIT C ([C],False) \ unit,s,False\ \* \Val v',(h',l',sh'),b'\" + and Addr: "new_Addr h' = None" and cls_C: "is_class P C" +shows "P \ \new C,s,False\ \* \THROW OutOfMemory,(h',l',sh'),False\" +(*<*)(is "(?a, ?c) \ (red P)\<^sup>*") +proof - + let ?b = "(INIT C ([C],False) \ new C,s,False)" + let ?b' = "(new C,(h', l', sh'),icheck P C (new C::expr))" + have b'c: "(?b', ?c) \ (red P)\<^sup>*" + using RedNewFail[OF Addr cls_C, THEN r_into_rtrancl] by simp + obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp + have "(?a, ?b) \ (red P)\<^sup>*" + using NewInitRed[OF _ cls_C] nDone by fastforce + also have "(?b, ?c) \ (red P)\<^sup>*" + by(rule InitSeqReds[OF INIT_steps b'c]) simp + ultimately show ?thesis by simp +qed (*>*) 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 +assumes nDone: "\sfs. shp s C = Some (sfs, Done)" + and cls_C: "is_class P C" + and INIT_steps: "P \ \INIT C ([C],False) \ unit,s,False\ \* \throw a,s',b'\" +shows "P \ \new C,s,False\ \* \throw a,s',b'\" +(*<*)(is "(?a, ?c) \ (red P)\<^sup>*") +proof - + let ?b = "(INIT C ([C],False) \ new C,s,False)" + obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp + have "(?a, ?b) \ (red P)\<^sup>*" + using NewInitRed[OF _ cls_C] nDone by fastforce + also have "(?b, ?c) \ (red P)\<^sup>*" + using InitSeqThrowReds[OF INIT_steps] by simp + ultimately show ?thesis by simp +qed (*>*) 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 + by(rule rtrancl_into_rtrancl[OF step(3) CastRed[OF step(2)]]) +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 -(*>*) +(*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCastNull])(*>*) 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 -(*>*) +(*<*)by(cases s', simp) (rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCast])(*>*) 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 -(*>*) +(*<*)by(cases s', simp) (rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCastFail])(*>*) 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 -(*>*) +(*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ red_reds.CastThrow])(*>*) 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 + by(rule rtrancl_into_rtrancl[OF step(3) LAssRed[OF step(2)]]) +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 -(*>*) +(*<*)by(rule LAssReds[THEN rtrancl_into_rtrancl, OF _ RedLAss])(*>*) 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 -(*>*) +(*<*)by(rule LAssReds[THEN rtrancl_into_rtrancl, OF _ red_reds.LAssThrow])(*>*) 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 + by(rule rtrancl_into_rtrancl[OF step(3) BinOpRed1[OF step(2)]]) +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 + by(rule rtrancl_into_rtrancl[OF step(3) BinOpRed2[OF step(2)]]) +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 +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1,b\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v\<^sub>2,s\<^sub>2,b\<^sub>2\" + and op: "binop(bop,v\<^sub>1,v\<^sub>2) = Some v" +shows "P \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \Val v,s\<^sub>2,b\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(Val v\<^sub>1 \bop\ e\<^sub>2, s\<^sub>1,b\<^sub>1)" + let ?y' = "(Val v\<^sub>1 \bop\ Val v\<^sub>2, s\<^sub>2,b\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule BinOp1Reds[OF e\<^sub>1_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule BinOp2Reds[OF e\<^sub>2_steps]) + also have "(?y', ?z) \ (red P)" by(rule RedBinOp[OF op]) + ultimately show ?thesis by simp +qed (*>*) 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 -(*>*) +(*<*)by(rule BinOp1Reds[THEN rtrancl_into_rtrancl, OF _ red_reds.BinOpThrow1])(*>*) 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 +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1,b\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \throw e,s\<^sub>2,b\<^sub>2\" +shows "P \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \throw e,s\<^sub>2,b\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(Val v\<^sub>1 \bop\ e\<^sub>2, s\<^sub>1,b\<^sub>1)" + let ?y' = "(Val v\<^sub>1 \bop\ throw e, s\<^sub>2,b\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule BinOp1Reds[OF e\<^sub>1_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule BinOp2Reds[OF e\<^sub>2_steps]) + also have "(?y', ?z) \ (red P)" by(rule red_reds.BinOpThrow2) + ultimately show ?thesis by simp +qed (*>*) 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 + by(rule rtrancl_into_rtrancl[OF step(3) FAccRed[OF step(2)]]) +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 -(*>*) +(*<*)by(cases s',simp) (rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAcc])(*>*) 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 -(*>*) +(*<*)by(rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAccNull])(*>*) 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 -(*>*) +(*<*)by(cases s',simp) (auto intro: FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAccNone])(*>*) 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 -(*>*) +(*<*)by(cases s',simp) (rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAccStatic])(*>*) 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 -(*>*) +(*<*)by(rule FAccReds[THEN rtrancl_into_rtrancl, OF _ red_reds.FAccThrow])(*>*) 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 -(*>*) +(*<*)by(cases s,simp) (rule RedSFAcc[THEN r_into_rtrancl])(*>*) 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 -(*>*) +(*<*)by(cases s,simp) (auto intro: RedSFAccNone[THEN r_into_rtrancl])(*>*) 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 -(*>*) +(*<*)by(cases s,simp) (rule RedSFAccNonStatic[THEN r_into_rtrancl])(*>*) 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 +assumes cF: "P \ C has F,Static:t in D" + and shp: "shp s D = Some (sfs,Done)" and sfs: "sfs F = Some v" +shows "P \ \C\\<^sub>sF{D}, s, b\ \* \Val v, s, False\" +(*<*)(is "(?a, ?c) \ (red P)\<^sup>*") +proof - + obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp + show ?thesis proof(cases b) + case True + then show ?thesis using assms + by simp (rule RedSFAcc[THEN r_into_rtrancl]) + next + case False + let ?b = "(C\\<^sub>sF{D},s,True)" + have "(?a, ?b) \ (red P)\<^sup>*" + using False SFAccInitDoneRed[where sh="shp s", OF cF shp] by fastforce + also have "(?b, ?c) \ (red P)\<^sup>*" by(rule SFAccReds[OF assms]) + ultimately show ?thesis by simp + qed +qed (*>*) 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 +assumes cF: "P \ C has F,Static:t in D" + and nDone: "\sfs. shp s D = Some (sfs,Done)" + and INIT_steps: "P \ \INIT D ([D],False) \ unit,s,False\ \* \Val v',s',b'\" + and shp': "shp s' D = Some (sfs,i)" and sfs: "sfs F = Some v" +shows "P \ \C\\<^sub>sF{D},s,False\ \* \Val v,s',False\" +(*<*)(is "(?a, ?c) \ (red P)\<^sup>*") +proof - + let ?b = "(INIT D ([D],False) \ C\\<^sub>sF{D},s,False)" + let ?b' = "(C\\<^sub>sF{D},s',icheck P D (C\\<^sub>sF{D}::expr))" + obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp + obtain h' l' sh' where [simp]: "s' = (h', l', sh')" by(cases s') simp + have icheck: "icheck P D (C\\<^sub>sF{D}) = True" using cF by fastforce + then have b'c: "(?b', ?c) \ (red P)\<^sup>*" + using RedSFAcc[THEN r_into_rtrancl, OF cF] shp' sfs by simp + have "(?a, ?b) \ (red P)" using SFAccInitRed[OF cF] nDone by simp + also have "(?b, ?c) \ (red P)\<^sup>*" + by(rule InitSeqReds[OF INIT_steps b'c]) simp + ultimately show ?thesis by simp +qed (*>*) 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 +by(cases s, simp) + (auto elim: converse_rtrancl_into_rtrancl[OF SFAccInitRed InitSeqThrowReds]) (*>*) 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 + by(rule rtrancl_into_rtrancl[OF step(3) FAssRed1[OF step(2)]]) +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 + by(rule rtrancl_into_rtrancl[OF step(3) FAssRed2[OF step(2)]]) +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 +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \addr a,s\<^sub>1,b\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" + and cF: "P \ C has F,NonStatic:t in D" and hC: "Some(C,fs) = h\<^sub>2 a" +shows "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\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(addr a\F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)" + let ?y' = "(addr a\F{D}:=Val v::expr,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) + also have "(?y', ?z) \ (red P)" using RedFAss[OF cF] hC by simp + ultimately show ?thesis by simp +qed (*>*) 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 +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \null,s\<^sub>1,b\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v,s\<^sub>2,b\<^sub>2\" +shows "P \ \e\<^sub>1\F{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\ \* \THROW NullPointer, s\<^sub>2, b\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(null\F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)" + let ?y' = "(null\F{D}:=Val v::expr,s\<^sub>2,b\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) + also have "(?y', ?z) \ (red P)" by(rule RedFAssNull) + ultimately show ?thesis by simp +qed (*>*) 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 -(*>*) +(*<*)by(rule FAssReds1[THEN rtrancl_into_rtrancl, OF _ red_reds.FAssThrow1])(*>*) 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 +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Val v,s\<^sub>1,b\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \throw e,s\<^sub>2,b\<^sub>2\" +shows "P \ \e\<^sub>1\F{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \throw e,s\<^sub>2,b\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(Val v\F{D}:=e\<^sub>2,s\<^sub>1,b\<^sub>1)" + let ?y' = "(Val v\F{D}:=throw e,s\<^sub>2,b\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) + also have "(?y', ?z) \ (red P)" by(rule red_reds.FAssThrow2) + ultimately show ?thesis by simp +qed (*>*) 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 +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \addr a,s\<^sub>1,b\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" + and hC: "h\<^sub>2 a = Some(C,fs)" and ncF: "\(\b t. P \ C has F,b:t in D)" +shows "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\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(addr a\F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)" + let ?y' = "(addr a\F{D}:=Val v::expr,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) + also have "(?y', ?z) \ (red P)" + using RedFAssNone[OF _ ncF] hC by simp + ultimately show ?thesis by simp +qed (*>*) 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 +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \addr a,s\<^sub>1,b\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" + and hC: "h\<^sub>2 a = Some(C,fs)" and cF_Static: "P \ C has F,Static:t in D" +shows "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\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(addr a\F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)" + let ?y' = "(addr a\F{D}:=Val v::expr,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps]) + also have "(?y', ?z) \ (red P)" + using RedFAssStatic[OF _ cF_Static] hC by simp + ultimately show ?thesis by simp +qed (*>*) 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 + by(rule rtrancl_into_rtrancl[OF step(3) SFAssRed[OF step(2)]]) +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 +assumes e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\" + and cF: "P \ C has F,Static:t in D" + and shD: "sh\<^sub>2 D = \(sfs,Done)\" +shows "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\" +(*<*)(is "(?a, ?c) \ (red P)\<^sup>*") +proof - + let ?b = "(C\\<^sub>sF{D}:=Val v::expr, (h\<^sub>2,l\<^sub>2,sh\<^sub>2), b\<^sub>2)" + have "(?a, ?b) \ (red P)\<^sup>*" by(rule SFAssReds[OF e\<^sub>2_steps]) + also have "(?b, ?c) \ (red P)\<^sup>*" proof(cases b\<^sub>2) + case True + then show ?thesis + using RedSFAss[THEN r_into_rtrancl, OF cF] shD by simp + next + case False + let ?b' = "(C\\<^sub>sF{D}:=Val v::expr, (h\<^sub>2,l\<^sub>2,sh\<^sub>2), True)" + have "(?b, ?b') \ (red P)\<^sup>*" + using False SFAssInitDoneRed[THEN converse_rtrancl_into_rtrancl, OF cF] shD + by simp + also have "(?b', ?c) \ (red P)\<^sup>*" + using RedSFAss[THEN r_into_rtrancl, OF cF] shD by simp + ultimately show ?thesis by simp + qed + ultimately show ?thesis by simp +qed (*>*) 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 -(*>*) +(*<*)by(rule SFAssReds[THEN rtrancl_into_rtrancl, OF _ red_reds.SFAssThrow])(*>*) 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 -(*>*) +(*<*)by(rule SFAssReds[THEN rtrancl_into_rtrancl, OF _ RedSFAssNone])(*>*) 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 -(*>*) +(*<*)by(rule SFAssReds[THEN rtrancl_into_rtrancl, OF _ RedSFAssNonStatic])(*>*) 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 +assumes e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\" + and cF: "P \ C has F,Static:t in D" + and nDone: "\sfs. sh\<^sub>2 D = Some (sfs, Done)" + and INIT_steps: "P \ \INIT D ([D],False) \ unit,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\ \* \Val v',(h',l',sh'),b'\" + and sh'D: "sh' D = Some(sfs,i)" + and sfs': "sfs' = sfs(F\v)" and sh'': "sh'' = sh'(D\(sfs',i))" +shows "P \ \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \unit,(h',l',sh''),False\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(C\\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)" + let ?y' = "(INIT D ([D],False) \ C\\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)" + let ?y'' = "(C\\<^sub>sF{D} := Val v::expr,(h', l', sh'),icheck P D (C\\<^sub>sF{D} := Val v::expr))" + have icheck: "icheck P D (C\\<^sub>sF{D} := Val v::expr)" using cF by fastforce + then have y''z: "(?y'',?z) \ (red P)" + using RedSFAss[OF cF _ sfs' sh''] sh'D by simp + have "(?x, ?y) \ (red P)\<^sup>*" by(rule SFAssReds[OF e\<^sub>2_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" + using SFAssInitRed[THEN converse_rtrancl_into_rtrancl, OF cF] nDone + by simp + also have "(?y', ?z) \ (red P)\<^sup>*" + using InitSeqReds[OF INIT_steps y''z[THEN r_into_rtrancl]] by simp + ultimately show ?thesis by simp +qed (*>*) 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 +assumes e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\" + and cF: "P \ C has F,Static:t in D" + and nDone: "\sfs. sh\<^sub>2 D = Some (sfs, Done)" + and INIT_steps: "P \ \INIT D ([D],False) \ unit,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\ \* \throw a,s',b'\" +shows "P \ \C\\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\ \* \throw a,s',b'\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(C\\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)" + let ?y' = "(INIT D ([D],False) \ C\\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule SFAssReds[OF e\<^sub>2_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" + using SFAssInitRed[THEN converse_rtrancl_into_rtrancl, OF cF] nDone + by simp + also have "(?y', ?z) \ (red P)\<^sup>*" + using InitSeqThrowReds[OF INIT_steps] by simp + ultimately show ?thesis by simp +qed (*>*) 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 + by(rule rtrancl_into_rtrancl[OF step(3) SeqRed[OF step(2)]]) +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 -(*>*) +(*<*)by(rule SeqReds[THEN rtrancl_into_rtrancl, OF _ red_reds.SeqThrow])(*>*) 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 +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1,b\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2,s\<^sub>1,b\<^sub>1\ \* \e\<^sub>2',s\<^sub>2,b\<^sub>2\" +shows "P \ \e\<^sub>1;;e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \e\<^sub>2',s\<^sub>2,b\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(Val v\<^sub>1;; e\<^sub>2,s\<^sub>1,b\<^sub>1)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule SeqReds[OF e\<^sub>1_steps]) + also have "(?y, ?z) \ (red P)\<^sup>*" + by(rule RedSeq[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>2_steps]) + ultimately show ?thesis by simp +qed (*>*) 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 + by(rule rtrancl_into_rtrancl[OF step(3) CondRed[OF step(2)]]) +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 -(*>*) +(*<*)by(rule CondReds[THEN rtrancl_into_rtrancl, OF _ red_reds.CondThrow])(*>*) 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 +assumes e_steps: "P \ \e,s\<^sub>0,b\<^sub>0\ \* \true,s\<^sub>1,b\<^sub>1\" + and e\<^sub>1_steps: "P \ \e\<^sub>1, s\<^sub>1,b\<^sub>1\ \* \e',s\<^sub>2,b\<^sub>2\" +shows "P \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \e',s\<^sub>2,b\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(if (true) e\<^sub>1 else e\<^sub>2,s\<^sub>1,b\<^sub>1)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF e_steps]) + also have "(?y, ?z) \ (red P)\<^sup>*" + by(rule RedCondT[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>1_steps]) + ultimately show ?thesis by simp +qed (*>*) 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 +assumes e_steps: "P \ \e,s\<^sub>0,b\<^sub>0\ \* \false,s\<^sub>1,b\<^sub>1\" + and e\<^sub>2_steps: "P \ \e\<^sub>2, s\<^sub>1,b\<^sub>1\ \* \e',s\<^sub>2,b\<^sub>2\" +shows "P \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0,b\<^sub>0\ \* \e',s\<^sub>2,b\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(if (false) e\<^sub>1 else e\<^sub>2,s\<^sub>1,b\<^sub>1)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF e_steps]) + also have "(?y, ?z) \ (red P)\<^sup>*" + by(rule RedCondF[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>2_steps]) + ultimately show ?thesis by simp +qed (*>*) subsubsection "While" lemma WhileFReds: - "P \ \b,s,b\<^sub>0\ \* \false,s',b'\ \ P \ \while (b) c,s,b\<^sub>0\ \* \unit,s',b'\" +assumes b_steps: "P \ \b,s,b\<^sub>0\ \* \false,s',b'\" +shows "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 +by(rule RedWhile[THEN converse_rtrancl_into_rtrancl, + OF CondReds[THEN rtrancl_into_rtrancl, + OF b_steps RedCondF]]) (*>*) lemma WhileRedsThrow: - "P \ \b,s,b\<^sub>0\ \* \throw e,s',b'\ \ P \ \while (b) c,s,b\<^sub>0\ \* \throw e,s',b'\" +assumes b_steps: "P \ \b,s,b\<^sub>0\ \* \throw e,s',b'\" +shows "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 +by(rule RedWhile[THEN converse_rtrancl_into_rtrancl, + OF CondReds[THEN rtrancl_into_rtrancl, + OF b_steps red_reds.CondThrow]]) (*>*) 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 +assumes b_steps: "P \ \b,s\<^sub>0,b\<^sub>0\ \* \true,s\<^sub>1,b\<^sub>1\" + and c_steps: "P \ \c,s\<^sub>1,b\<^sub>1\ \* \Val v\<^sub>1,s\<^sub>2,b\<^sub>2\" + and while_steps: "P \ \while (b) c,s\<^sub>2,b\<^sub>2\ \* \e,s\<^sub>3,b\<^sub>3\" +shows "P \ \while (b) c,s\<^sub>0,b\<^sub>0\ \* \e,s\<^sub>3,b\<^sub>3\" +(*<*)(is "(?a, ?c) \ (red P)\<^sup>*") +proof - + let ?b = "(if (b) (c;; while (b) c) else unit,s\<^sub>0,b\<^sub>0)" + let ?y = "(if (true) (c;; while (b) c) else unit,s\<^sub>1,b\<^sub>1)" + let ?b' = "(c;; while (b) c,s\<^sub>1,b\<^sub>1)" + let ?y' = "(Val v\<^sub>1;; while (b) c,s\<^sub>2,b\<^sub>2)" + have "(?a, ?b) \ (red P)\<^sup>*" + using RedWhile[THEN converse_rtrancl_into_rtrancl] by simp + also have "(?b, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF b_steps]) + also have "(?y, ?b') \ (red P)\<^sup>*" + using RedCondT[THEN converse_rtrancl_into_rtrancl] by simp + also have "(?b', ?y') \ (red P)\<^sup>*" by(rule SeqReds[OF c_steps]) + also have "(?y', ?c) \ (red P)\<^sup>*" + by(rule RedSeq[THEN converse_rtrancl_into_rtrancl, OF while_steps]) + ultimately show ?thesis by simp +qed (*>*) 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 +assumes b_steps: "P \ \b,s\<^sub>0,b\<^sub>0\ \* \true,s\<^sub>1,b\<^sub>1\" + and c_steps: "P \ \c,s\<^sub>1,b\<^sub>1\ \* \throw e,s\<^sub>2,b\<^sub>2\" +shows "P \ \while (b) c,s\<^sub>0,b\<^sub>0\ \* \throw e,s\<^sub>2,b\<^sub>2\" +(*<*)(is "(?a, ?c) \ (red P)\<^sup>*") +proof - + let ?b = "(if (b) (c;; while (b) c) else unit,s\<^sub>0,b\<^sub>0)" + let ?y = "(if (true) (c;; while (b) c) else unit,s\<^sub>1,b\<^sub>1)" + let ?b' = "(c;; while (b) c,s\<^sub>1,b\<^sub>1)" + let ?y' = "(throw e;; while (b) c,s\<^sub>2,b\<^sub>2)" + have "(?a, ?b) \ (red P)\<^sup>*" + using RedWhile[THEN converse_rtrancl_into_rtrancl] by simp + also have "(?b, ?y) \ (red P)\<^sup>*" by(rule CondReds[OF b_steps]) + also have "(?y, ?b') \ (red P)\<^sup>*" + using RedCondT[THEN converse_rtrancl_into_rtrancl] by simp + also have "(?b', ?y') \ (red P)\<^sup>*" by(rule SeqReds[OF c_steps]) + also have "(?y', ?c) \ (red P)" by(rule red_reds.SeqThrow) + ultimately show ?thesis by simp +qed (*>*) 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 + by(rule rtrancl_into_rtrancl[OF step(3) ThrowRed[OF step(2)]]) +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 -(*>*) +(*<*)by(rule ThrowReds[THEN rtrancl_into_rtrancl, OF _ RedThrowNull])(*>*) 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 -(*>*) +(*<*)by(rule ThrowReds[THEN rtrancl_into_rtrancl, OF _ red_reds.ThrowThrow])(*>*) 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 +proof(induct rule: converse_rtrancl_induct3) + case refl then show ?case + by(fastforce simp: fun_upd_same simp del:fun_upd_apply) +next + case (step e0 s0 b0 e1 s1 b1) + obtain h1 l1 sh1 where s1[simp]: "s1 = (h1, l1, sh1)" by(cases s1) simp + { fix h0 l0 sh0 h2 l2 sh2 v0 + assume [simp]: "s0 = (h0, l0(V \ v0), sh0)" and s'[simp]: "s' = (h2, l2, sh2)" + then have "V \ dom l1" using step(1) by(auto dest!: red_lcl_incr) + then obtain v1 where l1V[simp]: "l1 V = \v1\" by blast + + let ?a = "({V:T; V:=Val v0;; e0},(h0, l0, sh0),b0)" + let ?b = "({V:T; V:=Val v1;; e1},(h1, l1(V := l0 V), sh1),b1)" + let ?c = "({V:T; V:=Val (the (l2 V));; e'},(h2, l2(V := l0 V), sh2),b')" + let ?l = "l1(V := l0 V)" and ?v = v1 + + have e0_steps: "P \ \e0,(h0, l0(V \ v0), sh0),b0\ \ \e1,(h1, l1, sh1),b1\" + using step(1) by simp + + have lv: "\l v. l1 = l(V \ v) \ + P \ \{V:T; V:=Val v;; e1},(h1, l, sh1),b1\ \* + \{V:T; V:=Val (the (l2 V));; e'},(h2, l2(V := l V), sh2),b'\" + using step(3) s' s1 by blast + moreover have "l1 = ?l(V \ ?v)" by(rule ext) (simp add:fun_upd_def) + ultimately have "(?b, ?c) \ (red P)\<^sup>*" using lv[of ?l ?v] by simp + then have "(?a, ?c) \ (red P)\<^sup>*" + by(rule InitBlockRed[THEN converse_rtrancl_into_rtrancl, OF e0_steps l1V]) + } + then show ?case by blast +qed (*>*) 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 +by(fast elim!:InitBlockReds[THEN rtrancl_into_rtrancl] finalE + intro:RedInitBlock InitBlockThrow) (*>*) 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 + by(rule rtrancl_into_rtrancl[OF step(3) TryRed[OF step(2)]]) +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 -(*>*) +(*<*)by(rule TryReds[THEN rtrancl_into_rtrancl, OF _ RedTry])(*>*) 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 +assumes e\<^sub>1_steps: "P \ \e\<^sub>1,s\<^sub>0,b\<^sub>0\ \* \Throw a,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b\<^sub>1\" + and h\<^sub>1a: "h\<^sub>1 a = Some(D,fs)" and sub: "P \ D \\<^sup>* C" + and e\<^sub>2_steps: "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\" + and final: "final e\<^sub>2'" +shows "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::locals)(V := l\<^sub>1 V),sh\<^sub>2),b\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(try Throw a catch(C V) e\<^sub>2,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b\<^sub>1)" + let ?b = "({V:Class C; V:=addr a;; e\<^sub>2},(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b\<^sub>1)" + have bz: "(?b, ?z) \ (red P)\<^sup>*" + by(rule InitBlockRedsFinal[OF e\<^sub>2_steps final]) + have hp: "hp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) a = \(D, fs)\" using h\<^sub>1a by simp + have "(?x, ?y) \ (red P)\<^sup>*" by(rule TryReds[OF e\<^sub>1_steps]) + also have "(?y, ?z) \ (red P)\<^sup>*" + by(rule RedTryCatch[THEN converse_rtrancl_into_rtrancl, OF hp sub bz]) + ultimately show ?thesis by simp +qed (*>*) 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 -(*>*) +(*<*)by(fastforce intro!: TryReds[THEN rtrancl_into_rtrancl, OF _ RedTryFail])(*>*) 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 + by(rule rtrancl_into_rtrancl[OF step(3) ListRed1[OF step(2)]]) +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 + by(rule rtrancl_into_rtrancl[OF step(3) ListRed2[OF step(2)]]) +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 -(*>*) +(*<*)by(rule rtrancl_trans[OF ListReds1 ListReds2])(*>*) 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" +assumes wf: "wwf_J_prog P" +shows "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 +proof(induct rule: converse_rtrancl_induct_red) + case 1 then show ?case by blast +next + case 2 then show ?case using wf by(blast dest: Red_fv Red_dom_lcl) +qed (*>*) 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 +(*<*)by(rule ext) (simp add:override_on_def) 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 + by(rule rtrancl_into_rtrancl[OF step(3) CallObj[OF step(2)]]) +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 + by(rule rtrancl_into_rtrancl[OF step(3) CallParams[OF step(2)]]) +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 +assumes e_steps: "P \ \e,s\<^sub>0,b\<^sub>0\ \* \Val v,s\<^sub>1,b\<^sub>1\" + and es_steps: "P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \map Val vs\<^sub>1 @ throw a # es\<^sub>2,s\<^sub>2,b\<^sub>2\" +shows "P \ \e\M(es),s\<^sub>0,b\<^sub>0\ \* \throw a,s\<^sub>2,b\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(Val v\M(es),s\<^sub>1,b\<^sub>1)" + let ?y' = "(Val v\M(map Val vs\<^sub>1 @ throw a # es\<^sub>2),s\<^sub>2,b\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps]) + also have "(?y', ?z) \ (red P)\<^sup>*" using CallThrowParams by fast + ultimately show ?thesis by simp +qed (*>*) 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 -(*>*) +(*<*)by(rule CallRedsObj[THEN rtrancl_into_rtrancl, OF _ CallThrowObj])(*>*) 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 +assumes e_steps: "P \ \e,s\<^sub>0,b\<^sub>0\ \* \null,s\<^sub>1,b\<^sub>1\" + and es_steps: "P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \map Val vs,s\<^sub>2,b\<^sub>2\" +shows "P \ \e\M(es),s\<^sub>0,b\<^sub>0\ \* \THROW NullPointer,s\<^sub>2,b\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(null\M(es),s\<^sub>1,b\<^sub>1)" + let ?y' = "(null\M(map Val vs),s\<^sub>2,b\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps]) + also have "(?y', ?z) \ (red P)" by(rule RedCallNull) + ultimately show ?thesis by simp +qed (*>*) 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 +assumes e_steps: "P \ \e,s,b\ \* \addr a,s\<^sub>1,b\<^sub>1\" + and es_steps: "P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \map Val vs,s\<^sub>2,b\<^sub>2\" + and hp\<^sub>2a: "hp s\<^sub>2 a = Some(C,fs)" + and ncM: "\(\b Ts T m D. P \ C sees M,b:Ts\T = m in D)" +shows "P \ \e\M(es),s,b\ \* \THROW NoSuchMethodError,s\<^sub>2,b\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(addr a\M(es),s\<^sub>1,b\<^sub>1)" + let ?y' = "(addr a\M(map Val vs),s\<^sub>2,b\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps]) + also have "(?y', ?z) \ (red P)" using RedCallNone[OF _ ncM] hp\<^sub>2a + by(cases s\<^sub>2) simp + ultimately show ?thesis by simp +qed (*>*) 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 +assumes e_steps: "P \ \e,s,b\ \* \addr a,s\<^sub>1,b\<^sub>1\" + and es_steps: "P \ \es,s\<^sub>1,b\<^sub>1\ [\]* \map Val vs,s\<^sub>2,b\<^sub>2\" + and hp\<^sub>2a: "hp s\<^sub>2 a = Some(C,fs)" + and cM_Static: "P \ C sees M,Static:Ts\T = m in D" +shows "P \ \e\M(es),s,b\ \* \THROW IncompatibleClassChangeError,s\<^sub>2,b\<^sub>2\" +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") +proof - + let ?y = "(addr a\M(es),s\<^sub>1,b\<^sub>1)" + let ?y' = "(addr a\M(map Val vs),s\<^sub>2,b\<^sub>2)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps]) + also have "(?y, ?y') \ (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps]) + also have "(?y', ?z) \ (red P)" using RedCallStatic[OF _ cM_Static] hp\<^sub>2a + by(cases s\<^sub>2) simp + ultimately show ?thesis by simp +qed (*>*) 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 + by(rule rtrancl_into_rtrancl[OF step(3) SCallParams[OF step(2)]]) +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 +by(erule SCallRedsParams[THEN rtrancl_into_rtrancl, OF _ SCallThrowParams]) + simp (*>*) 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 +by(erule SCallRedsParams[THEN rtrancl_into_rtrancl, OF _ RedSCallNone]) + simp (*>*) 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 +by(erule SCallRedsParams[THEN rtrancl_into_rtrancl, OF _ RedSCallNonStatic]) + simp (*>*) 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\" -(*<*) +(*<*)(is "(?x, ?z) \ (red P)\<^sup>*") 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\" + let ?y = "(C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False)" + let ?y' = "(INIT D ([D],False) \ C\\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False)" + have "(?x, ?y) \ (red P)\<^sup>*" by(rule SCallRedsParams[OF assms(2)]) + also have "(?y, ?y') \ (red P)\<^sup>*" 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)]) + also have "(?y', ?z) \ (red P)\<^sup>*" 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 +by(auto dest!: eval_init_return [where C'=C'] + append_butlast_last_id[THEN sym]) 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 +by(auto dest!: eval_init_return [where C'=C'] + append_butlast_last_id[THEN sym]) \ \ 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'\" +shows "\ P \ \e,s,b\ \* \e'',s'',b''\; P \ \e'',s''\ \ \e',s'\; + iconf (shp s) e; P,shp s \\<^sub>b (e::expr,b) \ \ + \ 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'\" +shows "\ P \ \es,s,b\ [\]* \es'',s'',b''\; P \ \es'',s''\ [\] \es',s'\; + iconfs (shp s) es; P,shp s \\<^sub>b (es::expr list,b) \ \ + \ 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/Expr.thy b/thys/JinjaDCI/J/Expr.thy --- a/thys/JinjaDCI/J/Expr.thy +++ b/thys/JinjaDCI/J/Expr.thy @@ -1,572 +1,579 @@ (* Title: JinjaDCI/J/Expr.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory J/Expr.thy by Tobias Nipkow *) section \ Expressions \ theory Expr imports "../Common/Exceptions" begin datatype bop = Eq | Add \ \names of binary operations\ datatype 'a exp = new cname \ \class instance creation\ | Cast cname "('a exp)" \ \type cast\ | Val val \ \value\ | BinOp "('a exp)" bop "('a exp)" ("_ \_\ _" [80,0,81] 80) \ \binary operation\ | Var 'a \ \local variable (incl. parameter)\ | LAss 'a "('a exp)" ("_:=_" [90,90]90) \ \local assignment\ | FAcc "('a exp)" vname cname ("_\_{_}" [10,90,99]90) \ \field access\ | SFAcc cname vname cname ("_\\<^sub>s_{_}" [10,90,99]90) \ \static field access\ | FAss "('a exp)" vname cname "('a exp)" ("_\_{_} := _" [10,90,99,90]90) \ \field assignment\ | SFAss cname vname cname "('a exp)" ("_\\<^sub>s_{_} := _" [10,90,99,90]90) \ \static field assignment\ | Call "('a exp)" mname "('a exp list)" ("_\_'(_')" [90,99,0] 90) \ \method call\ | SCall cname mname "('a exp list)" ("_\\<^sub>s_'(_')" [90,99,0] 90) \ \static method call\ | Block 'a ty "('a exp)" ("'{_:_; _}") | Seq "('a exp)" "('a exp)" ("_;;/ _" [61,60]60) | Cond "('a exp)" "('a exp)" "('a exp)" ("if '(_') _/ else _" [80,79,79]70) | While "('a exp)" "('a exp)" ("while '(_') _" [80,79]70) | throw "('a exp)" | TryCatch "('a exp)" cname 'a "('a exp)" ("try _/ catch'(_ _') _" [0,99,80,79] 70) | INIT cname "cname list" bool "('a exp)" ("INIT _ '(_,_') \ _" [60,60,60,60] 60) \ \internal initialization command: class, list of superclasses to initialize, preparation flag; command on hold\ | RI cname "('a exp)" "cname list" "('a exp)" ("RI '(_,_') ; _ \ _" [60,60,60,60] 60) \ \running of the initialization procedure for class with expression, classes still to initialize command on hold\ type_synonym expr = "vname exp" \ \Jinja expression\ type_synonym J_mb = "vname list \ expr" \ \Jinja method body: parameter names and expression\ type_synonym J_prog = "J_mb prog" \ \Jinja program\ type_synonym init_stack = "expr list \ bool" \ \Stack of expressions waiting on initialization in small step; indicator boolean True if current expression has been init checked\ text\The semantics of binary operators: \ fun binop :: "bop \ val \ val \ val option" where "binop(Eq,v\<^sub>1,v\<^sub>2) = Some(Bool (v\<^sub>1 = v\<^sub>2))" | "binop(Add,Intg i\<^sub>1,Intg i\<^sub>2) = Some(Intg(i\<^sub>1+i\<^sub>2))" | "binop(bop,v\<^sub>1,v\<^sub>2) = None" lemma [simp]: "(binop(Add,v\<^sub>1,v\<^sub>2) = Some v) = (\i\<^sub>1 i\<^sub>2. v\<^sub>1 = Intg i\<^sub>1 \ v\<^sub>2 = Intg i\<^sub>2 \ v = Intg(i\<^sub>1+i\<^sub>2))" -(*<*) -apply(cases v\<^sub>1) -apply auto -apply(cases v\<^sub>2) -apply auto -done -(*>*) +(*<*)by(cases v\<^sub>1; cases v\<^sub>2) auto(*>*) lemma map_Val_throw_eq: "map Val vs @ throw ex # es = map Val vs' @ throw ex' # es' \ ex = ex'" -apply(induct vs arbitrary: vs') - apply(case_tac vs', auto)+ -done +(*<*)by(induct vs arbitrary: vs'; case_tac vs') auto(*>*) lemma map_Val_nthrow_neq: "map Val vs = map Val vs' @ throw ex' # es' \ False" -apply(induct vs arbitrary: vs') - apply(case_tac vs', auto)+ -done +(*<*)by(induct vs arbitrary: vs'; case_tac vs') auto(*>*) lemma map_Val_eq: "map Val vs = map Val vs' \ vs = vs'" -apply(induct vs arbitrary: vs') - apply(case_tac vs', auto)+ -done +(*<*)by(induct vs arbitrary: vs'; case_tac vs') auto(*>*) lemma init_rhs_neq [simp]: "e \ INIT C (Cs,b) \ e" proof - have "size e \ size (INIT C (Cs,b) \ e)" by auto then show ?thesis by fastforce qed lemma init_rhs_neq' [simp]: "INIT C (Cs,b) \ e \ e" proof - have "size e \ size (INIT C (Cs,b) \ e)" by auto then show ?thesis by fastforce qed lemma ri_rhs_neq [simp]: "e \ RI(C,e');Cs \ e" proof - have "size e \ size (RI(C,e');Cs \ e)" by auto then show ?thesis by fastforce qed lemma ri_rhs_neq' [simp]: "RI(C,e');Cs \ e \ e" proof - have "size e \ size (RI(C,e');Cs \ e)" by auto then show ?thesis by fastforce qed subsection "Syntactic sugar" abbreviation (input) InitBlock:: "'a \ ty \ 'a exp \ 'a exp \ 'a exp" ("(1'{_:_ := _;/ _})") where "InitBlock V T e1 e2 == {V:T; V := e1;; e2}" abbreviation unit where "unit == Val Unit" abbreviation null where "null == Val Null" abbreviation "addr a == Val(Addr a)" abbreviation "true == Val(Bool True)" abbreviation "false == Val(Bool False)" abbreviation Throw :: "addr \ 'a exp" where "Throw a == throw(Val(Addr a))" abbreviation THROW :: "cname \ 'a exp" where "THROW xc == Throw(addr_of_sys_xcpt xc)" subsection\Free Variables\ primrec fv :: "expr \ vname set" and fvs :: "expr list \ vname set" where "fv(new C) = {}" | "fv(Cast C e) = fv e" | "fv(Val v) = {}" | "fv(e\<^sub>1 \bop\ e\<^sub>2) = fv e\<^sub>1 \ fv e\<^sub>2" | "fv(Var V) = {V}" | "fv(LAss V e) = {V} \ fv e" | "fv(e\F{D}) = fv e" | "fv(C\\<^sub>sF{D}) = {}" | "fv(e\<^sub>1\F{D}:=e\<^sub>2) = fv e\<^sub>1 \ fv e\<^sub>2" | "fv(C\\<^sub>sF{D}:=e\<^sub>2) = fv e\<^sub>2" | "fv(e\M(es)) = fv e \ fvs es" | "fv(C\\<^sub>sM(es)) = fvs es" | "fv({V:T; e}) = fv e - {V}" | "fv(e\<^sub>1;;e\<^sub>2) = fv e\<^sub>1 \ fv e\<^sub>2" | "fv(if (b) e\<^sub>1 else e\<^sub>2) = fv b \ fv e\<^sub>1 \ fv e\<^sub>2" | "fv(while (b) e) = fv b \ fv e" | "fv(throw e) = fv e" | "fv(try e\<^sub>1 catch(C V) e\<^sub>2) = fv e\<^sub>1 \ (fv e\<^sub>2 - {V})" | "fv(INIT C (Cs,b) \ e) = fv e" | "fv(RI (C,e);Cs \ e') = fv e \ fv e'" | "fvs([]) = {}" | "fvs(e#es) = fv e \ fvs es" lemma [simp]: "fvs(es\<^sub>1 @ es\<^sub>2) = fvs es\<^sub>1 \ fvs es\<^sub>2" (*<*)by (induct es\<^sub>1 type:list) auto(*>*) lemma [simp]: "fvs(map Val vs) = {}" (*<*)by (induct vs) auto(*>*) subsection\Accessing expression constructor arguments\ fun val_of :: "'a exp \ val option" where "val_of (Val v) = Some v" | "val_of _ = None" lemma val_of_spec: "val_of e = Some v \ e = Val v" proof(cases e) qed(auto) fun lass_val_of :: "'a exp \ ('a \ val) option" where "lass_val_of (V:=Val v) = Some (V, v)" | "lass_val_of _ = None" lemma lass_val_of_spec: assumes "lass_val_of e = \a\" shows "e = (fst a:=Val (snd a))" using assms proof(cases e) case (LAss V e') then show ?thesis using assms proof(cases e')qed(auto) qed(auto) fun map_vals_of :: "'a exp list \ val list option" where "map_vals_of (e#es) = (case val_of e of Some v \ (case map_vals_of es of Some vs \ Some (v#vs) | _ \ None) | _ \ None)" | "map_vals_of [] = Some []" lemma map_vals_of_spec: "map_vals_of es = Some vs \ es = map Val vs" proof(induct es arbitrary: vs) qed(auto simp: val_of_spec) lemma map_vals_of_Vals[simp]: "map_vals_of (map Val vs) = \vs\" by(induct vs, auto) lemma map_vals_of_throw[simp]: "map_vals_of (map Val vs @ throw e # es') = None" by(induct vs, auto) fun bool_of :: "'a exp \ bool option" where "bool_of true = Some True" | "bool_of false = Some False" | "bool_of _ = None" lemma bool_of_specT: assumes "bool_of e = Some True" shows "e = true" proof - have "bool_of e = Some True" by fact then show ?thesis proof(cases e) case (Val x3) with assms show ?thesis proof(cases x3) case (Bool x) with assms Val show ?thesis proof(cases x)qed(auto) qed(simp_all) qed(auto) qed lemma bool_of_specF: assumes "bool_of e = Some False" shows "e = false" proof - have "bool_of e = Some False" by fact then show ?thesis proof(cases e) case (Val x3) with assms show ?thesis proof(cases x3) case (Bool x) with assms Val show ?thesis proof(cases x)qed(auto) qed(simp_all) qed(auto) qed fun throw_of :: "'a exp \ 'a exp option" where "throw_of (throw e') = Some e'" | "throw_of _ = None" lemma throw_of_spec: "throw_of e = Some e' \ e = throw e'" proof(cases e) qed(auto) fun init_exp_of :: "'a exp \ 'a exp option" where "init_exp_of (INIT C (Cs,b) \ e) = Some e" | "init_exp_of (RI(C,e');Cs \ e) = Some e" | "init_exp_of _ = None" lemma init_exp_of_neq [simp]: "init_exp_of e = \e'\ \ e' \ e" by(cases e, auto) lemma init_exp_of_neq'[simp]: "init_exp_of e = \e'\ \ e \ e'" by(cases e, auto) subsection\Class initialization\ text \ This section defines a few functions that return information about an expression's current initialization status. \ \ \ True if expression contains @{text INIT}, @{text RI}, or a call to a static method @{term clinit} \ primrec sub_RI :: "'a exp \ bool" and sub_RIs :: "'a exp list \ bool" where "sub_RI(new C) = False" | "sub_RI(Cast C e) = sub_RI e" | "sub_RI(Val v) = False" | "sub_RI(e\<^sub>1 \bop\ e\<^sub>2) = (sub_RI e\<^sub>1 \ sub_RI e\<^sub>2)" | "sub_RI(Var V) = False" | "sub_RI(LAss V e) = sub_RI e" | "sub_RI(e\F{D}) = sub_RI e" | "sub_RI(C\\<^sub>sF{D}) = False" | "sub_RI(e\<^sub>1\F{D}:=e\<^sub>2) = (sub_RI e\<^sub>1 \ sub_RI e\<^sub>2)" | "sub_RI(C\\<^sub>sF{D}:=e\<^sub>2) = sub_RI e\<^sub>2" | "sub_RI(e\M(es)) = (sub_RI e \ sub_RIs es)" | "sub_RI(C\\<^sub>sM(es)) = (M = clinit \ sub_RIs es)" | "sub_RI({V:T; e}) = sub_RI e" | "sub_RI(e\<^sub>1;;e\<^sub>2) = (sub_RI e\<^sub>1 \ sub_RI e\<^sub>2)" | "sub_RI(if (b) e\<^sub>1 else e\<^sub>2) = (sub_RI b \ sub_RI e\<^sub>1 \ sub_RI e\<^sub>2)" | "sub_RI(while (b) e) = (sub_RI b \ sub_RI e)" | "sub_RI(throw e) = sub_RI e" | "sub_RI(try e\<^sub>1 catch(C V) e\<^sub>2) = (sub_RI e\<^sub>1 \ sub_RI e\<^sub>2)" | "sub_RI(INIT C (Cs,b) \ e) = True" | "sub_RI(RI (C,e);Cs \ e') = True" | "sub_RIs([]) = False" | "sub_RIs(e#es) = (sub_RI e \ sub_RIs es)" lemmas sub_RI_sub_RIs_induct = sub_RI.induct sub_RIs.induct lemma nsub_RIs_def[simp]: "\sub_RIs es \ \e \ set es. \sub_RI e" by(induct es, auto) lemma sub_RI_base: "e = INIT C (Cs, b) \ e' \ e = RI(C,e\<^sub>0);Cs \ e' \ sub_RI e" by(cases e, auto) lemma nsub_RI_Vals[simp]: "\sub_RIs (map Val vs)" by(induct vs, auto) lemma lass_val_of_nsub_RI: "lass_val_of e = \a\ \ \sub_RI e" by(drule lass_val_of_spec, simp) \ \ is not currently initializing class @{text C'} (point past checking flag) \ primrec not_init :: "cname \ 'a exp \ bool" and not_inits :: "cname \ 'a exp list \ bool" where "not_init C' (new C) = True" | "not_init C' (Cast C e) = not_init C' e" | "not_init C' (Val v) = True" | "not_init C' (e\<^sub>1 \bop\ e\<^sub>2) = (not_init C' e\<^sub>1 \ not_init C' e\<^sub>2)" | "not_init C' (Var V) = True" | "not_init C' (LAss V e) = not_init C' e" | "not_init C' (e\F{D}) = not_init C' e" | "not_init C' (C\\<^sub>sF{D}) = True" | "not_init C' (e\<^sub>1\F{D}:=e\<^sub>2) = (not_init C' e\<^sub>1 \ not_init C' e\<^sub>2)" | "not_init C' (C\\<^sub>sF{D}:=e\<^sub>2) = not_init C' e\<^sub>2" | "not_init C' (e\M(es)) = (not_init C' e \ not_inits C' es)" | "not_init C' (C\\<^sub>sM(es)) = not_inits C' es" | "not_init C' ({V:T; e}) = not_init C' e" | "not_init C' (e\<^sub>1;;e\<^sub>2) = (not_init C' e\<^sub>1 \ not_init C' e\<^sub>2)" | "not_init C' (if (b) e\<^sub>1 else e\<^sub>2) = (not_init C' b \ not_init C' e\<^sub>1 \ not_init C' e\<^sub>2)" | "not_init C' (while (b) e) = (not_init C' b \ not_init C' e)" | "not_init C' (throw e) = not_init C' e" | "not_init C' (try e\<^sub>1 catch(C V) e\<^sub>2) = (not_init C' e\<^sub>1 \ not_init C' e\<^sub>2)" | "not_init C' (INIT C (Cs,b) \ e) = ((b \ Cs = Nil \ C' \ hd Cs) \ C' \ set(tl Cs) \ not_init C' e)" | "not_init C' (RI (C,e);Cs \ e') = (C' \ set (C#Cs) \ not_init C' e \ not_init C' e')" | "not_inits C' ([]) = True" | "not_inits C' (e#es) = (not_init C' e \ not_inits C' es)" lemma not_inits_def'[simp]: "not_inits C es \ \e \ set es. not_init C e" by(induct es, auto) lemma nsub_RIs_not_inits_aux: "\e \ set es. \sub_RI e \ not_init C e \ \sub_RIs es \ not_inits C es" by(induct es, auto) lemma nsub_RI_not_init: "\sub_RI e \ not_init C e" proof(induct e) qed(auto intro: nsub_RIs_not_inits_aux) lemma nsub_RIs_not_inits: "\sub_RIs es \ not_inits C es" -apply(rule nsub_RIs_not_inits_aux) - apply(simp_all add: nsub_RI_not_init) -done +by(rule nsub_RIs_not_inits_aux) (simp_all add: nsub_RI_not_init) subsection\Subexpressions\ \ \ all strictly smaller subexpressions; does not include self \ primrec subexp :: "'a exp \ 'a exp set" and subexps :: "'a exp list \ 'a exp set" where "subexp(new C) = {}" | "subexp(Cast C e) = {e} \ subexp e" | "subexp(Val v) = {}" | "subexp(e\<^sub>1 \bop\ e\<^sub>2) = {e\<^sub>1, e\<^sub>2} \ subexp e\<^sub>1 \ subexp e\<^sub>2" | "subexp(Var V) = {}" | "subexp(LAss V e) = {e} \ subexp e" | "subexp(e\F{D}) = {e} \ subexp e" | "subexp(C\\<^sub>sF{D}) = {}" | "subexp(e\<^sub>1\F{D}:=e\<^sub>2) = {e\<^sub>1, e\<^sub>2} \ subexp e\<^sub>1 \ subexp e\<^sub>2" | "subexp(C\\<^sub>sF{D}:=e\<^sub>2) = {e\<^sub>2} \subexp e\<^sub>2" | "subexp(e\M(es)) = {e} \ set es \ subexp e \ subexps es" | "subexp(C\\<^sub>sM(es)) = set es \ subexps es" | "subexp({V:T; e}) = {e} \ subexp e" | "subexp(e\<^sub>1;;e\<^sub>2) = {e\<^sub>1, e\<^sub>2} \ subexp e\<^sub>1 \ subexp e\<^sub>2" | "subexp(if (b) e\<^sub>1 else e\<^sub>2) = {b, e\<^sub>1, e\<^sub>2} \ subexp b \ subexp e\<^sub>1 \ subexp e\<^sub>2" | "subexp(while (b) e) = {b, e} \ subexp b \ subexp e" | "subexp(throw e) = {e} \ subexp e" | "subexp(try e\<^sub>1 catch(C V) e\<^sub>2) = {e\<^sub>1, e\<^sub>2} \ subexp e\<^sub>1 \ subexp e\<^sub>2" | "subexp(INIT C (Cs,b) \ e) = {e} \ subexp e" | "subexp(RI (C,e);Cs \ e') = {e, e'} \ subexp e \ subexp e'" | "subexps([]) = {}" | "subexps(e#es) = {e} \ subexp e \ subexps es" lemmas subexp_subexps_induct = subexp.induct subexps.induct abbreviation subexp_of :: "'a exp \ 'a exp \ bool" where "subexp_of e e' \ e \ subexp e'" lemma subexp_size_le: "(e' \ subexp e \ size e' < size e) \ (e' \ subexps es \ size e' < size_list size es)" proof(induct rule: subexp_subexps.induct) case Call:11 then show ?case using not_less_eq size_list_estimation by fastforce next case SCall:12 then show ?case using not_less_eq size_list_estimation by fastforce qed(auto) lemma subexps_def2: "subexps es = set es \ (\e \ set es. subexp e)" by(induct es, auto) \ \ strong induction \ lemma shows subexp_induct[consumes 1]: "(\e. subexp e = {} \ R e) \ (\e. (\e'. e' \ subexp e \ R e') \ R e) \ (\es. (\e'. e' \ subexps es \ R e') \ Rs es) \ (\e'. e' \ subexp e \ R e') \ R e" and subexps_induct[consumes 1]: "(\es. subexps es = {} \ Rs es) \ (\e. (\e'. e' \ subexp e \ R e') \ R e) \ (\es. (\e'. e' \ subexps es \ R e') \ Rs es) \ (\e'. e' \ subexps es \ R e') \ Rs es" proof(induct rule: subexp_subexps_induct) case (Cast x1 x2) then have "(\e'. subexp_of e' x2 \ R e') \ R x2" by fast then have "(\e'. subexp_of e' (Cast x1 x2) \ R e')" by auto then show ?case using Cast.prems(2) by fast next case (BinOp x1 x2 x3) then have "(\e'. subexp_of e' x1 \ R e') \ R x1" and "(\e'. subexp_of e' x3 \ R e') \ R x3" by fast+ then have "(\e'. subexp_of e' (x1 \x2\ x3) \ R e')" by auto then show ?case using BinOp.prems(2) by fast next case (LAss x1 x2) then have "(\e'. subexp_of e' x2 \ R e') \ R x2" by fast then have "(\e'. subexp_of e' (LAss x1 x2) \ R e')" by auto then show ?case using LAss.prems(2) by fast next case (FAcc x1 x2 x3) then have "(\e'. subexp_of e' x1 \ R e') \ R x1" by fast then have "(\e'. subexp_of e' (x1\x2{x3}) \ R e')" by auto then show ?case using FAcc.prems(2) by fast next case (FAss x1 x2 x3 x4) then have "(\e'. subexp_of e' x1 \ R e') \ R x1" and "(\e'. subexp_of e' x4 \ R e') \ R x4" by fast+ then have "(\e'. subexp_of e' (x1\x2{x3} := x4) \ R e')" by auto then show ?case using FAss.prems(2) by fast next case (SFAss x1 x2 x3 x4) then have "(\e'. subexp_of e' x4 \ R e') \ R x4" by fast then have "(\e'. subexp_of e' (x1\\<^sub>sx2{x3} := x4) \ R e')" by auto then show ?case using SFAss.prems(2) by fast next case (Call x1 x2 x3) then have "(\e'. subexp_of e' x1 \ R e') \ R x1" and "(\e'. e' \ subexps x3 \ R e') \ Rs x3" by fast+ then have "(\e'. subexp_of e' (x1\x2(x3)) \ R e')" using subexps_def2 by auto then show ?case using Call.prems(2) by fast next case (SCall x1 x2 x3) then have "(\e'. e' \ subexps x3 \ R e') \ Rs x3" by fast then have "(\e'. subexp_of e' (x1\\<^sub>sx2(x3)) \ R e')" using subexps_def2 by auto then show ?case using SCall.prems(2) by fast next case (Block x1 x2 x3) then have "(\e'. subexp_of e' x3 \ R e') \ R x3" by fast then have "(\e'. subexp_of e' {x1:x2; x3} \ R e')" by auto then show ?case using Block.prems(2) by fast next case (Seq x1 x2) then have "(\e'. subexp_of e' x1 \ R e') \ R x1" and "(\e'. subexp_of e' x2 \ R e') \ R x2" by fast+ then have "(\e'. subexp_of e' (x1;; x2) \ R e')" by auto then show ?case using Seq.prems(2) by fast next case (Cond x1 x2 x3) then have "(\e'. subexp_of e' x1 \ R e') \ R x1" and "(\e'. subexp_of e' x2 \ R e') \ R x2" and "(\e'. subexp_of e' x3 \ R e') \ R x3" by fast+ then have "(\e'. subexp_of e' (if (x1) x2 else x3) \ R e')" by auto then show ?case using Cond.prems(2) by fast next case (While x1 x2) then have "(\e'. subexp_of e' x1 \ R e') \ R x1" and "(\e'. subexp_of e' x2 \ R e') \ R x2" by fast+ then have "(\e'. subexp_of e' (while (x1) x2) \ R e')" by auto then show ?case using While.prems(2) by fast next case (throw x) then have "(\e'. subexp_of e' x \ R e') \ R x" by fast then have "(\e'. subexp_of e' (throw x) \ R e')" by auto then show ?case using throw.prems(2) by fast next case (TryCatch x1 x2 x3 x4) then have "(\e'. subexp_of e' x1 \ R e') \ R x1" and "(\e'. subexp_of e' x4 \ R e') \ R x4" by fast+ then have "(\e'. subexp_of e' (try x1 catch(x2 x3) x4) \ R e')" by auto then show ?case using TryCatch.prems(2) by fast next case (INIT x1 x2 x3 x4) then have "(\e'. subexp_of e' x4 \ R e') \ R x4" by fast then have "(\e'. subexp_of e' (INIT x1 (x2,x3) \ x4) \ R e')" by auto then show ?case using INIT.prems(2) by fast next case (RI x1 x2 x3 x4) then have "(\e'. subexp_of e' x2 \ R e') \ R x2" and "(\e'. subexp_of e' x4 \ R e') \ R x4" by fast+ then have "(\e'. subexp_of e' (RI (x1,x2) ; x3 \ x4) \ R e')" by auto then show ?case using RI.prems(2) by fast next case (Cons_exp x1 x2) then have "(\e'. subexp_of e' x1 \ R e') \ R x1" and "(\e'. e' \ subexps x2 \ R e') \ Rs x2" by fast+ then have "(\e'. e' \ subexps (x1 # x2) \ R e')" using subexps_def2 by auto then show ?case using Cons_exp.prems(3) by fast qed(auto) subsection"Final expressions" (* these definitions and most of the lemmas were in BigStep.thy in the original Jinja *) definition final :: "'a exp \ bool" where "final e \ (\v. e = Val v) \ (\a. e = Throw a)" definition finals:: "'a exp list \ bool" where "finals es \ (\vs. es = map Val vs) \ (\vs a es'. es = map Val vs @ Throw a # es')" lemma [simp]: "final(Val v)" (*<*)by(simp add:final_def)(*>*) lemma [simp]: "final(throw e) = (\a. e = addr a)" (*<*)by(simp add:final_def)(*>*) lemma finalE: "\ final e; \v. e = Val v \ R; \a. e = Throw a \ R \ \ R" (*<*)by(auto simp:final_def)(*>*) lemma final_fv[iff]: "final e \ fv e = {}" by (auto simp: final_def) lemma finalsE: "\ finals es; \vs. es = map Val vs \ R; \vs a es'. es = map Val vs @ Throw a # es' \ R \ \ R" (*<*)by(auto simp:finals_def)(*>*) lemma [iff]: "finals []" (*<*)by(simp add:finals_def)(*>*) lemma [iff]: "finals (Val v # es) = finals es" (*<*) -apply(clarsimp simp add: finals_def) -apply(rule iffI) - apply(erule disjE) - apply simp - apply(rule disjI2) - apply clarsimp - apply(case_tac vs) - apply simp - apply fastforce -apply(erule disjE) - apply clarsimp -apply(rule disjI2) -apply clarsimp -apply(rule_tac x = "v#vs" in exI) -apply simp -done +proof(rule iffI) + assume "finals (Val v # es)" + moreover { + fix vs a es' + assume "\vs a es'. es \ map Val vs @ Throw a # es'" + and "Val v # es = map Val vs @ Throw a # es'" + then have "\vs. es = map Val vs" by(case_tac vs; simp) + } + ultimately show "finals es" by(clarsimp simp add: finals_def) +next + assume "finals es" + moreover { + fix vs a es' + assume "es = map Val vs @ Throw a # es'" + then have "\vs' a' es''. Val v # map Val vs @ Throw a # es' = map Val vs' @ Throw a' # es''" + by(rule_tac x = "v#vs" in exI) simp + } + ultimately show "finals (Val v # es)" by(clarsimp simp add: finals_def) +qed (*>*) lemma finals_app_map[iff]: "finals (map Val vs @ es) = finals es" (*<*)by(induct_tac vs, auto)(*>*) lemma [iff]: "finals (map Val vs)" (*<*)using finals_app_map[of vs "[]"]by(simp)(*>*) lemma [iff]: "finals (throw e # es) = (\a. e = addr a)" (*<*) -apply(simp add:finals_def) -apply(rule iffI) - apply clarsimp - apply(case_tac vs) - apply simp - apply fastforce -apply clarsimp -apply(rule_tac x = "[]" in exI) -apply simp -done +proof(rule iffI) + assume "finals (throw e # es)" + moreover { + fix vs a es' + assume "throw e # es = map Val vs @ Throw a # es'" + then have "\a. e = addr a" by(case_tac vs; simp) + } + ultimately show "\a. e = addr a" by(clarsimp simp add: finals_def) +next + assume "\a. e = addr a" + moreover { + fix vs a es' + assume "e = addr a" + then have "\vs aa es'. Throw a # es = map Val vs @ Throw aa # es'" + by(rule_tac x = "[]" in exI) simp + } + ultimately show "finals (throw e # es)" by(clarsimp simp add: finals_def) +qed (*>*) lemma not_finals_ConsI: "\ final e \ \ finals(e#es)" - (*<*) -apply(clarsimp simp add:finals_def final_def) -apply(case_tac vs) -apply auto -done +(*<*) +proof - + assume "\ final e" + moreover { + fix vs a es' + assume "\v. e \ Val v" and "\a. e \ Throw a" + then have "e # es \ map Val vs @ Throw a # es'" by(case_tac vs; simp) + } + ultimately show ?thesis by(clarsimp simp add:finals_def final_def) +qed (*>*) lemma not_finals_ConsI2: "e = Val v \ \ finals es \ \ finals(e#es)" - (*<*) -apply(clarsimp simp add:finals_def final_def) -apply(case_tac vs) -apply auto -done +(*<*) +proof - + assume [simp]: "e = Val v" and "\ finals es" + moreover { + fix vs a es' + assume "\vs. es \ map Val vs" and "\vs a es'. es \ map Val vs @ Throw a # es'" + then have "e # es \ map Val vs @ Throw a # es'" by(case_tac vs; simp) + } + ultimately show ?thesis by(clarsimp simp add:finals_def final_def) +qed (*>*) end diff --git a/thys/JinjaDCI/J/JWellForm.thy b/thys/JinjaDCI/J/JWellForm.thy --- a/thys/JinjaDCI/J/JWellForm.thy +++ b/thys/JinjaDCI/J/JWellForm.thy @@ -1,82 +1,73 @@ (* Title: JinjaDCI/J/JWellForm.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory J/JWellForm.thy by Tobias Nipkow *) section \ Well-formedness Constraints \ theory JWellForm imports "../Common/WellForm" WWellForm WellType DefAss begin definition wf_J_mdecl :: "J_prog \ cname \ J_mb mdecl \ bool" where "wf_J_mdecl P C \ \(M,b,Ts,T,(pns,body)). length Ts = length pns \ distinct pns \ \sub_RI body \ (case b of NonStatic \ this \ set pns \ (\T'. P,[this\Class C,pns[\]Ts] \ body :: T' \ P \ T' \ T) \ \ body \{this} \ set pns\ | Static \ (\T'. P,[pns[\]Ts] \ body :: T' \ P \ T' \ T) \ \ body \set pns\)" lemma wf_J_mdecl_NonStatic[simp]: "wf_J_mdecl P C (M,NonStatic,Ts,T,pns,body) \ (length Ts = length pns \ distinct pns \ \sub_RI body \ this \ set pns \ (\T'. P,[this\Class C,pns[\]Ts] \ body :: T' \ P \ T' \ T) \ \ body \{this} \ set pns\)" (*<*)by(simp add:wf_J_mdecl_def)(*>*) lemma wf_J_mdecl_Static[simp]: "wf_J_mdecl P C (M,Static,Ts,T,pns,body) \ (length Ts = length pns \ distinct pns \ \sub_RI body \ (\T'. P,[pns[\]Ts] \ body :: T' \ P \ T' \ T) \ \ body \set pns\)" (*<*)by(simp add:wf_J_mdecl_def)(*>*) abbreviation wf_J_prog :: "J_prog \ bool" where "wf_J_prog == wf_prog wf_J_mdecl" lemma wf_J_prog_wf_J_mdecl: "\ wf_J_prog P; (C, D, fds, mths) \ set P; jmdcl \ set mths \ \ wf_J_mdecl P C jmdcl" -(*<*) -apply (simp add: wf_prog_def) -apply (simp add: wf_cdecl_def) -apply (erule conjE)+ -apply (drule bspec, assumption) -apply simp -apply (erule conjE)+ -apply (drule bspec, assumption) -apply (simp add: wf_mdecl_def split_beta) -done -(*>*) +(*<*)by(fastforce simp: wf_prog_def wf_cdecl_def wf_mdecl_def)(*>*) lemma wf_mdecl_wwf_mdecl: "wf_J_mdecl P C Md \ wwf_J_mdecl P C Md" (*<*) -apply(clarsimp simp:wwf_J_mdecl_def) apply(rename_tac M b Ts T pns body) -apply (case_tac b) - by (fastforce dest!:WT_fv)+ +proof - + obtain M b Ts T pns body where "Md = (M, b, Ts, T, pns, body)" by(cases Md) simp + then show "wf_J_mdecl P C Md \ wwf_J_mdecl P C Md" + by(case_tac b) (fastforce simp:wwf_J_mdecl_def dest!:WT_fv)+ +qed (*>*) lemma wf_prog_wwf_prog: "wf_J_prog P \ wwf_J_prog P" (*<*) -apply(simp add:wf_prog_def wf_cdecl_def wf_mdecl_def) -apply(fast intro:wf_mdecl_wwf_mdecl) -done +by (simp add:wf_prog_def wf_cdecl_def wf_mdecl_def) + (fast intro:wf_mdecl_wwf_mdecl) (*>*) end diff --git a/thys/JinjaDCI/J/Progress.thy b/thys/JinjaDCI/J/Progress.thy --- a/thys/JinjaDCI/J/Progress.thy +++ b/thys/JinjaDCI/J/Progress.thy @@ -1,887 +1,880 @@ (* Title: JinjaDCI/J/Progress.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory J/Progress.thy by Tobias Nipkow *) section \ Progress of Small Step Semantics \ theory Progress imports WellTypeRT DefAss "../Common/Conform" EConform begin lemma final_addrE: "\ P,E,h,sh \ e : Class C; final e; \a. e = addr a \ R; \a. e = Throw a \ R \ \ R" (*<*)by(auto simp:final_def)(*>*) lemma finalRefE: "\ P,E,h,sh \ e : T; is_refT T; final e; e = null \ R; \a C. \ e = addr a; T = Class C \ \ R; \a. e = Throw a \ R \ \ R" (*<*)by(auto simp:final_def is_refT_def)(*>*) text\ Derivation of new induction scheme for well typing: \ inductive WTrt' :: "[J_prog,heap,sheap,env,expr,ty] \ bool" and WTrts' :: "[J_prog,heap,sheap,env,expr list, ty list] \ bool" and WTrt2' :: "[J_prog,env,heap,sheap,expr,ty] \ bool" ("_,_,_,_ \ _ :'' _" [51,51,51,51]50) and WTrts2' :: "[J_prog,env,heap,sheap,expr list, ty list] \ bool" ("_,_,_,_ \ _ [:''] _" [51,51,51,51]50) for P :: J_prog and h :: heap and sh :: sheap where "P,E,h,sh \ e :' T \ WTrt' P h sh E e T" | "P,E,h,sh \ es [:'] Ts \ WTrts' P h sh E es Ts" | "is_class P C \ P,E,h,sh \ new C :' Class C" | "\ P,E,h,sh \ e :' T; is_refT T; is_class P C \ \ P,E,h,sh \ Cast C e :' Class C" | "typeof\<^bsub>h\<^esub> v = Some T \ P,E,h,sh \ Val v :' T" | "E v = Some T \ P,E,h,sh \ Var v :' T" | "\ P,E,h,sh \ e\<^sub>1 :' T\<^sub>1; P,E,h,sh \ e\<^sub>2 :' T\<^sub>2 \ \ P,E,h,sh \ e\<^sub>1 \Eq\ e\<^sub>2 :' Boolean" | "\ P,E,h,sh \ e\<^sub>1 :' Integer; P,E,h,sh \ e\<^sub>2 :' Integer \ \ P,E,h,sh \ e\<^sub>1 \Add\ e\<^sub>2 :' Integer" | "\ P,E,h,sh \ Var V :' T; P,E,h,sh \ e :' T'; P \ T' \ T \ \ P,E,h,sh \ V:=e :' Void" | "\ P,E,h,sh \ e :' Class C; P \ C has F,NonStatic:T in D \ \ P,E,h,sh \ e\F{D} :' T" | "P,E,h,sh \ e :' NT \ P,E,h,sh \ e\F{D} :' T" | "\ P \ C has F,Static:T in D \ \ P,E,h,sh \ C\\<^sub>sF{D} :' T" | "\ P,E,h,sh \ e\<^sub>1 :' Class C; P \ C has F,NonStatic:T in D; P,E,h,sh \ e\<^sub>2 :' T\<^sub>2; P \ T\<^sub>2 \ T \ \ P,E,h,sh \ e\<^sub>1\F{D}:=e\<^sub>2 :' Void" | "\ P,E,h,sh \ e\<^sub>1:'NT; P,E,h,sh \ e\<^sub>2 :' T\<^sub>2 \ \ P,E,h,sh \ e\<^sub>1\F{D}:=e\<^sub>2 :' Void" | "\ P \ C has F,Static:T in D; P,E,h,sh \ e\<^sub>2 :' T\<^sub>2; P \ T\<^sub>2 \ T \ \ P,E,h,sh \ C\\<^sub>sF{D}:=e\<^sub>2 :' Void" | "\ P,E,h,sh \ e :' Class C; P \ C sees M,NonStatic:Ts \ T = (pns,body) in D; P,E,h,sh \ es [:'] Ts'; P \ Ts' [\] Ts \ \ P,E,h,sh \ e\M(es) :' T" | "\ P,E,h,sh \ e :' NT; P,E,h,sh \ es [:'] Ts \ \ P,E,h,sh \ e\M(es) :' T" | "\ P \ C sees M,Static:Ts \ T = (pns,body) in D; P,E,h,sh \ es [:'] Ts'; P \ Ts' [\] Ts; M = clinit \ sh D = \(sfs,Processing)\ \ es = map Val vs \ \ P,E,h,sh \ C\\<^sub>sM(es) :' T" | "P,E,h,sh \ [] [:'] []" | "\ P,E,h,sh \ e :' T; P,E,h,sh \ es [:'] Ts \ \ P,E,h,sh \ e#es [:'] T#Ts" | "\ typeof\<^bsub>h\<^esub> v = Some T\<^sub>1; P \ T\<^sub>1 \ T; P,E(V\T),h,sh \ e\<^sub>2 :' T\<^sub>2 \ \ P,E,h,sh \ {V:T := Val v; e\<^sub>2} :' T\<^sub>2" | "\ P,E(V\T),h,sh \ e :' T'; \ assigned V e \ \ P,E,h,sh \ {V:T; e} :' T'" | "\ P,E,h,sh \ e\<^sub>1:' T\<^sub>1; P,E,h,sh \ e\<^sub>2:'T\<^sub>2 \ \ P,E,h,sh \ e\<^sub>1;;e\<^sub>2 :' T\<^sub>2" | "\ P,E,h,sh \ e :' Boolean; P,E,h,sh \ e\<^sub>1:' T\<^sub>1; P,E,h,sh \ e\<^sub>2:' T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1; P \ T\<^sub>1 \ T\<^sub>2 \ T = T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ T = T\<^sub>1 \ \ P,E,h,sh \ if (e) e\<^sub>1 else e\<^sub>2 :' T" | "\ P,E,h,sh \ e :' Boolean; P,E,h,sh \ c:' T \ \ P,E,h,sh \ while(e) c :' Void" | "\ P,E,h,sh \ e :' T\<^sub>r; is_refT T\<^sub>r \ \ P,E,h,sh \ throw e :' T" | "\ P,E,h,sh \ e\<^sub>1 :' T\<^sub>1; P,E(V \ Class C),h,sh \ e\<^sub>2 :' T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ \ P,E,h,sh \ try e\<^sub>1 catch(C V) e\<^sub>2 :' T\<^sub>2" | "\ P,E,h,sh \ e :' T; \C' \ set (C#Cs). is_class P C'; \sub_RI e; \C' \ set (tl Cs). \sfs. sh C' = \(sfs,Processing)\; b \ (\C' \ set Cs. \sfs. sh C' = \(sfs,Processing)\); distinct Cs; supercls_lst P Cs \ \ P,E,h,sh \ INIT C (Cs, b) \ e :' T" | "\ P,E,h,sh \ e :' T; P,E,h,sh \ e' :' T'; \C' \ set (C#Cs). is_class P C'; \sub_RI e'; \C' \ set (C#Cs). not_init C' e; \C' \ set Cs. \sfs. sh C' = \(sfs,Processing)\; \sfs. sh C = \(sfs, Processing)\ \ (sh C = \(sfs, Error)\ \ e = THROW NoClassDefFoundError); distinct (C#Cs); supercls_lst P (C#Cs) \ \ P,E,h,sh \ RI(C, e);Cs \ e' :' T'" (*<*) lemmas WTrt'_induct = WTrt'_WTrts'.induct [split_format (complete)] and WTrt'_inducts = WTrt'_WTrts'.inducts [split_format (complete)] inductive_cases WTrt'_elim_cases[elim!]: "P,E,h,sh \ V :=e :' T" (*>*) lemma [iff]: "P,E,h,sh \ e\<^sub>1;;e\<^sub>2 :' T\<^sub>2 = (\T\<^sub>1. P,E,h,sh \ e\<^sub>1:' T\<^sub>1 \ P,E,h,sh \ e\<^sub>2:' T\<^sub>2)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros) -done -(*>*) +(*<*)by(rule iffI) (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros)(*>*) lemma [iff]: "P,E,h,sh \ Val v :' T = (typeof\<^bsub>h\<^esub> v = Some T)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros) -done -(*>*) +(*<*)by(rule iffI) (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros)(*>*) lemma [iff]: "P,E,h,sh \ Var v :' T = (E v = Some T)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros) -done -(*>*) +(*<*)by(rule iffI) (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros)(*>*) lemma wt_wt': "P,E,h,sh \ e : T \ P,E,h,sh \ e :' T" and wts_wts': "P,E,h,sh \ es [:] Ts \ P,E,h,sh \ es [:'] Ts" (*<*) -apply (induct rule:WTrt_inducts) -prefer 17 -apply(case_tac "assigned V e") -apply(clarsimp simp add:fun_upd_same assigned_def simp del:fun_upd_apply) -apply(erule (2) WTrt'_WTrts'.intros) -apply(erule (1) WTrt'_WTrts'.intros) -apply(blast intro:WTrt'_WTrts'.intros)+ -done +proof(induct rule:WTrt_inducts) + case (WTrtBlock E V T e T') + then show ?case + proof(cases "assigned V e") + case True then show ?thesis using WTrtBlock.hyps(2) + by(clarsimp simp add:fun_upd_same assigned_def WTrt'_WTrts'.intros + simp del:fun_upd_apply) + next + case False then show ?thesis + by (simp add: WTrtBlock.hyps(2) WTrt'_WTrts'.intros) + qed +qed (blast intro:WTrt'_WTrts'.intros)+ (*>*) lemma wt'_wt: "P,E,h,sh \ e :' T \ P,E,h,sh \ e : T" and wts'_wts: "P,E,h,sh \ es [:'] Ts \ P,E,h,sh \ es [:] Ts" (*<*) -apply (induct rule:WTrt'_inducts) -prefer 19 -apply(rule WTrt_WTrts.intros) -apply(rule WTrt_WTrts.intros) -apply(rule WTrt_WTrts.intros) -apply simp -apply(erule (2) WTrt_WTrts.intros) -apply(blast intro:WTrt_WTrts.intros)+ -done +proof(induct rule:WTrt'_inducts) + case Block: (19 v T\<^sub>1 T E V e\<^sub>2 T\<^sub>2) + let ?E = "E(V \ T)" + have "P,?E,h,sh \ Val v : T\<^sub>1" using Block.hyps(1) by simp + moreover have "P \ T\<^sub>1 \ T" by(rule Block.hyps(2)) + ultimately have "P,?E,h,sh \ V:=Val v : Void" using WTrtLAss by simp + moreover have "P,?E,h,sh \ e\<^sub>2 : T\<^sub>2" by(rule Block.hyps(4)) + ultimately have "P,?E,h,sh \ V:=Val v;; e\<^sub>2 : T\<^sub>2" by blast + then show ?case by simp +qed (blast intro:WTrt_WTrts.intros)+ (*>*) corollary wt'_iff_wt: "(P,E,h,sh \ e :' T) = (P,E,h,sh \ e : T)" (*<*)by(blast intro:wt_wt' wt'_wt)(*>*) corollary wts'_iff_wts: "(P,E,h,sh \ es [:'] Ts) = (P,E,h,sh \ es [:] Ts)" (*<*)by(blast intro:wts_wts' wts'_wts)(*>*) (*<*) lemmas WTrt_inducts2 = WTrt'_inducts [unfolded wt'_iff_wt wts'_iff_wts, case_names WTrtNew WTrtCast WTrtVal WTrtVar WTrtBinOpEq WTrtBinOpAdd WTrtLAss WTrtFAcc WTrtFAccNT WTrtSFAcc WTrtFAss WTrtFAssNT WTrtSFAss WTrtCall WTrtCallNT WTrtSCall WTrtNil WTrtCons WTrtInitBlock WTrtBlock WTrtSeq WTrtCond WTrtWhile WTrtThrow WTrtTry WTrtInit WTrtRI, consumes 1] (*>*) theorem assumes wf: "wwf_J_prog P" and hconf: "P \ h \" and shconf: "P,h \\<^sub>s sh \" shows progress: "P,E,h,sh \ e : T \ (\l. \ \ e \dom l\; P,sh \\<^sub>b (e,b) \; \ final e \ \ \e' s' b'. P \ \e,(h,l,sh),b\ \ \e',s',b'\)" and "P,E,h,sh \ es [:] Ts \ (\l. \ \s es \dom l\; P,sh \\<^sub>b (es,b) \; \ finals es \ \ \es' s' b'. P \ \es,(h,l,sh),b\ [\] \es',s',b'\)" (*<*) proof (induct rule:WTrt_inducts2) case (WTrtNew C) show ?case proof (cases b) case True then show ?thesis proof cases assume "\a. h a = None" with assms WTrtNew True show ?thesis by (fastforce del:exE intro!:RedNew simp add:new_Addr_def elim!:wf_Fields_Ex[THEN exE]) next assume "\(\a. h a = None)" with assms WTrtNew True show ?thesis by(fastforce intro:RedNewFail simp:new_Addr_def) qed next case False then show ?thesis proof cases assume "\sfs. sh C = Some (sfs, Done)" with assms WTrtNew False show ?thesis by(fastforce intro:NewInitDoneRed simp:new_Addr_def) next assume "\sfs. sh C = Some (sfs, Done)" with assms WTrtNew False show ?thesis by(fastforce intro:NewInitRed simp:new_Addr_def) qed qed next case (WTrtCast E e T C) have wte: "P,E,h,sh \ e : T" and ref: "is_refT T" and IH: "\l. \\ e \dom l\; P,sh \\<^sub>b (e,b) \; \ final e\ \ \e' s' b'. P \ \e,(h,l,sh),b\ \ \e',s',b'\" and D: "\ (Cast C e) \dom l\" and castconf: "P,sh \\<^sub>b (Cast C e,b) \" by fact+ from D have De: "\ e \dom l\" by auto have bconf: "P,sh \\<^sub>b (e,b) \" using castconf bconf_Cast by fast show ?case proof cases assume "final e" with wte ref show ?thesis proof (rule finalRefE) assume "e = null" thus ?case by(fastforce intro:RedCastNull) next fix D a assume A: "T = Class D" "e = addr a" show ?thesis proof cases assume "P \ D \\<^sup>* C" thus ?thesis using A wte by(fastforce intro:RedCast) next assume "\ P \ D \\<^sup>* C" thus ?thesis using A wte by(fastforce elim!:RedCastFail) qed next fix a assume "e = Throw a" thus ?thesis by(blast intro!:red_reds.CastThrow) qed next assume nf: "\ final e" from IH[OF De bconf nf] show ?thesis by (blast intro:CastRed) qed next case WTrtVal thus ?case by(simp add:final_def) next case WTrtVar thus ?case by(fastforce intro:RedVar simp:hyper_isin_def) next case (WTrtBinOpEq E e1 T1 e2 T2) show ?case proof cases assume "final e1" thus ?thesis proof (rule finalE) fix v1 assume eV[simp]: "e1 = Val v1" show ?thesis proof cases assume "final e2" thus ?thesis proof (rule finalE) fix v2 assume "e2 = Val v2" thus ?thesis using WTrtBinOpEq by(fastforce intro:RedBinOp) next fix a assume "e2 = Throw a" thus ?thesis using eV by(blast intro:red_reds.BinOpThrow2) qed next assume nf: "\ final e2" then have "P,sh \\<^sub>b (e2,b) \" using WTrtBinOpEq.prems(2) by(simp add:bconf_BinOp) with WTrtBinOpEq nf show ?thesis by simp (fast intro!:BinOpRed2) qed next fix a assume "e1 = Throw a" thus ?thesis by (fast intro:red_reds.BinOpThrow1) qed next assume nf: "\ final e1" then have e1: "val_of e1 = None" proof(cases e1)qed(auto) then have "P,sh \\<^sub>b (e1,b) \" using WTrtBinOpEq.prems(2) by(simp add:bconf_BinOp) with WTrtBinOpEq nf e1 show ?thesis by simp (fast intro:BinOpRed1) qed next case (WTrtBinOpAdd E e1 e2) show ?case proof cases assume "final e1" thus ?thesis proof (rule finalE) fix v1 assume eV[simp]: "e1 = Val v1" show ?thesis proof cases assume "final e2" thus ?thesis proof (rule finalE) fix v2 assume eV2:"e2 = Val v2" then obtain i1 i2 where "v1 = Intg i1 \ v2 = Intg i2" using WTrtBinOpAdd by clarsimp thus ?thesis using WTrtBinOpAdd eV eV2 by(fastforce intro:RedBinOp) next fix a assume "e2 = Throw a" thus ?thesis using eV by(blast intro:red_reds.BinOpThrow2) qed next assume nf:"\ final e2" then have "P,sh \\<^sub>b (e2,b) \" using WTrtBinOpAdd.prems(2) by(simp add:bconf_BinOp) with WTrtBinOpAdd nf show ?thesis by simp (fast intro!:BinOpRed2) qed next fix a assume "e1 = Throw a" thus ?thesis by(fast intro:red_reds.BinOpThrow1) qed next assume nf: "\ final e1" then have e1: "val_of e1 = None" proof(cases e1)qed(auto) then have "P,sh \\<^sub>b (e1,b) \" using WTrtBinOpAdd.prems(2) by(simp add:bconf_BinOp) with WTrtBinOpAdd nf e1 show ?thesis by simp (fast intro:BinOpRed1) qed next case (WTrtLAss E V T e T') then have bconf: "P,sh \\<^sub>b (e,b) \" using bconf_LAss by fast show ?case proof cases assume "final e" with WTrtLAss show ?thesis by(fastforce simp:final_def intro: red_reds.RedLAss red_reds.LAssThrow) next assume "\ final e" with WTrtLAss bconf show ?thesis by simp (fast intro:LAssRed) qed next case (WTrtFAcc E e C F T D) then have bconf: "P,sh \\<^sub>b (e,b) \" using bconf_FAcc by fast have wte: "P,E,h,sh \ e : Class C" and field: "P \ C has F,NonStatic:T in D" by fact+ show ?case proof cases assume "final e" with wte show ?thesis proof (rule final_addrE) fix a assume e: "e = addr a" with wte obtain fs where hp: "h a = Some(C,fs)" by auto with hconf have "P,h \ (C,fs) \" using hconf_def by fastforce then obtain v where "fs(F,D) = Some v" using field by(fastforce dest:has_fields_fun simp:oconf_def has_field_def) with hp e show ?thesis by (meson field red_reds.RedFAcc) next fix a assume "e = Throw a" thus ?thesis by(fastforce intro:red_reds.FAccThrow) qed next assume "\ final e" with WTrtFAcc bconf show ?thesis by(fastforce intro!:FAccRed) qed next case (WTrtFAccNT E e F D T) then have bconf: "P,sh \\<^sub>b (e,b) \" using bconf_FAcc by fast show ?case proof cases assume "final e" \ \@{term e} is @{term null} or @{term throw}\ with WTrtFAccNT show ?thesis by(fastforce simp:final_def intro: red_reds.RedFAccNull red_reds.FAccThrow) next assume "\ final e" \ \@{term e} reduces by IH\ with WTrtFAccNT bconf show ?thesis by simp (fast intro:FAccRed) qed next case (WTrtSFAcc C F T D E) then show ?case proof (cases b) case True then obtain sfs i where shD: "sh D = \(sfs,i)\" using bconf_def[of P sh "C\\<^sub>sF{D}" b] WTrtSFAcc.prems(2) initPD_def by auto with shconf have "P,h,D \\<^sub>s sfs \" using shconf_def[of P h sh] by auto then obtain v where sfsF: "sfs F = Some v" using WTrtSFAcc.hyps by(unfold soconf_def) (auto dest:has_field_idemp) then show ?thesis using WTrtSFAcc.hyps shD sfsF True by(fastforce elim:RedSFAcc) next case False with assms WTrtSFAcc show ?thesis by(metis (full_types) SFAccInitDoneRed SFAccInitRed) qed next case (WTrtFAss E e1 C F T D e2 T2) have wte1: "P,E,h,sh \ e1 : Class C" and field: "P \ C has F,NonStatic:T in D" by fact+ show ?case proof cases assume "final e1" with wte1 show ?thesis proof (rule final_addrE) fix a assume e1: "e1 = addr a" show ?thesis proof cases assume "final e2" thus ?thesis proof (rule finalE) fix v assume "e2 = Val v" thus ?thesis using e1 wte1 by(fastforce intro: RedFAss[OF field]) next fix a assume "e2 = Throw a" thus ?thesis using e1 by(fastforce intro:red_reds.FAssThrow2) qed next assume nf: "\ final e2" then have "P,sh \\<^sub>b (e2,b) \" using WTrtFAss.prems(2) e1 by(simp add:bconf_FAss) with WTrtFAss e1 nf show ?thesis by simp (fast intro!:FAssRed2) qed next fix a assume "e1 = Throw a" thus ?thesis by(fastforce intro:red_reds.FAssThrow1) qed next assume nf: "\ final e1" then have e1: "val_of e1 = None" proof(cases e1)qed(auto) then have "P,sh \\<^sub>b (e1,b) \" using WTrtFAss.prems(2) by(simp add:bconf_FAss) with WTrtFAss nf e1 show ?thesis by simp (blast intro!:FAssRed1) qed next case (WTrtFAssNT E e\<^sub>1 e\<^sub>2 T\<^sub>2 F D) show ?case proof cases assume e1: "final e\<^sub>1" \ \@{term e\<^sub>1} is @{term null} or @{term throw}\ show ?thesis proof cases assume "final e\<^sub>2" \ \@{term e\<^sub>2} is @{term Val} or @{term throw}\ with WTrtFAssNT e1 show ?thesis by(fastforce simp:final_def intro: red_reds.RedFAssNull red_reds.FAssThrow1 red_reds.FAssThrow2) next assume nf: "\ final e\<^sub>2" \ \@{term e\<^sub>2} reduces by IH\ show ?thesis proof (rule finalE[OF e1]) fix v assume ev: "e\<^sub>1 = Val v" then have "P,sh \\<^sub>b (e\<^sub>2,b) \" using WTrtFAssNT.prems(2) nf by(simp add:bconf_FAss) with WTrtFAssNT ev nf show ?thesis by auto (meson red_reds.FAssRed2) next fix a assume et: "e\<^sub>1 = Throw a" then have "P,sh \\<^sub>b (e\<^sub>1,b) \" using WTrtFAssNT.prems(2) nf by(simp add:bconf_FAss) with WTrtFAssNT et nf show ?thesis by(fastforce intro: red_reds.FAssThrow1) qed qed next assume nf: "\ final e\<^sub>1" \ \@{term e\<^sub>1} reduces by IH\ then have e1: "val_of e\<^sub>1 = None" proof(cases e\<^sub>1)qed(auto) then have "P,sh \\<^sub>b (e\<^sub>1,b) \" using WTrtFAssNT.prems(2) by(simp add:bconf_FAss) with WTrtFAssNT nf e1 show ?thesis by simp (blast intro!:FAssRed1) qed next case (WTrtSFAss C F T D E e2 T\<^sub>2) have field: "P \ C has F,Static:T in D" by fact+ show ?case proof cases assume "final e2" thus ?thesis proof (rule finalE) fix v assume ev: "e2 = Val v" then show ?case proof (cases b) case True then obtain sfs i where shD: "sh D = \(sfs,i)\" using bconf_def[of P _ "C\\<^sub>sF{D} := e2"] WTrtSFAss.prems(2) initPD_def ev by auto with shconf have "P,h,D \\<^sub>s sfs \" using shconf_def[of P] by auto then obtain v where sfsF: "sfs F = Some v" using field by(unfold soconf_def) (auto dest:has_field_idemp) then show ?thesis using WTrtSFAss.hyps shD sfsF True ev by(fastforce elim:RedSFAss) next case False with assms WTrtSFAss ev show ?thesis by(metis (full_types) SFAssInitDoneRed SFAssInitRed) qed next fix a assume "e2 = Throw a" thus ?thesis by(fastforce intro:red_reds.SFAssThrow) qed next assume nf: "\ final e2" then have "val_of e2 = None" using final_def val_of_spec by fastforce then have "P,sh \\<^sub>b (e2,b) \" using WTrtSFAss.prems(2) by(simp add:bconf_SFAss) with WTrtSFAss nf show ?thesis by simp (fast intro!:SFAssRed) qed next case (WTrtCall E e C M Ts T pns body D es Ts') have 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 [:] Ts'"and sub: "P \ Ts' [\] Ts" and IHes: "\l. \\s es \dom l\; P,sh \\<^sub>b (es,b) \; \ finals es\ \ \es' s' b'. P \ \es,(h,l,sh),b\ [\] \es',s',b'\" and D: "\ (e\M(es)) \dom l\" by fact+ show ?case proof cases assume "final e" with wte show ?thesis proof (rule final_addrE) fix a assume e_addr: "e = addr a" show ?thesis proof cases assume es: "\vs. es = map Val vs" from wte e_addr obtain fs where ha: "h a = Some(C,fs)" by auto show ?thesis using e_addr ha "method" WTrts_same_length[OF wtes] sub es sees_wf_mdecl[OF wf "method"] by(fastforce intro!: RedCall simp:list_all2_iff wf_mdecl_def) next assume "\(\vs. es = map Val vs)" hence not_all_Val: "\(\e \ set es. \v. e = Val v)" by(simp add:ex_map_conv) let ?ves = "takeWhile (\e. \v. e = Val v) es" let ?rest = "dropWhile (\e. \v. e = Val v) es" let ?ex = "hd ?rest" let ?rst = "tl ?rest" from not_all_Val have nonempty: "?rest \ []" by auto hence es: "es = ?ves @ ?ex # ?rst" by simp have "\e \ set ?ves. \v. e = Val v" by(fastforce dest:set_takeWhileD) then obtain vs where ves: "?ves = map Val vs" using ex_map_conv by blast show ?thesis proof cases assume "final ?ex" moreover from nonempty have "\(\v. ?ex = Val v)" by(auto simp:neq_Nil_conv simp del:dropWhile_eq_Nil_conv) (simp add:dropWhile_eq_Cons_conv) ultimately obtain b where ex_Throw: "?ex = Throw b" by(fast elim!:finalE) show ?thesis using e_addr es ex_Throw ves by(fastforce intro:CallThrowParams) next assume not_fin: "\ final ?ex" have "finals es = finals(?ves @ ?ex # ?rst)" using es by(rule arg_cong) also have "\ = finals(?ex # ?rst)" using ves by simp finally have "finals es = finals(?ex # ?rst)" . hence fes: "\ finals es" using not_finals_ConsI[OF not_fin] by blast have "P,sh \\<^sub>b (es,b) \" using bconf_Call WTrtCall.prems(2) by (metis e_addr option.simps(5) val_of.simps(1)) thus ?thesis using fes e_addr D IHes by(fastforce intro!:CallParams) qed qed next fix a assume "e = Throw a" with WTrtCall.prems show ?thesis by(fast intro!:CallThrowObj) qed next assume nf: "\ final e" then have e1: "val_of e = None" proof(cases e)qed(auto) then have "P,sh \\<^sub>b (e,b) \" using WTrtCall.prems(2) by(simp add:bconf_Call) with WTrtCall nf e1 show ?thesis by simp (blast intro!:CallObj) qed next case (WTrtCallNT E e es Ts M T) show ?case proof cases assume "final e" moreover { fix v assume e: "e = Val v" hence "e = null" using WTrtCallNT by simp have ?case proof cases assume "finals es" moreover { fix vs assume "es = map Val vs" with WTrtCallNT e have ?thesis by(fastforce intro: RedCallNull) } moreover { fix vs a es' assume "es = map Val vs @ Throw a # es'" with WTrtCallNT e have ?thesis by(fastforce intro: CallThrowParams) } ultimately show ?thesis by(fastforce simp:finals_def) next assume nf: "\ finals es" \ \@{term es} reduces by IH\ have "P,sh \\<^sub>b (es,b) \" using WTrtCallNT.prems(2) e by (simp add: bconf_Call) with WTrtCallNT e nf show ?thesis by(fastforce intro: CallParams) qed } moreover { fix a assume "e = Throw a" with WTrtCallNT have ?case by(fastforce intro: CallThrowObj) } ultimately show ?thesis by(fastforce simp:final_def) next assume nf: "\ final e" \ \@{term e} reduces by IH\ then have "val_of e = None" proof(cases e)qed(auto) then have "P,sh \\<^sub>b (e,b) \" using WTrtCallNT.prems(2) by(simp add:bconf_Call) with WTrtCallNT nf show ?thesis by (fastforce intro:CallObj) qed next case (WTrtSCall C M Ts T pns body D E es Ts' sfs vs) have "method": "P \ C sees M,Static:Ts\T = (pns,body) in D" and wtes: "P,E,h,sh \ es [:] Ts'"and sub: "P \ Ts' [\] Ts" and IHes: "\l. \\s es \dom l\; P,sh \\<^sub>b (es,b) \; \ finals es\ \ \es' s' b'. P \ \es,(h,l,sh),b\ [\] \es',s',b'\" and clinit: "M = clinit \ sh D = \(sfs, Processing)\ \ es = map Val vs" and D: "\ (C\\<^sub>sM(es)) \dom l\" by fact+ show ?case proof cases assume es: "\vs. es = map Val vs" show ?thesis proof (cases b) case True then show ?thesis using "method" WTrts_same_length[OF wtes] sub es sees_wf_mdecl[OF wf "method"] True by(fastforce intro!: RedSCall simp:list_all2_iff wf_mdecl_def) next case False show ?thesis using "method" clinit WTrts_same_length[OF wtes] sub es False by (metis (full_types) red_reds.SCallInitDoneRed red_reds.SCallInitRed) qed next assume nmap: "\(\vs. es = map Val vs)" hence not_all_Val: "\(\e \ set es. \v. e = Val v)" by(simp add:ex_map_conv) let ?ves = "takeWhile (\e. \v. e = Val v) es" let ?rest = "dropWhile (\e. \v. e = Val v) es" let ?ex = "hd ?rest" let ?rst = "tl ?rest" from not_all_Val have nonempty: "?rest \ []" by auto hence es: "es = ?ves @ ?ex # ?rst" by simp have "\e \ set ?ves. \v. e = Val v" by(fastforce dest:set_takeWhileD) then obtain vs where ves: "?ves = map Val vs" using ex_map_conv by blast show ?thesis proof cases assume "final ?ex" moreover from nonempty have "\(\v. ?ex = Val v)" by(auto simp:neq_Nil_conv simp del:dropWhile_eq_Nil_conv) (simp add:dropWhile_eq_Cons_conv) ultimately obtain b where ex_Throw: "?ex = Throw b" by(fast elim!:finalE) show ?thesis using es ex_Throw ves by(fastforce intro:SCallThrowParams) next assume not_fin: "\ final ?ex" have "finals es = finals(?ves @ ?ex # ?rst)" using es by(rule arg_cong) also have "\ = finals(?ex # ?rst)" using ves by simp finally have "finals es = finals(?ex # ?rst)" . hence fes: "\ finals es" using not_finals_ConsI[OF not_fin] by blast have "P,sh \\<^sub>b (es,b) \" by (meson WTrtSCall.prems(2) nmap bconf_SCall map_vals_of_spec not_None_eq) thus ?thesis using fes D IHes by(fastforce intro!:SCallParams) qed qed next case WTrtNil thus ?case by simp next case (WTrtCons E e T es Ts) have IHe: "\l. \\ e \dom l\; P,sh \\<^sub>b (e,b) \; \ final e\ \ \e' s' b'. P \ \e,(h,l,sh),b\ \ \e',s',b'\" and IHes: "\l. \\s es \dom l\; P,sh \\<^sub>b (es,b) \; \ finals es\ \ \es' s' b'. P \ \es,(h,l,sh),b\ [\] \es',s',b'\" and D: "\s (e#es) \dom l\" and not_fins: "\ finals(e # es)" by fact+ have De: "\ e \dom l\" and Des: "\s es (\dom l\ \ \ e)" using D by auto show ?case proof cases assume "final e" thus ?thesis proof (rule finalE) fix v assume e: "e = Val v" hence Des': "\s es \dom l\" using De Des by auto have bconfs: "P,sh \\<^sub>b (es,b) \" using WTrtCons.prems(2) e by(simp add: bconf_Cons) have not_fins_tl: "\ finals es" using not_fins e by simp show ?thesis using e IHes[OF Des' bconfs not_fins_tl] by (blast intro!:ListRed2) next fix a assume "e = Throw a" hence False using not_fins by simp thus ?thesis .. qed next assume nf:"\ final e" then have "val_of e = None" proof(cases e)qed(auto) then have bconf: "P,sh \\<^sub>b (e,b) \" using WTrtCons.prems(2) by(simp add: bconf_Cons) with IHe[OF De bconf nf] show ?thesis by(fast intro!:ListRed1) qed next case (WTrtInitBlock v T\<^sub>1 T E V e\<^sub>2 T\<^sub>2) have IH2: "\l. \\ e\<^sub>2 \dom l\; P,sh \\<^sub>b (e\<^sub>2,b) \; \ final e\<^sub>2\ \ \e' s' b'. P \ \e\<^sub>2,(h,l,sh),b\ \ \e',s',b'\" and D: "\ {V:T := Val v; e\<^sub>2} \dom l\" by fact+ show ?case proof cases assume "final e\<^sub>2" then show ?thesis proof (rule finalE) fix v\<^sub>2 assume "e\<^sub>2 = Val v\<^sub>2" thus ?thesis by(fast intro:RedInitBlock) next fix a assume "e\<^sub>2 = Throw a" thus ?thesis by(fast intro:red_reds.InitBlockThrow) qed next assume not_fin2: "\ final e\<^sub>2" then have "val_of e\<^sub>2 = None" proof(cases e\<^sub>2)qed(auto) from D have D2: "\ e\<^sub>2 \dom(l(V\v))\" by (auto simp:hyperset_defs) have e2conf: "P,sh \\<^sub>b (e\<^sub>2,b) \" using WTrtInitBlock.prems(2) by(simp add: bconf_InitBlock) from IH2[OF D2 e2conf not_fin2] obtain h' l' sh' e' b' where red2: "P \ \e\<^sub>2,(h, l(V\v),sh),b\ \ \e',(h', l',sh'),b'\" by auto from red_lcl_incr[OF red2] have "V \ dom l'" by auto with red2 show ?thesis by(fastforce intro:InitBlockRed) qed next case (WTrtBlock E V T e T') have IH: "\l. \\ e \dom l\; P,sh \\<^sub>b (e,b) \; \ final e\ \ \e' s' b'. P \ \e,(h,l,sh),b\ \ \e',s',b'\" and unass: "\ assigned V e" and D: "\ {V:T; e} \dom l\" by fact+ show ?case proof cases assume "final e" thus ?thesis proof (rule finalE) fix v assume "e = Val v" thus ?thesis by(fast intro:RedBlock) next fix a assume "e = Throw a" thus ?thesis by(fast intro:red_reds.BlockThrow) qed next assume not_fin: "\ final e" then have "val_of e = None" proof(cases e)qed(auto) from D have De: "\ e \dom(l(V:=None))\" by(simp add:hyperset_defs) have bconf: "P,sh \\<^sub>b (e,b) \" using WTrtBlock by(simp add: bconf_Block) from IH[OF De bconf not_fin] obtain h' l' sh' e' b' where red: "P \ \e,(h,l(V:=None),sh),b\ \ \e',(h',l',sh'),b'\" by auto show ?thesis proof (cases "l' V") assume "l' V = None" with red unass show ?thesis by(blast intro: BlockRedNone) next fix v assume "l' V = Some v" with red unass show ?thesis by(blast intro: BlockRedSome) qed qed next case (WTrtSeq E e1 T1 e2 T2) show ?case proof cases assume "final e1" thus ?thesis by(fast elim:finalE intro:RedSeq red_reds.SeqThrow) next assume nf: "\ final e1" then have e1: "val_of e1 = None" proof(cases e1)qed(auto) then show ?thesis proof(cases "lass_val_of e1") case None then have "P,sh \\<^sub>b (e1,b) \" using WTrtSeq.prems(2) e1 by(simp add: bconf_Seq) with WTrtSeq nf e1 None show ?thesis by simp (blast intro:SeqRed) next case (Some p) obtain V v where "e1 = V:=Val v" using lass_val_of_spec[OF Some] by simp then show ?thesis using SeqRed[OF RedLAss] by blast qed qed next case (WTrtCond E e e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 T) have wt: "P,E,h,sh \ e : Boolean" by fact show ?case proof cases assume "final e" thus ?thesis proof (rule finalE) fix v assume val: "e = Val v" then obtain b where v: "v = Bool b" using wt by auto show ?thesis proof (cases b) case True with val v show ?thesis by(fastforce intro:RedCondT simp: prod_cases3) next case False with val v show ?thesis by(fastforce intro:RedCondF simp: prod_cases3) qed next fix a assume "e = Throw a" thus ?thesis by(fast intro:red_reds.CondThrow) qed next assume nf: "\ final e" then have "bool_of e = None" proof(cases e)qed(auto) then have "P,sh \\<^sub>b (e,b) \" using WTrtCond.prems(2) by(simp add: bconf_Cond) with WTrtCond nf show ?thesis by simp (blast intro:CondRed) qed next case WTrtWhile show ?case by(fast intro:RedWhile) next case (WTrtThrow E e T\<^sub>r T) show ?case proof cases assume "final e" \ \Then @{term e} must be @{term throw} or @{term null}\ with WTrtThrow show ?thesis by(fastforce simp:final_def is_refT_def intro:red_reds.ThrowThrow red_reds.RedThrowNull) next assume nf: "\ final e" \ \Then @{term e} must reduce\ then have "val_of e = None" proof(cases e)qed(auto) then have "P,sh \\<^sub>b (e,b) \" using WTrtThrow.prems(2) by(simp add: bconf_Throw) with WTrtThrow nf show ?thesis by simp (blast intro:ThrowRed) qed next case (WTrtTry E e1 T1 V C e2 T2) have wt1: "P,E,h,sh \ e1 : T1" by fact show ?case proof cases assume "final e1" thus ?thesis proof (rule finalE) fix v assume "e1 = Val v" thus ?thesis by(fast intro:RedTry) next fix a assume e1_Throw: "e1 = Throw a" with wt1 obtain D fs where ha: "h a = Some(D,fs)" by fastforce show ?thesis proof cases assume "P \ D \\<^sup>* C" with e1_Throw ha show ?thesis by(fastforce intro!:RedTryCatch) next assume "\ P \ D \\<^sup>* C" with e1_Throw ha show ?thesis by(fastforce intro!:RedTryFail) qed qed next assume nf: "\ final e1" then have "val_of e1 = None" proof(cases e1)qed(auto) then have "P,sh \\<^sub>b (e1,b) \" using WTrtTry.prems(2) by(simp add: bconf_Try) with WTrtTry nf show ?thesis by simp (fast intro:TryRed) qed next case (WTrtInit E e T\<^sub>r C Cs b) show ?case proof(cases Cs) case Nil then show ?thesis using WTrtInit by(fastforce intro!: RedInit) next case (Cons C' Cs') show ?thesis proof(cases b) case True then show ?thesis using Cons by(fastforce intro!: RedInitRInit) next case False show ?thesis proof(cases "sh C'") case None then show ?thesis using False Cons by(fastforce intro!: InitNoneRed) next case (Some sfsi) then obtain sfs i where sfsi:"sfsi = (sfs,i)" by(cases sfsi) show ?thesis proof(cases i) case Done then show ?thesis using False Some sfsi Cons by(fastforce intro!: RedInitDone) next case Processing then show ?thesis using False Some sfsi Cons by(fastforce intro!: RedInitProcessing) next case Error then show ?thesis using False Some sfsi Cons by(fastforce intro!: RedInitError) next case Prepared show ?thesis proof cases assume "C' = Object" then show ?thesis using False Some sfsi Prepared Cons by(fastforce intro: InitObjectRed) next assume "C' \ Object" then show ?thesis using False Some sfsi Prepared WTrtInit.hyps(3) Cons by(simp only: is_class_def)(fastforce intro!: InitNonObjectSuperRed) qed qed qed qed qed next case (WTrtRI E e T\<^sub>r e' T\<^sub>r' C Cs) obtain sfs i where shC: "sh C = \(sfs, i)\" using WTrtRI.hyps(9) by blast show ?case proof cases assume fin: "final e" then show ?thesis proof (rule finalE) fix v assume e: "e = Val v" then show ?thesis using shC e by(fast intro:RedRInit) next fix a assume eThrow: "e = Throw a" show ?thesis proof(cases Cs) case Nil then show ?thesis using eThrow shC by(fastforce intro!: RInitThrow) next case Cons then show ?thesis using eThrow shC by(fastforce intro!: RInitInitThrow) qed qed next assume nf: "\ final e" then have "val_of e = None" proof(cases e)qed(auto) then have "P,sh \\<^sub>b (e,b) \" using WTrtRI.prems(2) by(simp add: bconf_RI) with WTrtRI nf show ?thesis by simp (meson red_reds.RInitRed) qed qed (*>*) end diff --git a/thys/JinjaDCI/J/SmallStep.thy b/thys/JinjaDCI/J/SmallStep.thy --- a/thys/JinjaDCI/J/SmallStep.thy +++ b/thys/JinjaDCI/J/SmallStep.thy @@ -1,634 +1,646 @@ (* Title: JinjaDCI/J/SmallStep.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory J/SmallStep.thy by Tobias Nipkow *) section \ Small Step Semantics \ theory SmallStep imports Expr State WWellForm begin fun blocks :: "vname list * ty list * val list * expr \ expr" where "blocks(V#Vs, T#Ts, v#vs, e) = {V:T := Val v; blocks(Vs,Ts,vs,e)}" |"blocks([],[],[],e) = e" lemmas blocks_induct = blocks.induct[split_format (complete)] lemma [simp]: "\ size vs = size Vs; size Ts = size Vs \ \ fv(blocks(Vs,Ts,vs,e)) = fv e - set Vs" (*<*) by (induct rule:blocks_induct) auto (*>*) lemma sub_RI_blocks_body[iff]: "length vs = length pns \ length Ts = length pns \ sub_RI (blocks (pns, Ts, vs, body)) \ sub_RI body" proof(induct pns arbitrary: Ts vs) case Nil then show ?case by simp next case Cons then show ?case by(cases vs; cases Ts) auto qed definition assigned :: "'a \ 'a exp \ bool" where "assigned V e \ \v e'. e = (V := Val v;; e')" \ \ expression is okay to go the right side of @{text INIT} or @{text "RI \"} or to have indicator Boolean be True (in latter case, given that class is also verified initialized) \ fun icheck :: "'m prog \ cname \ 'a exp \ bool" where "icheck P C' (new C) = (C' = C)" | "icheck P D' (C\\<^sub>sF{D}) = ((D' = D) \ (\T. P \ C has F,Static:T in D))" | "icheck P D' (C\\<^sub>sF{D}:=(Val v)) = ((D' = D) \ (\T. P \ C has F,Static:T in D))" | "icheck P D (C\\<^sub>sM(es)) = ((\vs. es = map Val vs) \ (\Ts T m. P \ C sees M,Static:Ts\T = m in D))" | "icheck P _ _ = False" lemma nicheck_SFAss_nonVal: "val_of e\<^sub>2 = None \ \icheck P C' (C\\<^sub>sF{D} := (e\<^sub>2::'a exp))" by(rule notI, cases e\<^sub>2, auto) inductive_set red :: "J_prog \ ((expr \ state \ bool) \ (expr \ state \ bool)) set" and reds :: "J_prog \ ((expr list \ state \ bool) \ (expr list \ state \ bool)) set" and red' :: "J_prog \ expr \ state \ bool \ expr \ state \ bool \ bool" ("_ \ ((1\_,/_,/_\) \/ (1\_,/_,/_\))" [51,0,0,0,0,0,0] 81) and reds' :: "J_prog \ expr list \ state \ bool \ expr list \ state \ bool \ bool" ("_ \ ((1\_,/_,/_\) [\]/ (1\_,/_,/_\))" [51,0,0,0,0,0,0] 81) for P :: J_prog where "P \ \e,s,b\ \ \e',s',b'\ \ ((e,s,b), e',s',b') \ red P" | "P \ \es,s,b\ [\] \es',s',b'\ \ ((es,s,b), es',s',b') \ reds P" | RedNew: "\ new_Addr h = Some a; P \ C has_fields FDTs; h' = h(a\blank P C) \ \ P \ \new C, (h,l,sh), True\ \ \addr a, (h',l,sh), False\" | RedNewFail: "\ new_Addr h = None; is_class P C \ \ P \ \new C, (h,l,sh), True\ \ \THROW OutOfMemory, (h,l,sh), False\" | NewInitDoneRed: "sh C = Some (sfs, Done) \ P \ \new C, (h,l,sh), False\ \ \new C, (h,l,sh), True\" | NewInitRed: "\ \sfs. sh C = Some (sfs, Done); is_class P C \ \ P \ \new C,(h,l,sh),False\ \ \INIT C ([C],False) \ new C,(h,l,sh),False\" | CastRed: "P \ \e,s,b\ \ \e',s',b'\ \ P \ \Cast C e, s, b\ \ \Cast C e', s', b'\" | RedCastNull: "P \ \Cast C null, s, b\ \ \null,s,b\" | RedCast: "\ h a = Some(D,fs); P \ D \\<^sup>* C \ \ P \ \Cast C (addr a), (h,l,sh), b\ \ \addr a, (h,l,sh), b\" | RedCastFail: "\ h a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \ \Cast C (addr a), (h,l,sh), b\ \ \THROW ClassCast, (h,l,sh), b\" | BinOpRed1: "P \ \e,s,b\ \ \e',s',b'\ \ P \ \e \bop\ e\<^sub>2, s, b\ \ \e' \bop\ e\<^sub>2, s', b'\" | BinOpRed2: "P \ \e,s,b\ \ \e',s',b'\ \ P \ \(Val v\<^sub>1) \bop\ e, s, b\ \ \(Val v\<^sub>1) \bop\ e', s', b'\" | RedBinOp: "binop(bop,v\<^sub>1,v\<^sub>2) = Some v \ P \ \(Val v\<^sub>1) \bop\ (Val v\<^sub>2), s, b\ \ \Val v,s,b\" | RedVar: "l V = Some v \ P \ \Var V,(h,l,sh),b\ \ \Val v,(h,l,sh),b\" | LAssRed: "P \ \e,s,b\ \ \e',s',b'\ \ P \ \V:=e,s,b\ \ \V:=e',s',b'\" | RedLAss: "P \ \V:=(Val v), (h,l,sh), b\ \ \unit, (h,l(V\v),sh), b\" | FAccRed: "P \ \e,s,b\ \ \e',s',b'\ \ P \ \e\F{D}, s, b\ \ \e'\F{D}, s', b'\" | RedFAcc: "\ h a = Some(C,fs); fs(F,D) = Some v; P \ C has F,NonStatic:t in D \ \ P \ \(addr a)\F{D}, (h,l,sh), b\ \ \Val v,(h,l,sh),b\" | RedFAccNull: "P \ \null\F{D}, s, b\ \ \THROW NullPointer, s, b\" | RedFAccNone: "\ h a = Some(C,fs); \(\b t. P \ C has F,b:t in D) \ \ P \ \(addr a)\F{D},(h,l,sh),b\ \ \THROW NoSuchFieldError,(h,l,sh),b\" | RedFAccStatic: "\ h a = Some(C,fs); P \ C has F,Static:t in D \ \ P \ \(addr a)\F{D},(h,l,sh),b\ \ \THROW IncompatibleClassChangeError,(h,l,sh),b\" | RedSFAcc: "\ P \ C has F,Static:t in D; sh D = Some (sfs,i); sfs F = Some v \ \ P \ \C\\<^sub>sF{D},(h,l,sh),True\ \ \Val v,(h,l,sh),False\" | SFAccInitDoneRed: "\ P \ C has F,Static:t in D; sh D = Some (sfs,Done) \ \ P \ \C\\<^sub>sF{D},(h,l,sh),False\ \ \C\\<^sub>sF{D},(h,l,sh),True\" | SFAccInitRed: "\ P \ C has F,Static:t in D; \sfs. sh D = Some (sfs,Done) \ \ P \ \C\\<^sub>sF{D},(h,l,sh),False\ \ \INIT D ([D],False) \ C\\<^sub>sF{D},(h,l,sh),False\" | RedSFAccNone: "\(\b t. P \ C has F,b:t in D) \ P \ \C\\<^sub>sF{D},(h,l,sh),b\ \ \THROW NoSuchFieldError,(h,l,sh),False\" | RedSFAccNonStatic: "P \ C has F,NonStatic:t in D \ P \ \C\\<^sub>sF{D},(h,l,sh),b\ \ \THROW IncompatibleClassChangeError,(h,l,sh),False\" | FAssRed1: "P \ \e,s,b\ \ \e',s',b'\ \ P \ \e\F{D}:=e\<^sub>2, s, b\ \ \e'\F{D}:=e\<^sub>2, s', b'\" | FAssRed2: "P \ \e,s,b\ \ \e',s',b'\ \ P \ \Val v\F{D}:=e, s, b\ \ \Val v\F{D}:=e', s', b'\" | RedFAss: "\ P \ C has F,NonStatic:t in D; h a = Some(C,fs) \ \ P \ \(addr a)\F{D}:=(Val v), (h,l,sh), b\ \ \unit, (h(a \ (C,fs((F,D) \ v))),l,sh), b\" | RedFAssNull: "P \ \null\F{D}:=Val v, s, b\ \ \THROW NullPointer, s, b\" | RedFAssNone: "\ h a = Some(C,fs); \(\b t. P \ C has F,b:t in D) \ \ P \ \(addr a)\F{D}:=(Val v),(h,l,sh),b\ \ \THROW NoSuchFieldError,(h,l,sh),b\" | RedFAssStatic: "\ h a = Some(C,fs); P \ C has F,Static:t in D \ \ P \ \(addr a)\F{D}:=(Val v),(h,l,sh),b\ \ \THROW IncompatibleClassChangeError,(h,l,sh),b\" | SFAssRed: "P \ \e,s,b\ \ \e',s',b'\ \ P \ \C\\<^sub>sF{D}:=e, s, b\ \ \C\\<^sub>sF{D}:=e', s', b'\" | RedSFAss: "\ P \ C has F,Static:t in D; sh D = Some(sfs,i); sfs' = sfs(F\v); sh' = sh(D\(sfs',i)) \ \ P \ \C\\<^sub>sF{D}:=(Val v),(h,l,sh),True\ \ \unit,(h,l,sh'),False\" | SFAssInitDoneRed: "\ P \ C has F,Static:t in D; sh D = Some(sfs,Done) \ \ P \ \C\\<^sub>sF{D}:=(Val v),(h,l,sh),False\ \ \C\\<^sub>sF{D}:=(Val v),(h,l,sh),True\" | SFAssInitRed: "\ P \ C has F,Static:t in D; \sfs. sh D = Some(sfs,Done) \ \ P \ \C\\<^sub>sF{D}:=(Val v),(h,l,sh),False\ \ \INIT D ([D],False)\ C\\<^sub>sF{D}:=(Val v),(h,l,sh),False\" | RedSFAssNone: "\(\b t. P \ C has F,b:t in D) \ P \ \C\\<^sub>sF{D}:=(Val v),s,b\ \ \THROW NoSuchFieldError,s,False\" | RedSFAssNonStatic: "P \ C has F,NonStatic:t in D \ P \ \C\\<^sub>sF{D}:=(Val v),s,b\ \ \THROW IncompatibleClassChangeError,s,False\" | CallObj: "P \ \e,s,b\ \ \e',s',b'\ \ P \ \e\M(es),s,b\ \ \e'\M(es),s',b'\" | CallParams: "P \ \es,s,b\ [\] \es',s',b'\ \ P \ \(Val v)\M(es),s,b\ \ \(Val v)\M(es'),s',b'\" | RedCall: "\ h a = Some(C,fs); P \ C sees M,NonStatic:Ts\T = (pns,body) in D; size vs = size pns; size Ts = size pns \ \ P \ \(addr a)\M(map Val vs), (h,l,sh), b\ \ \blocks(this#pns, Class D#Ts, Addr a#vs, body), (h,l,sh), b\" | RedCallNull: "P \ \null\M(map Val vs),s,b\ \ \THROW NullPointer,s,b\" | RedCallNone: "\ h a = Some(C,fs); \(\b Ts T m D. P \ C sees M,b:Ts\T = m in D) \ \ P \ \(addr a)\M(map Val vs),(h,l,sh),b\ \ \THROW NoSuchMethodError,(h,l,sh),b\" | RedCallStatic: "\ h a = Some(C,fs); P \ C sees M,Static:Ts\T = m in D \ \ P \ \(addr a)\M(map Val vs),(h,l,sh),b\ \ \THROW IncompatibleClassChangeError,(h,l,sh),b\" | SCallParams: "P \ \es,s,b\ [\] \es',s',b'\ \ P \ \C\\<^sub>sM(es),s,b\ \ \C\\<^sub>sM(es'),s',b'\" | RedSCall: "\ P \ C sees M,Static:Ts\T = (pns,body) in D; length vs = length pns; size Ts = size pns \ \ P \ \C\\<^sub>sM(map Val vs),s,True\ \ \blocks(pns, Ts, vs, body), s, False\" | SCallInitDoneRed: "\ P \ C sees M,Static:Ts\T = (pns,body) in D; sh D = Some(sfs,Done) \ (M = clinit \ sh D = Some(sfs,Processing)) \ \ P \ \C\\<^sub>sM(map Val vs),(h,l,sh), False\ \ \C\\<^sub>sM(map Val vs),(h,l,sh), True\" | SCallInitRed: "\ P \ C sees M,Static:Ts\T = (pns,body) in D; \sfs. sh D = Some(sfs,Done); M \ clinit \ \ P \ \C\\<^sub>sM(map Val vs),(h,l,sh), False\ \ \INIT D ([D],False) \ C\\<^sub>sM(map Val vs),(h,l,sh),False\" | RedSCallNone: "\ \(\b Ts T m D. P \ C sees M,b:Ts\T = m in D) \ \ P \ \C\\<^sub>sM(map Val vs),s,b\ \ \THROW NoSuchMethodError,s,False\" | RedSCallNonStatic: "\ P \ C sees M,NonStatic:Ts\T = m in D \ \ P \ \C\\<^sub>sM(map Val vs),s,b\ \ \THROW IncompatibleClassChangeError,s,False\" | BlockRedNone: "\ P \ \e, (h,l(V:=None),sh), b\ \ \e', (h',l',sh'), b'\; l' V = None; \ assigned V e \ \ P \ \{V:T; e}, (h,l,sh), b\ \ \{V:T; e'}, (h',l'(V := l V),sh'), b'\" | BlockRedSome: "\ P \ \e, (h,l(V:=None),sh), b\ \ \e', (h',l',sh'), b'\; l' V = Some v;\ assigned V e \ \ P \ \{V:T; e}, (h,l,sh), b\ \ \{V:T := Val v; e'}, (h',l'(V := l V),sh'), b'\" | InitBlockRed: "\ P \ \e, (h,l(V\v),sh), b\ \ \e', (h',l',sh'), b'\; l' V = Some v' \ \ P \ \{V:T := Val v; e}, (h,l,sh), b\ \ \{V:T := Val v'; e'}, (h',l'(V := l V),sh'), b'\" | RedBlock: "P \ \{V:T; Val u}, s, b\ \ \Val u, s, b\" | RedInitBlock: "P \ \{V:T := Val v; Val u}, s, b\ \ \Val u, s, b\" | SeqRed: "P \ \e,s,b\ \ \e',s',b'\ \ P \ \e;;e\<^sub>2, s, b\ \ \e';;e\<^sub>2, s', b'\" | RedSeq: "P \ \(Val v);;e\<^sub>2, s, b\ \ \e\<^sub>2, s, b\" | CondRed: "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'\" | RedCondT: "P \ \if (true) e\<^sub>1 else e\<^sub>2, s, b\ \ \e\<^sub>1, s, b\" | RedCondF: "P \ \if (false) e\<^sub>1 else e\<^sub>2, s, b\ \ \e\<^sub>2, s, b\" | RedWhile: "P \ \while(b) c, s, b'\ \ \if(b) (c;;while(b) c) else unit, s, b'\" | ThrowRed: "P \ \e,s,b\ \ \e',s',b'\ \ P \ \throw e, s, b\ \ \throw e', s', b'\" | RedThrowNull: "P \ \throw null, s, b\ \ \THROW NullPointer, s, b\" | TryRed: "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'\" | RedTry: "P \ \try (Val v) catch(C V) e\<^sub>2, s, b\ \ \Val v, s, b\" | RedTryCatch: "\ hp s a = Some(D,fs); P \ D \\<^sup>* C \ \ P \ \try (Throw a) catch(C V) e\<^sub>2, s, b\ \ \{V:Class C := addr a; e\<^sub>2}, s, b\" | RedTryFail: "\ hp s a = Some(D,fs); \ P \ D \\<^sup>* C \ \ P \ \try (Throw a) catch(C V) e\<^sub>2, s, b\ \ \Throw a, s, b\" | ListRed1: "P \ \e,s,b\ \ \e',s',b'\ \ P \ \e#es,s,b\ [\] \e'#es,s',b'\" | ListRed2: "P \ \es,s,b\ [\] \es',s',b'\ \ P \ \Val v # es,s,b\ [\] \Val v # es',s',b'\" \ \Initialization procedure\ | RedInit: "\sub_RI e \ P \ \INIT C (Nil,b) \ e,s,b'\ \ \e,s,icheck P C e\" | InitNoneRed: "sh C = None \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \ \INIT C' (C#Cs,False) \ e,(h,l,sh(C \ (sblank P C, Prepared))),b\" | RedInitDone: "sh C = Some(sfs,Done) \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \ \INIT C' (Cs,True) \ e,(h,l,sh),b\" | RedInitProcessing: "sh C = Some(sfs,Processing) \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \ \INIT C' (Cs,True) \ e,(h,l,sh),b\" | RedInitError: "sh C = Some(sfs,Error) \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \ \RI (C,THROW NoClassDefFoundError);Cs \ e,(h,l,sh),b\" | InitObjectRed: "\ sh C = Some(sfs,Prepared); C = Object; sh' = sh(C \ (sfs,Processing)) \ \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \ \INIT C' (C#Cs,True) \ e,(h,l,sh'),b\" | InitNonObjectSuperRed: "\ sh C = Some(sfs,Prepared); C \ Object; class P C = Some (D,r); sh' = sh(C \ (sfs,Processing)) \ \ P \ \INIT C' (C#Cs,False) \ e,(h,l,sh),b\ \ \INIT C' (D#C#Cs,False) \ e,(h,l,sh'),b\" | RedInitRInit: "P \ \INIT C' (C#Cs,True) \ e,(h,l,sh),b\ \ \RI (C,C\\<^sub>sclinit([]));Cs \ e,(h,l,sh),b\" | RInitRed: "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'\" | RedRInit: "\ sh C = Some (sfs, i); sh' = sh(C \ (sfs,Done)); C' = last(C#Cs) \ \ P \ \RI (C, Val v);Cs \ e, (h,l,sh), b\ \ \INIT C' (Cs,True) \ e, (h,l,sh'), b\" \ \Exception propagation\ | CastThrow: "P \ \Cast C (throw e), s, b\ \ \throw e, s, b\" | BinOpThrow1: "P \ \(throw e) \bop\ e\<^sub>2, s, b\ \ \throw e, s, b\" | BinOpThrow2: "P \ \(Val v\<^sub>1) \bop\ (throw e), s, b\ \ \throw e, s, b\" | LAssThrow: "P \ \V:=(throw e), s, b\ \ \throw e, s, b\" | FAccThrow: "P \ \(throw e)\F{D}, s, b\ \ \throw e, s, b\" | FAssThrow1: "P \ \(throw e)\F{D}:=e\<^sub>2, s, b\ \ \throw e, s, b\" | FAssThrow2: "P \ \Val v\F{D}:=(throw e), s, b\ \ \throw e, s, b\" | SFAssThrow: "P \ \C\\<^sub>sF{D}:=(throw e), s, b\ \ \throw e, s, b\" | CallThrowObj: "P \ \(throw e)\M(es), s, b\ \ \throw e, s, b\" | CallThrowParams: "\ es = map Val vs @ throw e # es' \ \ P \ \(Val v)\M(es), s, b\ \ \throw e, s, b\" | SCallThrowParams: "\ es = map Val vs @ throw e # es' \ \ P \ \C\\<^sub>sM(es), s, b\ \ \throw e, s, b\" | BlockThrow: "P \ \{V:T; Throw a}, s, b\ \ \Throw a, s, b\" | InitBlockThrow: "P \ \{V:T := Val v; Throw a}, s, b\ \ \Throw a, s, b\" | SeqThrow: "P \ \(throw e);;e\<^sub>2, s, b\ \ \throw e, s, b\" | CondThrow: "P \ \if (throw e) e\<^sub>1 else e\<^sub>2, s, b\ \ \throw e, s, b\" | ThrowThrow: "P \ \throw(throw e), s, b\ \ \throw e, s, b\" | RInitInitThrow: "\ sh C = Some(sfs,i); sh' = sh(C \ (sfs,Error)) \ \ P \ \RI (C,Throw a);D#Cs \ e,(h,l,sh),b\ \ \RI (D,Throw a);Cs \ e,(h,l,sh'),b\" | RInitThrow: "\ sh C = Some(sfs, i); sh' = sh(C \ (sfs,Error)) \ \ P \ \RI (C,Throw a);Nil \ e,(h,l,sh),b\ \ \Throw a,(h,l,sh'),b\" (*<*) lemmas red_reds_induct = red_reds.induct [split_format (complete)] and red_reds_inducts = red_reds.inducts [split_format (complete)] inductive_cases [elim!]: "P \ \V:=e,s,b\ \ \e',s',b'\" "P \ \e1;;e2,s,b\ \ \e',s',b'\" (*>*) subsection\ The reflexive transitive closure \ abbreviation Step :: "J_prog \ expr \ state \ bool \ expr \ state \ bool \ bool" ("_ \ ((1\_,/_,/_\) \*/ (1\_,/_,/_\))" [51,0,0,0,0,0,0] 81) where "P \ \e,s,b\ \* \e',s',b'\ \ ((e,s,b), e',s',b') \ (red P)\<^sup>*" abbreviation Steps :: "J_prog \ expr list \ state \ bool \ expr list \ state \ bool \ bool" ("_ \ ((1\_,/_,/_\) [\]*/ (1\_,/_,/_\))" [51,0,0,0,0,0,0] 81) where "P \ \es,s,b\ [\]* \es',s',b'\ \ ((es,s,b), es',s',b') \ (reds P)\<^sup>*" lemmas converse_rtrancl_induct3 = converse_rtrancl_induct [of "(ax, ay, az)" "(bx, by, bz)", split_format (complete), consumes 1, case_names refl step] lemma converse_rtrancl_induct_red[consumes 1]: assumes "P \ \e,(h,l,sh),b\ \* \e',(h',l',sh'),b'\" and "\e h l sh b. R e h l sh b e h l sh b" and "\e\<^sub>0 h\<^sub>0 l\<^sub>0 sh\<^sub>0 b\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 b\<^sub>1 e' h' l' sh' b'. \ P \ \e\<^sub>0,(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\ \ \e\<^sub>1,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b\<^sub>1\; R e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 b\<^sub>1 e' h' l' sh' b' \ \ R e\<^sub>0 h\<^sub>0 l\<^sub>0 sh\<^sub>0 b\<^sub>0 e' h' l' sh' b'" shows "R e h l sh b e' h' l' sh' b'" (*<*) proof - { fix s s' assume reds: "P \ \e,s,b\ \* \e',s',b'\" and base: "\e s b. R e (hp s) (lcl s) (shp s) b e (hp s) (lcl s) (shp s) b" and red\<^sub>1: "\e\<^sub>0 s\<^sub>0 b\<^sub>0 e\<^sub>1 s\<^sub>1 b\<^sub>1 e' s' b'. \ P \ \e\<^sub>0,s\<^sub>0,b\<^sub>0\ \ \e\<^sub>1,s\<^sub>1,b\<^sub>1\; R e\<^sub>1 (hp s\<^sub>1) (lcl s\<^sub>1) (shp s\<^sub>1) b\<^sub>1 e' (hp s') (lcl s') (shp s') b' \ \ R e\<^sub>0 (hp s\<^sub>0) (lcl s\<^sub>0) (shp s\<^sub>0) b\<^sub>0 e' (hp s') (lcl s') (shp s') b'" from reds have "R e (hp s) (lcl s) (shp s) b e' (hp s') (lcl s') (shp s') b'" proof (induct rule:converse_rtrancl_induct3) case refl show ?case by(rule base) next case step thus ?case by(blast intro:red\<^sub>1) qed } with assms show ?thesis by fastforce qed (*>*) subsection\Some easy lemmas\ lemma [iff]: "\ P \ \[],s,b\ [\] \es',s',b'\" (*<*)by(blast elim: reds.cases)(*>*) lemma [iff]: "\ P \ \Val v,s,b\ \ \e',s',b'\" (*<*)by(fastforce elim: red.cases)(*>*) lemma val_no_step: "val_of e = \v\ \ \ P \ \e,s,b\ \ \e',s',b'\" (*<*)by(drule val_of_spec, simp)(*>*) lemma [iff]: "\ P \ \Throw a,s,b\ \ \e',s',b'\" (*<*)by(fastforce elim: red.cases)(*>*) lemma map_Vals_no_step [iff]: "\ P \ \map Val vs,s,b\ [\] \es',s',b'\" (*<*) -apply(induct vs arbitrary: es', simp) -apply(rule notI) -apply(erule reds.cases, auto) -done +proof(induct vs arbitrary: es') + case (Cons a vs) + { assume "P \ \map Val (a # vs),s,b\ [\] \es',s',b'\" + then have False by(rule reds.cases) (insert Cons, auto) + } + then show ?case by clarsimp +qed simp (*>*) lemma vals_no_step: "map_vals_of es = \vs\ \ \ P \ \es,s,b\ [\] \es',s',b'\" (*<*)by(drule map_vals_of_spec, simp)(*>*) lemma vals_throw_no_step [iff]: "\ P \ \map Val vs @ Throw a # es,s,b\ [\] \es',s',b'\" (*<*) -apply(induct vs arbitrary: es', auto) -apply(erule reds.cases, auto) -apply(erule reds.cases, auto) -done +proof(induct vs arbitrary: es') + case Nil + { assume "P \ \Throw a # es,s,b\ [\] \es',s',b'\" + then have False by(rule reds.cases) (insert Cons, auto) + } + then show ?case by clarsimp +next + case (Cons a' vs') + { assume "P \ \map Val (a'#vs') @ Throw a # es,s,b\ [\] \es',s',b'\" + then have False by(rule reds.cases) (insert Cons, auto) + } + then show ?case by clarsimp +qed (*>*) lemma lass_val_of_red: "\ lass_val_of e = \a\; P \ \e,(h, l, sh),b\ \ \e',(h', l', sh'),b'\ \ \ e' = unit \ h' = h \ l' = l(fst a\snd a) \ sh' = sh \ b = b'" (*<*)by(drule lass_val_of_spec, auto)(*>*) lemma final_no_step [iff]: "final e \ \ P \ \e,s,b\ \ \e',s',b'\" (*<*)by(erule finalE, simp+)(*>*) lemma finals_no_step [iff]: "finals es \ \ P \ \es,s,b\ [\] \es',s',b'\" (*<*)by(erule finalsE, simp+)(*>*) lemma reds_final_same: "P \ \e,s,b\ \* \e',s',b'\ \ final e \ e = e' \ s = s' \ b = b'" proof(induct rule:converse_rtrancl_induct3) case refl show ?case by simp next case (step e0 s0 b0 e1 s1 b1) show ?case proof(rule finalE[OF step.prems(1)]) fix v assume "e0 = Val v" then show ?thesis using step by simp next fix a assume "e0 = Throw a" then show ?thesis using step by simp qed qed lemma reds_throw: "P \ \e,s,b\ \* \e',s',b'\ \ (\e\<^sub>t. throw_of e = \e\<^sub>t\ \ \e\<^sub>t'. throw_of e' = \e\<^sub>t'\)" proof(induct rule:converse_rtrancl_induct3) case refl then show ?case by simp next case (step e0 s0 b0 e1 s1 b1) then show ?case by(auto elim: red.cases) qed lemma red_hext_incr: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ h \ h'" and reds_hext_incr: "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ h \ h'" (*<*) proof(induct rule:red_reds_inducts) case RedNew thus ?case by(fastforce dest:new_Addr_SomeD simp:hext_def split:if_splits) next case RedFAss thus ?case by(simp add:hext_def split:if_splits) qed simp_all (*>*) lemma red_lcl_incr: "P \ \e,(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\ \ \e',(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b'\ \ dom l\<^sub>0 \ dom l\<^sub>1" and reds_lcl_incr: "P \ \es,(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\ [\] \es',(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b'\ \ dom l\<^sub>0 \ dom l\<^sub>1" (*<*)by(induct rule: red_reds_inducts)(auto simp del:fun_upd_apply)(*>*) lemma red_lcl_add: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ (\l\<^sub>0. P \ \e,(h,l\<^sub>0++l,sh),b\ \ \e',(h',l\<^sub>0++l',sh'),b'\)" and reds_lcl_add: "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ (\l\<^sub>0. P \ \es,(h,l\<^sub>0++l,sh),b\ [\] \es',(h',l\<^sub>0++l',sh'),b'\)" (*<*) proof (induct rule:red_reds_inducts) case RedCast thus ?case by(fastforce intro:red_reds.intros) next case RedCastFail thus ?case by(force intro:red_reds.intros) next case RedFAcc thus ?case by(fastforce intro:red_reds.intros) next case RedCall thus ?case by(fastforce intro:red_reds.intros) next case (InitBlockRed e h l V v sh b e' h' l' sh' b' v' T l\<^sub>0) have IH: "\l\<^sub>0. P \ \e,(h, l\<^sub>0 ++ l(V \ v), sh),b\ \ \e',(h', l\<^sub>0 ++ l', sh'),b'\" and l'V: "l' V = Some v'" by fact+ from IH have IH': "P \ \e,(h, (l\<^sub>0 ++ l)(V \ v), sh),b\ \ \e',(h', l\<^sub>0 ++ l', sh'),b'\" by simp have "(l\<^sub>0 ++ l')(V := (l\<^sub>0 ++ l) V) = l\<^sub>0 ++ l'(V := l V)" by(rule ext)(simp add:map_add_def) with red_reds.InitBlockRed[OF IH'] l'V show ?case by(simp del:fun_upd_apply) next case (BlockRedNone e h l V sh b e' h' l' sh' b' T l\<^sub>0) have IH: "\l\<^sub>0. P \ \e,(h, l\<^sub>0 ++ l(V := None), sh),b\ \ \e',(h', l\<^sub>0 ++ l', sh'),b'\" and l'V: "l' V = None" and unass: "\ assigned V e" by fact+ have "l\<^sub>0(V := None) ++ l(V := None) = (l\<^sub>0 ++ l)(V := None)" by(simp add:fun_eq_iff map_add_def) hence IH': "P \ \e,(h, (l\<^sub>0++l)(V := None), sh),b\ \ \e',(h', l\<^sub>0(V := None) ++ l', sh'),b'\" using IH[of "l\<^sub>0(V := None)"] by simp have "(l\<^sub>0(V := None) ++ l')(V := (l\<^sub>0 ++ l) V) = l\<^sub>0 ++ l'(V := l V)" by(simp add:fun_eq_iff map_add_def) with red_reds.BlockRedNone[OF IH' _ unass] l'V show ?case by(simp add: map_add_def) next case (BlockRedSome e h l V sh b e' h' l' sh' b' v T l\<^sub>0) have IH: "\l\<^sub>0. P \ \e,(h, l\<^sub>0 ++ l(V := None), sh),b\ \ \e',(h', l\<^sub>0 ++ l', sh'),b'\" and l'V: "l' V = Some v" and unass: "\ assigned V e" by fact+ have "l\<^sub>0(V := None) ++ l(V := None) = (l\<^sub>0 ++ l)(V := None)" by(simp add:fun_eq_iff map_add_def) hence IH': "P \ \e,(h, (l\<^sub>0++l)(V := None), sh),b\ \ \e',(h', l\<^sub>0(V := None) ++ l', sh'),b'\" using IH[of "l\<^sub>0(V := None)"] by simp have "(l\<^sub>0(V := None) ++ l')(V := (l\<^sub>0 ++ l) V) = l\<^sub>0 ++ l'(V := l V)" by(simp add:fun_eq_iff map_add_def) with red_reds.BlockRedSome[OF IH' _ unass] l'V show ?case by(simp add:map_add_def) next case RedTryCatch thus ?case by(fastforce intro:red_reds.intros) next case RedTryFail thus ?case by(force intro!:red_reds.intros) qed (simp_all add:red_reds.intros) (*>*) lemma Red_lcl_add: assumes "P \ \e,(h,l,sh), b\ \* \e',(h',l',sh'), b'\" shows "P \ \e,(h,l\<^sub>0++l,sh),b\ \* \e',(h',l\<^sub>0++l',sh'),b'\" (*<*) using assms proof(induct rule:converse_rtrancl_induct_red) case 1 thus ?case by simp next case 2 thus ?case by (blast dest: red_lcl_add intro: converse_rtrancl_into_rtrancl) qed (*>*) lemma assumes wf: "wwf_J_prog P" shows red_proc_pres: "P \ \e,(h,l,sh),b\ \ \e',(h',l',sh'),b'\ \ not_init C e \ sh C = \(sfs, Processing)\ \ not_init C e' \ (\sfs'. sh' C = \(sfs', Processing)\)" and reds_proc_pres: "P \ \es,(h,l,sh),b\ [\] \es',(h',l',sh'),b'\ \ not_inits C es \ sh C = \(sfs, Processing)\ \ not_inits C es' \ (\sfs'. sh' C = \(sfs', Processing)\)" (*<*) proof(induct rule:red_reds_inducts) case RedCall then show ?case using sees_wwf_nsub_RI[OF wf RedCall.hyps(2)] sub_RI_blocks_body nsub_RI_not_init by auto next case RedSCall then show ?case using sees_wwf_nsub_RI[OF wf RedSCall.hyps(1)] sub_RI_blocks_body nsub_RI_not_init by auto next case (RedInitDone sh C sfs C' Cs e h l b) then show ?case by(cases Cs, auto) next case (RedInitProcessing sh C sfs C' Cs e h l b) then show ?case by(cases Cs, auto) next case (RedRInit sh C sfs i sh' C' Cs v e h l b) then show ?case by(cases Cs, auto) next case (CallThrowParams es vs e es' v M h l sh b) then show ?case by(auto dest: not_inits_def') next case (SCallThrowParams es vs e es' C M h l sh b) then show ?case by(auto dest: not_inits_def') qed(auto) 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 diff --git a/thys/JinjaDCI/J/WWellForm.thy b/thys/JinjaDCI/J/WWellForm.thy --- a/thys/JinjaDCI/J/WWellForm.thy +++ b/thys/JinjaDCI/J/WWellForm.thy @@ -1,42 +1,40 @@ (* Title: JinjaDCI/J/WWellForm.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory J/WWellForm.thy by Tobias Nipkow *) section \ Weak well-formedness of Jinja programs \ theory WWellForm imports "../Common/WellForm" Expr begin definition wwf_J_mdecl :: "J_prog \ cname \ J_mb mdecl \ bool" where "wwf_J_mdecl P C \ \(M,b,Ts,T,(pns,body)). length Ts = length pns \ distinct pns \ \sub_RI body \ (case b of NonStatic \ this \ set pns \ fv body \ {this} \ set pns | Static \ fv body \ set pns)" lemma wwf_J_mdecl_NonStatic[simp]: "wwf_J_mdecl P C (M,NonStatic,Ts,T,pns,body) = (length Ts = length pns \ distinct pns \ \sub_RI body \ this \ set pns \ fv body \ {this} \ set pns)" (*<*)by(simp add:wwf_J_mdecl_def)(*>*) lemma wwf_J_mdecl_Static[simp]: "wwf_J_mdecl P C (M,Static,Ts,T,pns,body) = (length Ts = length pns \ distinct pns \ \sub_RI body \ fv body \ set pns)" (*<*)by(simp add:wwf_J_mdecl_def)(*>*) abbreviation wwf_J_prog :: "J_prog \ bool" where "wwf_J_prog \ wf_prog wwf_J_mdecl" lemma sees_wwf_nsub_RI: "\ wwf_J_prog P; P \ C sees M,b : Ts\T = (pns, body) in D \ \ \sub_RI body" -apply(drule sees_wf_mdecl, simp) -apply(unfold wwf_J_mdecl_def wf_mdecl_def, simp) -done +(*<*)by(auto dest!: sees_wf_mdecl simp: wwf_J_mdecl_def wf_mdecl_def)(*>*) end diff --git a/thys/JinjaDCI/J/WellType.thy b/thys/JinjaDCI/J/WellType.thy --- a/thys/JinjaDCI/J/WellType.thy +++ b/thys/JinjaDCI/J/WellType.thy @@ -1,255 +1,219 @@ (* Title: JinjaDCI/J/WellType.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory J/WellType.thy by Tobias Nipkow *) section \ Well-typedness of Jinja expressions \ theory WellType imports "../Common/Objects" Expr begin type_synonym env = "vname \ ty" inductive WT :: "[J_prog,env, expr , ty ] \ bool" ("_,_ \ _ :: _" [51,51,51]50) and WTs :: "[J_prog,env, expr list, ty list] \ bool" ("_,_ \ _ [::] _" [51,51,51]50) for P :: J_prog where WTNew: "is_class P C \ P,E \ new C :: Class C" | WTCast: "\ P,E \ e :: Class D; is_class P C; P \ C \\<^sup>* D \ P \ D \\<^sup>* C \ \ P,E \ Cast C e :: Class C" | WTVal: "typeof v = Some T \ P,E \ Val v :: T" | WTVar: "E V = Some T \ P,E \ Var V :: T" | WTBinOpEq: "\ P,E \ e\<^sub>1 :: T\<^sub>1; P,E \ e\<^sub>2 :: T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1 \ \ P,E \ e\<^sub>1 \Eq\ e\<^sub>2 :: Boolean" | WTBinOpAdd: "\ P,E \ e\<^sub>1 :: Integer; P,E \ e\<^sub>2 :: Integer \ \ P,E \ e\<^sub>1 \Add\ e\<^sub>2 :: Integer" | WTLAss: "\ E V = Some T; P,E \ e :: T'; P \ T' \ T; V \ this \ \ P,E \ V:=e :: Void" | WTFAcc: "\ P,E \ e :: Class C; P \ C sees F,NonStatic:T in D \ \ P,E \ e\F{D} :: T" | WTSFAcc: "\ P \ C sees F,Static:T in D \ \ P,E \ C\\<^sub>sF{D} :: T" | WTFAss: "\ P,E \ e\<^sub>1 :: Class C; P \ C sees F,NonStatic:T in D; P,E \ e\<^sub>2 :: T'; P \ T' \ T \ \ P,E \ e\<^sub>1\F{D}:=e\<^sub>2 :: Void" | WTSFAss: "\ P \ C sees F,Static:T in D; P,E \ e\<^sub>2 :: T'; P \ T' \ T \ \ P,E \ C\\<^sub>sF{D}:=e\<^sub>2 :: Void" | WTCall: "\ P,E \ e :: Class C; P \ C sees M,NonStatic:Ts \ T = (pns,body) in D; P,E \ es [::] Ts'; P \ Ts' [\] Ts \ \ P,E \ e\M(es) :: T" | WTSCall: "\ P \ C sees M,Static:Ts \ T = (pns,body) in D; P,E \ es [::] Ts'; P \ Ts' [\] Ts; M \ clinit \ \ P,E \ C\\<^sub>sM(es) :: T" | WTBlock: "\ is_type P T; P,E(V \ T) \ e :: T' \ \ P,E \ {V:T; e} :: T'" | WTSeq: "\ P,E \ e\<^sub>1::T\<^sub>1; P,E \ e\<^sub>2::T\<^sub>2 \ \ P,E \ e\<^sub>1;;e\<^sub>2 :: T\<^sub>2" | WTCond: "\ P,E \ e :: Boolean; P,E \ e\<^sub>1::T\<^sub>1; P,E \ e\<^sub>2::T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1; P \ T\<^sub>1 \ T\<^sub>2 \ T = T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ T = T\<^sub>1 \ \ P,E \ if (e) e\<^sub>1 else e\<^sub>2 :: T" | WTWhile: "\ P,E \ e :: Boolean; P,E \ c::T \ \ P,E \ while (e) c :: Void" | WTThrow: "P,E \ e :: Class C \ P,E \ throw e :: Void" | WTTry: "\ P,E \ e\<^sub>1 :: T; P,E(V \ Class C) \ e\<^sub>2 :: T; is_class P C \ \ P,E \ try e\<^sub>1 catch(C V) e\<^sub>2 :: T" \ \well-typed expression lists\ | WTNil: "P,E \ [] [::] []" | WTCons: "\ P,E \ e :: T; P,E \ es [::] Ts \ \ P,E \ e#es [::] T#Ts" (*<*) declare WT_WTs.intros[intro!] (* WTNil[iff] *) lemmas WT_WTs_induct = WT_WTs.induct [split_format (complete)] and WT_WTs_inducts = WT_WTs.inducts [split_format (complete)] (*>*) lemma init_nwt [simp]:"\P,E \ INIT C (Cs,b) \ e :: T" by(auto elim:WT.cases) lemma ri_nwt [simp]:"\P,E \ RI(C,e);Cs \ e' :: T" by(auto elim:WT.cases) lemma [iff]: "(P,E \ [] [::] Ts) = (Ts = [])" -(*<*) -apply(rule iffI) -apply (auto elim: WTs.cases) -done -(*>*) +(*<*)by (rule iffI) (auto elim: WTs.cases)(*>*) lemma [iff]: "(P,E \ e#es [::] T#Ts) = (P,E \ e :: T \ P,E \ es [::] Ts)" -(*<*) -apply(rule iffI) -apply (auto elim: WTs.cases) -done -(*>*) +(*<*)by (rule iffI) (auto elim: WTs.cases)(*>*) lemma [iff]: "(P,E \ (e#es) [::] Ts) = (\U Us. Ts = U#Us \ P,E \ e :: U \ P,E \ es [::] Us)" -(*<*) -apply(rule iffI) -apply (auto elim: WTs.cases) -done -(*>*) +(*<*)by (rule iffI) (auto elim: WTs.cases)(*>*) lemma [iff]: "\Ts. (P,E \ es\<^sub>1 @ es\<^sub>2 [::] Ts) = (\Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \ P,E \ es\<^sub>1 [::] Ts\<^sub>1 \ P,E \ es\<^sub>2[::]Ts\<^sub>2)" (*<*) -apply(induct es\<^sub>1 type:list) - apply simp -apply clarsimp -apply(erule thin_rl) -apply (rule iffI) - apply clarsimp - apply(rule exI)+ - apply(rule conjI) - prefer 2 apply blast - apply simp -apply fastforce -done +proof(induct es\<^sub>1 type:list) + case (Cons a list) + let ?lhs = "(\U Us. Ts = U # Us \ P,E \ a :: U \ + (\Ts\<^sub>1 Ts\<^sub>2. Us = Ts\<^sub>1 @ Ts\<^sub>2 \ P,E \ list [::] Ts\<^sub>1 \ P,E \ es\<^sub>2 [::] Ts\<^sub>2))" + let ?rhs = "(\Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \ + (\U Us. Ts\<^sub>1 = U # Us \ P,E \ a :: U \ P,E \ list [::] Us) \ P,E \ es\<^sub>2 [::] Ts\<^sub>2)" + { assume ?lhs + then have ?rhs by (auto intro: Cons_eq_appendI) + } + moreover { + assume ?rhs + then have ?lhs by fastforce + } + ultimately have "?lhs = ?rhs" by(rule iffI) + then show ?case by (clarsimp simp: Cons) +qed simp (*>*) lemma [iff]: "P,E \ Val v :: T = (typeof v = Some T)" -(*<*) -apply(rule iffI) -apply (auto elim: WT.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*) lemma [iff]: "P,E \ Var V :: T = (E V = Some T)" -(*<*) -apply(rule iffI) -apply (auto elim: WT.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*) lemma [iff]: "P,E \ e\<^sub>1;;e\<^sub>2 :: T\<^sub>2 = (\T\<^sub>1. P,E \ e\<^sub>1::T\<^sub>1 \ P,E \ e\<^sub>2::T\<^sub>2)" -(*<*) -apply(rule iffI) -apply (auto elim: WT.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*) lemma [iff]: "(P,E \ {V:T; e} :: T') = (is_type P T \ P,E(V\T) \ e :: T')" -(*<*) -apply(rule iffI) -apply (auto elim: WT.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*) (*<*) inductive_cases WT_elim_cases[elim!]: "P,E \ V :=e :: T" "P,E \ if (e) e\<^sub>1 else e\<^sub>2 :: T" "P,E \ while (e) c :: T" "P,E \ throw e :: T" "P,E \ try e\<^sub>1 catch(C V) e\<^sub>2 :: T" "P,E \ Cast D e :: T" "P,E \ a\F{D} :: T" "P,E \ C\\<^sub>sF{D} :: T" "P,E \ a\F{D} := v :: T" "P,E \ C\\<^sub>sF{D} := v :: T" "P,E \ e\<^sub>1 \bop\ e\<^sub>2 :: T" "P,E \ new C :: T" "P,E \ e\M(ps) :: T" "P,E \ C\\<^sub>sM(ps) :: T" (*>*) lemma wt_env_mono: "P,E \ e :: T \ (\E'. E \\<^sub>m E' \ P,E' \ e :: T)" and "P,E \ es [::] Ts \ (\E'. E \\<^sub>m E' \ P,E' \ es [::] Ts)" (*<*) -apply(induct rule: WT_WTs_inducts) -apply(simp add: WTNew) -apply(fastforce simp: WTCast) -apply(fastforce simp: WTVal) -apply(simp add: WTVar map_le_def dom_def) -apply(fastforce simp: WTBinOpEq) -apply(fastforce simp: WTBinOpAdd) -apply(force simp:map_le_def) -apply(fastforce simp: WTFAcc) -apply(fastforce) -apply(fastforce simp: WTFAss del:WT_WTs.intros WT_elim_cases) -apply(fastforce) -apply(fastforce simp: WTCall) -apply(fastforce) -apply(fastforce simp: map_le_def WTBlock) -apply(fastforce simp: WTSeq) -apply(fastforce simp: WTCond) -apply(fastforce simp: WTWhile) -apply(fastforce simp: WTThrow) -apply(fastforce simp: WTTry map_le_def dom_def) -apply(simp add: WTNil) -apply(simp add: WTCons) -done +proof(induct rule: WT_WTs_inducts) + case WTVar then show ?case by(simp add: map_le_def dom_def) +next + case WTLAss then show ?case by(force simp:map_le_def) +qed fastforce+ (*>*) lemma WT_fv: "P,E \ e :: T \ fv e \ dom E" and "P,E \ es [::] Ts \ fvs es \ dom E" (*<*) -apply(induct rule:WT_WTs.inducts) -apply(simp_all del: fun_upd_apply) -apply fast+ -done +proof(induct rule:WT_WTs.inducts) + case WTVar then show ?case by fastforce +next + case WTLAss then show ?case by fastforce +next + case WTBlock then show ?case by fastforce +next + case WTTry then show ?case by fastforce +qed simp_all +(*>*) lemma WT_nsub_RI: "P,E \ e :: T \ \sub_RI e" and WTs_nsub_RIs: "P,E \ es [::] Ts \ \sub_RIs es" -proof(induct rule: WT_WTs.inducts) qed(simp_all) +(*<*)proof(induct rule: WT_WTs.inducts) qed(simp_all) end (*>*) diff --git a/thys/JinjaDCI/J/WellTypeRT.thy b/thys/JinjaDCI/J/WellTypeRT.thy --- a/thys/JinjaDCI/J/WellTypeRT.thy +++ b/thys/JinjaDCI/J/WellTypeRT.thy @@ -1,353 +1,293 @@ (* Title: JinjaDCI/J/WellTypeRT.thy Author: Tobias Nipkow, Susannah Mansky Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC Based on the Jinja theory J/WellTypeRT.thy by Tobias Nipkow *) section \ Runtime Well-typedness \ theory WellTypeRT imports WellType begin inductive WTrt :: "J_prog \ heap \ sheap \ env \ expr \ ty \ bool" and WTrts :: "J_prog \ heap \ sheap \ env \ expr list \ ty list \ bool" and WTrt2 :: "[J_prog,env,heap,sheap,expr,ty] \ bool" ("_,_,_,_ \ _ : _" [51,51,51,51]50) and WTrts2 :: "[J_prog,env,heap,sheap,expr list, ty list] \ bool" ("_,_,_,_ \ _ [:] _" [51,51,51,51]50) for P :: J_prog and h :: heap and sh :: sheap where "P,E,h,sh \ e : T \ WTrt P h sh E e T" | "P,E,h,sh \ es[:]Ts \ WTrts P h sh E es Ts" | WTrtNew: "is_class P C \ P,E,h,sh \ new C : Class C" | WTrtCast: "\ P,E,h,sh \ e : T; is_refT T; is_class P C \ \ P,E,h,sh \ Cast C e : Class C" | WTrtVal: "typeof\<^bsub>h\<^esub> v = Some T \ P,E,h,sh \ Val v : T" | WTrtVar: "E V = Some T \ P,E,h,sh \ Var V : T" | WTrtBinOpEq: "\ P,E,h,sh \ e\<^sub>1 : T\<^sub>1; P,E,h,sh \ e\<^sub>2 : T\<^sub>2 \ \ P,E,h,sh \ e\<^sub>1 \Eq\ e\<^sub>2 : Boolean" | WTrtBinOpAdd: "\ P,E,h,sh \ e\<^sub>1 : Integer; P,E,h,sh \ e\<^sub>2 : Integer \ \ P,E,h,sh \ e\<^sub>1 \Add\ e\<^sub>2 : Integer" | WTrtLAss: "\ E V = Some T; P,E,h,sh \ e : T'; P \ T' \ T \ \ P,E,h,sh \ V:=e : Void" | WTrtFAcc: "\ P,E,h,sh \ e : Class C; P \ C has F,NonStatic:T in D \ \ P,E,h,sh \ e\F{D} : T" | WTrtFAccNT: "P,E,h,sh \ e : NT \ P,E,h,sh \ e\F{D} : T" | WTrtSFAcc: "\ P \ C has F,Static:T in D \ \ P,E,h,sh \ C\\<^sub>sF{D} : T" | WTrtFAss: "\ P,E,h,sh \ e\<^sub>1 : Class C; P \ C has F,NonStatic:T in D; P,E,h,sh \ e\<^sub>2 : T\<^sub>2; P \ T\<^sub>2 \ T \ \ P,E,h,sh \ e\<^sub>1\F{D}:=e\<^sub>2 : Void" | WTrtFAssNT: "\ P,E,h,sh \ e\<^sub>1:NT; P,E,h,sh \ e\<^sub>2 : T\<^sub>2 \ \ P,E,h,sh \ e\<^sub>1\F{D}:=e\<^sub>2 : Void" | WTrtSFAss: "\ P,E,h,sh \ e\<^sub>2 : T\<^sub>2; P \ C has F,Static:T in D; P \ T\<^sub>2 \ T \ \ P,E,h,sh \ C\\<^sub>sF{D}:=e\<^sub>2 : Void" | WTrtCall: "\ P,E,h,sh \ e : Class C; P \ C sees M,NonStatic:Ts \ T = (pns,body) in D; P,E,h,sh \ es [:] Ts'; P \ Ts' [\] Ts \ \ P,E,h,sh \ e\M(es) : T" | WTrtCallNT: "\ P,E,h,sh \ e : NT; P,E,h,sh \ es [:] Ts \ \ P,E,h,sh \ e\M(es) : T" | WTrtSCall: "\ P \ C sees M,Static:Ts \ T = (pns,body) in D; P,E,h,sh \ es [:] Ts'; P \ Ts' [\] Ts; M = clinit \ sh D = \(sfs,Processing)\ \ es = map Val vs \ \ P,E,h,sh \ C\\<^sub>sM(es) : T" | WTrtBlock: "P,E(V\T),h,sh \ e : T' \ P,E,h,sh \ {V:T; e} : T'" | WTrtSeq: "\ P,E,h,sh \ e\<^sub>1:T\<^sub>1; P,E,h,sh \ e\<^sub>2:T\<^sub>2 \ \ P,E,h,sh \ e\<^sub>1;;e\<^sub>2 : T\<^sub>2" | WTrtCond: "\ P,E,h,sh \ e : Boolean; P,E,h,sh \ e\<^sub>1:T\<^sub>1; P,E,h,sh \ e\<^sub>2:T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1; P \ T\<^sub>1 \ T\<^sub>2 \ T = T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ T = T\<^sub>1 \ \ P,E,h,sh \ if (e) e\<^sub>1 else e\<^sub>2 : T" | WTrtWhile: "\ P,E,h,sh \ e : Boolean; P,E,h,sh \ c:T \ \ P,E,h,sh \ while(e) c : Void" | WTrtThrow: "\ P,E,h,sh \ e : T\<^sub>r; is_refT T\<^sub>r \ \ P,E,h,sh \ throw e : T" | WTrtTry: "\ P,E,h,sh \ e\<^sub>1 : T\<^sub>1; P,E(V \ Class C),h,sh \ e\<^sub>2 : T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ \ P,E,h,sh \ try e\<^sub>1 catch(C V) e\<^sub>2 : T\<^sub>2" | WTrtInit: "\ P,E,h,sh \ e : T; \C' \ set (C#Cs). is_class P C'; \sub_RI e; \C' \ set (tl Cs). \sfs. sh C' = \(sfs,Processing)\; b \ (\C' \ set Cs. \sfs. sh C' = \(sfs,Processing)\); distinct Cs; supercls_lst P Cs \ \ P,E,h,sh \ INIT C (Cs, b) \ e : T" | WTrtRI: "\ P,E,h,sh \ e : T; P,E,h,sh \ e' : T'; \C' \ set (C#Cs). is_class P C'; \sub_RI e'; \C' \ set (C#Cs). not_init C' e; \C' \ set Cs. \sfs. sh C' = \(sfs,Processing)\; \sfs. sh C = \(sfs, Processing)\ \ (sh C = \(sfs, Error)\ \ e = THROW NoClassDefFoundError); distinct (C#Cs); supercls_lst P (C#Cs) \ \ P,E,h,sh \ RI(C, e);Cs \ e' : T'" \ \well-typed expression lists\ | WTrtNil: "P,E,h,sh \ [] [:] []" | WTrtCons: "\ P,E,h,sh \ e : T; P,E,h,sh \ es [:] Ts \ \ P,E,h,sh \ e#es [:] T#Ts" (*<*) declare WTrt_WTrts.intros[intro!] WTrtNil[iff] declare WTrtFAcc[rule del] WTrtFAccNT[rule del] WTrtSFAcc[rule del] WTrtFAss[rule del] WTrtFAssNT[rule del] WTrtSFAss[rule del] WTrtCall[rule del] WTrtCallNT[rule del] WTrtSCall[rule del] lemmas WTrt_induct = WTrt_WTrts.induct [split_format (complete)] and WTrt_inducts = WTrt_WTrts.inducts [split_format (complete)] (*>*) subsection\Easy consequences\ lemma [iff]: "(P,E,h,sh \ [] [:] Ts) = (Ts = [])" -(*<*) -apply(rule iffI) -apply (auto elim: WTrts.cases) -done -(*>*) +(*<*)by(rule iffI) (auto elim: WTrts.cases)(*>*) lemma [iff]: "(P,E,h,sh \ e#es [:] T#Ts) = (P,E,h,sh \ e : T \ P,E,h,sh \ es [:] Ts)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrts.cases) -done -(*>*) +(*<*)by(rule iffI) (auto elim: WTrts.cases)(*>*) lemma [iff]: "(P,E,h,sh \ (e#es) [:] Ts) = (\U Us. Ts = U#Us \ P,E,h,sh \ e : U \ P,E,h,sh \ es [:] Us)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrts.cases) -done -(*>*) +(*<*)by(rule iffI) (auto elim: WTrts.cases)(*>*) lemma [simp]: "\Ts. (P,E,h,sh \ es\<^sub>1 @ es\<^sub>2 [:] Ts) = (\Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \ P,E,h,sh \ es\<^sub>1 [:] Ts\<^sub>1 & P,E,h,sh \ es\<^sub>2[:]Ts\<^sub>2)" (*<*) -apply(induct_tac es\<^sub>1) - apply simp -apply clarsimp -apply(erule thin_rl) -apply (rule iffI) - apply clarsimp - apply(rule exI)+ - apply(rule conjI) - prefer 2 apply blast - apply simp -apply fastforce -done +proof(induct es\<^sub>1) + case (Cons a list) + let ?lhs = "\Ts. (\U Us. Ts = U # Us \ P,E,h,sh \ a : U \ + (\Ts\<^sub>1 Ts\<^sub>2. Us = Ts\<^sub>1 @ Ts\<^sub>2 \ P,E,h,sh \ list [:] Ts\<^sub>1 \ P,E,h,sh \ es\<^sub>2 [:] Ts\<^sub>2))" + let ?rhs = "\Ts. (\Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \ + (\U Us. Ts\<^sub>1 = U # Us \ P,E,h,sh \ a : U \ P,E,h,sh \ list [:] Us) \ P,E,h,sh \ es\<^sub>2 [:] Ts\<^sub>2)" + { fix Ts assume "?lhs Ts" + then have "?rhs Ts" by (auto intro: Cons_eq_appendI) + } + moreover { + fix Ts assume "?rhs Ts" + then have "?lhs Ts" by fastforce + } + ultimately have "\Ts. ?lhs Ts = ?rhs Ts" by(rule iffI) + then show ?case by (clarsimp simp: Cons) +qed simp (*>*) lemma [iff]: "P,E,h,sh \ Val v : T = (typeof\<^bsub>h\<^esub> v = Some T)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrt.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*) lemma [iff]: "P,E,h,sh \ Var v : T = (E v = Some T)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrt.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*) lemma [iff]: "P,E,h,sh \ e\<^sub>1;;e\<^sub>2 : T\<^sub>2 = (\T\<^sub>1. P,E,h,sh \ e\<^sub>1:T\<^sub>1 \ P,E,h,sh \ e\<^sub>2:T\<^sub>2)" -(*<*) -apply(rule iffI) -apply (auto elim: WTrt.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*) lemma [iff]: "P,E,h,sh \ {V:T; e} : T' = (P,E(V\T),h,sh \ e : T')" -(*<*) -apply(rule iffI) -apply (auto elim: WTrt.cases) -done -(*>*) +(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*) + (*<*) inductive_cases WTrt_elim_cases[elim!]: "P,E,h,sh \ v :=e : T" "P,E,h,sh \ if (e) e\<^sub>1 else e\<^sub>2 : T" "P,E,h,sh \ while(e) c : T" "P,E,h,sh \ throw e : T" "P,E,h,sh \ try e\<^sub>1 catch(C V) e\<^sub>2 : T" "P,E,h,sh \ Cast D e : T" "P,E,h,sh \ e\F{D} : T" "P,E,h,sh \ C\\<^sub>sF{D} : T" "P,E,h,sh \ e\F{D} := v : T" "P,E,h,sh \ C\\<^sub>sF{D} := v : T" "P,E,h,sh \ e\<^sub>1 \bop\ e\<^sub>2 : T" "P,E,h,sh \ new C : T" "P,E,h,sh \ e\M{D}(es) : T" "P,E,h,sh \ C\\<^sub>sM{D}(es) : T" "P,E,h,sh \ INIT C (Cs,b) \ e : T" "P,E,h,sh \ RI(C,e);Cs \ e' : T" (*>*) subsection\Some interesting lemmas\ lemma WTrts_Val[simp]: "\Ts. (P,E,h,sh \ map Val vs [:] Ts) = (map (typeof\<^bsub>h\<^esub>) vs = map Some Ts)" (*<*) -apply(induct vs) - apply simp -apply(case_tac Ts) - apply simp -apply simp -done +proof(induct vs) + case (Cons a vs) + then show ?case by(case_tac Ts; simp) +qed simp (*>*) lemma WTrts_same_length: "\Ts. P,E,h,sh \ es [:] Ts \ length es = length Ts" (*<*)by(induct es type:list)auto(*>*) lemma WTrt_env_mono: "P,E,h,sh \ e : T \ (\E'. E \\<^sub>m E' \ P,E',h,sh \ e : T)" and "P,E,h,sh \ es [:] Ts \ (\E'. E \\<^sub>m E' \ P,E',h,sh \ es [:] Ts)" (*<*) proof(induct rule: WTrt_inducts) - case (WTrtVar E V T) - then show ?case by(simp add: WTrtVar map_le_def dom_def) + case WTrtVar then show ?case by(simp add: map_le_def dom_def) next - case (WTrtLAss E V T e T') - then show ?case by(force simp: map_le_def) + case WTrtLAss then show ?case by(force simp:map_le_def) qed(fastforce intro: WTrt_WTrts.intros)+ (*>*) lemma WTrt_hext_mono: "P,E,h,sh \ e : T \ h \ h' \ P,E,h',sh \ e : T" and WTrts_hext_mono: "P,E,h,sh \ es [:] Ts \ h \ h' \ P,E,h',sh \ es [:] Ts" (*<*) -apply(induct rule: WTrt_inducts) -apply(simp add: WTrtNew) -apply(fastforce simp: WTrtCast) -apply(fastforce simp: WTrtVal dest:hext_typeof_mono) -apply(simp add: WTrtVar) -apply(fastforce simp add: WTrtBinOpEq) -apply(fastforce simp add: WTrtBinOpAdd) -apply(fastforce simp add: WTrtLAss) -apply(fast intro: WTrtFAcc) -apply(simp add: WTrtFAccNT) -apply(fast intro: WTrtSFAcc) -apply(fastforce simp: WTrtFAss del:WTrt_WTrts.intros WTrt_elim_cases) -apply(fastforce simp: WTrtFAssNT) -apply(fastforce simp: WTrtSFAss del:WTrt_WTrts.intros WTrt_elim_cases) -apply(fastforce simp: WTrtCall) -apply(fastforce simp: WTrtCallNT) -using WTrtSCall apply blast -apply(fastforce) -apply(fastforce simp add: WTrtSeq) -apply(fastforce simp add: WTrtCond) -apply(fastforce simp add: WTrtWhile) -apply(fastforce simp add: WTrtThrow) -apply(fastforce simp: WTrtTry) -apply(simp add: WTrtInit) -apply(simp add: WTrtRI) -apply(simp add: WTrtNil) -apply(simp add: WTrtCons) -done +proof(induct rule: WTrt_inducts) + case WTrtVal then show ?case by(simp add: hext_typeof_mono) +qed(fastforce intro: WTrt_WTrts.intros)+ (*>*) lemma WTrt_shext_mono: "P,E,h,sh \ e : T \ sh \\<^sub>s sh' \ \sub_RI e \ P,E,h,sh' \ e : T" and WTrts_shext_mono: "P,E,h,sh \ es [:] Ts \ sh \\<^sub>s sh' \ \sub_RIs es \ P,E,h,sh' \ es [:] Ts" (*<*) by(induct rule: WTrt_inducts) (auto simp add: WTrt_WTrts.intros) (*>*) lemma WTrt_hext_shext_mono: "P,E,h,sh \ e : T \ h \ h' \ sh \\<^sub>s sh' \ \sub_RI e \ P,E,h',sh' \ e : T" by(auto intro: WTrt_hext_mono WTrt_shext_mono) lemma WTrts_hext_shext_mono: "P,E,h,sh \ es [:] Ts \ h \ h' \ sh \\<^sub>s sh' \ \sub_RIs es \ P,E,h',sh' \ es [:] Ts" by(auto intro: WTrts_hext_mono WTrts_shext_mono) lemma WT_implies_WTrt: "P,E \ e :: T \ P,E,h,sh \ e : T" and WTs_implies_WTrts: "P,E \ es [::] Ts \ P,E,h,sh \ es [:] Ts" (*<*) -apply(induct rule: WT_WTs_inducts) -apply fast -apply (fast) -apply(fastforce dest:typeof_lit_typeof) -apply(simp) -apply(fastforce) -apply(fastforce) -apply(fastforce) -apply(meson WTrtFAcc has_visible_field) -apply(meson WTrtSFAcc has_visible_field) -apply(meson WTrtFAss has_visible_field) -apply(meson WTrtSFAss has_visible_field) -apply(fastforce simp: WTrtCall) -apply(fastforce simp: WTrtSCall) -apply(fastforce) -apply(fastforce) -apply(fastforce simp: WTrtCond) -apply(fastforce) -apply(fastforce) -apply(fastforce) -apply(simp) -apply(simp) -done +proof(induct rule: WT_WTs_inducts) + case WTVal then show ?case by(fastforce dest: typeof_lit_typeof) +next + case WTFAcc + then show ?case by(fastforce simp: WTrtFAcc has_visible_field) +next + case WTSFAcc + then show ?case by(fastforce simp: WTrtSFAcc has_visible_field) +next + case WTFAss + then show ?case by(fastforce simp: WTrtFAss dest: has_visible_field) +next + case WTSFAss + then show ?case by(fastforce simp: WTrtSFAss dest: has_visible_field) +qed(fastforce intro: WTrt_WTrts.intros)+ (*>*) end