Up to index of Isabelle/HOL/objinit
theory BVNoTypeError = JVMDefensive + BVSpecTypeSafe:(* Title: HOL/MicroJava/BV/BVNoTypeErrors.thy ID: $Id: BVNoTypeError.html,v 1.1 2002/11/28 14:12:09 kleing Exp $ Author: Gerwin Klein Copyright GPL *) header {* \isaheader{Welltyped Programs produce no Type Errors} *} theory BVNoTypeError = JVMDefensive + BVSpecTypeSafe: text {* Some simple lemmas about the type testing functions of the defensive JVM: *} lemma typeof_NoneD [simp,dest]: "typeof (\<lambda>v. None) v = Some x \<Longrightarrow> ¬isAddr v" by (cases v) auto lemma isRef_def2: "isRef v = (v = Null \<or> (\<exists>loc. v = Addr loc))" by (cases v) (auto simp add: isRef_def) lemma isRef [simp]: "G,hp,ihp \<turnstile> v ::\<preceq>i Init (RefT T) \<Longrightarrow> isRef v" by (cases v) (auto simp add: iconf_def conf_def isRef_def) lemma isIntg [simp]: "G,hp,ihp \<turnstile> v ::\<preceq>i Init (PrimT Integer) \<Longrightarrow> isIntg v" by (cases v) (auto simp add: iconf_def conf_def) declare approx_loc_len [simp] approx_stk_len [simp] (* fixme: move to List.thy *) lemma list_all2I: "\<forall>(x,y) \<in> set (zip a b). P x y \<Longrightarrow> length a = length b \<Longrightarrow> list_all2 P a b" by (simp add: list_all2_def) text {* The main theorem: welltyped programs do not produce type errors if they are started in a conformant state. *} theorem assumes welltyped: "wt_jvm_prog G Phi" and conforms: "G,Phi \<turnstile>JVM s \<surd>" shows no_type_error: "exec_d G (Normal s) \<noteq> TypeError" proof - from welltyped obtain mb where wf: "wf_prog mb G" by (fast dest: wt_jvm_progD) obtain xcp hp ihp frs where s [simp]: "s = (xcp, hp, ihp, frs)" by (cases s) from conforms have "xcp \<noteq> None \<or> frs = [] \<Longrightarrow> check G s" by (unfold correct_state_def check_def) auto moreover { assume "¬(xcp \<noteq> None \<or> frs = [])" then obtain stk loc C sig pc r frs' where xcp [simp]: "xcp = None" and frs [simp]: "frs = (stk,loc,C,sig,pc,r)#frs'" by (clarsimp simp add: neq_Nil_conv) fast from conforms obtain ST LT z rT maxs maxl ins et where hconf: "G \<turnstile>h hp \<surd>" and class: "is_class G C" and meth: "method (G, C) sig = Some (C, rT, maxs, maxl, ins, et)" and phi: "Phi C sig ! pc = Some ((ST,LT), z)" and frame: "correct_frame G hp ihp (ST,LT) maxl ins (stk,loc,C,sig,pc,r)" and frames: "correct_frames G hp ihp Phi rT sig z r frs'" by simp (rule correct_stateE) from frame obtain stk: "approx_stk G hp ihp stk ST" and loc: "approx_loc G hp ihp loc LT" and init: "fst sig = init \<longrightarrow> corresponds stk loc (ST, LT) ihp (fst r) (PartInit C) \<and> (\<exists>l. fst r = Addr l \<and> hp l \<noteq> None \<and> (ihp l = PartInit C \<or> (\<exists>C'. ihp l = Init (Class C'))))" and pc: "pc < length ins" and len: "length loc = length (snd sig) + maxl + 1" by (rule correct_frameE) note approx_val_def [simp] from welltyped meth conforms have "wt_instr (ins!pc) G C rT (Phi C sig) maxs (fst sig=init) (length ins) et pc" by simp (rule wt_jvm_prog_impl_wt_instr_cor) then obtain app: "app (ins!pc) G C pc maxs rT (fst sig=init) et (Some ((ST,LT),z))" and pc': "\<forall>pc' \<in> set (succs (ins!pc) pc). pc' < length ins" by (simp add: wt_instr_def phi eff_def) blast with stk loc have "check_instr (ins!pc) G hp ihp stk loc C sig pc r (length ins) frs'" proof (cases "ins!pc") case (Getfield F C) with app stk loc obtain v vT stk' where class: "is_class G C" and field: "field (G, C) F = Some (C, vT)" and stk: "stk = v # stk'" and conf: "G,hp,ihp \<turnstile> v ::\<preceq>i Init (Class C)" by clarsimp (blast dest: iconf_widen [OF _ _ wf]) from conf have isRef: "isRef v" by simp moreover { assume "v \<noteq> Null" with conf isRef have "\<exists>D vs. hp (the_Addr v) = Some (D, vs) \<and> is_init hp ihp v \<and> G \<turnstile> D \<preceq>C C" by (fastsimp simp add: iconf_def conf_def isRef_def2) } ultimately show ?thesis using Getfield field class stk hconf apply clarsimp apply (fastsimp dest!: hconfD widen_cfs_fields [OF _ _ wf] oconf_objD) done next case (Putfield F C) with app stk loc obtain v ref vT stk' where class: "is_class G C" and field: "field (G, C) F = Some (C, vT)" and stk: "stk = v # ref # stk'" and confv: "G,hp,ihp \<turnstile> v ::\<preceq>i Init vT" and confr: "G,hp,ihp \<turnstile> ref ::\<preceq>i Init (Class C)" by clarsimp (blast dest: iconf_widen [OF _ _ wf]) from confr have isRef: "isRef ref" by simp moreover from confv have "is_init hp ihp v" by (simp add: iconf_def) moreover { assume "ref \<noteq> Null" with confr isRef have "\<exists>D vs. hp (the_Addr ref) = Some (D, vs) \<and> is_init hp ihp ref \<and> G \<turnstile> D \<preceq>C C" by (fastsimp simp add: iconf_def conf_def isRef_def2) } ultimately show ?thesis using Putfield field class stk confv by (clarsimp simp add: iconf_def) next case (Invoke C mn ps) with stk app show ?thesis apply clarsimp apply (clarsimp dest!: approx_stk_append_lemma simp add: nth_append) apply (drule iconf_widen [OF _ _ wf], assumption) apply (clarsimp simp add: iconf_def) apply (drule non_npD, assumption) apply clarsimp apply (drule widen_methd [OF _ wf], assumption) apply (clarsimp simp add: approx_stk_rev [symmetric]) apply (drule list_all2I, assumption) apply (unfold approx_stk_def approx_loc_def) apply (simp add: list_all2_approx) apply (drule list_all2_iconf_widen [OF wf], assumption+) done next case (Invoke_special C mn ps) with stk app show ?thesis apply clarsimp apply (clarsimp dest!: approx_stk_append_lemma simp add: nth_append) apply (erule disjE) apply (clarsimp simp add: iconf_def isRef_def) apply (clarsimp simp add: approx_stk_rev [symmetric]) apply (drule list_all2I, assumption) apply (simp add: list_all2_approx approx_stk_def approx_loc_def) apply (drule list_all2_iconf_widen [OF wf], assumption+) apply (clarsimp simp add: iconf_def isRef_def) apply (clarsimp simp add: approx_stk_rev [symmetric]) apply (drule list_all2I, assumption) apply (unfold approx_stk_def approx_loc_def) apply (simp add: list_all2_approx) apply (drule list_all2_iconf_widen [OF wf], assumption+) done next case Return with stk app init meth frames show ?thesis apply clarsimp apply (drule iconf_widen [OF _ _ wf], assumption) apply (clarsimp simp add: iconf_def neq_Nil_conv constructor_ok_def is_init_def isRef_def2) done qed auto hence "check G s" by (simp add: check_def meth) } ultimately have "check G s" by blast thus "exec_d G (Normal s) \<noteq> TypeError" .. qed text {* The theorem above tells us that, in welltyped programs, the defensive machine reaches the same result as the aggressive one (after arbitrarily many steps). *} theorem welltyped_aggressive_imp_defensive: "wt_jvm_prog G Phi \<Longrightarrow> G,Phi \<turnstile>JVM s \<surd> \<Longrightarrow> G \<turnstile> s -jvm\<rightarrow> t \<Longrightarrow> G \<turnstile> (Normal s) -jvmd\<rightarrow> (Normal t)" apply (unfold exec_all_def) apply (erule rtrancl_induct) apply (simp add: exec_all_d_def) apply simp apply (fold exec_all_def) apply (frule BV_correct, assumption+) apply (drule no_type_error, assumption, drule no_type_error_commutes, simp) apply (simp add: exec_all_d_def) apply (rule rtrancl_trans, assumption) apply blast done text {* As corollary we get that the aggresive and the defensive machine are equivalent for welltyped programs (if started in a conformant state, or in the canonical start state) *} corollary welltyped_commutes: fixes G ("\<Gamma>") and Phi ("\<Phi>") assumes "wt_jvm_prog \<Gamma> \<Phi>" and "\<Gamma>,\<Phi> \<turnstile>JVM s \<surd>" shows "\<Gamma> \<turnstile> (Normal s) -jvmd\<rightarrow> (Normal t) = \<Gamma> \<turnstile> s -jvm\<rightarrow> t" by rule (erule defensive_imp_aggressive,rule welltyped_aggressive_imp_defensive) corollary welltyped_initial_commutes: fixes G ("\<Gamma>") and Phi ("\<Phi>") assumes "wt_jvm_prog \<Gamma> \<Phi>" assumes "is_class \<Gamma> C" "method (\<Gamma>,C) (m,[]) = Some (C, b)" "m \<noteq> init" defines start: "s \<equiv> start_state \<Gamma> C m" shows "\<Gamma> \<turnstile> (Normal s) -jvmd\<rightarrow> (Normal t) = \<Gamma> \<turnstile> s -jvm\<rightarrow> t" proof - have "\<Gamma>,\<Phi> \<turnstile>JVM s \<surd>" by (unfold start, rule BV_correct_initial) thus ?thesis by - (rule welltyped_commutes) qed end
lemma typeof_NoneD:
typeof (%v. None) v = Some x ==> ¬ isAddr v
lemma isRef_def2:
isRef v = (v = Null | (EX loc. v = Addr loc))
lemma isRef:
G,hp,ihp \<turnstile> v ::\<preceq>i Init (RefT T) ==> isRef v
lemma isIntg:
G,hp,ihp \<turnstile> v ::\<preceq>i Init (PrimT Integer) ==> isIntg v
lemma list_all2I:
[| Ball (set (zip a b)) (split P); length a = length b |] ==> list_all2 P a b
theorem no_type_error:
[| wt_jvm_prog G Phi; G,Phi |-JVM s [ok] |] ==> exec_d G (Normal s) ~= TypeError
theorem welltyped_aggressive_imp_defensive:
[| wt_jvm_prog G Phi; G,Phi |-JVM s [ok]; G |- s -jvm-> t |] ==> G |- Normal s -jvmd-> Normal t
corollary
[| wt_jvm_prog G Phi; G,Phi |-JVM s [ok] |] ==> G |- Normal s -jvmd-> Normal t = G |- s -jvm-> t
corollary
[| wt_jvm_prog G Phi; is_class G C; method (G, C) (m, []) = Some (C, b); m ~= init |] ==> G |- Normal (start_state G C m) -jvmd-> Normal t = G |- start_state G C m -jvm-> t