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,460 +1,460 @@ (* 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 "../DFA/Typing_Framework_2" 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" context start_context begin lemma Cons_less_Conss3 [simp]: "x#xs [\\<^bsub>r\<^esub>] y#ys = (x \\<^bsub>r\<^esub> y \ xs [\\<^bsub>r\<^esub>] ys \ x = y \ xs [\\<^bsub>r\<^esub>] ys)" apply (unfold lesssub_def ) apply auto apply (insert sup_state_opt_err) apply (unfold lesssub_def lesub_def sup_state_opt_def sup_state_def sup_ty_opt_def) apply (simp only: JVM_le_unfold ) apply fastforce done lemma acc_le_listI3 [intro!]: " acc r \ acc (Listn.le r)" apply (unfold acc_def) apply (subgoal_tac "wf(UN n. {(ys,xs). size xs = n \ size ys = n \ xs <_(Listn.le r) ys})") apply (erule wf_subset) apply (blast intro: lesssub_lengthD) apply (rule wf_UN) prefer 2 apply (rename_tac m n) apply (case_tac "m=n") apply simp apply (fast intro!: equals0I dest: not_sym) apply (rename_tac n) apply (induct_tac n) apply (simp add: lesssub_def cong: conj_cong) apply (rename_tac k) apply (simp add: wf_eq_minimal del: r_def f_def step_def A_def) apply (simp (no_asm) add: length_Suc_conv cong: conj_cong del: r_def f_def step_def A_def) apply clarify apply (rename_tac M m) apply (case_tac "\x xs. size xs = k \ x#xs \ M") prefer 2 apply (erule thin_rl) apply (erule thin_rl) apply blast apply (erule_tac x = "{a. \xs. size xs = k \ a#xs:M}" in allE) apply (erule impE) apply blast apply (thin_tac "\x xs. P x xs" for P) apply clarify apply (rename_tac maxA xs) apply (erule_tac x = "{ys. size ys = size xs \ maxA#ys \ M}" in allE) apply (erule impE) apply blast apply clarify apply (thin_tac "m \ M") apply (thin_tac "maxA#xs \ M") apply (rule bexI) prefer 2 apply assumption apply clarify apply (simp del: r_def f_def step_def A_def) apply blast done lemma wf_jvm: " wf {(ss', ss). ss [\\<^bsub>r\<^esub>] ss'}" apply (insert acc_le_listI3 acc_JVM [OF wf]) by (simp add: acc_def r_def) lemma iter_properties_bv[rule_format]: -shows "\ \p\w0. p < n; ss0 \ list n A; \p w0 \ stable r step ss0 p \ \ +shows "\ \p\w0. p < n; ss0 \ nlists n A; \p w0 \ stable r step ss0 p \ \ iter f step ss0 w0 = (ss',w') \ - ss' \ list n A \ stables r step ss' \ ss0 [\\<^sub>r] ss' \ - (\ts\list n A. ss0 [\\<^sub>r] ts \ stables r step ts \ ss' [\\<^sub>r] ts)" + ss' \ nlists n A \ stables r step ss' \ ss0 [\\<^sub>r] ss' \ + (\ts\nlists n A. ss0 [\\<^sub>r] ts \ stables r step ts \ ss' [\\<^sub>r] ts)" (*<*) (is "PROP ?P") proof - show "PROP ?P" apply (insert semi bounded_step exec_pres_type step_mono[OF wf]) apply (unfold iter_def stables_def) apply (rule_tac P = "\(ss,w). - ss \ list n A \ (\p w \ stable r step ss p) \ ss0 [\\<^sub>r] ss \ - (\ts\list n A. ss0 [\\<^sub>r] ts \ stables r step ts \ ss [\\<^sub>r] ts) \ + ss \ nlists n A \ (\p w \ stable r step ss p) \ ss0 [\\<^sub>r] ss \ + (\ts\nlists n A. ss0 [\\<^sub>r] ts \ stables r step ts \ ss [\\<^sub>r] ts) \ (\p\w. p < n)" and r = "{(ss',ss) . ss [\\<^sub>r] ss'} <*lex*> finite_psubset" in while_rule) \ \Invariant holds initially:\ apply (simp add:stables_def semilat_Def del: n_def A_def r_def f_def step_def) apply (blast intro:le_list_refl') \ \Invariant is preserved:\ apply(simp add: stables_def split_paired_all del: A_def r_def f_def step_def n_def) apply(rename_tac ss w) apply(subgoal_tac "(SOME p. p \ w) \ w") prefer 2 apply (fast intro: someI) apply(subgoal_tac "\(q,t) \ set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length ss \ t \ A") prefer 2 apply clarify apply (rule conjI) apply(clarsimp, blast dest!: boundedD) apply (subgoal_tac "(SOME p. p \ w) < n") apply (subgoal_tac "(ss ! (SOME p. p \ w)) \ A") apply (fastforce simp only:n_def dest:pres_typeD ) apply simp apply simp apply (subst decomp_propa) apply blast apply (simp del:A_def r_def f_def step_def n_def) apply (rule conjI) apply (rule Semilat.merges_preserves_type[OF Semilat.intro, OF semi]) apply blast apply clarify apply (rule conjI) apply(clarsimp, blast dest!: boundedD) apply (erule pres_typeD) prefer 3 apply assumption - apply (erule listE_nth_in) + apply (erule nlistsE_nth_in) apply blast apply (simp only:n_def) apply (rule conjI) apply clarify - apply (subgoal_tac "ss \ list (length is) A" "\p\w. p < (length is) " "\p<(length is). p \ w \ stable r step ss p " + apply (subgoal_tac "ss \ nlists (length is) A" "\p\w. p < (length is) " "\p<(length is). p \ w \ stable r step ss p " "p < length is") apply (blast intro!: Semilat.stable_pres_lemma[OF Semilat.intro, OF semi]) apply (simp only:n_def) apply (simp only:n_def) apply (simp only:n_def) apply (simp only:n_def) apply (rule conjI) - apply (subgoal_tac "ss \ list (length is) A" + apply (subgoal_tac "ss \ nlists (length is) A" "\(q,t)\set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length is \ t \ A" - "ss [\\<^bsub>r\<^esub>] merges f (step (SOME p. p \ w) (ss ! (SOME p. p \ w))) ss" "ss0\ list (size is) A" - "merges f (step (SOME p. p \ w) (ss ! (SOME p. p \ w))) ss \ list (size is) A" - "ss \list (size is) A" "order r A" "ss0 [\\<^bsub>r\<^esub>] ss ") + "ss [\\<^bsub>r\<^esub>] merges f (step (SOME p. p \ w) (ss ! (SOME p. p \ w))) ss" "ss0\ nlists (size is) A" + "merges f (step (SOME p. p \ w) (ss ! (SOME p. p \ w))) ss \ nlists (size is) A" + "ss \nlists (size is) A" "order r A" "ss0 [\\<^bsub>r\<^esub>] ss ") apply (blast dest: le_list_trans) apply simp apply (simp only:semilat_Def) apply (simp only:n_def) apply (fastforce simp only: n_def dest:Semilat.merges_preserves_type[OF Semilat.intro, OF semi] ) apply (simp only:n_def) apply (blast intro:Semilat.merges_incr[OF Semilat.intro, OF semi]) apply (subgoal_tac "length ss = n") apply (simp only:n_def) - apply (subgoal_tac "ss \list n A") + apply (subgoal_tac "ss \nlists n A") defer apply simp apply (simp only:n_def) prefer 5 - apply (simp only:listE_length n_def) + apply (simp only:nlistsE_length n_def) apply(rule conjI) apply (clarsimp simp del: A_def r_def f_def step_def) apply (blast intro!: Semilat.merges_bounded_lemma[OF Semilat.intro, OF semi]) apply (subgoal_tac "bounded step n") apply (blast dest!: boundedD) apply (simp only:n_def) \ \Postcondition holds upon termination:\ apply(clarsimp simp add: stables_def split_paired_all) \ \Well-foundedness of the termination relation:\ apply (rule wf_lex_prod) apply (simp only:wf_jvm) apply (rule wf_finite_psubset) \ \Loop decreases along termination relation:\ apply(simp add: stables_def split_paired_all del: A_def r_def f_def step_def) apply(rename_tac ss w) apply(subgoal_tac "(SOME p. p \ w) \ w") prefer 2 apply (fast intro: someI) apply(subgoal_tac "\(q,t) \ set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length ss \ t \ A") prefer 2 apply clarify apply (rule conjI) apply(clarsimp, blast dest!: boundedD) apply (erule pres_typeD) prefer 3 apply assumption - apply (erule listE_nth_in) + apply (erule nlistsE_nth_in) apply blast apply blast apply (subst decomp_propa) apply blast apply clarify - apply (simp del: listE_length A_def r_def f_def step_def + apply (simp del: nlistsE_length A_def r_def f_def step_def add: lex_prod_def finite_psubset_def bounded_nat_set_is_finite) apply (rule termination_lemma) apply (insert Semilat.intro) apply assumption+ defer apply assumption defer apply clarsimp done qed (*>*) lemma kildall_properties_bv: -shows "\ ss0 \ list n A \ \ - kildall r f step ss0 \ list n A \ +shows "\ ss0 \ nlists n A \ \ + kildall r f step ss0 \ nlists n A \ stables r step (kildall r f step ss0) \ ss0 [\\<^sub>r] kildall r f step ss0 \ - (\ts\list n A. ss0 [\\<^sub>r] ts \ stables r step ts \ + (\ts\nlists n A. ss0 [\\<^sub>r] ts \ stables r step ts \ kildall r f step ss0 [\\<^sub>r] ts)" (*<*) (is "PROP ?P") proof - show "PROP ?P" apply (unfold kildall_def) apply(case_tac "iter f step ss0 (unstables r step ss0)") apply (simp del:r_def f_def n_def step_def A_def) apply (rule iter_properties_bv) apply (simp_all add: unstables_def stable_def) done qed end lemma (in start_context) is_bcv_kiljvm: shows "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 n_def) apply(unfold is_bcv_def wt_step_def) apply(insert semi kildall_properties_bv) apply(simp only:stables_def) apply clarify - apply(subgoal_tac "kildall r f step \s\<^sub>0 \ list n A") + apply(subgoal_tac "kildall r f step \s\<^sub>0 \ nlists n A") prefer 2 apply (fastforce intro: kildall_properties_bv) apply (rule iffI) apply (rule_tac x = "kildall r f step \s\<^sub>0" in bexI) apply (rule conjI) apply (fastforce intro: kildall_properties_bv) apply (force intro: kildall_properties_bv) apply simp apply clarify apply(subgoal_tac "kildall r f step \s\<^sub>0!pa <=_r \s!pa") defer apply (blast intro!: le_listD less_lengthI) apply (subgoal_tac "\s!pa \ Err") defer apply simp apply (rule ccontr) apply (simp only:top_def r_def JVM_le_unfold) apply fastforce done (* 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" + "0 < size is \ start \ nlists (size is) A" using Ts C (*<*) apply (simp add: JVM_states_unfold) - apply (force intro!: listI list_appendI dest!: in_set_replicate) + apply (force intro!: nlistsI nlists_appendI dest!: in_set_replicate) done (*>*) 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" .. + from instrs have "start \ nlists (size is) A" .. with bcv success result have - "\ts\list (size is) A. start [\\<^sub>r] ts \ wt_step r Err step ts" + "\ts\nlists (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 + in_A: "\s' \ nlists (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 from instrs have "start \ nlists (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) + have "?\s \ nlists (size is) A" + by (auto intro!: nlistsI 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 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" \ \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 qed (*>*) end diff --git a/thys/Jinja/BV/JVM_SemiType.thy b/thys/Jinja/BV/JVM_SemiType.thy --- a/thys/Jinja/BV/JVM_SemiType.thy +++ b/thys/Jinja/BV/JVM_SemiType.thy @@ -1,375 +1,375 @@ (* Title: HOL/MicroJava/BV/JVM.thy Author: Gerwin Klein Copyright 2000 TUM *) section \The JVM Type System as Semilattice\ theory JVM_SemiType imports SemiType begin type_synonym ty\<^sub>l = "ty err list" type_synonym ty\<^sub>s = "ty list" type_synonym ty\<^sub>i = "ty\<^sub>s \ ty\<^sub>l" type_synonym ty\<^sub>i' = "ty\<^sub>i option" type_synonym ty\<^sub>m = "ty\<^sub>i' list" type_synonym ty\<^sub>P = "mname \ cname \ ty\<^sub>m" definition stk_esl :: "'c prog \ nat \ ty\<^sub>s esl" where "stk_esl P mxs \ upto_esl mxs (SemiType.esl P)" definition loc_sl :: "'c prog \ nat \ ty\<^sub>l sl" where "loc_sl P mxl \ Listn.sl mxl (Err.sl (SemiType.esl P))" definition sl :: "'c prog \ nat \ nat \ ty\<^sub>i' err sl" where "sl P mxs mxl \ Err.sl(Opt.esl(Product.esl (stk_esl P mxs) (Err.esl(loc_sl P mxl))))" definition states :: "'c prog \ nat \ nat \ ty\<^sub>i' err set" where "states P mxs mxl \ fst(sl P mxs mxl)" definition le :: "'c prog \ nat \ nat \ ty\<^sub>i' err ord" where "le P mxs mxl \ fst(snd(sl P mxs mxl))" definition sup :: "'c prog \ nat \ nat \ ty\<^sub>i' err binop" where "sup P mxs mxl \ snd(snd(sl P mxs mxl))" definition sup_ty_opt :: "['c prog,ty err,ty err] \ bool" ("_ \ _ \\<^sub>\ _" [71,71,71] 70) where "sup_ty_opt P \ Err.le (subtype P)" definition sup_state :: "['c prog,ty\<^sub>i,ty\<^sub>i] \ bool" ("_ \ _ \\<^sub>i _" [71,71,71] 70) where "sup_state P \ Product.le (Listn.le (subtype P)) (Listn.le (sup_ty_opt P))" definition sup_state_opt :: "['c prog,ty\<^sub>i',ty\<^sub>i'] \ bool" ("_ \ _ \'' _" [71,71,71] 70) where "sup_state_opt P \ Opt.le (sup_state P)" abbreviation sup_loc :: "['c prog,ty\<^sub>l,ty\<^sub>l] \ bool" ("_ \ _ [\\<^sub>\] _" [71,71,71] 70) where "P \ LT [\\<^sub>\] LT' \ list_all2 (sup_ty_opt P) LT LT'" notation (ASCII) sup_ty_opt ("_ |- _ <=T _" [71,71,71] 70) and sup_state ("_ |- _ <=i _" [71,71,71] 70) and sup_state_opt ("_ |- _ <=' _" [71,71,71] 70) and sup_loc ("_ |- _ [<=T] _" [71,71,71] 70) subsection "Unfolding" lemma JVM_states_unfold: - "states P mxs mxl \ err(opt((Union {list n (types P) |n. n <= mxs}) \ - list mxl (err(types P))))" + "states P mxs mxl \ err(opt((Union {nlists n (types P) |n. n <= mxs}) \ + nlists mxl (err(types P))))" (*<*) apply (unfold states_def sl_def Opt.esl_def Err.sl_def stk_esl_def loc_sl_def Product.esl_def Listn.sl_def upto_esl_def SemiType.esl_def Err.esl_def) apply simp done (*>*) lemma JVM_le_unfold: "le P m n \ Err.le(Opt.le(Product.le(Listn.le(subtype P))(Listn.le(Err.le(subtype P)))))" (*<*) apply (unfold le_def sl_def Opt.esl_def Err.sl_def stk_esl_def loc_sl_def Product.esl_def Listn.sl_def upto_esl_def SemiType.esl_def Err.esl_def) apply simp done (*>*) lemma sl_def2: "JVM_SemiType.sl P mxs mxl \ (states P mxs mxl, JVM_SemiType.le P mxs mxl, JVM_SemiType.sup P mxs mxl)" (*<*) by (unfold JVM_SemiType.sup_def states_def JVM_SemiType.le_def) simp (*>*) lemma JVM_le_conv: "le P m n (OK t1) (OK t2) = P \ t1 \' t2" (*<*) by (simp add: JVM_le_unfold Err.le_def lesub_def sup_state_opt_def sup_state_def sup_ty_opt_def) (*>*) lemma JVM_le_Err_conv: "le P m n = Err.le (sup_state_opt P)" (*<*) by (unfold sup_state_opt_def sup_state_def sup_ty_opt_def JVM_le_unfold) simp (*>*) lemma err_le_unfold [iff]: "Err.le r (OK a) (OK b) = r a b" (*<*) by (simp add: Err.le_def lesub_def) (*>*) subsection \Semilattice\ lemma order_sup_state_opt' [intro, simp]: "wf_prog wf_mb P \ - order (sup_state_opt P) (opt ((\ {list n (types P) |n. n \ mxs} ) \ list (Suc (length Ts + mxl\<^sub>0)) (err (types P))))" + order (sup_state_opt P) (opt ((\ {nlists n (types P) |n. n \ mxs} ) \ nlists (Suc (length Ts + mxl\<^sub>0)) (err (types P))))" (*<*) apply (unfold sup_state_opt_def sup_state_def sup_ty_opt_def ) apply (blast intro:order_le_prodI) \\ use Listn.thy.order_listI2 \ done (*<*) lemma semilat_JVM [intro?]: "wf_prog wf_mb P \ semilat (JVM_SemiType.sl P mxs mxl)" (*<*) apply (unfold JVM_SemiType.sl_def stk_esl_def loc_sl_def) apply (blast intro: err_semilat_Product_esl err_semilat_upto_esl Listn_sl err_semilat_JType_esl) done (*>*) subsection \Widening with \\\\ lemma subtype_refl: "subtype P t t" (*<*) by (simp add: fun_of_def) (*>*) lemma sup_ty_opt_refl [iff]: "P \ T \\<^sub>\ T" (*<*) apply (unfold sup_ty_opt_def) apply (fold lesub_def) apply (rule le_err_refl) apply (simp add: lesub_def) done (*>*) lemma Err_any_conv [iff]: "P \ Err \\<^sub>\ T = (T = Err)" (*<*) by (unfold sup_ty_opt_def) (rule Err_le_conv [simplified lesub_def]) (*>*) lemma any_Err [iff]: "P \ T \\<^sub>\ Err" (*<*) by (unfold sup_ty_opt_def) (rule le_Err [simplified lesub_def]) (*>*) lemma OK_OK_conv [iff]: "P \ OK T \\<^sub>\ OK T' = P \ T \ T'" (*<*) by (simp add: sup_ty_opt_def fun_of_def) (*>*) lemma any_OK_conv [iff]: "P \ X \\<^sub>\ OK T' = (\T. X = OK T \ P \ T \ T')" (*<*) apply (unfold sup_ty_opt_def) apply (rule le_OK_conv [simplified lesub_def]) done (*>*) lemma OK_any_conv: "P \ OK T \\<^sub>\ X = (X = Err \ (\T'. X = OK T' \ P \ T \ T'))" (*<*) apply (unfold sup_ty_opt_def) apply (rule OK_le_conv [simplified lesub_def]) done (*>*) lemma sup_ty_opt_trans [intro?, trans]: "\P \ a \\<^sub>\ b; P \ b \\<^sub>\ c\ \ P \ a \\<^sub>\ c" (*<*) by (auto intro: widen_trans simp add: sup_ty_opt_def Err.le_def lesub_def fun_of_def split: err.splits) (*>*) subsection "Stack and Registers" lemma stk_convert: "P \ ST [\] ST' = Listn.le (subtype P) ST ST'" (*<*) by (simp add: Listn.le_def lesub_def) (*>*) lemma sup_loc_refl [iff]: "P \ LT [\\<^sub>\] LT" (*<*) by (rule list_all2_refl) simp (*>*) lemmas sup_loc_Cons1 [iff] = list_all2_Cons1 [of "sup_ty_opt P"] for P lemma sup_loc_def: "P \ LT [\\<^sub>\] LT' \ Listn.le (sup_ty_opt P) LT LT'" (*<*) by (simp add: Listn.le_def lesub_def) (*>*) lemma sup_loc_widens_conv [iff]: "P \ map OK Ts [\\<^sub>\] map OK Ts' = P \ Ts [\] Ts'" (*<*) by (simp add: list_all2_map1 list_all2_map2) (*>*) lemma sup_loc_trans [intro?, trans]: "\P \ a [\\<^sub>\] b; P \ b [\\<^sub>\] c\ \ P \ a [\\<^sub>\] c" (*<*) by (rule list_all2_trans, rule sup_ty_opt_trans) (*>*) subsection "State Type" lemma sup_state_conv [iff]: "P \ (ST,LT) \\<^sub>i (ST',LT') = (P \ ST [\] ST' \ P \ LT [\\<^sub>\] LT')" (*<*) by (auto simp add: sup_state_def stk_convert lesub_def Product.le_def sup_loc_def) (*>*) lemma sup_state_conv2: "P \ s1 \\<^sub>i s2 = (P \ fst s1 [\] fst s2 \ P \ snd s1 [\\<^sub>\] snd s2)" (*<*) by (cases s1, cases s2) simp (*>*) lemma sup_state_refl [iff]: "P \ s \\<^sub>i s" (*<*) by (auto simp add: sup_state_conv2) (*>*) lemma sup_state_trans [intro?, trans]: "\P \ a \\<^sub>i b; P \ b \\<^sub>i c\ \ P \ a \\<^sub>i c" (*<*) by (auto intro: sup_loc_trans widens_trans simp add: sup_state_conv2) (*>*) lemma sup_state_opt_None_any [iff]: "P \ None \' s" (*<*) by (simp add: sup_state_opt_def Opt.le_def) (*>*) lemma sup_state_opt_any_None [iff]: "P \ s \' None = (s = None)" (*<*) by (simp add: sup_state_opt_def Opt.le_def) (*>*) lemma sup_state_opt_Some_Some [iff]: "P \ Some a \' Some b = P \ a \\<^sub>i b" (*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*) lemma sup_state_opt_any_Some: "P \ (Some s) \' X = (\s'. X = Some s' \ P \ s \\<^sub>i s')" (*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*) lemma sup_state_opt_refl [iff]: "P \ s \' s" (*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*) lemma sup_state_opt_trans [intro?, trans]: "\P \ a \' b; P \ b \' c\ \ P \ a \' c" (*<*) apply (unfold sup_state_opt_def Opt.le_def lesub_def) apply (simp del: split_paired_All) apply (rule sup_state_trans, assumption+) done (*>*) lemma sup_state_opt_err : "(Err.le (sup_state_opt P)) s s" apply (unfold JVM_le_unfold Product.le_def Opt.le_def Err.le_def lesssub_def lesub_def Listn.le_def) apply (auto split: err.splits) done lemma Cons_less_Conss1 [simp]: "x#xs [\\<^bsub>subtype P\<^esub>] y#ys = (x \\<^bsub>subtype P\<^esub> y \ xs [\\<^bsub>subtype P\<^esub>] ys \ x = y \ xs [\\<^bsub>subtype P\<^esub>] ys)" apply (unfold lesssub_def ) apply auto apply (simp add:lesssub_def lesub_def) \\widen_refl, subtype_refl \ done lemma Cons_less_Conss2 [simp]: "x#xs [\\<^bsub>Err.le (subtype P)\<^esub>] y#ys = (x \\<^bsub>Err.le (subtype P)\<^esub> y \ xs [\\<^bsub>Err.le (subtype P)\<^esub>] ys \ x = y \ xs [\\<^bsub>Err.le (subtype P)\<^esub>] ys)" apply (unfold lesssub_def ) apply auto apply (simp add:lesssub_def lesub_def Err.le_def split: err.splits) done lemma acc_le_listI1 [intro!]: " acc (subtype P) \ acc (Listn.le (subtype P))" (*<*) apply (unfold acc_def) apply (subgoal_tac "wf(UN n. {(ys,xs). size xs = n \ size ys = n \ xs <_(Listn.le (subtype P)) ys})") apply (erule wf_subset) apply (blast intro: lesssub_lengthD) apply (rule wf_UN) prefer 2 apply (rename_tac m n) apply (case_tac "m=n") apply simp apply (fast intro!: equals0I dest: not_sym) apply (rename_tac n) apply (induct_tac n) apply (simp add: lesssub_def cong: conj_cong) apply (rename_tac k) apply (simp add: wf_eq_minimal) apply (simp (no_asm) add: length_Suc_conv cong: conj_cong) apply clarify apply (rename_tac M m) apply (case_tac "\x xs. size xs = k \ x#xs \ M") prefer 2 apply (erule thin_rl) apply (erule thin_rl) apply blast apply (erule_tac x = "{a. \xs. size xs = k \ a#xs:M}" in allE) apply (erule impE) apply blast apply (thin_tac "\x xs. P x xs" for P) apply clarify apply (rename_tac maxA xs) apply (erule_tac x = "{ys. size ys = size xs \ maxA#ys \ M}" in allE) apply (erule impE) apply blast apply clarify apply (thin_tac "m \ M") apply (thin_tac "maxA#xs \ M") apply (rule bexI) prefer 2 apply assumption apply clarify apply simp apply blast done lemma acc_le_listI2 [intro!]: " acc (Err.le (subtype P)) \ acc (Listn.le (Err.le (subtype P)))" (*<*) apply (unfold acc_def) apply (subgoal_tac "wf(UN n. {(ys,xs). size xs = n \ size ys = n \ xs <_(Listn.le (Err.le (subtype P))) ys})") apply (erule wf_subset) apply (blast intro: lesssub_lengthD) apply (rule wf_UN) prefer 2 apply (rename_tac m n) apply (case_tac "m=n") apply simp apply (fast intro!: equals0I dest: not_sym) apply (rename_tac n) apply (induct_tac n) apply (simp add: lesssub_def cong: conj_cong) apply (rename_tac k) apply (simp add: wf_eq_minimal) apply (simp (no_asm) add: length_Suc_conv cong: conj_cong) apply clarify apply (rename_tac M m) apply (case_tac "\x xs. size xs = k \ x#xs \ M") prefer 2 apply (erule thin_rl) apply (erule thin_rl) apply blast apply (erule_tac x = "{a. \xs. size xs = k \ a#xs:M}" in allE) apply (erule impE) apply blast apply (thin_tac "\x xs. P x xs" for P) apply clarify apply (rename_tac maxA xs) apply (erule_tac x = "{ys. size ys = size xs \ maxA#ys \ M}" in allE) apply (erule impE) apply blast apply clarify apply (thin_tac "m \ M") apply (thin_tac "maxA#xs \ M") apply (rule bexI) prefer 2 apply assumption apply clarify apply simp apply blast done lemma acc_JVM [intro]: "wf_prog wf_mb P \ acc (JVM_SemiType.le P mxs mxl)" (*<*) by (unfold JVM_le_unfold) blast (*>*) \\ use acc_listI1, acc_listI2 \ 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,223 @@ (* 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" + shows "\\s \ nlists (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 + list: "\s \ nlists (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 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" 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 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 (*>*) 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,278 +1,278 @@ (* 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 and n 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" defines [simp]: "n \ size is" 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 start_context) semi: "semilat (A, r, f)" apply (insert semilat_JVM[OF wf]) apply (unfold A_def r_def f_def JVM_SemiType.le_def JVM_SemiType.sup_def states_def) apply auto done 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 ex_in_nlists [iff]: + "(\n. ST \ nlists n A \ n \ mxs) = (set ST \ A \ size ST \ mxs)" + by (unfold nlists_def) auto -lemma singleton_list: - "(\n. [Class C] \ list n (types P) \ n \ mxs) = (is_class P C \ 0 < mxs)" +lemma singleton_nlists: + "(\n. [Class C] \ nlists 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 +lemma in_nlistsE: + "\ xs \ nlists n A; \size xs = n; set xs \ A\ \ P \ \ P" + by (unfold nlists_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 (frule nlistsE_nth_in, assumption) apply fastforce \ \Store\ apply fastforce \ \Push\ apply (fastforce simp add: typeof_lit_is_type) \ \New\ apply clarsimp apply fastforce \ \Getfield\ apply clarsimp apply (fastforce dest: sees_field_is_type) \ \Putfield\ apply clarsimp apply fastforce \ \Checkcast\ apply clarsimp apply fastforce defer \ \Return\ apply fastforce \ \Pop\ apply fastforce \ \IAdd\ apply fastforce \ \Goto\ apply fastforce \ \CmpEq\ apply fastforce \ \IfFalse\ apply fastforce \ \Throw\ apply clarsimp 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 (*>*) 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 (*>*) 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)" (*<*) 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 (*>*) 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 (*>*) 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 \\ order_sup_state_opt' \ 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 (*>*) lemma (in start_context) first_in_A [iff]: "OK first \ A" - using Ts C by (force intro!: list_appendI simp add: JVM_states_unfold) + using Ts C by (force intro!: nlists_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 (*>*) 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,1489 +1,1489 @@ (* 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" (*<*)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)(*>*) + "set E \ types P \ ty\<^sub>l E A \ nlists mxl (err (types P))" +(*<*)by(auto simp add:ty\<^sub>l_def intro!:nlistsI 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)" (*<*) 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: 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" (*<*) 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) + (blast intro!:nlistsI 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" (*<*) by(simp add:ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types) - (blast intro!:listI) + (blast intro!:nlistsI) (*>*) 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}" (*<*) 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" (*<*) 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)" (*<*) 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)" (*<*) 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 \))" (*<*)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: 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" (*<*) 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: 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" (*<*) 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: 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" (*<*) 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: 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" (*<*) 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]: 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" (*<*) 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]: 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" (*<*) 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]: assumes wtis: "\ is,xt [::] \s@[\]" and ss: "P \ \ \' \'" shows "\ is,xt [::] \s@[\']" (*<*) 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]: 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" proof(induct n) case (Suc n) then show ?case by(simp add: drop_Suc) qed simp lemma drop_mess: 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, ?\']" proof-{ show ?thesis apply (rule wt_IfFalse) apply simp apply simp apply (subgoal_tac "length (compE\<^sub>2 c) = length (compT E (A \ \ e) ST c) + 1" "nat (int (length (compE\<^sub>2 c)) + 3 - 2) > length (compT E (A \ \ e) ST c)") apply simp+ done } qed 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" proof-{ let ?t1 = "ty\<^sub>i' ST E A" let ?t2 = "ty\<^sub>i' (Boolean # ST) E (A \ \ e)" let ?t3 = "ty\<^sub>i' ST E (A \ \ e)" let ?t4 = "[ty\<^sub>i' (Tc # ST) E (A \ \ e \ \ c), ty\<^sub>i' ST E (A \ \ e \ \ c), ty\<^sub>i' ST E (A \ \ e), ty\<^sub>i' (Void # ST) E (A \ \ e)]" let ?c = " compT E (A \ \ e) ST c" let ?e = "compT E A ST e" assume ass: "P \ ty\<^sub>i' ST E (A \ \ e \ \ c) \' ?t1" show ?thesis apply (rule wt_Goto) apply simp apply simp apply simp proof-{ let ?s1 = "((?t1 # ?e @ [?t2]) @ ?t3 # ?c)" have "length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2 > length ?s1" by auto hence "(?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) = ?t4 ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2 - length ?s1)" by (auto simp only:nth_append split: if_splits) hence "(?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) = ty\<^sub>i' ST E (A \ \ e \ \ c)" using compT_sizes by auto hence "P \ (?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) \' ty\<^sub>i' ST E A" using ass by simp thus "P \ ((?t1 # ?e @ [?t2]) @ ?t3 # ?c @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) \' ((?t1 # ?e @ [?t2]) @ ?t3 # ?c @ ?t4) ! nat (int (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) + - int (length (compE\<^sub>2 c) + length (compE\<^sub>2 e) + 2))" by simp }qed }qed 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" proof-{ show ?thesis apply clarsimp apply (rule wt_IfFalse) apply (simp only:nat_add_distrib) apply simp apply (auto simp only:nth_append split:if_splits) apply (simp only:compT_sizes) apply simp done }qed 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, \)" (*<*)(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" (*<*) 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 \" (*<*) by (simp add: app_def xcpt_app_def eff_def xcpt_eff_def norm_eff_def) (fastforce simp add: image_def) (*>*) lemma [simp]: assumes "app i P mpc T pc mxl xt \" shows "eff i (compP f P) pc xt \ = eff i P pc xt \" (*<*) 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" (*<*)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 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)" (*<*)(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: assumes wf: "wf_J\<^sub>1_prog P" shows "wf_jvm_prog (compP\<^sub>2 P)" (*<*) 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)" (*<*) by(simp only:o_def J2JVM_def) (blast intro:wt_compP\<^sub>2 compP\<^sub>1_pres_wf) end diff --git a/thys/Jinja/DFA/Kildall_1.thy b/thys/Jinja/DFA/Kildall_1.thy --- a/thys/Jinja/DFA/Kildall_1.thy +++ b/thys/Jinja/DFA/Kildall_1.thy @@ -1,143 +1,116 @@ (* Title: HOL/MicroJava/BV/Kildall.thy Author: Tobias Nipkow, Gerwin Klein Copyright 2000 TUM Kildall's algorithm. *) section \Kildall's Algorithm \label{sec:Kildall}\ theory Kildall_1 imports SemilatAlg begin - - - - - - - - - - - - - - - - - - - - - - - - - - - primrec merges :: "'s binop \ (nat \ 's) list \ 's list \ 's list" where "merges f [] \s = \s" | "merges f (p'#ps) \s = (let (p,\) = p' in merges f ps (\s[p := \ \\<^bsub>f\<^esub> \s!p]))" lemmas [simp] = Let_def Semilat.le_iff_plus_unchanged [OF Semilat.intro, symmetric] lemma (in Semilat) nth_merges: - "\ss. \p < length ss; ss \ list n A; \(p,t)\set ps. p t\A \ \ + "\ss. \p < length ss; ss \ nlists n A; \(p,t)\set ps. p t\A \ \ (merges f ps ss)!p = map snd [(p',t') \ ps. p'=p] \\<^bsub>f\<^esub> ss!p" (is "\ss. \_; _; ?steptype ps\ \ ?P ss ps") (*<*) proof (induct ps) show "\ss. ?P ss []" by simp fix ss p' ps' - assume ss: "ss \ list n A" + assume ss: "ss \ nlists n A" assume l: "p < length ss" assume "?steptype (p'#ps')" then obtain a b where p': "p'=(a,b)" and ab: "aA" and ps': "?steptype ps'" by (cases p') auto - assume "\ss. p< length ss \ ss \ list n A \ ?steptype ps' \ ?P ss ps'" - hence IH: "\ss. ss \ list n A \ p < length ss \ ?P ss ps'" using ps' by iprover + assume "\ss. p< length ss \ ss \ nlists n A \ ?steptype ps' \ ?P ss ps'" + hence IH: "\ss. ss \ nlists n A \ p < length ss \ ?P ss ps'" using ps' by iprover from ss ab - have "ss[a := b \\<^bsub>f\<^esub> ss!a] \ list n A" by (simp add: closedD) + have "ss[a := b \\<^bsub>f\<^esub> ss!a] \ nlists n A" by (simp add: closedD) moreover with l have "p < length (ss[a := b \\<^bsub>f\<^esub> ss!a])" by simp ultimately have "?P (ss[a := b \\<^bsub>f\<^esub> ss!a]) ps'" by (rule IH) with p' l show "?P ss (p'#ps')" by simp qed (*>*) (** merges **) lemma length_merges [simp]: "\ss. size(merges f ps ss) = size ss" (*<*) by (induct ps, auto) (*>*) lemma (in Semilat) merges_preserves_type_lemma: -shows "\xs. xs \ list n A \ (\(p,x) \ set ps. p x\A) - \ merges f ps xs \ list n A" +shows "\xs. xs \ nlists n A \ (\(p,x) \ set ps. p x\A) + \ merges f ps xs \ nlists n A" (*<*) apply (insert closedI) apply (unfold closed_def) apply (induct ps) apply simp apply clarsimp done (*>*) lemma (in Semilat) merges_preserves_type [simp]: - "\ xs \ list n A; \(p,x) \ set ps. p x\A \ - \ merges f ps xs \ list n A" + "\ xs \ nlists n A; \(p,x) \ set ps. p x\A \ + \ merges f ps xs \ nlists n A" by (simp add: merges_preserves_type_lemma) lemma (in Semilat) list_update_le_listI [rule_format]: "set xs \ A \ set ys \ A \ xs [\\<^bsub>r\<^esub>] ys \ p < size xs \ x \\<^bsub>r\<^esub> ys!p \ x\A \ xs[p := x \\<^bsub>f\<^esub> xs!p] [\\<^bsub>r\<^esub>] ys" (*<*) apply (insert semilat) apply (simp only: Listn.le_def lesub_def semilat_def) apply (simp add: list_all2_conv_all_nth nth_list_update) done (*>*) lemma (in Semilat) merges_pres_le_ub: assumes "set ts \ A" "set ss \ A" "\(p,t)\set ps. t \\<^bsub>r\<^esub> ts!p \ t \ A \ p < size ts" "ss [\\<^bsub>r\<^esub>] ts" shows "merges f ps ss [\\<^bsub>r\<^esub>] ts" (*<*) proof - { fix t ts ps have "\qs. \set ts \ A; \(p,t)\set ps. t \\<^bsub>r\<^esub> ts!p \ t \ A \ p< size ts \ \ set qs \ set ps \ (\ss. set ss \ A \ ss [\\<^bsub>r\<^esub>] ts \ merges f qs ss [\\<^bsub>r\<^esub>] ts)" apply (induct_tac qs) apply simp apply (simp (no_asm_simp)) apply clarify apply simp apply (erule allE, erule impE, erule_tac [2] mp) apply (drule bspec, assumption) apply (simp add: closedD) apply (drule bspec, assumption) apply (simp add: list_update_le_listI) done } note this [dest] from assms show ?thesis by blast qed (*>*) end diff --git a/thys/Jinja/DFA/Kildall_2.thy b/thys/Jinja/DFA/Kildall_2.thy --- a/thys/Jinja/DFA/Kildall_2.thy +++ b/thys/Jinja/DFA/Kildall_2.thy @@ -1,288 +1,288 @@ (* Title: HOL/MicroJava/BV/Kildall.thy Author: Tobias Nipkow, Gerwin Klein Copyright 2000 TUM Kildall's algorithm. *) section \Kildall's Algorithm \label{sec:Kildall}\ theory Kildall_2 imports SemilatAlg Kildall_1 begin primrec propa :: "'s binop \ (nat \ 's) list \ 's list \ nat set \ 's list * nat set" where "propa f [] \s w = (\s,w)" | "propa f (q'#qs) \s w = (let (q,\) = q'; u = \ \\<^bsub>f\<^esub> \s!q; w' = (if u = \s!q then w else insert q w) in propa f qs (\s[q := u]) w')" definition iter :: "'s binop \ 's step_type \ 's list \ nat set \ 's list \ nat set" where "iter f step \s w = while (\(\s,w). w \ {}) (\(\s,w). let p = SOME p. p \ w in propa f (step p (\s!p)) \s (w-{p})) (\s,w)" definition unstables :: "'s ord \ 's step_type \ 's list \ nat set" where "unstables r step \s = {p. p < size \s \ \stable r step \s p}" definition kildall :: "'s ord \ 's binop \ 's step_type \ 's list \ 's list" where "kildall r f step \s = fst(iter f step \s (unstables r step \s))" (** propa **) lemma (in Semilat) merges_incr_lemma: - "\xs. xs \ list n A \ (\(p,x)\set ps. p x \ A) \ xs [\\<^bsub>r\<^esub>] merges f ps xs" + "\xs. xs \ nlists n A \ (\(p,x)\set ps. p x \ A) \ xs [\\<^bsub>r\<^esub>] merges f ps xs" apply (induct ps) apply simp apply(insert orderI) apply (fastforce intro:le_list_refl') apply simp apply clarify apply (rule order_trans) defer apply (erule list_update_incr) apply simp+ done (*>*) lemma (in Semilat) merges_incr: - "\ xs \ list n A; \(p,x)\set ps. p x \ A \ + "\ xs \ nlists n A; \(p,x)\set ps. p x \ A \ \ xs [\\<^bsub>r\<^esub>] merges f ps xs" by (simp add: merges_incr_lemma) lemma (in Semilat) merges_same_conv [rule_format]: - "(\xs. xs \ list n A \ (\(p,x)\set ps. p x\A) \ + "(\xs. xs \ nlists n A \ (\(p,x)\set ps. p x\A) \ (merges f ps xs = xs) = (\(p,x)\set ps. x \\<^bsub>r\<^esub> xs!p))" (*<*) apply (induct_tac ps) apply simp apply clarsimp apply (rename_tac p x ps xs) apply (rule iffI) apply (rule context_conjI) apply (subgoal_tac "xs[p := x \\<^bsub>f\<^esub> xs!p] [\\<^bsub>r\<^esub>] xs") apply (fastforce dest!: le_listD) \\apply (force dest!: le_listD simp add: nth_list_update) \ apply (erule subst, rule merges_incr) - apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in]) + apply (blast intro!: nlistsE_set intro: closedD nlistsE_length [THEN nth_in]) apply clarify apply (rule conjI) apply simp apply (blast dest: boundedD) apply blast apply clarify apply (erule allE) apply (erule impE) apply assumption apply (drule bspec) apply assumption apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2]) apply blast apply clarify apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2]) done (*>*) lemma decomp_propa: "\ss w. (\(q,t)\set qs. q < size ss) \ propa f qs ss w = (merges f qs ss, {q. \t.(q,t)\set qs \ t \\<^bsub>f\<^esub> ss!q \ ss!q} \ w)" (*<*) apply (induct qs) apply simp apply (simp (no_asm)) apply clarify apply simp apply (rule conjI) apply blast apply (simp add: nth_list_update) apply blast done (*>*) lemma (in Semilat) stable_pres_lemma: shows "\pres_type step n A; bounded step n; - ss \ list n A; p \ w; \q\w. q < n; + ss \ nlists n A; p \ w; \q\w. q < n; \q. q < n \ q \ w \ stable r step ss q; q < n; \s'. (q,s') \ set (step p (ss!p)) \ s' \\<^bsub>f\<^esub> ss!q = ss!q; q \ w \ q = p \ \ stable r step (merges f (step p (ss!p)) ss) q" (*<*) apply (unfold stable_def) apply (subgoal_tac "\s'. (q,s') \ set (step p (ss!p)) \ s' : A") prefer 2 apply clarify apply (erule pres_typeD) prefer 3 apply assumption - apply (rule listE_nth_in) + apply (rule nlistsE_nth_in) apply assumption apply simp apply simp apply simp apply clarify apply (subst nth_merges) apply simp apply (blast dest: boundedD) apply assumption apply clarify apply (rule conjI) apply (blast dest: boundedD) apply (erule pres_typeD) prefer 3 apply assumption apply simp apply simp apply(subgoal_tac "q < length ss") prefer 2 apply simp apply (frule nth_merges [of q _ _ "step p (ss!p)"]) (* fixme: why does method subst not work?? *) apply assumption apply clarify apply (rule conjI) apply (blast dest: boundedD) apply (erule pres_typeD) prefer 3 apply assumption apply simp apply simp apply (drule_tac P = "\x. (a, b) \ set (step q x)" in subst) apply assumption apply (simp add: plusplus_empty) apply (cases "q \ w") apply simp apply (rule ub1') apply (rule Semilat.intro) apply (rule semilat) apply clarify apply (rule pres_typeD) apply assumption prefer 3 apply assumption - apply (blast intro: listE_nth_in dest: boundedD) + apply (blast intro: nlistsE_nth_in dest: boundedD) apply (blast intro: pres_typeD dest: boundedD) - apply (blast intro: listE_nth_in dest: boundedD) + apply (blast intro: nlistsE_nth_in dest: boundedD) apply assumption apply simp apply (erule allE, erule impE, assumption, erule impE, assumption) apply (rule order_trans) apply fastforce defer apply (rule pp_ub2) apply simp apply clarify apply simp apply (rule pres_typeD) apply assumption prefer 3 apply assumption - apply (blast intro: listE_nth_in) + apply (blast intro: nlistsE_nth_in) apply (blast) - apply (blast intro: listE_nth_in dest: boundedD) + apply (blast intro: nlistsE_nth_in dest: boundedD) prefer 4 apply fastforce - apply (blast intro: listE_nth_in dest: pres_typeD) - apply (blast intro: listE_nth_in dest: boundedD) + apply (blast intro: nlistsE_nth_in dest: pres_typeD) + apply (blast intro: nlistsE_nth_in dest: boundedD) apply(subgoal_tac "\(q,t) \ set (step p (ss!p)). q < n \ t \ A") - apply (subgoal_tac "merges f (step p (ss!p)) ss \ list n A") + apply (subgoal_tac "merges f (step p (ss!p)) ss \ nlists n A") defer apply (blast dest:merges_preserves_type) defer apply (subgoal_tac "a < n") defer - apply (blast intro: listE_nth_in boundedD ) + apply (blast intro: nlistsE_nth_in boundedD ) defer apply (subgoal_tac "merges f (step p (ss ! p)) ss ! a = map snd (filter (\(p', t'). p' = a) (step p (ss ! p))) \\<^bsub>f\<^esub> ss ! a") - apply (fastforce dest: listE_nth_in ) (* slow *) + apply (fastforce dest: nlistsE_nth_in ) (* slow *) apply (subgoal_tac "a < length ss") apply (fastforce dest:nth_merges) apply simp apply (intro strip) apply auto defer apply (auto intro:boundedD) - apply (fastforce intro: listE_nth_in dest:pres_typeD) + apply (fastforce intro: nlistsE_nth_in dest:pres_typeD) done (*>*) lemma (in Semilat) merges_bounded_lemma: "\ mono r step n A; bounded step n; pres_type step n A; - \(p',s') \ set (step p (ss!p)). s' \ A; ss \ list n A; ts \ list n A; p < n; + \(p',s') \ set (step p (ss!p)). s' \ A; ss \ nlists n A; ts \ nlists n A; p < n; ss [\\<^sub>r] ts; \p. p < n \ stable r step ts p \ \ merges f (step p (ss!p)) ss [\\<^sub>r] ts" (*<*) apply (unfold stable_def) apply (rule merges_pres_le_ub) apply simp apply simp prefer 2 apply assumption apply clarsimp apply (drule boundedD, assumption+) apply (erule allE, erule impE, assumption) apply (drule bspec, assumption) apply simp apply (drule monoD [of _ _ _ _ p "ss!p" "ts!p"]) apply assumption apply simp apply (simp add: le_listD) apply (drule lesub_step_typeD, assumption) apply clarify apply (drule bspec, assumption) apply simp apply (rule order_trans) apply fastforce apply simp apply simp apply simp apply (auto intro:boundedD) - apply (blast intro: listE_nth_in dest:pres_typeD) + apply (blast intro: nlistsE_nth_in dest:pres_typeD) done (*>*) (** iter **) lemma termination_lemma: assumes "Semilat A r f" -shows "\ ss \ list n A; \(q,t)\set qs. q t\A; p\w \ \ +shows "\ ss \ nlists n A; \(q,t)\set qs. q t\A; p\w \ \ ss [\\<^sub>r] merges f qs ss \ merges f qs ss = ss \ {q. \t. (q,t)\set qs \ t \\<^bsub>f\<^esub> ss!q \ ss!q} \ (w-{p}) \ w" (*<*) (is "PROP ?P") proof - interpret Semilat A r f by fact show "PROP ?P" apply(insert semilat) apply (unfold lesssub_def) apply (simp (no_asm_simp) add: merges_incr) apply (rule impI) apply (rule merges_same_conv [THEN iffD1, elim_format]) apply assumption+ defer apply (rule sym, assumption) defer apply simp apply (subgoal_tac "\q t. \((q, t) \ set qs \ t \\<^bsub>f\<^esub> ss ! q \ ss ! q)") apply (blast intro!: psubsetI elim: equalityE) apply clarsimp apply (drule bspec, assumption) apply (drule bspec, assumption) apply clarsimp done qed (*>*) end diff --git a/thys/Jinja/DFA/LBVCorrect.thy b/thys/Jinja/DFA/LBVCorrect.thy --- a/thys/Jinja/DFA/LBVCorrect.thy +++ b/thys/Jinja/DFA/LBVCorrect.thy @@ -1,221 +1,221 @@ (* Author: Gerwin Klein Copyright 1999 Technische Universitaet Muenchen *) section \Correctness of the LBV\ theory LBVCorrect imports LBVSpec Typing_Framework_1 begin locale lbvs = lbv + fixes s\<^sub>0 :: 'a fixes c :: "'a list" fixes ins :: "'b list" fixes \s :: "'a list" defines phi_def: "\s \ map (\pc. if c!pc = \ then wtl (take pc ins) c 0 s\<^sub>0 else c!pc) [0.. \ A" assumes pres: "pres_type step (size ins) A" lemma (in lbvs) phi_None [intro?]: "\ pc < size ins; c!pc = \ \ \ \s!pc = wtl (take pc ins) c 0 s\<^sub>0" (*<*) by (simp add: phi_def) (*>*) lemma (in lbvs) phi_Some [intro?]: "\ pc < size ins; c!pc \ \ \ \ \s!pc = c!pc" (*<*) by (simp add: phi_def) (*>*) lemma (in lbvs) phi_len [simp]: "size \s = size ins" (*<*) by (simp add: phi_def) (*>*) lemma (in lbvs) wtl_suc_pc: assumes all: "wtl ins c 0 s\<^sub>0 \ \" assumes pc: "pc+1 < size ins" assumes sA: "s\<^sub>0 \ A" shows "wtl (take (pc+1) ins) c 0 s\<^sub>0 \\<^sub>r \s!(pc+1)" (*<*) proof - from all pc have "wtc c (pc+1) (wtl (take (pc+1) ins) c 0 s\<^sub>0) \ T" by (rule wtl_all) with pc show ?thesis using sA pres cert all wtl_pres by (simp add: phi_def wtc split: if_split_asm) qed (*>*) lemma (in lbvs) wtl_stable: assumes wtl: "wtl ins c 0 s\<^sub>0 \ \" assumes s\<^sub>0: "s\<^sub>0 \ A" and pc: "pc < size ins" shows "stable r step \s pc" (*<*) proof (unfold stable_def, clarify) fix pc' s' assume step: "(pc',s') \ set (step pc (\s ! pc))" (is "(pc',s') \ set (?step pc)") from bounded pc step have pc': "pc' < size ins" by (rule boundedD) have tkpc: "wtl (take pc ins) c 0 s\<^sub>0 \ \" (is "?s\<^sub>1 \ _") using wtl by (rule wtl_take) have s\<^sub>2: "wtl (take (pc+1) ins) c 0 s\<^sub>0 \ \" (is "?s\<^sub>2 \ _") using wtl by (rule wtl_take) from wtl pc have wt_s\<^sub>1: "wtc c pc ?s\<^sub>1 \ \" by (rule wtl_all) have c_Some: "\pc t. pc < size ins \ c!pc \ \ \ \s!pc = c!pc" by (simp add: phi_def) have c_None: "c!pc = \ \ \s!pc = ?s\<^sub>1" using pc .. from wt_s\<^sub>1 pc c_None c_Some have inst: "wtc c pc ?s\<^sub>1 = wti c pc (\s!pc)" by (simp add: wtc split: if_split_asm) have "?s\<^sub>1 \ A" using pres cert s\<^sub>0 wtl pc by (rule wtl_pres) with pc c_Some cert c_None have "\s!pc \ A" by (cases "c!pc = \") (auto dest: cert_okD1) with pc pres have step_in_A: "snd`set (?step pc) \ A" by (auto dest: pres_typeD2) then have inA1: "s' \ A" using step by auto show "s' \\<^sub>r \s!pc'" proof (cases "pc' = pc+1") case True with pc' cert have cert_in_A: "c!(pc+1) \ A" by (auto dest: cert_okD1) from True pc' have pc1: "pc+1 < size ins" by simp with pres cert s\<^sub>0 wtl have inA2: "wtl (take (pc + 1) ins) c 0 s\<^sub>0 \ A" by (auto dest:wtl_pres) have c_None': "c!(pc +1)= \ \ \s!(pc + 1)= ?s\<^sub>2" using pc1 .. have "?s\<^sub>2 \ A" using pres cert s\<^sub>0 wtl pc1 by (rule wtl_pres) with pc1 c_Some cert c_None' have inA3: "\s!(pc+1) \ A" by (cases "c!(pc+1) = \") (auto dest: cert_okD1) from pc1 tkpc have "?s\<^sub>2 = wtc c pc ?s\<^sub>1" by - (rule wtl_Suc) with inst have merge: "?s\<^sub>2 = merge c pc (?step pc) (c!(pc+1))" by (simp add: wti) also from s\<^sub>2 merge have "\ \ \" (is "?merge \ _") by simp with cert_in_A step_in_A have "?merge = (map snd [(p',t') \ ?step pc. p'=pc+1] \\<^bsub>f\<^esub> c!(pc+1))" by (rule merge_not_top_s) finally have "s' \\<^sub>r ?s\<^sub>2" using step_in_A cert_in_A True step by (auto intro: pp_ub1') also from wtl pc1 have "?s\<^sub>2 \\<^sub>r \s!(pc+1)" using s\<^sub>0 by (auto dest: wtl_suc_pc) also note True [symmetric] finally show ?thesis using inA1 inA2 inA3 by simp next case False from wt_s\<^sub>1 inst have "merge c pc (?step pc) (c!(pc+1)) \ \" by (simp add: wti) with step_in_A have "\(pc', s')\set (?step pc). pc'\pc+1 \ s' \\<^sub>r c!pc'" by - (rule merge_not_top) with step False have ok: "s' \\<^sub>r c!pc'" by blast moreover from ok have "c!pc' = \ \ s' = \" using inA1 by simp moreover from c_Some pc' have "c!pc' \ \ \ \s!pc' = c!pc'" by auto ultimately show ?thesis by (cases "c!pc' = \") auto qed qed (*>*) lemma (in lbvs) phi_not_top: assumes wtl: "wtl ins c 0 s\<^sub>0 \ \" and pc: "pc < size ins" shows "\s!pc \ \" (*<*) proof (cases "c!pc = \") case False with pc have "\s!pc = c!pc" .. also from cert pc have "\ \ \" by (rule cert_okD4) finally show ?thesis . next case True with pc have "\s!pc = wtl (take pc ins) c 0 s\<^sub>0" .. also from wtl have "\ \ \" by (rule wtl_take) finally show ?thesis . qed (*>*) lemma (in lbvs) phi_in_A: assumes wtl: "wtl ins c 0 s\<^sub>0 \ \" and s\<^sub>0: "s\<^sub>0 \ A" - shows "\s \ list (size ins) A" + shows "\s \ nlists (size ins) A" (*<*) proof - { fix x assume "x \ set \s" then obtain xs ys where "\s = xs @ x # ys" by (auto simp add: in_set_conv_decomp) then obtain pc where pc: "pc < size \s" and x: "\s!pc = x" by (simp add: that [of "size xs"] nth_append) from pres cert wtl s\<^sub>0 pc have "wtl (take pc ins) c 0 s\<^sub>0 \ A" by (auto intro!: wtl_pres) moreover from pc have "pc < size ins" by simp with cert have "c!pc \ A" .. ultimately have "\s!pc \ A" using pc by (simp add: phi_def) hence "x \ A" using x by simp } hence "set \s \ A" .. - thus ?thesis by (unfold list_def) simp + thus ?thesis by (unfold nlists_def) simp qed (*>*) lemma (in lbvs) phi0: assumes wtl: "wtl ins c 0 s\<^sub>0 \ \" and 0: "0 < size ins" and s\<^sub>0: "s\<^sub>0 \ A" shows "s\<^sub>0 \\<^sub>r \s!0" (*<*) proof (cases "c!0 = \") case True with 0 have "\s!0 = wtl (take 0 ins) c 0 s\<^sub>0" .. moreover have "wtl (take 0 ins) c 0 s\<^sub>0 = s\<^sub>0" by simp ultimately have "\s!0 = s\<^sub>0" by simp thus ?thesis using s\<^sub>0 by simp next case False with 0 have "\s!0 = c!0" .. moreover have "wtl (take 1 ins) c 0 s\<^sub>0 \ \" using wtl by (rule wtl_take) with 0 False have "s\<^sub>0 \\<^sub>r c!0" by (auto simp add: neq_Nil_conv wtc split: if_split_asm) ultimately show ?thesis by simp qed (*>*) theorem (in lbvs) wtl_sound: assumes wtl: "wtl ins c 0 s\<^sub>0 \ \" and s\<^sub>0: "s\<^sub>0 \ A" shows "\\s. wt_step r \ step \s" (*<*) proof - have "wt_step r \ step \s" proof (unfold wt_step_def, intro strip conjI) fix pc assume "pc < size \s" then obtain pc: "pc < size ins" by simp with wtl show "\s!pc \ \" by (rule phi_not_top) from wtl s\<^sub>0 pc show "stable r step \s pc" by (rule wtl_stable) qed thus ?thesis .. qed (*>*) theorem (in lbvs) wtl_sound_strong: assumes wtl: "wtl ins c 0 s\<^sub>0 \ \" assumes s\<^sub>0: "s\<^sub>0 \ A" and ins: "0 < size ins" - shows "\\s \ list (size ins) A. wt_step r \ step \s \ s\<^sub>0 \\<^sub>r \s!0" + shows "\\s \ nlists (size ins) A. wt_step r \ step \s \ s\<^sub>0 \\<^sub>r \s!0" (*<*) proof - - have "\s \ list (size ins) A" using wtl s\<^sub>0 by (rule phi_in_A) + have "\s \ nlists (size ins) A" using wtl s\<^sub>0 by (rule phi_in_A) moreover have "wt_step r \ step \s" proof (unfold wt_step_def, intro strip conjI) fix pc assume "pc < size \s" then obtain pc: "pc < size ins" by simp with wtl show "\s!pc \ \" by (rule phi_not_top) from wtl s\<^sub>0 and pc show "stable r step \s pc" by (rule wtl_stable) qed moreover from wtl ins have "s\<^sub>0 \\<^sub>r \s!0" using s\<^sub>0 by (rule phi0) ultimately show ?thesis by fast qed (*>*) end diff --git a/thys/Jinja/DFA/Listn.thy b/thys/Jinja/DFA/Listn.thy --- a/thys/Jinja/DFA/Listn.thy +++ b/thys/Jinja/DFA/Listn.thy @@ -1,710 +1,584 @@ (* Title: HOL/MicroJava/BV/Listn.thy Author: Tobias Nipkow Copyright 2000 TUM Lists of a fixed length. *) section \Fixed Length Lists\ theory Listn -imports Err +imports Err "HOL-Library.NList" begin -definition list :: "nat \ 'a set \ 'a list set" -where - "list n A = {xs. size xs = n \ set xs \ A}" - definition le :: "'a ord \ ('a list)ord" where "le r = list_all2 (\x y. x \\<^sub>r y)" abbreviation lesublist :: "'a list \ 'a ord \ 'a list \ bool" ("(_ /[\\<^bsub>_\<^esub>] _)" [50, 0, 51] 50) where "x [\\<^bsub>r\<^esub>] y == x <=_(Listn.le r) y" abbreviation lesssublist :: "'a list \ 'a ord \ 'a list \ bool" ("(_ /[\\<^bsub>_\<^esub>] _)" [50, 0, 51] 50) where "x [\\<^bsub>r\<^esub>] y == x <_(Listn.le r) y" (*<*) notation (ASCII) lesublist ("(_ /[<=_] _)" [50, 0, 51] 50) and lesssublist ("(_ /[<_] _)" [50, 0, 51] 50) abbreviation (input) lesublist2 :: "'a list \ 'a ord \ 'a list \ bool" ("(_ /[\\<^sub>_] _)" [50, 0, 51] 50) where "x [\\<^sub>r] y == x [\\<^bsub>r\<^esub>] y" abbreviation (input) lesssublist2 :: "'a list \ 'a ord \ 'a list \ bool" ("(_ /[\\<^sub>_] _)" [50, 0, 51] 50) where "x [\\<^sub>r] y == x [\\<^bsub>r\<^esub>] y" (*>*) abbreviation plussublist :: "'a list \ ('a \ 'b \ 'c) \ 'b list \ 'c list" ("(_ /[\\<^bsub>_\<^esub>] _)" [65, 0, 66] 65) where "x [\\<^bsub>f\<^esub>] y == x \\<^bsub>map2 f\<^esub> y" (*<*) notation (ASCII) plussublist ("(_ /[+_] _)" [65, 0, 66] 65) abbreviation (input) plussublist2 :: "'a list \ ('a \ 'b \ 'c) \ 'b list \ 'c list" ("(_ /[\\<^sub>_] _)" [65, 0, 66] 65) where "x [\\<^sub>f] y == x [\\<^bsub>f\<^esub>] y" (*>*) primrec coalesce :: "'a err list \ 'a list err" where "coalesce [] = OK[]" | "coalesce (ex#exs) = Err.sup (#) ex (coalesce exs)" definition sl :: "nat \ 'a sl \ 'a list sl" where - "sl n = (\(A,r,f). (list n A, le r, map2 f))" + "sl n = (\(A,r,f). (nlists n A, le r, map2 f))" definition sup :: "('a \ 'b \ 'c err) \ 'a list \ 'b list \ 'c list err" where "sup f = (\xs ys. if size xs = size ys then coalesce(xs [\\<^bsub>f\<^esub>] ys) else Err)" definition upto_esl :: "nat \ 'a esl \ 'a list esl" where - "upto_esl m = (\(A,r,f). (Union{list n A |n. n \ m}, le r, sup f))" + "upto_esl m = (\(A,r,f). (Union{nlists n A |n. n \ m}, le r, sup f))" lemmas [simp] = set_update_subsetI lemma unfold_lesub_list: "xs [\\<^bsub>r\<^esub>] ys = Listn.le r xs ys" (*<*) by (simp add: lesub_def) (*>*) lemma Nil_le_conv [iff]: "([] [\\<^bsub>r\<^esub>] ys) = (ys = [])" (*<*) apply (unfold lesub_def Listn.le_def) apply simp done (*>*) lemma Cons_notle_Nil [iff]: "\ x#xs [\\<^bsub>r\<^esub>] []" (*<*) apply (unfold lesub_def Listn.le_def) apply simp done (*>*) lemma Cons_le_Cons [iff]: "x#xs [\\<^bsub>r\<^esub>] y#ys = (x \\<^sub>r y \ xs [\\<^bsub>r\<^esub>] ys)" (*<*) by (simp add: lesub_def Listn.le_def) (*>*) - - - - - - - - lemma list_update_le_cong: "\ i\<^bsub>r\<^esub>] ys; x \\<^sub>r y \ \ xs[i:=x] [\\<^bsub>r\<^esub>] ys[i:=y]" (*<*) apply (unfold unfold_lesub_list) apply (unfold Listn.le_def) apply (simp add: list_all2_update_cong) done (*>*) lemma le_listD: "\ xs [\\<^bsub>r\<^esub>] ys; p < size xs \ \ xs!p \\<^sub>r ys!p" (*<*) by (simp add: Listn.le_def lesub_def list_all2_nthD) (*>*) lemma le_list_refl: "\x. x \\<^sub>r x \ xs [\\<^bsub>r\<^esub>] xs" (*<*) apply (simp add: unfold_lesub_list lesub_def Listn.le_def list_all2_refl) done (*>*) lemma le_list_trans: assumes ord: "order r A" - and xs: "xs \ list n A" and ys: "ys \ list n A" and zs: "zs \ list n A" + and xs: "xs \ nlists n A" and ys: "ys \ nlists n A" and zs: "zs \ nlists n A" and "xs [\\<^bsub>r\<^esub>] ys" and "ys [\\<^bsub>r\<^esub>] zs" shows "xs [\\<^bsub>r\<^esub>] zs" (*<*) using assms proof (unfold le_def lesssub_def lesub_def) assume "list_all2 r xs ys" and "list_all2 r ys zs" hence xs_ys_zs: "\i < length xs. r (xs!i) (ys!i) \ r (ys!i) (zs!i)" and len_xs_zs: "length xs = length zs" by (auto simp add: list_all2_conv_all_nth) - from xs ys zs have inA: "\i A \ ys!i \ A \ zs!i \ A" by (unfold list_def) auto + from xs ys zs have inA: "\i A \ ys!i \ A \ zs!i \ A" by (unfold nlists_def) auto from ord have "\x\A. \y\A. \z\A. x \\<^sub>r y \ y \\<^sub>r z \ x \\<^sub>r z" by (unfold order_def) blast hence "\x\A. \y\A. \z\A. r x y \ r y z \ r x z" by (unfold lesssub_def lesub_def) with xs_ys_zs inA have "\i list n A" and ys: "ys \ list n A" + and xs: "xs \ nlists n A" and ys: "ys \ nlists n A" and "xs [\\<^bsub>r\<^esub>] ys" and "ys [\\<^bsub>r\<^esub>] xs" shows "xs = ys" (*<*) using assms proof (simp add: le_def lesssub_def lesub_def) assume "list_all2 r xs ys" and "list_all2 r ys xs" hence xs_ys: "\i < length xs. r (xs!i) (ys!i) \ r (ys!i) (xs!i)" and len_xs_ys: "length xs = length ys" by (auto simp add: list_all2_conv_all_nth) - from xs ys have inA: "\i A \ ys!i \ A" by (unfold list_def) auto + from xs ys have inA: "\i A \ ys!i \ A" by (unfold nlists_def) auto from ord have "\x\A. \y\A. x \\<^sub>r y \ y \\<^sub>r x \ x = y" by (unfold order_def) blast hence "\x\A. \y\A. r x y \ r y x \ x = y" by (unfold lesssub_def lesub_def) with xs_ys inA have "\i*) -lemma listE_set [simp]: "xs \ list n A \ set xs \ A" -(*<*) -apply (unfold list_def) -apply blast -done -(*>*) - - lemma nth_in [rule_format, simp]: "\i n. size xs = n \ set xs \ A \ i < n \ (xs!i) \ A" (*<*) apply (induct "xs") apply simp apply (simp add: nth_Cons split: nat.split) done (*>*) -(* FIXME: remove simp *) -lemma listE_length [simp]: "xs \ list n A \ size xs = n" -(*<*) -apply (unfold list_def) -apply blast -done -(*>*) - -lemma listE_nth_in: "\ xs \ list n A; i < n \ \ xs!i \ A" -(*<*) by auto (*>*) - - -lemma le_list_refl': "order r A \ xs \ list n A \ xs \\<^bsub>Listn.le r\<^esub> xs" +lemma le_list_refl': "order r A \ xs \ nlists n A \ xs \\<^bsub>Listn.le r\<^esub> xs" apply (unfold le_def lesssub_def lesub_def list_all2_conv_all_nth) apply auto apply (subgoal_tac "xs ! i \ A") apply (subgoal_tac "(xs ! i) \\<^sub>r (xs ! i)") apply (simp only: lesssub_def lesub_def) - apply (fastforce dest: listE_nth_in) + apply (fastforce dest: nlistsE_nth_in) apply (fastforce dest: order_refl) done -lemma order_listI [simp, intro!]: "order r A \ order(Listn.le r) (list n A)" +lemma order_listI [simp, intro!]: "order r A \ order(Listn.le r) (nlists n A)" (*<*) proof- assume ord: "order r A" let ?r = "Listn.le r" - let ?A = "list n A" + let ?A = "nlists n A" have "\x\?A. x \\<^bsub>?r\<^esub> x" using ord le_list_refl' by auto moreover have "\x\?A. \y\?A. x \\<^bsub>?r\<^esub> y \ y \\<^bsub>?r\<^esub> x \ x=y" using ord le_list_antisym by auto moreover have "\x\?A. \y\?A. \z\?A. x \\<^bsub>?r\<^esub>y \ y \\<^bsub>?r\<^esub> z \ x \\<^bsub>?r\<^esub> z" using ord le_list_trans by auto ultimately show ?thesis by (auto simp only: order_def) qed (*>*) -lemma le_list_refl2': "order r A \ xs \ (\{list n A |n. n \ mxs})\ xs \\<^bsub>Listn.le r\<^esub> xs" +lemma le_list_refl2': "order r A \ xs \ (\{nlists n A |n. n \ mxs})\ xs \\<^bsub>Listn.le r\<^esub> xs" by (auto simp add:le_list_refl') lemma le_list_trans2: assumes "order r A" - and "xs \ (\{list n A |n. n \ mxs})" and "ys \(\{list n A |n. n \ mxs})" and "zs \(\{list n A |n. n \ mxs})" + and "xs \ (\{nlists n A |n. n \ mxs})" and "ys \(\{nlists n A |n. n \ mxs})" and "zs \(\{nlists n A |n. n \ mxs})" and "xs [\\<^bsub>r\<^esub>] ys" and "ys [\\<^bsub>r\<^esub>] zs" shows "xs [\\<^bsub>r\<^esub>] zs" (*<*)using assms -proof(auto simp only:list_def le_def lesssub_def lesub_def) +proof(auto simp only:nlists_def le_def lesssub_def lesub_def) assume xA: "set xs \ A" and "length xs \ mxs " and " length ys \ mxs " and yA: "set ys \ A" and len_zs: "length zs \ mxs" and zA: "set zs \ A" and xy: "list_all2 r xs ys" and yz: "list_all2 r ys zs" and ord: "order r A" from xy yz have xs_ys: "length xs = length ys" and ys_zs: "length ys = length zs" by (auto simp add:list_all2_conv_all_nth) from ord have "\x\A. \y\A. \z\A. x \\<^sub>r y \ y \\<^sub>r z \ x \\<^sub>r z" by (unfold order_def) blast hence trans: "\x\A. \y\A. \z\A. r x y \ r y z \ r x z" by (simp only:lesssub_def lesub_def) from xA yA zA have x_inA: "\i A" and y_inA: "\i A" and z_inA: "\i A" using xs_ys ys_zs by auto from xy yz have "\i < length xs. r (xs!i) (ys!i)" and "\i < length ys. r (ys!i) (zs!i)" by (auto simp add:list_all2_conv_all_nth) hence "\i < length xs. r (xs!i) (ys!i) \ r (ys!i) (zs!i)" using xs_ys ys_zs by auto with x_inA y_inA z_inA have "\i(\{list n A |n. n \ mxs})" and "ys \(\{list n A |n. n \ mxs})" + and "xs \(\{nlists n A |n. n \ mxs})" and "ys \(\{nlists n A |n. n \ mxs})" and "xs [\\<^bsub>r\<^esub>] ys" and "ys [\\<^bsub>r\<^esub>] xs" shows " xs = ys" (*<*) using assms -proof(auto simp only:list_def le_def lesssub_def lesub_def) +proof(auto simp only:nlists_def le_def lesssub_def lesub_def) assume xA: "set xs \ A" and len_ys: "length ys \ mxs" and len_xs: "length xs \ mxs" and yA: "set ys \ A" and xy: "list_all2 r xs ys" and yx: "list_all2 r ys xs" and ord: "order r A" from xy have len_eq_xs_ys: "length xs = length ys" by (simp add:list_all2_conv_all_nth) from ord have antisymm:"\x\A. \y\A. r x y \ r y x\ x=y " by (auto simp only:lesssub_def lesub_def order_def) from xA yA have x_inA: "\i A" and y_inA: "\i A" by auto from xy yx have "\i < length xs. r (xs!i) (ys!i)" and "\i < length ys. r (ys!i) (xs!i)" by (auto simp add:list_all2_conv_all_nth) with x_inA y_inA have "\i order(Listn.le r) (\{list n A |n. n \ mxs})" +lemma order_listI2[intro!] : "order r A \ order(Listn.le r) (\{nlists n A |n. n \ mxs})" (*<*) proof- assume ord: "order r A" let ?r = "Listn.le r" - let ?A = "(\{list n A |n. n \ mxs})" + let ?A = "(\{nlists n A |n. n \ mxs})" have "\x\?A. x \\<^bsub>?r\<^esub> x" using ord le_list_refl2' by auto \\ use order_listI \ moreover have "\x\?A. \y\?A. x \\<^bsub>?r\<^esub> y \ y \\<^bsub>?r\<^esub> x \ x=y" using ord le_list_antisym2 by blast moreover have "\x\?A. \y\?A. \z\?A. x \\<^bsub>?r\<^esub>y \ y \\<^bsub>?r\<^esub> z \ x \\<^bsub>?r\<^esub> z" using ord le_list_trans2 by blast ultimately show ?thesis by (auto simp only: order_def) qed lemma lesub_list_impl_same_size [simp]: "xs [\\<^bsub>r\<^esub>] ys \ size ys = size xs" (*<*) apply (unfold Listn.le_def lesub_def) apply (simp add: list_all2_lengthD) done (*>*) lemma lesssub_lengthD: "xs [\\<^sub>r] ys \ size ys = size xs" (*<*) apply (unfold lesssub_def) apply auto done (*>*) lemma le_list_appendI: "a [\\<^bsub>r\<^esub>] b \ c [\\<^bsub>r\<^esub>] d \ a@c [\\<^bsub>r\<^esub>] b@d" (*<*) apply (unfold Listn.le_def lesub_def) apply (rule list_all2_appendI, assumption+) done (*>*) lemma le_listI: assumes "length a = length b" assumes "\n. n < length a \ a!n \\<^sub>r b!n" shows "a [\\<^bsub>r\<^esub>] b" (*<*) proof - from assms have "list_all2 r a b" by (simp add: list_all2_all_nthI lesub_def) then show ?thesis by (simp add: Listn.le_def lesub_def) qed (*>*) -lemma listI: "\ size xs = n; set xs \ A \ \ xs \ list n A" -(*<*) -apply (unfold list_def) -apply blast -done -(*>*) - -lemma less_lengthI: "\ xs \ list n A; p < n \ \ p < size xs" -(*<*) by simp (*>*) - -lemma list_0 [simp]: "list 0 A = {[]}" -(*<*) -apply (unfold list_def) -apply auto -done -(*>*) - -lemma in_list_Suc_iff: - "(xs \ list (Suc n) A) = (\y\A. \ys \ list n A. xs = y#ys)" -(*<*) -apply (unfold list_def) -apply (case_tac "xs") -apply auto -done -(*>*) - - -lemma Cons_in_list_Suc [iff]: - "(x#xs \ list (Suc n) A) = (x\A \ xs \ list n A)" -(*<*) -apply (simp add: in_list_Suc_iff) -done -(*>*) - -lemma list_not_empty: - "\a. a\A \ \xs. xs \ list n A" -(*<*) -apply (induct "n") - apply simp -apply (simp add: in_list_Suc_iff) -apply blast -done -(*>*) - -lemma listn_Cons_Suc [elim!]: - "l#xs \ list n A \ (\n'. n = Suc n' \ l \ A \ xs \ list n' A \ P) \ P" -(*<*) by (cases n) auto (*>*) - -lemma listn_appendE [elim!]: - "a@b \ list n A \ (\n1 n2. n=n1+n2 \ a \ list n1 A \ b \ list n2 A \ P) \ P" -(*<*) -proof - - have "\n. a@b \ list n A \ \n1 n2. n=n1+n2 \ a \ list n1 A \ b \ list n2 A" - (is "\n. ?list a n \ \n1 n2. ?P a n n1 n2") - proof (induct a) - fix n assume "?list [] n" - hence "?P [] n 0 n" by simp - thus "\n1 n2. ?P [] n n1 n2" by fast - next - fix n l ls - assume "?list (l#ls) n" - then obtain n' where n: "n = Suc n'" "l \ A" and n': "ls@b \ list n' A" by fastforce - assume "\n. ls @ b \ list n A \ \n1 n2. n = n1 + n2 \ ls \ list n1 A \ b \ list n2 A" - from this and n' have "\n1 n2. n' = n1 + n2 \ ls \ list n1 A \ b \ list n2 A" . - then obtain n1 n2 where "n' = n1 + n2" "ls \ list n1 A" "b \ list n2 A" by fast - with n have "?P (l#ls) n (n1+1) n2" by simp - thus "\n1 n2. ?P (l#ls) n n1 n2" by fastforce - qed - moreover - assume "a@b \ list n A" "\n1 n2. n=n1+n2 \ a \ list n1 A \ b \ list n2 A \ P" - ultimately - show ?thesis by blast -qed -(*>*) - - -lemma listt_update_in_list [simp, intro!]: - "\ xs \ list n A; x\A \ \ xs[i := x] \ list n A" -(*<*) -apply (unfold list_def) -apply simp -done -(*>*) - -lemma list_appendI [intro?]: - "\ a \ list n A; b \ list m A \ \ a @ b \ list (n+m) A" -(*<*) by (unfold list_def) auto (*>*) - -lemma list_map [simp]: "(map f xs \ list (size xs) A) = (f ` set xs \ A)" -(*<*) by (unfold list_def) simp (*>*) - -lemma list_replicateI [intro]: "x \ A \ replicate n x \ list n A" -(*<*) by (induct n) auto (*>*) - lemma plus_list_Nil [simp]: "[] [\\<^bsub>f\<^esub>] xs = []" (*<*) apply (unfold plussub_def) apply simp done (*>*) lemma plus_list_Cons [simp]: "(x#xs) [\\<^bsub>f\<^esub>] ys = (case ys of [] \ [] | y#ys \ (x \\<^sub>f y)#(xs [\\<^bsub>f\<^esub>] ys))" (*<*) by (simp add: plussub_def split: list.split) (*>*) lemma length_plus_list [rule_format, simp]: "\ys. size(xs [\\<^bsub>f\<^esub>] ys) = min(size xs) (size ys)" (*<*) apply (induct xs) apply simp apply clarify apply (simp (no_asm_simp) split: list.split) done (*>*) lemma nth_plus_list [rule_format, simp]: "\xs ys i. size xs = n \ size ys = n \ i (xs [\\<^bsub>f\<^esub>] ys)!i = (xs!i) \\<^sub>f (ys!i)" (*<*) apply (induct n) apply simp apply clarify apply (case_tac xs) apply simp apply (force simp add: nth_Cons split: list.split nat.split) done (*>*) lemma (in Semilat) plus_list_ub1 [rule_format]: "\ set xs \ A; set ys \ A; size xs = size ys \ \ xs [\\<^bsub>r\<^esub>] xs [\\<^bsub>f\<^esub>] ys" (*<*) apply (unfold unfold_lesub_list) apply (simp add: Listn.le_def list_all2_conv_all_nth) done (*>*) lemma (in Semilat) plus_list_ub2: "\set xs \ A; set ys \ A; size xs = size ys \ \ ys [\\<^bsub>r\<^esub>] xs [\\<^bsub>f\<^esub>] ys" (*<*) apply (unfold unfold_lesub_list) apply (simp add: Listn.le_def list_all2_conv_all_nth) done (*>*) lemma (in Semilat) plus_list_lub [rule_format]: shows "\xs ys zs. set xs \ A \ set ys \ A \ set zs \ A \ size xs = n \ size ys = n \ xs [\\<^bsub>r\<^esub>] zs \ ys [\\<^bsub>r\<^esub>] zs \ xs [\\<^bsub>f\<^esub>] ys [\\<^bsub>r\<^esub>] zs" (*<*) apply (unfold unfold_lesub_list) apply (simp add: Listn.le_def list_all2_conv_all_nth) done (*>*) lemma (in Semilat) list_update_incr [rule_format]: "x\A \ set xs \ A \ (\i. i xs [\\<^bsub>r\<^esub>] xs[i := x \\<^sub>f xs!i])" (*<*) apply (unfold unfold_lesub_list) apply (simp add: Listn.le_def list_all2_conv_all_nth) apply (induct xs) apply simp -apply (simp add: in_list_Suc_iff) +apply (simp add: in_nlists_Suc_iff) apply clarify apply (simp add: nth_Cons split: nat.split) done (*>*) -lemma closed_listI: - "closed S f \ closed (list n S) (map2 f)" +lemma closed_nlistsI: + "closed S f \ closed (nlists n S) (map2 f)" (*<*) apply (unfold closed_def) apply (induct n) apply simp apply clarify -apply (simp add: in_list_Suc_iff) +apply (simp add: in_nlists_Suc_iff) apply clarify apply simp done (*>*) lemma Listn_sl_aux: assumes "Semilat A r f" shows "semilat (Listn.sl n (A,r,f))" (*<*) proof - interpret Semilat A r f by fact show ?thesis apply (unfold Listn.sl_def) apply (simp (no_asm) only: semilat_Def split_conv) apply (rule conjI) apply simp apply (rule conjI) - apply (simp only: closedI closed_listI) - apply (simp (no_asm) only: list_def) + apply (simp only: closedI closed_nlistsI) + apply (simp (no_asm) only: nlists_def) apply (simp (no_asm_simp) add: plus_list_ub1 plus_list_ub2 plus_list_lub) done qed (*>*) lemma Listn_sl: "semilat L \ semilat (Listn.sl n L)" (*<*) apply (cases L) apply simp apply (drule Semilat.intro) by (simp add: Listn_sl_aux split_tupled_all) (*>*) lemma coalesce_in_err_list [rule_format]: - "\xes. xes \ list n (err A) \ coalesce xes \ err(list n A)" + "\xes. xes \ nlists n (err A) \ coalesce xes \ err(nlists n A)" (*<*) apply (induct n) apply simp apply clarify -apply (simp add: in_list_Suc_iff) +apply (simp add: in_nlists_Suc_iff) apply clarify apply (simp (no_asm) add: plussub_def Err.sup_def lift2_def split: err.split) apply force done (*>*) lemma lem: "\x xs. x \\<^bsub>(#)\<^esub> xs = x#xs" (*<*) by (simp add: plussub_def) (*>*) lemma coalesce_eq_OK1_D [rule_format]: "semilat(err A, Err.le r, lift2 f) \ - \xs. xs \ list n A \ (\ys. ys \ list n A \ + \xs. xs \ nlists n A \ (\ys. ys \ nlists n A \ (\zs. coalesce (xs [\\<^bsub>f\<^esub>] ys) = OK zs \ xs [\\<^bsub>r\<^esub>] zs))" (*<*) apply (induct n) apply simp apply clarify -apply (simp add: in_list_Suc_iff) +apply (simp add: in_nlists_Suc_iff) apply clarify apply (simp split: err.split_asm add: lem Err.sup_def lift2_def) apply (force simp add: semilat_le_err_OK1) done (*>*) lemma coalesce_eq_OK2_D [rule_format]: "semilat(err A, Err.le r, lift2 f) \ - \xs. xs \ list n A \ (\ys. ys \ list n A \ + \xs. xs \ nlists n A \ (\ys. ys \ nlists n A \ (\zs. coalesce (xs [\\<^bsub>f\<^esub>] ys) = OK zs \ ys [\\<^bsub>r\<^esub>] zs))" (*<*) apply (induct n) apply simp apply clarify -apply (simp add: in_list_Suc_iff) +apply (simp add: in_nlists_Suc_iff) apply clarify apply (simp split: err.split_asm add: lem Err.sup_def lift2_def) apply (force simp add: semilat_le_err_OK2) done (*>*) lemma lift2_le_ub: "\ semilat(err A, Err.le r, lift2 f); x\A; y\A; x \\<^sub>f y = OK z; u\A; x \\<^sub>r u; y \\<^sub>r u \ \ z \\<^sub>r u" (*<*) apply (unfold semilat_Def plussub_def err_def') apply (simp add: lift2_def) apply clarify apply (rotate_tac -3) apply (erule thin_rl) apply (erule thin_rl) apply force done (*>*) lemma coalesce_eq_OK_ub_D [rule_format]: "semilat(err A, Err.le r, lift2 f) \ - \xs. xs \ list n A \ (\ys. ys \ list n A \ + \xs. xs \ nlists n A \ (\ys. ys \ nlists n A \ (\zs us. coalesce (xs [\\<^bsub>f\<^esub>] ys) = OK zs \ xs [\\<^bsub>r\<^esub>] us \ ys [\\<^bsub>r\<^esub>] us - \ us \ list n A \ zs [\\<^bsub>r\<^esub>] us))" + \ us \ nlists n A \ zs [\\<^bsub>r\<^esub>] us))" (*<*) apply (induct n) apply simp apply clarify -apply (simp add: in_list_Suc_iff) +apply (simp add: in_nlists_Suc_iff) apply clarify apply (simp (no_asm_use) split: err.split_asm add: lem Err.sup_def lift2_def) apply clarify apply (rule conjI) apply (blast intro: lift2_le_ub) apply blast done (*>*) lemma lift2_eq_ErrD: "\ x \\<^sub>f y = Err; semilat(err A, Err.le r, lift2 f); x\A; y\A \ \ \(\u\A. x \\<^sub>r u \ y \\<^sub>r u)" (*<*) by (simp add: OK_plus_OK_eq_Err_conv [THEN iffD1]) (*>*) lemma coalesce_eq_Err_D [rule_format]: "\ semilat(err A, Err.le r, lift2 f) \ - \ \xs. xs \ list n A \ (\ys. ys \ list n A \ + \ \xs. xs \ nlists n A \ (\ys. ys \ nlists n A \ coalesce (xs [\\<^bsub>f\<^esub>] ys) = Err \ - \(\zs \ list n A. xs [\\<^bsub>r\<^esub>] zs \ ys [\\<^bsub>r\<^esub>] zs))" + \(\zs \ nlists n A. xs [\\<^bsub>r\<^esub>] zs \ ys [\\<^bsub>r\<^esub>] zs))" (*<*) apply (induct n) apply simp apply clarify -apply (simp add: in_list_Suc_iff) +apply (simp add: in_nlists_Suc_iff) apply clarify apply (simp split: err.split_asm add: lem Err.sup_def lift2_def) apply (blast dest: lift2_eq_ErrD) done (*>*) lemma closed_err_lift2_conv: "closed (err A) (lift2 f) = (\x\A. \y\A. x \\<^sub>f y \ err A)" (*<*) apply (unfold closed_def) apply (simp add: err_def') done (*>*) lemma closed_map2_list [rule_format]: "closed (err A) (lift2 f) \ - \xs. xs \ list n A \ (\ys. ys \ list n A \ - map2 f xs ys \ list n (err A))" + \xs. xs \ nlists n A \ (\ys. ys \ nlists n A \ + map2 f xs ys \ nlists n (err A))" (*<*) apply (induct n) apply simp apply clarify -apply (simp add: in_list_Suc_iff) +apply (simp add: in_nlists_Suc_iff) apply clarify apply (simp add: plussub_def closed_err_lift2_conv) done (*>*) lemma closed_lift2_sup: "closed (err A) (lift2 f) \ - closed (err (list n A)) (lift2 (sup f))" + closed (err (nlists n A)) (lift2 (sup f))" (*<*) by (fastforce simp add: closed_def plussub_def sup_def lift2_def coalesce_in_err_list closed_map2_list split: err.split) (*>*) lemma err_semilat_sup: "err_semilat (A,r,f) \ - err_semilat (list n A, Listn.le r, sup f)" + err_semilat (nlists n A, Listn.le r, sup f)" (*<*) apply (unfold Err.sl_def) apply (simp only: split_conv) apply (simp (no_asm) only: semilat_Def plussub_def) apply (simp (no_asm_simp) only: Semilat.closedI [OF Semilat.intro] closed_lift2_sup) apply (rule conjI) apply (drule Semilat.orderI [OF Semilat.intro]) apply simp apply (simp (no_asm) only: unfold_lesub_err Err.le_def err_def' sup_def lift2_def) apply (simp (no_asm_simp) add: coalesce_eq_OK1_D coalesce_eq_OK2_D split: err.split) apply (blast intro: coalesce_eq_OK_ub_D dest: coalesce_eq_Err_D) done (*>*) lemma err_semilat_upto_esl: "\L. err_semilat L \ err_semilat(upto_esl m L)" (*<*) apply (unfold Listn.upto_esl_def) apply (simp (no_asm_simp) only: split_tupled_all) apply simp apply (fastforce intro!: err_semilat_UnionI err_semilat_sup dest: lesub_list_impl_same_size simp add: plussub_def Listn.sup_def) done (*>*) end diff --git a/thys/Jinja/DFA/Typing_Framework_2.thy b/thys/Jinja/DFA/Typing_Framework_2.thy --- a/thys/Jinja/DFA/Typing_Framework_2.thy +++ b/thys/Jinja/DFA/Typing_Framework_2.thy @@ -1,19 +1,19 @@ (* Title: HOL/MicroJava/BV/Typing_Framework.thy Author: Tobias Nipkow Copyright 2000 TUM *) section \Typing and Dataflow Analysis Framework\ theory Typing_Framework_2 imports Typing_Framework_1 begin text \ The relationship between dataflow analysis and a welltyped-instruction predicate. \ definition is_bcv :: "'s ord \ 's \ 's step_type \ nat \ 's set \ ('s list \ 's list) \ bool" where - "is_bcv r T step n A bcv \ (\\s\<^sub>0 \ list n A. - (\ps\<^sub>0)!p \ T) = (\\s \ list n A. \s\<^sub>0 [\\<^sub>r] \s \ wt_step r T step \s))" + "is_bcv r T step n A bcv \ (\\s\<^sub>0 \ nlists n A. + (\ps\<^sub>0)!p \ T) = (\\s \ nlists n A. \s\<^sub>0 [\\<^sub>r] \s \ wt_step r T step \s))" 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,506 +1,506 @@ (* 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 Jinja.Typing_Framework_2 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 \ (b = Static \ b = NonStatic) \ 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" context start_context begin lemma Cons_less_Conss3 [simp]: "x#xs [\\<^bsub>r\<^esub>] y#ys = (x \\<^bsub>r\<^esub> y \ xs [\\<^bsub>r\<^esub>] ys \ x = y \ xs [\\<^bsub>r\<^esub>] ys)" apply (unfold lesssub_def ) apply auto apply (insert sup_state_opt_err) apply (unfold lesssub_def lesub_def sup_state_opt_def sup_state_def sup_ty_opt_def) apply (simp only: JVM_le_unfold ) apply fastforce done lemma acc_le_listI3 [intro!]: " acc r \ acc (Listn.le r)" apply (unfold acc_def) apply (subgoal_tac "wf(UN n. {(ys,xs). size xs = n \ size ys = n \ xs <_(Listn.le r) ys})") apply (erule wf_subset) apply (blast intro: lesssub_lengthD) apply (rule wf_UN) prefer 2 apply (rename_tac m n) apply (case_tac "m=n") apply simp apply (fast intro!: equals0I dest: not_sym) apply (rename_tac n) apply (induct_tac n) apply (simp add: lesssub_def cong: conj_cong) apply (rename_tac k) apply (simp add: wf_eq_minimal del: r_def f_def step_def A_def) apply (simp (no_asm) add: length_Suc_conv cong: conj_cong del: r_def f_def step_def A_def) apply clarify apply (rename_tac M m) apply (case_tac "\x xs. size xs = k \ x#xs \ M") prefer 2 apply (erule thin_rl) apply (erule thin_rl) apply blast apply (erule_tac x = "{a. \xs. size xs = k \ a#xs:M}" in allE) apply (erule impE) apply blast apply (thin_tac "\x xs. P x xs" for P) apply clarify apply (rename_tac maxA xs) apply (erule_tac x = "{ys. size ys = size xs \ maxA#ys \ M}" in allE) apply (erule impE) apply blast apply clarify apply (thin_tac "m \ M") apply (thin_tac "maxA#xs \ M") apply (rule bexI) prefer 2 apply assumption apply clarify apply (simp del: r_def f_def step_def A_def) apply blast done lemma wf_jvm: " wf {(ss', ss). ss [\\<^bsub>r\<^esub>] ss'}" apply (insert acc_le_listI3 acc_JVM [OF wf]) by (simp add: acc_def r_def) lemma iter_properties_bv[rule_format]: -shows "\ \p\w0. p < n; ss0 \ list n A; \p w0 \ stable r step ss0 p \ \ +shows "\ \p\w0. p < n; ss0 \ nlists n A; \p w0 \ stable r step ss0 p \ \ iter f step ss0 w0 = (ss',w') \ - ss' \ list n A \ stables r step ss' \ ss0 [\\<^sub>r] ss' \ - (\ts\list n A. ss0 [\\<^sub>r] ts \ stables r step ts \ ss' [\\<^sub>r] ts)" + ss' \ nlists n A \ stables r step ss' \ ss0 [\\<^sub>r] ss' \ + (\ts\nlists n A. ss0 [\\<^sub>r] ts \ stables r step ts \ ss' [\\<^sub>r] ts)" (*<*) (is "PROP ?P") proof - show "PROP ?P" apply (insert semi bounded_step exec_pres_type step_mono[OF wf]) apply (unfold iter_def stables_def) apply (rule_tac P = "\(ss,w). - ss \ list n A \ (\p w \ stable r step ss p) \ ss0 [\\<^sub>r] ss \ - (\ts\list n A. ss0 [\\<^sub>r] ts \ stables r step ts \ ss [\\<^sub>r] ts) \ + ss \ nlists n A \ (\p w \ stable r step ss p) \ ss0 [\\<^sub>r] ss \ + (\ts\nlists n A. ss0 [\\<^sub>r] ts \ stables r step ts \ ss [\\<^sub>r] ts) \ (\p\w. p < n)" and r = "{(ss',ss) . ss [\\<^sub>r] ss'} <*lex*> finite_psubset" in while_rule) \ \Invariant holds initially:\ apply (simp add:stables_def semilat_Def del: n_def A_def r_def f_def step_def) apply (blast intro:le_list_refl') \ \Invariant is preserved:\ apply(simp add: stables_def split_paired_all del: A_def r_def f_def step_def n_def) apply(rename_tac ss w) apply(subgoal_tac "(SOME p. p \ w) \ w") prefer 2 apply (fast intro: someI) apply(subgoal_tac "\(q,t) \ set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length ss \ t \ A") prefer 2 apply clarify apply (rule conjI) apply(clarsimp, blast dest!: boundedD) apply (subgoal_tac "(SOME p. p \ w) < n") apply (subgoal_tac "(ss ! (SOME p. p \ w)) \ A") apply (fastforce simp only:n_def dest:pres_typeD ) apply simp apply simp apply (subst decomp_propa) apply blast apply (simp del:A_def r_def f_def step_def n_def) apply (rule conjI) apply (rule Semilat.merges_preserves_type[OF Semilat.intro, OF semi]) apply blast apply clarify apply (rule conjI) apply(clarsimp, blast dest!: boundedD) apply (erule pres_typeD) prefer 3 apply assumption - apply (erule listE_nth_in) + apply (erule nlistsE_nth_in) apply blast apply (simp only:n_def) apply (rule conjI) apply clarify - apply (subgoal_tac "ss \ list (length is) A" "\p\w. p < (length is) " "\p<(length is). p \ w \ stable r step ss p " + apply (subgoal_tac "ss \ nlists (length is) A" "\p\w. p < (length is) " "\p<(length is). p \ w \ stable r step ss p " "p < length is") apply (blast intro!: Semilat.stable_pres_lemma[OF Semilat.intro, OF semi]) apply (simp only:n_def) apply (simp only:n_def) apply (simp only:n_def) apply (simp only:n_def) apply (rule conjI) - apply (subgoal_tac "ss \ list (length is) A" + apply (subgoal_tac "ss \ nlists (length is) A" "\(q,t)\set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length is \ t \ A" - "ss [\\<^bsub>r\<^esub>] merges f (step (SOME p. p \ w) (ss ! (SOME p. p \ w))) ss" "ss0\ list (size is) A" - "merges f (step (SOME p. p \ w) (ss ! (SOME p. p \ w))) ss \ list (size is) A" - "ss \list (size is) A" "order r A" "ss0 [\\<^bsub>r\<^esub>] ss ") + "ss [\\<^bsub>r\<^esub>] merges f (step (SOME p. p \ w) (ss ! (SOME p. p \ w))) ss" "ss0\ nlists (size is) A" + "merges f (step (SOME p. p \ w) (ss ! (SOME p. p \ w))) ss \ nlists (size is) A" + "ss \nlists (size is) A" "order r A" "ss0 [\\<^bsub>r\<^esub>] ss ") apply (blast dest: le_list_trans) apply simp apply (simp only:semilat_Def) apply (simp only:n_def) apply (fastforce simp only: n_def dest:Semilat.merges_preserves_type[OF Semilat.intro, OF semi] ) apply (simp only:n_def) apply (blast intro:Semilat.merges_incr[OF Semilat.intro, OF semi]) apply (subgoal_tac "length ss = n") apply (simp only:n_def) - apply (subgoal_tac "ss \list n A") + apply (subgoal_tac "ss \nlists n A") defer apply simp apply (simp only:n_def) prefer 5 - apply (simp only:listE_length n_def) + apply (simp only:nlistsE_length n_def) apply(rule conjI) apply (clarsimp simp del: A_def r_def f_def step_def) apply (blast intro!: Semilat.merges_bounded_lemma[OF Semilat.intro, OF semi]) apply (subgoal_tac "bounded step n") apply (blast dest!: boundedD) apply (simp only:n_def) \ \Postcondition holds upon termination:\ apply(clarsimp simp add: stables_def split_paired_all) \ \Well-foundedness of the termination relation:\ apply (rule wf_lex_prod) apply (simp only:wf_jvm) apply (rule wf_finite_psubset) \ \Loop decreases along termination relation:\ apply(simp add: stables_def split_paired_all del: A_def r_def f_def step_def) apply(rename_tac ss w) apply(subgoal_tac "(SOME p. p \ w) \ w") prefer 2 apply (fast intro: someI) apply(subgoal_tac "\(q,t) \ set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length ss \ t \ A") prefer 2 apply clarify apply (rule conjI) apply(clarsimp, blast dest!: boundedD) apply (erule pres_typeD) prefer 3 apply assumption - apply (erule listE_nth_in) + apply (erule nlistsE_nth_in) apply blast apply blast apply (subst decomp_propa) apply blast apply clarify - apply (simp del: listE_length A_def r_def f_def step_def + apply (simp del: nlistsE_length A_def r_def f_def step_def add: lex_prod_def finite_psubset_def bounded_nat_set_is_finite) apply (rule termination_lemma) apply (insert Semilat.intro) apply assumption+ defer apply assumption defer apply clarsimp done qed (*>*) lemma kildall_properties_bv: -shows "\ ss0 \ list n A \ \ - kildall r f step ss0 \ list n A \ +shows "\ ss0 \ nlists n A \ \ + kildall r f step ss0 \ nlists n A \ stables r step (kildall r f step ss0) \ ss0 [\\<^sub>r] kildall r f step ss0 \ - (\ts\list n A. ss0 [\\<^sub>r] ts \ stables r step ts \ + (\ts\nlists n A. ss0 [\\<^sub>r] ts \ stables r step ts \ kildall r f step ss0 [\\<^sub>r] ts)" (*<*) (is "PROP ?P") proof - show "PROP ?P" apply (unfold kildall_def) apply(case_tac "iter f step ss0 (unstables r step ss0)") apply (simp del:r_def f_def n_def step_def A_def) apply (rule iter_properties_bv) apply (simp_all add: unstables_def stable_def) done qed end 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 n_def) apply(unfold is_bcv_def wt_step_def) apply(insert semi kildall_properties_bv) apply(simp only:stables_def) apply clarify - apply(subgoal_tac "kildall r f step \s\<^sub>0 \ list n A") + apply(subgoal_tac "kildall r f step \s\<^sub>0 \ nlists n A") prefer 2 apply (fastforce intro: kildall_properties_bv) apply (rule iffI) apply (rule_tac x = "kildall r f step \s\<^sub>0" in bexI) apply (rule conjI) apply (fastforce intro: kildall_properties_bv) apply (force intro: kildall_properties_bv) apply simp apply clarify apply(subgoal_tac "kildall r f step \s\<^sub>0!pa <=_r \s!pa") defer apply (blast intro!: le_listD less_lengthI) apply (subgoal_tac "\s!pa \ Err") defer apply simp apply (rule ccontr) apply (simp only:top_def r_def JVM_le_unfold) apply fastforce 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" + "0 < size is \ start \ nlists (size is) A" using Ts C (*<*) proof(cases b) -qed (force simp: JVM_states_unfold intro!: listI list_appendI +qed (force simp: JVM_states_unfold intro!: nlistsI nlists_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" and stab: "b = Static \ b = NonStatic" 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" .. + from instrs have "start \ nlists (size is) A" .. with bcv success result have - "\ts\list (size is) A. start [\\<^sub>r] ts \ wt_step r Err step ts" + "\ts\nlists (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 + in_A: "\s' \ nlists (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 stab 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" and stab: "b = Static \ b = NonStatic" 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 from instrs have "start \ nlists (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) + have "?\s \ nlists (size is) A" + by (auto intro!: nlistsI 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" using start_context_intro_auxi[OF staticb] using stab by (unfold wt_kildall_def) simp qed (*>*) theorem jvm_kildall_correct: "wf_jvm_prog\<^sub>k P = wf_jvm_prog P" (*<*) 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))" show ?thesis proof(rule iffI) \ \soundness\ assume wt: "wf_jvm_prog\<^sub>k P" then have wt': "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 ass1: "wf_prog wf_md P" and sees: "P \ Ca sees M, b : Ts\T = m in Ca" and ass2: "set Ts \ types P" and ass3: "bd = (M, b, Ts, T, m)" and ass4: "?A P Ca bd" from ass3 ass4 have stab: "b = Static \ b = NonStatic" by (simp add:wt_kildall_def) from ass1 sees ass2 ass3 ass4 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_auxi[OF stab]] 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 ass1: "wf_prog wf_md P" and sees: "P \ Ca sees M, b : Ts\T = m in Ca" and ass2: "set Ts \ types P" and ass3: "bd = (M, b, Ts, T, m)" and ass4: "(?B\<^sub>\ \) P Ca bd" from ass3 ass4 have stab: "b = Static \ b = NonStatic" by (simp add:wt_method_def) from ass1 sees ass2 ass3 ass4 have "?A P Ca bd" using sees_method_is_class[OF sees] using JVM_sl.staticb by (auto intro!: start_context.wt_kil_complete start_context_intro_auxi[OF stab]) } 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/ClassAdd.thy b/thys/JinjaDCI/BV/ClassAdd.thy --- a/thys/JinjaDCI/BV/ClassAdd.thy +++ b/thys/JinjaDCI/BV/ClassAdd.thy @@ -1,685 +1,676 @@ (* Title: JinjaDCI/BV/ClassAdd.thy Author: Susannah Mansky 2019-20, UIUC *) section \ Property preservation under @{text "class_add"} \ theory ClassAdd imports BVConform begin lemma err_mono: "A \ B \ err A \ err B" by(unfold err_def) auto lemma opt_mono: "A \ B \ opt A \ opt B" by(unfold opt_def) auto -lemma list_mono: -assumes "A \ B" shows "list n A \ list n B" -proof(rule) - fix xs assume "xs \ list n A" - then obtain size: "size xs = n" and inA: "set xs \ A" by simp - with assms have "set xs \ B" by simp - with size show "xs \ list n B" by(clarsimp intro!: listI) -qed - (****************************************************************) \ \ adding a class in the simplest way \ abbreviation class_add :: "jvm_prog \ jvm_method cdecl \ jvm_prog" where "class_add P cd \ cd#P" subsection "Fields" lemma class_add_has_fields: assumes fs: "P \ D has_fields FDTs" and nc: "\is_class P C" shows "class_add P (C, cdec) \ D has_fields FDTs" using assms proof(induct rule: Fields.induct) case (has_fields_Object D fs ms FDTs) from has_fields_is_class_Object[OF fs] nc have "C \ Object" by fast with has_fields_Object show ?case by(auto simp: class_def fun_upd_apply intro!: TypeRel.has_fields_Object) next case rec: (has_fields_rec C1 D fs ms FDTs FDTs') with has_fields_is_class have [simp]: "D \ C" by auto with rec have "C1 \ C" by(clarsimp simp: is_class_def) with rec show ?case by(auto simp: class_def fun_upd_apply intro: TypeRel.has_fields_rec) qed lemma class_add_has_fields_rev: "\ class_add P (C, cdec) \ D has_fields FDTs; \P \ D \\<^sup>* C \ \ P \ D has_fields FDTs" proof(induct rule: Fields.induct) case (has_fields_Object D fs ms FDTs) then show ?case by(auto simp: class_def fun_upd_apply intro!: TypeRel.has_fields_Object) next case rec: (has_fields_rec C1 D fs ms FDTs FDTs') then have sub1: "P \ C1 \\<^sup>1 D" by(auto simp: class_def fun_upd_apply intro!: subcls1I split: if_split_asm) with rec.prems have cls: "\ P \ D \\<^sup>* C" by (meson converse_rtrancl_into_rtrancl) with cls rec show ?case by(auto simp: class_def fun_upd_apply intro: TypeRel.has_fields_rec split: if_split_asm) qed lemma class_add_has_field: assumes "P \ C\<^sub>0 has F,b:T in D" and "\ is_class P C" shows "class_add P (C, cdec) \ C\<^sub>0 has F,b:T in D" using assms by(auto simp: has_field_def dest!: class_add_has_fields[of P C\<^sub>0]) lemma class_add_has_field_rev: assumes has: "class_add P (C, cdec) \ C\<^sub>0 has F,b:T in D" and ncp: "\D'. P \ C\<^sub>0 \\<^sup>* D' \ D' \ C" shows "P \ C\<^sub>0 has F,b:T in D" using assms by(auto simp: has_field_def dest!: class_add_has_fields_rev) lemma class_add_sees_field: assumes "P \ C\<^sub>0 sees F,b:T in D" and "\ is_class P C" shows "class_add P (C, cdec) \ C\<^sub>0 sees F,b:T in D" using assms by(auto simp: sees_field_def dest!: class_add_has_fields[of P C\<^sub>0]) lemma class_add_sees_field_rev: assumes has: "class_add P (C, cdec) \ C\<^sub>0 sees F,b:T in D" and ncp: "\D'. P \ C\<^sub>0 \\<^sup>* D' \ D' \ C" shows "P \ C\<^sub>0 sees F,b:T in D" using assms by(auto simp: sees_field_def dest!: class_add_has_fields_rev) lemma class_add_field: assumes fd: "P \ C\<^sub>0 sees F,b:T in D" and "\ is_class P C" shows "field P C\<^sub>0 F = field (class_add P (C, cdec)) C\<^sub>0 F" using class_add_sees_field[OF assms, of cdec] fd by simp subsection "Methods" lemma class_add_sees_methods: assumes ms: "P \ D sees_methods Mm" and nc: "\is_class P C" shows "class_add P (C, cdec) \ D sees_methods Mm" using assms proof(induct rule: Methods.induct) case (sees_methods_Object D fs ms Mm) from sees_methods_is_class_Object[OF ms] nc have "C \ Object" by fast with sees_methods_Object show ?case by(auto simp: class_def fun_upd_apply intro!: TypeRel.sees_methods_Object) next case rec: (sees_methods_rec C1 D fs ms Mm Mm') with sees_methods_is_class have [simp]: "D \ C" by auto with rec have "C1 \ C" by(clarsimp simp: is_class_def) with rec show ?case by(auto simp: class_def fun_upd_apply intro: TypeRel.sees_methods_rec) qed lemma class_add_sees_methods_rev: "\ class_add P (C, cdec) \ D sees_methods Mm; \D'. P \ D \\<^sup>* D' \ D' \ C \ \ P \ D sees_methods Mm" proof(induct rule: Methods.induct) case (sees_methods_Object D fs ms Mm) then show ?case by(auto simp: class_def fun_upd_apply intro!: TypeRel.sees_methods_Object) next case rec: (sees_methods_rec C1 D fs ms Mm Mm') then have sub1: "P \ C1 \\<^sup>1 D" by(auto simp: class_def fun_upd_apply intro!: subcls1I) have cls: "\D'. P \ D \\<^sup>* D' \ D' \ C" proof - fix D' assume "P \ D \\<^sup>* D'" with sub1 have "P \ C1 \\<^sup>* D'" by simp with rec.prems show "D' \ C" by simp qed with cls rec show ?case by(auto simp: class_def fun_upd_apply intro: TypeRel.sees_methods_rec) qed lemma class_add_sees_methods_Obj: assumes "P \ Object sees_methods Mm" and nObj: "C \ Object" shows "class_add P (C, cdec) \ Object sees_methods Mm" proof - from assms obtain C' fs ms where cls: "class P Object = Some(C',fs,ms)" by(auto elim!: Methods.cases) with nObj have cls': "class (class_add P (C, cdec)) Object = Some(C',fs,ms)" by(simp add: class_def fun_upd_apply) from assms cls have "Mm = map_option (\m. (m, Object)) \ map_of ms" by(auto elim!: Methods.cases) with assms cls' show ?thesis by(auto simp: is_class_def fun_upd_apply intro!: sees_methods_Object) qed lemma class_add_sees_methods_rev_Obj: assumes "class_add P (C, cdec) \ Object sees_methods Mm" and nObj: "C \ Object" shows "P \ Object sees_methods Mm" proof - from assms obtain C' fs ms where cls: "class (class_add P (C, cdec)) Object = Some(C',fs,ms)" by(auto elim!: Methods.cases) with nObj have cls': "class P Object = Some(C',fs,ms)" by(simp add: class_def fun_upd_apply) from assms cls have "Mm = map_option (\m. (m, Object)) \ map_of ms" by(auto elim!: Methods.cases) with assms cls' show ?thesis by(auto simp: is_class_def fun_upd_apply intro!: sees_methods_Object) qed lemma class_add_sees_method: assumes "P \ C\<^sub>0 sees M\<^sub>0, b : Ts\T = m in D" and "\ is_class P C" shows "class_add P (C, cdec) \ C\<^sub>0 sees M\<^sub>0, b : Ts\T = m in D" using assms by(auto simp: Method_def dest!: class_add_sees_methods[of P C\<^sub>0]) lemma class_add_method: assumes md: "P \ C\<^sub>0 sees M\<^sub>0, b : Ts\T = m in D" and "\ is_class P C" shows "method P C\<^sub>0 M\<^sub>0 = method (class_add P (C, cdec)) C\<^sub>0 M\<^sub>0" using class_add_sees_method[OF assms, of cdec] md by simp lemma class_add_sees_method_rev: "\ class_add P (C, cdec) \ C\<^sub>0 sees M\<^sub>0, b : Ts\T = m in D; \ P \ C\<^sub>0 \\<^sup>* C \ \ P \ C\<^sub>0 sees M\<^sub>0, b : Ts\T = m in D" by(auto simp: Method_def dest!: class_add_sees_methods_rev) lemma class_add_sees_method_Obj: "\ P \ Object sees M\<^sub>0, b : Ts\T = m in D; C \ Object \ \ class_add P (C, cdec) \ Object sees M\<^sub>0, b : Ts\T = m in D" by(auto simp: Method_def dest!: class_add_sees_methods_Obj[where P=P]) lemma class_add_sees_method_rev_Obj: "\ class_add P (C, cdec) \ Object sees M\<^sub>0, b : Ts\T = m in D; C \ Object \ \ P \ Object sees M\<^sub>0, b : Ts\T = m in D" by(auto simp: Method_def dest!: class_add_sees_methods_rev_Obj[where P=P]) subsection "Types and states" lemma class_add_is_type: "is_type P T \ is_type (class_add P (C, cdec)) T" by(cases cdec, simp add: is_type_def is_class_def class_def fun_upd_apply split: ty.splits) lemma class_add_types: "types P \ types (class_add P (C, cdec))" using class_add_is_type by(cases cdec, clarsimp) lemma class_add_states: "states P mxs mxl \ states (class_add P (C, cdec)) mxs mxl" proof - let ?A = "types P" and ?B = "types (class_add P (C, cdec))" have ab: "?A \ ?B" by(rule class_add_types) - moreover have "\n. list n ?A \ list n ?B" using ab by(rule list_mono) - moreover have "list mxl (err ?A) \ list mxl (err ?B)" using err_mono[OF ab] by(rule list_mono) + moreover have "\n. nlists n ?A \ nlists n ?B" using ab by(rule nlists_mono) + moreover have "nlists mxl (err ?A) \ nlists mxl (err ?B)" using err_mono[OF ab] by(rule nlists_mono) ultimately show ?thesis by(auto simp: JVM_states_unfold intro!: err_mono opt_mono) qed lemma class_add_check_types: "check_types P mxs mxl \s \ check_types (class_add P (C, cdec)) mxs mxl \s" using class_add_states by(fastforce simp: check_types_def) subsection "Subclasses and subtypes" lemma class_add_subcls: "\ P \ D \\<^sup>* D'; \ is_class P C \ \ class_add P (C, cdec) \ D \\<^sup>* D'" proof(induct rule: rtrancl.induct) case (rtrancl_into_rtrancl a b c) then have "b \ C" by(clarsimp simp: is_class_def dest!: subcls1D) with rtrancl_into_rtrancl show ?case by(fastforce dest!: subcls1D simp: class_def fun_upd_apply intro!: rtrancl_trans[of a b] subcls1I) qed(simp) lemma class_add_subcls_rev: "\ class_add P (C, cdec) \ D \\<^sup>* D'; \P \ D \\<^sup>* C \ \ P \ D \\<^sup>* D'" proof(induct rule: rtrancl.induct) case (rtrancl_into_rtrancl a b c) then have "b \ C" by(clarsimp simp: is_class_def dest!: subcls1D) with rtrancl_into_rtrancl show ?case by(fastforce dest!: subcls1D simp: class_def fun_upd_apply intro!: rtrancl_trans[of a b] subcls1I) qed(simp) lemma class_add_subtype: "\ subtype P x y; \ is_class P C \ \ subtype (class_add P (C, cdec)) x y" proof(induct rule: widen.induct) case (widen_subcls C D) then show ?case using class_add_subcls by simp qed(simp+) lemma class_add_widens: "\ P \ Ts [\] Ts'; \ is_class P C \ \ (class_add P (C, cdec)) \ Ts [\] Ts'" using class_add_subtype by (metis (no_types) list_all2_mono) lemma class_add_sup_ty_opt: "\ P \ l1 \\<^sub>\ l2; \ is_class P C \ \ class_add P (C, cdec) \ l1 \\<^sub>\ l2" using class_add_subtype by(auto simp: sup_ty_opt_def Err.le_def lesub_def split: err.splits) lemma class_add_sup_loc: "\ P \ LT [\\<^sub>\] LT'; \ is_class P C \ \ class_add P (C, cdec) \ LT [\\<^sub>\] LT'" using class_add_sup_ty_opt[where P=P and C=C] by (simp add: list.rel_mono_strong) lemma class_add_sup_state: "\ P \ \ \\<^sub>i \'; \ is_class P C \ \ class_add P (C, cdec) \ \ \\<^sub>i \'" using class_add_subtype class_add_sup_ty_opt by(auto simp: sup_state_def Listn.le_def Product.le_def lesub_def class_add_widens class_add_sup_ty_opt list_all2_mono) lemma class_add_sup_state_opt: "\ P \ \ \' \'; \ is_class P C \ \ class_add P (C, cdec) \ \ \' \'" by(auto simp: sup_state_opt_def Opt.le_def lesub_def class_add_widens class_add_sup_ty_opt list_all2_mono) subsection "Effect" lemma class_add_is_relevant_class: "\ is_relevant_class i P C\<^sub>0; \ is_class P C \ \ is_relevant_class i (class_add P (C, cdec)) C\<^sub>0" by(cases i, auto simp: class_add_subcls) lemma class_add_is_relevant_class_rev: assumes irc: "is_relevant_class i (class_add P (C, cdec)) C\<^sub>0" and ncp: "\cd D'. cd \ set P \ \P \ fst cd \\<^sup>* C" and wfxp: "wf_syscls P" shows "is_relevant_class i P C\<^sub>0" using assms proof(cases i) case (Getfield F D) with assms show ?thesis by(fastforce simp: wf_syscls_def sys_xcpts_def dest!: class_add_subcls_rev) next case (Putfield F D) with assms show ?thesis by(fastforce simp: wf_syscls_def sys_xcpts_def dest!: class_add_subcls_rev) next case (Checkcast D) with assms show ?thesis by(fastforce simp: wf_syscls_def sys_xcpts_def dest!: class_add_subcls_rev) qed(simp_all) lemma class_add_is_relevant_entry: "\ is_relevant_entry P i pc e; \ is_class P C \ \ is_relevant_entry (class_add P (C, cdec)) i pc e" by(clarsimp simp: is_relevant_entry_def class_add_is_relevant_class) lemma class_add_is_relevant_entry_rev: "\ is_relevant_entry (class_add P (C, cdec)) i pc e; \cd D'. cd \ set P \ \P \ fst cd \\<^sup>* C; wf_syscls P \ \ is_relevant_entry P i pc e" by(auto simp: is_relevant_entry_def dest!: class_add_is_relevant_class_rev) lemma class_add_relevant_entries: "\ is_class P C \ set (relevant_entries P i pc xt) \ set (relevant_entries (class_add P (C, cdec)) i pc xt)" by(clarsimp simp: relevant_entries_def class_add_is_relevant_entry) lemma class_add_relevant_entries_eq: assumes wf: "wf_prog wf_md P" and nclass: "\ is_class P C" shows "relevant_entries P i pc xt = relevant_entries (class_add P (C, cdec)) i pc xt" proof - have ncp: "\cd D'. cd \ set P \ \P \ fst cd \\<^sup>* C" by(rule wf_subcls_nCls'[OF assms]) moreover from wf have wfsys: "wf_syscls P" by(simp add: wf_prog_def) moreover note class_add_is_relevant_entry[OF _ nclass, of i pc _ cdec] class_add_is_relevant_entry_rev[OF _ ncp wfsys, of cdec i pc] ultimately show ?thesis by (metis filter_cong relevant_entries_def) qed lemma class_add_norm_eff_pc: assumes ne: "\(pc',\') \ set (norm_eff i P pc \). pc' < mpc" shows "\(pc',\') \ set (norm_eff i (class_add P (C, cdec)) pc \). pc' < mpc" using assms by(cases i, auto simp: norm_eff_def) lemma class_add_norm_eff_sup_state_opt: assumes ne: "\(pc',\') \ set (norm_eff i P pc \). P \ \' \' \s!pc'" and nclass: "\ is_class P C" and app: "app\<^sub>i (i, P, pc, mxs, T, \)" shows "\(pc',\') \ set (norm_eff i (class_add P (C, cdec)) pc \). (class_add P (C, cdec)) \ \' \' \s!pc'" proof - obtain ST LT where "\ = (ST,LT)" by(cases \) with assms show ?thesis proof(cases i) qed(fastforce simp: norm_eff_def dest!: class_add_field[where cdec=cdec] class_add_method[where cdec=cdec] class_add_sup_loc[OF _ nclass] class_add_subtype[OF _ nclass] class_add_widens[OF _ nclass] class_add_sup_state_opt[OF _ nclass])+ qed lemma class_add_xcpt_eff_eq: assumes wf: "wf_prog wf_md P" and nclass: "\ is_class P C" shows "xcpt_eff i P pc \ xt = xcpt_eff i (class_add P (C, cdec)) pc \ xt" using class_add_relevant_entries_eq[OF assms, of i pc xt cdec] by(cases \, simp add: xcpt_eff_def) lemma class_add_eff_pc: assumes eff: "\(pc',\') \ set (eff i P pc xt (Some \)). pc' < mpc" and wf: "wf_prog wf_md P" and nclass: "\ is_class P C" shows "\(pc',\') \ set (eff i (class_add P (C, cdec)) pc xt (Some \)). pc' < mpc" using eff class_add_norm_eff_pc class_add_xcpt_eff_eq[OF wf nclass] by(auto simp: norm_eff_def eff_def) lemma class_add_eff_sup_state_opt: assumes eff: "\(pc',\') \ set (eff i P pc xt (Some \)). P \ \' \' \s!pc'" and wf: "wf_prog wf_md P"and nclass: "\ is_class P C" and app: "app\<^sub>i (i, P, pc, mxs, T, \)" shows "\(pc',\') \ set (eff i (class_add P (C, cdec)) pc xt (Some \)). (class_add P (C, cdec)) \ \' \' \s!pc'" proof - from eff have ne: "\(pc', \')\set (norm_eff i P pc \). P \ \' \' \s ! pc'" by(simp add: norm_eff_def eff_def) from eff have "\(pc', \')\set (xcpt_eff i P pc \ xt). P \ \' \' \s ! pc'" by(simp add: xcpt_eff_def eff_def) with class_add_norm_eff_sup_state_opt[OF ne nclass app] class_add_xcpt_eff_eq[OF wf nclass]class_add_sup_state_opt[OF _ nclass] show ?thesis by(cases cdec, auto simp: eff_def norm_eff_def xcpt_app_def) qed lemma class_add_app\<^sub>i: assumes "app\<^sub>i (i, P, pc, mxs, T\<^sub>r, ST, LT)" and "\ is_class P C" shows "app\<^sub>i (i, class_add P (C, cdec), pc, mxs, T\<^sub>r, ST, LT)" using assms proof(cases i) case New then show ?thesis using assms by(fastforce simp: is_class_def class_def fun_upd_apply) next case Getfield then show ?thesis using assms by(auto simp: class_add_subtype dest!: class_add_sees_field[where P=P]) next case Getstatic then show ?thesis using assms by(auto dest!: class_add_sees_field[where P=P]) next case Putfield then show ?thesis using assms by(auto dest!: class_add_subtype[where P=P] class_add_sees_field[where P=P]) next case Putstatic then show ?thesis using assms by(auto dest!: class_add_subtype[where P=P] class_add_sees_field[where P=P]) next case Checkcast then show ?thesis using assms by(clarsimp simp: is_class_def class_def fun_upd_apply) next case Invoke then show ?thesis using assms by(fastforce dest!: class_add_widens[where P=P] class_add_sees_method[where P=P]) next case Invokestatic then show ?thesis using assms by(fastforce dest!: class_add_widens[where P=P] class_add_sees_method[where P=P]) next case Return then show ?thesis using assms by(clarsimp simp: class_add_subtype) qed(simp+) lemma class_add_xcpt_app: assumes xa: "xcpt_app i P pc mxs xt \" and wf: "wf_prog wf_md P" and nclass: "\ is_class P C" shows "xcpt_app i (class_add P (C, cdec)) pc mxs xt \" using xa class_add_relevant_entries_eq[OF wf nclass] nclass by(auto simp: xcpt_app_def is_class_def class_def fun_upd_apply) auto lemma class_add_app: assumes app: "app i P mxs T pc mpc xt t" and wf: "wf_prog wf_md P" and nclass: "\ is_class P C" shows "app i (class_add P (C, cdec)) mxs T pc mpc xt t" proof(cases t) case (Some \) let ?P = "class_add P (C, cdec)" from assms Some have eff: "\(pc', \')\set (eff i P pc xt \\\). pc' < mpc" by(simp add: app_def) from assms Some have app\<^sub>i: "app\<^sub>i (i,P,pc,mxs,T,\)" by(simp add: app_def) with class_add_app\<^sub>i[OF _ nclass] Some have "app\<^sub>i (i,?P,pc,mxs,T,\)" by(cases \,simp) moreover from app class_add_xcpt_app[OF _ wf nclass] Some have "xcpt_app i ?P pc mxs xt \" by(simp add: app_def del: xcpt_app_def) moreover from app class_add_eff_pc[OF eff wf nclass] Some have "\(pc',\') \ set (eff i ?P pc xt t). pc' < mpc" by auto moreover note app Some ultimately show ?thesis by(simp add: app_def) qed(simp) subsection "Well-formedness and well-typedness" lemma class_add_wf_mdecl: "\ wf_mdecl wf_md P C\<^sub>0 md; \C\<^sub>0 md. wf_md P C\<^sub>0 md \ wf_md (class_add P (C, cdec)) C\<^sub>0 md \ \ wf_mdecl wf_md (class_add P (C, cdec)) C\<^sub>0 md" by(clarsimp simp: wf_mdecl_def class_add_is_type) lemma class_add_wf_mdecl': assumes wfd: "wf_mdecl wf_md P C\<^sub>0 md" and ms: "(C\<^sub>0,S,fs,ms) \ set P" and md: "md \ set ms" and wf_md': "\C\<^sub>0 S fs ms m.\(C\<^sub>0,S,fs,ms) \ set P; m \ set ms\ \ wf_md' (class_add P (C, cdec)) C\<^sub>0 m" shows "wf_mdecl wf_md' (class_add P (C, cdec)) C\<^sub>0 md" using assms by(clarsimp simp: wf_mdecl_def class_add_is_type) lemma class_add_wf_cdecl: assumes wfcd: "wf_cdecl wf_md P cd" and cdP: "cd \ set P" and ncp: "\ P \ fst cd \\<^sup>* C" and dist: "distinct_fst P" and wfmd: "\C\<^sub>0 md. wf_md P C\<^sub>0 md \ wf_md (class_add P (C, cdec)) C\<^sub>0 md" and nclass: "\ is_class P C" shows "wf_cdecl wf_md (class_add P (C, cdec)) cd" proof - let ?P = "class_add P (C, cdec)" obtain C1 D fs ms where [simp]: "cd = (C1,(D,fs,ms))" by(cases cd) from wfcd have "\f\set fs. wf_fdecl ?P f" by(auto simp: wf_cdecl_def wf_fdecl_def class_add_is_type) moreover from wfcd wfmd class_add_wf_mdecl have "\m\set ms. wf_mdecl wf_md ?P C1 m" by(auto simp: wf_cdecl_def) moreover have "C1 \ Object \ is_class ?P D \ \ ?P \ D \\<^sup>* C1 \ (\(M,b,Ts,T,m)\set 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 - assume nObj[simp]: "C1 \ Object" with cdP dist have sub1: "P \ C1 \\<^sup>1 D" by(auto simp: class_def intro!: subcls1I map_of_SomeI) with ncp have ncp': "\ P \ D \\<^sup>* C" by(auto simp: converse_rtrancl_into_rtrancl) with wfcd have clsD: "is_class ?P D" by(auto simp: wf_cdecl_def is_class_def class_def fun_upd_apply) moreover from wfcd sub1 have "\ ?P \ D \\<^sup>* C1" by(auto simp: wf_cdecl_def dest!: class_add_subcls_rev[OF _ ncp']) moreover have "\M b Ts T m D' b' Ts' T' m'. (M,b,Ts,T,m) \ set ms \ ?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 ms: "(M,b,Ts,T,m) \ set ms" and meth': "?P \ D sees M,b':Ts' \ T' = m' in D'" with sub1 have "P \ D sees M,b':Ts' \ T' = m' in D'" by(fastforce dest!: class_add_sees_method_rev[OF _ ncp']) moreover with wfcd ms meth' have "b = b' \ P \ Ts' [\] Ts \ P \ T \ T'" by(cases m', fastforce simp: wf_cdecl_def elim!: ballE[where x="(M,b,Ts,T,m)"]) ultimately show "b = b' \ ?P \ Ts' [\] Ts \ ?P \ T \ T'" by(auto dest!: class_add_subtype[OF _ nclass] class_add_widens[OF _ nclass]) qed ultimately show ?thesis by clarsimp qed moreover note wfcd ultimately show ?thesis by(simp add: wf_cdecl_def) qed lemma class_add_wf_cdecl': assumes wfcd: "wf_cdecl wf_md P cd" and cdP: "cd \ set P" and ncp: "\P \ fst cd \\<^sup>* C" and dist: "distinct_fst P" and wfmd: "\C\<^sub>0 S fs ms m.\(C\<^sub>0,S,fs,ms) \ set P; m \ set ms\ \ wf_md' (class_add P (C, cdec)) C\<^sub>0 m" and nclass: "\ is_class P C" shows "wf_cdecl wf_md' (class_add P (C, cdec)) cd" proof - let ?P = "class_add P (C, cdec)" obtain C1 D fs ms where [simp]: "cd = (C1,(D,fs,ms))" by(cases cd) from wfcd have "\f\set fs. wf_fdecl ?P f" by(auto simp: wf_cdecl_def wf_fdecl_def class_add_is_type) moreover from cdP wfcd wfmd have "\m\set ms. wf_mdecl wf_md' ?P C1 m" by(auto simp: wf_cdecl_def wf_mdecl_def class_add_is_type) moreover have "C1 \ Object \ is_class ?P D \ \ ?P \ D \\<^sup>* C1 \ (\(M,b,Ts,T,m)\set 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 - assume nObj[simp]: "C1 \ Object" with cdP dist have sub1: "P \ C1 \\<^sup>1 D" by(auto simp: class_def intro!: subcls1I map_of_SomeI) with ncp have ncp': "\ P \ D \\<^sup>* C" by(auto simp: converse_rtrancl_into_rtrancl) with wfcd have clsD: "is_class ?P D" by(auto simp: wf_cdecl_def is_class_def class_def fun_upd_apply) moreover from wfcd sub1 have "\ ?P \ D \\<^sup>* C1" by(auto simp: wf_cdecl_def dest!: class_add_subcls_rev[OF _ ncp']) moreover have "\M b Ts T m D' b' Ts' T' m'. (M,b,Ts,T,m) \ set ms \ ?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 ms: "(M,b,Ts,T,m) \ set ms" and meth': "?P \ D sees M,b':Ts' \ T' = m' in D'" with sub1 have "P \ D sees M,b':Ts' \ T' = m' in D'" by(fastforce dest!: class_add_sees_method_rev[OF _ ncp']) moreover with wfcd ms meth' have "b = b' \ P \ Ts' [\] Ts \ P \ T \ T'" by(cases m', fastforce simp: wf_cdecl_def elim!: ballE[where x="(M,b,Ts,T,m)"]) ultimately show "b = b' \ ?P \ Ts' [\] Ts \ ?P \ T \ T'" by(auto dest!: class_add_subtype[OF _ nclass] class_add_widens[OF _ nclass]) qed ultimately show ?thesis by clarsimp qed moreover note wfcd ultimately show ?thesis by(simp add: wf_cdecl_def) qed lemma class_add_wt_start: "\ wt_start P C\<^sub>0 b Ts mxl \s; \ is_class P C \ \ wt_start (class_add P (C, cdec)) C\<^sub>0 b Ts mxl \s" using class_add_sup_state_opt by(clarsimp simp: wt_start_def split: staticb.splits) lemma class_add_wt_instr: assumes wti: "P,T,mxs,mpc,xt \ i,pc :: \s" and wf: "wf_prog wf_md P" and nclass: "\ is_class P C" shows "class_add P (C, cdec),T,mxs,mpc,xt \ i,pc :: \s" proof - let ?P = "class_add P (C, cdec)" from wti have eff: "\(pc', \')\set (eff i P pc xt (\s ! pc)). P \ \' \' \s ! pc'" by(simp add: wt_instr_def) from wti have app\<^sub>i: "\s!pc \ None \ app\<^sub>i (i,P,pc,mxs,T,the (\s!pc))" by(simp add: wt_instr_def app_def) from wti class_add_app[OF _ wf nclass] have "app i ?P mxs T pc mpc xt (\s!pc)" by(simp add: wt_instr_def) moreover have "\(pc',\') \ set (eff i ?P pc xt (\s!pc)). ?P \ \' \' \s!pc'" proof(cases "\s!pc") case Some with eff class_add_eff_sup_state_opt[OF _ wf nclass app\<^sub>i] show ?thesis by auto qed(simp add: eff_def) moreover note wti ultimately show ?thesis by(clarsimp simp: wt_instr_def) qed lemma class_add_wt_method: assumes wtm: "wt_method P C\<^sub>0 b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C\<^sub>0 M\<^sub>0)" and wf: "wf_prog wf_md P" and nclass: "\ is_class P C" shows "wt_method (class_add P (C, cdec)) C\<^sub>0 b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C\<^sub>0 M\<^sub>0)" proof - let ?P = "class_add P (C, cdec)" let ?\s = "\ C\<^sub>0 M\<^sub>0" from wtm class_add_check_types have "check_types ?P mxs ((case b of Static \ 0 | NonStatic \ 1)+size Ts+mxl\<^sub>0) (map OK ?\s)" by(simp add: wt_method_def) moreover from wtm class_add_wt_start nclass have "wt_start ?P C\<^sub>0 b Ts mxl\<^sub>0 ?\s" by(simp add: wt_method_def) moreover from wtm class_add_wt_instr[OF _ wf nclass] have "\pc < size is. ?P,T\<^sub>r,mxs,size is,xt \ is!pc,pc :: ?\s" by(clarsimp simp: wt_method_def) moreover note wtm ultimately show ?thesis by(clarsimp simp: wt_method_def) qed lemma class_add_wt_method': "\ (\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 C\<^sub>0 md; wf_prog wf_md P; \ is_class P C \ \ (\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)) (class_add P (C, cdec)) C\<^sub>0 md" by(clarsimp simp: class_add_wt_method) subsection \ @{text "distinct_fst"} \ lemma class_add_distinct_fst: "\ distinct_fst P; \ is_class P C \ \ distinct_fst (class_add P (C, cdec))" by(clarsimp simp: distinct_fst_def is_class_def class_def) subsection "Conformance" lemma class_add_conf: "\ P,h \ v :\ T; \ is_class P C \ \ class_add P (C, cdec),h \ v :\ T" by(clarsimp simp: conf_def class_add_subtype) lemma class_add_oconf: fixes obj::obj assumes oc: "P,h \ obj \" and ns: "\ is_class P C" and ncp: "\D'. P \ fst(obj) \\<^sup>* D' \ D' \ C" shows "(class_add P (C, cdec)),h \ obj \" proof - obtain C\<^sub>0 fs where [simp]: "obj=(C\<^sub>0,fs)" by(cases obj) from oc have oc': "\F D T. P \ C\<^sub>0 has F,NonStatic:T in D \ (\v. fs (F, D) = \v\ \ P,h \ v :\ T)" by(simp add: oconf_def) have "\F D T. class_add P (C, cdec) \ C\<^sub>0 has F,NonStatic:T in D \ \v. fs(F,D) = Some v \ class_add P (C, cdec),h \ v :\ T" proof - fix F D T assume "class_add P (C, cdec) \ C\<^sub>0 has F,NonStatic:T in D" with class_add_has_field_rev[OF _ ncp] have meth: "P \ C\<^sub>0 has F,NonStatic:T in D" by simp then show "\v. fs(F,D) = Some v \ class_add P (C, cdec),h \ v :\ T" using oc'[OF meth] class_add_conf[OF _ ns] by(fastforce simp: oconf_def) qed then show ?thesis by(simp add: oconf_def) qed lemma class_add_soconf: assumes soc: "P,h,C\<^sub>0 \\<^sub>s sfs \" and ns: "\ is_class P C" and ncp: "\D'. P \ C\<^sub>0 \\<^sup>* D' \ D' \ C" shows "(class_add P (C, cdec)),h,C\<^sub>0 \\<^sub>s sfs \" proof - from soc have oc': "\F T. P \ C\<^sub>0 has F,Static:T in C\<^sub>0 \ (\v. sfs F = \v\ \ P,h \ v :\ T)" by(simp add: soconf_def) have "\F T. class_add P (C, cdec) \ C\<^sub>0 has F,Static:T in C\<^sub>0 \ \v. sfs F = Some v \ class_add P (C, cdec),h \ v :\ T" proof - fix F T assume "class_add P (C, cdec) \ C\<^sub>0 has F,Static:T in C\<^sub>0" with class_add_has_field_rev[OF _ ncp] have meth: "P \ C\<^sub>0 has F,Static:T in C\<^sub>0" by simp then show "\v. sfs F = Some v \ class_add P (C, cdec),h \ v :\ T" using oc'[OF meth] class_add_conf[OF _ ns] by(fastforce simp: soconf_def) qed then show ?thesis by(simp add: soconf_def) qed lemma class_add_hconf: assumes "P \ h \" and "\ is_class P C" and "\a obj D'. h a = Some obj \ P \ fst(obj) \\<^sup>* D' \ D' \ C" shows "class_add P (C, cdec) \ h \" using assms by(auto simp: hconf_def intro!: class_add_oconf) lemma class_add_hconf_wf: assumes wf: "wf_prog wf_md P" and "P \ h \" and "\ is_class P C" and "\a obj. h a = Some obj \ fst(obj) \ C" shows "class_add P (C, cdec) \ h \" using wf_subcls_nCls[OF wf] assms by(fastforce simp: hconf_def intro!: class_add_oconf) lemma class_add_shconf: assumes "P,h \\<^sub>s sh \" and ns: "\ is_class P C" and "\C sobj D'. sh C = Some sobj \ P \ C \\<^sup>* D' \ D' \ C" shows "class_add P (C, cdec),h \\<^sub>s sh \" using assms by(fastforce simp: shconf_def) lemma class_add_shconf_wf: assumes wf: "wf_prog wf_md P" and "P,h \\<^sub>s sh \" and "\ is_class P C" and "\C sobj. sh C = Some sobj \ C \ C" shows "class_add P (C, cdec),h \\<^sub>s sh \" using wf_subcls_nCls[OF wf] assms by(fastforce simp: shconf_def) end \ No newline at end of file diff --git a/thys/JinjaDCI/BV/JVM_SemiType.thy b/thys/JinjaDCI/BV/JVM_SemiType.thy --- a/thys/JinjaDCI/BV/JVM_SemiType.thy +++ b/thys/JinjaDCI/BV/JVM_SemiType.thy @@ -1,388 +1,388 @@ (* Title: HOL/MicroJava/BV/JVM.thy Author: Gerwin Klein Copyright 2000 TUM *) section \ The JVM Type System as Semilattice \ theory JVM_SemiType imports SemiType begin type_synonym ty\<^sub>l = "ty err list" type_synonym ty\<^sub>s = "ty list" type_synonym ty\<^sub>i = "ty\<^sub>s \ ty\<^sub>l" type_synonym ty\<^sub>i' = "ty\<^sub>i option" type_synonym ty\<^sub>m = "ty\<^sub>i' list" type_synonym ty\<^sub>P = "mname \ cname \ ty\<^sub>m" definition stk_esl :: "'c prog \ nat \ ty\<^sub>s esl" where "stk_esl P mxs \ upto_esl mxs (SemiType.esl P)" definition loc_sl :: "'c prog \ nat \ ty\<^sub>l sl" where "loc_sl P mxl \ Listn.sl mxl (Err.sl (SemiType.esl P))" definition sl :: "'c prog \ nat \ nat \ ty\<^sub>i' err sl" where "sl P mxs mxl \ Err.sl(Opt.esl(Product.esl (stk_esl P mxs) (Err.esl(loc_sl P mxl))))" definition states :: "'c prog \ nat \ nat \ ty\<^sub>i' err set" where "states P mxs mxl \ fst(sl P mxs mxl)" definition le :: "'c prog \ nat \ nat \ ty\<^sub>i' err ord" where "le P mxs mxl \ fst(snd(sl P mxs mxl))" definition sup :: "'c prog \ nat \ nat \ ty\<^sub>i' err binop" where "sup P mxs mxl \ snd(snd(sl P mxs mxl))" definition sup_ty_opt :: "['c prog,ty err,ty err] \ bool" ("_ \ _ \\<^sub>\ _" [71,71,71] 70) where "sup_ty_opt P \ Err.le (subtype P)" definition sup_state :: "['c prog,ty\<^sub>i,ty\<^sub>i] \ bool" ("_ \ _ \\<^sub>i _" [71,71,71] 70) where "sup_state P \ Product.le (Listn.le (subtype P)) (Listn.le (sup_ty_opt P))" definition sup_state_opt :: "['c prog,ty\<^sub>i',ty\<^sub>i'] \ bool" ("_ \ _ \'' _" [71,71,71] 70) where "sup_state_opt P \ Opt.le (sup_state P)" abbreviation sup_loc :: "['c prog,ty\<^sub>l,ty\<^sub>l] \ bool" ("_ \ _ [\\<^sub>\] _" [71,71,71] 70) where "P \ LT [\\<^sub>\] LT' \ list_all2 (sup_ty_opt P) LT LT'" notation (ASCII) sup_ty_opt ("_ |- _ <=T _" [71,71,71] 70) and sup_state ("_ |- _ <=i _" [71,71,71] 70) and sup_state_opt ("_ |- _ <=' _" [71,71,71] 70) and sup_loc ("_ |- _ [<=T] _" [71,71,71] 70) subsection "Unfolding" lemma JVM_states_unfold: - "states P mxs mxl \ err(opt((Union {list n (types P) |n. n <= mxs}) \ - list mxl (err(types P))))" + "states P mxs mxl \ err(opt((Union {nlists n (types P) |n. n <= mxs}) \ + nlists mxl (err(types P))))" (*<*) apply (unfold states_def sl_def Opt.esl_def Err.sl_def stk_esl_def loc_sl_def Product.esl_def Listn.sl_def upto_esl_def SemiType.esl_def Err.esl_def) apply simp done (*>*) lemma JVM_le_unfold: "le P m n \ Err.le(Opt.le(Product.le(Listn.le(subtype P))(Listn.le(Err.le(subtype P)))))" (*<*) apply (unfold le_def sl_def Opt.esl_def Err.sl_def stk_esl_def loc_sl_def Product.esl_def Listn.sl_def upto_esl_def SemiType.esl_def Err.esl_def) apply simp done (*>*) lemma sl_def2: "JVM_SemiType.sl P mxs mxl \ (states P mxs mxl, JVM_SemiType.le P mxs mxl, JVM_SemiType.sup P mxs mxl)" (*<*) by (unfold JVM_SemiType.sup_def states_def JVM_SemiType.le_def) simp (*>*) lemma JVM_le_conv: "le P m n (OK t1) (OK t2) = P \ t1 \' t2" (*<*) by (simp add: JVM_le_unfold Err.le_def lesub_def sup_state_opt_def sup_state_def sup_ty_opt_def) (*>*) lemma JVM_le_Err_conv: "le P m n = Err.le (sup_state_opt P)" (*<*) by (unfold sup_state_opt_def sup_state_def sup_ty_opt_def JVM_le_unfold) simp (*>*) lemma err_le_unfold [iff]: "Err.le r (OK a) (OK b) = r a b" (*<*) by (simp add: Err.le_def lesub_def) (*>*) subsection \ Semilattice \ lemma order_sup_state_opt' [intro, simp]: "wf_prog wf_mb P \ - order (sup_state_opt P) (opt ((\ {list n (types P) |n. n \ mxs} ) \ list (Suc (length Ts + mxl\<^sub>0)) (err (types P))))" + order (sup_state_opt P) (opt ((\ {nlists n (types P) |n. n \ mxs} ) \ nlists (Suc (length Ts + mxl\<^sub>0)) (err (types P))))" (*<*) apply (unfold sup_state_opt_def sup_state_def sup_ty_opt_def ) apply (blast intro:order_le_prodI) \\ use Listn.thy.order_listI2 \ done (*<*) lemma order_sup_state_opt'' [intro, simp]: "wf_prog wf_mb P \ - order (sup_state_opt P) (opt ((\ {list n (types P) |n. n \ mxs} ) \ list ((length Ts + mxl\<^sub>0)) (err (types P))))" + order (sup_state_opt P) (opt ((\ {nlists n (types P) |n. n \ mxs} ) \ nlists ((length Ts + mxl\<^sub>0)) (err (types P))))" (*<*) apply (unfold sup_state_opt_def sup_state_def sup_ty_opt_def ) apply (blast intro:order_le_prodI) \\ use Listn.thy.order_listI2 \ done (*<*) (* lemma order_sup_state_opt [intro, simp]: "wf_prog wf_mb P \ order (sup_state_opt P)" (*<*) by (unfold sup_state_opt_def sup_state_def sup_ty_opt_def) blast (*>*) *) lemma semilat_JVM [intro?]: "wf_prog wf_mb P \ semilat (JVM_SemiType.sl P mxs mxl)" (*<*) apply (unfold JVM_SemiType.sl_def stk_esl_def loc_sl_def) apply (blast intro: err_semilat_Product_esl err_semilat_upto_esl Listn_sl err_semilat_JType_esl) done (*>*) subsection \ Widening with @{text "\"} \ lemma subtype_refl[iff]: "subtype P t t" (*<*) by (simp add: fun_of_def) (*>*) lemma sup_ty_opt_refl [iff]: "P \ T \\<^sub>\ T" (*<*) apply (unfold sup_ty_opt_def) apply (fold lesub_def) apply (rule le_err_refl) apply (simp add: lesub_def) done (*>*) lemma Err_any_conv [iff]: "P \ Err \\<^sub>\ T = (T = Err)" (*<*) by (unfold sup_ty_opt_def) (rule Err_le_conv [simplified lesub_def]) (*>*) lemma any_Err [iff]: "P \ T \\<^sub>\ Err" (*<*) by (unfold sup_ty_opt_def) (rule le_Err [simplified lesub_def]) (*>*) lemma OK_OK_conv [iff]: "P \ OK T \\<^sub>\ OK T' = P \ T \ T'" (*<*) by (simp add: sup_ty_opt_def fun_of_def) (*>*) lemma any_OK_conv [iff]: "P \ X \\<^sub>\ OK T' = (\T. X = OK T \ P \ T \ T')" (*<*) apply (unfold sup_ty_opt_def) apply (rule le_OK_conv [simplified lesub_def]) done (*>*) lemma OK_any_conv: "P \ OK T \\<^sub>\ X = (X = Err \ (\T'. X = OK T' \ P \ T \ T'))" (*<*) apply (unfold sup_ty_opt_def) apply (rule OK_le_conv [simplified lesub_def]) done (*>*) lemma sup_ty_opt_trans [intro?, trans]: "\P \ a \\<^sub>\ b; P \ b \\<^sub>\ c\ \ P \ a \\<^sub>\ c" (*<*) by (auto intro: widen_trans simp add: sup_ty_opt_def Err.le_def lesub_def fun_of_def split: err.splits) (*>*) subsection "Stack and Registers" lemma stk_convert: "P \ ST [\] ST' = Listn.le (subtype P) ST ST'" (*<*) by (simp add: Listn.le_def lesub_def) (*>*) lemma sup_loc_refl [iff]: "P \ LT [\\<^sub>\] LT" (*<*) by (rule list_all2_refl) simp (*>*) lemmas sup_loc_Cons1 [iff] = list_all2_Cons1 [of "sup_ty_opt P"] for P lemma sup_loc_def: "P \ LT [\\<^sub>\] LT' \ Listn.le (sup_ty_opt P) LT LT'" (*<*) by (simp add: Listn.le_def lesub_def) (*>*) lemma sup_loc_widens_conv [iff]: "P \ map OK Ts [\\<^sub>\] map OK Ts' = P \ Ts [\] Ts'" (*<*) by (simp add: list_all2_map1 list_all2_map2) (*>*) lemma sup_loc_trans [intro?, trans]: "\P \ a [\\<^sub>\] b; P \ b [\\<^sub>\] c\ \ P \ a [\\<^sub>\] c" (*<*) by (rule list_all2_trans, rule sup_ty_opt_trans) (*>*) subsection "State Type" lemma sup_state_conv [iff]: "P \ (ST,LT) \\<^sub>i (ST',LT') = (P \ ST [\] ST' \ P \ LT [\\<^sub>\] LT')" (*<*) by (auto simp add: sup_state_def stk_convert lesub_def Product.le_def sup_loc_def) (*>*) lemma sup_state_conv2: "P \ s1 \\<^sub>i s2 = (P \ fst s1 [\] fst s2 \ P \ snd s1 [\\<^sub>\] snd s2)" (*<*) by (cases s1, cases s2) simp (*>*) lemma sup_state_refl [iff]: "P \ s \\<^sub>i s" (*<*) by (auto simp add: sup_state_conv2) (*>*) lemma sup_state_trans [intro?, trans]: "\P \ a \\<^sub>i b; P \ b \\<^sub>i c\ \ P \ a \\<^sub>i c" (*<*) by (auto intro: sup_loc_trans widens_trans simp add: sup_state_conv2) (*>*) lemma sup_state_opt_None_any [iff]: "P \ None \' s" (*<*) by (simp add: sup_state_opt_def Opt.le_def) (*>*) lemma sup_state_opt_any_None [iff]: "P \ s \' None = (s = None)" (*<*) by (simp add: sup_state_opt_def Opt.le_def) (*>*) lemma sup_state_opt_Some_Some [iff]: "P \ Some a \' Some b = P \ a \\<^sub>i b" (*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*) lemma sup_state_opt_any_Some: "P \ (Some s) \' X = (\s'. X = Some s' \ P \ s \\<^sub>i s')" (*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*) lemma sup_state_opt_refl [iff]: "P \ s \' s" (*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*) lemma sup_state_opt_trans [intro?, trans]: "\P \ a \' b; P \ b \' c\ \ P \ a \' c" (*<*) apply (unfold sup_state_opt_def Opt.le_def lesub_def) apply (simp del: split_paired_All) apply (rule sup_state_trans, assumption+) done (*>*) lemma sup_state_opt_err : "(Err.le (sup_state_opt P)) s s" apply (unfold JVM_le_unfold Product.le_def Opt.le_def Err.le_def lesssub_def lesub_def Listn.le_def) apply (auto split: err.splits) done lemma Cons_less_Conss1 [simp]: "x#xs [\\<^bsub>subtype P\<^esub>] y#ys = (x \\<^bsub>subtype P\<^esub> y \ xs [\\<^bsub>subtype P\<^esub>] ys \ x = y \ xs [\\<^bsub>subtype P\<^esub>] ys)" apply (unfold lesssub_def ) apply auto apply (simp add:lesssub_def lesub_def) \\widen_refl, subtype_refl \ done lemma Cons_less_Conss2 [simp]: "x#xs [\\<^bsub>Err.le (subtype P)\<^esub>] y#ys = (x \\<^bsub>Err.le (subtype P)\<^esub> y \ xs [\\<^bsub>Err.le (subtype P)\<^esub>] ys \ x = y \ xs [\\<^bsub>Err.le (subtype P)\<^esub>] ys)" apply (unfold lesssub_def ) apply auto apply (simp add:lesssub_def lesub_def Err.le_def split: err.splits) done lemma acc_le_listI1 [intro!]: " acc (subtype P) \ acc (Listn.le (subtype P))" (*<*) apply (unfold acc_def) apply (subgoal_tac "wf(UN n. {(ys,xs). size xs = n \ size ys = n \ xs <_(Listn.le (subtype P)) ys})") apply (erule wf_subset) apply (blast intro: lesssub_lengthD) apply (rule wf_UN) prefer 2 apply (rename_tac m n) apply (case_tac "m=n") apply simp apply (fast intro!: equals0I dest: not_sym) apply (rename_tac n) apply (induct_tac n) apply (simp add: lesssub_def cong: conj_cong) apply (rename_tac k) apply (simp add: wf_eq_minimal) apply (simp (no_asm) add: length_Suc_conv cong: conj_cong) apply clarify apply (rename_tac M m) apply (case_tac "\x xs. size xs = k \ x#xs \ M") prefer 2 apply (erule thin_rl) apply (erule thin_rl) apply blast apply (erule_tac x = "{a. \xs. size xs = k \ a#xs:M}" in allE) apply (erule impE) apply blast apply (thin_tac "\x xs. P x xs" for P) apply clarify apply (rename_tac maxA xs) apply (erule_tac x = "{ys. size ys = size xs \ maxA#ys \ M}" in allE) apply (erule impE) apply blast apply clarify apply (thin_tac "m \ M") apply (thin_tac "maxA#xs \ M") apply (rule bexI) prefer 2 apply assumption apply clarify apply simp apply blast done lemma acc_le_listI2 [intro!]: " acc (Err.le (subtype P)) \ acc (Listn.le (Err.le (subtype P)))" (*<*) apply (unfold acc_def) apply (subgoal_tac "wf(UN n. {(ys,xs). size xs = n \ size ys = n \ xs <_(Listn.le (Err.le (subtype P))) ys})") apply (erule wf_subset) apply (blast intro: lesssub_lengthD) apply (rule wf_UN) prefer 2 apply (rename_tac m n) apply (case_tac "m=n") apply simp apply (fast intro!: equals0I dest: not_sym) apply (rename_tac n) apply (induct_tac n) apply (simp add: lesssub_def cong: conj_cong) apply (rename_tac k) apply (simp add: wf_eq_minimal) apply (simp (no_asm) add: length_Suc_conv cong: conj_cong) apply clarify apply (rename_tac M m) apply (case_tac "\x xs. size xs = k \ x#xs \ M") prefer 2 apply (erule thin_rl) apply (erule thin_rl) apply blast apply (erule_tac x = "{a. \xs. size xs = k \ a#xs:M}" in allE) apply (erule impE) apply blast apply (thin_tac "\x xs. P x xs" for P) apply clarify apply (rename_tac maxA xs) apply (erule_tac x = "{ys. size ys = size xs \ maxA#ys \ M}" in allE) apply (erule impE) apply blast apply clarify apply (thin_tac "m \ M") apply (thin_tac "maxA#xs \ M") apply (rule bexI) prefer 2 apply assumption apply clarify apply simp apply blast done lemma acc_JVM [intro]: "wf_prog wf_mb P \ acc (JVM_SemiType.le P mxs mxl)" (*<*) by (unfold JVM_le_unfold) blast (*>*) \\ use acc_listI1, acc_listI2 \ 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,250 +1,250 @@ (* 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 \ (b = Static \ b = NonStatic) \ 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" + shows "\\s \ nlists (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 \ []" and stab: "b = Static \ b = NonStatic" by (auto simp add: wt_lbv_def) moreover from wf lbv C Ts obtain \s where - list: "\s \ list (size is) A" and + list: "\s \ nlists (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" by (fastforce simp: make_cert_def check_types_def JVM_states_unfold dest!: nth_mem) moreover note 0 size ultimately show ?thesis using staticb 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" 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 ass1: "wf_prog wf_md P" and sees: "P \ Ca sees M, b : Ts\T = m in Ca" and ass2: "set Ts \ types P" and ass3: "bd = (M, b, Ts, T, m)" and ass4: "?A P Ca bd" from ass3 ass4 have stab: "b = Static \ b = NonStatic" by (simp add:wt_lbv_def) from ass1 sees ass2 ass3 ass4 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_auxi[OF stab]] 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 \)" (*<*) 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 ass1: "wf_prog wf_md P" and sees: "P \ Ca sees M, b : Ts\T = m in Ca" and ass2: "set Ts \ types P" and ass3: "bd = (M, b, Ts, T, m)" and ass4: "?A P Ca bd" from ass3 ass4 have stab: "b = Static \ b = NonStatic" by (simp add:wt_method_def) from ass1 sees ass2 ass3 ass4 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_auxi[OF stab]) } 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 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,355 +1,355 @@ (* 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 and n 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" defines [simp]: "n \ size is" assumes staticb: "b = Static \ b = NonStatic" 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)" thm start_context.intro lemma start_context_intro_auxi: fixes P b Ts p C assumes "b = Static \ b = NonStatic " and "wf_prog p P" and "is_class P C" and "set Ts \ types P" shows " start_context P b Ts p C" using start_context.intro[OF JVM_sl.intro] start_context_axioms_def assms by auto subsection \ Connecting JVM and Framework \ lemma (in start_context) semi: "semilat (A, r, f)" apply (insert semilat_JVM[OF wf]) apply (unfold A_def r_def f_def JVM_SemiType.le_def JVM_SemiType.sup_def states_def) apply auto done 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 + "(\n. ST \ nlists n A \ n \ mxs) = (set ST \ A \ size ST \ mxs)" + by (unfold nlists_def) auto -lemma singleton_list: - "(\n. [Class C] \ list n (types P) \ n \ mxs) = (is_class P C \ 0 < mxs)" +lemma singleton_nlists: + "(\n. [Class C] \ nlists 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 +lemma in_nlistsE: + "\ xs \ nlists n A; \size xs = n; set xs \ A\ \ P \ \ P" + by (unfold nlists_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" (*<*) 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)))" + let ?A = "opt((Union {nlists n (types P) |n. n <= mxs}) \ + nlists ?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) + dest: nlistsE_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" (*<*) 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: 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)" (*<*) 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)" (*<*) 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 (subgoal_tac "b = Static \ b = NonStatic") apply (fastforce split:if_splits)\\ order_sup_state_opt' order_sup_state_opt'' \ apply (simp only:staticb) 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 ?A" 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) + using Ts C by (cases b; force intro!: nlists_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 \ [] \ (b = Static \ b = NonStatic) \ 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)" (*<*)using staticb 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/TypeComp.thy b/thys/JinjaDCI/Compiler/TypeComp.thy --- a/thys/JinjaDCI/Compiler/TypeComp.thy +++ b/thys/JinjaDCI/Compiler/TypeComp.thy @@ -1,1686 +1,1686 @@ (* 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" (*<*)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)(*>*) + "set E \ types P \ ty\<^sub>l E A \ nlists mxl (err (types P))" +(*<*)by(auto simp add:ty\<^sub>l_def intro!:nlistsI 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)" (*<*) 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: 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" (*<*) 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) + (blast intro!:nlistsI 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" (*<*) by(simp add:ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types) - (blast intro!:listI) + (blast intro!:nlistsI) (*>*) 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}" (*<*) 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" (*<*) 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)" (*<*) 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)" (*<*) 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 \))" (*<*)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: 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" (*<*) 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: 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" (*<*) 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: 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" (*<*) 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: 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" (*<*) 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]: 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" (*<*) 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]: 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" (*<*) 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]: assumes wtis: "\ is,xt [::] \s@[\]" and ss: "P \ \ \' \'" shows "\ is,xt [::] \s@[\']" (*<*) 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]: 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" proof(induct n) case (Suc n) then show ?case by(simp add: drop_Suc) qed simp lemma drop_mess: 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)" thm compT_sizes(1) let ?if = "IfFalse (int ?n\<^sub>c + 3)" from wtc wte have "compE\<^sub>2 c \ []" and "compE\<^sub>2 e \ []" using WT\<^sub>1_nsub_RI by auto then have compe2c: "length (compE\<^sub>2 c) > 0" and compe2e: "length (compE\<^sub>2 e) > 0" by auto have Suc_pred'_auxi: "\n. 0 < n \ n = Suc (n - Suc 0)" using Suc_pred'[OF compe2c] by auto have "\ [?if],[] [::] ?\\<^sub>e # ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c, ?\\<^sub>2, ?\\<^sub>1, ?\']" proof-{ show ?thesis apply (rule wt_IfFalse) apply simp apply simp apply (subgoal_tac "length (compE\<^sub>2 c) = length (compT E (A \ \ e) ST c) + 1" "nat (int (length (compE\<^sub>2 c)) + 3 - 2) > length (compT E (A \ \ e) ST c)") apply (insert Suc_pred'_auxi[OF compe2c]) apply (simp add:compT_sizes(1)[OF wtc] )+ done } qed 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 proof-{ let ?t1 = "ty\<^sub>i' ST E A" let ?t2 = "ty\<^sub>i' (Boolean # ST) E (A \ \ e)" let ?t3 = "ty\<^sub>i' ST E (A \ \ e)" let ?t4 = "[ty\<^sub>i' (Tc # ST) E (A \ \ e \ \ c), ty\<^sub>i' ST E (A \ \ e \ \ c), ty\<^sub>i' ST E (A \ \ e), ty\<^sub>i' (Void # ST) E (A \ \ e)]" let ?c = " compT E (A \ \ e) ST c" let ?e = "compT E A ST e" assume ass: "P \ ty\<^sub>i' ST E (A \ \ e \ \ c) \' ?t1" show ?thesis apply (rule wt_Goto) apply simp apply simp apply simp proof-{ let ?s1 = "((?t1 # ?e @ [?t2]) @ ?t3 # ?c)" have len1: "length ?c = length (compE\<^sub>2 c) - 1" using compT_sizes(1)[OF wtc] by simp have len2: "length ?e = length (compE\<^sub>2 e) - 1" using compT_sizes(1)[OF wte] by simp hence "length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2 > length ?s1" using len1 compe2c compe2e by auto hence "(?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) = ?t4 ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2 - length ?s1)" by (auto simp only:nth_append split: if_splits) hence "(?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) = ty\<^sub>i' ST E (A \ \ e \ \ c)" using len1 len2 compe2c compe2e by auto hence "P \ (?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) \' ty\<^sub>i' ST E A" using ass by simp thus "P \ ((?t1 # ?e @ [?t2]) @ ?t3 # ?c @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) \' ((?t1 # ?e @ [?t2]) @ ?t3 # ?c @ ?t4) ! nat (int (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) + - int (length (compE\<^sub>2 c) + length (compE\<^sub>2 e) + 2))" by simp }qed }qed 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 from wte wt\<^sub>1 wt\<^sub>2 have "compE\<^sub>2 e\<^sub>1 \ []" and "compE\<^sub>2 e\<^sub>2 \ []" and "compE\<^sub>2 e \ []" using WT\<^sub>1_nsub_RI by auto then have compe: "length (compE\<^sub>2 e) > 0" and compe1: "length (compE\<^sub>2 e\<^sub>1) > 0" and compe2: "length (compE\<^sub>2 e\<^sub>2) > 0" 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 len1: "length ?\s\<^sub>1 = length (compE\<^sub>2 e\<^sub>1) - 1" using compT_sizes(1)[OF wt\<^sub>1] by simp have len2: "length ?\s\<^sub>2 = length (compE\<^sub>2 e\<^sub>2) - 1" using compT_sizes(1)[OF wt\<^sub>2] by simp have len_auxi: "length (compE\<^sub>2 e\<^sub>1) - (length (compE\<^sub>2 e\<^sub>1) - Suc 0) = Suc 0" using compe1 by (simp add:diff_Suc split:nat.split) have "\ [?g\<^sub>1],[] [::] [?\\<^sub>1] @ ?\s\<^sub>1\<^sub>2" proof-{ show ?thesis apply clarsimp apply (rule wt_IfFalse) apply (simp only:nat_add_distrib) apply simp apply (auto simp only:nth_append split:if_splits) apply (simp add:len1) apply (simp only: len1 compe1) apply (simp add:len1 len2 compe1 compe2 ) apply (insert len_auxi) apply simp done }qed 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, \)" (*<*)(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" (*<*) 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 \" (*<*) by (simp add: app_def xcpt_app_def eff_def xcpt_eff_def norm_eff_def) (fastforce simp add: image_def) (*>*) lemma [simp]: assumes "app i P mpc T pc mxl xt \" shows "eff i (compP f P) pc xt \ = eff i P pc xt \" (*<*) 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" (*<*)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 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)" (*<*)(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 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)" (*<*)(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: assumes wf: "wf_J\<^sub>1_prog P" shows "wf_jvm_prog (compP\<^sub>2 P)" (*<*) 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)" (*<*) by(simp only:o_def J2JVM_def) (blast intro:wt_compP\<^sub>2 compP\<^sub>1_pres_wf) end