| author | blanchet |
| Fri, 26 Feb 2010 18:38:23 +0100 | |
| changeset 35388 | 42d39948cace |
| parent 35355 | 613e133966ea |
| child 35417 | 47ee18b6ae32 |
| permissions | -rwxr-xr-x |
| 13631 | 1 |
(* Title: HOL/MicroJava/JVM/JVMDefensive.thy |
2 |
Author: Gerwin Klein |
|
3 |
*) |
|
4 |
||
5 |
header {* \isaheader{A Defensive JVM} *}
|
|
6 |
||
| 16417 | 7 |
theory JVMDefensive imports JVMExec begin |
| 13631 | 8 |
|
9 |
text {*
|
|
10 |
Extend the state space by one element indicating a type error (or |
|
11 |
other abnormal termination) *} |
|
12 |
datatype 'a type_error = TypeError | Normal 'a |
|
13 |
||
14 |
||
| 35102 | 15 |
abbreviation |
16 |
fifth :: "'a \<times> 'b \<times> 'c \<times> 'd \<times> 'e \<times> 'f \<Rightarrow> 'e" |
|
17 |
where "fifth x == fst(snd(snd(snd(snd x))))" |
|
| 13631 | 18 |
|
19 |
||
20 |
consts isAddr :: "val \<Rightarrow> bool" |
|
21 |
recdef isAddr "{}"
|
|
22 |
"isAddr (Addr loc) = True" |
|
23 |
"isAddr v = False" |
|
24 |
||
25 |
consts isIntg :: "val \<Rightarrow> bool" |
|
26 |
recdef isIntg "{}"
|
|
27 |
"isIntg (Intg i) = True" |
|
28 |
"isIntg v = False" |
|
29 |
||
30 |
constdefs |
|
31 |
isRef :: "val \<Rightarrow> bool" |
|
32 |
"isRef v \<equiv> v = Null \<or> isAddr v" |
|
33 |
||
34 |
||
35 |
consts |
|
36 |
check_instr :: "[instr, jvm_prog, aheap, opstack, locvars, |
|
| 13822 | 37 |
cname, sig, p_count, nat, frame list] \<Rightarrow> bool" |
| 13631 | 38 |
primrec |
| 13822 | 39 |
"check_instr (Load idx) G hp stk vars C sig pc mxs frs = |
40 |
(idx < length vars \<and> size stk < mxs)" |
|
| 13631 | 41 |
|
| 13822 | 42 |
"check_instr (Store idx) G hp stk vars Cl sig pc mxs frs = |
| 13631 | 43 |
(0 < length stk \<and> idx < length vars)" |
44 |
||
| 13822 | 45 |
"check_instr (LitPush v) G hp stk vars Cl sig pc mxs frs = |
46 |
(\<not>isAddr v \<and> size stk < mxs)" |
|
| 13631 | 47 |
|
| 13822 | 48 |
"check_instr (New C) G hp stk vars Cl sig pc mxs frs = |
49 |
(is_class G C \<and> size stk < mxs)" |
|
| 13631 | 50 |
|
| 13822 | 51 |
"check_instr (Getfield F C) G hp stk vars Cl sig pc mxs frs = |
| 13631 | 52 |
(0 < length stk \<and> is_class G C \<and> field (G,C) F \<noteq> None \<and> |
53 |
(let (C', T) = the (field (G,C) F); ref = hd stk in |
|
54 |
C' = C \<and> isRef ref \<and> (ref \<noteq> Null \<longrightarrow> |
|
55 |
hp (the_Addr ref) \<noteq> None \<and> |
|
56 |
(let (D,vs) = the (hp (the_Addr ref)) in |
|
57 |
G \<turnstile> D \<preceq>C C \<and> vs (F,C) \<noteq> None \<and> G,hp \<turnstile> the (vs (F,C)) ::\<preceq> T))))" |
|
58 |
||
| 13822 | 59 |
"check_instr (Putfield F C) G hp stk vars Cl sig pc mxs frs = |
| 13631 | 60 |
(1 < length stk \<and> is_class G C \<and> field (G,C) F \<noteq> None \<and> |
61 |
(let (C', T) = the (field (G,C) F); v = hd stk; ref = hd (tl stk) in |
|
62 |
C' = C \<and> isRef ref \<and> (ref \<noteq> Null \<longrightarrow> |
|
63 |
hp (the_Addr ref) \<noteq> None \<and> |
|
64 |
(let (D,vs) = the (hp (the_Addr ref)) in |
|
65 |
G \<turnstile> D \<preceq>C C \<and> G,hp \<turnstile> v ::\<preceq> T))))" |
|
66 |
||
| 13822 | 67 |
"check_instr (Checkcast C) G hp stk vars Cl sig pc mxs frs = |
| 13631 | 68 |
(0 < length stk \<and> is_class G C \<and> isRef (hd stk))" |
69 |
||
| 13822 | 70 |
"check_instr (Invoke C mn ps) G hp stk vars Cl sig pc mxs frs = |
| 13631 | 71 |
(length ps < length stk \<and> |
72 |
(let n = length ps; v = stk!n in |
|
73 |
isRef v \<and> (v \<noteq> Null \<longrightarrow> |
|
74 |
hp (the_Addr v) \<noteq> None \<and> |
|
75 |
method (G,cname_of hp v) (mn,ps) \<noteq> None \<and> |
|
76 |
list_all2 (\<lambda>v T. G,hp \<turnstile> v ::\<preceq> T) (rev (take n stk)) ps)))" |
|
77 |
||
| 13822 | 78 |
"check_instr Return G hp stk0 vars Cl sig0 pc mxs frs = |
| 13631 | 79 |
(0 < length stk0 \<and> (0 < length frs \<longrightarrow> |
80 |
method (G,Cl) sig0 \<noteq> None \<and> |
|
81 |
(let v = hd stk0; (C, rT, body) = the (method (G,Cl) sig0) in |
|
82 |
Cl = C \<and> G,hp \<turnstile> v ::\<preceq> rT)))" |
|
83 |
||
| 13822 | 84 |
"check_instr Pop G hp stk vars Cl sig pc mxs frs = |
| 13631 | 85 |
(0 < length stk)" |
86 |
||
| 13822 | 87 |
"check_instr Dup G hp stk vars Cl sig pc mxs frs = |
88 |
(0 < length stk \<and> size stk < mxs)" |
|
| 13631 | 89 |
|
| 13822 | 90 |
"check_instr Dup_x1 G hp stk vars Cl sig pc mxs frs = |
91 |
(1 < length stk \<and> size stk < mxs)" |
|
92 |
||
93 |
"check_instr Dup_x2 G hp stk vars Cl sig pc mxs frs = |
|
94 |
(2 < length stk \<and> size stk < mxs)" |
|
95 |
||
96 |
"check_instr Swap G hp stk vars Cl sig pc mxs frs = |
|
| 13631 | 97 |
(1 < length stk)" |
98 |
||
| 13822 | 99 |
"check_instr IAdd G hp stk vars Cl sig pc mxs frs = |
| 13631 | 100 |
(1 < length stk \<and> isIntg (hd stk) \<and> isIntg (hd (tl stk)))" |
101 |
||
| 13822 | 102 |
"check_instr (Ifcmpeq b) G hp stk vars Cl sig pc mxs frs = |
| 13677 | 103 |
(1 < length stk \<and> 0 \<le> int pc+b)" |
| 13631 | 104 |
|
| 13822 | 105 |
"check_instr (Goto b) G hp stk vars Cl sig pc mxs frs = |
| 13677 | 106 |
(0 \<le> int pc+b)" |
| 13631 | 107 |
|
| 13822 | 108 |
"check_instr Throw G hp stk vars Cl sig pc mxs frs = |
| 13631 | 109 |
(0 < length stk \<and> isRef (hd stk))" |
110 |
||
111 |
constdefs |
|
112 |
check :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> bool" |
|
113 |
"check G s \<equiv> let (xcpt, hp, frs) = s in |
|
114 |
(case frs of [] \<Rightarrow> True | (stk,loc,C,sig,pc)#frs' \<Rightarrow> |
|
| 13822 | 115 |
(let (C',rt,mxs,mxl,ins,et) = the (method (G,C) sig); i = ins!pc in |
116 |
pc < size ins \<and> |
|
117 |
check_instr i G hp stk loc C sig pc mxs frs'))" |
|
| 13631 | 118 |
|
119 |
||
120 |
exec_d :: "jvm_prog \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state option type_error" |
|
121 |
"exec_d G s \<equiv> case s of |
|
122 |
TypeError \<Rightarrow> TypeError |
|
123 |
| Normal s' \<Rightarrow> if check G s' then Normal (exec (G, s')) else TypeError" |
|
124 |
||
125 |
||
|
35355
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents:
35102
diff
changeset
|
126 |
constdefs |
|
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents:
35102
diff
changeset
|
127 |
exec_all_d :: "jvm_prog \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state type_error \<Rightarrow> bool" |
| 13631 | 128 |
("_ |- _ -jvmd-> _" [61,61,61]60)
|
|
35355
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents:
35102
diff
changeset
|
129 |
"G |- s -jvmd-> t \<equiv> |
|
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents:
35102
diff
changeset
|
130 |
(s,t) \<in> ({(s,t). exec_d G s = TypeError \<and> t = TypeError} \<union>
|
|
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents:
35102
diff
changeset
|
131 |
{(s,t). \<exists>t'. exec_d G s = Normal (Some t') \<and> t = Normal t'})\<^sup>*"
|
| 13631 | 132 |
|
|
35355
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents:
35102
diff
changeset
|
133 |
notation (xsymbols) |
|
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents:
35102
diff
changeset
|
134 |
exec_all_d ("_ \<turnstile> _ -jvmd\<rightarrow> _" [61,61,61]60)
|
| 13631 | 135 |
|
136 |
||
137 |
declare split_paired_All [simp del] |
|
138 |
declare split_paired_Ex [simp del] |
|
139 |
||
140 |
lemma [dest!]: |
|
141 |
"(if P then A else B) \<noteq> B \<Longrightarrow> P" |
|
142 |
by (cases P, auto) |
|
143 |
||
144 |
lemma exec_d_no_errorI [intro]: |
|
145 |
"check G s \<Longrightarrow> exec_d G (Normal s) \<noteq> TypeError" |
|
146 |
by (unfold exec_d_def) simp |
|
147 |
||
148 |
theorem no_type_error_commutes: |
|
149 |
"exec_d G (Normal s) \<noteq> TypeError \<Longrightarrow> |
|
150 |
exec_d G (Normal s) = Normal (exec (G, s))" |
|
151 |
by (unfold exec_d_def, auto) |
|
152 |
||
153 |
||
154 |
lemma defensive_imp_aggressive: |
|
155 |
"G \<turnstile> (Normal s) -jvmd\<rightarrow> (Normal t) \<Longrightarrow> G \<turnstile> s -jvm\<rightarrow> t" |
|
156 |
proof - |
|
157 |
have "\<And>x y. G \<turnstile> x -jvmd\<rightarrow> y \<Longrightarrow> \<forall>s t. x = Normal s \<longrightarrow> y = Normal t \<longrightarrow> G \<turnstile> s -jvm\<rightarrow> t" |
|
158 |
apply (unfold exec_all_d_def) |
|
159 |
apply (erule rtrancl_induct) |
|
160 |
apply (simp add: exec_all_def) |
|
161 |
apply (fold exec_all_d_def) |
|
162 |
apply simp |
|
163 |
apply (intro allI impI) |
|
164 |
apply (erule disjE, simp) |
|
165 |
apply (elim exE conjE) |
|
166 |
apply (erule allE, erule impE, assumption) |
|
167 |
apply (simp add: exec_all_def exec_d_def split: type_error.splits split_if_asm) |
|
168 |
apply (rule rtrancl_trans, assumption) |
|
169 |
apply blast |
|
170 |
done |
|
171 |
moreover |
|
172 |
assume "G \<turnstile> (Normal s) -jvmd\<rightarrow> (Normal t)" |
|
173 |
ultimately |
|
174 |
show "G \<turnstile> s -jvm\<rightarrow> t" by blast |
|
175 |
qed |
|
176 |
||
177 |
end |