| author | chaieb | 
| Mon, 25 Feb 2008 11:27:02 +0100 | |
| changeset 26118 | 6f94eb10adad | 
| parent 22271 | 51a80e238b29 | 
| child 26450 | 158b924b5153 | 
| permissions | -rw-r--r-- | 
| 12516 | 1 | (* Title: HOL/MicroJava/BV/JVM.thy | 
| 10592 | 2 | ID: $Id$ | 
| 12516 | 3 | Author: Tobias Nipkow, Gerwin Klein | 
| 10592 | 4 | Copyright 2000 TUM | 
| 5 | *) | |
| 6 | ||
| 12911 | 7 | header {* \isaheader{Kildall for the JVM}\label{sec:JVM} *}
 | 
| 10592 | 8 | |
| 16417 | 9 | theory JVM imports Kildall Typing_Framework_JVM begin | 
| 13066 | 10 | |
| 10592 | 11 | |
| 12 | constdefs | |
| 13006 | 13 | kiljvm :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> | 
| 14 | instr list \<Rightarrow> state list \<Rightarrow> state list" | |
| 12516 | 15 | "kiljvm G maxs maxr rT et bs == | 
| 16 | kildall (JVMType.le G maxs maxr) (JVMType.sup G maxs maxr) (exec G maxs rT et bs)" | |
| 10592 | 17 | |
| 12516 | 18 | wt_kil :: "jvm_prog \<Rightarrow> cname \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> | 
| 19 | exception_table \<Rightarrow> instr list \<Rightarrow> bool" | |
| 20 | "wt_kil G C pTs rT mxs mxl et ins == | |
| 13066 | 21 | check_bounded ins et \<and> 0 < size ins \<and> | 
| 10637 | 22 | (let first = Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err)); | 
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11299diff
changeset | 23 | start = OK first#(replicate (size ins - 1) (OK None)); | 
| 12516 | 24 | result = kiljvm G mxs (1+size pTs+mxl) rT et ins start | 
| 10637 | 25 | in \<forall>n < size ins. result!n \<noteq> Err)" | 
| 26 | ||
| 13006 | 27 | wt_jvm_prog_kildall :: "jvm_prog \<Rightarrow> bool" | 
| 10637 | 28 | "wt_jvm_prog_kildall G == | 
| 12516 | 29 | wf_prog (\<lambda>G C (sig,rT,(maxs,maxl,b,et)). wt_kil G C (snd sig) rT maxs maxl et b) G" | 
| 10592 | 30 | |
| 31 | ||
| 32 | theorem is_bcv_kiljvm: | |
| 13066 | 33 | "\<lbrakk> wf_prog wf_mb G; bounded (exec G maxs rT et bs) (size bs) \<rbrakk> \<Longrightarrow> | 
| 12516 | 34 | is_bcv (JVMType.le G maxs maxr) Err (exec G maxs rT et bs) | 
| 35 | (size bs) (states G maxs maxr) (kiljvm G maxs maxr rT et bs)" | |
| 10592 | 36 | apply (unfold kiljvm_def sl_triple_conv) | 
| 37 | apply (rule is_bcv_kildall) | |
| 38 | apply (simp (no_asm) add: sl_triple_conv [symmetric]) | |
| 14045 | 39 | apply (force intro!: semilat_JVM_slI dest: wf_acyclic | 
| 40 | simp add: symmetric sl_triple_conv) | |
| 10592 | 41 | apply (simp (no_asm) add: JVM_le_unfold) | 
| 42 | apply (blast intro!: order_widen wf_converse_subcls1_impl_acc_subtype | |
| 22271 | 43 | dest: wf_subcls1 wfP_acyclicP wf_prog_ws_prog) | 
| 10592 | 44 | apply (simp add: JVM_le_unfold) | 
| 45 | apply (erule exec_pres_type) | |
| 13066 | 46 | apply assumption | 
| 14045 | 47 | apply (drule wf_prog_ws_prog, erule exec_mono, assumption) | 
| 10592 | 48 | done | 
| 49 | ||
| 13224 | 50 | lemma subset_replicate: "set (replicate n x) \<subseteq> {x}"
 | 
| 51 | by (induct n) auto | |
| 52 | ||
| 53 | lemma in_set_replicate: | |
| 54 | "x \<in> set (replicate n y) \<Longrightarrow> x = y" | |
| 55 | proof - | |
| 56 | assume "x \<in> set (replicate n y)" | |
| 57 |   also have "set (replicate n y) \<subseteq> {y}" by (rule subset_replicate)
 | |
| 58 |   finally have "x \<in> {y}" .
 | |
| 59 | thus ?thesis by simp | |
| 60 | qed | |
| 10637 | 61 | |
| 62 | theorem wt_kil_correct: | |
| 13224 | 63 | assumes wf: "wf_prog wf_mb G" | 
| 64 | assumes C: "is_class G C" | |
| 65 | assumes pTs: "set pTs \<subseteq> types G" | |
| 66 | ||
| 67 | assumes wtk: "wt_kil G C pTs rT maxs mxl et bs" | |
| 68 | ||
| 69 | shows "\<exists>phi. wt_method G C pTs rT maxs mxl bs et phi" | |
| 10592 | 70 | proof - | 
| 71 | let ?start = "OK (Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err))) | |
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11299diff
changeset | 72 | #(replicate (size bs - 1) (OK None))" | 
| 10592 | 73 | |
| 13224 | 74 | from wtk obtain maxr r where | 
| 13066 | 75 | bounded: "check_bounded bs et" and | 
| 12516 | 76 | result: "r = kiljvm G maxs maxr rT et bs ?start" and | 
| 10637 | 77 | success: "\<forall>n < size bs. r!n \<noteq> Err" and | 
| 78 | instrs: "0 < size bs" and | |
| 79 | maxr: "maxr = Suc (length pTs + mxl)" | |
| 80 | by (unfold wt_kil_def) simp | |
| 10592 | 81 | |
| 13066 | 82 | from bounded have "bounded (exec G maxs rT et bs) (size bs)" | 
| 83 | by (unfold exec_def) (intro bounded_lift check_bounded_is_bounded) | |
| 84 | with wf have bcv: | |
| 12516 | 85 | "is_bcv (JVMType.le G maxs maxr) Err (exec G maxs rT et bs) | 
| 13066 | 86 | (size bs) (states G maxs maxr) (kiljvm G maxs maxr rT et bs)" | 
| 10592 | 87 | by (rule is_bcv_kiljvm) | 
| 13066 | 88 | |
| 13224 | 89 | from C pTs instrs maxr | 
| 10592 | 90 | have "?start \<in> list (length bs) (states G maxs maxr)" | 
| 13224 | 91 | apply (unfold JVM_states_unfold) | 
| 92 | apply (rule listI) | |
| 93 | apply (auto intro: list_appendI dest!: in_set_replicate) | |
| 94 | apply force | |
| 95 | done | |
| 96 | ||
| 12516 | 97 | with bcv success result have | 
| 10592 | 98 | "\<exists>ts\<in>list (length bs) (states G maxs maxr). | 
| 10812 
ead84e90bfeb
merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
 kleing parents: 
10657diff
changeset | 99 | ?start <=[JVMType.le G maxs maxr] ts \<and> | 
| 12516 | 100 | wt_step (JVMType.le G maxs maxr) Err (exec G maxs rT et bs) ts" | 
| 10592 | 101 | by (unfold is_bcv_def) auto | 
| 102 | then obtain phi' where | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 103 | phi': "phi' \<in> list (length bs) (states G maxs maxr)" and | 
| 10812 
ead84e90bfeb
merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
 kleing parents: 
10657diff
changeset | 104 | s: "?start <=[JVMType.le G maxs maxr] phi'" and | 
| 12516 | 105 | w: "wt_step (JVMType.le G maxs maxr) Err (exec G maxs rT et bs) phi'" | 
| 10592 | 106 | by blast | 
| 13066 | 107 | hence wt_err_step: | 
| 108 | "wt_err_step (sup_state_opt G) (exec G maxs rT et bs) phi'" | |
| 109 | by (simp add: wt_err_step_def exec_def JVM_le_Err_conv) | |
| 10592 | 110 | |
| 12516 | 111 | from s have le: "JVMType.le G maxs maxr (?start ! 0) (phi'!0)" | 
| 10592 | 112 | by (drule_tac p=0 in le_listD) (simp add: lesub_def)+ | 
| 113 | ||
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 114 | from phi' have l: "size phi' = size bs" by simp | 
| 12516 | 115 | with instrs w have "phi' ! 0 \<noteq> Err" by (unfold wt_step_def) simp | 
| 116 | with instrs l have phi0: "OK (map ok_val phi' ! 0) = phi' ! 0" | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 117 | by (clarsimp simp add: not_Err_eq) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 118 | |
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 119 | from phi' have "check_types G maxs maxr phi'" by(simp add: check_types_def) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 120 | also from w have "phi' = map OK (map ok_val phi')" | 
| 13224 | 121 | apply (clarsimp simp add: wt_step_def not_Err_eq) | 
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 122 | apply (rule map_id [symmetric]) | 
| 13224 | 123 | apply (erule allE, erule impE, assumption) | 
| 124 | apply clarsimp | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 125 | done | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 126 | finally | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 127 | have check_types: | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 128 | "check_types G maxs maxr (map OK (map ok_val phi'))" . | 
| 10592 | 129 | |
| 12516 | 130 | from l bounded | 
| 13067 | 131 | have "bounded (\<lambda>pc. eff (bs!pc) G pc et) (length phi')" | 
| 13066 | 132 | by (simp add: exec_def check_bounded_is_bounded) | 
| 13067 | 133 | hence bounded': "bounded (exec G maxs rT et bs) (length bs)" | 
| 134 | by (auto intro: bounded_lift simp add: exec_def l) | |
| 13066 | 135 | with wt_err_step | 
| 136 | have "wt_app_eff (sup_state_opt G) (\<lambda>pc. app (bs!pc) G maxs rT pc et) | |
| 137 | (\<lambda>pc. eff (bs!pc) G pc et) (map ok_val phi')" | |
| 13067 | 138 | by (auto intro: wt_err_imp_wt_app_eff simp add: l exec_def) | 
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 139 | with instrs l le bounded bounded' check_types maxr | 
| 12516 | 140 | have "wt_method G C pTs rT maxs mxl bs et (map ok_val phi')" | 
| 13066 | 141 | apply (unfold wt_method_def wt_app_eff_def) | 
| 10592 | 142 | apply simp | 
| 143 | apply (rule conjI) | |
| 144 | apply (unfold wt_start_def) | |
| 145 | apply (rule JVM_le_convert [THEN iffD1]) | |
| 146 | apply (simp (no_asm) add: phi0) | |
| 147 | apply clarify | |
| 148 | apply (erule allE, erule impE, assumption) | |
| 149 | apply (elim conjE) | |
| 150 | apply (clarsimp simp add: lesub_def wt_instr_def) | |
| 13067 | 151 | apply (simp add: exec_def) | 
| 152 | apply (drule bounded_err_stepD, assumption+) | |
| 153 | apply blast | |
| 10592 | 154 | done | 
| 155 | ||
| 156 | thus ?thesis by blast | |
| 157 | qed | |
| 158 | ||
| 10651 | 159 | |
| 160 | theorem wt_kil_complete: | |
| 13224 | 161 | assumes wf: "wf_prog wf_mb G" | 
| 162 | assumes C: "is_class G C" | |
| 163 | assumes pTs: "set pTs \<subseteq> types G" | |
| 164 | ||
| 165 | assumes wtm: "wt_method G C pTs rT maxs mxl bs et phi" | |
| 166 | ||
| 167 | shows "wt_kil G C pTs rT maxs mxl et bs" | |
| 10651 | 168 | proof - | 
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 169 | let ?mxr = "1+size pTs+mxl" | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 170 | |
| 13224 | 171 | from wtm obtain | 
| 10651 | 172 | instrs: "0 < length bs" and | 
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 173 | len: "length phi = length bs" and | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 174 | bounded: "check_bounded bs et" and | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 175 | ck_types: "check_types G maxs ?mxr (map OK phi)" and | 
| 10651 | 176 | wt_start: "wt_start G C pTs mxl phi" and | 
| 177 | wt_ins: "\<forall>pc. pc < length bs \<longrightarrow> | |
| 12516 | 178 | wt_instr (bs ! pc) G rT phi maxs (length bs) et pc" | 
| 10651 | 179 | by (unfold wt_method_def) simp | 
| 180 | ||
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 181 | from ck_types len | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 182 | have istype_phi: | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 183 | "map OK phi \<in> list (length bs) (states G maxs (1+size pTs+mxl))" | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 184 | by (auto simp add: check_types_def intro!: listI) | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 185 | |
| 12516 | 186 | let ?eff = "\<lambda>pc. eff (bs!pc) G pc et" | 
| 187 | let ?app = "\<lambda>pc. app (bs!pc) G maxs rT pc et" | |
| 10651 | 188 | |
| 13066 | 189 | from bounded | 
| 190 | have bounded_exec: "bounded (exec G maxs rT et bs) (size bs)" | |
| 191 | by (unfold exec_def) (intro bounded_lift check_bounded_is_bounded) | |
| 12516 | 192 | |
| 10651 | 193 | from wt_ins | 
| 13066 | 194 | have "wt_app_eff (sup_state_opt G) ?app ?eff phi" | 
| 195 | apply (unfold wt_app_eff_def wt_instr_def lesub_def) | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 196 | apply (simp (no_asm) only: len) | 
| 10651 | 197 | apply blast | 
| 198 | done | |
| 13066 | 199 | with bounded_exec | 
| 200 | have "wt_err_step (sup_state_opt G) (err_step (size phi) ?app ?eff) (map OK phi)" | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 201 | by - (erule wt_app_eff_imp_wt_err,simp add: exec_def len) | 
| 13066 | 202 | hence wt_err: | 
| 203 | "wt_err_step (sup_state_opt G) (exec G maxs rT et bs) (map OK phi)" | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 204 | by (unfold exec_def) (simp add: len) | 
| 10651 | 205 | |
| 12516 | 206 | from wf bounded_exec | 
| 10651 | 207 | have is_bcv: | 
| 13224 | 208 | "is_bcv (JVMType.le G maxs ?mxr) Err (exec G maxs rT et bs) | 
| 209 | (size bs) (states G maxs ?mxr) (kiljvm G maxs ?mxr rT et bs)" | |
| 10651 | 210 | by (rule is_bcv_kiljvm) | 
| 211 | ||
| 212 | let ?start = "OK (Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err))) | |
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11299diff
changeset | 213 | #(replicate (size bs - 1) (OK None))" | 
| 10651 | 214 | |
| 13224 | 215 | from C pTs instrs | 
| 216 | have start: "?start \<in> list (length bs) (states G maxs ?mxr)" | |
| 217 | apply (unfold JVM_states_unfold) | |
| 218 | apply (rule listI) | |
| 219 | apply (auto intro!: list_appendI dest!: in_set_replicate) | |
| 220 | apply force | |
| 221 | done | |
| 10651 | 222 | |
| 12516 | 223 | let ?phi = "map OK phi" | 
| 13224 | 224 | have less_phi: "?start <=[JVMType.le G maxs ?mxr] ?phi" | 
| 10657 | 225 | proof - | 
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 226 | from len instrs | 
| 12516 | 227 | have "length ?start = length (map OK phi)" by simp | 
| 228 | moreover | |
| 10657 | 229 |     { fix n
 | 
| 230 | from wt_start | |
| 231 | have "G \<turnstile> ok_val (?start!0) <=' phi!0" | |
| 232 | by (simp add: wt_start_def) | |
| 233 | moreover | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 234 | from instrs len | 
| 10657 | 235 | have "0 < length phi" by simp | 
| 236 | ultimately | |
| 13224 | 237 | have "JVMType.le G maxs ?mxr (?start!0) (?phi!0)" | 
| 10657 | 238 | by (simp add: JVM_le_Err_conv Err.le_def lesub_def) | 
| 239 | moreover | |
| 240 |       { fix n'
 | |
| 13224 | 241 | have "JVMType.le G maxs ?mxr (OK None) (?phi!n)" | 
| 10657 | 242 | by (auto simp add: JVM_le_Err_conv Err.le_def lesub_def | 
| 243 | split: err.splits) | |
| 13006 | 244 | hence "\<lbrakk> n = Suc n'; n < length ?start \<rbrakk> | 
| 13224 | 245 | \<Longrightarrow> JVMType.le G maxs ?mxr (?start!n) (?phi!n)" | 
| 10657 | 246 | by simp | 
| 247 | } | |
| 248 | ultimately | |
| 13224 | 249 | have "n < length ?start \<Longrightarrow> (?start!n) <=_(JVMType.le G maxs ?mxr) (?phi!n)" | 
| 12516 | 250 | by (unfold lesub_def) (cases n, blast+) | 
| 251 | } | |
| 252 | ultimately show ?thesis by (rule le_listI) | |
| 10657 | 253 | qed | 
| 10651 | 254 | |
| 13066 | 255 | from wt_err | 
| 13224 | 256 | have "wt_step (JVMType.le G maxs ?mxr) Err (exec G maxs rT et bs) ?phi" | 
| 13066 | 257 | by (simp add: wt_err_step_def JVM_le_Err_conv) | 
| 12516 | 258 | with start istype_phi less_phi is_bcv | 
| 13224 | 259 | have "\<forall>p. p < length bs \<longrightarrow> kiljvm G maxs ?mxr rT et bs ?start ! p \<noteq> Err" | 
| 12516 | 260 | by (unfold is_bcv_def) auto | 
| 13066 | 261 | with bounded instrs | 
| 12516 | 262 | show "wt_kil G C pTs rT maxs mxl et bs" by (unfold wt_kil_def) simp | 
| 263 | qed | |
| 10651 | 264 | |
| 10637 | 265 | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 266 | theorem jvm_kildall_sound_complete: | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 267 | "wt_jvm_prog_kildall G = (\<exists>Phi. wt_jvm_prog G Phi)" | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 268 | proof | 
| 12516 | 269 | let ?Phi = "\<lambda>C sig. let (C,rT,(maxs,maxl,ins,et)) = the (method (G,C) sig) in | 
| 270 | SOME phi. wt_method G C (snd sig) rT maxs maxl ins et phi" | |
| 13224 | 271 | |
| 272 | assume "wt_jvm_prog_kildall G" | |
| 273 | hence "wt_jvm_prog G ?Phi" | |
| 274 | apply (unfold wt_jvm_prog_def wt_jvm_prog_kildall_def) | |
| 275 | apply (erule jvm_prog_lift) | |
| 276 | apply (auto dest!: wt_kil_correct intro: someI) | |
| 10637 | 277 | done | 
| 13224 | 278 | thus "\<exists>Phi. wt_jvm_prog G Phi" by fast | 
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 279 | next | 
| 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 280 | assume "\<exists>Phi. wt_jvm_prog G Phi" | 
| 13224 | 281 | thus "wt_jvm_prog_kildall G" | 
| 282 | apply (clarify) | |
| 283 | apply (unfold wt_jvm_prog_def wt_jvm_prog_kildall_def) | |
| 284 | apply (erule jvm_prog_lift) | |
| 285 | apply (auto intro: wt_kil_complete) | |
| 13214 
2aa33ed5f526
wt_method now checks bounded+types ==> wt_kildall <=> wt_method
 kleing parents: 
13074diff
changeset | 286 | done | 
| 10637 | 287 | qed | 
| 288 | ||
| 10592 | 289 | end |