Theory BVExec
section ‹Kildall for the JVM \label{sec:JVM}›
theory BVExec
imports
"../DFA/Abstract_BV"
TF_JVM
begin
definition kiljvm :: "'addr jvm_prog ⇒ nat ⇒ nat ⇒ ty ⇒
'addr instr list ⇒ ex_table ⇒ ty⇩i' err list ⇒ ty⇩i' err list"
where
"kiljvm P mxs mxl T⇩r is xt ≡
kildall (JVM_SemiType.le P mxs mxl) (JVM_SemiType.sup P mxs mxl)
(exec P mxs T⇩r xt is)"
definition wt_kildall :: "'addr jvm_prog ⇒ cname ⇒ ty list ⇒ ty ⇒ nat ⇒ nat ⇒
'addr instr list ⇒ ex_table ⇒ bool"
where
"wt_kildall P C' Ts T⇩r mxs mxl⇩0 is xt ≡
0 < size is ∧
(let first = Some ([],[OK (Class C')]@(map OK Ts)@(replicate mxl⇩0 Err));
start = OK first#(replicate (size is - 1) (OK None));
result = kiljvm P mxs (1+size Ts+mxl⇩0) T⇩r is xt start
in ∀n < size is. result!n ≠ Err)"
definition wf_jvm_prog⇩k :: "'addr jvm_prog ⇒ bool"
where
"wf_jvm_prog⇩k P ≡
wf_prog (λP C' (M,Ts,T⇩r,(mxs,mxl⇩0,is,xt)). wt_kildall P C' Ts T⇩r mxs mxl⇩0 is xt) P"
theorem (in start_context) is_bcv_kiljvm:
"is_bcv r Err step (size is) A (kiljvm P mxs mxl T⇩r is xt)"
apply (insert wf)
apply (unfold kiljvm_def)
apply (fold r_def f_def step_def_exec)
apply (rule is_bcv_kildall)
apply simp
apply (rule Semilat.intro)
apply (fold sl_def2, erule semilat_JVM)
apply simp
apply blast
apply (simp add: JVM_le_unfold)
apply (rule exec_pres_type)
apply (rule bounded_step)
apply (erule step_mono)
done
lemma subset_replicate [intro?]: "set (replicate n x) ⊆ {x}"
by (induct n) auto
lemma in_set_replicate:
assumes "x ∈ set (replicate n y)"
shows "x = y"
proof -
note assms
also have "set (replicate n y) ⊆ {y}" ..
finally show ?thesis by simp
qed
lemma (in start_context) start_in_A [intro?]:
"0 < size is ⟹ start ∈ list (size is) A"
using Ts C
apply (simp add: JVM_states_unfold)
apply (force intro!: listI list_appendI dest!: in_set_replicate)
done
theorem (in start_context) wt_kil_correct:
assumes wtk: "wt_kildall P C Ts T⇩r mxs mxl⇩0 is xt"
shows "∃τs. wt_method P C Ts T⇩r mxs mxl⇩0 is xt τs"
proof -
from wtk obtain res where
result: "res = kiljvm P mxs mxl T⇩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⇩r is xt)"
by (rule is_bcv_kiljvm)
from instrs have "start ∈ list (size is) A" ..
with bcv success result have
"∃ts∈list (size is) A. start [⊑⇩r] ts ∧ wt_step r Err step ts"
by (unfold is_bcv_def) blast
then obtain τs' where
in_A: "τs' ∈ list (size is) A" and
s: "start [⊑⇩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 ⊑⇩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⇩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⇩r mxs mxl⇩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⇩r mxs mxl⇩0 is xt τs"
shows "wt_kildall P C Ts T⇩r mxs mxl⇩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⇩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⇩r is xt)"
by (rule is_bcv_kiljvm)
moreover from instrs have "start ∈ list (size is) A" ..
moreover
let ?τs = "map OK τs"
have less_τs: "start [⊑⇩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 ⊑⇩r ?τs!0"
by (simp add: JVM_le_Err_conv lesub_def)
moreover {
fix n'
have "OK None ⊑⇩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 ⊑⇩r ?τs!n" by simp
}
ultimately
show "n < size start ⟹ start!n ⊑⇩r ?τs!n" by (cases n, blast+)
qed
moreover
from ck_type length
have "?τs ∈ list (size is) A"
by (auto intro!: listI simp add: check_types_def)
moreover
from wt_err have "wt_step r Err step ?τs"
by (simp add: wt_err_step_def JVM_le_Err_conv)
ultimately
have "∀p. p < size is ⟶ kiljvm P mxs mxl T⇩r is xt start ! p ≠ Err"
by (unfold is_bcv_def) blast
with instrs
show "wt_kildall P C Ts T⇩r mxs mxl⇩0 is xt" by (unfold wt_kildall_def) simp
qed
theorem jvm_kildall_correct:
"wf_jvm_prog⇩k P = wf_jvm_prog P"
proof
let ?Φ = "λC M. let (C,Ts,T⇩r,meth) = method P C M; (mxs,mxl⇩0,is,xt) = the meth in
SOME τs. wt_method P C Ts T⇩r mxs mxl⇩0 is xt τs"
assume wt: "wf_jvm_prog⇩k P"
hence "wf_jvm_prog⇘?Φ⇙ P"
apply (unfold wf_jvm_prog_phi_def wf_jvm_prog⇩k_def)
apply (erule wf_prog_lift)
apply (auto intro: someI_ex[OF start_context.wt_kil_correct [OF start_context.intro]])
done
thus "wf_jvm_prog P" by (unfold wf_jvm_prog_def) fast
next
assume wt: "wf_jvm_prog P"
thus "wf_jvm_prog⇩k P"
apply (unfold wf_jvm_prog_def wf_jvm_prog_phi_def wf_jvm_prog⇩k_def)
apply (clarify)
apply (erule wf_prog_lift)
apply (auto intro!: start_context.wt_kil_complete start_context.intro)
done
qed
end