| author | wenzelm | 
| Tue, 09 Sep 2008 16:29:34 +0200 | |
| changeset 28177 | 8c0335bc9336 | 
| parent 27680 | 5a557a5afc48 | 
| child 33954 | 1bc3b688548c | 
| permissions | -rw-r--r-- | 
| 8011 | 1 | (* Title: HOL/MicroJava/BV/BVSpec.thy | 
| 2 | ID: $Id$ | |
| 12516 | 3 | Author: Cornelia Pusch, Gerwin Klein | 
| 8011 | 4 | Copyright 1999 Technische Universitaet Muenchen | 
| 5 | ||
| 6 | *) | |
| 7 | ||
| 12911 | 8 | header {* \isaheader{The Bytecode Verifier}\label{sec:BVSpec} *}
 | 
| 9757 
1024a2d80ac0
functional LBV style, dead code, type safety -> Isar
 kleing parents: 
9559diff
changeset | 9 | |
| 27680 | 10 | theory BVSpec | 
| 11 | imports Effect | |
| 12 | begin | |
| 12516 | 13 | |
| 14 | text {*
 | |
| 15 | This theory contains a specification of the BV. The specification | |
| 16 | describes correct typings of method bodies; it corresponds | |
| 17 |   to type \emph{checking}.
 | |
| 18 | *} | |
| 8011 | 19 | |
| 20 | constdefs | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 21 | -- "The program counter will always be inside the method:" | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 22 | check_bounded :: "instr list \<Rightarrow> exception_table \<Rightarrow> bool" | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 23 | "check_bounded ins et \<equiv> | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 24 | (\<forall>pc < length ins. \<forall>pc' \<in> set (succs (ins!pc) pc). pc' < length ins) \<and> | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 25 | (\<forall>e \<in> set et. fst (snd (snd e)) < length ins)" | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 26 | |
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 27 | -- "The method type only contains declared classes:" | 
| 26453 
758c6fef7b94
avoid ambiguity of State.state vs. JVMType.state;
 wenzelm parents: 
16417diff
changeset | 28 | check_types :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> JVMType.state list \<Rightarrow> bool" | 
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 29 | "check_types G mxs mxr phi \<equiv> set phi \<subseteq> states G mxs mxr" | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 30 | |
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 31 | -- "An instruction is welltyped if it is applicable and its effect" | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 32 | -- "is compatible with the type at all successor instructions:" | 
| 12516 | 33 | wt_instr :: "[instr,jvm_prog,ty,method_type,nat,p_count, | 
| 13006 | 34 | exception_table,p_count] \<Rightarrow> bool" | 
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 35 | "wt_instr i G rT phi mxs max_pc et pc \<equiv> | 
| 12516 | 36 | app i G mxs rT pc et (phi!pc) \<and> | 
| 37 | (\<forall>(pc',s') \<in> set (eff i G pc et (phi!pc)). pc' < max_pc \<and> G \<turnstile> s' <=' phi!pc')" | |
| 9549 
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
 kleing parents: 
9376diff
changeset | 38 | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 39 |   -- {* The type at @{text "pc=0"} conforms to the method calling convention: *}
 | 
| 13006 | 40 | wt_start :: "[jvm_prog,cname,ty list,nat,method_type] \<Rightarrow> bool" | 
| 12516 | 41 | "wt_start G C pTs mxl phi == | 
| 42 | G \<turnstile> Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err)) <=' phi!0" | |
| 8011 | 43 | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 44 | -- "A method is welltyped if the body is not empty, if execution does not" | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 45 | -- "leave the body, if the method type covers all instructions and mentions" | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 46 | -- "declared classes only, if the method calling convention is respected, and" | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 47 | -- "if all instructions are welltyped." | 
| 12516 | 48 | wt_method :: "[jvm_prog,cname,ty list,ty,nat,nat,instr list, | 
| 13006 | 49 | exception_table,method_type] \<Rightarrow> bool" | 
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 50 | "wt_method G C pTs rT mxs mxl ins et phi \<equiv> | 
| 12516 | 51 | let max_pc = length ins in | 
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 52 | 0 < max_pc \<and> | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 53 | length phi = length ins \<and> | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 54 | check_bounded ins et \<and> | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 55 | check_types G mxs (1+length pTs+mxl) (map OK phi) \<and> | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 56 | wt_start G C pTs mxl phi \<and> | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 57 | (\<forall>pc. pc<max_pc \<longrightarrow> wt_instr (ins!pc) G rT phi mxs max_pc et pc)" | 
| 8011 | 58 | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 59 | -- "A program is welltyped if it is wellformed and all methods are welltyped" | 
| 13006 | 60 | wt_jvm_prog :: "[jvm_prog,prog_type] \<Rightarrow> bool" | 
| 12516 | 61 | "wt_jvm_prog G phi == | 
| 62 | wf_prog (\<lambda>G C (sig,rT,(maxs,maxl,b,et)). | |
| 63 | wt_method G C (snd sig) rT maxs maxl b et (phi C sig)) G" | |
| 9559 | 64 | |
| 65 | ||
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 66 | lemma check_boundedD: | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 67 | "\<lbrakk> check_bounded ins et; pc < length ins; | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 68 | (pc',s') \<in> set (eff (ins!pc) G pc et s) \<rbrakk> \<Longrightarrow> | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 69 | pc' < length ins" | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 70 | apply (unfold eff_def) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 71 | apply simp | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 72 | apply (unfold check_bounded_def) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 73 | apply clarify | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 74 | apply (erule disjE) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 75 | apply blast | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 76 | apply (erule allE, erule impE, assumption) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 77 | apply (unfold xcpt_eff_def) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 78 | apply clarsimp | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 79 | apply (drule xcpt_names_in_et) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 80 | apply clarify | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 81 | apply (drule bspec, assumption) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 82 | apply simp | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 83 | done | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 84 | |
| 9549 
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
 kleing parents: 
9376diff
changeset | 85 | lemma wt_jvm_progD: | 
| 13006 | 86 | "wt_jvm_prog G phi \<Longrightarrow> (\<exists>wt. wf_prog wt G)" | 
| 12516 | 87 | by (unfold wt_jvm_prog_def, blast) | 
| 9549 
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
 kleing parents: 
9376diff
changeset | 88 | |
| 10629 | 89 | lemma wt_jvm_prog_impl_wt_instr: | 
| 13006 | 90 | "\<lbrakk> wt_jvm_prog G phi; is_class G C; | 
| 91 | method (G,C) sig = Some (C,rT,maxs,maxl,ins,et); pc < length ins \<rbrakk> | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 92 | \<Longrightarrow> wt_instr (ins!pc) G rT (phi C sig) maxs (length ins) et pc" | 
| 12516 | 93 | by (unfold wt_jvm_prog_def, drule method_wf_mdecl, | 
| 94 | simp, simp, simp add: wf_mdecl_def wt_method_def) | |
| 9549 
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
 kleing parents: 
9376diff
changeset | 95 | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 96 | text {*
 | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 97 |   We could leave out the check @{term "pc' < max_pc"} in the 
 | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 98 |   definition of @{term wt_instr} in the context of @{term wt_method}.
 | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 99 | *} | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 100 | lemma wt_instr_def2: | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 101 | "\<lbrakk> wt_jvm_prog G Phi; is_class G C; | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 102 | method (G,C) sig = Some (C,rT,maxs,maxl,ins,et); pc < length ins; | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 103 | i = ins!pc; phi = Phi C sig; max_pc = length ins \<rbrakk> | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 104 | \<Longrightarrow> wt_instr i G rT phi maxs max_pc et pc = | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 105 | (app i G maxs rT pc et (phi!pc) \<and> | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 106 | (\<forall>(pc',s') \<in> set (eff i G pc et (phi!pc)). G \<turnstile> s' <=' phi!pc'))" | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 107 | apply (simp add: wt_instr_def) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 108 | apply (unfold wt_jvm_prog_def) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 109 | apply (drule method_wf_mdecl) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 110 | apply (simp, simp, simp add: wf_mdecl_def wt_method_def) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 111 | apply (auto dest: check_boundedD) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 112 | done | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13006diff
changeset | 113 | |
| 10629 | 114 | lemma wt_jvm_prog_impl_wt_start: | 
| 13006 | 115 | "\<lbrakk> wt_jvm_prog G phi; is_class G C; | 
| 116 | method (G,C) sig = Some (C,rT,maxs,maxl,ins,et) \<rbrakk> \<Longrightarrow> | |
| 12516 | 117 | 0 < (length ins) \<and> wt_start G C (snd sig) maxl (phi C sig)" | 
| 118 | by (unfold wt_jvm_prog_def, drule method_wf_mdecl, | |
| 119 | simp, simp, simp add: wf_mdecl_def wt_method_def) | |
| 9549 
40d64cb4f4e6
BV and LBV specified in terms of app and step functions
 kleing parents: 
9376diff
changeset | 120 | |
| 8011 | 121 | end |