Up to index of Isabelle/HOL/objinit
theory Effect = JVMType + JVMExec:(* Title: HOL/MicroJava/BV/Effect.thy
ID: $Id: Effect.html,v 1.1 2002/11/28 14:12:09 kleing Exp $
Author: Gerwin Klein
Copyright 2000 Technische Universitaet Muenchen
*)
header {* \isaheader{Effect of Instructions on the State Type} *}
theory Effect = JVMType + JVMExec:
types
succ_type = "(p_count × state_bool option) list"
text {* Program counter of successor instructions: *}
consts
succs :: "instr \<Rightarrow> p_count \<Rightarrow> p_count list"
primrec
"succs (Load idx) pc = [pc+1]"
"succs (Store idx) pc = [pc+1]"
"succs (LitPush v) pc = [pc+1]"
"succs (Getfield F C) pc = [pc+1]"
"succs (Putfield F C) pc = [pc+1]"
"succs (New C) pc = [pc+1]"
"succs (Checkcast C) pc = [pc+1]"
"succs Pop pc = [pc+1]"
"succs Dup pc = [pc+1]"
"succs Dup_x1 pc = [pc+1]"
"succs Dup_x2 pc = [pc+1]"
"succs Swap pc = [pc+1]"
"succs IAdd pc = [pc+1]"
"succs (Ifcmpeq b) pc = [pc+1, nat (int pc + b)]"
"succs (Goto b) pc = [nat (int pc + b)]"
"succs Return pc = [pc]"
"succs (Invoke C mn fpTs) pc = [pc+1]"
"succs (Invoke_special C mn fpTs) pc
= [pc+1]"
"succs Throw pc = [pc]"
consts theClass :: "init_ty \<Rightarrow> ty"
primrec
"theClass (PartInit C) = Class C"
"theClass (UnInit C pc) = Class C"
text "Effect of instruction on the state type:"
consts
eff' :: "instr × jvm_prog × p_count × state_type \<Rightarrow> state_type"
recdef eff' "{}"
"eff' (Load idx, G, pc, (ST, LT)) = (ok_val (LT ! idx) # ST, LT)"
"eff' (Store idx, G, pc, (ts#ST, LT)) = (ST, LT[idx:= OK ts])"
"eff' (LitPush v, G, pc, (ST, LT)) = (Init (the (typeof (\<lambda>v. None) v))#ST, LT)"
"eff' (Getfield F C, G, pc, (oT#ST, LT)) = (Init (snd (the (field (G,C) F)))#ST, LT)"
"eff' (Putfield F C, G, pc, (vT#oT#ST, LT)) = (ST,LT)"
"eff' (New C, G, pc, (ST,LT)) = (UnInit C pc # ST, replace (OK (UnInit C pc)) Err LT)"
"eff' (Checkcast C,G,pc,(Init (RefT t)#ST,LT)) = (Init (Class C) # ST,LT)"
"eff' (Pop, G, pc, (ts#ST,LT)) = (ST,LT)"
"eff' (Dup, G, pc, (ts#ST,LT)) = (ts#ts#ST,LT)"
"eff' (Dup_x1, G, pc, (ts1#ts2#ST,LT)) = (ts1#ts2#ts1#ST,LT)"
"eff' (Dup_x2, G, pc, (ts1#ts2#ts3#ST,LT)) = (ts1#ts2#ts3#ts1#ST,LT)"
"eff' (Swap, G, pc, (ts1#ts2#ST,LT)) = (ts2#ts1#ST,LT)"
"eff' (IAdd, G, pc, (t1#t2#ST,LT)) = (Init (PrimT Integer)#ST,LT)"
"eff' (Ifcmpeq b, G, pc, (ts1#ts2#ST,LT)) = (ST,LT)"
"eff' (Goto b, G, pc, s) = s"
-- "Return has no successor instruction in the same method:"
"eff' (Return, G, pc, s) = s"
-- "Throw always terminates abruptly:"
"eff' (Throw, G, pc, s) = s"
"eff' (Invoke C mn fpTs, G, pc, (ST,LT)) =
(let ST' = drop (length fpTs) ST;
X = hd ST';
ST'' = tl ST';
rT = fst (snd (the (method (G,C) (mn,fpTs))))
in ((Init rT)#ST'', LT))"
"eff' (Invoke_special C mn fpTs, G, pc, (ST,LT)) =
(let ST' = drop (length fpTs) ST;
X = hd ST';
N = Init (theClass X);
ST'' = replace X N (tl ST');
LT' = replace (OK X) (OK N) LT;
rT = fst (snd (the (method (G,C) (mn,fpTs))))
in ((Init rT)#ST'', LT'))"
text {*
For @{term Invoke_special} only: mark when invoking a
constructor on a partly initialized class. app will check that we
call the right constructor.
*}
constdefs
eff_bool :: "instr \<Rightarrow> jvm_prog \<Rightarrow> p_count \<Rightarrow> state_bool \<Rightarrow> state_bool"
"eff_bool i G pc == \<lambda>((ST,LT),z). (eff'(i,G,pc,(ST,LT)),
if \<exists>C p D. i = Invoke_special C init p \<and> ST!length p = PartInit D then True else z)"
text {*
For exception handling:
*}
consts
match_any :: "jvm_prog \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> cname list"
primrec
"match_any G pc [] = []"
"match_any G pc (e#es) = (let (start_pc, end_pc, handler_pc, catch_type) = e;
es' = match_any G pc es
in
if start_pc <= pc \<and> pc < end_pc then catch_type#es' else es')"
consts
match :: "jvm_prog \<Rightarrow> cname \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> cname list"
primrec
"match G X pc [] = []"
"match G X pc (e#es) =
(if match_exception_entry G X pc e then [X] else match G X pc es)"
lemma match_some_entry:
"match G X pc et = (if \<exists>e \<in> set et. match_exception_entry G X pc e then [X] else [])"
by (induct et) auto
consts
xcpt_names :: "instr × jvm_prog × p_count × exception_table \<Rightarrow> cname list"
recdef xcpt_names "{}"
"xcpt_names (Getfield F C, G, pc, et) = match G (Xcpt NullPointer) pc et"
"xcpt_names (Putfield F C, G, pc, et) = match G (Xcpt NullPointer) pc et"
"xcpt_names (New C, G, pc, et) = match G (Xcpt OutOfMemory) pc et"
"xcpt_names (Checkcast C, G, pc, et) = match G (Xcpt ClassCast) pc et"
"xcpt_names (Throw, G, pc, et) = match_any G pc et"
"xcpt_names (Invoke C m p, G, pc, et) = match_any G pc et"
"xcpt_names (Invoke_special C m p, G, pc, et) = match_any G pc et"
"xcpt_names (i, G, pc, et) = []"
constdefs
xcpt_eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> p_count \<Rightarrow> state_bool option \<Rightarrow> exception_table \<Rightarrow> succ_type"
"xcpt_eff i G pc s et ==
map (\<lambda>C. (the (match_exception_table G C pc et), case s of
None \<Rightarrow> None | Some s' \<Rightarrow> Some (([Init (Class C)], snd (fst s')),snd s') ))
(xcpt_names (i,G,pc,et))"
norm_eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> p_count \<Rightarrow> state_bool option \<Rightarrow> state_bool option"
"norm_eff i G pc == option_map (eff_bool i G pc)"
text {*
Putting it all together:
*}
constdefs
eff :: "instr \<Rightarrow> jvm_prog \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> state_bool option \<Rightarrow> succ_type"
"eff i G pc et s == (map (\<lambda>pc'. (pc',norm_eff i G pc s)) (succs i pc)) @ (xcpt_eff i G pc s et)"
text {*
Some small helpers for direct executability
*}
constdefs
isPrimT :: "ty \<Rightarrow> bool"
"isPrimT T == case T of PrimT T' \<Rightarrow> True | RefT T' \<Rightarrow> False"
isRefT :: "ty \<Rightarrow> bool"
"isRefT T == case T of PrimT T' \<Rightarrow> False | RefT T' \<Rightarrow> True"
lemma isPrimT [simp]:
"isPrimT T = (\<exists>T'. T = PrimT T')" by (simp add: isPrimT_def split: ty.splits)
lemma isRefT [simp]:
"isRefT T = (\<exists>T'. T = RefT T')" by (simp add: isRefT_def split: ty.splits)
lemma "list_all2 P a b \<Longrightarrow> \<forall>(x,y) \<in> set (zip a b). P x y"
by (simp add: list_all2_def)
text "Conditions under which eff is applicable:"
consts
app' :: "instr × jvm_prog × cname × p_count × nat × ty × state_type \<Rightarrow> bool"
recdef app' "{}"
"app' (Load idx, G, C', pc, maxs, rT, s)
= (idx < length (snd s) \<and> (snd s) ! idx \<noteq> Err \<and> length (fst s) < maxs)"
"app' (Store idx, G, C', pc, maxs, rT, (ts#ST, LT))
= (idx < length LT)"
"app' (LitPush v, G, C', pc, maxs, rT, s)
= (length (fst s) < maxs \<and> typeof (\<lambda>t. None) v \<noteq> None)"
"app' (Getfield F C, G, C', pc, maxs, rT, (oT#ST, LT))
= (is_class G C \<and> field (G,C) F \<noteq> None \<and> fst (the (field (G,C) F)) = C \<and>
G \<turnstile> oT \<preceq>i Init (Class C))"
"app' (Putfield F C, G, C', pc, maxs, rT, (vT#oT#ST, LT))
= (is_class G C \<and> field (G,C) F \<noteq> None \<and> fst (the (field (G,C) F)) = C \<and>
G \<turnstile> oT \<preceq>i Init (Class C) \<and> G \<turnstile> vT \<preceq>i Init (snd (the (field (G,C) F))))"
"app' (New C, G, C', pc, maxs, rT, s)
= (is_class G C \<and> length (fst s) < maxs \<and> UnInit C pc \<notin> set (fst s))"
"app' (Checkcast C, G, C', pc, maxs, rT, (Init (RefT rt)#ST,LT))
= is_class G C"
"app' (Pop, G, C', pc, maxs, rT, (ts#ST,LT)) = True"
"app' (Dup, G, C', pc, maxs, rT, (ts#ST,LT)) = (1+length ST < maxs)"
"app' (Dup_x1, G, C', pc, maxs, rT, (ts1#ts2#ST,LT)) = (2+length ST < maxs)"
"app' (Dup_x2, G, C', pc, maxs, rT, (ts1#ts2#ts3#ST,LT)) = (3+length ST < maxs)"
"app' (Swap, G, C', pc, maxs, rT, (ts1#ts2#ST,LT)) = True"
"app' (IAdd, G, C', pc, maxs, rT, (t1#t2#ST,LT))
= (t1 = Init (PrimT Integer) \<and> t1 = t2)"
"app' (Ifcmpeq b, G, C', pc, maxs, rT, (Init ts#Init ts'#ST,LT))
= (0 \<le> int pc + b \<and> (isPrimT ts \<longrightarrow> ts' = ts) \<and> (isRefT ts \<longrightarrow> isRefT ts'))"
"app' (Goto b, G, C', pc, maxs, rT, s) = (0 \<le> int pc + b)"
"app' (Return, G, C', pc, maxs, rT, (T#ST,LT)) = (G \<turnstile> T \<preceq>i Init rT)"
"app' (Throw, G, C', pc, maxs, rT, (Init T#ST,LT)) = isRefT T"
"app' (Invoke C mn fpTs, G, C', pc, maxs, rT, s) =
(length fpTs < length (fst s) \<and> mn \<noteq> init \<and>
(let apTs = rev (take (length fpTs) (fst s));
X = hd (drop (length fpTs) (fst s))
in is_class G C \<and>
list_all2 (\<lambda>aT fT. G \<turnstile> aT \<preceq>i (Init fT)) apTs fpTs \<and>
G \<turnstile> X \<preceq>i Init (Class C)) \<and>
method (G,C) (mn,fpTs) \<noteq> None)"
"app' (Invoke_special C mn fpTs, G, C', pc, maxs, rT, s) =
(length fpTs < length (fst s) \<and> mn = init \<and>
(let apTs = rev (take (length fpTs) (fst s));
X = (fst s)!length fpTs
in is_class G C \<and>
list_all2 (\<lambda>aT fT. G \<turnstile> aT \<preceq>i (Init fT)) apTs fpTs \<and>
(\<exists>rT' b. method (G,C) (mn,fpTs) = Some (C,rT',b)) \<and>
((\<exists>pc. X = UnInit C pc) \<or> (X = PartInit C' \<and> G \<turnstile> C' \<prec>C1 C))))"
-- "@{text C'} is the current class, the constructor must be called on the"
-- "superclass (if partly initialized) or on the exact class that is"
-- "to be constructed (if not yet initialized at all)."
-- "In JCVM @{text Invoke_special} may also call another constructor of the same"
-- {* class (@{text "C = C' \<or> C = super C'"}) *}
"app' (i,G,pc,maxs,rT,s) = False"
constdefs
xcpt_app :: "instr \<Rightarrow> jvm_prog \<Rightarrow> nat \<Rightarrow> exception_table \<Rightarrow> bool"
"xcpt_app i G pc et \<equiv> \<forall>C\<in>set(xcpt_names (i,G,pc,et)). is_class G C"
constdefs
app :: "instr \<Rightarrow> jvm_prog \<Rightarrow> cname \<Rightarrow> p_count \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> bool \<Rightarrow>
exception_table \<Rightarrow> state_bool option \<Rightarrow> bool"
"app i G C' pc maxs rT ini et s \<equiv> case s of None \<Rightarrow> True | Some t \<Rightarrow>
let (s,z) = t in
xcpt_app i G pc et \<and>
app' (i,G,C',pc,maxs,rT,s) \<and>
(ini \<and> i = Return \<longrightarrow> z) \<and>
(\<forall>C m p. i = Invoke_special C m p \<and> (fst s)!length p = PartInit C' \<longrightarrow> ¬z)"
lemma match_any_match_table:
"C \<in> set (match_any G pc et) \<Longrightarrow> match_exception_table G C pc et \<noteq> None"
apply (induct et)
apply simp
apply simp
apply clarify
apply (simp split: split_if_asm)
apply (auto simp add: match_exception_entry_def)
done
lemma match_X_match_table:
"C \<in> set (match G X pc et) \<Longrightarrow> match_exception_table G C pc et \<noteq> None"
apply (induct et)
apply simp
apply (simp split: split_if_asm)
done
lemma xcpt_names_in_et:
"C \<in> set (xcpt_names (i,G,pc,et)) \<Longrightarrow>
\<exists>e \<in> set et. the (match_exception_table G C pc et) = fst (snd (snd e))"
apply (cases i)
apply (auto dest!: match_any_match_table match_X_match_table
dest: match_exception_table_in_et)
done
lemma 1: "2 < length a \<Longrightarrow> (\<exists>l l' l'' ls. a = l#l'#l''#ls)"
proof (cases a)
fix x xs assume "a = x#xs" "2 < length a"
thus ?thesis by - (cases xs, simp, cases "tl xs", auto)
qed auto
lemma 2: "¬(2 < length a) \<Longrightarrow> a = [] \<or> (\<exists> l. a = [l]) \<or> (\<exists> l l'. a = [l,l'])"
proof -
assume "¬(2 < length a)"
hence "length a < (Suc (Suc (Suc 0)))" by simp
hence * : "length a = 0 \<or> length a = Suc 0 \<or> length a = Suc (Suc 0)"
by (auto simp add: less_Suc_eq)
{ fix x assume "length x = Suc 0"
hence "\<exists> l. x = [l]" by - (cases x, auto)
} note 0 = this
have "length a = Suc (Suc 0) \<Longrightarrow> \<exists>l l'. a = [l,l']" by (cases a, auto dest: 0)
with * show ?thesis by (auto dest: 0)
qed
lemmas [simp] = app_def xcpt_app_def
text {*
\medskip
simp rules for @{term app}
*}
lemma appNone[simp]: "app i G C' maxs rT pc ini et None = True" by simp
lemma appLoad[simp]:
"app (Load idx) G C' pc maxs rT ini et (Some s) =
(\<exists>ST LT z. s = ((ST,LT),z) \<and> idx < length LT \<and> LT!idx \<noteq> Err \<and> length ST < maxs)"
by (cases s, auto)
lemma appStore[simp]:
"app (Store idx) G C' pc maxs rT ini et (Some s) = (\<exists>ts ST LT z. s = ((ts#ST,LT),z) \<and> idx < length LT)"
by (cases s, cases "2 < length (fst (fst s))", auto dest: 1 2)
lemma appLitPush[simp]:
"app (LitPush v) G C' pc maxs rT ini et (Some s) = (\<exists>ST LT z. s = ((ST,LT),z) \<and> length ST < maxs \<and> typeof (\<lambda>v. None) v \<noteq> None)"
by (cases s, auto)
lemma appGetField[simp]:
"app (Getfield F C) G C' pc maxs rT ini et (Some s) =
(\<exists> oT vT ST LT z. s = ((oT#ST, LT),z) \<and> is_class G C \<and>
field (G,C) F = Some (C,vT) \<and> G \<turnstile> oT \<preceq>i (Init (Class C)) \<and>
(\<forall>x \<in> set (match G (Xcpt NullPointer) pc et). is_class G x))"
by (cases s, cases "2 <length (fst (fst s))", auto dest!: 1 2)
lemma appPutField[simp]:
"app (Putfield F C) G C' pc maxs rT ini et (Some s) =
(\<exists> vT vT' oT ST LT z. s = ((vT#oT#ST, LT),z) \<and> is_class G C \<and>
field (G,C) F = Some (C, vT') \<and>
G \<turnstile> oT \<preceq>i Init (Class C) \<and> G \<turnstile> vT \<preceq>i Init vT' \<and>
(\<forall>x \<in> set (match G (Xcpt NullPointer) pc et). is_class G x))"
by (cases s, cases "2 <length (fst (fst s))", auto dest!: 1 2)
lemma appNew[simp]:
"app (New C) G C' pc maxs rT ini et (Some s) =
(\<exists>ST LT z. s=((ST,LT),z) \<and>
is_class G C \<and> length ST < maxs \<and>
UnInit C pc \<notin> set ST \<and>
(\<forall>x \<in> set (match G (Xcpt OutOfMemory) pc et). is_class G x))"
by (cases s, auto)
lemma appCheckcast[simp]:
"app (Checkcast C) G C' pc maxs rT ini et (Some s) =
(\<exists>rT ST LT z. s = ((Init (RefT rT)#ST,LT),z) \<and> is_class G C \<and>
(\<forall>x \<in> set (match G (Xcpt ClassCast) pc et). is_class G x))"
proof -
{ fix t ST LT z assume "s = ((Init t#ST,LT),z)"
hence ?thesis by (cases t, auto)
} thus ?thesis
by (cases s, cases "fst s", cases "fst (fst s)", simp,
cases "hd (fst (fst s))", auto)
qed
lemma appPop[simp]:
"app Pop G C' pc maxs rT ini et (Some s) = (\<exists>ts ST LT z. s = ((ts#ST,LT),z))"
by (cases s, cases "2 <length (fst (fst s))", auto dest: 1 2)
lemma appDup[simp]:
"app Dup G C' pc maxs rT ini et (Some s) =
(\<exists>ts ST LT z. s = ((ts#ST,LT),z) \<and> 1+length ST < maxs)"
by (cases s, cases "2 < length (fst (fst s))", auto dest: 1 2)
lemma appDup_x1[simp]:
"app Dup_x1 G C' pc maxs rT ini et (Some s) =
(\<exists>ts1 ts2 ST LT z. s = ((ts1#ts2#ST,LT),z) \<and> 2+length ST < maxs)"
by (cases s, cases "2 < length (fst (fst s))", auto dest: 1 2)
lemma appDup_x2[simp]:
"app Dup_x2 G C' pc maxs rT ini et (Some s) =
(\<exists>ts1 ts2 ts3 ST LT z. s = ((ts1#ts2#ts3#ST,LT),z) \<and> 3+length ST < maxs)"
by (cases s, cases "2 < length (fst (fst s))", auto dest: 1 2)
lemma appSwap[simp]:
"app Swap G C' pc maxs rT ini et (Some s) =
(\<exists>ts1 ts2 ST LT z. s = ((ts1#ts2#ST,LT),z))"
by (cases s, cases "2 < length (fst (fst s))", auto dest: 1 2)
lemma appIAdd[simp]:
"app IAdd G C' pc maxs rT ini et (Some s) =
(\<exists> ST LT z. s = ((Init (PrimT Integer)#Init (PrimT Integer)#ST,LT),z))"
by (cases s, cases "2 < length (fst (fst s))", auto dest: 1 2)
lemma appIfcmpeq[simp]:
"app (Ifcmpeq b) G C' pc maxs rT ini et (Some s) =
(\<exists>ts1 ts2 ST LT z. s = ((Init ts1#Init ts2#ST,LT),z) \<and> 0 \<le> b + int pc \<and>
((\<exists>p. ts1 = PrimT p \<and> ts2 = PrimT p) \<or>
(\<exists>r r'. ts1 = RefT r \<and> ts2 = RefT r')))"
apply (cases s)
apply (cases "fst s")
(* fixme *)
apply (case_tac aa)
apply simp
apply (case_tac list)
apply (case_tac ab, simp, simp, simp)
apply (case_tac ab)
apply auto
apply (case_tac ac)
defer
apply simp
apply simp
apply (case_tac ty)
apply auto
done
lemma appReturn[simp]:
"app Return G C' pc maxs rT ini et (Some s) =
(\<exists>T ST LT z. s = ((T#ST,LT),z) \<and> (G \<turnstile> T \<preceq>i Init rT) \<and> (ini \<longrightarrow> z))"
by (cases s, cases "2 <length (fst (fst s))", auto dest: 1 2)
lemma appGoto[simp]:
"app (Goto b) G C' pc maxs rT ini et (Some s) = (0 \<le> int pc + b)" by simp
lemma appThrow[simp]:
"app Throw G C' pc maxs rT ini et (Some s) =
(\<exists>ST LT z r. s=((Init (RefT r)#ST,LT),z) \<and> (\<forall>C \<in> set (match_any G pc et). is_class G C))"
apply (cases s)
apply (cases "fst s")
(* fixme *)
apply (case_tac aa)
apply simp
apply (case_tac ab)
apply auto
done
lemma appInvoke[simp]:
"app (Invoke C mn fpTs) G C' pc maxs rT ini et (Some s) =
(\<exists>apTs X ST LT mD' rT' b' z.
s = (((rev apTs) @ (X # ST), LT), z) \<and> mn \<noteq> init \<and>
length apTs = length fpTs \<and> is_class G C \<and>
(\<forall>(aT,fT)\<in>set(zip apTs fpTs). G \<turnstile> aT \<preceq>i (Init fT)) \<and>
method (G,C) (mn,fpTs) = Some (mD', rT', b') \<and> (G \<turnstile> X \<preceq>i Init (Class C)) \<and>
(\<forall>C \<in> set (match_any G pc et). is_class G C))"
(is "?app s = ?P s")
proof -
note list_all2_def[simp]
{ fix a b z
have "?app ((a,b),z) \<Longrightarrow> ?P ((a,b),z)"
proof -
assume app: "?app ((a,b),z)"
hence "a = (rev (rev (take (length fpTs) a))) @ (drop (length fpTs) a) \<and>
length fpTs < length a" (is "?a \<and> ?l")
by (auto simp add: app_def)
hence "?a \<and> 0 < length (drop (length fpTs) a)" (is "?a \<and> ?l")
by auto
hence "?a \<and> ?l \<and> length (rev (take (length fpTs) a)) = length fpTs"
by (auto simp add: min_def)
then obtain apTs ST where
"a = rev apTs @ ST \<and> length apTs = length fpTs \<and> 0 < length ST"
by blast
hence "a = rev apTs @ ST \<and> length apTs = length fpTs \<and> ST \<noteq> []"
by blast
then obtain X ST' where
"a = rev apTs @ X # ST'" "length apTs = length fpTs"
by (simp add: neq_Nil_conv) blast
with app show ?thesis by clarsimp blast
qed }
moreover obtain a b z where "s = ((a,b),z)" by (cases s, cases "fst s", simp)
ultimately have "?app s \<Longrightarrow> ?P s" by (simp only:)
moreover
have "?P s \<Longrightarrow> ?app s" by (clarsimp simp add: min_def)
ultimately
show ?thesis by (rule iffI)
qed
lemma appInvoke_special[simp]:
"app (Invoke_special C mn fpTs) G C' pc maxs rT ini et (Some s) =
(\<exists>apTs X ST LT rT' b' z.
s = (((rev apTs) @ X # ST, LT), z) \<and> mn = init \<and>
length apTs = length fpTs \<and> is_class G C \<and>
(\<forall>(aT,fT)\<in>set(zip apTs fpTs). G \<turnstile> aT \<preceq>i (Init fT)) \<and>
method (G,C) (mn,fpTs) = Some (C, rT', b') \<and>
((\<exists>pc. X = UnInit C pc) \<or> (X = PartInit C' \<and> G \<turnstile> C' \<prec>C1 C \<and> ¬z)) \<and>
(\<forall>C \<in> set (match_any G pc et). is_class G C))"
(is "?app s = ?P s")
proof -
note list_all2_def [simp]
{ fix a b z
have "?app ((a,b),z) \<Longrightarrow> ?P ((a,b),z)"
proof -
assume app: "?app ((a,b),z)"
hence "a = (rev (rev (take (length fpTs) a))) @ (drop (length fpTs) a) \<and>
length fpTs < length a" (is "?a \<and> ?l")
by (auto simp add: app_def)
hence "?a \<and> 0 < length (drop (length fpTs) a)" (is "?a \<and> ?l")
by auto
hence "?a \<and> ?l \<and> length (rev (take (length fpTs) a)) = length fpTs"
by (auto simp add: min_def)
then obtain apTs ST where
"a = rev apTs @ ST \<and> length apTs = length fpTs \<and> 0 < length ST"
by blast
hence "a = rev apTs @ ST \<and> length apTs = length fpTs \<and> ST \<noteq> []"
by blast
then obtain X ST' where
"a = rev apTs @ X # ST'" "length apTs = length fpTs"
by (simp add: neq_Nil_conv) blast
with app show ?thesis by (simp add: nth_append) blast
qed }
moreover obtain a b z where "s = ((a,b),z)" by (cases s, cases "fst s", simp)
ultimately have "?app s \<Longrightarrow> ?P s" by (simp only:)
moreover
have "?P s \<Longrightarrow> ?app s" by (clarsimp simp add: nth_append min_def) blast
ultimately
show ?thesis by (rule iffI)
qed
(* this is here, because it needs OK and replace, and is used
in StepMono and Correct *)
lemma replace_map_OK:
"replace (OK x) (OK y) (map OK l) = map OK (replace x y l)"
proof -
have "inj OK" by (blast intro: datatype_injI)
thus ?thesis by (rule replace_map)
qed
lemma effNone:
"(pc', s') \<in> set (eff i G pc et None) \<Longrightarrow> s' = None"
by (auto simp add: eff_def xcpt_eff_def norm_eff_def)
text {* some more helpers to make the specification directly executable: *}
declare list_all2_Nil [code]
declare list_all2_Cons [code]
lemma xcpt_app_lemma [code]:
"xcpt_app i G pc et = list_all (is_class G) (xcpt_names (i, G, pc, et))"
by (simp add: list_all_conv)
lemmas [simp del] = app_def xcpt_app_def
end
lemma match_some_entry:
match G X pc et = (if Bex (set et) (match_exception_entry G X pc) then [X] else [])
lemma isPrimT:
isPrimT T = (EX T'. T = PrimT T')
lemma isRefT:
isRefT T = (EX T'. T = RefT T')
lemma
list_all2 P a b ==> Ball (set (zip a b)) (split P)
lemma match_any_match_table:
C : set (match_any G pc et) ==> match_exception_table G C pc et ~= None
lemma match_X_match_table:
C : set (match G X pc et) ==> match_exception_table G C pc et ~= None
lemma xcpt_names_in_et:
C : set (xcpt_names (i, G, pc, et)) ==> EX e:set et. the (match_exception_table G C pc et) = fst (snd (snd e))
lemma 1:
2 < length a ==> EX l l' l'' ls. a = l # l' # l'' # ls
lemma 2:
¬ 2 < length a ==> a = [] | (EX l. a = [l]) | (EX l l'. a = [l, l'])
lemmas
app i G C' pc maxs rT ini et s ==
case s of None => True
| Some t =>
let (s, z) = t
in xcpt_app i G pc et &
app' (i, G, C', pc, maxs, rT, s) &
(ini & i = Return --> z) &
(ALL C m p.
i = Invoke_special C m p & fst s ! length p = PartInit C' --> ¬ z)
xcpt_app i G pc et == Ball (set (xcpt_names (i, G, pc, et))) (is_class G)
lemma appNone:
app i G C' maxs rT pc ini et None = True
lemma appLoad:
app (Load idx) G C' pc maxs rT ini et (Some s) =
(EX ST LT z.
s = ((ST, LT), z) & idx < length LT & LT ! idx ~= Err & length ST < maxs)
lemma appStore:
app (Store idx) G C' pc maxs rT ini et (Some s) = (EX ts ST LT z. s = ((ts # ST, LT), z) & idx < length LT)
lemma appLitPush:
app (LitPush v) G C' pc maxs rT ini et (Some s) = (EX ST LT z. s = ((ST, LT), z) & length ST < maxs & typeof (%v. None) v ~= None)
lemma appGetField:
app (Getfield F C) G C' pc maxs rT ini et (Some s) =
(EX oT vT ST LT z.
s = ((oT # ST, LT), z) &
is_class G C &
field (G, C) F = Some (C, vT) &
G |- oT <=i Init (Class C) &
Ball (set (match G (Xcpt NullPointer) pc et)) (is_class G))
lemma appPutField:
app (Putfield F C) G C' pc maxs rT ini et (Some s) =
(EX vT vT' oT ST LT z.
s = ((vT # oT # ST, LT), z) &
is_class G C &
field (G, C) F = Some (C, vT') &
G |- oT <=i Init (Class C) &
G |- vT <=i Init vT' &
Ball (set (match G (Xcpt NullPointer) pc et)) (is_class G))
lemma appNew:
app (New C) G C' pc maxs rT ini et (Some s) =
(EX ST LT z.
s = ((ST, LT), z) &
is_class G C &
length ST < maxs &
UnInit C pc ~: set ST &
Ball (set (match G (Xcpt OutOfMemory) pc et)) (is_class G))
lemma appCheckcast:
app (Checkcast C) G C' pc maxs rT ini et (Some s) =
(EX rT ST LT z.
s = ((Init (RefT rT) # ST, LT), z) &
is_class G C & Ball (set (match G (Xcpt ClassCast) pc et)) (is_class G))
lemma appPop:
app Pop G C' pc maxs rT ini et (Some s) = (EX ts ST LT z. s = ((ts # ST, LT), z))
lemma appDup:
app Dup G C' pc maxs rT ini et (Some s) = (EX ts ST LT z. s = ((ts # ST, LT), z) & 1 + length ST < maxs)
lemma appDup_x1:
app Dup_x1 G C' pc maxs rT ini et (Some s) = (EX ts1 ts2 ST LT z. s = ((ts1 # ts2 # ST, LT), z) & 2 + length ST < maxs)
lemma appDup_x2:
app Dup_x2 G C' pc maxs rT ini et (Some s) =
(EX ts1 ts2 ts3 ST LT z.
s = ((ts1 # ts2 # ts3 # ST, LT), z) & 3 + length ST < maxs)
lemma appSwap:
app Swap G C' pc maxs rT ini et (Some s) = (EX ts1 ts2 ST LT z. s = ((ts1 # ts2 # ST, LT), z))
lemma appIAdd:
app IAdd G C' pc maxs rT ini et (Some s) = (EX ST LT z. s = ((Init (PrimT Integer) # Init (PrimT Integer) # ST, LT), z))
lemma appIfcmpeq:
app (Ifcmpeq b) G C' pc maxs rT ini et (Some s) =
(EX ts1 ts2 ST LT z.
s = ((Init ts1 # Init ts2 # ST, LT), z) &
0 <= b + int pc &
((EX p. ts1 = PrimT p & ts2 = PrimT p) |
(EX r r'. ts1 = RefT r & ts2 = RefT r')))
lemma appReturn:
app Return G C' pc maxs rT ini et (Some s) = (EX T ST LT z. s = ((T # ST, LT), z) & G |- T <=i Init rT & (ini --> z))
lemma appGoto:
app (Goto b) G C' pc maxs rT ini et (Some s) = (0 <= int pc + b)
lemma appThrow:
app Throw G C' pc maxs rT ini et (Some s) =
(EX ST LT z r.
s = ((Init (RefT r) # ST, LT), z) &
Ball (set (match_any G pc et)) (is_class G))
lemma appInvoke:
app (Invoke C mn fpTs) G C' pc maxs rT ini et (Some s) =
(EX apTs X ST LT mD' rT' b' z.
s = ((rev apTs @ X # ST, LT), z) &
mn ~= init &
length apTs = length fpTs &
is_class G C &
(ALL (aT, fT):set (zip apTs fpTs). G |- aT <=i Init fT) &
method (G, C) (mn, fpTs) = Some (mD', rT', b') &
G |- X <=i Init (Class C) & Ball (set (match_any G pc et)) (is_class G))
lemma appInvoke_special:
app (Invoke_special C mn fpTs) G C' pc maxs rT ini et (Some s) =
(EX apTs X ST LT rT' b' z.
s = ((rev apTs @ X # ST, LT), z) &
mn = init &
length apTs = length fpTs &
is_class G C &
(ALL (aT, fT):set (zip apTs fpTs). G |- aT <=i Init fT) &
method (G, C) (mn, fpTs) = Some (C, rT', b') &
((EX pc. X = UnInit C pc) | X = PartInit C' & G |- C' <=C1 C & ¬ z) &
Ball (set (match_any G pc et)) (is_class G))
lemma replace_map_OK:
replace (OK x) (OK y) (map OK l) = map OK (replace x y l)
lemma effNone:
(pc', s') : set (eff i G pc et None) ==> s' = None
lemma xcpt_app_lemma:
xcpt_app i G pc et = list_all (is_class G) (xcpt_names (i, G, pc, et))
lemmas
app i G C' pc maxs rT ini et s ==
case s of None => True
| Some t =>
let (s, z) = t
in xcpt_app i G pc et &
app' (i, G, C', pc, maxs, rT, s) &
(ini & i = Return --> z) &
(ALL C m p.
i = Invoke_special C m p & fst s ! length p = PartInit C' --> ¬ z)
xcpt_app i G pc et == Ball (set (xcpt_names (i, G, pc, et))) (is_class G)