| author | blanchet | 
| Thu, 20 Sep 2012 02:42:48 +0200 | |
| changeset 49456 | fa8302c8dea1 | 
| parent 42463 | f270e3e18be5 | 
| child 58886 | 8a6cac7c7247 | 
| permissions | -rw-r--r-- | 
| 42150 | 1 | (* Title: HOL/MicroJava/BV/LBVJVM.thy | 
| 13215 | 2 | Author: Tobias Nipkow, Gerwin Klein | 
| 3 | Copyright 2000 TUM | |
| 4 | *) | |
| 5 | ||
| 6 | header {* \isaheader{LBV for the JVM}\label{sec:JVM} *}
 | |
| 7 | ||
| 17090 | 8 | theory LBVJVM | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: 
33639diff
changeset | 9 | imports Typing_Framework_JVM | 
| 17090 | 10 | begin | 
| 13215 | 11 | |
| 42463 | 12 | type_synonym prog_cert = "cname \<Rightarrow> sig \<Rightarrow> JVMType.state list" | 
| 13215 | 13 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
34227diff
changeset | 14 | definition check_cert :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> JVMType.state list \<Rightarrow> bool" where | 
| 13215 | 15 | "check_cert G mxs mxr n cert \<equiv> check_types G mxs mxr cert \<and> length cert = n+1 \<and> | 
| 16 | (\<forall>i<n. cert!i \<noteq> Err) \<and> cert!n = OK None" | |
| 17 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
34227diff
changeset | 18 | definition lbvjvm :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
34227diff
changeset | 19 | JVMType.state list \<Rightarrow> instr list \<Rightarrow> JVMType.state \<Rightarrow> JVMType.state" where | 
| 13215 | 20 | "lbvjvm G maxs maxr rT et cert bs \<equiv> | 
| 21 | wtl_inst_list bs cert (JVMType.sup G maxs maxr) (JVMType.le G maxs maxr) Err (OK None) (exec G maxs rT et bs) 0" | |
| 22 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
34227diff
changeset | 23 | definition wt_lbv :: "jvm_prog \<Rightarrow> cname \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
34227diff
changeset | 24 | exception_table \<Rightarrow> JVMType.state list \<Rightarrow> instr list \<Rightarrow> bool" where | 
| 13215 | 25 | "wt_lbv G C pTs rT mxs mxl et cert ins \<equiv> | 
| 26 | check_bounded ins et \<and> | |
| 27 | check_cert G mxs (1+size pTs+mxl) (length ins) cert \<and> | |
| 28 | 0 < size ins \<and> | |
| 29 | (let start = Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err)); | |
| 30 | result = lbvjvm G mxs (1+size pTs+mxl) rT et cert ins (OK start) | |
| 31 | in result \<noteq> Err)" | |
| 32 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
34227diff
changeset | 33 | definition wt_jvm_prog_lbv :: "jvm_prog \<Rightarrow> prog_cert \<Rightarrow> bool" where | 
| 13215 | 34 | "wt_jvm_prog_lbv G cert \<equiv> | 
| 35 | wf_prog (\<lambda>G C (sig,rT,(maxs,maxl,b,et)). wt_lbv G C (snd sig) rT maxs maxl et (cert C sig) b) G" | |
| 36 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
34227diff
changeset | 37 | definition mk_cert :: "jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> instr list | 
| 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
34227diff
changeset | 38 | \<Rightarrow> method_type \<Rightarrow> JVMType.state list" where | 
| 13215 | 39 | "mk_cert G maxs rT et bs phi \<equiv> make_cert (exec G maxs rT et bs) (map OK phi) (OK None)" | 
| 40 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
34227diff
changeset | 41 | definition prg_cert :: "jvm_prog \<Rightarrow> prog_type \<Rightarrow> prog_cert" where | 
| 13215 | 42 | "prg_cert G phi C sig \<equiv> let (C,rT,(maxs,maxl,ins,et)) = the (method (G,C) sig) in | 
| 43 | mk_cert G maxs rT et ins (phi C sig)" | |
| 44 | ||
| 45 | ||
| 46 | lemma wt_method_def2: | |
| 47 | fixes pTs and mxl and G and mxs and rT and et and bs and phi | |
| 48 | defines [simp]: "mxr \<equiv> 1 + length pTs + mxl" | |
| 49 | defines [simp]: "r \<equiv> sup_state_opt G" | |
| 50 | defines [simp]: "app0 \<equiv> \<lambda>pc. app (bs!pc) G mxs rT pc et" | |
| 51 | defines [simp]: "step0 \<equiv> \<lambda>pc. eff (bs!pc) G pc et" | |
| 52 | ||
| 53 | shows | |
| 54 | "wt_method G C pTs rT mxs mxl bs et phi = | |
| 55 | (bs \<noteq> [] \<and> | |
| 56 | length phi = length bs \<and> | |
| 57 | check_bounded bs et \<and> | |
| 58 | check_types G mxs mxr (map OK phi) \<and> | |
| 59 | wt_start G C pTs mxl phi \<and> | |
| 60 | wt_app_eff r app0 step0 phi)" | |
| 61 | by (auto simp add: wt_method_def wt_app_eff_def wt_instr_def lesub_def | |
| 62 | dest: check_bounded_is_bounded boundedD) | |
| 63 | ||
| 64 | ||
| 13224 | 65 | lemma check_certD: | 
| 66 | "check_cert G mxs mxr n cert \<Longrightarrow> cert_ok cert n Err (OK None) (states G mxs mxr)" | |
| 67 | apply (unfold cert_ok_def check_cert_def check_types_def) | |
| 17090 | 68 | apply (auto simp add: list_all_iff) | 
| 13224 | 69 | done | 
| 70 | ||
| 13215 | 71 | |
| 72 | lemma wt_lbv_wt_step: | |
| 73 | assumes wf: "wf_prog wf_mb G" | |
| 74 | assumes lbv: "wt_lbv G C pTs rT mxs mxl et cert ins" | |
| 75 | assumes C: "is_class G C" | |
| 76 | assumes pTs: "set pTs \<subseteq> types G" | |
| 77 | ||
| 78 | defines [simp]: "mxr \<equiv> 1+length pTs+mxl" | |
| 79 | ||
| 80 | shows "\<exists>ts \<in> list (size ins) (states G mxs mxr). | |
| 81 | wt_step (JVMType.le G mxs mxr) Err (exec G mxs rT et ins) ts | |
| 82 | \<and> OK (Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err))) <=_(JVMType.le G mxs mxr) ts!0" | |
| 83 | proof - | |
| 84 | let ?step = "exec G mxs rT et ins" | |
| 85 | let ?r = "JVMType.le G mxs mxr" | |
| 86 | let ?f = "JVMType.sup G mxs mxr" | |
| 87 | let ?A = "states G mxs mxr" | |
| 88 | ||
| 14045 | 89 | have "semilat (JVMType.sl G mxs mxr)" | 
| 23467 | 90 | by (rule semilat_JVM_slI, rule wf_prog_ws_prog, rule wf) | 
| 13215 | 91 | hence "semilat (?A, ?r, ?f)" by (unfold sl_triple_conv) | 
| 92 | moreover | |
| 93 | have "top ?r Err" by (simp add: JVM_le_unfold) | |
| 94 | moreover | |
| 95 | have "Err \<in> ?A" by (simp add: JVM_states_unfold) | |
| 96 | moreover | |
| 97 | have "bottom ?r (OK None)" | |
| 98 | by (simp add: JVM_le_unfold bottom_def) | |
| 99 | moreover | |
| 100 | have "OK None \<in> ?A" by (simp add: JVM_states_unfold) | |
| 101 | moreover | |
| 102 | from lbv | |
| 103 | have "bounded ?step (length ins)" | |
| 104 | by (clarsimp simp add: wt_lbv_def exec_def) | |
| 105 | (intro bounded_lift check_bounded_is_bounded) | |
| 106 | moreover | |
| 107 | from lbv | |
| 108 | have "cert_ok cert (length ins) Err (OK None) ?A" | |
| 109 | by (unfold wt_lbv_def) (auto dest: check_certD) | |
| 110 | moreover | |
| 23467 | 111 | from wf have "pres_type ?step (length ins) ?A" by (rule exec_pres_type) | 
| 13215 | 112 | moreover | 
| 113 | let ?start = "OK (Some ([],(OK (Class C))#(map OK pTs)@(replicate mxl Err)))" | |
| 114 | from lbv | |
| 115 | have "wtl_inst_list ins cert ?f ?r Err (OK None) ?step 0 ?start \<noteq> Err" | |
| 116 | by (simp add: wt_lbv_def lbvjvm_def) | |
| 117 | moreover | |
| 118 | from C pTs have "?start \<in> ?A" | |
| 119 | by (unfold JVM_states_unfold) (auto intro: list_appendI, force) | |
| 120 | moreover | |
| 121 | from lbv have "0 < length ins" by (simp add: wt_lbv_def) | |
| 122 | ultimately | |
| 27681 | 123 | show ?thesis by (rule lbvs.wtl_sound_strong [OF lbvs.intro, OF lbv.intro lbvs_axioms.intro, OF Semilat.intro lbv_axioms.intro]) | 
| 13215 | 124 | qed | 
| 125 | ||
| 126 | lemma wt_lbv_wt_method: | |
| 127 | assumes wf: "wf_prog wf_mb G" | |
| 128 | assumes lbv: "wt_lbv G C pTs rT mxs mxl et cert ins" | |
| 129 | assumes C: "is_class G C" | |
| 130 | assumes pTs: "set pTs \<subseteq> types G" | |
| 131 | ||
| 132 | shows "\<exists>phi. wt_method G C pTs rT mxs mxl ins et phi" | |
| 133 | proof - | |
| 134 | let ?mxr = "1 + length pTs + mxl" | |
| 135 | let ?step = "exec G mxs rT et ins" | |
| 136 | let ?r = "JVMType.le G mxs ?mxr" | |
| 137 | let ?f = "JVMType.sup G mxs ?mxr" | |
| 138 | let ?A = "states G mxs ?mxr" | |
| 139 | let ?start = "OK (Some ([],(OK (Class C))#(map OK pTs)@(replicate mxl Err)))" | |
| 140 | ||
| 141 | from lbv have l: "ins \<noteq> []" by (simp add: wt_lbv_def) | |
| 142 | moreover | |
| 143 | from wf lbv C pTs | |
| 144 | obtain phi where | |
| 145 | list: "phi \<in> list (length ins) ?A" and | |
| 146 | step: "wt_step ?r Err ?step phi" and | |
| 147 | start: "?start <=_?r phi!0" | |
| 148 | by (blast dest: wt_lbv_wt_step) | |
| 149 | from list have [simp]: "length phi = length ins" by simp | |
| 150 | have "length (map ok_val phi) = length ins" by simp | |
| 151 | moreover | |
| 152 | from l have 0: "0 < length phi" by simp | |
| 153 | with step obtain phi0 where "phi!0 = OK phi0" | |
| 154 | by (unfold wt_step_def) blast | |
| 155 | with start 0 | |
| 156 | have "wt_start G C pTs mxl (map ok_val phi)" | |
| 157 | by (simp add: wt_start_def JVM_le_Err_conv lesub_def) | |
| 158 | moreover | |
| 159 | from lbv have chk_bounded: "check_bounded ins et" | |
| 160 | by (simp add: wt_lbv_def) | |
| 161 |   moreover {
 | |
| 162 | from list | |
| 163 | have "check_types G mxs ?mxr phi" | |
| 164 | by (simp add: check_types_def) | |
| 165 | also from step | |
| 13224 | 166 | have [symmetric]: "map OK (map ok_val phi) = phi" | 
| 33639 
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
 hoelzl parents: 
27681diff
changeset | 167 | by (auto intro!: nth_equalityI simp add: wt_step_def) | 
| 13215 | 168 | finally have "check_types G mxs ?mxr (map OK (map ok_val phi))" . | 
| 169 | } | |
| 170 |   moreover {  
 | |
| 171 | let ?app = "\<lambda>pc. app (ins!pc) G mxs rT pc et" | |
| 172 | let ?eff = "\<lambda>pc. eff (ins!pc) G pc et" | |
| 173 | ||
| 174 | from chk_bounded | |
| 175 | have "bounded (err_step (length ins) ?app ?eff) (length ins)" | |
| 176 | by (blast dest: check_bounded_is_bounded boundedD intro: bounded_err_stepI) | |
| 177 | moreover | |
| 178 | from step | |
| 179 | have "wt_err_step (sup_state_opt G) ?step phi" | |
| 180 | by (simp add: wt_err_step_def JVM_le_Err_conv) | |
| 181 | ultimately | |
| 182 | have "wt_app_eff (sup_state_opt G) ?app ?eff (map ok_val phi)" | |
| 183 | by (auto intro: wt_err_imp_wt_app_eff simp add: exec_def) | |
| 184 | } | |
| 185 | ultimately | |
| 186 | have "wt_method G C pTs rT mxs mxl ins et (map ok_val phi)" | |
| 187 | by - (rule wt_method_def2 [THEN iffD2], simp) | |
| 188 | thus ?thesis .. | |
| 189 | qed | |
| 190 | ||
| 191 | ||
| 192 | lemma wt_method_wt_lbv: | |
| 193 | assumes wf: "wf_prog wf_mb G" | |
| 194 | assumes wt: "wt_method G C pTs rT mxs mxl ins et phi" | |
| 195 | assumes C: "is_class G C" | |
| 196 | assumes pTs: "set pTs \<subseteq> types G" | |
| 197 | ||
| 198 | defines [simp]: "cert \<equiv> mk_cert G mxs rT et ins phi" | |
| 199 | ||
| 200 | shows "wt_lbv G C pTs rT mxs mxl et cert ins" | |
| 201 | proof - | |
| 202 | let ?mxr = "1 + length pTs + mxl" | |
| 203 | let ?step = "exec G mxs rT et ins" | |
| 204 | let ?app = "\<lambda>pc. app (ins!pc) G mxs rT pc et" | |
| 205 | let ?eff = "\<lambda>pc. eff (ins!pc) G pc et" | |
| 206 | let ?r = "JVMType.le G mxs ?mxr" | |
| 207 | let ?f = "JVMType.sup G mxs ?mxr" | |
| 208 | let ?A = "states G mxs ?mxr" | |
| 209 | let ?phi = "map OK phi" | |
| 210 | let ?cert = "make_cert ?step ?phi (OK None)" | |
| 211 | ||
| 34227 | 212 | from wt have | 
| 13215 | 213 | 0: "0 < length ins" and | 
| 214 | length: "length ins = length ?phi" and | |
| 215 | ck_bounded: "check_bounded ins et" and | |
| 216 | ck_types: "check_types G mxs ?mxr ?phi" and | |
| 217 | wt_start: "wt_start G C pTs mxl phi" and | |
| 218 | app_eff: "wt_app_eff (sup_state_opt G) ?app ?eff phi" | |
| 34227 | 219 | by (simp_all add: wt_method_def2) | 
| 13215 | 220 | |
| 14045 | 221 | have "semilat (JVMType.sl G mxs ?mxr)" | 
| 23467 | 222 | by (rule semilat_JVM_slI) (rule wf_prog_ws_prog [OF wf]) | 
| 13215 | 223 | hence "semilat (?A, ?r, ?f)" by (unfold sl_triple_conv) | 
| 224 | moreover | |
| 225 | have "top ?r Err" by (simp add: JVM_le_unfold) | |
| 226 | moreover | |
| 227 | have "Err \<in> ?A" by (simp add: JVM_states_unfold) | |
| 228 | moreover | |
| 229 | have "bottom ?r (OK None)" | |
| 230 | by (simp add: JVM_le_unfold bottom_def) | |
| 231 | moreover | |
| 232 | have "OK None \<in> ?A" by (simp add: JVM_states_unfold) | |
| 233 | moreover | |
| 234 | from ck_bounded | |
| 235 | have bounded: "bounded ?step (length ins)" | |
| 236 | by (clarsimp simp add: exec_def) | |
| 237 | (intro bounded_lift check_bounded_is_bounded) | |
| 238 | with wf | |
| 14045 | 239 | have "mono ?r ?step (length ins) ?A" | 
| 240 | by (rule wf_prog_ws_prog [THEN exec_mono]) | |
| 13215 | 241 | hence "mono ?r ?step (length ?phi) ?A" by (simp add: length) | 
| 242 | moreover | |
| 23467 | 243 | from wf have "pres_type ?step (length ins) ?A" by (rule exec_pres_type) | 
| 13215 | 244 | hence "pres_type ?step (length ?phi) ?A" by (simp add: length) | 
| 245 | moreover | |
| 246 | from ck_types | |
| 247 | have "set ?phi \<subseteq> ?A" by (simp add: check_types_def) | |
| 248 | hence "\<forall>pc. pc < length ?phi \<longrightarrow> ?phi!pc \<in> ?A \<and> ?phi!pc \<noteq> Err" by auto | |
| 249 | moreover | |
| 250 | from bounded | |
| 251 | have "bounded (exec G mxs rT et ins) (length ?phi)" by (simp add: length) | |
| 252 | moreover | |
| 253 | have "OK None \<noteq> Err" by simp | |
| 254 | moreover | |
| 255 | from bounded length app_eff | |
| 256 | have "wt_err_step (sup_state_opt G) ?step ?phi" | |
| 257 | by (auto intro: wt_app_eff_imp_wt_err simp add: exec_def) | |
| 258 | hence "wt_step ?r Err ?step ?phi" | |
| 259 | by (simp add: wt_err_step_def JVM_le_Err_conv) | |
| 260 | moreover | |
| 261 | let ?start = "OK (Some ([],(OK (Class C))#(map OK pTs)@(replicate mxl Err)))" | |
| 262 | from 0 length have "0 < length phi" by auto | |
| 263 | hence "?phi!0 = OK (phi!0)" by simp | |
| 264 | with wt_start have "?start <=_?r ?phi!0" | |
| 265 | by (clarsimp simp add: wt_start_def lesub_def JVM_le_Err_conv) | |
| 266 | moreover | |
| 267 | from C pTs have "?start \<in> ?A" | |
| 268 | by (unfold JVM_states_unfold) (auto intro: list_appendI, force) | |
| 269 | moreover | |
| 270 | have "?start \<noteq> Err" by simp | |
| 271 | moreover | |
| 272 | note length | |
| 273 | ultimately | |
| 274 | have "wtl_inst_list ins ?cert ?f ?r Err (OK None) ?step 0 ?start \<noteq> Err" | |
| 27681 | 275 | by (rule lbvc.wtl_complete [OF lbvc.intro, OF lbv.intro lbvc_axioms.intro, OF Semilat.intro lbv_axioms.intro]) | 
| 13215 | 276 | moreover | 
| 277 | from 0 length have "phi \<noteq> []" by auto | |
| 278 | moreover | |
| 279 | from ck_types | |
| 280 | have "check_types G mxs ?mxr ?cert" | |
| 281 | by (auto simp add: make_cert_def check_types_def JVM_states_unfold) | |
| 282 | moreover | |
| 283 | note ck_bounded 0 length | |
| 284 | ultimately | |
| 285 | show ?thesis | |
| 286 | by (simp add: wt_lbv_def lbvjvm_def mk_cert_def | |
| 287 | check_cert_def make_cert_def nth_append) | |
| 288 | qed | |
| 289 | ||
| 290 | ||
| 13224 | 291 | |
| 292 | theorem jvm_lbv_correct: | |
| 293 | "wt_jvm_prog_lbv G Cert \<Longrightarrow> \<exists>Phi. wt_jvm_prog G Phi" | |
| 294 | proof - | |
| 295 | let ?Phi = "\<lambda>C sig. let (C,rT,(maxs,maxl,ins,et)) = the (method (G,C) sig) in | |
| 296 | SOME phi. wt_method G C (snd sig) rT maxs maxl ins et phi" | |
| 297 | ||
| 298 | assume "wt_jvm_prog_lbv G Cert" | |
| 299 | hence "wt_jvm_prog G ?Phi" | |
| 300 | apply (unfold wt_jvm_prog_def wt_jvm_prog_lbv_def) | |
| 301 | apply (erule jvm_prog_lift) | |
| 302 | apply (auto dest: wt_lbv_wt_method intro: someI) | |
| 303 | done | |
| 304 | thus ?thesis by blast | |
| 305 | qed | |
| 306 | ||
| 13215 | 307 | theorem jvm_lbv_complete: | 
| 308 | "wt_jvm_prog G Phi \<Longrightarrow> wt_jvm_prog_lbv G (prg_cert G Phi)" | |
| 13224 | 309 | apply (unfold wt_jvm_prog_def wt_jvm_prog_lbv_def) | 
| 310 | apply (erule jvm_prog_lift) | |
| 19437 | 311 | apply (auto simp add: prg_cert_def intro: wt_method_wt_lbv) | 
| 13224 | 312 | done | 
| 13215 | 313 | |
| 314 | end |