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