1 (* Title: HOL/Bali/WellForm.thy
3 Author: David von Oheimb and Norbert Schirmer
6 header {* Well-formedness of Java programs *}
7 theory WellForm imports DefiniteAssignment begin
10 For static checks on expressions and statements, see WellType.thy
12 improvements over Java Specification 1.0 (cf. 8.4.6.3, 8.4.6.4, 9.4.1):
14 \item a method implementing or overwriting another method may have a result
15 type that widens to the result type of the other method
16 (instead of identical type)
17 \item if a method hides another method (both methods have to be static!)
18 there are no restrictions to the result type
19 since the methods have to be static and there is no dynamic binding of
21 \item if an interface inherits more than one method with the same signature, the
22 methods need not have identical return types
26 \item Object and standard exceptions are assumed to be declared like normal
31 section "well-formed field declarations"
32 text {* well-formed field declaration (common part for classes and interfaces),
36 wf_fdecl :: "prog \<Rightarrow> pname \<Rightarrow> fdecl \<Rightarrow> bool"
37 "wf_fdecl G P \<equiv> \<lambda>(fn,f). is_acc_type G P (type f)"
39 lemma wf_fdecl_def2: "\<And>fd. wf_fdecl G P fd = is_acc_type G P (type (snd fd))"
40 apply (unfold wf_fdecl_def)
46 section "well-formed method declarations"
47 (*well-formed method declaration,cf. 8.4, 8.4.1, 8.4.3, 8.4.5, 14.3.2, (9.4)*)
48 (* cf. 14.15, 15.7.2, for scope issues cf. 8.4.1 and 14.3.2 *)
51 A method head is wellformed if:
53 \item the signature and the method head agree in the number of parameters
54 \item all types of the parameters are visible
55 \item the result type is visible
56 \item the parameter names are unique
60 wf_mhead :: "prog \<Rightarrow> pname \<Rightarrow> sig \<Rightarrow> mhead \<Rightarrow> bool"
61 "wf_mhead G P \<equiv> \<lambda> sig mh. length (parTs sig) = length (pars mh) \<and>
62 \<spacespace> ( \<forall>T\<in>set (parTs sig). is_acc_type G P T) \<and>
63 is_acc_type G P (resTy mh) \<and>
64 \<spacespace> distinct (pars mh)"
68 A method declaration is wellformed if:
70 \item the method head is wellformed
71 \item the names of the local variables are unique
72 \item the types of the local variables must be accessible
73 \item the local variables don't shadow the parameters
74 \item the class of the method is defined
75 \item the body statement is welltyped with respect to the
76 modified environment of local names, were the local variables,
77 the parameters the special result variable (Res) and This are assoziated
82 constdefs callee_lcl:: "qtname \<Rightarrow> sig \<Rightarrow> methd \<Rightarrow> lenv"
84 \<equiv> \<lambda> k. (case k of
86 \<Rightarrow> (case e of
88 \<Rightarrow>(table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v
89 | Res \<Rightarrow> Some (resTy m))
90 | This \<Rightarrow> if is_static m then None else Some (Class C))"
92 constdefs parameters :: "methd \<Rightarrow> lname set"
93 "parameters m \<equiv> set (map (EName \<circ> VNam) (pars m))
94 \<union> (if (static m) then {} else {This})"
97 wf_mdecl :: "prog \<Rightarrow> qtname \<Rightarrow> mdecl \<Rightarrow> bool"
98 "wf_mdecl G C \<equiv>
100 wf_mhead G (pid C) sig (mhead m) \<and>
101 unique (lcls (mbody m)) \<and>
102 (\<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T) \<and>
103 (\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None) \<and>
104 jumpNestingOkS {Ret} (stmt (mbody m)) \<and>
106 \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd> \<and>
107 (\<exists> A. \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>
108 \<turnstile> parameters m \<guillemotright>\<langle>stmt (mbody m)\<rangle>\<guillemotright> A
109 \<and> Result \<in> nrm A)"
111 lemma callee_lcl_VNam_simp [simp]:
112 "callee_lcl C sig m (EName (VNam v))
113 = (table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v"
114 by (simp add: callee_lcl_def)
116 lemma callee_lcl_Res_simp [simp]:
117 "callee_lcl C sig m (EName Res) = Some (resTy m)"
118 by (simp add: callee_lcl_def)
120 lemma callee_lcl_This_simp [simp]:
121 "callee_lcl C sig m (This) = (if is_static m then None else Some (Class C))"
122 by (simp add: callee_lcl_def)
124 lemma callee_lcl_This_static_simp:
125 "is_static m \<Longrightarrow> callee_lcl C sig m (This) = None"
128 lemma callee_lcl_This_not_static_simp:
129 "\<not> is_static m \<Longrightarrow> callee_lcl C sig m (This) = Some (Class C)"
133 "\<lbrakk>length (parTs sig) = length (pars m); \<forall>T\<in>set (parTs sig). is_acc_type G P T;
134 is_acc_type G P (resTy m); distinct (pars m)\<rbrakk> \<Longrightarrow>
136 apply (unfold wf_mhead_def)
137 apply (simp (no_asm_simp))
140 lemma wf_mdeclI: "\<lbrakk>
141 wf_mhead G (pid C) sig (mhead m); unique (lcls (mbody m));
142 (\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None);
143 \<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T;
144 jumpNestingOkS {Ret} (stmt (mbody m));
146 \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd>;
147 (\<exists> A. \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr> \<turnstile> parameters m \<guillemotright>\<langle>stmt (mbody m)\<rangle>\<guillemotright> A
148 \<and> Result \<in> nrm A)
149 \<rbrakk> \<Longrightarrow>
150 wf_mdecl G C (sig,m)"
151 apply (unfold wf_mdecl_def)
155 lemma wf_mdeclE [consumes 1]:
156 "\<lbrakk>wf_mdecl G C (sig,m);
157 \<lbrakk>wf_mhead G (pid C) sig (mhead m); unique (lcls (mbody m));
158 \<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None;
159 \<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T;
160 jumpNestingOkS {Ret} (stmt (mbody m));
162 \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd>;
163 (\<exists> A. \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile> parameters m \<guillemotright>\<langle>stmt (mbody m)\<rangle>\<guillemotright> A
164 \<and> Result \<in> nrm A)
165 \<rbrakk> \<Longrightarrow> P
166 \<rbrakk> \<Longrightarrow> P"
167 by (unfold wf_mdecl_def) simp
171 "wf_mdecl G C (sig,m) \<Longrightarrow>
172 wf_mhead G (pid C) sig (mhead m) \<and> unique (lcls (mbody m)) \<and>
173 (\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None) \<and>
174 (\<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T)"
175 apply (unfold wf_mdecl_def)
179 lemma wf_mdecl_bodyD:
180 "wf_mdecl G C (sig,m) \<Longrightarrow>
181 (\<exists>T. \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile>Body C (stmt (mbody m))\<Colon>-T \<and>
182 G\<turnstile>T\<preceq>(resTy m))"
183 apply (unfold wf_mdecl_def)
185 apply (rule_tac x="(resTy m)" in exI)
186 apply (unfold wf_mhead_def)
187 apply (auto simp add: wf_mhead_def is_acc_type_def intro: wt.Body )
192 lemma static_Object_methodsE [elim!]:
193 "\<lbrakk>wf_mdecl G Object (sig, m);static m\<rbrakk> \<Longrightarrow> R"
194 apply (unfold wf_mdecl_def)
199 lemma rT_is_acc_type:
200 "wf_mhead G P sig m \<Longrightarrow> is_acc_type G P (resTy m)"
201 apply (unfold wf_mhead_def)
205 section "well-formed interface declarations"
206 (* well-formed interface declaration, cf. 9.1, 9.1.2.1, 9.1.3, 9.4 *)
209 A interface declaration is wellformed if:
211 \item the interface hierarchy is wellstructured
212 \item there is no class with the same name
213 \item the method heads are wellformed and not static and have Public access
214 \item the methods are uniquely named
215 \item all superinterfaces are accessible
216 \item the result type of a method overriding a method of Object widens to the
217 result type of the overridden method.
218 Shadowing static methods is forbidden.
219 \item the result type of a method overriding a set of methods defined in the
220 superinterfaces widens to each of the corresponding result types
224 wf_idecl :: "prog \<Rightarrow> idecl \<Rightarrow> bool"
227 ws_idecl G I (isuperIfs i) \<and>
228 \<not>is_class G I \<and>
229 (\<forall>(sig,mh)\<in>set (imethods i). wf_mhead G (pid I) sig mh \<and>
230 \<not>is_static mh \<and>
231 accmodi mh = Public) \<and>
232 unique (imethods i) \<and>
233 (\<forall> J\<in>set (isuperIfs i). is_acc_iface G (pid I) J) \<and>
234 (table_of (imethods i)
235 hiding (methd G Object)
236 under (\<lambda> new old. accmodi old \<noteq> Private)
237 entails (\<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old \<and>
238 is_static new = is_static old)) \<and>
239 (o2s \<circ> table_of (imethods i)
240 hidings Un_tables((\<lambda>J.(imethds G J))`set (isuperIfs i))
241 entails (\<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old))"
243 lemma wf_idecl_mhead: "\<lbrakk>wf_idecl G (I,i); (sig,mh)\<in>set (imethods i)\<rbrakk> \<Longrightarrow>
244 wf_mhead G (pid I) sig mh \<and> \<not>is_static mh \<and> accmodi mh = Public"
245 apply (unfold wf_idecl_def)
249 lemma wf_idecl_hidings:
250 "wf_idecl G (I, i) \<Longrightarrow>
251 (\<lambda>s. o2s (table_of (imethods i) s))
252 hidings Un_tables ((\<lambda>J. imethds G J) ` set (isuperIfs i))
253 entails \<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old"
254 apply (unfold wf_idecl_def o_def)
258 lemma wf_idecl_hiding:
259 "wf_idecl G (I, i) \<Longrightarrow>
260 (table_of (imethods i)
261 hiding (methd G Object)
262 under (\<lambda> new old. accmodi old \<noteq> Private)
263 entails (\<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old \<and>
264 is_static new = is_static old))"
265 apply (unfold wf_idecl_def)
270 "\<lbrakk>wf_idecl G (I,i); J \<in> set (isuperIfs i)\<rbrakk>
271 \<Longrightarrow> is_acc_iface G (pid I) J \<and> (J, I) \<notin> (subint1 G)^+"
272 apply (unfold wf_idecl_def ws_idecl_def)
276 section "well-formed class declarations"
277 (* well-formed class declaration, cf. 8.1, 8.1.2.1, 8.1.2.2, 8.1.3, 8.1.4 and
278 class method declaration, cf. 8.4.3.3, 8.4.6.1, 8.4.6.2, 8.4.6.3, 8.4.6.4 *)
281 A class declaration is wellformed if:
283 \item there is no interface with the same name
284 \item all superinterfaces are accessible and for all methods implementing
285 an interface method the result type widens to the result type of
286 the interface method, the method is not static and offers at least
288 (this actually means that the method has Public access, since all
289 interface methods have public access)
290 \item all field declarations are wellformed and the field names are unique
291 \item all method declarations are wellformed and the method names are unique
292 \item the initialization statement is welltyped
293 \item the classhierarchy is wellstructured
294 \item Unless the class is Object:
296 \item the superclass is accessible
297 \item for all methods overriding another method (of a superclass )the
298 result type widens to the result type of the overridden method,
299 the access modifier of the new method provides at least as much
300 access as the overwritten one.
301 \item for all methods hiding a method (of a superclass) the hidden
302 method must be static and offer at least as much access rights.
303 Remark: In contrast to the Java Language Specification we don't
304 restrict the result types of the method
305 (as in case of overriding), because there seems to be no reason,
306 since there is no dynamic binding of static methods.
307 (cf. 8.4.6.3 vs. 15.12.1).
308 Stricly speaking the restrictions on the access rights aren't
309 necessary to, since the static type and the access rights
310 together determine which method is to be called statically.
311 But if a class gains more then one static method with the
312 same signature due to inheritance, it is confusing when the
313 method selection depends on the access rights only:
315 Class C declares static public method foo().
316 Class D is subclass of C and declares static method foo()
317 with default package access.
318 D.foo() ? if this call is in the same package as D then
319 foo of class D is called, otherwise foo of class C.
325 constdefs entails:: "('a,'b) table \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool"
327 "t entails P \<equiv> \<forall>k. \<forall> x \<in> t k: P x"
330 "\<lbrakk>t entails P; t k = Some x\<rbrakk> \<Longrightarrow> P x"
331 by (simp add: entails_def)
333 lemma empty_entails[simp]: "empty entails P"
334 by (simp add: entails_def)
337 wf_cdecl :: "prog \<Rightarrow> cdecl \<Rightarrow> bool"
340 \<not>is_iface G C \<and>
341 (\<forall>I\<in>set (superIfs c). is_acc_iface G (pid C) I \<and>
342 (\<forall>s. \<forall> im \<in> imethds G I s.
343 (\<exists> cm \<in> methd G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and>
344 \<not> is_static cm \<and>
345 accmodi im \<le> accmodi cm))) \<and>
346 (\<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f) \<and> unique (cfields c) \<and>
347 (\<forall>m\<in>set (methods c). wf_mdecl G C m) \<and> unique (methods c) \<and>
348 jumpNestingOkS {} (init c) \<and>
349 (\<exists> A. \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile> {} \<guillemotright>\<langle>init c\<rangle>\<guillemotright> A) \<and>
350 \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd> \<and> ws_cdecl G C (super c) \<and>
351 (C \<noteq> Object \<longrightarrow>
352 (is_acc_class G (pid C) (super c) \<and>
353 (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c))
354 entails (\<lambda> new. \<forall> old sig.
355 (G,sig\<turnstile>new overrides\<^sub>S old
356 \<longrightarrow> (G\<turnstile>resTy new\<preceq>resTy old \<and>
357 accmodi old \<le> accmodi new \<and>
358 \<not>is_static old)) \<and>
359 (G,sig\<turnstile>new hides old
360 \<longrightarrow> (accmodi old \<le> accmodi new \<and>
366 wf_cdecl :: "prog \<Rightarrow> cdecl \<Rightarrow> bool"
369 \<not>is_iface G C \<and>
370 (\<forall>I\<in>set (superIfs c). is_acc_iface G (pid C) I \<and>
371 (\<forall>s. \<forall> im \<in> imethds G I s.
372 (\<exists> cm \<in> methd G C s: G\<turnstile>resTy (mthd cm)\<preceq>resTy (mthd im) \<and>
373 \<not> is_static cm \<and>
374 accmodi im \<le> accmodi cm))) \<and>
375 (\<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f) \<and> unique (cfields c) \<and>
376 (\<forall>m\<in>set (methods c). wf_mdecl G C m) \<and> unique (methods c) \<and>
377 \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd> \<and> ws_cdecl G C (super c) \<and>
378 (C \<noteq> Object \<longrightarrow>
379 (is_acc_class G (pid C) (super c) \<and>
380 (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c))
381 hiding methd G (super c)
382 under (\<lambda> new old. G\<turnstile>new overrides old)
383 entails (\<lambda> new old.
384 (G\<turnstile>resTy (mthd new)\<preceq>resTy (mthd old) \<and>
385 accmodi old \<le> accmodi new \<and>
386 \<not> is_static old))) \<and>
387 (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c))
388 hiding methd G (super c)
389 under (\<lambda> new old. G\<turnstile>new hides old)
390 entails (\<lambda> new old. is_static old \<and>
391 accmodi old \<le> accmodi new)) \<and>
392 (table_of (cfields c) hiding accfield G C (super c)
393 entails (\<lambda> newF oldF. accmodi oldF \<le> access newF))))"
396 lemma wf_cdeclE [consumes 1]:
397 "\<lbrakk>wf_cdecl G (C,c);
398 \<lbrakk>\<not>is_iface G C;
399 (\<forall>I\<in>set (superIfs c). is_acc_iface G (pid C) I \<and>
400 (\<forall>s. \<forall> im \<in> imethds G I s.
401 (\<exists> cm \<in> methd G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and>
402 \<not> is_static cm \<and>
403 accmodi im \<le> accmodi cm)));
404 \<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f; unique (cfields c);
405 \<forall>m\<in>set (methods c). wf_mdecl G C m; unique (methods c);
406 jumpNestingOkS {} (init c);
407 \<exists> A. \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile> {} \<guillemotright>\<langle>init c\<rangle>\<guillemotright> A;
408 \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd>;
409 ws_cdecl G C (super c);
410 (C \<noteq> Object \<longrightarrow>
411 (is_acc_class G (pid C) (super c) \<and>
412 (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c))
413 entails (\<lambda> new. \<forall> old sig.
414 (G,sig\<turnstile>new overrides\<^sub>S old
415 \<longrightarrow> (G\<turnstile>resTy new\<preceq>resTy old \<and>
416 accmodi old \<le> accmodi new \<and>
417 \<not>is_static old)) \<and>
418 (G,sig\<turnstile>new hides old
419 \<longrightarrow> (accmodi old \<le> accmodi new \<and>
421 ))\<rbrakk> \<Longrightarrow> P
422 \<rbrakk> \<Longrightarrow> P"
423 by (unfold wf_cdecl_def) simp
425 lemma wf_cdecl_unique:
426 "wf_cdecl G (C,c) \<Longrightarrow> unique (cfields c) \<and> unique (methods c)"
427 apply (unfold wf_cdecl_def)
431 lemma wf_cdecl_fdecl:
432 "\<lbrakk>wf_cdecl G (C,c); f\<in>set (cfields c)\<rbrakk> \<Longrightarrow> wf_fdecl G (pid C) f"
433 apply (unfold wf_cdecl_def)
437 lemma wf_cdecl_mdecl:
438 "\<lbrakk>wf_cdecl G (C,c); m\<in>set (methods c)\<rbrakk> \<Longrightarrow> wf_mdecl G C m"
439 apply (unfold wf_cdecl_def)
444 "\<lbrakk>wf_cdecl G (C,c); I\<in>set (superIfs c)\<rbrakk>
445 \<Longrightarrow> is_acc_iface G (pid C) I \<and>
446 (\<forall>s. \<forall>im \<in> imethds G I s.
447 (\<exists>cm \<in> methd G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and> \<not>is_static cm \<and>
448 accmodi im \<le> accmodi cm))"
449 apply (unfold wf_cdecl_def)
454 "\<lbrakk>wf_cdecl G (C,c); C \<noteq> Object\<rbrakk> \<Longrightarrow>
455 is_acc_class G (pid C) (super c) \<and> (super c,C) \<notin> (subcls1 G)^+ \<and>
456 (table_of (map (\<lambda> (s,m). (s,C,m)) (methods c))
457 entails (\<lambda> new. \<forall> old sig.
458 (G,sig\<turnstile>new overrides\<^sub>S old
459 \<longrightarrow> (G\<turnstile>resTy new\<preceq>resTy old \<and>
460 accmodi old \<le> accmodi new \<and>
461 \<not>is_static old)) \<and>
462 (G,sig\<turnstile>new hides old
463 \<longrightarrow> (accmodi old \<le> accmodi new \<and>
465 apply (unfold wf_cdecl_def ws_cdecl_def)
470 lemma wf_cdecl_overrides_SomeD:
471 "\<lbrakk>wf_cdecl G (C,c); C \<noteq> Object; table_of (methods c) sig = Some newM;
472 G,sig\<turnstile>(C,newM) overrides\<^sub>S old
473 \<rbrakk> \<Longrightarrow> G\<turnstile>resTy newM\<preceq>resTy old \<and>
474 accmodi old \<le> accmodi newM \<and>
475 \<not> is_static old"
476 apply (drule (1) wf_cdecl_supD)
478 apply (drule entailsD)
479 apply (blast intro: table_of_map_SomeI)
480 apply (drule_tac x="old" in spec)
481 apply (auto dest: overrides_eq_sigD simp add: msig_def)
484 lemma wf_cdecl_hides_SomeD:
485 "\<lbrakk>wf_cdecl G (C,c); C \<noteq> Object; table_of (methods c) sig = Some newM;
486 G,sig\<turnstile>(C,newM) hides old
487 \<rbrakk> \<Longrightarrow> accmodi old \<le> access newM \<and>
489 apply (drule (1) wf_cdecl_supD)
491 apply (drule entailsD)
492 apply (blast intro: table_of_map_SomeI)
493 apply (drule_tac x="old" in spec)
494 apply (auto dest: hides_eq_sigD simp add: msig_def)
497 lemma wf_cdecl_wt_init:
498 "wf_cdecl G (C, c) \<Longrightarrow> \<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>init c\<Colon>\<surd>"
499 apply (unfold wf_cdecl_def)
504 section "well-formed programs"
505 (* well-formed program, cf. 8.1, 9.1 *)
508 A program declaration is wellformed if:
510 \item the class ObjectC of Object is defined
511 \item every method of Object has an access modifier distinct from Package.
513 necessary since every interface automatically inherits from Object.
514 We must know, that every time a Object method is "overriden" by an
515 interface method this is also overriden by the class implementing the
516 the interface (see @{text "implement_dynmethd and class_mheadsD"})
517 \item all standard Exceptions are defined
518 \item all defined interfaces are wellformed
519 \item all defined classes are wellformed
523 wf_prog :: "prog \<Rightarrow> bool"
524 "wf_prog G \<equiv> let is = ifaces G; cs = classes G in
525 ObjectC \<in> set cs \<and>
526 (\<forall> m\<in>set Object_mdecls. accmodi m \<noteq> Package) \<and>
527 (\<forall>xn. SXcptC xn \<in> set cs) \<and>
528 (\<forall>i\<in>set is. wf_idecl G i) \<and> unique is \<and>
529 (\<forall>c\<in>set cs. wf_cdecl G c) \<and> unique cs"
531 lemma wf_prog_idecl: "\<lbrakk>iface G I = Some i; wf_prog G\<rbrakk> \<Longrightarrow> wf_idecl G (I,i)"
532 apply (unfold wf_prog_def Let_def)
534 apply (fast dest: map_of_SomeD)
537 lemma wf_prog_cdecl: "\<lbrakk>class G C = Some c; wf_prog G\<rbrakk> \<Longrightarrow> wf_cdecl G (C,c)"
538 apply (unfold wf_prog_def Let_def)
540 apply (fast dest: map_of_SomeD)
543 lemma wf_prog_Object_mdecls:
544 "wf_prog G \<Longrightarrow> (\<forall> m\<in>set Object_mdecls. accmodi m \<noteq> Package)"
545 apply (unfold wf_prog_def Let_def)
549 lemma wf_prog_acc_superD:
550 "\<lbrakk>wf_prog G; class G C = Some c; C \<noteq> Object \<rbrakk>
551 \<Longrightarrow> is_acc_class G (pid C) (super c)"
552 by (auto dest: wf_prog_cdecl wf_cdecl_supD)
554 lemma wf_ws_prog [elim!,simp]: "wf_prog G \<Longrightarrow> ws_prog G"
555 apply (unfold wf_prog_def Let_def)
556 apply (rule ws_progI)
557 apply (simp_all (no_asm))
558 apply (auto simp add: is_acc_class_def is_acc_iface_def
559 dest!: wf_idecl_supD wf_cdecl_supD )+
562 lemma class_Object [simp]:
563 "wf_prog G \<Longrightarrow>
564 class G Object = Some \<lparr>access=Public,cfields=[],methods=Object_mdecls,
565 init=Skip,super=undefined,superIfs=[]\<rparr>"
566 apply (unfold wf_prog_def Let_def ObjectC_def)
567 apply (fast dest!: map_of_SomeI)
570 lemma methd_Object[simp]: "wf_prog G \<Longrightarrow> methd G Object =
571 table_of (map (\<lambda>(s,m). (s, Object, m)) Object_mdecls)"
572 apply (subst methd_rec)
573 apply (auto simp add: Let_def)
576 lemma wf_prog_Object_methd:
577 "\<lbrakk>wf_prog G; methd G Object sig = Some m\<rbrakk> \<Longrightarrow> accmodi m \<noteq> Package"
578 by (auto dest!: wf_prog_Object_mdecls) (auto dest!: map_of_SomeD)
580 lemma wf_prog_Object_is_public[intro]:
581 "wf_prog G \<Longrightarrow> is_public G Object"
582 by (auto simp add: is_public_def dest: class_Object)
584 lemma class_SXcpt [simp]:
585 "wf_prog G \<Longrightarrow>
586 class G (SXcpt xn) = Some \<lparr>access=Public,cfields=[],methods=SXcpt_mdecls,
588 super=if xn = Throwable then Object
589 else SXcpt Throwable,
591 apply (unfold wf_prog_def Let_def SXcptC_def)
592 apply (fast dest!: map_of_SomeI)
595 lemma wf_ObjectC [simp]:
596 "wf_cdecl G ObjectC = (\<not>is_iface G Object \<and> Ball (set Object_mdecls)
597 (wf_mdecl G Object) \<and> unique Object_mdecls)"
598 apply (unfold wf_cdecl_def ws_cdecl_def ObjectC_def)
599 apply (auto intro: da.Skip)
602 lemma Object_is_class [simp,elim!]: "wf_prog G \<Longrightarrow> is_class G Object"
603 apply (simp (no_asm_simp))
606 lemma Object_is_acc_class [simp,elim!]: "wf_prog G \<Longrightarrow> is_acc_class G S Object"
607 apply (simp (no_asm_simp) add: is_acc_class_def is_public_def
608 accessible_in_RefT_simp)
611 lemma SXcpt_is_class [simp,elim!]: "wf_prog G \<Longrightarrow> is_class G (SXcpt xn)"
612 apply (simp (no_asm_simp))
615 lemma SXcpt_is_acc_class [simp,elim!]:
616 "wf_prog G \<Longrightarrow> is_acc_class G S (SXcpt xn)"
617 apply (simp (no_asm_simp) add: is_acc_class_def is_public_def
618 accessible_in_RefT_simp)
621 lemma fields_Object [simp]: "wf_prog G \<Longrightarrow> DeclConcepts.fields G Object = []"
622 by (force intro: fields_emptyI)
624 lemma accfield_Object [simp]:
625 "wf_prog G \<Longrightarrow> accfield G S Object = empty"
626 apply (unfold accfield_def)
627 apply (simp (no_asm_simp) add: Let_def)
630 lemma fields_Throwable [simp]:
631 "wf_prog G \<Longrightarrow> DeclConcepts.fields G (SXcpt Throwable) = []"
632 by (force intro: fields_emptyI)
634 lemma fields_SXcpt [simp]: "wf_prog G \<Longrightarrow> DeclConcepts.fields G (SXcpt xn) = []"
635 apply (case_tac "xn = Throwable")
636 apply (simp (no_asm_simp))
637 by (force intro: fields_emptyI)
639 lemmas widen_trans = ws_widen_trans [OF _ _ wf_ws_prog, elim]
640 lemma widen_trans2 [elim]: "\<lbrakk>G\<turnstile>U\<preceq>T; G\<turnstile>S\<preceq>U; wf_prog G\<rbrakk> \<Longrightarrow> G\<turnstile>S\<preceq>T"
641 apply (erule (2) widen_trans)
644 lemma Xcpt_subcls_Throwable [simp]:
645 "wf_prog G \<Longrightarrow> G\<turnstile>SXcpt xn\<preceq>\<^sub>C SXcpt Throwable"
646 apply (rule SXcpt_subcls_Throwable_lemma)
651 "\<lbrakk>is_class G C; wf_prog G\<rbrakk> \<Longrightarrow> unique (DeclConcepts.fields G C)"
652 apply (erule ws_unique_fields)
653 apply (erule wf_ws_prog)
654 apply (erule (1) wf_prog_cdecl [THEN wf_cdecl_unique [THEN conjunct1]])
658 "\<lbrakk>table_of (DeclConcepts.fields G C) fn = Some f; G\<turnstile>D\<preceq>\<^sub>C C;
659 is_class G D; wf_prog G\<rbrakk>
660 \<Longrightarrow> table_of (DeclConcepts.fields G D) fn = Some f"
661 apply (rule map_of_SomeI)
662 apply (erule (1) unique_fields)
663 apply (erule (1) map_of_SomeD [THEN fields_mono_lemma])
664 apply (erule wf_ws_prog)
668 lemma fields_is_type [elim]:
669 "\<lbrakk>table_of (DeclConcepts.fields G C) m = Some f; wf_prog G; is_class G C\<rbrakk> \<Longrightarrow>
671 apply (frule wf_ws_prog)
672 apply (force dest: fields_declC [THEN conjunct1]
673 wf_prog_cdecl [THEN wf_cdecl_fdecl]
674 simp add: wf_fdecl_def2 is_acc_type_def)
677 lemma imethds_wf_mhead [rule_format (no_asm)]:
678 "\<lbrakk>m \<in> imethds G I sig; wf_prog G; is_iface G I\<rbrakk> \<Longrightarrow>
679 wf_mhead G (pid (decliface m)) sig (mthd m) \<and>
680 \<not> is_static m \<and> accmodi m = Public"
681 apply (frule wf_ws_prog)
682 apply (drule (2) imethds_declI [THEN conjunct1])
684 apply (frule_tac I="(decliface m)" in wf_prog_idecl,assumption)
685 apply (drule wf_idecl_mhead)
686 apply (erule map_of_SomeD)
687 apply (cases m, simp)
690 lemma methd_wf_mdecl:
691 "\<lbrakk>methd G C sig = Some m; wf_prog G; class G C = Some y\<rbrakk> \<Longrightarrow>
692 G\<turnstile>C\<preceq>\<^sub>C (declclass m) \<and> is_class G (declclass m) \<and>
693 wf_mdecl G (declclass m) (sig,(mthd m))"
694 apply (frule wf_ws_prog)
695 apply (drule (1) methd_declC)
698 apply (frule (1) wf_prog_cdecl, erule wf_cdecl_mdecl, erule map_of_SomeD)
702 This lemma doesn't hold!
703 lemma methd_rT_is_acc_type:
704 "\<lbrakk>wf_prog G;methd G C C sig = Some (D,m);
705 class G C = Some y\<rbrakk>
706 \<Longrightarrow> is_acc_type G (pid C) (resTy m)"
707 The result Type is only visible in the scope of defining class D
708 "is_vis_type G (pid D) (resTy m)" but not necessarily in scope of class C!
709 (The same is true for the type of pramaters of a method)
713 lemma methd_rT_is_type:
714 "\<lbrakk>wf_prog G;methd G C sig = Some m;
715 class G C = Some y\<rbrakk>
716 \<Longrightarrow> is_type G (resTy m)"
717 apply (drule (2) methd_wf_mdecl)
719 apply (drule wf_mdeclD1)
721 apply (drule rT_is_acc_type)
722 apply (cases m, simp add: is_acc_type_def)
725 lemma accmethd_rT_is_type:
726 "\<lbrakk>wf_prog G;accmethd G S C sig = Some m;
727 class G C = Some y\<rbrakk>
728 \<Longrightarrow> is_type G (resTy m)"
729 by (auto simp add: accmethd_def
730 intro: methd_rT_is_type)
732 lemma methd_Object_SomeD:
733 "\<lbrakk>wf_prog G;methd G Object sig = Some m\<rbrakk>
734 \<Longrightarrow> declclass m = Object"
735 by (auto dest: class_Object simp add: methd_rec )
738 "\<lbrakk>im \<in> imethds G I sig;wf_prog G; is_iface G I\<rbrakk>
739 \<Longrightarrow> \<not>is_static im \<and> accmodi im = Public"
741 assume asm: "wf_prog G" "is_iface G I" "im \<in> imethds G I sig"
742 have "wf_prog G \<longrightarrow>
743 (\<forall> i im. iface G I = Some i \<longrightarrow> im \<in> imethds G I sig
744 \<longrightarrow> \<not>is_static im \<and> accmodi im = Public)" (is "?P G I")
745 proof (rule iface_rec.induct,intro allI impI)
747 assume hyp: "\<forall> J i. J \<in> set (isuperIfs i) \<and> ws_prog G \<and> iface G I = Some i
748 \<longrightarrow> ?P G J"
749 assume wf: "wf_prog G" and if_I: "iface G I = Some i" and
750 im: "im \<in> imethds G I sig"
751 show "\<not>is_static im \<and> accmodi im = Public"
753 let ?inherited = "Un_tables (imethds G ` set (isuperIfs i))"
754 let ?new = "(o2s \<circ> table_of (map (\<lambda>(s, mh). (s, I, mh)) (imethods i)))"
755 from if_I wf im have imethds:"im \<in> (?inherited \<oplus>\<oplus> ?new) sig"
756 by (simp add: imethds_rec)
758 wf_supI: "\<forall> J. J \<in> set (isuperIfs i) \<longrightarrow> (\<exists> j. iface G J = Some j)"
759 by (blast dest: wf_prog_idecl wf_idecl_supD is_acc_ifaceD)
761 "\<forall> im \<in> set (imethods i). \<not> is_static im \<and> accmodi im = Public"
762 by (auto dest!: wf_prog_idecl wf_idecl_mhead)
763 then have new_ok: "\<forall> im. table_of (imethods i) sig = Some im
764 \<longrightarrow> \<not> is_static im \<and> accmodi im = Public"
765 by (auto dest!: table_of_Some_in_set)
767 proof (cases "?new sig = {}")
769 from True wf wf_supI if_I imethds hyp
770 show ?thesis by (auto simp del: split_paired_All)
773 from False wf wf_supI if_I imethds new_ok hyp
774 show ?thesis by (auto dest: wf_idecl_hidings hidings_entailsD)
778 with asm show ?thesis by (auto simp del: split_paired_All)
781 lemma wf_prog_hidesD:
782 assumes hides: "G \<turnstile>new hides old" and wf: "wf_prog G"
784 "accmodi old \<le> accmodi new \<and>
789 clsNew: "class G (declclass new) = Some c" and
790 neqObj: "declclass new \<noteq> Object"
791 by (auto dest: hidesD declared_in_classD)
792 with hides obtain newM oldM where
793 newM: "table_of (methods c) (msig new) = Some newM" and
794 new: "new = (declclass new,(msig new),newM)" and
795 old: "old = (declclass old,(msig old),oldM)" and
796 "msig new = msig old"
797 by (cases new,cases old)
799 simp add: cdeclaredmethd_def declared_in_def)
802 "G,(msig new)\<turnstile>(declclass new,newM) hides (declclass old,oldM)"
805 have "wf_cdecl G (declclass new,c)" by (blast intro: wf_prog_cdecl)
806 note wf_cdecl_hides_SomeD [OF this neqObj newM hides']
809 by (cases new, cases old) auto
812 text {* Compare this lemma about static
813 overriding @{term "G \<turnstile>new overrides\<^sub>S old"} with the definition of
814 dynamic overriding @{term "G \<turnstile>new overrides old"}.
815 Conforming result types and restrictions on the access modifiers of the old
816 and the new method are not part of the predicate for static overriding. But
817 they are enshured in a wellfromed program. Dynamic overriding has
818 no restrictions on the access modifiers but enforces confrom result types
819 as precondition. But with some efford we can guarantee the access modifier
820 restriction for dynamic overriding, too. See lemma
821 @{text wf_prog_dyn_override_prop}.
823 lemma wf_prog_stat_overridesD:
824 assumes stat_override: "G \<turnstile>new overrides\<^sub>S old" and wf: "wf_prog G"
826 "G\<turnstile>resTy new\<preceq>resTy old \<and>
827 accmodi old \<le> accmodi new \<and>
828 \<not> is_static old"
832 clsNew: "class G (declclass new) = Some c" and
833 neqObj: "declclass new \<noteq> Object"
834 by (auto dest: stat_overrides_commonD declared_in_classD)
835 with stat_override obtain newM oldM where
836 newM: "table_of (methods c) (msig new) = Some newM" and
837 new: "new = (declclass new,(msig new),newM)" and
838 old: "old = (declclass old,(msig old),oldM)" and
839 "msig new = msig old"
840 by (cases new,cases old)
841 (auto dest: stat_overrides_commonD
842 simp add: cdeclaredmethd_def declared_in_def)
845 "G,(msig new)\<turnstile>(declclass new,newM) overrides\<^sub>S (declclass old,oldM)"
848 have "wf_cdecl G (declclass new,c)" by (blast intro: wf_prog_cdecl)
849 note wf_cdecl_overrides_SomeD [OF this neqObj newM stat_override']
852 by (cases new, cases old) auto
855 lemma static_to_dynamic_overriding:
856 assumes stat_override: "G\<turnstile>new overrides\<^sub>S old" and wf : "wf_prog G"
857 shows "G\<turnstile>new overrides old"
860 show ?thesis (is "?Overrides new old")
862 case (Direct new old superNew)
863 then have stat_override:"G\<turnstile>new overrides\<^sub>S old"
864 by (rule stat_overridesR.Direct)
865 from stat_override wf
866 have resTy_widen: "G\<turnstile>resTy new\<preceq>resTy old" and
867 not_static_old: "\<not> is_static old"
868 by (auto dest: wf_prog_stat_overridesD)
869 have not_private_new: "accmodi new \<noteq> Private"
872 have "accmodi old \<noteq> Private"
873 by (rule no_Private_stat_override)
875 from stat_override wf
876 have "accmodi old \<le> accmodi new"
877 by (auto dest: wf_prog_stat_overridesD)
880 by (auto dest: acc_modi_bottom)
882 with Direct resTy_widen not_static_old
883 show "?Overrides new old"
884 by (auto intro: overridesR.Direct stat_override_declclasses_relation)
886 case (Indirect new inter old)
887 then show "?Overrides new old"
888 by (blast intro: overridesR.Indirect)
892 lemma non_Package_instance_method_inheritance:
893 assumes old_inheritable: "G\<turnstile>Method old inheritable_in (pid C)" and
894 accmodi_old: "accmodi old \<noteq> Package" and
895 instance_method: "\<not> is_static old" and
896 subcls: "G\<turnstile>C \<prec>\<^sub>C declclass old" and
897 old_declared: "G\<turnstile>Method old declared_in (declclass old)" and
899 shows "G\<turnstile>Method old member_of C \<or>
900 (\<exists> new. G\<turnstile> new overrides\<^sub>S old \<and> G\<turnstile>Method new member_of C)"
902 from wf have ws: "ws_prog G" by auto
903 from old_declared have iscls_declC_old: "is_class G (declclass old)"
904 by (auto simp add: declared_in_def cdeclaredmethd_def)
905 from subcls have iscls_C: "is_class G C"
906 by (blast dest: subcls_is_class)
907 from iscls_C ws old_inheritable subcls
908 show ?thesis (is "?P C old")
909 proof (induct rule: ws_class_induct')
911 assume "G\<turnstile>Object\<prec>\<^sub>C declclass old"
912 then show "?P Object old"
916 assume cls_C: "class G C = Some c" and
917 neq_C_Obj: "C \<noteq> Object" and
918 hyp: "\<lbrakk>G \<turnstile>Method old inheritable_in pid (super c);
919 G\<turnstile>super c\<prec>\<^sub>C declclass old\<rbrakk> \<Longrightarrow> ?P (super c) old" and
920 inheritable: "G \<turnstile>Method old inheritable_in pid C" and
921 subclsC: "G\<turnstile>C\<prec>\<^sub>C declclass old"
923 have super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c"
925 from wf cls_C neq_C_Obj
926 have accessible_super: "G\<turnstile>(Class (super c)) accessible_in (pid C)"
927 by (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD)
930 assume member_super: "G\<turnstile>Method old member_of (super c)"
931 assume inheritable: "G \<turnstile>Method old inheritable_in pid C"
932 assume instance_method: "\<not> is_static old"
934 have old_declared: "G\<turnstile>Method old declared_in (declclass old)"
935 by (cases old) (auto dest: member_of_declC)
937 proof (cases "G\<turnstile>mid (msig old) undeclared_in C")
939 with inheritable super accessible_super member_super
940 have "G\<turnstile>Method old member_of C"
941 by (cases old) (auto intro: members.Inherited)
946 then obtain new_member where
947 "G\<turnstile>new_member declared_in C" and
948 "mid (msig old) = memberid new_member"
949 by (auto dest: not_undeclared_declared)
950 then obtain new where
951 new: "G\<turnstile>Method new declared_in C" and
952 eq_sig: "msig old = msig new" and
953 declC_new: "declclass new = C"
954 by (cases new_member) auto
955 then have member_new: "G\<turnstile>Method new member_of C"
956 by (cases new) (auto intro: members.Immediate)
957 from declC_new super member_super
958 have subcls_new_old: "G\<turnstile>declclass new \<prec>\<^sub>C declclass old"
959 by (auto dest!: member_of_subclseq_declC
960 dest: r_into_trancl intro: trancl_rtrancl_trancl)
962 proof (cases "is_static new")
964 with eq_sig declC_new new old_declared inheritable
965 super member_super subcls_new_old
966 have "G\<turnstile>new overrides\<^sub>S old"
967 by (auto intro!: stat_overridesR.Direct)
968 with member_new show ?thesis
972 with eq_sig declC_new subcls_new_old new old_declared inheritable
973 have "G\<turnstile>new hides old"
974 by (auto intro: hidesI)
977 by (blast dest: wf_prog_hidesD)
983 } note hyp_member_super = this
985 have "G\<turnstile>(super c)\<preceq>\<^sub>C declclass old"
986 by (rule subcls_superD)
989 proof (cases rule: subclseq_cases)
991 assume "super c = declclass old"
993 have "G\<turnstile>Method old member_of (super c)"
994 by (cases old) (auto intro: members.Immediate)
995 with inheritable instance_method
997 by (blast dest: hyp_member_super)
1000 assume "G\<turnstile>super c\<prec>\<^sub>C declclass old"
1002 from inheritable accmodi_old
1003 have "G \<turnstile>Method old inheritable_in pid (super c)"
1004 by (cases "accmodi old") (auto simp add: inheritable_in_def)
1006 have "?P (super c) old"
1007 by (blast dest: hyp)
1010 assume "G \<turnstile>Method old member_of super c"
1011 with inheritable instance_method
1013 by (blast dest: hyp_member_super)
1015 assume "\<exists>new. G \<turnstile> new overrides\<^sub>S old \<and> G \<turnstile>Method new member_of super c"
1016 then obtain super_new where
1017 super_new_override: "G \<turnstile> super_new overrides\<^sub>S old" and
1018 super_new_member: "G \<turnstile>Method super_new member_of super c"
1020 from super_new_override wf
1021 have "accmodi old \<le> accmodi super_new"
1022 by (auto dest: wf_prog_stat_overridesD)
1023 with inheritable accmodi_old
1024 have "G \<turnstile>Method super_new inheritable_in pid C"
1025 by (auto simp add: inheritable_in_def
1026 split: acc_modi.splits
1027 dest: acc_modi_le_Dests)
1029 from super_new_override
1030 have "\<not> is_static super_new"
1031 by (auto dest: stat_overrides_commonD)
1033 note super_new_member
1034 ultimately have "?P C super_new"
1035 by (auto dest: hyp_member_super)
1038 assume "G \<turnstile>Method super_new member_of C"
1039 with super_new_override
1043 assume "\<exists>new. G \<turnstile> new overrides\<^sub>S super_new \<and>
1044 G \<turnstile>Method new member_of C"
1045 with super_new_override show ?thesis
1046 by (blast intro: stat_overridesR.Indirect)
1053 lemma non_Package_instance_method_inheritance_cases [consumes 6,
1054 case_names Inheritance Overriding]:
1055 assumes old_inheritable: "G\<turnstile>Method old inheritable_in (pid C)" and
1056 accmodi_old: "accmodi old \<noteq> Package" and
1057 instance_method: "\<not> is_static old" and
1058 subcls: "G\<turnstile>C \<prec>\<^sub>C declclass old" and
1059 old_declared: "G\<turnstile>Method old declared_in (declclass old)" and
1061 inheritance: "G\<turnstile>Method old member_of C \<Longrightarrow> P" and
1062 overriding: "\<And> new.
1063 \<lbrakk>G\<turnstile> new overrides\<^sub>S old;G\<turnstile>Method new member_of C\<rbrakk>
1064 \<Longrightarrow> P"
1067 from old_inheritable accmodi_old instance_method subcls old_declared wf
1068 inheritance overriding
1070 by (auto dest: non_Package_instance_method_inheritance)
1073 lemma dynamic_to_static_overriding:
1074 assumes dyn_override: "G\<turnstile> new overrides old" and
1075 accmodi_old: "accmodi old \<noteq> Package" and
1077 shows "G\<turnstile> new overrides\<^sub>S old"
1079 from dyn_override accmodi_old
1080 show ?thesis (is "?Overrides new old")
1081 proof (induct rule: overridesR.induct)
1082 case (Direct new old)
1083 assume new_declared: "G\<turnstile>Method new declared_in declclass new"
1084 assume eq_sig_new_old: "msig new = msig old"
1085 assume subcls_new_old: "G\<turnstile>declclass new \<prec>\<^sub>C declclass old"
1086 assume "G \<turnstile>Method old inheritable_in pid (declclass new)" and
1087 "accmodi old \<noteq> Package" and
1088 "\<not> is_static old" and
1089 "G\<turnstile>declclass new\<prec>\<^sub>C declclass old" and
1090 "G\<turnstile>Method old declared_in declclass old"
1092 show "?Overrides new old"
1093 proof (cases rule: non_Package_instance_method_inheritance_cases)
1095 assume "G \<turnstile>Method old member_of declclass new"
1096 then have "G\<turnstile>mid (msig old) undeclared_in declclass new"
1099 with subcls_new_old wf show ?thesis
1100 by (auto dest: subcls_irrefl)
1106 with eq_sig_new_old new_declared
1108 by (cases old,cases new) (auto dest!: declared_not_undeclared)
1110 case (Overriding new')
1111 assume stat_override_new': "G \<turnstile> new' overrides\<^sub>S old"
1112 then have "msig new' = msig old"
1113 by (auto dest: stat_overrides_commonD)
1114 with eq_sig_new_old have eq_sig_new_new': "msig new=msig new'"
1116 assume "G \<turnstile>Method new' member_of declclass new"
1120 then have declC_new: "declclass new' = declclass new"
1123 have "G\<turnstile>Method new' declared_in declclass new"
1124 by (cases new') auto
1125 with new_declared eq_sig_new_new' declC_new
1127 by (cases new, cases new') (auto dest: unique_declared_in)
1128 with stat_override_new'
1133 then have "G\<turnstile>mid (msig new') undeclared_in declclass new"
1134 by (cases new') (auto)
1135 with eq_sig_new_new' new_declared
1137 by (cases new,cases new') (auto dest!: declared_not_undeclared)
1141 case (Indirect new inter old)
1142 assume accmodi_old: "accmodi old \<noteq> Package"
1143 assume "accmodi old \<noteq> Package \<Longrightarrow> G \<turnstile> inter overrides\<^sub>S old"
1145 have stat_override_inter_old: "G \<turnstile> inter overrides\<^sub>S old"
1148 assume hyp_inter: "accmodi inter \<noteq> Package \<Longrightarrow> G \<turnstile> new overrides\<^sub>S inter"
1150 have "accmodi inter \<noteq> Package"
1152 from stat_override_inter_old wf
1153 have "accmodi old \<le> accmodi inter"
1154 by (auto dest: wf_prog_stat_overridesD)
1155 with stat_override_inter_old accmodi_old
1157 by (auto dest!: no_Private_stat_override
1158 split: acc_modi.splits
1159 dest: acc_modi_le_Dests)
1161 ultimately show "?Overrides new old"
1162 by (blast intro: stat_overridesR.Indirect)
1166 lemma wf_prog_dyn_override_prop:
1167 assumes dyn_override: "G \<turnstile> new overrides old" and
1169 shows "accmodi old \<le> accmodi new"
1170 proof (cases "accmodi old = Package")
1172 note old_Package = this
1174 proof (cases "accmodi old \<le> accmodi new")
1175 case True then show ?thesis .
1179 have "accmodi new = Private"
1180 by (cases "accmodi new") (auto simp add: le_acc_def less_acc_def)
1183 by (auto dest: overrides_commonD)
1187 with dyn_override wf
1188 have "G \<turnstile> new overrides\<^sub>S old"
1189 by (blast intro: dynamic_to_static_overriding)
1192 by (blast dest: wf_prog_stat_overridesD)
1195 lemma overrides_Package_old:
1196 assumes dyn_override: "G \<turnstile> new overrides old" and
1197 accmodi_new: "accmodi new = Package" and
1199 shows "accmodi old = Package"
1200 proof (cases "accmodi old")
1202 with dyn_override show ?thesis
1203 by (simp add: no_Private_override)
1209 with dyn_override wf
1210 have "G \<turnstile> new overrides\<^sub>S old"
1211 by (auto intro: dynamic_to_static_overriding)
1213 have "accmodi old \<le> accmodi new"
1214 by (auto dest: wf_prog_stat_overridesD)
1215 with Protected accmodi_new
1217 by (simp add: less_acc_def le_acc_def)
1220 with dyn_override wf
1221 have "G \<turnstile> new overrides\<^sub>S old"
1222 by (auto intro: dynamic_to_static_overriding)
1224 have "accmodi old \<le> accmodi new"
1225 by (auto dest: wf_prog_stat_overridesD)
1226 with Public accmodi_new
1228 by (simp add: less_acc_def le_acc_def)
1231 lemma dyn_override_Package:
1232 assumes dyn_override: "G \<turnstile> new overrides old" and
1233 accmodi_old: "accmodi old = Package" and
1234 accmodi_new: "accmodi new = Package" and
1236 shows "pid (declclass old) = pid (declclass new)"
1238 from dyn_override accmodi_old accmodi_new
1239 show ?thesis (is "?EqPid old new")
1240 proof (induct rule: overridesR.induct)
1241 case (Direct new old)
1242 assume "accmodi old = Package"
1243 "G \<turnstile>Method old inheritable_in pid (declclass new)"
1244 then show "pid (declclass old) = pid (declclass new)"
1245 by (auto simp add: inheritable_in_def)
1247 case (Indirect new inter old)
1248 assume accmodi_old: "accmodi old = Package" and
1249 accmodi_new: "accmodi new = Package"
1250 assume "G \<turnstile> new overrides inter"
1252 have "accmodi inter = Package"
1253 by (auto intro: overrides_Package_old)
1255 show "pid (declclass old) = pid (declclass new)"
1260 lemma dyn_override_Package_escape:
1261 assumes dyn_override: "G \<turnstile> new overrides old" and
1262 accmodi_old: "accmodi old = Package" and
1263 outside_pack: "pid (declclass old) \<noteq> pid (declclass new)" and
1265 shows "\<exists> inter. G \<turnstile> new overrides inter \<and> G \<turnstile> inter overrides old \<and>
1266 pid (declclass old) = pid (declclass inter) \<and>
1267 Protected \<le> accmodi inter"
1269 from dyn_override accmodi_old outside_pack
1270 show ?thesis (is "?P new old")
1271 proof (induct rule: overridesR.induct)
1272 case (Direct new old)
1273 assume accmodi_old: "accmodi old = Package"
1274 assume outside_pack: "pid (declclass old) \<noteq> pid (declclass new)"
1275 assume "G \<turnstile>Method old inheritable_in pid (declclass new)"
1277 have "pid (declclass old) = pid (declclass new)"
1278 by (simp add: inheritable_in_def)
1283 case (Indirect new inter old)
1284 assume accmodi_old: "accmodi old = Package"
1285 assume outside_pack: "pid (declclass old) \<noteq> pid (declclass new)"
1286 assume override_new_inter: "G \<turnstile> new overrides inter"
1287 assume override_inter_old: "G \<turnstile> inter overrides old"
1288 assume hyp_new_inter: "\<lbrakk>accmodi inter = Package;
1289 pid (declclass inter) \<noteq> pid (declclass new)\<rbrakk>
1290 \<Longrightarrow> ?P new inter"
1291 assume hyp_inter_old: "\<lbrakk>accmodi old = Package;
1292 pid (declclass old) \<noteq> pid (declclass inter)\<rbrakk>
1293 \<Longrightarrow> ?P inter old"
1295 proof (cases "pid (declclass old) = pid (declclass inter)")
1297 note same_pack_old_inter = this
1299 proof (cases "pid (declclass inter) = pid (declclass new)")
1301 with same_pack_old_inter outside_pack
1306 note diff_pack_inter_new = this
1308 proof (cases "accmodi inter = Package")
1310 with diff_pack_inter_new hyp_new_inter
1311 obtain newinter where
1312 over_new_newinter: "G \<turnstile> new overrides newinter" and
1313 over_newinter_inter: "G \<turnstile> newinter overrides inter" and
1314 eq_pid: "pid (declclass inter) = pid (declclass newinter)" and
1315 accmodi_newinter: "Protected \<le> accmodi newinter"
1317 from over_newinter_inter override_inter_old
1318 have "G\<turnstile>newinter overrides old"
1319 by (rule overridesR.Indirect)
1321 from eq_pid same_pack_old_inter
1322 have "pid (declclass old) = pid (declclass newinter)"
1325 note over_new_newinter accmodi_newinter
1326 ultimately show ?thesis
1330 with override_new_inter
1331 have "Protected \<le> accmodi inter"
1332 by (cases "accmodi inter") (auto dest: no_Private_override)
1333 with override_new_inter override_inter_old same_pack_old_inter
1340 with accmodi_old hyp_inter_old
1341 obtain newinter where
1342 over_inter_newinter: "G \<turnstile> inter overrides newinter" and
1343 over_newinter_old: "G \<turnstile> newinter overrides old" and
1344 eq_pid: "pid (declclass old) = pid (declclass newinter)" and
1345 accmodi_newinter: "Protected \<le> accmodi newinter"
1347 from override_new_inter over_inter_newinter
1348 have "G \<turnstile> new overrides newinter"
1349 by (rule overridesR.Indirect)
1350 with eq_pid over_newinter_old accmodi_newinter
1357 lemma declclass_widen[rule_format]:
1359 \<longrightarrow> (\<forall>c m. class G C = Some c \<longrightarrow> methd G C sig = Some m
1360 \<longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m)" (is "?P G C")
1361 proof (rule class_rec.induct,intro allI impI)
1363 assume Hyp: "\<forall>c. C \<noteq> Object \<and> ws_prog G \<and> class G C = Some c
1364 \<longrightarrow> ?P G (super c)"
1365 assume wf: "wf_prog G" and cls_C: "class G C = Some c" and
1366 m: "methd G C sig = Some m"
1367 show "G\<turnstile>C\<preceq>\<^sub>C declclass m"
1368 proof (cases "C=Object")
1370 with wf m show ?thesis by (simp add: methd_Object_SomeD)
1372 let ?filter="filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m)"
1373 let ?table = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c))"
1376 have methd_C: "(?filter (methd G (super c)) ++ ?table) sig = Some m "
1377 by (simp add: methd_rec)
1379 proof (cases "?table sig")
1381 from this methd_C have "?filter (methd G (super c)) sig = Some m"
1384 from wf cls_C False obtain sup where "class G (super c) = Some sup"
1385 by (blast dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
1386 moreover note wf False cls_C
1387 ultimately have "G\<turnstile>super c \<preceq>\<^sub>C declclass m"
1388 by (auto intro: Hyp [rule_format])
1389 moreover from cls_C False have "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c" by (rule subcls1I)
1390 ultimately show ?thesis by - (rule rtrancl_into_rtrancl2)
1393 from this wf False cls_C methd_C show ?thesis by auto
1398 lemma declclass_methd_Object:
1399 "\<lbrakk>wf_prog G; methd G Object sig = Some m\<rbrakk> \<Longrightarrow> declclass m = Object"
1402 lemma methd_declaredD:
1403 "\<lbrakk>wf_prog G; is_class G C;methd G C sig = Some m\<rbrakk>
1404 \<Longrightarrow> G\<turnstile>(mdecl (sig,mthd m)) declared_in (declclass m)"
1406 assume wf: "wf_prog G"
1407 then have ws: "ws_prog G" ..
1408 assume clsC: "is_class G C"
1410 show "methd G C sig = Some m
1411 \<Longrightarrow> G\<turnstile>(mdecl (sig,mthd m)) declared_in (declclass m)"
1413 proof (induct ?P C rule: ws_class_induct')
1415 assume "methd G Object sig = Some m"
1416 with wf show ?thesis
1417 by - (rule method_declared_inI, auto)
1421 assume clsC: "class G C = Some c"
1422 and m: "methd G C sig = Some m"
1423 and hyp: "methd G (super c) sig = Some m \<Longrightarrow> ?thesis"
1424 let ?newMethods = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c))"
1426 proof (cases "?newMethods sig")
1428 from None ws clsC m hyp
1429 show ?thesis by (auto intro: method_declared_inI simp add: methd_rec)
1433 show ?thesis by (auto intro: method_declared_inI simp add: methd_rec)
1438 lemma methd_rec_Some_cases [consumes 4, case_names NewMethod InheritedMethod]:
1439 assumes methd_C: "methd G C sig = Some m" and
1441 clsC: "class G C = Some c" and
1442 neq_C_Obj: "C\<noteq>Object"
1444 "\<lbrakk>table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig = Some m \<Longrightarrow> P;
1445 \<lbrakk>G\<turnstile>C inherits (method sig m); methd G (super c) sig = Some m\<rbrakk> \<Longrightarrow> P
1446 \<rbrakk> \<Longrightarrow> P"
1448 let ?inherited = "filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m)
1449 (methd G (super c))"
1450 let ?new = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c))"
1451 from ws clsC neq_C_Obj methd_C
1452 have methd_unfold: "(?inherited ++ ?new) sig = Some m"
1453 by (simp add: methd_rec)
1454 assume NewMethod: "?new sig = Some m \<Longrightarrow> P"
1455 assume InheritedMethod: "\<lbrakk>G\<turnstile>C inherits (method sig m);
1456 methd G (super c) sig = Some m\<rbrakk> \<Longrightarrow> P"
1458 proof (cases "?new sig")
1460 with methd_unfold have "?inherited sig = Some m"
1462 with InheritedMethod show P by blast
1465 with methd_unfold have "?new sig = Some m"
1467 with NewMethod show P by blast
1472 lemma methd_member_of:
1473 assumes wf: "wf_prog G"
1475 "\<lbrakk>is_class G C; methd G C sig = Some m\<rbrakk> \<Longrightarrow> G\<turnstile>Methd sig m member_of C"
1476 (is "?Class C \<Longrightarrow> ?Method C \<Longrightarrow> ?MemberOf C")
1478 from wf have ws: "ws_prog G" ..
1479 assume defC: "is_class G C"
1481 show "?Class C \<Longrightarrow> ?Method C \<Longrightarrow> ?MemberOf C"
1482 proof (induct rule: ws_class_induct')
1484 with wf have declC: "Object = declclass m"
1485 by (simp add: declclass_methd_Object)
1486 from Object wf have "G\<turnstile>Methd sig m declared_in Object"
1487 by (auto intro: methd_declaredD simp add: declC)
1489 show "?MemberOf Object"
1490 by (auto intro!: members.Immediate
1491 simp del: methd_Object)
1494 assume clsC: "class G C = Some c" and
1495 neq_C_Obj: "C \<noteq> Object"
1496 assume methd: "?Method C"
1497 from methd ws clsC neq_C_Obj
1499 proof (cases rule: methd_rec_Some_cases)
1501 with clsC show ?thesis
1502 by (auto dest: method_declared_inI intro!: members.Immediate)
1504 case InheritedMethod
1506 by (blast dest: inherits_member)
1511 lemma current_methd:
1512 "\<lbrakk>table_of (methods c) sig = Some new;
1513 ws_prog G; class G C = Some c; C \<noteq> Object;
1514 methd G (super c) sig = Some old\<rbrakk>
1515 \<Longrightarrow> methd G C sig = Some (C,new)"
1516 by (auto simp add: methd_rec
1517 intro: filter_tab_SomeI map_add_find_right table_of_map_SomeI)
1519 lemma wf_prog_staticD:
1520 assumes wf: "wf_prog G" and
1521 clsC: "class G C = Some c" and
1522 neq_C_Obj: "C \<noteq> Object" and
1523 old: "methd G (super c) sig = Some old" and
1524 accmodi_old: "Protected \<le> accmodi old" and
1525 new: "table_of (methods c) sig = Some new"
1526 shows "is_static new = is_static old"
1529 have wf_cdecl: "wf_cdecl G (C,c)" by (rule wf_prog_cdecl)
1530 from wf clsC neq_C_Obj
1531 have is_cls_super: "is_class G (super c)"
1532 by (blast dest: wf_prog_acc_superD is_acc_classD)
1533 from wf is_cls_super old
1534 have old_member_of: "G\<turnstile>Methd sig old member_of (super c)"
1535 by (rule methd_member_of)
1536 from old wf is_cls_super
1537 have old_declared: "G\<turnstile>Methd sig old declared_in (declclass old)"
1538 by (auto dest: methd_declared_in_declclass)
1540 have new_declared: "G\<turnstile>Methd sig (C,new) declared_in C"
1541 by (auto intro: method_declared_inI)
1542 note trancl_rtrancl_tranc = trancl_rtrancl_trancl [trans] (* ### in Basis *)
1544 have subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c"
1546 then have "G\<turnstile>C \<prec>\<^sub>C super c" ..
1547 also from old wf is_cls_super
1548 have "G\<turnstile>super c \<preceq>\<^sub>C (declclass old)" by (auto dest: methd_declC)
1549 finally have subcls_C_old: "G\<turnstile>C \<prec>\<^sub>C (declclass old)" .
1551 have inheritable: "G\<turnstile>Methd sig old inheritable_in pid C"
1552 by (auto simp add: inheritable_in_def
1553 dest: acc_modi_le_Dests)
1555 proof (cases "is_static new")
1557 with subcls_C_old new_declared old_declared inheritable
1558 have "G,sig\<turnstile>(C,new) hides old"
1559 by (auto intro: hidesI)
1560 with True wf_cdecl neq_C_Obj new
1562 by (auto dest: wf_cdecl_hides_SomeD)
1565 with subcls_C_old new_declared old_declared inheritable subcls1_C_super
1567 have "G,sig\<turnstile>(C,new) overrides\<^sub>S old"
1568 by (auto intro: stat_overridesR.Direct)
1569 with False wf_cdecl neq_C_Obj new
1571 by (auto dest: wf_cdecl_overrides_SomeD)
1575 lemma inheritable_instance_methd:
1576 assumes subclseq_C_D: "G\<turnstile>C \<preceq>\<^sub>C D" and
1577 is_cls_D: "is_class G D" and
1579 old: "methd G D sig = Some old" and
1580 accmodi_old: "Protected \<le> accmodi old" and
1581 not_static_old: "\<not> is_static old"
1583 "\<exists>new. methd G C sig = Some new \<and>
1584 (new = old \<or> G,sig\<turnstile>new overrides\<^sub>S old)"
1585 (is "(\<exists>new. (?Constraint C new old))")
1587 from subclseq_C_D is_cls_D
1588 have is_cls_C: "is_class G C" by (rule subcls_is_class2)
1590 have ws: "ws_prog G" ..
1591 from is_cls_C ws subclseq_C_D
1592 show "\<exists>new. ?Constraint C new old"
1593 proof (induct rule: ws_class_induct')
1595 then have eq_D_Obj: "D=Object" by auto
1597 have "?Constraint Object old old"
1600 show "\<exists> new. ?Constraint Object new old" by auto
1603 assume hyp: "G\<turnstile>super c\<preceq>\<^sub>C D \<Longrightarrow> \<exists>new. ?Constraint (super c) new old"
1604 assume clsC: "class G C = Some c"
1605 assume neq_C_Obj: "C\<noteq>Object"
1607 have wf_cdecl: "wf_cdecl G (C,c)"
1608 by (rule wf_prog_cdecl)
1609 from ws clsC neq_C_Obj
1610 have is_cls_super: "is_class G (super c)"
1611 by (auto dest: ws_prog_cdeclD)
1612 from clsC wf neq_C_Obj
1613 have superAccessible: "G\<turnstile>(Class (super c)) accessible_in (pid C)" and
1614 subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c"
1615 by (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD
1617 show "\<exists>new. ?Constraint C new old"
1618 proof (cases "G\<turnstile>super c\<preceq>\<^sub>C D")
1622 by (auto dest: subclseq_superD)
1624 have "?Constraint C old old"
1627 show "\<exists> new. ?Constraint C new old" by auto
1630 with hyp obtain super_method
1631 where super: "?Constraint (super c) super_method old" by blast
1632 from super not_static_old
1633 have not_static_super: "\<not> is_static super_method"
1634 by (auto dest!: stat_overrides_commonD)
1635 from super old wf accmodi_old
1636 have accmodi_super_method: "Protected \<le> accmodi super_method"
1637 by (auto dest!: wf_prog_stat_overridesD)
1638 from super accmodi_old wf
1639 have inheritable: "G\<turnstile>Methd sig super_method inheritable_in (pid C)"
1640 by (auto dest!: wf_prog_stat_overridesD
1642 simp add: inheritable_in_def)
1643 from super wf is_cls_super
1644 have member: "G\<turnstile>Methd sig super_method member_of (super c)"
1645 by (auto intro: methd_member_of)
1647 have decl_super_method:
1648 "G\<turnstile>Methd sig super_method declared_in (declclass super_method)"
1649 by (auto dest: member_of_declC)
1650 from super subcls1_C_super ws is_cls_super
1651 have subcls_C_super: "G\<turnstile>C \<prec>\<^sub>C (declclass super_method)"
1652 by (auto intro: rtrancl_into_trancl2 dest: methd_declC)
1653 show "\<exists> new. ?Constraint C new old"
1654 proof (cases "methd G C sig")
1656 have "methd G (super c) sig = None"
1659 have no_new: "table_of (methods c) sig = None"
1660 by (auto simp add: methd_rec)
1662 have undeclared: "G\<turnstile>mid sig undeclared_in C"
1663 by (auto simp add: undeclared_in_def cdeclaredmethd_def)
1664 with inheritable member superAccessible subcls1_C_super
1665 have inherits: "G\<turnstile>C inherits (method sig super_method)"
1666 by (auto simp add: inherits_def)
1667 with clsC ws no_new super neq_C_Obj
1668 have "methd G C sig = Some super_method"
1669 by (auto simp add: methd_rec map_add_def intro: filter_tab_SomeI)
1670 with None show ?thesis
1673 with super show ?thesis by auto
1676 from this ws clsC neq_C_Obj
1678 proof (cases rule: methd_rec_Some_cases)
1679 case InheritedMethod
1680 with super Some show ?thesis
1684 assume new: "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig
1687 have declcls_new: "declclass new = C"
1689 from wf clsC neq_C_Obj super new not_static_super accmodi_super_method
1690 have not_static_new: "\<not> is_static new"
1691 by (auto dest: wf_prog_staticD)
1693 have decl_new: "G\<turnstile>Methd sig new declared_in C"
1694 by (auto simp add: declared_in_def cdeclaredmethd_def)
1695 from not_static_new decl_new decl_super_method
1696 member subcls1_C_super inheritable declcls_new subcls_C_super
1697 have "G,sig\<turnstile> new overrides\<^sub>S super_method"
1698 by (auto intro: stat_overridesR.Direct)
1701 by (auto intro: stat_overridesR.Indirect)
1708 lemma inheritable_instance_methd_cases [consumes 6
1709 , case_names Inheritance Overriding]:
1710 assumes subclseq_C_D: "G\<turnstile>C \<preceq>\<^sub>C D" and
1711 is_cls_D: "is_class G D" and
1713 old: "methd G D sig = Some old" and
1714 accmodi_old: "Protected \<le> accmodi old" and
1715 not_static_old: "\<not> is_static old" and
1716 inheritance: "methd G C sig = Some old \<Longrightarrow> P" and
1717 overriding: "\<And> new. \<lbrakk>methd G C sig = Some new;
1718 G,sig\<turnstile>new overrides\<^sub>S old\<rbrakk> \<Longrightarrow> P"
1722 from subclseq_C_D is_cls_D wf old accmodi_old not_static_old
1724 by (auto dest: inheritable_instance_methd intro: inheritance overriding)
1727 lemma inheritable_instance_methd_props:
1728 assumes subclseq_C_D: "G\<turnstile>C \<preceq>\<^sub>C D" and
1729 is_cls_D: "is_class G D" and
1731 old: "methd G D sig = Some old" and
1732 accmodi_old: "Protected \<le> accmodi old" and
1733 not_static_old: "\<not> is_static old"
1735 "\<exists>new. methd G C sig = Some new \<and>
1736 \<not> is_static new \<and> G\<turnstile>resTy new\<preceq>resTy old \<and> accmodi old \<le>accmodi new"
1737 (is "(\<exists>new. (?Constraint C new old))")
1739 from subclseq_C_D is_cls_D wf old accmodi_old not_static_old
1741 proof (cases rule: inheritable_instance_methd_cases)
1743 with not_static_old accmodi_old show ?thesis by auto
1745 case (Overriding new)
1746 then have "\<not> is_static new" by (auto dest: stat_overrides_commonD)
1747 with Overriding not_static_old accmodi_old wf
1749 by (auto dest!: wf_prog_stat_overridesD)
1754 lemma bexI': "x \<in> A \<Longrightarrow> P x \<Longrightarrow> \<exists>x\<in>A. P x" by blast
1755 lemma ballE': "\<forall>x\<in>A. P x \<Longrightarrow> (x \<notin> A \<Longrightarrow> Q) \<Longrightarrow> (P x \<Longrightarrow> Q) \<Longrightarrow> Q" by blast
1757 lemma subint_widen_imethds:
1758 "\<lbrakk>G\<turnstile>I\<preceq>I J; wf_prog G; is_iface G J; jm \<in> imethds G J sig\<rbrakk> \<Longrightarrow>
1759 \<exists> im \<in> imethds G I sig. is_static im = is_static jm \<and>
1760 accmodi im = accmodi jm \<and>
1761 G\<turnstile>resTy im\<preceq>resTy jm"
1763 assume irel: "G\<turnstile>I\<preceq>I J" and
1765 is_iface: "is_iface G J"
1766 from irel show "jm \<in> imethds G J sig \<Longrightarrow> ?thesis"
1767 (is "PROP ?P I" is "PROP ?Prem J \<Longrightarrow> ?Concl I")
1768 proof (induct ?P I rule: converse_rtrancl_induct)
1770 assume "jm \<in> imethds G J sig"
1771 then show "?Concl J" by (blast elim: bexI')
1775 assume subint1_I_SI: "G\<turnstile>I \<prec>I1 SI" and
1776 subint_SI_J: "G\<turnstile>SI \<preceq>I J" and
1777 hyp: "PROP ?P SI" and
1778 jm: "jm \<in> imethds G J sig"
1781 ifI: "iface G I = Some i" and
1782 SI: "SI \<in> set (isuperIfs i)"
1783 by (blast dest: subint1D)
1786 = "(o2s \<circ> table_of (map (\<lambda>(sig, mh). (sig, I, mh)) (imethods i)))"
1788 proof (cases "?newMethods sig = {}")
1790 with ifI SI hyp wf jm
1792 by (auto simp add: imethds_rec)
1796 have imethds: "imethds G I sig = ?newMethods sig"
1797 by (simp add: imethds_rec)
1800 imdef: "im \<in> ?newMethods sig"
1803 have im: "im \<in> imethds G I sig"
1807 imStatic: "\<not> is_static im" and
1808 imPublic: "accmodi im = Public"
1809 by (auto dest!: imethds_wf_mhead)
1811 have wf_I: "wf_idecl G (I,i)"
1812 by (rule wf_prog_idecl)
1815 ifSI: "iface G SI = Some si" and
1816 wf_SI: "wf_idecl G (SI,si)"
1817 by (auto dest!: wf_idecl_supD is_acc_ifaceD
1818 dest: wf_prog_idecl)
1820 obtain sim::"qtname \<times> mhead" where
1821 sim: "sim \<in> imethds G SI sig" and
1822 eq_static_sim_jm: "is_static sim = is_static jm" and
1823 eq_access_sim_jm: "accmodi sim = accmodi jm" and
1824 resTy_widen_sim_jm: "G\<turnstile>resTy sim\<preceq>resTy jm"
1826 with wf_I SI imdef sim
1827 have "G\<turnstile>resTy im\<preceq>resTy sim"
1828 by (auto dest!: wf_idecl_hidings hidings_entailsD)
1829 with wf resTy_widen_sim_jm
1830 have resTy_widen_im_jm: "G\<turnstile>resTy im\<preceq>resTy jm"
1831 by (blast intro: widen_trans)
1834 simStatic: "\<not> is_static sim" and
1835 simPublic: "accmodi sim = Public"
1836 by (auto dest!: imethds_wf_mhead)
1838 imStatic simStatic eq_static_sim_jm
1839 imPublic simPublic eq_access_sim_jm
1847 (* Tactical version *)
1849 lemma subint_widen_imethds: "\<lbrakk>G\<turnstile>I\<preceq>I J; wf_prog G; is_iface G J\<rbrakk> \<Longrightarrow>
1850 \<forall> jm \<in> imethds G J sig.
1851 \<exists> im \<in> imethds G I sig. static (mthd im)=static (mthd jm) \<and>
1852 access (mthd im)= access (mthd jm) \<and>
1853 G\<turnstile>resTy (mthd im)\<preceq>resTy (mthd jm)"
1854 apply (erule converse_rtrancl_induct)
1855 apply (clarsimp elim!: bexI')
1856 apply (frule subint1D)
1858 apply (erule ballE')
1860 apply (erule_tac V = "?x \<in> imethds G J sig" in thin_rl)
1862 apply (subst imethds_rec, assumption, erule wf_ws_prog)
1863 apply (unfold overrides_t_def)
1864 apply (drule (1) wf_prog_idecl)
1865 apply (frule (3) imethds_wf_mhead [OF _ _ wf_idecl_supD [THEN conjunct1
1866 [THEN is_acc_ifaceD [THEN conjunct1]]]])
1867 apply (case_tac "(o2s \<circ> table_of (map (\<lambda>(s, mh). (s, y, mh)) (imethods i)))
1874 apply (frule wf_idecl_hidings [THEN hidings_entailsD])
1879 apply (drule table_of_map_SomeI [of _ "sig"])
1882 apply (frule wf_idecl_mhead [of _ _ _ "sig"])
1883 apply (rule table_of_Some_in_set)
1891 lemma implmt1_methd:
1892 "\<And>sig. \<lbrakk>G\<turnstile>C\<leadsto>1I; wf_prog G; im \<in> imethds G I sig\<rbrakk> \<Longrightarrow>
1893 \<exists>cm \<in>methd G C sig: \<not> is_static cm \<and> \<not> is_static im \<and>
1894 G\<turnstile>resTy cm\<preceq>resTy im \<and>
1895 accmodi im = Public \<and> accmodi cm = Public"
1896 apply (drule implmt1D)
1898 apply (drule (2) wf_prog_cdecl [THEN wf_cdecl_impD])
1899 apply (frule (1) imethds_wf_mhead)
1900 apply (simp add: is_acc_iface_def)
1906 lemma implmt_methd [rule_format (no_asm)]:
1907 "\<lbrakk>wf_prog G; G\<turnstile>C\<leadsto>I\<rbrakk> \<Longrightarrow> is_iface G I \<longrightarrow>
1908 (\<forall> im \<in>imethds G I sig.
1909 \<exists> cm\<in>methd G C sig: \<not>is_static cm \<and> \<not> is_static im \<and>
1910 G\<turnstile>resTy cm\<preceq>resTy im \<and>
1911 accmodi im = Public \<and> accmodi cm = Public)"
1912 apply (frule implmt_is_class)
1913 apply (erule implmt.induct)
1915 apply (drule (2) implmt1_methd)
1917 apply (drule (1) subint_widen_imethds)
1921 apply (drule (2) implmt1_methd)
1923 apply (frule subcls1D)
1924 apply (drule (1) bspec)
1926 apply (drule (3) r_into_rtrancl [THEN inheritable_instance_methd_props,
1927 OF _ implmt_is_class])
1931 lemma mheadsD [rule_format (no_asm)]:
1932 "emh \<in> mheads G S t sig \<longrightarrow> wf_prog G \<longrightarrow>
1933 (\<exists>C D m. t = ClassT C \<and> declrefT emh = ClassT D \<and>
1934 accmethd G S C sig = Some m \<and>
1935 (declclass m = D) \<and> mhead (mthd m) = (mhd emh)) \<or>
1936 (\<exists>I. t = IfaceT I \<and> ((\<exists>im. im \<in> accimethds G (pid S) I sig \<and>
1937 mthd im = mhd emh) \<or>
1938 (\<exists>m. G\<turnstile>Iface I accessible_in (pid S) \<and> accmethd G S Object sig = Some m \<and>
1939 accmodi m \<noteq> Private \<and>
1940 declrefT emh = ClassT Object \<and> mhead (mthd m) = mhd emh))) \<or>
1941 (\<exists>T m. t = ArrayT T \<and> G\<turnstile>Array T accessible_in (pid S) \<and>
1942 accmethd G S Object sig = Some m \<and> accmodi m \<noteq> Private \<and>
1943 declrefT emh = ClassT Object \<and> mhead (mthd m) = mhd emh)"
1944 apply (rule_tac ref_ty1="t" in ref_ty_ty.induct [THEN conjunct1])
1946 apply (auto simp add: cmheads_def accObjectmheads_def Objectmheads_def)
1947 apply (auto dest!: accmethd_SomeD)
1950 lemma mheads_cases [consumes 2, case_names Class_methd
1951 Iface_methd Iface_Object_methd Array_Object_methd]:
1952 "\<lbrakk>emh \<in> mheads G S t sig; wf_prog G;
1953 \<And> C D m. \<lbrakk>t = ClassT C;declrefT emh = ClassT D; accmethd G S C sig = Some m;
1954 (declclass m = D); mhead (mthd m) = (mhd emh)\<rbrakk> \<Longrightarrow> P emh;
1955 \<And> I im. \<lbrakk>t = IfaceT I; im \<in> accimethds G (pid S) I sig; mthd im = mhd emh\<rbrakk>
1956 \<Longrightarrow> P emh;
1957 \<And> I m. \<lbrakk>t = IfaceT I; G\<turnstile>Iface I accessible_in (pid S);
1958 accmethd G S Object sig = Some m; accmodi m \<noteq> Private;
1959 declrefT emh = ClassT Object; mhead (mthd m) = mhd emh\<rbrakk> \<Longrightarrow> P emh;
1960 \<And> T m. \<lbrakk>t = ArrayT T;G\<turnstile>Array T accessible_in (pid S);
1961 accmethd G S Object sig = Some m; accmodi m \<noteq> Private;
1962 declrefT emh = ClassT Object; mhead (mthd m) = mhd emh\<rbrakk> \<Longrightarrow> P emh
1963 \<rbrakk> \<Longrightarrow> P emh"
1964 by (blast dest!: mheadsD)
1966 lemma declclassD[rule_format]:
1967 "\<lbrakk>wf_prog G;class G C = Some c; methd G C sig = Some m;
1968 class G (declclass m) = Some d\<rbrakk>
1969 \<Longrightarrow> table_of (methods d) sig = Some (mthd m)"
1971 assume wf: "wf_prog G"
1972 then have ws: "ws_prog G" ..
1973 assume clsC: "class G C = Some c"
1975 show "\<And> m d. \<lbrakk>methd G C sig = Some m; class G (declclass m) = Some d\<rbrakk>
1976 \<Longrightarrow> table_of (methods d) sig = Some (mthd m)"
1978 proof (induct ?P C rule: ws_class_induct)
1981 assume "methd G Object sig = Some m"
1982 "class G (declclass m) = Some d"
1983 with wf show "?thesis m d" by auto
1987 assume hyp: "PROP ?P (super c)"
1988 and m: "methd G C sig = Some m"
1989 and declC: "class G (declclass m) = Some d"
1990 and clsC: "class G C = Some c"
1991 and nObj: "C \<noteq> Object"
1992 let ?newMethods = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig"
1994 proof (cases "?newMethods")
1996 from None clsC nObj ws m declC
1997 show "?thesis" by (auto simp add: methd_rec) (rule hyp)
2000 from Some clsC nObj ws m declC
2002 by (auto simp add: methd_rec
2003 dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
2008 (* Tactical version *)
2010 lemma declclassD[rule_format]:
2011 "wf_prog G \<longrightarrow>
2012 (\<forall> c d m. class G C = Some c \<longrightarrow> methd G C sig = Some m \<longrightarrow>
2013 class G (declclass m) = Some d
2014 \<longrightarrow> table_of (methods d) sig = Some (mthd m))"
2015 apply (rule class_rec.induct)
2019 apply (case_tac "C=Object")
2020 apply (force simp add: methd_rec)
2022 apply (subst methd_rec)
2023 apply (blast dest: wf_ws_prog)+
2024 apply (case_tac "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig")
2025 apply (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
2029 lemma dynmethd_Object:
2030 assumes statM: "methd G Object sig = Some statM" and
2031 private: "accmodi statM = Private" and
2032 is_cls_C: "is_class G C" and
2034 shows "dynmethd G Object C sig = Some statM"
2037 have subclseq: "G\<turnstile>C \<preceq>\<^sub>C Object"
2038 by (auto intro: subcls_ObjectI)
2039 from wf have ws: "ws_prog G"
2042 have is_cls_Obj: "is_class G Object"
2044 from statM subclseq is_cls_Obj ws private
2046 proof (cases rule: dynmethd_cases)
2047 case Static then show ?thesis .
2050 with private show ?thesis
2051 by (auto dest: no_Private_override)
2055 lemma wf_imethds_hiding_objmethdsD:
2056 assumes old: "methd G Object sig = Some old" and
2057 is_if_I: "is_iface G I" and
2059 not_private: "accmodi old \<noteq> Private" and
2060 new: "new \<in> imethds G I sig"
2061 shows "G\<turnstile>resTy new\<preceq>resTy old \<and> is_static new = is_static old" (is "?P new")
2063 from wf have ws: "ws_prog G" by simp
2066 assume ifI: "iface G I = Some i"
2067 assume new: "table_of (imethods i) sig = Some new"
2068 from ifI new not_private wf old
2070 by (auto dest!: wf_prog_idecl wf_idecl_hiding cond_hiding_entailsD
2071 simp del: methd_Object)
2072 } note hyp_newmethod = this
2075 proof (induct rule: ws_interface_induct)
2077 assume ifI: "iface G I = Some i"
2078 assume new: "new \<in> imethds G I sig"
2080 have hyp: "\<forall> J \<in> set (isuperIfs i). (new \<in> imethds G J sig \<longrightarrow> ?P new)"
2084 proof (cases rule: imethds_cases)
2086 with ifI hyp_newmethod
2090 case (InheritedMethod J)
2091 assume "J \<in> set (isuperIfs i)"
2092 "new \<in> imethds G J sig"
2101 Which dynamic classes are valid to look up a member of a distinct static type?
2102 We have to distinct class members (named static members in Java)
2103 from instance members. Class members are global to all Objects of a class,
2104 instance members are local to a single Object instance. If a member is
2105 equipped with the static modifier it is a class member, else it is an
2107 The following table gives an overview of the current framework. We assume
2108 to have a reference with static type statT and a dynamic class dynC. Between
2109 both of these types the widening relation holds
2110 @{term "G\<turnstile>Class dynC\<preceq> statT"}. Unfortunately this ordinary widening relation
2111 isn't enough to describe the valid lookup classes, since we must cope the
2112 special cases of arrays and interfaces,too. If we statically expect an array or
2113 inteface we may lookup a field or a method in Object which isn't covered in
2114 the widening relation.
2116 statT field instance method static (class) method
2117 ------------------------------------------------------------------------
2120 Class dynC dynC dynC
2121 Array / Object Object
2123 In most cases we con lookup the member in the dynamic class. But as an
2124 interface can't declare new static methods, nor an array can define new
2125 methods at all, we have to lookup methods in the base class Object.
2127 The limitation to classes in the field column is artificial and comes out
2128 of the typing rule for the field access (see rule @{text "FVar"} in the
2129 welltyping relation @{term "wt"} in theory WellType).
2130 I stems out of the fact, that Object
2131 indeed has no non private fields. So interfaces and arrays can actually
2132 have no fields at all and a field access would be senseless. (In Java
2133 interfaces are allowed to declare new fields but in current Bali not!).
2134 So there is no principal reason why we should not allow Objects to declare
2135 non private fields. Then we would get the following column:
2144 consts valid_lookup_cls:: "prog \<Rightarrow> ref_ty \<Rightarrow> qtname \<Rightarrow> bool \<Rightarrow> bool"
2145 ("_,_ \<turnstile> _ valid'_lookup'_cls'_for _" [61,61,61,61] 60)
2147 "G,NullT \<turnstile> dynC valid_lookup_cls_for static_membr = False"
2148 "G,IfaceT I \<turnstile> dynC valid_lookup_cls_for static_membr
2151 else G\<turnstile>Class dynC\<preceq> Iface I)"
2152 "G,ClassT C \<turnstile> dynC valid_lookup_cls_for static_membr = G\<turnstile>Class dynC\<preceq> Class C"
2153 "G,ArrayT T \<turnstile> dynC valid_lookup_cls_for static_membr = (dynC=Object)"
2155 lemma valid_lookup_cls_is_class:
2156 assumes dynC: "G,statT \<turnstile> dynC valid_lookup_cls_for static_membr" and
2157 ty_statT: "isrtype G statT" and
2159 shows "is_class G dynC"
2162 with dynC ty_statT show ?thesis
2163 by (auto dest: widen_NT2)
2166 with dynC wf show ?thesis
2167 by (auto dest: implmt_is_class)
2170 with dynC ty_statT show ?thesis
2171 by (auto dest: subcls_is_class2)
2174 with dynC wf show ?thesis
2178 declare split_paired_All [simp del] split_paired_Ex [simp del]
2179 declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *}
2180 declaration {* K (Classical.map_cs (fn cs => cs delSWrapper "split_all_tac")) *}
2182 lemma dynamic_mheadsD:
2183 "\<lbrakk>emh \<in> mheads G S statT sig;
2184 G,statT \<turnstile> dynC valid_lookup_cls_for (is_static emh);
2185 isrtype G statT; wf_prog G
2186 \<rbrakk> \<Longrightarrow> \<exists>m \<in> dynlookup G statT dynC sig:
2187 is_static m=is_static emh \<and> G\<turnstile>resTy m\<preceq>resTy emh"
2189 assume emh: "emh \<in> mheads G S statT sig"
2191 and dynC_Prop: "G,statT \<turnstile> dynC valid_lookup_cls_for (is_static emh)"
2192 and istype: "isrtype G statT"
2193 from dynC_Prop istype wf
2195 dynC: "class G dynC = Some y"
2196 by (auto dest: valid_lookup_cls_is_class)
2197 from emh wf show ?thesis
2198 proof (cases rule: mheads_cases)
2200 fix statC statDeclC sm
2201 assume statC: "statT = ClassT statC"
2202 assume "accmethd G S statC sig = Some sm"
2203 then have sm: "methd G statC sig = Some sm"
2204 by (blast dest: accmethd_SomeD)
2205 assume eq_mheads: "mhead (mthd sm) = mhd emh"
2207 have dynlookup: "dynlookup G statT dynC sig = dynmethd G statC dynC sig"
2208 by (simp add: dynlookup_def)
2209 from wf statC istype dynC_Prop sm
2211 "dynmethd G statC dynC sig = Some dm"
2212 "is_static dm = is_static sm"
2213 "G\<turnstile>resTy dm\<preceq>resTy sm"
2214 by (force dest!: ws_dynmethd accmethd_SomeD)
2215 with dynlookup eq_mheads
2217 by (cases emh type: *) (auto)
2221 assume statI: "statT = IfaceT I" and
2222 eq_mheads: "mthd im = mhd emh" and
2223 "im \<in> accimethds G (pid S) I sig"
2224 then have im: "im \<in> imethds G I sig"
2225 by (blast dest: accimethdsD)
2226 with istype statI eq_mheads wf
2227 have not_static_emh: "\<not> is_static emh"
2228 by (cases emh) (auto dest: wf_prog_idecl imethds_wf_mhead)
2230 have dynlookup: "dynlookup G statT dynC sig = methd G dynC sig"
2231 by (auto simp add: dynlookup_def dynimethd_def)
2232 from wf dynC_Prop statI istype im not_static_emh
2234 "methd G dynC sig = Some dm"
2235 "is_static dm = is_static im"
2236 "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mthd im)"
2237 by (force dest: implmt_methd)
2238 with dynlookup eq_mheads
2240 by (cases emh type: *) (auto)
2242 case Iface_Object_methd
2244 assume statI: "statT = IfaceT I" and
2245 sm: "accmethd G S Object sig = Some sm" and
2246 eq_mheads: "mhead (mthd sm) = mhd emh" and
2247 nPriv: "accmodi sm \<noteq> Private"
2249 proof (cases "imethds G I sig = {}")
2252 have dynlookup: "dynlookup G statT dynC sig = dynmethd G Object dynC sig"
2253 by (simp add: dynlookup_def dynimethd_def)
2255 have subclsObj: "G\<turnstile>dynC \<preceq>\<^sub>C Object"
2256 by (auto intro: subcls_ObjectI)
2257 from wf dynC dynC_Prop istype sm subclsObj
2259 "dynmethd G Object dynC sig = Some dm"
2260 "is_static dm = is_static sm"
2261 "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mthd sm)"
2262 by (auto dest!: ws_dynmethd accmethd_SomeD
2263 intro: class_Object [OF wf] intro: that)
2264 with dynlookup eq_mheads
2266 by (cases emh type: *) (auto)
2270 have dynlookup: "dynlookup G statT dynC sig = methd G dynC sig"
2271 by (simp add: dynlookup_def dynimethd_def)
2275 with wf sm nPriv False
2277 im: "im \<in> imethds G I sig" and
2278 eq_stat: "is_static im = is_static sm" and
2279 resProp: "G\<turnstile>resTy (mthd im)\<preceq>resTy (mthd sm)"
2280 by (auto dest: wf_imethds_hiding_objmethdsD accmethd_SomeD)
2281 from im wf statI istype eq_stat eq_mheads
2282 have not_static_sm: "\<not> is_static emh"
2283 by (cases emh) (auto dest: wf_prog_idecl imethds_wf_mhead)
2284 from im wf dynC_Prop dynC istype statI not_static_sm
2286 "methd G dynC sig = Some dm"
2287 "is_static dm = is_static im"
2288 "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mthd im)"
2289 by (auto dest: implmt_methd)
2290 with wf eq_stat resProp dynlookup eq_mheads
2292 by (cases emh type: *) (auto intro: widen_trans)
2295 case Array_Object_methd
2297 assume statArr: "statT = ArrayT T" and
2298 sm: "accmethd G S Object sig = Some sm" and
2299 eq_mheads: "mhead (mthd sm) = mhd emh"
2300 from statArr dynC_Prop wf
2301 have dynlookup: "dynlookup G statT dynC sig = methd G Object sig"
2302 by (auto simp add: dynlookup_def dynmethd_C_C)
2303 with sm eq_mheads sm
2305 by (cases emh type: *) (auto dest: accmethd_SomeD)
2308 declare split_paired_All [simp] split_paired_Ex [simp]
2309 declaration {* K (Classical.map_cs (fn cs => cs addSbefore ("split_all_tac", split_all_tac))) *}
2310 declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *}
2312 (* Tactical version *)
2314 lemma dynamic_mheadsD: "
2315 \<lbrakk>emh \<in> mheads G S statT sig; wf_prog G; class G dynC = Some y;
2316 if (\<exists>T. statT=ArrayT T) then dynC=Object else G\<turnstile>Class dynC\<preceq>RefT statT;
2317 isrtype G statT\<rbrakk> \<Longrightarrow>
2318 \<exists>m \<in> dynlookup G statT dynC sig:
2319 static (mthd m)=static (mhd emh) \<and> G\<turnstile>resTy (mthd m)\<preceq>resTy (mhd emh)"
2320 apply (drule mheadsD)
2322 -- reftype statT is a class
2323 apply (case_tac "\<exists>T. ClassT C = ArrayT T")
2326 apply (clarsimp simp add: dynlookup_def )
2327 apply (frule_tac statC="C" and dynC="dynC" and sig="sig"
2330 apply (case_tac "emh")
2331 apply (force dest: accmethd_SomeD)
2333 -- reftype statT is a interface, method defined in interface
2334 apply (clarsimp simp add: dynlookup_def)
2335 apply (drule (1) implmt_methd)
2339 apply (unfold dynimethd_def)
2340 apply (rule_tac x="cm" in bexI)
2341 apply (case_tac "emh")
2346 -- reftype statT is a interface, method defined in Object
2347 apply (simp add: dynlookup_def)
2348 apply (simp only: dynimethd_def)
2349 apply (case_tac "imethds G I sig = {}")
2351 apply (frule_tac statC="Object" and dynC="dynC" and sig="sig"
2353 apply (blast intro: subcls_ObjectI wf_ws_prog)
2354 apply (blast dest: class_Object)
2355 apply (case_tac "emh")
2356 apply (force dest: accmethd_SomeD)
2359 apply (subgoal_tac "\<exists> im. im \<in> imethds G I sig")
2360 prefer 2 apply blast
2362 apply (frule (1) implmt_methd)
2365 apply (clarify dest!: accmethd_SomeD)
2366 apply (frule (4) iface_overrides_Object)
2368 apply (case_tac emh)
2371 -- reftype statT is a array
2372 apply (simp add: dynlookup_def)
2373 apply (case_tac emh)
2374 apply (force dest: accmethd_SomeD simp add: dynmethd_def)
2378 (* FIXME occasionally convert to ws_class_induct*)
2379 lemma methd_declclass:
2380 "\<lbrakk>class G C = Some c; wf_prog G; methd G C sig = Some m\<rbrakk>
2381 \<Longrightarrow> methd G (declclass m) sig = Some m"
2383 assume asm: "class G C = Some c" "wf_prog G" "methd G C sig = Some m"
2384 have "wf_prog G \<longrightarrow>
2385 (\<forall> c m. class G C = Some c \<longrightarrow> methd G C sig = Some m
2386 \<longrightarrow> methd G (declclass m) sig = Some m)" (is "?P G C")
2387 proof (rule class_rec.induct,intro allI impI)
2389 assume hyp: "\<forall>c. C \<noteq> Object \<and> ws_prog G \<and> class G C = Some c \<longrightarrow>
2391 assume wf: "wf_prog G" and cls_C: "class G C = Some c" and
2392 m: "methd G C sig = Some m"
2393 show "methd G (declclass m) sig = Some m"
2394 proof (cases "C=Object")
2396 with wf m show ?thesis by (auto intro: table_of_map_SomeI)
2398 let ?filter="filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m)"
2399 let ?table = "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c))"
2402 have methd_C: "(?filter (methd G (super c)) ++ ?table) sig = Some m "
2403 by (simp add: methd_rec)
2405 proof (cases "?table sig")
2407 from this methd_C have "?filter (methd G (super c)) sig = Some m"
2410 from wf cls_C False obtain sup where "class G (super c) = Some sup"
2411 by (blast dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
2412 moreover note wf False cls_C
2413 ultimately show ?thesis by (auto intro: hyp [rule_format])
2416 from this methd_C m show ?thesis by auto
2420 with asm show ?thesis by auto
2423 lemma dynmethd_declclass:
2424 "\<lbrakk>dynmethd G statC dynC sig = Some m;
2425 wf_prog G; is_class G statC
2426 \<rbrakk> \<Longrightarrow> methd G (declclass m) sig = Some m"
2427 by (auto dest: dynmethd_declC)
2429 lemma dynlookup_declC:
2430 "\<lbrakk>dynlookup G statT dynC sig = Some m; wf_prog G;
2431 is_class G dynC;isrtype G statT
2432 \<rbrakk> \<Longrightarrow> G\<turnstile>dynC \<preceq>\<^sub>C (declclass m) \<and> is_class G (declclass m)"
2434 (auto simp add: dynlookup_def dynimethd_def
2435 dest: methd_declC dynmethd_declC)
2437 lemma dynlookup_Array_declclassD [simp]:
2438 "\<lbrakk>dynlookup G (ArrayT T) Object sig = Some dm;wf_prog G\<rbrakk>
2439 \<Longrightarrow> declclass dm = Object"
2441 assume dynL: "dynlookup G (ArrayT T) Object sig = Some dm"
2442 assume wf: "wf_prog G"
2443 from wf have ws: "ws_prog G" by auto
2444 from wf have is_cls_Obj: "is_class G Object" by auto
2447 by (auto simp add: dynlookup_def dynmethd_C_C [OF is_cls_Obj ws]
2448 dest: methd_Object_SomeD)
2452 declare split_paired_All [simp del] split_paired_Ex [simp del]
2453 declaration {* K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac")) *}
2454 declaration {* K (Classical.map_cs (fn cs => cs delSWrapper "split_all_tac")) *}
2456 lemma wt_is_type: "E,dt\<Turnstile>v\<Colon>T \<Longrightarrow> wf_prog (prg E) \<longrightarrow>
2457 dt=empty_dt \<longrightarrow> (case T of
2458 Inl T \<Rightarrow> is_type (prg E) T
2459 | Inr Ts \<Rightarrow> Ball (set Ts) (is_type (prg E)))"
2460 apply (unfold empty_dt_def)
2461 apply (erule wt.induct)
2462 apply (auto split del: split_if_asm simp del: snd_conv
2463 simp add: is_acc_class_def is_acc_type_def)
2464 apply (erule typeof_empty_is_type)
2465 apply (frule (1) wf_prog_cdecl [THEN wf_cdecl_supD],
2466 force simp del: snd_conv, clarsimp simp add: is_acc_class_def)
2467 apply (drule (1) max_spec2mheads [THEN conjunct1, THEN mheadsD])
2468 apply (drule_tac [2] accfield_fields)
2469 apply (frule class_Object)
2470 apply (auto dest: accmethd_rT_is_type
2471 imethds_wf_mhead [THEN conjunct1, THEN rT_is_acc_type]
2473 simp del: class_Object
2474 simp add: is_acc_type_def
2477 declare split_paired_All [simp] split_paired_Ex [simp]
2478 declaration {* K (Classical.map_cs (fn cs => cs addSbefore ("split_all_tac", split_all_tac))) *}
2479 declaration {* K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac))) *}
2481 lemma ty_expr_is_type:
2482 "\<lbrakk>E\<turnstile>e\<Colon>-T; wf_prog (prg E)\<rbrakk> \<Longrightarrow> is_type (prg E) T"
2483 by (auto dest!: wt_is_type)
2484 lemma ty_var_is_type:
2485 "\<lbrakk>E\<turnstile>v\<Colon>=T; wf_prog (prg E)\<rbrakk> \<Longrightarrow> is_type (prg E) T"
2486 by (auto dest!: wt_is_type)
2487 lemma ty_exprs_is_type:
2488 "\<lbrakk>E\<turnstile>es\<Colon>\<doteq>Ts; wf_prog (prg E)\<rbrakk> \<Longrightarrow> Ball (set Ts) (is_type (prg E))"
2489 by (auto dest!: wt_is_type)
2492 lemma static_mheadsD:
2493 "\<lbrakk> emh \<in> mheads G S t sig; wf_prog G; E\<turnstile>e\<Colon>-RefT t; prg E=G ;
2494 invmode (mhd emh) e \<noteq> IntVir
2495 \<rbrakk> \<Longrightarrow> \<exists>m. ( (\<exists> C. t = ClassT C \<and> accmethd G S C sig = Some m)
2496 \<or> (\<forall> C. t \<noteq> ClassT C \<and> accmethd G S Object sig = Some m )) \<and>
2497 declrefT emh = ClassT (declclass m) \<and> mhead (mthd m) = (mhd emh)"
2498 apply (subgoal_tac "is_static emh \<or> e = Super")
2499 defer apply (force simp add: invmode_def)
2500 apply (frule ty_expr_is_type)
2502 apply (case_tac "is_static emh")
2503 apply (frule (1) mheadsD)
2507 apply (auto dest!: imethds_wf_mhead
2510 simp add: accObjectmheads_def Objectmheads_def)
2512 apply (erule wt_elim_cases)
2513 apply (force simp add: cmheads_def)
2517 "\<lbrakk>methd G C sig = Some m; wf_prog G;
2518 class G C = Some c\<rbrakk> \<Longrightarrow>
2519 \<exists>T. \<lparr>prg=G,cls=(declclass m),
2520 lcl=callee_lcl (declclass m) sig (mthd m)\<rparr>\<turnstile> Methd C sig\<Colon>-T \<and> G\<turnstile>T\<preceq>resTy m"
2521 apply (frule (2) methd_wf_mdecl, clarify)
2522 apply (force dest!: wf_mdecl_bodyD intro!: wt.Methd)
2525 subsection "accessibility concerns"
2527 lemma mheads_type_accessible:
2528 "\<lbrakk>emh \<in> mheads G S T sig; wf_prog G\<rbrakk>
2529 \<Longrightarrow> G\<turnstile>RefT T accessible_in (pid S)"
2530 by (erule mheads_cases)
2531 (auto dest: accmethd_SomeD accessible_from_commonD accimethdsD)
2533 lemma static_to_dynamic_accessible_from_aux:
2534 "\<lbrakk>G\<turnstile>m of C accessible_from accC;wf_prog G\<rbrakk>
2535 \<Longrightarrow> G\<turnstile>m in C dyn_accessible_from accC"
2536 proof (induct rule: accessible_fromR.induct)
2537 qed (auto intro: dyn_accessible_fromR.intros
2538 member_of_to_member_in
2539 static_to_dynamic_overriding)
2541 lemma static_to_dynamic_accessible_from:
2542 assumes stat_acc: "G\<turnstile>m of statC accessible_from accC" and
2543 subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
2545 shows "G\<turnstile>m in dynC dyn_accessible_from accC"
2547 from stat_acc subclseq
2548 show ?thesis (is "?Dyn_accessible m")
2549 proof (induct rule: accessible_fromR.induct)
2550 case (Immediate m statC)
2551 then show "?Dyn_accessible m"
2552 by (blast intro: dyn_accessible_fromR.Immediate
2554 permits_acc_inheritance)
2556 case (Overriding m _ _)
2557 with wf show "?Dyn_accessible m"
2558 by (blast intro: dyn_accessible_fromR.Overriding
2560 static_to_dynamic_overriding
2561 rtrancl_trancl_trancl
2562 static_to_dynamic_accessible_from_aux)
2566 lemma static_to_dynamic_accessible_from_static:
2567 assumes stat_acc: "G\<turnstile>m of statC accessible_from accC" and
2568 static: "is_static m" and
2570 shows "G\<turnstile>m in (declclass m) dyn_accessible_from accC"
2573 have "G\<turnstile>m in statC dyn_accessible_from accC"
2574 by (auto intro: static_to_dynamic_accessible_from)
2577 by (rule dyn_accessible_from_static_declC)
2580 lemma dynmethd_member_in:
2581 assumes m: "dynmethd G statC dynC sig = Some m" and
2582 iscls_statC: "is_class G statC" and
2584 shows "G\<turnstile>Methd sig m member_in dynC"
2587 have subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC"
2588 by (auto simp add: dynmethd_def)
2589 from subclseq iscls_statC
2590 have iscls_dynC: "is_class G dynC"
2591 by (rule subcls_is_class2)
2592 from iscls_dynC iscls_statC wf m
2593 have "G\<turnstile>dynC \<preceq>\<^sub>C (declclass m) \<and> is_class G (declclass m) \<and>
2594 methd G (declclass m) sig = Some m"
2595 by - (drule dynmethd_declC, auto)
2598 by (auto intro: member_inI dest: methd_member_of)
2601 lemma dynmethd_access_prop:
2602 assumes statM: "methd G statC sig = Some statM" and
2603 stat_acc: "G\<turnstile>Methd sig statM of statC accessible_from accC" and
2604 dynM: "dynmethd G statC dynC sig = Some dynM" and
2606 shows "G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
2608 from wf have ws: "ws_prog G" ..
2610 have subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC"
2611 by (auto simp add: dynmethd_def)
2613 have is_cls_statC: "is_class G statC"
2614 by (auto dest: accessible_from_commonD member_of_is_classD)
2616 have is_cls_dynC: "is_class G dynC"
2617 by (rule subcls_is_class2)
2618 from is_cls_statC statM wf
2619 have member_statC: "G\<turnstile>Methd sig statM member_of statC"
2620 by (auto intro: methd_member_of)
2622 have statC_acc: "G\<turnstile>Class statC accessible_in (pid accC)"
2623 by (auto dest: accessible_from_commonD)
2624 from statM subclseq is_cls_statC ws
2626 proof (cases rule: dynmethd_cases)
2628 assume dynmethd: "dynmethd G statC dynC sig = Some statM"
2629 with dynM have eq_dynM_statM: "dynM=statM"
2631 with stat_acc subclseq wf
2633 by (auto intro: static_to_dynamic_accessible_from)
2635 case (Overrides newM)
2636 assume dynmethd: "dynmethd G statC dynC sig = Some newM"
2637 assume override: "G,sig\<turnstile>newM overrides statM"
2638 assume neq: "newM\<noteq>statM"
2640 have eq_dynM_newM: "dynM=newM"
2642 from dynmethd eq_dynM_newM wf is_cls_statC
2643 have "G\<turnstile>Methd sig dynM member_in dynC"
2644 by (auto intro: dynmethd_member_in)
2647 have "G\<turnstile>dynC\<prec>\<^sub>C statC"
2648 proof (cases rule: subclseq_cases)
2652 from is_cls_statC obtain c
2653 where "class G statC = Some c"
2656 note statM ws dynmethd
2659 by (auto simp add: dynmethd_C_C)
2660 with neq show ?thesis
2663 case Subcls then show ?thesis .
2667 have "G\<turnstile>Methd sig statM in statC dyn_accessible_from accC"
2668 by (blast intro: static_to_dynamic_accessible_from)
2670 note override eq_dynM_newM
2671 ultimately show ?thesis
2672 by (cases dynM,cases statM) (auto intro: dyn_accessible_fromR.Overriding)
2676 lemma implmt_methd_access:
2678 assumes iface_methd: "imethds G I sig \<noteq> {}" and
2679 implements: "G\<turnstile>dynC\<leadsto>I" and
2680 isif_I: "is_iface G I" and
2682 shows "\<exists> dynM. methd G dynC sig = Some dynM \<and>
2683 G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
2686 have iscls_dynC: "is_class G dynC" by (rule implmt_is_class)
2689 where "im \<in> imethds G I sig"
2691 with wf implements isif_I
2693 where dynM: "methd G dynC sig = Some dynM" and
2694 pub: "accmodi dynM = Public"
2695 by (blast dest: implmt_methd)
2697 have "G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
2698 by (auto intro!: dyn_accessible_fromR.Immediate
2699 intro: methd_member_of member_of_to_member_in
2700 simp add: permits_acc_def)
2706 corollary implmt_dynimethd_access:
2708 assumes iface_methd: "imethds G I sig \<noteq> {}" and
2709 implements: "G\<turnstile>dynC\<leadsto>I" and
2710 isif_I: "is_iface G I" and
2712 shows "\<exists> dynM. dynimethd G I dynC sig = Some dynM \<and>
2713 G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
2716 have "dynimethd G I dynC sig = methd G dynC sig"
2717 by (simp add: dynimethd_def)
2718 with iface_methd implements isif_I wf
2721 (blast intro: implmt_methd_access)
2724 lemma dynlookup_access_prop:
2725 assumes emh: "emh \<in> mheads G accC statT sig" and
2726 dynM: "dynlookup G statT dynC sig = Some dynM" and
2727 dynC_prop: "G,statT \<turnstile> dynC valid_lookup_cls_for is_static emh" and
2728 isT_statT: "isrtype G statT" and
2730 shows "G \<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
2733 have statT_acc: "G\<turnstile>RefT statT accessible_in (pid accC)"
2734 by (rule mheads_type_accessible)
2735 from dynC_prop isT_statT wf
2736 have iscls_dynC: "is_class G dynC"
2737 by (rule valid_lookup_cls_is_class)
2738 from emh dynC_prop isT_statT wf dynM
2739 have eq_static: "is_static emh = is_static dynM"
2740 by (auto dest: dynamic_mheadsD)
2741 from emh wf show ?thesis
2742 proof (cases rule: mheads_cases)
2743 case (Class_methd statC _ statM)
2744 assume statT: "statT = ClassT statC"
2745 assume "accmethd G accC statC sig = Some statM"
2746 then have statM: "methd G statC sig = Some statM" and
2747 stat_acc: "G\<turnstile>Methd sig statM of statC accessible_from accC"
2748 by (auto dest: accmethd_SomeD)
2750 have dynM': "dynmethd G statC dynC sig = Some dynM"
2751 by (simp add: dynlookup_def)
2752 from statM stat_acc wf dynM'
2754 by (auto dest!: dynmethd_access_prop)
2756 case (Iface_methd I im)
2757 then have iface_methd: "imethds G I sig \<noteq> {}" and
2758 statT_acc: "G\<turnstile>RefT statT accessible_in (pid accC)"
2759 by (auto dest: accimethdsD)
2760 assume statT: "statT = IfaceT I"
2761 assume im: "im \<in> accimethds G (pid accC) I sig"
2762 assume eq_mhds: "mthd im = mhd emh"
2764 have dynM': "dynimethd G I dynC sig = Some dynM"
2765 by (simp add: dynlookup_def)
2766 from isT_statT statT
2767 have isif_I: "is_iface G I"
2770 proof (cases "is_static emh")
2772 with statT dynC_prop
2773 have widen_dynC: "G\<turnstile>Class dynC \<preceq> RefT statT"
2775 from statT widen_dynC
2776 have implmnt: "G\<turnstile>dynC\<leadsto>I"
2778 from eq_static False
2779 have not_static_dynM: "\<not> is_static dynM"
2781 from iface_methd implmnt isif_I wf dynM'
2783 by - (drule implmt_dynimethd_access, auto)
2786 assume "is_static emh"
2788 from wf isT_statT statT im
2789 have "\<not> is_static im"
2790 by (auto dest: accimethdsD wf_prog_idecl imethds_wf_mhead)
2791 moreover note eq_mhds
2792 ultimately show ?thesis
2796 case (Iface_Object_methd I statM)
2797 assume statT: "statT = IfaceT I"
2798 assume "accmethd G accC Object sig = Some statM"
2799 then have statM: "methd G Object sig = Some statM" and
2800 stat_acc: "G\<turnstile>Methd sig statM of Object accessible_from accC"
2801 by (auto dest: accmethd_SomeD)
2802 assume not_Private_statM: "accmodi statM \<noteq> Private"
2803 assume eq_mhds: "mhead (mthd statM) = mhd emh"
2805 have widen_dynC_Obj: "G\<turnstile>dynC \<preceq>\<^sub>C Object"
2806 by (auto intro: subcls_ObjectI)
2808 proof (cases "imethds G I sig = {}")
2810 from dynM statT True
2811 have dynM': "dynmethd G Object dynC sig = Some dynM"
2812 by (simp add: dynlookup_def dynimethd_def)
2814 have "G\<turnstile>RefT statT \<preceq>Class Object"
2816 with statM statT_acc stat_acc widen_dynC_Obj statT isT_statT
2817 wf dynM' eq_static dynC_prop
2819 by - (drule dynmethd_access_prop,force+)
2822 then obtain im where
2823 im: "im \<in> imethds G I sig"
2825 have not_static_emh: "\<not> is_static emh"
2827 from im statM statT isT_statT wf not_Private_statM
2828 have "is_static im = is_static statM"
2829 by (fastsimp dest: wf_imethds_hiding_objmethdsD)
2830 with wf isT_statT statT im
2831 have "\<not> is_static statM"
2832 by (auto dest: wf_prog_idecl imethds_wf_mhead)
2837 with statT dynC_prop
2838 have implmnt: "G\<turnstile>dynC\<leadsto>I"
2840 with isT_statT statT
2841 have isif_I: "is_iface G I"
2844 have dynM': "dynimethd G I dynC sig = Some dynM"
2845 by (simp add: dynlookup_def)
2846 from False implmnt isif_I wf dynM'
2848 by - (drule implmt_dynimethd_access, auto)
2851 case (Array_Object_methd T statM)
2852 assume statT: "statT = ArrayT T"
2853 assume "accmethd G accC Object sig = Some statM"
2854 then have statM: "methd G Object sig = Some statM" and
2855 stat_acc: "G\<turnstile>Methd sig statM of Object accessible_from accC"
2856 by (auto dest: accmethd_SomeD)
2857 from statT dynC_prop
2858 have dynC_Obj: "dynC = Object"
2861 have widen_dynC_Obj: "G\<turnstile>Class dynC \<preceq> Class Object"
2864 have dynM': "dynmethd G Object dynC sig = Some dynM"
2865 by (simp add: dynlookup_def)
2866 from statM statT_acc stat_acc dynM' wf widen_dynC_Obj
2869 by - (drule dynmethd_access_prop, simp+)
2873 lemma dynlookup_access:
2874 assumes emh: "emh \<in> mheads G accC statT sig" and
2875 dynC_prop: "G,statT \<turnstile> dynC valid_lookup_cls_for (is_static emh) " and
2876 isT_statT: "isrtype G statT" and
2878 shows "\<exists> dynM. dynlookup G statT dynC sig = Some dynM \<and>
2879 G\<turnstile>Methd sig dynM in dynC dyn_accessible_from accC"
2881 from dynC_prop isT_statT wf
2882 have is_cls_dynC: "is_class G dynC"
2883 by (auto dest: valid_lookup_cls_is_class)
2884 with emh wf dynC_prop isT_statT
2886 "dynlookup G statT dynC sig = Some dynM"
2887 by - (drule dynamic_mheadsD,auto)
2888 with emh dynC_prop isT_statT wf
2890 by (blast intro: dynlookup_access_prop)
2893 lemma stat_overrides_Package_old:
2894 assumes stat_override: "G \<turnstile> new overrides\<^sub>S old" and
2895 accmodi_new: "accmodi new = Package" and
2897 shows "accmodi old = Package"
2899 from stat_override wf
2900 have "accmodi old \<le> accmodi new"
2901 by (auto dest: wf_prog_stat_overridesD)
2902 with stat_override accmodi_new show ?thesis
2903 by (cases "accmodi old") (auto dest: no_Private_stat_override
2904 dest: acc_modi_le_Dests)
2907 subsubsection {* Properties of dynamic accessibility *}
2909 lemma dyn_accessible_Private:
2910 assumes dyn_acc: "G \<turnstile> m in C dyn_accessible_from accC" and
2911 priv: "accmodi m = Private"
2912 shows "accC = declclass m"
2917 case (Immediate m C)
2918 from `G \<turnstile> m in C permits_acc_from accC` and `accmodi m = Private`
2920 by (simp add: permits_acc_def)
2924 by (auto dest!: overrides_commonD)
2928 text {* @{text dyn_accessible_Package} only works with the @{text wf_prog} assumption.
2929 Without it. it is easy to leaf the Package!
2931 lemma dyn_accessible_Package:
2932 "\<lbrakk>G \<turnstile> m in C dyn_accessible_from accC; accmodi m = Package;
2934 \<Longrightarrow> pid accC = pid (declclass m)"
2936 assume wf: "wf_prog G "
2937 assume accessible: "G \<turnstile> m in C dyn_accessible_from accC"
2938 then show "accmodi m = Package
2939 \<Longrightarrow> pid accC = pid (declclass m)"
2940 (is "?Pack m \<Longrightarrow> ?P m")
2941 proof (induct rule: dyn_accessible_fromR.induct)
2942 case (Immediate m C)
2943 assume "G\<turnstile>m member_in C"
2944 "G \<turnstile> m in C permits_acc_from accC"
2945 "accmodi m = Package"
2947 by (auto simp add: permits_acc_def)
2949 case (Overriding new C declC newm old Sup)
2950 assume member_new: "G \<turnstile> new member_in C" and
2951 new: "new = (declC, mdecl newm)" and
2952 override: "G \<turnstile> (declC, newm) overrides old" and
2953 subcls_C_Sup: "G\<turnstile>C \<prec>\<^sub>C Sup" and
2954 acc_old: "G \<turnstile> methdMembr old in Sup dyn_accessible_from accC" and
2955 hyp: "?Pack (methdMembr old) \<Longrightarrow> ?P (methdMembr old)" and
2956 accmodi_new: "accmodi new = Package"
2957 from override accmodi_new new wf
2958 have accmodi_old: "accmodi old = Package"
2959 by (auto dest: overrides_Package_old)
2961 have P_sup: "?P (methdMembr old)"
2963 from wf override new accmodi_old accmodi_new
2964 have eq_pid_new_old: "pid (declclass new) = pid (declclass old)"
2965 by (auto dest: dyn_override_Package)
2966 with eq_pid_new_old P_sup show "?P new"
2971 text {* For fields we don't need the wellformedness of the program, since
2972 there is no overriding *}
2973 lemma dyn_accessible_field_Package:
2974 assumes dyn_acc: "G \<turnstile> f in C dyn_accessible_from accC" and
2975 pack: "accmodi f = Package" and
2977 shows "pid accC = pid (declclass f)"
2979 from dyn_acc pack field
2982 case (Immediate f C)
2983 from `G \<turnstile> f in C permits_acc_from accC` and `accmodi f = Package`
2985 by (simp add: permits_acc_def)
2988 then show ?case by (simp add: is_field_def)
2992 text {* @{text dyn_accessible_instance_field_Protected} only works for fields
2993 since methods can break the package bounds due to overriding
2995 lemma dyn_accessible_instance_field_Protected:
2996 assumes dyn_acc: "G \<turnstile> f in C dyn_accessible_from accC" and
2997 prot: "accmodi f = Protected" and
2998 field: "is_field f" and
2999 instance_field: "\<not> is_static f" and
3000 outside: "pid (declclass f) \<noteq> pid accC"
3001 shows "G\<turnstile> C \<preceq>\<^sub>C accC"
3003 from dyn_acc prot field instance_field outside
3006 case (Immediate f C)
3007 note `G \<turnstile> f in C permits_acc_from accC`
3009 assume "accmodi f = Protected" and "is_field f" and "\<not> is_static f" and
3010 "pid (declclass f) \<noteq> pid accC"
3012 show "G\<turnstile> C \<preceq>\<^sub>C accC"
3013 by (auto simp add: permits_acc_def)
3016 then show ?case by (simp add: is_field_def)
3020 lemma dyn_accessible_static_field_Protected:
3021 assumes dyn_acc: "G \<turnstile> f in C dyn_accessible_from accC" and
3022 prot: "accmodi f = Protected" and
3023 field: "is_field f" and
3024 static_field: "is_static f" and
3025 outside: "pid (declclass f) \<noteq> pid accC"
3026 shows "G\<turnstile> accC \<preceq>\<^sub>C declclass f \<and> G\<turnstile>C \<preceq>\<^sub>C declclass f"
3028 from dyn_acc prot field static_field outside
3031 case (Immediate f C)
3032 assume "accmodi f = Protected" and "is_field f" and "is_static f" and
3033 "pid (declclass f) \<noteq> pid accC"
3035 note `G \<turnstile> f in C permits_acc_from accC`
3037 have "G\<turnstile> accC \<preceq>\<^sub>C declclass f"
3038 by (auto simp add: permits_acc_def)
3040 from `G \<turnstile> f member_in C`
3041 have "G\<turnstile>C \<preceq>\<^sub>C declclass f"
3042 by (rule member_in_class_relation)
3043 ultimately show ?case
3047 then show ?case by (simp add: is_field_def)