Theory TypeComp
section ‹ Preservation of Well-Typedness ›
theory TypeComp
imports Compiler "../BV/BVSpec"
begin
declare nth_append[simp]
lemma max_stack1: "P,E ⊢⇩1 e :: T ⟹ 1 ≤ max_stack e"
using max_stack1'[OF WT⇩1_nsub_RI] by simp
locale TC0 =
fixes P :: "J⇩1_prog" and mxl :: nat
begin
definition "ty E e = (THE T. P,E ⊢⇩1 e :: T)"
definition "ty⇩l E A' = map (λi. if i ∈ A' ∧ i < size E then OK(E!i) else Err) [0..<mxl]"
definition "ty⇩i' ST E A = (case A of None ⇒ None | ⌊A'⌋ ⇒ Some(ST, ty⇩l E A'))"
definition "after E A ST e = ty⇩i' (ty E e # ST) E (A ⊔ 𝒜 e)"
end
lemma (in TC0) ty_def2 [simp]: "P,E ⊢⇩1 e :: T ⟹ ty E e = T"
by(unfold ty_def) (blast intro: the_equality WT⇩1_unique)
lemma (in TC0) [simp]: "ty⇩i' ST E None = None"
by (simp add: ty⇩i'_def)
lemma (in TC0) ty⇩l_app_diff[simp]:
"ty⇩l (E@[T]) (A - {size E}) = ty⇩l E A"
by(auto simp add:ty⇩l_def hyperset_defs)
lemma (in TC0) ty⇩i'_app_diff[simp]:
"ty⇩i' ST (E @ [T]) (A ⊖ size E) = ty⇩i' ST E A"
by(auto simp add:ty⇩i'_def hyperset_defs)
lemma (in TC0) ty⇩l_antimono:
"A ⊆ A' ⟹ P ⊢ ty⇩l E A' [≤⇩⊤] ty⇩l E A"
by(auto simp:ty⇩l_def list_all2_conv_all_nth)
lemma (in TC0) ty⇩i'_antimono:
"A ⊆ A' ⟹ P ⊢ ty⇩i' ST E ⌊A'⌋ ≤' ty⇩i' ST E ⌊A⌋"
by(auto simp:ty⇩i'_def ty⇩l_def list_all2_conv_all_nth)
lemma (in TC0) ty⇩l_env_antimono:
"P ⊢ ty⇩l (E@[T]) A [≤⇩⊤] ty⇩l E A"
by(auto simp:ty⇩l_def list_all2_conv_all_nth)
lemma (in TC0) ty⇩i'_env_antimono:
"P ⊢ ty⇩i' ST (E@[T]) A ≤' ty⇩i' ST E A"
by(auto simp:ty⇩i'_def ty⇩l_def list_all2_conv_all_nth)
lemma (in TC0) ty⇩i'_incr:
"P ⊢ ty⇩i' ST (E @ [T]) ⌊insert (size E) A⌋ ≤' ty⇩i' ST E ⌊A⌋"
by(auto simp:ty⇩i'_def ty⇩l_def list_all2_conv_all_nth)
lemma (in TC0) ty⇩l_incr:
"P ⊢ ty⇩l (E @ [T]) (insert (size E) A) [≤⇩⊤] ty⇩l E A"
by(auto simp: hyperset_defs ty⇩l_def list_all2_conv_all_nth)
lemma (in TC0) ty⇩l_in_types:
"set E ⊆ types P ⟹ ty⇩l E A ∈ nlists mxl (err (types P))"
by(auto simp add:ty⇩l_def intro!:nlistsI dest!: nth_mem)
locale TC1 = TC0
begin
primrec compT :: "ty list ⇒ nat hyperset ⇒ ty list ⇒ expr⇩1 ⇒ ty⇩i' list" and
compTs :: "ty list ⇒ nat hyperset ⇒ ty list ⇒ expr⇩1 list ⇒ ty⇩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⇩1 «bop» e⇩2) =
(let ST⇩1 = ty E e⇩1#ST; A⇩1 = A ⊔ 𝒜 e⇩1 in
compT E A ST e⇩1 @ [after E A ST e⇩1] @
compT E A⇩1 ST⇩1 e⇩2 @ [after E A⇩1 ST⇩1 e⇩2])"
| "compT E A ST (Var i) = []"
| "compT E A ST (i := e) = compT E A ST e @
[after E A ST e, ty⇩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∙⇩sF{D}) = []"
| "compT E A ST (e⇩1∙F{D} := e⇩2) =
(let ST⇩1 = ty E e⇩1#ST; A⇩1 = A ⊔ 𝒜 e⇩1; A⇩2 = A⇩1 ⊔ 𝒜 e⇩2 in
compT E A ST e⇩1 @ [after E A ST e⇩1] @
compT E A⇩1 ST⇩1 e⇩2 @ [after E A⇩1 ST⇩1 e⇩2] @
[ty⇩i' ST E A⇩2])"
| "compT E A ST (C∙⇩sF{D} := e⇩2) = compT E A ST e⇩2 @ [after E A ST e⇩2] @ [ty⇩i' ST E (A ⊔ 𝒜 e⇩2)]"
| "compT E A ST {i:T; e} = compT (E@[T]) (A⊖i) ST e"
| "compT E A ST (e⇩1;;e⇩2) =
(let A⇩1 = A ⊔ 𝒜 e⇩1 in
compT E A ST e⇩1 @ [after E A ST e⇩1, ty⇩i' ST E A⇩1] @
compT E A⇩1 ST e⇩2)"
| "compT E A ST (if (e) e⇩1 else e⇩2) =
(let A⇩0 = A ⊔ 𝒜 e; τ = ty⇩i' ST E A⇩0 in
compT E A ST e @ [after E A ST e, τ] @
compT E A⇩0 ST e⇩1 @ [after E A⇩0 ST e⇩1, τ] @
compT E A⇩0 ST e⇩2)"
| "compT E A ST (while (e) c) =
(let A⇩0 = A ⊔ 𝒜 e; A⇩1 = A⇩0 ⊔ 𝒜 c; τ = ty⇩i' ST E A⇩0 in
compT E A ST e @ [after E A ST e, τ] @
compT E A⇩0 ST c @ [after E A⇩0 ST c, ty⇩i' ST E A⇩1, ty⇩i' ST E A⇩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∙⇩sM(es)) = compTs E A ST es"
| "compT E A ST (try e⇩1 catch(C i) e⇩2) =
compT E A ST e⇩1 @ [after E A ST e⇩1] @
[ty⇩i' (Class C#ST) E A, ty⇩i' ST (E@[Class C]) (A ⊔ ⌊{i}⌋)] @
compT (E@[Class C]) (A ⊔ ⌊{i}⌋) ST e⇩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⇩a :: "ty list ⇒ nat hyperset ⇒ ty list ⇒ expr⇩1 ⇒ ty⇩i' list" where
"compT⇩a E A ST e = compT E A ST e @ [after E A ST e]"
end
lemma compE⇩2_not_Nil[simp]: "P,E ⊢⇩1 e :: T ⟹ compE⇩2 e ≠ []"
by(simp add: compE⇩2_not_Nil' WT⇩1_nsub_RI)
lemma (in TC1) compT_sizes':
shows "⋀E A ST. ¬sub_RI e ⟹ size(compT E A ST e) = size(compE⇩2 e) - 1"
and "⋀E A ST. ¬sub_RIs es ⟹ size(compTs E A ST es) = size(compEs⇩2 es)"
by(induct e and es rule: compE⇩2.induct compEs⇩2.induct)
(auto split:bop.splits nat_diff_split simp: compE⇩2_not_Nil')
lemma (in TC1) compT_sizes[simp]:
shows "⋀E A ST. P,E ⊢⇩1 e :: T ⟹ size(compT E A ST e) = size(compE⇩2 e) - 1"
and "⋀E A ST. P,E ⊢⇩1 es [::] Ts ⟹ size(compTs E A ST es) = size(compEs⇩2 es)"
using compT_sizes' WT⇩1_nsub_RI WTs⇩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⇩i'_conv:
"(⌊(ST, LT)⌋ = ty⇩i' ST⇩0 E A) =
(case A of None ⇒ False | Some A ⇒ (ST = ST⇩0 ∧ LT = ty⇩l E A))"
by(simp add:ty⇩i'_def)
lemma (in TC0) pair_conv_ty⇩i':
"⌊(ST, ty⇩l E A)⌋ = ty⇩i' ST E ⌊A⌋"
by(simp add:ty⇩i'_def)
declare (in TC0)
ty⇩i'_antimono [intro!] after_def[simp]
pair_conv_ty⇩i'[simp] pair_eq_ty⇩i'_conv[simp]
lemma (in TC1) compT_LT_prefix:
"⋀E A ST⇩0. ⟦ ⌊(ST,LT)⌋ ∈ set(compT E A ST⇩0 e); ℬ e (size E) ⟧
⟹ P ⊢ ⌊(ST,LT)⌋ ≤' ty⇩i' ST E A"
and
"⋀E A ST⇩0. ⟦ ⌊(ST,LT)⌋ ∈ set(compTs E A ST⇩0 es); ℬs es (size E) ⟧
⟹ P ⊢ ⌊(ST,LT)⌋ ≤' ty⇩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⇩i'_def simp del:pair_conv_ty⇩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⇩i'_incr
elim!:sup_state_opt_trans)
qed (auto simp:hyperset_defs)
declare (in TC0)
ty⇩i'_antimono [rule del] after_def[simp del]
pair_conv_ty⇩i'[simp del] pair_eq_ty⇩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 ⊢⇩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⇩i'_def JVM_states_unfold ty⇩l_in_types)
(blast intro!:nlistsI WT⇩1_is_type)
qed
lemma (in TC0) OK_ty⇩i'_in_statesI[simp]:
"⟦ set E ⊆ types P; set ST ⊆ types P; size ST ≤ mxs ⟧
⟹ OK (ty⇩i' ST E A) ∈ states P mxs mxl"
by(simp add:ty⇩i'_def JVM_states_unfold ty⇩l_in_types)
(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 ⊢⇩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 ⊢⇩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⇩1_is_type[OF wf] after_in_states[OF wf])
next
case SFAss thus ?case
by(auto simp:image_Un WT⇩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⇩1_is_type[OF wf] after_in_states[OF wf])
next
case Cond thus ?case
by(force simp:image_Un WT⇩1_is_type[OF wf] after_in_states[OF wf])
next
case While thus ?case
by(auto simp:image_Un WT⇩1_is_type[OF wf] after_in_states[OF wf])
next
case Block thus ?case by(auto)
next
case (TryCatch e⇩1 C i e⇩2)
moreover have "size ST + 1 ≤ mxs"
using TryCatch.prems max_stack1[where e=e⇩1] by fastforce
ultimately show ?case
by(auto simp:image_Un WT⇩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⇩1_is_type[OF wf] after_in_states[OF wf])
next
case throw thus ?case
by(auto simp: WT⇩1_is_type[OF wf] after_in_states[OF wf])
next
case Call thus ?case
by(auto simp:image_Un WT⇩1_is_type[OF wf] after_in_states[OF wf])
next
case SCall thus ?case
by(auto simp:image_Un WT⇩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⇩1 @ xt⇩2) = shift n xt⇩1 @ shift n xt⇩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..<t}))"
by(rule_tac x = "x-n" in exI) (force split:nat_diff_split)
}
then show ?thesis by(auto simp:shift_def pcs_def) fast
qed
lemma shift_compxE⇩2:
shows "⋀pc pc' d. shift pc (compxE⇩2 e pc' d) = compxE⇩2 e (pc' + pc) d"
and "⋀pc pc' d. shift pc (compxEs⇩2 es pc' d) = compxEs⇩2 es (pc' + pc) d"
by(induct e and es rule: compxE⇩2.induct compxEs⇩2.induct)
(auto simp:shift_def ac_simps)
lemma compxE⇩2_size_convs[simp]:
shows "n ≠ 0 ⟹ compxE⇩2 e n d = shift n (compxE⇩2 e 0 d)"
and "n ≠ 0 ⟹ compxEs⇩2 es n d = shift n (compxEs⇩2 es 0 d)"
by(simp_all add:shift_compxE⇩2)
locale TC2 = TC1 +
fixes T⇩r :: ty and mxs :: pc
begin
definition
wt_instrs :: "instr list ⇒ ex_table ⇒ ty⇩i' list ⇒ bool"
(‹(⊢ _, _ /[::]/ _)› [0,0,51] 50) where
"⊢ is,xt [::] τs ⟷ size is < size τs ∧ pcs xt ⊆ {0..<size is} ∧
(∀pc< size is. P,T⇩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⇩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⇩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⇩1: "⊢ is⇩1,xt⇩1 [::] τs⇩1@τs⇩2" and wt⇩2: "⊢ is⇩2,xt⇩2 [::] τs⇩2"
and τs_size: "size τs⇩1 = size is⇩1"
shows "⊢ is⇩1@is⇩2, xt⇩1 @ shift (size is⇩1) xt⇩2 [::] τs⇩1@τs⇩2"
proof -
let ?is = "is⇩1@is⇩2" and ?xt = "xt⇩1 @ shift (size is⇩1) xt⇩2"
and ?τs = "τs⇩1@τs⇩2"
have "size ?is < size ?τs" using wt⇩2 τs_size by(fastforce simp:wt_instrs_def)
moreover have "pcs ?xt ⊆ {0..<size ?is}" using wt⇩1 wt⇩2
by(fastforce simp:wt_instrs_def)
moreover {
fix pc assume pc: "pc<size ?is"
have "P,T⇩r,mxs,size ?τs,?xt ⊢ ?is!pc,pc :: ?τs"
proof(cases "pc < length is⇩1")
case True then show ?thesis using wt⇩1 pc
by(fastforce simp: wt_instrs_def wt_instr_appRx)
next
case False
then have "pc - length is⇩1 < length is⇩2" using pc by fastforce
then have "P,T⇩r,mxs,length τs⇩2,xt⇩2 ⊢ is⇩2 ! (pc - length is⇩1),pc - length is⇩1 :: τs⇩2"
using wt⇩2 by(clarsimp simp: wt_instrs_def)
moreover have "pc - length is⇩1 < length τs⇩2" using pc wt⇩2
by(clarsimp simp: wt_instrs_def) arith
moreover have "length τs⇩2 ≤ length τs⇩2" by simp
moreover have "pc - length is⇩1 + length τs⇩1 ∉ pcs xt⇩1" using wt⇩1 τs_size
by(fastforce simp: wt_instrs_def)
ultimately have "P,T⇩r,mxs,length τs⇩2 + length τs⇩1,xt⇩1 @ shift (length τs⇩1) xt⇩2
⊢ is⇩2 ! (pc - length is⇩1),pc - length is⇩1 + length τs⇩1 :: τs⇩1 @ τs⇩2"
by(rule wt_instr_appLx[OF wt_instr_appL[where τs' = "τs⇩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⇩2,xt⇩2 [::] τs⇩2; ⊢ is⇩1,xt⇩1 [::] τs⇩1@τs⇩2; size τs⇩1 = size is⇩1 ⟧
⟹ ⊢ is⇩1@is⇩2, xt⇩1 @ shift (size is⇩1) xt⇩2 [::] τs⇩1@τs⇩2"
by(rule wt_instrs_ext)
corollary (in TC2) wt_instrs_ext_prefix [trans]:
"⟦ ⊢ is⇩1,xt⇩1 [::] τs⇩1@τs⇩2; ⊢ is⇩2,xt⇩2 [::] τs⇩3;
size τs⇩1 = size is⇩1; prefix τs⇩3 τs⇩2 ⟧
⟹ ⊢ is⇩1@is⇩2, xt⇩1 @ shift (size is⇩1) xt⇩2 [::] τs⇩1@τs⇩2"
by(bestsimp simp:prefix_def elim: wt_instrs_ext dest:wt_instrs_extR)
corollary (in TC2) wt_instrs_app:
assumes is⇩1: "⊢ is⇩1,xt⇩1 [::] τs⇩1@[τ]"
assumes is⇩2: "⊢ is⇩2,xt⇩2 [::] τ#τs⇩2"
assumes s: "size τs⇩1 = size is⇩1"
shows "⊢ is⇩1@is⇩2, xt⇩1@shift (size is⇩1) xt⇩2 [::] τs⇩1@τ#τs⇩2"
proof -
from is⇩1 have "⊢ is⇩1,xt⇩1 [::] (τs⇩1@[τ])@τs⇩2"
by (rule wt_instrs_extR)
hence "⊢ is⇩1,xt⇩1 [::] τs⇩1@τ#τs⇩2" by simp
from this is⇩2 s show ?thesis by (rule wt_instrs_ext)
qed
corollary (in TC2) wt_instrs_app_last[trans]:
assumes "⊢ is⇩2,xt⇩2 [::] τ#τs⇩2" "⊢ is⇩1,xt⇩1 [::] τs⇩1"
"last τs⇩1 = τ" "size τs⇩1 = size is⇩1+1"
shows "⊢ is⇩1@is⇩2, xt⇩1@shift (size is⇩1) xt⇩2 [::] τs⇩1@τs⇩2"
using assms proof(cases τs⇩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⇩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 ⊆ {..<Suc (length is)}" using wtis by (fastforce simp:wt_instrs_def)
moreover {
fix pc' assume pc': "¬ pc' < length is" "pc' < Suc (length is)"
have "P,T⇩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⇩2,xt⇩2 [::] τ'#τs⇩2; ⊢ is⇩1,xt⇩1 [::] τ#τs⇩1@[τ'];
xt' = xt⇩1 @ shift (size is⇩1) xt⇩2; size τs⇩1+1 = size is⇩1 ⟧
⟹ ⊢ is⇩1@is⇩2,xt' [::] τ#τs⇩1@τ'#τs⇩2"
using wt_instrs_app[where ?τs⇩1.0 = "τ # τs⇩1"] by simp
corollary (in TC2) wt_instrs_app2_simp[trans,simp]:
"⟦ ⊢ is⇩2,xt⇩2 [::] τ'#τs⇩2; ⊢ is⇩1,xt⇩1 [::] τ#τs⇩1@[τ']; size τs⇩1+1 = size is⇩1 ⟧
⟹ ⊢ is⇩1@is⇩2, xt⇩1@shift (size is⇩1) xt⇩2 [::] τ#τs⇩1@τ'#τs⇩2"
using wt_instrs_app[where ?τs⇩1.0 = "τ # τs⇩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⇩1.0 = "[i]" and ?τs⇩1.0 = "[]" and ?is⇩2.0 = "is"
and ?xt⇩1.0 = "[]"]
by simp
corollary (in TC2) wt_instrs_Cons2[trans]:
assumes τs: "⊢ is,xt [::] τs"
assumes i: "P,T⇩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⇩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⇩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⇩i' ST E A, ty⇩i' (Class C#ST) E A]"
by(simp add:wt_defs ty⇩i'_def)
lemma (in TC2) wt_Cast:
"is_class P C ⟹
⊢ [Checkcast C],[] [::] [ty⇩i' (Class D # ST) E A, ty⇩i' (Class C # ST) E A]"
by(simp add: ty⇩i'_def wt_defs)
lemma (in TC2) wt_Push:
"⟦ size ST < mxs; typeof v = Some T ⟧
⟹ ⊢ [Push v],[] [::] [ty⇩i' ST E A, ty⇩i' (T#ST) E A]"
by(simp add: ty⇩i'_def wt_defs)
lemma (in TC2) wt_Pop:
"⊢ [Pop],[] [::] (ty⇩i' (T#ST) E A # ty⇩i' ST E A # τs)"
by(simp add: ty⇩i'_def wt_defs)
lemma (in TC2) wt_CmpEq:
"⟦ P ⊢ T⇩1 ≤ T⇩2 ∨ P ⊢ T⇩2 ≤ T⇩1⟧
⟹ ⊢ [CmpEq],[] [::] [ty⇩i' (T⇩2 # T⇩1 # ST) E A, ty⇩i' (Boolean # ST) E A]"
by(auto simp:ty⇩i'_def wt_defs elim!: refTE not_refTE)
lemma (in TC2) wt_IAdd:
"⊢ [IAdd],[] [::] [ty⇩i' (Integer#Integer#ST) E A, ty⇩i' (Integer#ST) E A]"
by(simp add:ty⇩i'_def wt_defs)
lemma (in TC2) wt_Load:
"⟦ size ST < mxs; size E ≤ mxl; i ∈∈ A; i < size E ⟧
⟹ ⊢ [Load i],[] [::] [ty⇩i' ST E A, ty⇩i' (E!i # ST) E A]"
by(auto simp add:ty⇩i'_def wt_defs ty⇩l_def hyperset_defs)
lemma (in TC2) wt_Store:
"⟦ P ⊢ T ≤ E!i; i < size E; size E ≤ mxl ⟧ ⟹
⊢ [Store i],[] [::] [ty⇩i' (T#ST) E A, ty⇩i' ST E (⌊{i}⌋ ⊔ A)]"
by(auto simp:hyperset_defs nth_list_update ty⇩i'_def wt_defs ty⇩l_def
intro:list_all2_all_nthI)
lemma (in TC2) wt_Get:
"⟦ P ⊢ C sees F,NonStatic:T in D ⟧ ⟹
⊢ [Getfield F D],[] [::] [ty⇩i' (Class C # ST) E A, ty⇩i' (T # ST) E A]"
by(auto simp: ty⇩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⇩i' ST E A, ty⇩i' (T # ST) E A]"
by(auto simp: ty⇩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⇩i' (T' # Class C # ST) E A, ty⇩i' ST E A]"
by(auto intro: sees_field_idemp sees_field_decl_above simp: ty⇩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⇩i' (T' # ST) E A, ty⇩i' ST E A]"
by(auto intro: sees_field_idemp sees_field_decl_above simp: ty⇩i'_def wt_defs)
lemma (in TC2) wt_Throw:
"⊢ [Throw],[] [::] [ty⇩i' (Class C # ST) E A, τ']"
by(auto simp: ty⇩i'_def wt_defs)
lemma (in TC2) wt_IfFalse:
"⟦ 2 ≤ i; nat i < size τs + 2; P ⊢ ty⇩i' ST E A ≤' τs ! nat(i - 2) ⟧
⟹ ⊢ [IfFalse i],[] [::] ty⇩i' (Boolean # ST) E A # ty⇩i' ST E A # τs"
by(simp add: ty⇩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⇩i' (rev Ts' @ Class C # ST) E A, ty⇩i' (T#ST) E A]"
by(fastforce simp add: ty⇩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⇩i' (rev Ts' @ ST) E A, ty⇩i' (T#ST) E A]"
by(fastforce simp add: ty⇩i'_def wt_defs)
corollary (in TC2) wt_instrs_app3[simp]:
"⟦ ⊢ is⇩2,[] [::] (τ' # τs⇩2); ⊢ is⇩1,xt⇩1 [::] τ # τs⇩1 @ [τ']; size τs⇩1+1 = size is⇩1⟧
⟹ ⊢ (is⇩1 @ is⇩2),xt⇩1 [::] τ # τs⇩1 @ τ' # τs⇩2"
using wt_instrs_app2[where ?xt⇩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⇩1 @ is⇩2, xt [::] τs⇩1 @ ty⇩i' (Class C # ST) E A # τs⇩2"
and P_τs⇩1: "∀τ ∈ set τs⇩1. ∀ST' LT'. τ = Some(ST',LT') ⟶
size ST ≤ size ST' ∧ P ⊢ Some (drop (size ST' - size ST) ST',LT') ≤' ty⇩i' ST E A"
and is_τs: "size is⇩1 = size τs⇩1" and PC: "is_class P C" and ST_mxs: "size ST < mxs"
shows "⊢ is⇩1 @ is⇩2, xt @ [(0,size is⇩1 - 1,C,size is⇩1,size ST)] [::] τs⇩1 @ ty⇩i' (Class C # ST) E A # τs⇩2"
(is "⊢ ?is, xt@[?xte] [::] ?τs" is "⊢ ?is, ?xt' [::] ?τs")
proof -
have P_τs⇩1': "⋀τ. τ ∈ set τs⇩1 ⟹ (∀ST' LT'. τ = Some(ST',LT') ⟶
size ST ≤ size ST' ∧ P ⊢ Some (drop (size ST' - size ST) ST',LT') ≤' ty⇩i' ST E A)"
using P_τs⇩1 by fast
have "size ?is < size ?τs" and "pcs ?xt' ⊆ {0..<size ?is}"
and P_pc: "⋀pc. pc< size ?is ⟹ P,T⇩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⇩r,mxs,?mpc,xt ⊢ ?i,pc :: ?τs" by(rule P_pc)
then have app: "app ?i P mxs T⇩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⇩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⇩i: "app⇩i (?i,P,pc,mxs,T⇩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⇩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⇩1" by simp
then have "τs⇩1 ! pc = ?t" by(fastforce simp: nth_append)
moreover have τs⇩1_pc: "τs⇩1 ! pc ∈ set τs⇩1" by(rule nth_mem[OF True'])
ultimately show ?thesis
using Some τ PC ST_mxs xcpt_app P_τs⇩1'[OF τs⇩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⇩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⇩1" by simp
have τs⇩1_pc: "τs⇩1 ! pc ∈ set τs⇩1" by(rule nth_mem[OF True'])
have "P ⊢ ⌊(Class C # drop (length ST' - length ST) ST', LT')⌋
≤' ty⇩i' (Class C # ST) E A"
using True' Some τ P_τs⇩1'[OF τs⇩1_pc]
by (fastforce simp: nth_append ty⇩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⇩i by(fastforce simp add: app_def)
qed simp
then have "P,T⇩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⇩0) ≤ length xs"
and "drop (length xs - Suc (length xs⇩0)) xs = x # xs⇩0"
shows "drop (length xs - length xs⇩0) xs = xs⇩0"
using assms proof(cases xs)
case (Cons a list) then show ?thesis using assms
proof(cases "length list - length xs⇩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⇩i'_conv[simp]
lemma (in TC1) compT_ST_prefix:
"⋀E A ST⇩0. ⌊(ST,LT)⌋ ∈ set(compT E A ST⇩0 e) ⟹
size ST⇩0 ≤ size ST ∧ drop (size ST - size ST⇩0) ST = ST⇩0"
and
"⋀E A ST⇩0. ⌊(ST,LT)⌋ ∈ set(compTs E A ST⇩0 es) ⟹
size ST⇩0 ≤ size ST ∧ drop (size ST - size ST⇩0) ST = ST⇩0"
proof(induct e and es rule: compT.induct compTs.induct)
case (FAss e⇩1 F D e⇩2)
moreover {
let ?ST⇩0 = "ty E e⇩1 # ST⇩0"
fix A assume "⌊(ST, LT)⌋ ∈ set (compT E A ?ST⇩0 e⇩2)"
with FAss
have "length ?ST⇩0 ≤ length ST ∧ drop (size ST - size ?ST⇩0) ST = ?ST⇩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⇩0 = "ty E e # ST⇩0"
fix A assume "⌊(ST, LT)⌋ ∈ set (compTs E A ?ST⇩0 es)"
with Call
have "length ?ST⇩0 ≤ length ST ∧ drop (size ST - size ?ST⇩0) ST = ?ST⇩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⇩0 = "ty E e # ST⇩0"
fix A assume "⌊(ST, LT)⌋ ∈ set (compTs E A ?ST⇩0 es)"
with Cons_exp
have "length ?ST⇩0 ≤ length ST ∧ drop (size ST - size ?ST⇩0) ST = ?ST⇩0" by blast
hence ?case by (clarsimp simp add: drop_mess)
}
ultimately show ?case by auto
next
case (BinOp e⇩1 bop e⇩2)
moreover {
let ?ST⇩0 = "ty E e⇩1 # ST⇩0"
fix A assume "⌊(ST, LT)⌋ ∈ set (compT E A ?ST⇩0 e⇩2)"
with BinOp
have "length ?ST⇩0 ≤ length ST ∧ drop (size ST - size ?ST⇩0) ST = ?ST⇩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⇩i'_conv[simp del]
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 ⊢⇩1 e :: T; 𝒟 e A; ℬ e (size E);
size ST + max_stack e ≤ mxs; size E + max_vars e ≤ mxl ⟧
⟹ ⊢ compE⇩2 e, compxE⇩2 e 0 (size ST) [::]
ty⇩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 ⊢⇩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⇩i' ST E A # compTs E A ST es in
⊢ compEs⇩2 es,compxEs⇩2 es 0 (size ST) [::] τs ∧
last τs = ty⇩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⇩2.induct compxEs⇩2.induct)
case (TryCatch e⇩1 C i e⇩2)
hence [simp]: "i = size E" by simp
have wt⇩1: "P,E ⊢⇩1 e⇩1 :: T" and wt⇩2: "P,E@[Class C] ⊢⇩1 e⇩2 :: T"
and "class": "is_class P C" using TryCatch by auto
let ?A⇩1 = "A ⊔ 𝒜 e⇩1" let ?A⇩i = "A ⊔ ⌊{i}⌋" let ?E⇩i = "E @ [Class C]"
let ?τ = "ty⇩i' ST E A" let ?τs⇩1 = "compT E A ST e⇩1"
let ?τ⇩1 = "ty⇩i' (T#ST) E ?A⇩1" let ?τ⇩2 = "ty⇩i' (Class C#ST) E A"
let ?τ⇩3 = "ty⇩i' ST ?E⇩i ?A⇩i" let ?τs⇩2 = "compT ?E⇩i ?A⇩i ST e⇩2"
let ?τ⇩2' = "ty⇩i' (T#ST) ?E⇩i (?A⇩i ⊔ 𝒜 e⇩2)"
let ?τ' = "ty⇩i' (T#ST) E (A ⊔ 𝒜 e⇩1 ⊓ (𝒜 e⇩2 ⊖ i))"
let ?go = "Goto (int(size(compE⇩2 e⇩2)) + 2)"
have "PROP ?P e⇩2 ?E⇩i T ?A⇩i ST" by fact
hence "⊢ compE⇩2 e⇩2,compxE⇩2 e⇩2 0 (size ST) [::] (?τ⇩3 # ?τs⇩2) @ [?τ⇩2']"
using TryCatch.prems by(auto simp:after_def)
also have "?A⇩i ⊔ 𝒜 e⇩2 = (A ⊔ 𝒜 e⇩2) ⊔ ⌊{size E}⌋"
by(fastforce simp:hyperset_defs)
also have "P ⊢ ty⇩i' (T#ST) ?E⇩i … ≤' ty⇩i' (T#ST) E (A ⊔ 𝒜 e⇩2)"
by(simp add:hyperset_defs ty⇩l_incr ty⇩i'_def)
also have "P ⊢ … ≤' ty⇩i' (T#ST) E (A ⊔ 𝒜 e⇩1 ⊓ (𝒜 e⇩2 ⊖ i))"
by(auto intro!: ty⇩l_antimono simp:hyperset_defs ty⇩i'_def)
also have "(?τ⇩3 # ?τs⇩2) @ [?τ'] = ?τ⇩3 # ?τs⇩2 @ [?τ']" by simp
also have "⊢ [Store i],[] [::] ?τ⇩2 # [] @ [?τ⇩3]"
using TryCatch.prems
by(auto simp:nth_list_update wt_defs ty⇩i'_def ty⇩l_def
list_all2_conv_all_nth hyperset_defs)
also have "[] @ (?τ⇩3 # ?τs⇩2 @ [?τ']) = (?τ⇩3 # ?τs⇩2 @ [?τ'])" by simp
also have "P,T⇩r,mxs,size(compE⇩2 e⇩2)+3,[] ⊢ ?go,0 :: ?τ⇩1#?τ⇩2#?τ⇩3#?τs⇩2 @ [?τ']" using wt⇩2
by (auto simp: hyperset_defs ty⇩i'_def wt_defs nth_Cons nat_add_distrib
fun_of_def intro: ty⇩l_antimono list_all2_refl split:nat.split)
also have "⊢ compE⇩2 e⇩1,compxE⇩2 e⇩1 0 (size ST) [::] ?τ # ?τs⇩1 @ [?τ⇩1]"
using TryCatch by(auto simp:after_def)
also have "?τ # ?τs⇩1 @ ?τ⇩1 # ?τ⇩2 # ?τ⇩3 # ?τs⇩2 @ [?τ'] =
(?τ # ?τs⇩1 @ [?τ⇩1]) @ ?τ⇩2 # ?τ⇩3 # ?τs⇩2 @ [?τ']" by simp
also have "compE⇩2 e⇩1 @ ?go # [Store i] @ compE⇩2 e⇩2 =
(compE⇩2 e⇩1 @ [?go]) @ (Store i # compE⇩2 e⇩2)" by simp
also
let "?Q τ" = "∀ST' LT'. τ = ⌊(ST', LT')⌋ ⟶
size ST ≤ size ST' ∧ P ⊢ Some (drop (size ST' - size ST) ST',LT') ≤' ty⇩i' ST E A"
{
have "?Q (ty⇩i' ST E A)" by (clarsimp simp add: ty⇩i'_def)
moreover have "?Q (ty⇩i' (T # ST) E ?A⇩1)"
by (fastforce simp add: ty⇩i'_def hyperset_defs intro!: ty⇩l_antimono)
moreover have "⋀τ. τ ∈ set (compT E A ST e⇩1) ⟹ ?Q τ" using TryCatch.prems
by clarsimp (frule compT_ST_prefix,
fastforce dest!: compT_LT_prefix simp add: ty⇩i'_def)
ultimately
have "∀τ∈set (ty⇩i' ST E A # compT E A ST e⇩1 @ [ty⇩i' (T # ST) E ?A⇩1]). ?Q τ"
by auto
}
also from TryCatch.prems max_stack1[OF wt⇩1] have "size ST + 1 ≤ mxs" by auto
ultimately show ?case using wt⇩1 wt⇩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⇩1 bop e⇩2)
let ?op = "case bop of Eq ⇒ [CmpEq] | Add ⇒ [IAdd]"
have T: "P,E ⊢⇩1 e⇩1 «bop» e⇩2 :: T" by fact
then obtain T⇩1 T⇩2 where T⇩1: "P,E ⊢⇩1 e⇩1 :: T⇩1" and T⇩2: "P,E ⊢⇩1 e⇩2 :: T⇩2" and
bopT: "case bop of Eq ⇒ (P ⊢ T⇩1 ≤ T⇩2 ∨ P ⊢ T⇩2 ≤ T⇩1) ∧ T = Boolean
| Add ⇒ T⇩1 = Integer ∧ T⇩2 = Integer ∧ T = Integer" by auto
let ?A⇩1 = "A ⊔ 𝒜 e⇩1" let ?A⇩2 = "?A⇩1 ⊔ 𝒜 e⇩2"
let ?τ = "ty⇩i' ST E A" let ?τs⇩1 = "compT E A ST e⇩1"
let ?τ⇩1 = "ty⇩i' (T⇩1#ST) E ?A⇩1" let ?τs⇩2 = "compT E ?A⇩1 (T⇩1#ST) e⇩2"
let ?τ⇩2 = "ty⇩i' (T⇩2#T⇩1#ST) E ?A⇩2" let ?τ' = "ty⇩i' (T#ST) E ?A⇩2"
from bopT have "⊢ ?op,[] [::] [?τ⇩2,?τ']"
by (cases bop) (auto simp add: wt_CmpEq wt_IAdd)
also have "PROP ?P e⇩2 E T⇩2 ?A⇩1 (T⇩1#ST)" by fact
with BinOp.prems T⇩2
have "⊢ compE⇩2 e⇩2, compxE⇩2 e⇩2 0 (size (T⇩1#ST)) [::] ?τ⇩1#?τs⇩2@[?τ⇩2]"
by (auto simp: after_def)
also from BinOp T⇩1 have "⊢ compE⇩2 e⇩1, compxE⇩2 e⇩1 0 (size ST) [::] ?τ#?τs⇩1@[?τ⇩1]"
by (auto simp: after_def)
finally show ?case using T T⇩1 T⇩2 by (simp add: after_def hyperUn_assoc)
next
case (Cons_exp e es)
have "P,E ⊢⇩1 e # es [::] Ts" by fact
then obtain T⇩e Ts' where
T⇩e: "P,E ⊢⇩1 e :: T⇩e" and Ts': "P,E ⊢⇩1 es [::] Ts'" and
Ts: "Ts = T⇩e#Ts'" by auto
let ?A⇩e = "A ⊔ 𝒜 e"
let ?τ = "ty⇩i' ST E A" let ?τs⇩e = "compT E A ST e"
let ?τ⇩e = "ty⇩i' (T⇩e#ST) E ?A⇩e" let ?τs' = "compTs E ?A⇩e (T⇩e#ST) es"
let ?τs = "?τ # ?τs⇩e @ (?τ⇩e # ?τs')"
have Ps: "PROP ?Ps es E Ts' ?A⇩e (T⇩e#ST)" by fact
with Cons_exp.prems T⇩e Ts'
have "⊢ compEs⇩2 es, compxEs⇩2 es 0 (size (T⇩e#ST)) [::] ?τ⇩e#?τs'" by (simp add: after_def)
also from Cons_exp T⇩e have "⊢ compE⇩2 e, compxE⇩2 e 0 (size ST) [::] ?τ#?τs⇩e@[?τ⇩e]"
by (auto simp: after_def)
moreover
from Ps Cons_exp.prems T⇩e Ts' Ts
have "last ?τs = ty⇩i' (rev Ts@ST) E (?A⇩e ⊔ 𝒜s es)" by simp
ultimately show ?case using T⇩e by (simp add: after_def hyperUn_assoc)
next
case (FAss e⇩1 F D e⇩2)
hence Void: "P,E ⊢⇩1 e⇩1∙F{D} := e⇩2 :: Void" by auto
then obtain C T T' where
C: "P,E ⊢⇩1 e⇩1 :: Class C" and sees: "P ⊢ C sees F,NonStatic:T in D" and
T': "P,E ⊢⇩1 e⇩2 :: T'" and T'_T: "P ⊢ T' ≤ T" by auto
let ?A⇩1 = "A ⊔ 𝒜 e⇩1" let ?A⇩2 = "?A⇩1 ⊔ 𝒜 e⇩2"
let ?τ = "ty⇩i' ST E A" let ?τs⇩1 = "compT E A ST e⇩1"
let ?τ⇩1 = "ty⇩i' (Class C#ST) E ?A⇩1" let ?τs⇩2 = "compT E ?A⇩1 (Class C#ST) e⇩2"
let ?τ⇩2 = "ty⇩i' (T'#Class C#ST) E ?A⇩2" let ?τ⇩3 = "ty⇩i' ST E ?A⇩2"
let ?τ' = "ty⇩i' (Void#ST) E ?A⇩2"
from FAss.prems sees T'_T
have "⊢ [Putfield F D,Push Unit],[] [::] [?τ⇩2,?τ⇩3,?τ']"
by (fastforce simp add: wt_Push wt_Put)
also have "PROP ?P e⇩2 E T' ?A⇩1 (Class C#ST)" by fact
with FAss.prems T'
have "⊢ compE⇩2 e⇩2, compxE⇩2 e⇩2 0 (size ST+1) [::] ?τ⇩1#?τs⇩2@[?τ⇩2]"
by (auto simp add: after_def hyperUn_assoc)
also from FAss C have "⊢ compE⇩2 e⇩1, compxE⇩2 e⇩1 0 (size ST) [::] ?τ#?τs⇩1@[?τ⇩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⇩2)
hence Void: "P,E ⊢⇩1 C∙⇩sF{D} := e⇩2 :: Void" by auto
then obtain T T' where
sees: "P ⊢ C sees F,Static:T in D" and
T': "P,E ⊢⇩1 e⇩2 :: T'" and T'_T: "P ⊢ T' ≤ T" by auto
let ?A⇩2 = "A ⊔ 𝒜 e⇩2"
let ?τ = "ty⇩i' ST E A" let ?τs⇩2 = "compT E A ST e⇩2"
let ?τ⇩2 = "ty⇩i' (T'#ST) E ?A⇩2" let ?τ⇩3 = "ty⇩i' ST E ?A⇩2"
let ?τ' = "ty⇩i' (Void#ST) E ?A⇩2"
from SFAss.prems sees T'_T max_stack1[OF T']
have "⊢ [Putstatic C F D,Push Unit],[] [::] [?τ⇩2,?τ⇩3,?τ']"
by (fastforce simp add: wt_Push wt_PutS)
also have "PROP ?P e⇩2 E T' A ST" by fact
with SFAss.prems T'
have "⊢ compE⇩2 e⇩2, compxE⇩2 e⇩2 0 (size ST) [::] ?τ#?τs⇩2@[?τ⇩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⇩i e)
let ?τs = "ty⇩i' ST E A # compT (E @ [T⇩i]) (A⊖i) ST e"
have IH: "PROP ?P e (E@[T⇩i]) T (A⊖i) ST" by fact
hence "⊢ compE⇩2 e, compxE⇩2 e 0 (size ST) [::]
?τs @ [ty⇩i' (T#ST) (E@[T⇩i]) (A⊖(size E) ⊔ 𝒜 e)]"
using Block.prems by (auto simp add: after_def)
also have "P ⊢ ty⇩i' (T # ST) (E@[T⇩i]) (A ⊖ size E ⊔ 𝒜 e) ≤'
ty⇩i' (T # ST) (E@[T⇩i]) ((A ⊔ 𝒜 e) ⊖ size E)"
by(auto simp add:hyperset_defs intro: ty⇩i'_antimono)
also have "… = ty⇩i' (T # ST) E (A ⊔ 𝒜 e)" by simp
also have "P ⊢ … ≤' ty⇩i' (T # ST) E (A ⊔ (𝒜 e ⊖ i))"
by(auto simp add:hyperset_defs intro: ty⇩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 ⊢⇩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 ⊢⇩1 e :: Boolean" and wtc: "P,E ⊢⇩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⇩0 = "A ⊔ 𝒜 e" let ?A⇩1 = "?A⇩0 ⊔ 𝒜 c"
let ?τ = "ty⇩i' ST E A" let ?τs⇩e = "compT E A ST e"
let ?τ⇩e = "ty⇩i' (Boolean#ST) E ?A⇩0" let ?τ⇩1 = "ty⇩i' ST E ?A⇩0"
let ?τs⇩c = "compT E ?A⇩0 ST c" let ?τ⇩c = "ty⇩i' (Tc#ST) E ?A⇩1"
let ?τ⇩2 = "ty⇩i' ST E ?A⇩1" let ?τ' = "ty⇩i' (Void#ST) E ?A⇩0"
let ?τs = "(?τ # ?τs⇩e @ [?τ⇩e]) @ ?τ⇩1 # ?τs⇩c @ [?τ⇩c, ?τ⇩2, ?τ⇩1, ?τ']"
have "⊢ [],[] [::] [] @ ?τs" by(simp add:wt_instrs_def)
also
have "PROP ?P e E Boolean A ST" by fact
hence "⊢ compE⇩2 e,compxE⇩2 e 0 (size ST) [::] ?τ # ?τs⇩e @ [?τ⇩e]"
using While.prems by (auto simp:after_def)
also
have "[] @ ?τs = (?τ # ?τs⇩e) @ ?τ⇩e # ?τ⇩1 # ?τs⇩c @ [?τ⇩c,?τ⇩2,?τ⇩1,?τ']" by simp
also
let ?n⇩e = "size(compE⇩2 e)" let ?n⇩c = "size(compE⇩2 c)"
thm compT_sizes(1)
let ?if = "IfFalse (int ?n⇩c + 3)"
from wtc wte have "compE⇩2 c ≠ []" and "compE⇩2 e ≠ []" using WT⇩1_nsub_RI by auto
then have compe2c: "length (compE⇩2 c) > 0" and compe2e: "length (compE⇩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],[] [::] ?τ⇩e # ?τ⇩1 # ?τs⇩c @ [?τ⇩c, ?τ⇩2, ?τ⇩1, ?τ']"
proof-{
show ?thesis
apply (rule wt_IfFalse)
apply simp
apply simp
apply (subgoal_tac "length (compE⇩2 c) = length (compT E (A ⊔ 𝒜 e) ST c) + 1"
"nat (int (length (compE⇩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⇩e) @ (?τ⇩e # ?τ⇩1 # ?τs⇩c @ [?τ⇩c, ?τ⇩2, ?τ⇩1, ?τ']) = ?τs" by simp
also
have "PROP ?P c E Tc ?A⇩0 ST" by fact
hence "⊢ compE⇩2 c,compxE⇩2 c 0 (size ST) [::] ?τ⇩1 # ?τs⇩c @ [?τ⇩c]"
using While.prems wtc by (auto simp:after_def)
also have "?τs = (?τ # ?τs⇩e @ [?τ⇩e,?τ⇩1] @ ?τs⇩c) @ [?τ⇩c,?τ⇩2,?τ⇩1,?τ']" by simp
also have "⊢ [Pop],[] [::] [?τ⇩c, ?τ⇩2]" by(simp add:wt_Pop)
also have "(?τ # ?τs⇩e @ [?τ⇩e,?τ⇩1] @ ?τs⇩c) @ [?τ⇩c,?τ⇩2,?τ⇩1,?τ'] = ?τs" by simp
also let ?go = "Goto (-int(?n⇩c+?n⇩e+2))"
have "P ⊢ ?τ⇩2 ≤' ?τ" by(fastforce intro: ty⇩i'_antimono simp: hyperset_defs)
hence "P,T⇩r,mxs,size ?τs,[] ⊢ ?go,?n⇩e+?n⇩c+2 :: ?τs" using wte wtc
proof-{
let ?t1 = "ty⇩i' ST E A" let ?t2 = "ty⇩i' (Boolean # ST) E (A ⊔ 𝒜 e)" let ?t3 = "ty⇩i' ST E (A ⊔ 𝒜 e)"
let ?t4 = "[ty⇩i' (Tc # ST) E (A ⊔ 𝒜 e ⊔ 𝒜 c), ty⇩i' ST E (A ⊔ 𝒜 e ⊔ 𝒜 c), ty⇩i' ST E (A ⊔ 𝒜 e), ty⇩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⇩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⇩2 c) - 1" using compT_sizes(1)[OF wtc] by simp
have len2: "length ?e = length (compE⇩2 e) - 1" using compT_sizes(1)[OF wte] by simp
hence "length (compE⇩2 e) + length (compE⇩2 c) + 2 > length ?s1"
using len1 compe2c compe2e by auto
hence "(?s1 @ ?t4) ! (length (compE⇩2 e) + length (compE⇩2 c) + 2) =
?t4 ! (length (compE⇩2 e) + length (compE⇩2 c) + 2 - length ?s1)"
by (auto simp only:nth_append split: if_splits)
hence "(?s1 @ ?t4) ! (length (compE⇩2 e) + length (compE⇩2 c) + 2) = ty⇩i' ST E (A ⊔ 𝒜 e ⊔ 𝒜 c)"
using len1 len2 compe2c compe2e by auto
hence "P ⊢ (?s1 @ ?t4) ! (length (compE⇩2 e) + length (compE⇩2 c) + 2) ≤' ty⇩i' ST E A"
using ass by simp
thus "P ⊢ ((?t1 # ?e @ [?t2]) @ ?t3 # ?c @ ?t4) ! (length (compE⇩2 e) + length (compE⇩2 c) + 2) ≤'
((?t1 # ?e @ [?t2]) @ ?t3 # ?c @ ?t4) ! nat (int (length (compE⇩2 e) + length (compE⇩2 c) + 2) +
- int (length (compE⇩2 c) + length (compE⇩2 e) + 2))"
by simp
}qed
}qed
also have "?τs = (?τ # ?τs⇩e @ [?τ⇩e,?τ⇩1] @ ?τs⇩c @ [?τ⇩c, ?τ⇩2]) @ [?τ⇩1, ?τ']"
by simp
also have "⊢ [Push Unit],[] [::] [?τ⇩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⇩1 e⇩2)
obtain T⇩1 T⇩2 where wte: "P,E ⊢⇩1 e :: Boolean"
and wt⇩1: "P,E ⊢⇩1 e⇩1 :: T⇩1" and wt⇩2: "P,E ⊢⇩1 e⇩2 :: T⇩2"
and sub⇩1: "P ⊢ T⇩1 ≤ T" and sub⇩2: "P ⊢ T⇩2 ≤ T"
using Cond by auto
from wte wt⇩1 wt⇩2 have "compE⇩2 e⇩1 ≠ []" and "compE⇩2 e⇩2 ≠ []" and "compE⇩2 e ≠ []" using WT⇩1_nsub_RI by auto
then have compe: "length (compE⇩2 e) > 0"
and compe1: "length (compE⇩2 e⇩1) > 0"
and compe2: "length (compE⇩2 e⇩2) > 0" by auto
have [simp]: "ty E (if (e) e⇩1 else e⇩2) = T" using Cond by simp
let ?A⇩0 = "A ⊔ 𝒜 e" let ?A⇩2 = "?A⇩0 ⊔ 𝒜 e⇩2" let ?A⇩1 = "?A⇩0 ⊔ 𝒜 e⇩1"
let ?A' = "?A⇩0 ⊔ 𝒜 e⇩1 ⊓ 𝒜 e⇩2"
let ?τ⇩2 = "ty⇩i' ST E ?A⇩0" let ?τ' = "ty⇩i' (T#ST) E ?A'"
let ?τs⇩2 = "compT E ?A⇩0 ST e⇩2"
have "PROP ?P e⇩2 E T⇩2 ?A⇩0 ST" by fact
hence "⊢ compE⇩2 e⇩2, compxE⇩2 e⇩2 0 (size ST) [::] (?τ⇩2#?τs⇩2) @ [ty⇩i' (T⇩2#ST) E ?A⇩2]"
using Cond.prems wt⇩2 by(auto simp add:after_def)
also have "P ⊢ ty⇩i' (T⇩2#ST) E ?A⇩2 ≤' ?τ'" using sub⇩2
by(auto simp add: hyperset_defs ty⇩i'_def intro!: ty⇩l_antimono)
also
let ?τ⇩3 = "ty⇩i' (T⇩1 # ST) E ?A⇩1"
let ?g⇩2 = "Goto(int (size (compE⇩2 e⇩2) + 1))"
from sub⇩1 have "P,T⇩r,mxs,size(compE⇩2 e⇩2)+2,[] ⊢ ?g⇩2,0 :: ?τ⇩3#(?τ⇩2#?τs⇩2)@[?τ']" using wt⇩2
by(auto simp: hyperset_defs wt_defs nth_Cons ty⇩i'_def
split:nat.split intro!: ty⇩l_antimono)
also
let ?τs⇩1 = "compT E ?A⇩0 ST e⇩1"
have "PROP ?P e⇩1 E T⇩1 ?A⇩0 ST" by fact
hence "⊢ compE⇩2 e⇩1,compxE⇩2 e⇩1 0 (size ST) [::] ?τ⇩2 # ?τs⇩1 @ [?τ⇩3]"
using Cond.prems wt⇩1 by(auto simp add:after_def)
also
let ?τs⇩1⇩2 = "?τ⇩2 # ?τs⇩1 @ ?τ⇩3 # (?τ⇩2 # ?τs⇩2) @ [?τ']"
let ?τ⇩1 = "ty⇩i' (Boolean#ST) E ?A⇩0"
let ?g⇩1 = "IfFalse(int (size (compE⇩2 e⇩1) + 2))"
let ?code = "compE⇩2 e⇩1 @ ?g⇩2 # compE⇩2 e⇩2"
have len1: "length ?τs⇩1 = length (compE⇩2 e⇩1) - 1" using compT_sizes(1)[OF wt⇩1] by simp
have len2: "length ?τs⇩2 = length (compE⇩2 e⇩2) - 1" using compT_sizes(1)[OF wt⇩2] by simp
have len_auxi: "length (compE⇩2 e⇩1) - (length (compE⇩2 e⇩1) - Suc 0) = Suc 0" using compe1
by (simp add:diff_Suc split:nat.split)
have "⊢ [?g⇩1],[] [::] [?τ⇩1] @ ?τs⇩1⇩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 "[?τ⇩1] @ ?τs⇩1⇩2 = ?τ⇩1 # ?τs⇩1⇩2" by simp also
let ?τ = "ty⇩i' ST E A"
have "PROP ?P e E Boolean A ST" by fact
hence "⊢ compE⇩2 e, compxE⇩2 e 0 (size ST) [::] ?τ # compT E A ST e @ [?τ⇩1]"
using Cond.prems wte by(auto simp add:after_def)
finally show ?case using wte wt⇩1 wt⇩2 by(simp add:after_def hyperUn_assoc)
next
case (Call e M es)
obtain C D Ts m Ts' where C: "P,E ⊢⇩1 e :: Class C"
and "method": "P ⊢ C sees M,NonStatic:Ts → T = m in D"
and wtes: "P,E ⊢⇩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⇩1_same_size)
let ?A⇩0 = "A ⊔ 𝒜 e" let ?A⇩1 = "?A⇩0 ⊔ 𝒜s es"
let ?τ = "ty⇩i' ST E A" let ?τs⇩e = "compT E A ST e"
let ?τ⇩e = "ty⇩i' (Class C # ST) E ?A⇩0"
let ?τs⇩e⇩s = "compTs E ?A⇩0 (Class C # ST) es"
let ?τ⇩1 = "ty⇩i' (rev Ts' @ Class C # ST) E ?A⇩1"
let ?τ' = "ty⇩i' (T # ST) E ?A⇩1"
have "⊢ [Invoke M (size es)],[] [::] [?τ⇩1,?τ']"
by(rule wt_Invoke[OF same_size "method" subs])
also
have "PROP ?Ps es E Ts' ?A⇩0 (Class C # ST)" by fact
hence "⊢ compEs⇩2 es,compxEs⇩2 es 0 (size ST+1) [::] ?τ⇩e # ?τs⇩e⇩s"
"last (?τ⇩e # ?τs⇩e⇩s) = ?τ⇩1"
using Call.prems wtes by(auto simp add:after_def)
also have "(?τ⇩e # ?τs⇩e⇩s) @ [?τ'] = ?τ⇩e # ?τs⇩e⇩s @ [?τ']" by simp
also have "⊢ compE⇩2 e,compxE⇩2 e 0 (size ST) [::] ?τ # ?τs⇩e @ [?τ⇩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 ⊢⇩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⇩1_same_size)
have mxs: "length ST < mxs" using WT⇩1_nsub_RI[OF SCall.prems(1)] SCall.prems(4) by simp
let ?A⇩1 = "A ⊔ 𝒜s es"
let ?τ = "ty⇩i' ST E A"
let ?τs⇩e⇩s = "compTs E A ST es"
let ?τ⇩1 = "ty⇩i' (rev Ts' @ ST) E ?A⇩1"
let ?τ' = "ty⇩i' (T # ST) E ?A⇩1"
have "⊢ [Invokestatic C M (size es)],[] [::] [?τ⇩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⇩2 es,compxEs⇩2 es 0 (size ST) [::] ?τ # ?τs⇩e⇩s"
"last (?τ # ?τs⇩e⇩s) = ?τ⇩1"
using SCall.prems wtes by(auto simp add:after_def)
also have "(?τ # ?τs⇩e⇩s) @ [?τ'] = ?τ # ?τs⇩e⇩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 ⊢⇩1 INIT C (Cs,b) ← e :: T" by fact
thus ?case using WT⇩1_nsub_RI by simp
next
case (RI C e' Cs e)
have "P,E ⊢⇩1 RI (C,e') ; Cs ← e :: T" by fact
thus ?case using WT⇩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⇩i (i, compP f P, pc, mpc, T, τ) = app⇩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⇩0
defines [simp]: "E ≡ Class C # Ts"
and [simp]: "A ≡ ⌊{..size Ts}⌋"
and [simp]: "A' ≡ A ⊔ 𝒜 e"
and [simp]: "mxl⇩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 ⊢⇩1 e :: T" and 𝒟: "𝒟 e A"
and ℬ: "ℬ e (size E)" and E_P: "set E ⊆ types P" and wid: "P ⊢ T ≤ T⇩r"
shows "wt_method (compP⇩2 P) C NonStatic Ts T⇩r mxs mxl⇩0 (compE⇩2 e @ [Return])
(compxE⇩2 e 0 0) (ty⇩i' [] E A # compT⇩a E A [] e)"
(is "wt_method ?P C ?b Ts T⇩r mxs mxl⇩0 ?is ?xt ?τs")
proof -
let ?n = "length E + mxl⇩0"
have wt_compE: "P,T⇩r,mxs ⊢ compE⇩2 e, compxE⇩2 e 0 (length []) [::]
TC0.ty⇩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⇩r] by simp
have "OK (ty⇩i' [T] E A') ∈ states P mxs mxl"
using mxs WT⇩1_is_type[OF wf wte E_P] max_stack1[OF wte] OK_ty⇩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⇩a_def after_def check_types_def)
moreover have "wt_start ?P C ?b Ts mxl⇩0 ?τs" using mxl
by (auto simp: wt_start_def ty⇩i'_def ty⇩l_def list_all2_conv_all_nth nth_Cons
split: nat.split)
moreover {
fix pc assume pc: "pc < size ?is"
then have "?P,T⇩r,mxs,size ?is,?xt ⊢ ?is!pc,pc :: ?τs"
proof(cases "pc < length (compE⇩2 e)")
case True
then show ?thesis using mxs wte wt_compE
by (clarsimp simp: compT⇩a_def mxl after_def wt_instrs_def)
next
case False
have "length (compE⇩2 e) = pc" using less_antisym[OF False] pc by simp
then show ?thesis using mxl wte E_P wid
by (clarsimp simp: compT⇩a_def after_def wt_defs xcpt_app_pcs xcpt_eff_pcs ty⇩i'_def)
qed
}
moreover have "0 < size ?is" and "size ?τs = size ?is"
using wte by (simp_all add: compT⇩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⇩0
defines [simp]: "E ≡ Ts"
and [simp]: "A ≡ ⌊{..<size Ts}⌋"
and [simp]: "A' ≡ A ⊔ 𝒜 e"
and [simp]: "mxl⇩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 ⊢⇩1 e :: T" and 𝒟: "𝒟 e A"
and ℬ: "ℬ e (size E)" and E_P: "set E ⊆ types P" and wid: "P ⊢ T ≤ T⇩r"
shows "wt_method (compP⇩2 P) C Static Ts T⇩r mxs mxl⇩0 (compE⇩2 e @ [Return])
(compxE⇩2 e 0 0) (ty⇩i' [] E A # compT⇩a E A [] e)"
(is "wt_method ?P C ?b Ts T⇩r mxs mxl⇩0 ?is ?xt ?τs")
proof -
let ?n = "length E + mxl⇩0"
have wt_compE: "P,T⇩r,mxs ⊢ compE⇩2 e, compxE⇩2 e 0 (length []) [::]
TC0.ty⇩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⇩r] by simp
have "OK (ty⇩i' [T] E A') ∈ states P mxs mxl"
using mxs WT⇩1_is_type[OF wf wte E_P] max_stack1[OF wte] OK_ty⇩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⇩a_def after_def check_types_def)
moreover have "wt_start ?P C ?b Ts mxl⇩0 ?τs" using mxl
by (auto simp: wt_start_def ty⇩i'_def ty⇩l_def list_all2_conv_all_nth nth_Cons
split: nat.split)
moreover {
fix pc assume pc: "pc < size ?is"
then have "?P,T⇩r,mxs,size ?is,?xt ⊢ ?is!pc,pc :: ?τs"
proof(cases "pc < length (compE⇩2 e)")
case True
then show ?thesis using mxs wte wt_compE
by (clarsimp simp: compT⇩a_def mxl after_def wt_instrs_def)
next
case False
have "length (compE⇩2 e) = pc" using less_antisym[OF False] pc by simp
then show ?thesis using mxl wte E_P wid
by (clarsimp simp: compT⇩a_def after_def wt_defs xcpt_app_pcs xcpt_eff_pcs ty⇩i'_def)
qed
}
moreover have "0 < size ?is" and "size ?τs = size ?is"
using wte by (simp_all add: compT⇩a_def)
ultimately show ?thesis by(simp add: wt_method_def)
qed
end
definition compTP :: "J⇩1_prog ⇒ ty⇩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 ⇒ ⌊{..<size Ts}⌋ | NonStatic ⇒ ⌊{..size Ts}⌋;
mxl = (case b of Static ⇒ 0 | NonStatic ⇒ 1) + size Ts + max_vars e
in (TC0.ty⇩i' mxl [] E A # TC1.compT⇩a P mxl E A [] e))"
theorem wt_compP⇩2:
assumes wf: "wf_J⇩1_prog P" shows "wf_jvm_prog (compP⇩2 P)"
proof -
let ?Φ = "compTP P" and ?f = compMb⇩2
let ?wf⇩2 = "λP C (M, b, Ts, T⇩r, mxs, mxl⇩0, is, xt).
wt_method P C b Ts T⇩r mxs mxl⇩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⇩1 mdecl"
assume cM: "P ⊢ C sees M, b : Ts→T = m in C"
and wfm: "wf_mdecl wf_J⇩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⇩1: "wf_J⇩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⇩2 ?P C (M,b,Ts,T,?f b m)"
proof(cases b)
case Static
then obtain T' where wte: "P,Ts ⊢⇩1 m :: T'" and wid: "P ⊢ T' ≤ T"
and 𝒟: "𝒟 m ⌊{..<size Ts}⌋" and ℬ: "ℬ m (size Ts)"
using wfm⇩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 ⊢⇩1 m :: T'" and wid: "P ⊢ T' ≤ T"
and 𝒟: "𝒟 m ⌊{..size Ts}⌋" and ℬ: "ℬ m (Suc (size Ts))"
using wfm⇩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⇩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⇩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⇩2 compP⇩1_pres_wf)
end