author | kleing |
Sat, 09 Mar 2002 20:39:46 +0100 | |
changeset 13052 | 3bf41c474a88 |
parent 13006 | 51c5f3f11d16 |
child 13062 | 4b1edf2f6bd2 |
permissions | -rw-r--r-- |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
1 |
(* Title: HOL/MicroJava/BV/LBVSpec.thy |
8245 | 2 |
ID: $Id$ |
3 |
Author: Gerwin Klein |
|
4 |
Copyright 1999 Technische Universitaet Muenchen |
|
9054 | 5 |
*) |
8245 | 6 |
|
12911 | 7 |
header {* \isaheader{The Lightweight Bytecode Verifier} *} |
9054 | 8 |
|
12516 | 9 |
theory LBVSpec = Effect + Kildall: |
8245 | 10 |
|
12516 | 11 |
text {* |
12 |
The Lightweight Bytecode Verifier with exceptions has not |
|
13 |
made it completely into the Isabelle 2001 release. Currently |
|
14 |
there is only the specification itself available. The proofs of |
|
15 |
soundness and completeness are broken (they still need to be |
|
16 |
ported to the exception version). Both theories are included |
|
17 |
for documentation (but they don't work for this specification), |
|
12978 | 18 |
please see the Isabelle 99-2 release for a working copy, or |
19 |
\url{http://isabelle.in.tum.de/verificard} for the most recent |
|
20 |
development of \mJava. |
|
12516 | 21 |
*} |
8245 | 22 |
|
23 |
types |
|
24 |
certificate = "state_type option list" |
|
13006 | 25 |
class_certificate = "sig \<Rightarrow> certificate" |
26 |
prog_certificate = "cname \<Rightarrow> class_certificate" |
|
8245 | 27 |
|
12516 | 28 |
consts |
29 |
merge :: "[jvm_prog, p_count, nat, nat, p_count, certificate, (nat \<times> (state_type option)) list, state] \<Rightarrow> state" |
|
30 |
primrec |
|
31 |
"merge G pc mxs mxr max_pc cert [] x = x" |
|
32 |
"merge G pc mxs mxr max_pc cert (s#ss) x = (let (pc',s') = s in |
|
33 |
if pc' < max_pc \<and> pc' = pc+1 then |
|
34 |
merge G pc mxs mxr max_pc cert ss (OK s' +_(sup G mxs mxr) x) |
|
35 |
else if pc' < max_pc \<and> G \<turnstile> s' <=' cert!pc' then |
|
36 |
merge G pc mxs mxr max_pc cert ss x |
|
37 |
else Err)" |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
38 |
|
12516 | 39 |
constdefs |
40 |
wtl_inst :: "[instr,jvm_prog,ty,state_type option,certificate,nat,nat,p_count,exception_table,p_count] |
|
13006 | 41 |
\<Rightarrow> state_type option err" |
12516 | 42 |
"wtl_inst i G rT s cert maxs maxr max_pc et pc == |
43 |
if app i G maxs rT pc et s then |
|
44 |
merge G pc maxs maxr max_pc cert (eff i G pc et s) (OK (cert!(pc+1))) |
|
10628 | 45 |
else Err" |
8245 | 46 |
|
12516 | 47 |
wtl_cert :: "[instr,jvm_prog,ty,state_type option,certificate,nat,nat,p_count,exception_table,p_count] |
13006 | 48 |
\<Rightarrow> state_type option err" |
12516 | 49 |
"wtl_cert i G rT s cert maxs maxr max_pc et pc == |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
50 |
case cert!pc of |
13006 | 51 |
None \<Rightarrow> wtl_inst i G rT s cert maxs maxr max_pc et pc |
52 |
| Some s' \<Rightarrow> if G \<turnstile> s <=' (Some s') then |
|
12516 | 53 |
wtl_inst i G rT (Some s') cert maxs maxr max_pc et pc |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
54 |
else Err" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
55 |
|
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
56 |
consts |
12516 | 57 |
wtl_inst_list :: "[instr list,jvm_prog,ty,certificate,nat,nat,p_count,exception_table,p_count, |
13006 | 58 |
state_type option] \<Rightarrow> state_type option err" |
8245 | 59 |
primrec |
12516 | 60 |
"wtl_inst_list [] G rT cert maxs maxr max_pc et pc s = OK s" |
61 |
"wtl_inst_list (i#is) G rT cert maxs maxr max_pc et pc s = |
|
62 |
(let s' = wtl_cert i G rT s cert maxs maxr max_pc et pc in |
|
63 |
strict (wtl_inst_list is G rT cert maxs maxr max_pc et (pc+1)) s')" |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
64 |
|
8245 | 65 |
|
66 |
constdefs |
|
13006 | 67 |
wtl_method :: "[jvm_prog,cname,ty list,ty,nat,nat,exception_table,instr list,certificate] \<Rightarrow> bool" |
12516 | 68 |
"wtl_method G C pTs rT mxs mxl et ins cert == |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
69 |
let max_pc = length ins |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
70 |
in |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
71 |
0 < max_pc \<and> |
12516 | 72 |
wtl_inst_list ins G rT cert mxs mxl max_pc et 0 |
10496 | 73 |
(Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err))) \<noteq> Err" |
8245 | 74 |
|
13006 | 75 |
wtl_jvm_prog :: "[jvm_prog,prog_certificate] \<Rightarrow> bool" |
10042 | 76 |
"wtl_jvm_prog G cert == |
12516 | 77 |
wf_prog (\<lambda>G C (sig,rT,maxs,maxl,b,et). wtl_method G C (snd sig) rT maxs maxl et b (cert C sig)) G" |
9012 | 78 |
|
10042 | 79 |
|
10812
ead84e90bfeb
merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents:
10638
diff
changeset
|
80 |
lemmas [iff] = not_Err_eq |
ead84e90bfeb
merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents:
10638
diff
changeset
|
81 |
|
12516 | 82 |
lemma if_eq_cases: |
83 |
"(P \<Longrightarrow> x = z) \<Longrightarrow> (\<not>P \<Longrightarrow> y = z) \<Longrightarrow> (if P then x else y) = z" |
|
84 |
by simp |
|
85 |
||
86 |
lemma merge_def: |
|
13006 | 87 |
"\<And>x. merge G pc mxs mxr max_pc cert ss x = |
12516 | 88 |
(if \<forall>(pc',s') \<in> set ss. pc' < max_pc \<and> (pc' \<noteq> pc+1 \<longrightarrow> G \<turnstile> s' <=' cert!pc') then |
89 |
map (OK \<circ> snd) [(p',t') \<in> ss. p'=pc+1] ++_(sup G mxs mxr) x |
|
13006 | 90 |
else Err)" (is "\<And>x. ?merge ss x = ?if ss x" is "\<And>x. ?P ss x") |
12516 | 91 |
proof (induct ss) |
13006 | 92 |
show "\<And>x. ?P [] x" by simp |
12516 | 93 |
next |
94 |
have OK_snd: "(\<lambda>u. OK (snd u)) = OK \<circ> snd" by (rule ext) auto |
|
95 |
fix x ss and s::"nat \<times> (state_type option)" |
|
96 |
assume IH: "\<And>x. ?P ss x" |
|
97 |
obtain pc' s' where s: "s = (pc',s')" by (cases s) |
|
98 |
hence "?merge (s#ss) x = ?merge ((pc',s')#ss) x" by hypsubst (rule refl) |
|
99 |
also |
|
100 |
have "\<dots> = (if pc' < max_pc \<and> pc' = pc+1 then |
|
101 |
?merge ss (OK s' +_(sup G mxs mxr) x) |
|
102 |
else if pc' < max_pc \<and> G \<turnstile> s' <=' cert!pc' then |
|
103 |
?merge ss x |
|
104 |
else Err)" |
|
105 |
(is "_ = (if ?pc' then ?merge ss (_ +_?f _) else if ?G then _ else _)") |
|
106 |
by simp |
|
107 |
also |
|
108 |
let "if ?all ss then _ else _" = "?if ss x" |
|
109 |
have "\<dots> = ?if ((pc',s')#ss) x" |
|
110 |
proof (cases "?pc'") |
|
111 |
case True |
|
112 |
hence "?all (((pc',s')#ss)) = (pc+1 < max_pc \<and> ?all ss)" by simp |
|
113 |
with True |
|
114 |
have "?if ss (OK s' +_?f x) = ?if ((pc',s')#ss) x" by (auto simp add: OK_snd) |
|
115 |
hence "?merge ss (OK s' +_?f x) = ?if ((pc',s')#ss) x" by (simp only: IH) |
|
116 |
with True show ?thesis by (fast intro: if_eq_cases) |
|
117 |
next |
|
118 |
case False |
|
119 |
have "(if ?G then ?merge ss x else Err) = ?if ((pc',s')#ss) x" |
|
120 |
proof (cases ?G) |
|
121 |
assume G: ?G with False |
|
122 |
have "?if ss x = ?if ((pc',s')#ss) x" by (auto simp add: OK_snd) |
|
123 |
hence "?merge ss x = ?if ((pc',s')#ss) x" by (simp only: IH) |
|
124 |
with G show ?thesis by (fast intro: if_eq_cases) |
|
125 |
next |
|
126 |
assume G: "\<not>?G" |
|
127 |
with False have "Err = ?if ((pc',s')#ss) x" by auto |
|
128 |
with G show ?thesis by (fast intro: if_eq_cases) |
|
129 |
qed |
|
130 |
with False show ?thesis by (fast intro: if_eq_cases) |
|
131 |
qed |
|
132 |
also from s have "\<dots> = ?if (s#ss) x" by hypsubst (rule refl) |
|
133 |
finally show "?P (s#ss) x" . |
|
134 |
qed |
|
135 |
||
10042 | 136 |
|
10496 | 137 |
lemma wtl_inst_OK: |
12516 | 138 |
"(wtl_inst i G rT s cert mxs mxr max_pc et pc = OK s') = ( |
139 |
app i G mxs rT pc et s \<and> |
|
140 |
(\<forall>(pc',r) \<in> set (eff i G pc et s). |
|
141 |
pc' < max_pc \<and> (pc' \<noteq> pc+1 \<longrightarrow> G \<turnstile> r <=' cert!pc')) \<and> |
|
142 |
map (OK \<circ> snd) [(p',t') \<in> (eff i G pc et s). p'=pc+1] |
|
143 |
++_(sup G mxs mxr) (OK (cert!(pc+1))) = OK s')" |
|
144 |
by (auto simp add: wtl_inst_def merge_def split: split_if_asm) |
|
9012 | 145 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
146 |
lemma wtl_Cons: |
12516 | 147 |
"wtl_inst_list (i#is) G rT cert maxs maxr max_pc et pc s \<noteq> Err = |
148 |
(\<exists>s'. wtl_cert i G rT s cert maxs maxr max_pc et pc = OK s' \<and> |
|
149 |
wtl_inst_list is G rT cert maxs maxr max_pc et (pc+1) s' \<noteq> Err)" |
|
10812
ead84e90bfeb
merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents:
10638
diff
changeset
|
150 |
by (auto simp del: split_paired_Ex) |
9549
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
kleing
parents:
9376
diff
changeset
|
151 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
152 |
lemma wtl_append: |
12516 | 153 |
"\<forall> s pc. (wtl_inst_list (a@b) G rT cert mxs mxr mpc et pc s = OK s') = |
154 |
(\<exists>s''. wtl_inst_list a G rT cert mxs mxr mpc et pc s = OK s'' \<and> |
|
155 |
wtl_inst_list b G rT cert mxs mxr mpc et (pc+length a) s'' = OK s')" |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
156 |
(is "\<forall> s pc. ?C s pc a" is "?P a") |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
157 |
proof (induct ?P a) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
158 |
show "?P []" by simp |
12516 | 159 |
next |
160 |
fix x xs assume IH: "?P xs" |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
161 |
show "?P (x#xs)" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
162 |
proof (intro allI) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
163 |
fix s pc |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
164 |
show "?C s pc (x#xs)" |
12516 | 165 |
proof (cases "wtl_cert x G rT s cert mxs mxr mpc et pc") |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
166 |
case Err thus ?thesis by simp |
9183 | 167 |
next |
12516 | 168 |
fix s0 assume "wtl_cert x G rT s cert mxs mxr mpc et pc = OK s0" |
169 |
with IH have "?C s0 (Suc pc) xs" by blast |
|
170 |
thus ?thesis by (simp!) |
|
9183 | 171 |
qed |
172 |
qed |
|
173 |
qed |
|
9012 | 174 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
175 |
lemma wtl_take: |
13006 | 176 |
"wtl_inst_list is G rT cert mxs mxr mpc et pc s = OK s'' \<Longrightarrow> |
12516 | 177 |
\<exists>s'. wtl_inst_list (take pc' is) G rT cert mxs mxr mpc et pc s = OK s'" |
9183 | 178 |
proof - |
12516 | 179 |
assume "wtl_inst_list is G rT cert mxs mxr mpc et pc s = OK s''" |
180 |
hence "wtl_inst_list (take pc' is @ drop pc' is) G rT cert mxs mxr mpc et pc s = OK s''" |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
181 |
by simp |
12516 | 182 |
thus ?thesis by (auto simp add: wtl_append simp del: append_take_drop_id) |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
183 |
qed |
9012 | 184 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
185 |
lemma take_Suc: |
13006 | 186 |
"\<forall>n. n < length l \<longrightarrow> take (Suc n) l = (take n l)@[l!n]" (is "?P l") |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
187 |
proof (induct l) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
188 |
show "?P []" by simp |
12516 | 189 |
next |
190 |
fix x xs assume IH: "?P xs" |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
191 |
show "?P (x#xs)" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
192 |
proof (intro strip) |
12516 | 193 |
fix n assume "n < length (x#xs)" |
194 |
with IH show "take (Suc n) (x # xs) = take n (x # xs) @ [(x # xs) ! n]" |
|
195 |
by (cases n, auto) |
|
9183 | 196 |
qed |
197 |
qed |
|
9012 | 198 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
199 |
lemma wtl_Suc: |
13006 | 200 |
"\<lbrakk> wtl_inst_list (take pc is) G rT cert maxs maxr (length is) et 0 s = OK s'; |
12516 | 201 |
wtl_cert (is!pc) G rT s' cert maxs maxr (length is) et pc = OK s''; |
13006 | 202 |
Suc pc < length is \<rbrakk> \<Longrightarrow> |
12516 | 203 |
wtl_inst_list (take (Suc pc) is) G rT cert maxs maxr (length is) et 0 s = OK s''" |
9183 | 204 |
proof - |
12516 | 205 |
assume "wtl_inst_list (take pc is) G rT cert maxs maxr (length is) et 0 s = OK s'" |
206 |
"wtl_cert (is!pc) G rT s' cert maxs maxr (length is) et pc = OK s''" |
|
207 |
"Suc pc < length is" |
|
208 |
hence "take (Suc pc) is = (take pc is)@[is!pc]" by (simp add: take_Suc) |
|
209 |
thus ?thesis by (simp! add: wtl_append min_def) |
|
9183 | 210 |
qed |
9012 | 211 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
212 |
lemma wtl_all: |
13006 | 213 |
"\<lbrakk> wtl_inst_list is G rT cert maxs maxr (length is) et 0 s \<noteq> Err; |
214 |
pc < length is \<rbrakk> \<Longrightarrow> |
|
12516 | 215 |
\<exists>s' s''. wtl_inst_list (take pc is) G rT cert maxs maxr (length is) et 0 s = OK s' \<and> |
216 |
wtl_cert (is!pc) G rT s' cert maxs maxr (length is) et pc = OK s''" |
|
9183 | 217 |
proof - |
12516 | 218 |
assume all: "wtl_inst_list is G rT cert maxs maxr (length is) et 0 s \<noteq> Err" |
9012 | 219 |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
220 |
assume pc: "pc < length is" |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
221 |
hence "0 < length (drop pc is)" by simp |
12516 | 222 |
then obtain i r where Cons: "drop pc is = i#r" |
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
223 |
by (auto simp add: neq_Nil_conv simp del: length_drop) |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
224 |
hence "i#r = drop pc is" .. |
12516 | 225 |
with all have take: |
226 |
"wtl_inst_list (take pc is@i#r) G rT cert maxs maxr (length is) et 0 s \<noteq> Err" |
|
9757
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
227 |
by simp |
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
kleing
parents:
9664
diff
changeset
|
228 |
|
12516 | 229 |
from pc have "is!pc = drop pc is ! 0" by simp |
230 |
with Cons have "is!pc = i" by simp |
|
231 |
with take pc show ?thesis by (auto simp add: wtl_append min_def) |
|
9183 | 232 |
qed |
9012 | 233 |
|
10042 | 234 |
|
9183 | 235 |
end |