12854
|
1 |
header {* Advanced concepts on Java declarations like overriding, inheritance,
|
|
2 |
dynamic method lookup*}
|
|
3 |
|
|
4 |
theory DeclConcepts = TypeRel:
|
|
5 |
|
|
6 |
section "access control (cf. 6.6), overriding and hiding (cf. 8.4.6.1)"
|
|
7 |
|
|
8 |
constdefs
|
|
9 |
is_public :: "prog \<Rightarrow> qtname \<Rightarrow> bool"
|
|
10 |
"is_public G qn \<equiv> (case class G qn of
|
|
11 |
None \<Rightarrow> (case iface G qn of
|
|
12 |
None \<Rightarrow> False
|
|
13 |
| Some iface \<Rightarrow> access iface = Public)
|
|
14 |
| Some class \<Rightarrow> access class = Public)"
|
|
15 |
|
|
16 |
subsection "accessibility of types (cf. 6.6.1)"
|
|
17 |
text {*
|
|
18 |
Primitive types are always accessible, interfaces and classes are accessible
|
|
19 |
in their package or if they are defined public, an array type is accessible if
|
|
20 |
its element type is accessible *}
|
|
21 |
|
|
22 |
consts accessible_in :: "prog \<Rightarrow> ty \<Rightarrow> pname \<Rightarrow> bool"
|
|
23 |
("_ \<turnstile> _ accessible'_in _" [61,61,61] 60)
|
|
24 |
rt_accessible_in:: "prog \<Rightarrow> ref_ty \<Rightarrow> pname \<Rightarrow> bool"
|
|
25 |
("_ \<turnstile> _ accessible'_in' _" [61,61,61] 60)
|
|
26 |
primrec
|
|
27 |
"G\<turnstile>(PrimT p) accessible_in pack = True"
|
|
28 |
accessible_in_RefT_simp:
|
|
29 |
"G\<turnstile>(RefT r) accessible_in pack = G\<turnstile>r accessible_in' pack"
|
|
30 |
|
|
31 |
"G\<turnstile>(NullT) accessible_in' pack = True"
|
|
32 |
"G\<turnstile>(IfaceT I) accessible_in' pack = ((pid I = pack) \<or> is_public G I)"
|
|
33 |
"G\<turnstile>(ClassT C) accessible_in' pack = ((pid C = pack) \<or> is_public G C)"
|
|
34 |
"G\<turnstile>(ArrayT ty) accessible_in' pack = G\<turnstile>ty accessible_in pack"
|
|
35 |
|
|
36 |
declare accessible_in_RefT_simp [simp del]
|
|
37 |
|
|
38 |
constdefs
|
|
39 |
is_acc_class :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> bool"
|
|
40 |
"is_acc_class G P C \<equiv> is_class G C \<and> G\<turnstile>(Class C) accessible_in P"
|
|
41 |
is_acc_iface :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> bool"
|
|
42 |
"is_acc_iface G P I \<equiv> is_iface G I \<and> G\<turnstile>(Iface I) accessible_in P"
|
|
43 |
is_acc_type :: "prog \<Rightarrow> pname \<Rightarrow> ty \<Rightarrow> bool"
|
|
44 |
"is_acc_type G P T \<equiv> is_type G T \<and> G\<turnstile>T accessible_in P"
|
|
45 |
is_acc_reftype :: "prog \<Rightarrow> pname \<Rightarrow> ref_ty \<Rightarrow> bool"
|
|
46 |
"is_acc_reftype G P T \<equiv> isrtype G T \<and> G\<turnstile>T accessible_in' P"
|
|
47 |
|
|
48 |
lemma is_acc_classD:
|
|
49 |
"is_acc_class G P C \<Longrightarrow> is_class G C \<and> G\<turnstile>(Class C) accessible_in P"
|
|
50 |
by (simp add: is_acc_class_def)
|
|
51 |
|
|
52 |
lemma is_acc_class_is_class: "is_acc_class G P C \<Longrightarrow> is_class G C"
|
|
53 |
by (auto simp add: is_acc_class_def)
|
|
54 |
|
|
55 |
lemma is_acc_ifaceD:
|
|
56 |
"is_acc_iface G P I \<Longrightarrow> is_iface G I \<and> G\<turnstile>(Iface I) accessible_in P"
|
|
57 |
by (simp add: is_acc_iface_def)
|
|
58 |
|
|
59 |
lemma is_acc_typeD:
|
|
60 |
"is_acc_type G P T \<equiv> is_type G T \<and> G\<turnstile>T accessible_in P"
|
|
61 |
by (simp add: is_acc_type_def)
|
|
62 |
|
|
63 |
lemma is_acc_reftypeD:
|
|
64 |
"is_acc_reftype G P T \<Longrightarrow> isrtype G T \<and> G\<turnstile>T accessible_in' P"
|
|
65 |
by (simp add: is_acc_reftype_def)
|
|
66 |
|
|
67 |
subsection "accessibility of members"
|
|
68 |
text {*
|
|
69 |
The accessibility of members is more involved as the accessibility of types.
|
|
70 |
We have to distinguish several cases to model the different effects of
|
|
71 |
accessibility during inheritance, overriding and ordinary member access
|
|
72 |
*}
|
|
73 |
|
|
74 |
subsubsection {* Various technical conversion and selection functions *}
|
|
75 |
|
|
76 |
text {* overloaded selector @{text accmodi} to select the access modifier
|
|
77 |
out of various HOL types *}
|
|
78 |
|
|
79 |
axclass has_accmodi < "type"
|
|
80 |
consts accmodi:: "'a::has_accmodi \<Rightarrow> acc_modi"
|
|
81 |
|
|
82 |
instance acc_modi::has_accmodi
|
|
83 |
by (intro_classes)
|
|
84 |
|
|
85 |
defs (overloaded)
|
|
86 |
acc_modi_accmodi_def: "accmodi (a::acc_modi) \<equiv> a"
|
|
87 |
|
|
88 |
lemma acc_modi_accmodi_simp[simp]: "accmodi (a::acc_modi) = a"
|
|
89 |
by (simp add: acc_modi_accmodi_def)
|
|
90 |
|
|
91 |
instance access_field_type:: ("type","type") has_accmodi
|
|
92 |
by (intro_classes)
|
|
93 |
|
|
94 |
defs (overloaded)
|
|
95 |
decl_acc_modi_def: "accmodi (d::('a:: type) decl_scheme) \<equiv> access d"
|
|
96 |
|
|
97 |
|
|
98 |
lemma decl_acc_modi_simp[simp]: "accmodi (d::('a::type) decl_scheme) = access d"
|
|
99 |
by (simp add: decl_acc_modi_def)
|
|
100 |
|
|
101 |
instance * :: ("type",has_accmodi) has_accmodi
|
|
102 |
by (intro_classes)
|
|
103 |
|
|
104 |
defs (overloaded)
|
|
105 |
pair_acc_modi_def: "accmodi p \<equiv> (accmodi (snd p))"
|
|
106 |
|
|
107 |
lemma pair_acc_modi_simp[simp]: "accmodi (x,a) = (accmodi a)"
|
|
108 |
by (simp add: pair_acc_modi_def)
|
|
109 |
|
|
110 |
instance memberdecl :: has_accmodi
|
|
111 |
by (intro_classes)
|
|
112 |
|
|
113 |
defs (overloaded)
|
|
114 |
memberdecl_acc_modi_def: "accmodi m \<equiv> (case m of
|
|
115 |
fdecl f \<Rightarrow> accmodi f
|
|
116 |
| mdecl m \<Rightarrow> accmodi m)"
|
|
117 |
|
|
118 |
lemma memberdecl_fdecl_acc_modi_simp[simp]:
|
|
119 |
"accmodi (fdecl m) = accmodi m"
|
|
120 |
by (simp add: memberdecl_acc_modi_def)
|
|
121 |
|
|
122 |
lemma memberdecl_mdecl_acc_modi_simp[simp]:
|
|
123 |
"accmodi (mdecl m) = accmodi m"
|
|
124 |
by (simp add: memberdecl_acc_modi_def)
|
|
125 |
|
|
126 |
text {* overloaded selector @{text declclass} to select the declaring class
|
|
127 |
out of various HOL types *}
|
|
128 |
|
|
129 |
axclass has_declclass < "type"
|
|
130 |
consts declclass:: "'a::has_declclass \<Rightarrow> qtname"
|
|
131 |
|
|
132 |
instance pid_field_type::("type","type") has_declclass
|
|
133 |
by (intro_classes)
|
|
134 |
|
|
135 |
defs (overloaded)
|
|
136 |
qtname_declclass_def: "declclass (q::qtname) \<equiv> q"
|
|
137 |
|
|
138 |
lemma qtname_declclass_simp[simp]: "declclass (q::qtname) = q"
|
|
139 |
by (simp add: qtname_declclass_def)
|
|
140 |
|
|
141 |
instance * :: ("has_declclass","type") has_declclass
|
|
142 |
by (intro_classes)
|
|
143 |
|
|
144 |
defs (overloaded)
|
|
145 |
pair_declclass_def: "declclass p \<equiv> declclass (fst p)"
|
|
146 |
|
|
147 |
lemma pair_declclass_simp[simp]: "declclass (c,x) = declclass c"
|
|
148 |
by (simp add: pair_declclass_def)
|
|
149 |
|
|
150 |
text {* overloaded selector @{text is_static} to select the static modifier
|
|
151 |
out of various HOL types *}
|
|
152 |
|
|
153 |
|
|
154 |
axclass has_static < "type"
|
|
155 |
consts is_static :: "'a::has_static \<Rightarrow> bool"
|
|
156 |
|
|
157 |
(*
|
|
158 |
consts is_static :: "'a \<Rightarrow> bool"
|
|
159 |
*)
|
|
160 |
|
|
161 |
instance access_field_type :: ("type","has_static") has_static
|
|
162 |
by (intro_classes)
|
|
163 |
|
|
164 |
defs (overloaded)
|
|
165 |
decl_is_static_def:
|
|
166 |
"is_static (m::('a::has_static) decl_scheme) \<equiv> is_static (Decl.decl.more m)"
|
|
167 |
|
|
168 |
instance static_field_type :: ("type","type") has_static
|
|
169 |
by (intro_classes)
|
|
170 |
|
|
171 |
defs (overloaded)
|
|
172 |
static_field_type_is_static_def:
|
|
173 |
"is_static (m::(bool,'b::type) static_field_type) \<equiv> static_val m"
|
|
174 |
|
|
175 |
lemma member_is_static_simp: "is_static (m::'a member_scheme) = static m"
|
|
176 |
apply (cases m)
|
|
177 |
apply (simp add: static_field_type_is_static_def
|
|
178 |
decl_is_static_def Decl.member.dest_convs)
|
|
179 |
done
|
|
180 |
|
|
181 |
instance * :: ("type","has_static") has_static
|
|
182 |
by (intro_classes)
|
|
183 |
|
|
184 |
defs (overloaded)
|
|
185 |
pair_is_static_def: "is_static p \<equiv> is_static (snd p)"
|
|
186 |
|
|
187 |
lemma pair_is_static_simp [simp]: "is_static (x,s) = is_static s"
|
|
188 |
by (simp add: pair_is_static_def)
|
|
189 |
|
|
190 |
lemma pair_is_static_simp1: "is_static p = is_static (snd p)"
|
|
191 |
by (simp add: pair_is_static_def)
|
|
192 |
|
|
193 |
instance memberdecl:: has_static
|
|
194 |
by (intro_classes)
|
|
195 |
|
|
196 |
defs (overloaded)
|
|
197 |
memberdecl_is_static_def:
|
|
198 |
"is_static m \<equiv> (case m of
|
|
199 |
fdecl f \<Rightarrow> is_static f
|
|
200 |
| mdecl m \<Rightarrow> is_static m)"
|
|
201 |
|
|
202 |
lemma memberdecl_is_static_fdecl_simp[simp]:
|
|
203 |
"is_static (fdecl f) = is_static f"
|
|
204 |
by (simp add: memberdecl_is_static_def)
|
|
205 |
|
|
206 |
lemma memberdecl_is_static_mdecl_simp[simp]:
|
|
207 |
"is_static (mdecl m) = is_static m"
|
|
208 |
by (simp add: memberdecl_is_static_def)
|
|
209 |
|
|
210 |
lemma mhead_static_simp [simp]: "is_static (mhead m) = is_static m"
|
|
211 |
by (cases m) (simp add: mhead_def member_is_static_simp)
|
|
212 |
|
|
213 |
constdefs (* some mnemotic selectors for (qtname \<times> ('a::more) decl_scheme)
|
|
214 |
* the first component is a class or interface name
|
|
215 |
* the second component is a method, field or method head *)
|
|
216 |
(* "declclass":: "(qtname \<times> ('a::more) decl_scheme) \<Rightarrow> qtname"*)
|
|
217 |
(* "declclass \<equiv> fst" *) (* get the class component *)
|
|
218 |
|
|
219 |
"decliface":: "(qtname \<times> ('a::type) decl_scheme) \<Rightarrow> qtname"
|
|
220 |
"decliface \<equiv> fst" (* get the interface component *)
|
|
221 |
|
|
222 |
(*
|
|
223 |
"member":: "(qtname \<times> ('a::type) decl_scheme) \<Rightarrow> ('a::type) decl_scheme"
|
|
224 |
*)
|
|
225 |
"mbr":: "(qtname \<times> memberdecl) \<Rightarrow> memberdecl"
|
|
226 |
"mbr \<equiv> snd" (* get the memberdecl component *)
|
|
227 |
|
|
228 |
"mthd":: "('b \<times> 'a) \<Rightarrow> 'a"
|
|
229 |
(* also used for mdecl,mhead *)
|
|
230 |
"mthd \<equiv> snd" (* get the method component *)
|
|
231 |
|
|
232 |
"fld":: "('b \<times> ('a::type) decl_scheme) \<Rightarrow> ('a::type) decl_scheme"
|
|
233 |
(* also used for ((vname \<times> qtname)\<times> field) *)
|
|
234 |
"fld \<equiv> snd" (* get the field component *)
|
|
235 |
|
|
236 |
(* "accmodi" :: "('b \<times> ('a::type) decl_scheme) \<Rightarrow> acc_modi"*)
|
|
237 |
(* also used for mdecl *)
|
|
238 |
(* "accmodi \<equiv> access \<circ> snd"*) (* get the access modifier *)
|
|
239 |
(*
|
|
240 |
"is_static" ::"('b \<times> ('a::type) member_scheme) \<Rightarrow> bool" *)
|
|
241 |
(* also defined for emhead cf. WellType *)
|
|
242 |
(*"is_static \<equiv> static \<circ> snd"*) (* get the static modifier *)
|
|
243 |
|
|
244 |
constdefs (* some mnemotic selectors for (vname \<times> qtname) *)
|
|
245 |
fname:: "(vname \<times> 'a) \<Rightarrow> vname" (* also used for fdecl *)
|
|
246 |
"fname \<equiv> fst"
|
|
247 |
|
|
248 |
declclassf:: "(vname \<times> qtname) \<Rightarrow> qtname"
|
|
249 |
"declclassf \<equiv> snd"
|
|
250 |
|
|
251 |
(*
|
|
252 |
lemma declclass_simp[simp]: "declclass (C,m) = C"
|
|
253 |
by (simp add: declclass_def)
|
|
254 |
*)
|
|
255 |
|
|
256 |
lemma decliface_simp[simp]: "decliface (I,m) = I"
|
|
257 |
by (simp add: decliface_def)
|
|
258 |
|
|
259 |
lemma mbr_simp[simp]: "mbr (C,m) = m"
|
|
260 |
by (simp add: mbr_def)
|
|
261 |
|
|
262 |
lemma access_mbr_simp [simp]: "(accmodi (mbr m)) = accmodi m"
|
|
263 |
by (cases m) (simp add: mbr_def)
|
|
264 |
|
|
265 |
lemma mthd_simp[simp]: "mthd (C,m) = m"
|
|
266 |
by (simp add: mthd_def)
|
|
267 |
|
|
268 |
lemma fld_simp[simp]: "fld (C,f) = f"
|
|
269 |
by (simp add: fld_def)
|
|
270 |
|
|
271 |
lemma accmodi_simp[simp]: "accmodi (C,m) = access m"
|
|
272 |
by (simp )
|
|
273 |
|
|
274 |
lemma access_mthd_simp [simp]: "(access (mthd m)) = accmodi m"
|
|
275 |
by (cases m) (simp add: mthd_def)
|
|
276 |
|
|
277 |
lemma access_fld_simp [simp]: "(access (fld f)) = accmodi f"
|
|
278 |
by (cases f) (simp add: fld_def)
|
|
279 |
|
|
280 |
(*
|
|
281 |
lemma is_static_simp[simp]: "is_static (C,m) = static m"
|
|
282 |
by (simp add: is_static_def)
|
|
283 |
*)
|
|
284 |
|
|
285 |
lemma static_mthd_simp[simp]: "static (mthd m) = is_static m"
|
|
286 |
by (cases m) (simp add: mthd_def member_is_static_simp)
|
|
287 |
|
|
288 |
lemma mthd_is_static_simp [simp]: "is_static (mthd m) = is_static m"
|
|
289 |
by (cases m) simp
|
|
290 |
|
|
291 |
lemma static_fld_simp[simp]: "static (fld f) = is_static f"
|
|
292 |
by (cases f) (simp add: fld_def member_is_static_simp)
|
|
293 |
|
|
294 |
lemma ext_field_simp [simp]: "(declclass f,fld f) = f"
|
|
295 |
by (cases f) (simp add: fld_def)
|
|
296 |
|
|
297 |
lemma ext_method_simp [simp]: "(declclass m,mthd m) = m"
|
|
298 |
by (cases m) (simp add: mthd_def)
|
|
299 |
|
|
300 |
lemma ext_mbr_simp [simp]: "(declclass m,mbr m) = m"
|
|
301 |
by (cases m) (simp add: mbr_def)
|
|
302 |
|
|
303 |
lemma fname_simp[simp]:"fname (n,c) = n"
|
|
304 |
by (simp add: fname_def)
|
|
305 |
|
|
306 |
lemma declclassf_simp[simp]:"declclassf (n,c) = c"
|
|
307 |
by (simp add: declclassf_def)
|
|
308 |
|
|
309 |
constdefs (* some mnemotic selectors for (vname \<times> qtname) *)
|
|
310 |
"fldname" :: "(vname \<times> qtname) \<Rightarrow> vname"
|
|
311 |
"fldname \<equiv> fst"
|
|
312 |
|
|
313 |
"fldclass" :: "(vname \<times> qtname) \<Rightarrow> qtname"
|
|
314 |
"fldclass \<equiv> snd"
|
|
315 |
|
|
316 |
lemma fldname_simp[simp]: "fldname (n,c) = n"
|
|
317 |
by (simp add: fldname_def)
|
|
318 |
|
|
319 |
lemma fldclass_simp[simp]: "fldclass (n,c) = c"
|
|
320 |
by (simp add: fldclass_def)
|
|
321 |
|
|
322 |
lemma ext_fieldname_simp[simp]: "(fldname f,fldclass f) = f"
|
|
323 |
by (simp add: fldname_def fldclass_def)
|
|
324 |
|
|
325 |
text {* Convert a qualified method declaration (qualified with its declaring
|
|
326 |
class) to a qualified member declaration: @{text methdMembr} *}
|
|
327 |
|
|
328 |
constdefs
|
|
329 |
methdMembr :: "(qtname \<times> mdecl) \<Rightarrow> (qtname \<times> memberdecl)"
|
|
330 |
"methdMembr m \<equiv> (fst m,mdecl (snd m))"
|
|
331 |
|
|
332 |
lemma methdMembr_simp[simp]: "methdMembr (c,m) = (c,mdecl m)"
|
|
333 |
by (simp add: methdMembr_def)
|
|
334 |
|
|
335 |
lemma accmodi_methdMembr_simp[simp]: "accmodi (methdMembr m) = accmodi m"
|
|
336 |
by (cases m) (simp add: methdMembr_def)
|
|
337 |
|
|
338 |
lemma is_static_methdMembr_simp[simp]: "is_static (methdMembr m) = is_static m"
|
|
339 |
by (cases m) (simp add: methdMembr_def)
|
|
340 |
|
|
341 |
lemma declclass_methdMembr_simp[simp]: "declclass (methdMembr m) = declclass m"
|
|
342 |
by (cases m) (simp add: methdMembr_def)
|
|
343 |
|
|
344 |
text {* Convert a qualified method (qualified with its declaring
|
|
345 |
class) to a qualified member declaration: @{text method} *}
|
|
346 |
|
|
347 |
constdefs
|
|
348 |
method :: "sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> memberdecl)"
|
|
349 |
"method sig m \<equiv> (declclass m, mdecl (sig, mthd m))"
|
|
350 |
|
|
351 |
lemma method_simp[simp]: "method sig (C,m) = (C,mdecl (sig,m))"
|
|
352 |
by (simp add: method_def)
|
|
353 |
|
|
354 |
lemma accmodi_method_simp[simp]: "accmodi (method sig m) = accmodi m"
|
|
355 |
by (simp add: method_def)
|
|
356 |
|
|
357 |
lemma declclass_method_simp[simp]: "declclass (method sig m) = declclass m"
|
|
358 |
by (simp add: method_def)
|
|
359 |
|
|
360 |
lemma is_static_method_simp[simp]: "is_static (method sig m) = is_static m"
|
|
361 |
by (cases m) (simp add: method_def)
|
|
362 |
|
|
363 |
lemma mbr_method_simp[simp]: "mbr (method sig m) = mdecl (sig,mthd m)"
|
|
364 |
by (simp add: mbr_def method_def)
|
|
365 |
|
|
366 |
lemma memberid_method_simp[simp]: "memberid (method sig m) = mid sig"
|
|
367 |
by (simp add: method_def)
|
|
368 |
|
|
369 |
constdefs
|
|
370 |
fieldm :: "vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> (qtname \<times> memberdecl)"
|
|
371 |
"fieldm n f \<equiv> (declclass f, fdecl (n, fld f))"
|
|
372 |
|
|
373 |
lemma fieldm_simp[simp]: "fieldm n (C,f) = (C,fdecl (n,f))"
|
|
374 |
by (simp add: fieldm_def)
|
|
375 |
|
|
376 |
lemma accmodi_fieldm_simp[simp]: "accmodi (fieldm n f) = accmodi f"
|
|
377 |
by (simp add: fieldm_def)
|
|
378 |
|
|
379 |
lemma declclass_fieldm_simp[simp]: "declclass (fieldm n f) = declclass f"
|
|
380 |
by (simp add: fieldm_def)
|
|
381 |
|
|
382 |
lemma is_static_fieldm_simp[simp]: "is_static (fieldm n f) = is_static f"
|
|
383 |
by (cases f) (simp add: fieldm_def)
|
|
384 |
|
|
385 |
lemma mbr_fieldm_simp[simp]: "mbr (fieldm n f) = fdecl (n,fld f)"
|
|
386 |
by (simp add: mbr_def fieldm_def)
|
|
387 |
|
|
388 |
lemma memberid_fieldm_simp[simp]: "memberid (fieldm n f) = fid n"
|
|
389 |
by (simp add: fieldm_def)
|
|
390 |
|
|
391 |
text {* Select the signature out of a qualified method declaration:
|
|
392 |
@{text msig} *}
|
|
393 |
|
|
394 |
constdefs msig:: "(qtname \<times> mdecl) \<Rightarrow> sig"
|
|
395 |
"msig m \<equiv> fst (snd m)"
|
|
396 |
|
|
397 |
lemma msig_simp[simp]: "msig (c,(s,m)) = s"
|
|
398 |
by (simp add: msig_def)
|
|
399 |
|
|
400 |
text {* Convert a qualified method (qualified with its declaring
|
|
401 |
class) to a qualified method declaration: @{text qmdecl} *}
|
|
402 |
|
|
403 |
constdefs qmdecl :: "sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> mdecl)"
|
|
404 |
"qmdecl sig m \<equiv> (declclass m, (sig,mthd m))"
|
|
405 |
|
|
406 |
lemma qmdecl_simp[simp]: "qmdecl sig (C,m) = (C,(sig,m))"
|
|
407 |
by (simp add: qmdecl_def)
|
|
408 |
|
|
409 |
lemma declclass_qmdecl_simp[simp]: "declclass (qmdecl sig m) = declclass m"
|
|
410 |
by (simp add: qmdecl_def)
|
|
411 |
|
|
412 |
lemma accmodi_qmdecl_simp[simp]: "accmodi (qmdecl sig m) = accmodi m"
|
|
413 |
by (simp add: qmdecl_def)
|
|
414 |
|
|
415 |
lemma is_static_qmdecl_simp[simp]: "is_static (qmdecl sig m) = is_static m"
|
|
416 |
by (cases m) (simp add: qmdecl_def)
|
|
417 |
|
|
418 |
lemma msig_qmdecl_simp[simp]: "msig (qmdecl sig m) = sig"
|
|
419 |
by (simp add: qmdecl_def)
|
|
420 |
|
|
421 |
lemma mdecl_qmdecl_simp[simp]:
|
|
422 |
"mdecl (mthd (qmdecl sig new)) = mdecl (sig, mthd new)"
|
|
423 |
by (simp add: qmdecl_def)
|
|
424 |
|
|
425 |
lemma methdMembr_qmdecl_simp [simp]:
|
|
426 |
"methdMembr (qmdecl sig old) = method sig old"
|
|
427 |
by (simp add: methdMembr_def qmdecl_def method_def)
|
|
428 |
|
|
429 |
text {* overloaded selector @{text resTy} to select the result type
|
|
430 |
out of various HOL types *}
|
|
431 |
|
|
432 |
axclass has_resTy < "type"
|
|
433 |
consts resTy:: "'a::has_resTy \<Rightarrow> ty"
|
|
434 |
|
|
435 |
instance access_field_type :: ("type","has_resTy") has_resTy
|
|
436 |
by (intro_classes)
|
|
437 |
|
|
438 |
defs (overloaded)
|
|
439 |
decl_resTy_def:
|
|
440 |
"resTy (m::('a::has_resTy) decl_scheme) \<equiv> resTy (Decl.decl.more m)"
|
|
441 |
|
|
442 |
instance static_field_type :: ("type","has_resTy") has_resTy
|
|
443 |
by (intro_classes)
|
|
444 |
|
|
445 |
defs (overloaded)
|
|
446 |
static_field_type_resTy_def:
|
|
447 |
"resTy (m::(bool,'b::has_resTy) static_field_type)
|
|
448 |
\<equiv> resTy (static_more m)"
|
|
449 |
|
|
450 |
instance pars_field_type :: ("type","has_resTy") has_resTy
|
|
451 |
by (intro_classes)
|
|
452 |
|
|
453 |
defs (overloaded)
|
|
454 |
pars_field_type_resTy_def:
|
|
455 |
"resTy (m::(vname list,'b::has_resTy) pars_field_type)
|
|
456 |
\<equiv> resTy (pars_more m)"
|
|
457 |
|
|
458 |
instance resT_field_type :: ("type","type") has_resTy
|
|
459 |
by (intro_classes)
|
|
460 |
|
|
461 |
defs (overloaded)
|
|
462 |
resT_field_type_resTy_def:
|
|
463 |
"resTy (m::(ty,'b::type) resT_field_type)
|
|
464 |
\<equiv> resT_val m"
|
|
465 |
|
|
466 |
lemma mhead_resTy_simp: "resTy (m::'a mhead_scheme) = resT m"
|
|
467 |
apply (cases m)
|
|
468 |
apply (simp add: decl_resTy_def static_field_type_resTy_def
|
|
469 |
pars_field_type_resTy_def resT_field_type_resTy_def
|
|
470 |
member.dest_convs mhead.dest_convs)
|
|
471 |
done
|
|
472 |
|
|
473 |
lemma resTy_mhead [simp]:"resTy (mhead m) = resTy m"
|
|
474 |
by (simp add: mhead_def mhead_resTy_simp)
|
|
475 |
|
|
476 |
instance * :: ("type","has_resTy") has_resTy
|
|
477 |
by (intro_classes)
|
|
478 |
|
|
479 |
defs (overloaded)
|
|
480 |
pair_resTy_def: "resTy p \<equiv> resTy (snd p)"
|
|
481 |
|
|
482 |
lemma pair_resTy_simp[simp]: "resTy (x,m) = resTy m"
|
|
483 |
by (simp add: pair_resTy_def)
|
|
484 |
|
|
485 |
lemma qmdecl_resTy_simp [simp]: "resTy (qmdecl sig m) = resTy m"
|
|
486 |
by (cases m) (simp)
|
|
487 |
|
|
488 |
lemma resTy_mthd [simp]:"resTy (mthd m) = resTy m"
|
|
489 |
by (cases m) (simp add: mthd_def )
|
|
490 |
|
|
491 |
subsubsection "inheritable-in"
|
|
492 |
text {*
|
|
493 |
@{text "G\<turnstile>m inheritable_in P"}: m can be inherited by
|
|
494 |
classes in package P if:
|
|
495 |
\begin{itemize}
|
|
496 |
\item the declaration class of m is accessible in P and
|
|
497 |
\item the member m is declared with protected or public access or if it is
|
|
498 |
declared with default (package) access, the package of the declaration
|
|
499 |
class of m is also P. If the member m is declared with private access
|
|
500 |
it is not accessible for inheritance at all.
|
|
501 |
\end{itemize}
|
|
502 |
*}
|
|
503 |
constdefs
|
|
504 |
inheritable_in::
|
|
505 |
"prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> pname \<Rightarrow> bool"
|
|
506 |
("_ \<turnstile> _ inheritable'_in _" [61,61,61] 60)
|
|
507 |
"G\<turnstile>membr inheritable_in pack
|
|
508 |
\<equiv> (case (accmodi membr) of
|
|
509 |
Private \<Rightarrow> False
|
|
510 |
| Package \<Rightarrow> (pid (declclass membr)) = pack
|
|
511 |
| Protected \<Rightarrow> True
|
|
512 |
| Public \<Rightarrow> True)"
|
|
513 |
|
|
514 |
syntax
|
|
515 |
Method_inheritable_in::
|
|
516 |
"prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> pname \<Rightarrow> bool"
|
|
517 |
("_ \<turnstile>Method _ inheritable'_in _ " [61,61,61] 60)
|
|
518 |
|
|
519 |
translations
|
|
520 |
"G\<turnstile>Method m inheritable_in p" == "G\<turnstile>methdMembr m inheritable_in p"
|
|
521 |
|
|
522 |
syntax
|
|
523 |
Methd_inheritable_in::
|
|
524 |
"prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> pname \<Rightarrow> bool"
|
|
525 |
("_ \<turnstile>Methd _ _ inheritable'_in _ " [61,61,61,61] 60)
|
|
526 |
|
|
527 |
translations
|
|
528 |
"G\<turnstile>Methd s m inheritable_in p" == "G\<turnstile>(method s m) inheritable_in p"
|
|
529 |
|
|
530 |
subsubsection "declared-in/undeclared-in"
|
|
531 |
|
|
532 |
constdefs cdeclaredmethd:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,methd) table"
|
|
533 |
"cdeclaredmethd G C
|
|
534 |
\<equiv> (case class G C of
|
|
535 |
None \<Rightarrow> \<lambda> sig. None
|
|
536 |
| Some c \<Rightarrow> table_of (methods c)
|
|
537 |
)"
|
|
538 |
|
|
539 |
constdefs
|
|
540 |
cdeclaredfield:: "prog \<Rightarrow> qtname \<Rightarrow> (vname,field) table"
|
|
541 |
"cdeclaredfield G C
|
|
542 |
\<equiv> (case class G C of
|
|
543 |
None \<Rightarrow> \<lambda> sig. None
|
|
544 |
| Some c \<Rightarrow> table_of (cfields c)
|
|
545 |
)"
|
|
546 |
|
|
547 |
|
|
548 |
constdefs
|
|
549 |
declared_in:: "prog \<Rightarrow> memberdecl \<Rightarrow> qtname \<Rightarrow> bool"
|
|
550 |
("_\<turnstile> _ declared'_in _" [61,61,61] 60)
|
|
551 |
"G\<turnstile>m declared_in C \<equiv> (case m of
|
|
552 |
fdecl (fn,f ) \<Rightarrow> cdeclaredfield G C fn = Some f
|
|
553 |
| mdecl (sig,m) \<Rightarrow> cdeclaredmethd G C sig = Some m)"
|
|
554 |
|
|
555 |
syntax
|
|
556 |
method_declared_in:: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> bool"
|
|
557 |
("_\<turnstile>Method _ declared'_in _" [61,61,61] 60)
|
|
558 |
translations
|
|
559 |
"G\<turnstile>Method m declared_in C" == "G\<turnstile>mdecl (mthd m) declared_in C"
|
|
560 |
|
|
561 |
syntax
|
|
562 |
methd_declared_in:: "prog \<Rightarrow> sig \<Rightarrow>(qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> bool"
|
|
563 |
("_\<turnstile>Methd _ _ declared'_in _" [61,61,61,61] 60)
|
|
564 |
translations
|
|
565 |
"G\<turnstile>Methd s m declared_in C" == "G\<turnstile>mdecl (s,mthd m) declared_in C"
|
|
566 |
|
|
567 |
lemma declared_in_classD:
|
|
568 |
"G\<turnstile>m declared_in C \<Longrightarrow> is_class G C"
|
|
569 |
by (cases m)
|
|
570 |
(auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def)
|
|
571 |
|
|
572 |
constdefs
|
|
573 |
undeclared_in:: "prog \<Rightarrow> memberid \<Rightarrow> qtname \<Rightarrow> bool"
|
|
574 |
("_\<turnstile> _ undeclared'_in _" [61,61,61] 60)
|
|
575 |
|
|
576 |
"G\<turnstile>m undeclared_in C \<equiv> (case m of
|
|
577 |
fid fn \<Rightarrow> cdeclaredfield G C fn = None
|
|
578 |
| mid sig \<Rightarrow> cdeclaredmethd G C sig = None)"
|
|
579 |
|
|
580 |
lemma method_declared_inI:
|
|
581 |
"\<lbrakk>class G C = Some c; table_of (methods c) sig = Some m\<rbrakk>
|
|
582 |
\<Longrightarrow> G\<turnstile>mdecl (sig,m) declared_in C"
|
|
583 |
by (auto simp add: declared_in_def cdeclaredmethd_def)
|
|
584 |
|
|
585 |
|
|
586 |
subsubsection "members"
|
|
587 |
|
|
588 |
consts
|
|
589 |
members:: "prog \<Rightarrow> (qtname \<times> (qtname \<times> memberdecl)) set"
|
|
590 |
(* Can't just take a function: prog \<Rightarrow> qtname \<Rightarrow> memberdecl set because
|
|
591 |
the class qtname changes to the superclass in the inductive definition
|
|
592 |
below
|
|
593 |
*)
|
|
594 |
|
|
595 |
syntax
|
|
596 |
member_of:: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
|
|
597 |
("_ \<turnstile> _ member'_of _" [61,61,61] 60)
|
|
598 |
|
|
599 |
translations
|
|
600 |
"G\<turnstile>m member_of C" \<rightleftharpoons> "(C,m) \<in> members G"
|
|
601 |
|
|
602 |
inductive "members G" intros
|
|
603 |
|
|
604 |
Immediate: "\<lbrakk>G\<turnstile>mbr m declared_in C;declclass m = C\<rbrakk> \<Longrightarrow> G\<turnstile>m member_of C"
|
|
605 |
Inherited: "\<lbrakk>G\<turnstile>m inheritable_in (pid C); G\<turnstile>memberid m undeclared_in C;
|
|
606 |
G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S; G\<turnstile>(Class S) accessible_in (pid C);G\<turnstile>m member_of S
|
|
607 |
\<rbrakk> \<Longrightarrow> G\<turnstile>m member_of C"
|
|
608 |
text {* Note that in the case of an inherited member only the members of the
|
|
609 |
direct superclass are concerned. If a member of a superclass of the direct
|
|
610 |
superclass isn't inherited in the direct superclass (not member of the
|
|
611 |
direct superclass) than it can't be a member of the class. E.g. If a
|
|
612 |
member of a class A is defined with package access it isn't member of a
|
|
613 |
subclass S if S isn't in the same package as A. Any further subclasses
|
|
614 |
of S will not inherit the member, regardless if they are in the same
|
|
615 |
package as A or not.*}
|
|
616 |
|
|
617 |
syntax
|
|
618 |
method_member_of:: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> bool"
|
|
619 |
("_ \<turnstile>Method _ member'_of _" [61,61,61] 60)
|
|
620 |
|
|
621 |
translations
|
|
622 |
"G\<turnstile>Method m member_of C" \<rightleftharpoons> "G\<turnstile>(methdMembr m) member_of C"
|
|
623 |
|
|
624 |
syntax
|
|
625 |
methd_member_of:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> bool"
|
|
626 |
("_ \<turnstile>Methd _ _ member'_of _" [61,61,61,61] 60)
|
|
627 |
|
|
628 |
translations
|
|
629 |
"G\<turnstile>Methd s m member_of C" \<rightleftharpoons> "G\<turnstile>(method s m) member_of C"
|
|
630 |
|
|
631 |
syntax
|
|
632 |
fieldm_member_of:: "prog \<Rightarrow> vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> qtname \<Rightarrow> bool"
|
|
633 |
("_ \<turnstile>Field _ _ member'_of _" [61,61,61] 60)
|
|
634 |
|
|
635 |
translations
|
|
636 |
"G\<turnstile>Field n f member_of C" \<rightleftharpoons> "G\<turnstile>fieldm n f member_of C"
|
|
637 |
|
|
638 |
constdefs
|
|
639 |
inherits:: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> bool"
|
|
640 |
("_ \<turnstile> _ inherits _" [61,61,61] 60)
|
|
641 |
"G\<turnstile>C inherits m
|
|
642 |
\<equiv> G\<turnstile>m inheritable_in (pid C) \<and> G\<turnstile>memberid m undeclared_in C \<and>
|
|
643 |
(\<exists> S. G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S \<and> G\<turnstile>(Class S) accessible_in (pid C) \<and> G\<turnstile>m member_of S)"
|
|
644 |
|
|
645 |
lemma inherits_member: "G\<turnstile>C inherits m \<Longrightarrow> G\<turnstile>m member_of C"
|
|
646 |
by (auto simp add: inherits_def intro: members.Inherited)
|
|
647 |
|
|
648 |
|
|
649 |
constdefs member_in::"prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
|
|
650 |
("_ \<turnstile> _ member'_in _" [61,61,61] 60)
|
|
651 |
"G\<turnstile>m member_in C \<equiv> \<exists> provC. G\<turnstile> C \<preceq>\<^sub>C provC \<and> G \<turnstile> m member_of provC"
|
|
652 |
text {* A member is in a class if it is member of the class or a superclass.
|
|
653 |
If a member is in a class we can select this member. This additional notion
|
|
654 |
is necessary since not all members are inherited to subclasses. So such
|
|
655 |
members are not member-of the subclass but member-in the subclass.*}
|
|
656 |
|
|
657 |
syntax
|
|
658 |
method_member_in:: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> bool"
|
|
659 |
("_ \<turnstile>Method _ member'_in _" [61,61,61] 60)
|
|
660 |
|
|
661 |
translations
|
|
662 |
"G\<turnstile>Method m member_in C" \<rightleftharpoons> "G\<turnstile>(methdMembr m) member_in C"
|
|
663 |
|
|
664 |
syntax
|
|
665 |
methd_member_in:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> bool"
|
|
666 |
("_ \<turnstile>Methd _ _ member'_in _" [61,61,61,61] 60)
|
|
667 |
|
|
668 |
translations
|
|
669 |
"G\<turnstile>Methd s m member_in C" \<rightleftharpoons> "G\<turnstile>(method s m) member_in C"
|
|
670 |
|
|
671 |
consts stat_overridesR::
|
|
672 |
"prog \<Rightarrow> ((qtname \<times> mdecl) \<times> (qtname \<times> mdecl)) set"
|
|
673 |
|
|
674 |
lemma member_inD: "G\<turnstile>m member_in C
|
|
675 |
\<Longrightarrow> \<exists> provC. G\<turnstile> C \<preceq>\<^sub>C provC \<and> G \<turnstile> m member_of provC"
|
|
676 |
by (auto simp add: member_in_def)
|
|
677 |
|
|
678 |
lemma member_inI: "\<lbrakk>G \<turnstile> m member_of provC;G\<turnstile> C \<preceq>\<^sub>C provC\<rbrakk> \<Longrightarrow> G\<turnstile>m member_in C"
|
|
679 |
by (auto simp add: member_in_def)
|
|
680 |
|
|
681 |
lemma member_of_to_member_in: "G \<turnstile> m member_of C \<Longrightarrow> G \<turnstile>m member_in C"
|
|
682 |
by (auto intro: member_inI)
|
|
683 |
|
|
684 |
|
|
685 |
subsubsection "overriding"
|
|
686 |
|
|
687 |
text {* Unfortunately the static notion of overriding (used during the
|
|
688 |
typecheck of the compiler) and the dynamic notion of overriding (used during
|
|
689 |
execution in the JVM) are not exactly the same.
|
|
690 |
*}
|
|
691 |
|
|
692 |
text {* Static overriding (used during the typecheck of the compiler) *}
|
|
693 |
syntax
|
|
694 |
stat_overrides:: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool"
|
|
695 |
("_ \<turnstile> _ overrides\<^sub>S _" [61,61,61] 60)
|
|
696 |
translations
|
|
697 |
"G\<turnstile>new overrides\<^sub>S old" == "(new,old) \<in> stat_overridesR G "
|
|
698 |
|
|
699 |
inductive "stat_overridesR G" intros
|
|
700 |
|
|
701 |
Direct: "\<lbrakk>\<not> is_static new; msig new = msig old;
|
|
702 |
G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old);
|
|
703 |
G\<turnstile>Method new declared_in (declclass new);
|
|
704 |
G\<turnstile>Method old declared_in (declclass old);
|
|
705 |
G\<turnstile>Method old inheritable_in pid (declclass new);
|
|
706 |
G\<turnstile>(declclass new) \<prec>\<^sub>C\<^sub>1 superNew;
|
|
707 |
G \<turnstile>Method old member_of superNew
|
|
708 |
\<rbrakk> \<Longrightarrow> G\<turnstile>new overrides\<^sub>S old"
|
|
709 |
|
|
710 |
Indirect: "\<lbrakk>G\<turnstile>new overrides\<^sub>S inter; G\<turnstile>inter overrides\<^sub>S old\<rbrakk>
|
|
711 |
\<Longrightarrow> G\<turnstile>new overrides\<^sub>S old"
|
|
712 |
|
|
713 |
text {* Dynamic overriding (used during the typecheck of the compiler) *}
|
|
714 |
consts overridesR::
|
|
715 |
"prog \<Rightarrow> ((qtname \<times> mdecl) \<times> (qtname \<times> mdecl)) set"
|
|
716 |
|
|
717 |
|
|
718 |
overrides:: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool"
|
|
719 |
("_ \<turnstile> _ overrides _" [61,61,61] 60)
|
|
720 |
translations
|
|
721 |
"G\<turnstile>new overrides old" == "(new,old) \<in> overridesR G "
|
|
722 |
|
|
723 |
inductive "overridesR G" intros
|
|
724 |
|
|
725 |
Direct: "\<lbrakk>\<not> is_static new; \<not> is_static old; accmodi new \<noteq> Private;
|
|
726 |
msig new = msig old;
|
|
727 |
G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old);
|
|
728 |
G\<turnstile>Method new declared_in (declclass new);
|
|
729 |
G\<turnstile>Method old declared_in (declclass old);
|
|
730 |
G\<turnstile>Method old inheritable_in pid (declclass new);
|
|
731 |
G\<turnstile>resTy new \<preceq> resTy old
|
|
732 |
\<rbrakk> \<Longrightarrow> G\<turnstile>new overrides old"
|
|
733 |
|
|
734 |
Indirect: "\<lbrakk>G\<turnstile>new overrides inter; G\<turnstile>inter overrides old\<rbrakk>
|
|
735 |
\<Longrightarrow> G\<turnstile>new overrides old"
|
|
736 |
|
|
737 |
syntax
|
|
738 |
sig_stat_overrides::
|
|
739 |
"prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> methd) \<Rightarrow> bool"
|
|
740 |
("_,_\<turnstile> _ overrides\<^sub>S _" [61,61,61,61] 60)
|
|
741 |
translations
|
|
742 |
"G,s\<turnstile>new overrides\<^sub>S old" \<rightharpoonup> "G\<turnstile>(qmdecl s new) overrides\<^sub>S (qmdecl s old)"
|
|
743 |
|
|
744 |
syntax
|
|
745 |
sig_overrides:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> methd) \<Rightarrow> bool"
|
|
746 |
("_,_\<turnstile> _ overrides _" [61,61,61,61] 60)
|
|
747 |
translations
|
|
748 |
"G,s\<turnstile>new overrides old" \<rightharpoonup> "G\<turnstile>(qmdecl s new) overrides (qmdecl s old)"
|
|
749 |
|
|
750 |
subsubsection "Hiding"
|
|
751 |
|
|
752 |
constdefs hides::
|
|
753 |
"prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool"
|
|
754 |
("_\<turnstile> _ hides _" [61,61,61] 60)
|
|
755 |
"G\<turnstile>new hides old
|
|
756 |
\<equiv> is_static new \<and> msig new = msig old \<and>
|
|
757 |
G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
|
|
758 |
G\<turnstile>Method new declared_in (declclass new) \<and>
|
|
759 |
G\<turnstile>Method old declared_in (declclass old) \<and>
|
|
760 |
G\<turnstile>Method old inheritable_in pid (declclass new)"
|
|
761 |
|
|
762 |
syntax
|
|
763 |
sig_hides:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool"
|
|
764 |
("_,_\<turnstile> _ hides _" [61,61,61,61] 60)
|
|
765 |
translations
|
|
766 |
"G,s\<turnstile>new hides old" \<rightharpoonup> "G\<turnstile>(qmdecl s new) hides (qmdecl s old)"
|
|
767 |
|
|
768 |
lemma hidesI:
|
|
769 |
"\<lbrakk>is_static new; msig new = msig old;
|
|
770 |
G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old);
|
|
771 |
G\<turnstile>Method new declared_in (declclass new);
|
|
772 |
G\<turnstile>Method old declared_in (declclass old);
|
|
773 |
G\<turnstile>Method old inheritable_in pid (declclass new)
|
|
774 |
\<rbrakk> \<Longrightarrow> G\<turnstile>new hides old"
|
|
775 |
by (auto simp add: hides_def)
|
|
776 |
|
|
777 |
lemma hidesD:
|
|
778 |
"\<lbrakk>G\<turnstile>new hides old\<rbrakk> \<Longrightarrow>
|
|
779 |
declclass new \<noteq> Object \<and> is_static new \<and> msig new = msig old \<and>
|
|
780 |
G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
|
|
781 |
G\<turnstile>Method new declared_in (declclass new) \<and>
|
|
782 |
G\<turnstile>Method old declared_in (declclass old)"
|
|
783 |
by (auto simp add: hides_def)
|
|
784 |
|
|
785 |
lemma overrides_commonD:
|
|
786 |
"\<lbrakk>G\<turnstile>new overrides old\<rbrakk> \<Longrightarrow>
|
|
787 |
declclass new \<noteq> Object \<and> \<not> is_static new \<and> \<not> is_static old \<and>
|
|
788 |
accmodi new \<noteq> Private \<and>
|
|
789 |
msig new = msig old \<and>
|
|
790 |
G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
|
|
791 |
G\<turnstile>Method new declared_in (declclass new) \<and>
|
|
792 |
G\<turnstile>Method old declared_in (declclass old)"
|
|
793 |
by (induct set: overridesR) (auto intro: trancl_trans)
|
|
794 |
|
|
795 |
lemma ws_overrides_commonD:
|
|
796 |
"\<lbrakk>G\<turnstile>new overrides old;ws_prog G\<rbrakk> \<Longrightarrow>
|
|
797 |
declclass new \<noteq> Object \<and> \<not> is_static new \<and> \<not> is_static old \<and>
|
|
798 |
accmodi new \<noteq> Private \<and> G\<turnstile>resTy new \<preceq> resTy old \<and>
|
|
799 |
msig new = msig old \<and>
|
|
800 |
G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
|
|
801 |
G\<turnstile>Method new declared_in (declclass new) \<and>
|
|
802 |
G\<turnstile>Method old declared_in (declclass old)"
|
|
803 |
by (induct set: overridesR) (auto intro: trancl_trans ws_widen_trans)
|
|
804 |
|
|
805 |
lemma stat_overrides_commonD:
|
|
806 |
"\<lbrakk>G\<turnstile>new overrides\<^sub>S old\<rbrakk> \<Longrightarrow>
|
|
807 |
declclass new \<noteq> Object \<and> \<not> is_static new \<and> msig new = msig old \<and>
|
|
808 |
G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
|
|
809 |
G\<turnstile>Method new declared_in (declclass new) \<and>
|
|
810 |
G\<turnstile>Method old declared_in (declclass old)"
|
|
811 |
by (induct set: stat_overridesR) (auto intro: trancl_trans)
|
|
812 |
|
|
813 |
lemma overrides_eq_sigD:
|
|
814 |
"\<lbrakk>G\<turnstile>new overrides old\<rbrakk> \<Longrightarrow> msig old=msig new"
|
|
815 |
by (auto dest: overrides_commonD)
|
|
816 |
|
|
817 |
lemma hides_eq_sigD:
|
|
818 |
"\<lbrakk>G\<turnstile>new hides old\<rbrakk> \<Longrightarrow> msig old=msig new"
|
|
819 |
by (auto simp add: hides_def)
|
|
820 |
|
|
821 |
subsubsection "permits access"
|
|
822 |
constdefs
|
|
823 |
permits_acc::
|
|
824 |
"prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
|
|
825 |
("_ \<turnstile> _ in _ permits'_acc'_to _" [61,61,61,61] 60)
|
|
826 |
|
|
827 |
"G\<turnstile>membr in class permits_acc_to accclass
|
|
828 |
\<equiv> (case (accmodi membr) of
|
|
829 |
Private \<Rightarrow> (declclass membr = accclass)
|
|
830 |
| Package \<Rightarrow> (pid (declclass membr) = pid accclass)
|
|
831 |
| Protected \<Rightarrow> (pid (declclass membr) = pid accclass)
|
|
832 |
\<or>
|
|
833 |
(G\<turnstile>accclass \<prec>\<^sub>C declclass membr \<and> G\<turnstile>class \<preceq>\<^sub>C accclass)
|
|
834 |
| Public \<Rightarrow> True)"
|
|
835 |
text {*
|
|
836 |
The subcondition of the @{term "Protected"} case:
|
|
837 |
@{term "G\<turnstile>accclass \<prec>\<^sub>C declclass membr"} could also be relaxed to:
|
|
838 |
@{term "G\<turnstile>accclass \<preceq>\<^sub>C declclass membr"} since in case both classes are the
|
|
839 |
same the other condition @{term "(pid (declclass membr) = pid accclass)"}
|
|
840 |
holds anyway.
|
|
841 |
*}
|
|
842 |
|
|
843 |
text {* Like in case of overriding, the static and dynamic accessibility
|
|
844 |
of members is not uniform.
|
|
845 |
\begin{itemize}
|
|
846 |
\item Statically the class/interface of the member must be accessible for the
|
|
847 |
member to be accessible. During runtime this is not necessary. For
|
|
848 |
Example, if a class is accessible and we are allowed to access a member
|
|
849 |
of this class (statically) we expect that we can access this member in
|
|
850 |
an arbitrary subclass (during runtime). It's not intended to restrict
|
|
851 |
the access to accessible subclasses during runtime.
|
|
852 |
\item Statically the member we want to access must be "member of" the class.
|
|
853 |
Dynamically it must only be "member in" the class.
|
|
854 |
\end{itemize}
|
|
855 |
*}
|
|
856 |
|
|
857 |
|
|
858 |
consts
|
|
859 |
accessible_fromR::
|
|
860 |
"prog \<Rightarrow> qtname \<Rightarrow> ((qtname \<times> memberdecl) \<times> qtname) set"
|
|
861 |
|
|
862 |
syntax
|
|
863 |
accessible_from::
|
|
864 |
"prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
|
|
865 |
("_ \<turnstile> _ of _ accessible'_from _" [61,61,61,61] 60)
|
|
866 |
|
|
867 |
translations
|
|
868 |
"G\<turnstile>membr of cls accessible_from accclass"
|
|
869 |
\<rightleftharpoons> "(membr,cls) \<in> accessible_fromR G accclass"
|
|
870 |
|
|
871 |
syntax
|
|
872 |
method_accessible_from::
|
|
873 |
"prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
|
|
874 |
("_ \<turnstile>Method _ of _ accessible'_from _" [61,61,61,61] 60)
|
|
875 |
|
|
876 |
translations
|
|
877 |
"G\<turnstile>Method m of cls accessible_from accclass"
|
|
878 |
\<rightleftharpoons> "G\<turnstile>methdMembr m of cls accessible_from accclass"
|
|
879 |
|
|
880 |
syntax
|
|
881 |
methd_accessible_from::
|
|
882 |
"prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
|
|
883 |
("_ \<turnstile>Methd _ _ of _ accessible'_from _" [61,61,61,61,61] 60)
|
|
884 |
|
|
885 |
translations
|
|
886 |
"G\<turnstile>Methd s m of cls accessible_from accclass"
|
|
887 |
\<rightleftharpoons> "G\<turnstile>(method s m) of cls accessible_from accclass"
|
|
888 |
|
|
889 |
syntax
|
|
890 |
field_accessible_from::
|
|
891 |
"prog \<Rightarrow> vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
|
|
892 |
("_ \<turnstile>Field _ _ of _ accessible'_from _" [61,61,61,61,61] 60)
|
|
893 |
|
|
894 |
translations
|
|
895 |
"G\<turnstile>Field fn f of C accessible_from accclass"
|
|
896 |
\<rightleftharpoons> "G\<turnstile>(fieldm fn f) of C accessible_from accclass"
|
|
897 |
|
|
898 |
inductive "accessible_fromR G accclass" intros
|
|
899 |
immediate: "\<lbrakk>G\<turnstile>membr member_of class;
|
|
900 |
G\<turnstile>(Class class) accessible_in (pid accclass);
|
|
901 |
G\<turnstile>membr in class permits_acc_to accclass
|
|
902 |
\<rbrakk> \<Longrightarrow> G\<turnstile>membr of class accessible_from accclass"
|
|
903 |
|
|
904 |
overriding: "\<lbrakk>G\<turnstile>membr member_of class;
|
|
905 |
G\<turnstile>(Class class) accessible_in (pid accclass);
|
|
906 |
membr=(C,mdecl new);
|
|
907 |
G\<turnstile>(C,new) overrides\<^sub>S old;
|
|
908 |
G\<turnstile>class \<prec>\<^sub>C sup;
|
|
909 |
G\<turnstile>Method old of sup accessible_from accclass
|
|
910 |
\<rbrakk>\<Longrightarrow> G\<turnstile>membr of class accessible_from accclass"
|
|
911 |
|
|
912 |
consts
|
|
913 |
dyn_accessible_fromR::
|
|
914 |
"prog \<Rightarrow> qtname \<Rightarrow> ((qtname \<times> memberdecl) \<times> qtname) set"
|
|
915 |
|
|
916 |
syntax
|
|
917 |
dyn_accessible_from::
|
|
918 |
"prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
|
|
919 |
("_ \<turnstile> _ in _ dyn'_accessible'_from _" [61,61,61,61] 60)
|
|
920 |
|
|
921 |
translations
|
|
922 |
"G\<turnstile>membr in C dyn_accessible_from accC"
|
|
923 |
\<rightleftharpoons> "(membr,C) \<in> dyn_accessible_fromR G accC"
|
|
924 |
|
|
925 |
syntax
|
|
926 |
method_dyn_accessible_from::
|
|
927 |
"prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
|
|
928 |
("_ \<turnstile>Method _ in _ dyn'_accessible'_from _" [61,61,61,61] 60)
|
|
929 |
|
|
930 |
translations
|
|
931 |
"G\<turnstile>Method m in C dyn_accessible_from accC"
|
|
932 |
\<rightleftharpoons> "G\<turnstile>methdMembr m in C dyn_accessible_from accC"
|
|
933 |
|
|
934 |
syntax
|
|
935 |
methd_dyn_accessible_from::
|
|
936 |
"prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
|
|
937 |
("_ \<turnstile>Methd _ _ in _ dyn'_accessible'_from _" [61,61,61,61,61] 60)
|
|
938 |
|
|
939 |
translations
|
|
940 |
"G\<turnstile>Methd s m in C dyn_accessible_from accC"
|
|
941 |
\<rightleftharpoons> "G\<turnstile>(method s m) in C dyn_accessible_from accC"
|
|
942 |
|
|
943 |
syntax
|
|
944 |
field_dyn_accessible_from::
|
|
945 |
"prog \<Rightarrow> vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
|
|
946 |
("_ \<turnstile>Field _ _ in _ dyn'_accessible'_from _" [61,61,61,61,61] 60)
|
|
947 |
|
|
948 |
translations
|
|
949 |
"G\<turnstile>Field fn f in dynC dyn_accessible_from accC"
|
|
950 |
\<rightleftharpoons> "G\<turnstile>(fieldm fn f) in dynC dyn_accessible_from accC"
|
|
951 |
|
|
952 |
(* #### Testet JVM noch über den Bytecodeverifier hinaus ob der
|
|
953 |
statische Typ accessible ist bevor es den Zugriff erlaubt
|
|
954 |
\<longrightarrow> Test mit Reflektion\<dots>
|
|
955 |
*)
|
|
956 |
inductive "dyn_accessible_fromR G accclass" intros
|
|
957 |
immediate: "\<lbrakk>G\<turnstile>membr member_in class;
|
|
958 |
G\<turnstile>membr in class permits_acc_to accclass
|
|
959 |
\<rbrakk> \<Longrightarrow> G\<turnstile>membr in class dyn_accessible_from accclass"
|
|
960 |
|
|
961 |
overriding: "\<lbrakk>G\<turnstile>membr member_in class;
|
|
962 |
membr=(C,mdecl new);
|
|
963 |
G\<turnstile>(C,new) overrides old;
|
|
964 |
G\<turnstile>class \<prec>\<^sub>C sup;
|
|
965 |
G\<turnstile>Method old in sup dyn_accessible_from accclass
|
|
966 |
\<rbrakk>\<Longrightarrow> G\<turnstile>membr in class dyn_accessible_from accclass"
|
|
967 |
|
|
968 |
|
|
969 |
lemma accessible_from_commonD: "G\<turnstile>m of C accessible_from S
|
|
970 |
\<Longrightarrow> G\<turnstile>m member_of C \<and> G\<turnstile>(Class C) accessible_in (pid S)"
|
|
971 |
by (auto elim: accessible_fromR.induct)
|
|
972 |
|
|
973 |
lemma declared_not_undeclared:
|
|
974 |
"G\<turnstile>m declared_in C \<Longrightarrow> \<not> G\<turnstile> memberid m undeclared_in C"
|
|
975 |
by (cases m) (auto simp add: declared_in_def undeclared_in_def)
|
|
976 |
|
|
977 |
lemma not_undeclared_declared:
|
|
978 |
"\<not> G\<turnstile> membr_id undeclared_in C \<Longrightarrow> (\<exists> m. G\<turnstile>m declared_in C \<and>
|
|
979 |
membr_id = memberid m)"
|
|
980 |
proof -
|
|
981 |
assume not_undecl:"\<not> G\<turnstile> membr_id undeclared_in C"
|
|
982 |
show ?thesis (is "?P membr_id")
|
|
983 |
proof (cases membr_id)
|
|
984 |
case (fid vname)
|
|
985 |
with not_undecl
|
|
986 |
obtain fld where
|
|
987 |
"G\<turnstile>fdecl (vname,fld) declared_in C"
|
|
988 |
by (auto simp add: undeclared_in_def declared_in_def
|
|
989 |
cdeclaredfield_def)
|
|
990 |
with fid show ?thesis
|
|
991 |
by auto
|
|
992 |
next
|
|
993 |
case (mid sig)
|
|
994 |
with not_undecl
|
|
995 |
obtain mthd where
|
|
996 |
"G\<turnstile>mdecl (sig,mthd) declared_in C"
|
|
997 |
by (auto simp add: undeclared_in_def declared_in_def
|
|
998 |
cdeclaredmethd_def)
|
|
999 |
with mid show ?thesis
|
|
1000 |
by auto
|
|
1001 |
qed
|
|
1002 |
qed
|
|
1003 |
|
|
1004 |
lemma unique_declared_in:
|
|
1005 |
"\<lbrakk>G\<turnstile>m declared_in C; G\<turnstile>n declared_in C; memberid m = memberid n\<rbrakk>
|
|
1006 |
\<Longrightarrow> m = n"
|
|
1007 |
by (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def
|
|
1008 |
split: memberdecl.splits)
|
|
1009 |
|
|
1010 |
lemma unique_member_of:
|
|
1011 |
(assumes n: "G\<turnstile>n member_of C" and
|
|
1012 |
m: "G\<turnstile>m member_of C" and
|
|
1013 |
eqid: "memberid n = memberid m"
|
|
1014 |
) "n=m"
|
|
1015 |
proof -
|
|
1016 |
from n m eqid
|
|
1017 |
show "n=m"
|
|
1018 |
proof (induct)
|
|
1019 |
case (Immediate C n)
|
|
1020 |
assume member_n: "G\<turnstile> mbr n declared_in C" "declclass n = C"
|
|
1021 |
assume eqid: "memberid n = memberid m"
|
|
1022 |
assume "G \<turnstile> m member_of C"
|
|
1023 |
then show "n=m"
|
|
1024 |
proof (cases)
|
|
1025 |
case (Immediate _ m')
|
|
1026 |
with eqid
|
|
1027 |
have "m=m'"
|
|
1028 |
"memberid n = memberid m"
|
|
1029 |
"G\<turnstile> mbr m declared_in C"
|
|
1030 |
"declclass m = C"
|
|
1031 |
by auto
|
|
1032 |
with member_n
|
|
1033 |
show ?thesis
|
|
1034 |
by (cases n, cases m)
|
|
1035 |
(auto simp add: declared_in_def
|
|
1036 |
cdeclaredmethd_def cdeclaredfield_def
|
|
1037 |
split: memberdecl.splits)
|
|
1038 |
next
|
|
1039 |
case (Inherited _ _ m')
|
|
1040 |
then have "G\<turnstile> memberid m undeclared_in C"
|
|
1041 |
by simp
|
|
1042 |
with eqid member_n
|
|
1043 |
show ?thesis
|
|
1044 |
by (cases n) (auto dest: declared_not_undeclared)
|
|
1045 |
qed
|
|
1046 |
next
|
|
1047 |
case (Inherited C S n)
|
|
1048 |
assume undecl: "G\<turnstile> memberid n undeclared_in C"
|
|
1049 |
assume super: "G\<turnstile>C\<prec>\<^sub>C\<^sub>1S"
|
|
1050 |
assume hyp: "\<lbrakk>G \<turnstile> m member_of S; memberid n = memberid m\<rbrakk> \<Longrightarrow> n = m"
|
|
1051 |
assume eqid: "memberid n = memberid m"
|
|
1052 |
assume "G \<turnstile> m member_of C"
|
|
1053 |
then show "n=m"
|
|
1054 |
proof (cases)
|
|
1055 |
case Immediate
|
|
1056 |
then have "G\<turnstile> mbr m declared_in C" by simp
|
|
1057 |
with eqid undecl
|
|
1058 |
show ?thesis
|
|
1059 |
by (cases m) (auto dest: declared_not_undeclared)
|
|
1060 |
next
|
|
1061 |
case Inherited
|
|
1062 |
with super have "G \<turnstile> m member_of S"
|
|
1063 |
by (auto dest!: subcls1D)
|
|
1064 |
with eqid hyp
|
|
1065 |
show ?thesis
|
|
1066 |
by blast
|
|
1067 |
qed
|
|
1068 |
qed
|
|
1069 |
qed
|
|
1070 |
|
|
1071 |
lemma member_of_is_classD: "G\<turnstile>m member_of C \<Longrightarrow> is_class G C"
|
|
1072 |
proof (induct set: members)
|
|
1073 |
case (Immediate C m)
|
|
1074 |
assume "G\<turnstile> mbr m declared_in C"
|
|
1075 |
then show "is_class G C"
|
|
1076 |
by (cases "mbr m")
|
|
1077 |
(auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def)
|
|
1078 |
next
|
|
1079 |
case (Inherited C S m)
|
|
1080 |
assume "G\<turnstile>C\<prec>\<^sub>C\<^sub>1S" and "is_class G S"
|
|
1081 |
then show "is_class G C"
|
|
1082 |
by - (rule subcls_is_class2,auto)
|
|
1083 |
qed
|
|
1084 |
|
|
1085 |
lemma member_of_declC:
|
|
1086 |
"G\<turnstile>m member_of C
|
|
1087 |
\<Longrightarrow> G\<turnstile>mbr m declared_in (declclass m)"
|
|
1088 |
by (induct set: members) auto
|
|
1089 |
|
|
1090 |
lemma member_of_member_of_declC:
|
|
1091 |
"G\<turnstile>m member_of C
|
|
1092 |
\<Longrightarrow> G\<turnstile>m member_of (declclass m)"
|
|
1093 |
by (auto dest: member_of_declC intro: members.Immediate)
|
|
1094 |
|
|
1095 |
lemma member_of_class_relation:
|
|
1096 |
"G\<turnstile>m member_of C \<Longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m"
|
|
1097 |
proof (induct set: members)
|
|
1098 |
case (Immediate C m)
|
|
1099 |
then show "G\<turnstile>C \<preceq>\<^sub>C declclass m" by simp
|
|
1100 |
next
|
|
1101 |
case (Inherited C S m)
|
|
1102 |
then show "G\<turnstile>C \<preceq>\<^sub>C declclass m"
|
|
1103 |
by (auto dest: r_into_rtrancl intro: rtrancl_trans)
|
|
1104 |
qed
|
|
1105 |
|
|
1106 |
lemma member_in_class_relation:
|
|
1107 |
"G\<turnstile>m member_in C \<Longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m"
|
|
1108 |
by (auto dest: member_inD member_of_class_relation
|
|
1109 |
intro: rtrancl_trans)
|
|
1110 |
|
|
1111 |
lemma member_of_Package:
|
|
1112 |
"\<lbrakk>G\<turnstile>m member_of C; accmodi m = Package\<rbrakk>
|
|
1113 |
\<Longrightarrow> pid (declclass m) = pid C"
|
|
1114 |
proof -
|
|
1115 |
assume member: "G\<turnstile>m member_of C"
|
|
1116 |
then show " accmodi m = Package \<Longrightarrow> ?thesis" (is "PROP ?P m C")
|
|
1117 |
proof (induct rule: members.induct)
|
|
1118 |
fix C m
|
|
1119 |
assume C: "declclass m = C"
|
|
1120 |
then show "pid (declclass m) = pid C"
|
|
1121 |
by simp
|
|
1122 |
next
|
|
1123 |
fix C S m
|
|
1124 |
assume inheritable: "G \<turnstile> m inheritable_in pid C"
|
|
1125 |
assume hyp: "PROP ?P m S" and
|
|
1126 |
package_acc: "accmodi m = Package"
|
|
1127 |
with inheritable package_acc hyp
|
|
1128 |
show "pid (declclass m) = pid C"
|
|
1129 |
by (auto simp add: inheritable_in_def)
|
|
1130 |
qed
|
|
1131 |
qed
|
|
1132 |
|
|
1133 |
lemma dyn_accessible_from_commonD: "G\<turnstile>m in C dyn_accessible_from S
|
|
1134 |
\<Longrightarrow> G\<turnstile>m member_in C"
|
|
1135 |
by (auto elim: dyn_accessible_fromR.induct)
|
|
1136 |
|
|
1137 |
(* ### Gilt nicht für wf_progs!dynmaisches Override,
|
|
1138 |
da die accmodi Bedingung nur für stat override gilt! *)
|
|
1139 |
(*
|
|
1140 |
lemma override_Package:
|
|
1141 |
"\<lbrakk>G\<turnstile>new overrides old;
|
|
1142 |
\<And> new old. G\<turnstile>new overrides old \<Longrightarrow> accmodi old \<le> accmodi new;
|
|
1143 |
accmodi old = Package; accmodi new = Package\<rbrakk>
|
|
1144 |
\<Longrightarrow> pid (declclass old) = pid (declclass new)"
|
|
1145 |
proof -
|
|
1146 |
assume wf: "\<And> new old. G\<turnstile>new overrides old \<Longrightarrow> accmodi old \<le> accmodi new"
|
|
1147 |
assume ovverride: "G\<turnstile>new overrides old"
|
|
1148 |
then show "\<lbrakk>accmodi old = Package;accmodi new = Package\<rbrakk> \<Longrightarrow> ?thesis"
|
|
1149 |
(is "?Pack old \<Longrightarrow> ?Pack new \<Longrightarrow> ?EqPid old new")
|
|
1150 |
proof (induct rule: overridesR.induct)
|
|
1151 |
case Direct
|
|
1152 |
fix new old
|
|
1153 |
assume "accmodi old = Package"
|
|
1154 |
"G \<turnstile> methdMembr old inheritable_in pid (declclass new)"
|
|
1155 |
then show "pid (declclass old) = pid (declclass new)"
|
|
1156 |
by (auto simp add: inheritable_in_def)
|
|
1157 |
next
|
|
1158 |
case (Indirect inter new old)
|
|
1159 |
assume accmodi_old: "accmodi old = Package" and
|
|
1160 |
accmodi_new: "accmodi new = Package"
|
|
1161 |
assume "G \<turnstile> new overrides inter"
|
|
1162 |
with wf have le_inter_new: "accmodi inter \<le> accmodi new"
|
|
1163 |
by blast
|
|
1164 |
assume "G \<turnstile> inter overrides old"
|
|
1165 |
with wf have le_old_inter: "accmodi old \<le> accmodi inter"
|
|
1166 |
by blast
|
|
1167 |
from accmodi_old accmodi_new le_inter_new le_old_inter
|
|
1168 |
have "accmodi inter = Package"
|
|
1169 |
by(auto simp add: le_acc_def less_acc_def)
|
|
1170 |
with Indirect accmodi_old accmodi_new
|
|
1171 |
show "?EqPid old new"
|
|
1172 |
by auto
|
|
1173 |
qed
|
|
1174 |
qed
|
|
1175 |
|
|
1176 |
lemma stat_override_Package:
|
|
1177 |
"\<lbrakk>G\<turnstile>new overrides\<^sub>S old;
|
|
1178 |
\<And> new old. G\<turnstile>new overrides\<^sub>S old \<Longrightarrow> accmodi old \<le> accmodi new;
|
|
1179 |
accmodi old = Package; accmodi new = Package\<rbrakk>
|
|
1180 |
\<Longrightarrow> pid (declclass old) = pid (declclass new)"
|
|
1181 |
proof -
|
|
1182 |
assume wf: "\<And> new old. G\<turnstile>new overrides\<^sub>S old \<Longrightarrow> accmodi old \<le> accmodi new"
|
|
1183 |
assume ovverride: "G\<turnstile>new overrides\<^sub>S old"
|
|
1184 |
then show "\<lbrakk>accmodi old = Package;accmodi new = Package\<rbrakk> \<Longrightarrow> ?thesis"
|
|
1185 |
(is "?Pack old \<Longrightarrow> ?Pack new \<Longrightarrow> ?EqPid old new")
|
|
1186 |
proof (induct rule: stat_overridesR.induct)
|
|
1187 |
case Direct
|
|
1188 |
fix new old
|
|
1189 |
assume "accmodi old = Package"
|
|
1190 |
"G \<turnstile> methdMembr old inheritable_in pid (declclass new)"
|
|
1191 |
then show "pid (declclass old) = pid (declclass new)"
|
|
1192 |
by (auto simp add: inheritable_in_def)
|
|
1193 |
next
|
|
1194 |
case (Indirect inter new old)
|
|
1195 |
assume accmodi_old: "accmodi old = Package" and
|
|
1196 |
accmodi_new: "accmodi new = Package"
|
|
1197 |
assume "G \<turnstile> new overrides\<^sub>S inter"
|
|
1198 |
with wf have le_inter_new: "accmodi inter \<le> accmodi new"
|
|
1199 |
by blast
|
|
1200 |
assume "G \<turnstile> inter overrides\<^sub>S old"
|
|
1201 |
with wf have le_old_inter: "accmodi old \<le> accmodi inter"
|
|
1202 |
by blast
|
|
1203 |
from accmodi_old accmodi_new le_inter_new le_old_inter
|
|
1204 |
have "accmodi inter = Package"
|
|
1205 |
by(auto simp add: le_acc_def less_acc_def)
|
|
1206 |
with Indirect accmodi_old accmodi_new
|
|
1207 |
show "?EqPid old new"
|
|
1208 |
by auto
|
|
1209 |
qed
|
|
1210 |
qed
|
|
1211 |
|
|
1212 |
*)
|
|
1213 |
lemma no_Private_stat_override:
|
|
1214 |
"\<lbrakk>G\<turnstile>new overrides\<^sub>S old\<rbrakk> \<Longrightarrow> accmodi old \<noteq> Private"
|
|
1215 |
by (induct set: stat_overridesR) (auto simp add: inheritable_in_def)
|
|
1216 |
|
|
1217 |
lemma no_Private_override: "\<lbrakk>G\<turnstile>new overrides old\<rbrakk> \<Longrightarrow> accmodi old \<noteq> Private"
|
|
1218 |
by (induct set: overridesR) (auto simp add: inheritable_in_def)
|
|
1219 |
|
|
1220 |
lemma permits_acc_inheritance:
|
|
1221 |
"\<lbrakk>G\<turnstile>m in statC permits_acc_to accC; G\<turnstile>dynC \<preceq>\<^sub>C statC
|
|
1222 |
\<rbrakk> \<Longrightarrow> G\<turnstile>m in dynC permits_acc_to accC"
|
|
1223 |
by (cases "accmodi m")
|
|
1224 |
(auto simp add: permits_acc_def
|
|
1225 |
intro: subclseq_trans)
|
|
1226 |
|
|
1227 |
lemma field_accessible_fromD:
|
|
1228 |
"\<lbrakk>G\<turnstile>membr of C accessible_from accC;is_field membr\<rbrakk>
|
|
1229 |
\<Longrightarrow> G\<turnstile>membr member_of C \<and>
|
|
1230 |
G\<turnstile>(Class C) accessible_in (pid accC) \<and>
|
|
1231 |
G\<turnstile>membr in C permits_acc_to accC"
|
|
1232 |
by (cases set: accessible_fromR)
|
|
1233 |
(auto simp add: is_field_def split: memberdecl.splits)
|
|
1234 |
|
|
1235 |
lemma field_accessible_from_permits_acc_inheritance:
|
|
1236 |
"\<lbrakk>G\<turnstile>membr of statC accessible_from accC; is_field membr; G \<turnstile> dynC \<preceq>\<^sub>C statC\<rbrakk>
|
|
1237 |
\<Longrightarrow> G\<turnstile>membr in dynC permits_acc_to accC"
|
|
1238 |
by (auto dest: field_accessible_fromD intro: permits_acc_inheritance)
|
|
1239 |
|
|
1240 |
|
|
1241 |
(*
|
|
1242 |
lemma accessible_Package:
|
|
1243 |
"\<lbrakk>G \<turnstile> m of C accessible_from S;accmodi m = Package;
|
|
1244 |
\<And> new old. G\<turnstile>new overrides\<^sub>S old \<Longrightarrow> accmodi old \<le> accmodi new\<rbrakk>
|
|
1245 |
\<Longrightarrow> pid S = pid C \<and> pid C = pid (declclass m)"
|
|
1246 |
proof -
|
|
1247 |
assume wf: "\<And> new old. G\<turnstile>new overrides\<^sub>S old \<Longrightarrow> accmodi old \<le> accmodi new"
|
|
1248 |
assume "G \<turnstile> m of C accessible_from S"
|
|
1249 |
then show "accmodi m = Package \<Longrightarrow> pid S = pid C \<and> pid C = pid (declclass m)"
|
|
1250 |
(is "?Pack m \<Longrightarrow> ?P C m")
|
|
1251 |
proof (induct rule: accessible_fromR.induct)
|
|
1252 |
fix C m
|
|
1253 |
assume "G\<turnstile>m member_of C"
|
|
1254 |
"G \<turnstile> m in C permits_acc_to S"
|
|
1255 |
"accmodi m = Package"
|
|
1256 |
then show "?P C m"
|
|
1257 |
by (auto dest: member_of_Package simp add: permits_acc_def)
|
|
1258 |
next
|
|
1259 |
fix declC C new newm old Sup
|
|
1260 |
assume member_new: "G \<turnstile> new member_of C" and
|
|
1261 |
acc_C: "G \<turnstile> Class C accessible_in pid S" and
|
|
1262 |
new: "new = (declC, mdecl newm)" and
|
|
1263 |
override: "G \<turnstile> (declC, newm) overrides\<^sub>S old" and
|
|
1264 |
subcls_C_Sup: "G\<turnstile>C \<prec>\<^sub>C Sup" and
|
|
1265 |
acc_old: "G \<turnstile> methdMembr old of Sup accessible_from S" and
|
|
1266 |
hyp: "?Pack (methdMembr old) \<Longrightarrow> ?P Sup (methdMembr old)" and
|
|
1267 |
accmodi_new: "accmodi new = Package"
|
|
1268 |
from override wf
|
|
1269 |
have accmodi_weaken: "accmodi old \<le> accmodi newm"
|
|
1270 |
by (cases old,cases newm) auto
|
|
1271 |
from override new
|
|
1272 |
have "accmodi old \<noteq> Private"
|
|
1273 |
by (simp add: no_Private_stat_override)
|
|
1274 |
with accmodi_weaken accmodi_new new
|
|
1275 |
have accmodi_old: "accmodi old = Package"
|
|
1276 |
by (cases "accmodi old") (auto simp add: le_acc_def less_acc_def)
|
|
1277 |
with hyp
|
|
1278 |
have P_sup: "?P Sup (methdMembr old)"
|
|
1279 |
by (simp)
|
|
1280 |
from wf override new accmodi_old accmodi_new
|
|
1281 |
have eq_pid_new_old: "pid (declclass new) = pid (declclass old)"
|
|
1282 |
by (auto dest: stat_override_Package)
|
|
1283 |
from member_new accmodi_new
|
|
1284 |
have "pid (declclass new) = pid C"
|
|
1285 |
by (auto dest: member_of_Package)
|
|
1286 |
with eq_pid_new_old P_sup show "?P C new"
|
|
1287 |
by auto
|
|
1288 |
qed
|
|
1289 |
qed
|
|
1290 |
*)
|
|
1291 |
lemma accessible_fieldD:
|
|
1292 |
"\<lbrakk>G\<turnstile>membr of C accessible_from accC; is_field membr\<rbrakk>
|
|
1293 |
\<Longrightarrow> G\<turnstile>membr member_of C \<and>
|
|
1294 |
G\<turnstile>(Class C) accessible_in (pid accC) \<and>
|
|
1295 |
G\<turnstile>membr in C permits_acc_to accC"
|
|
1296 |
by (induct rule: accessible_fromR.induct) (auto dest: is_fieldD)
|
|
1297 |
|
|
1298 |
(* lemmata:
|
|
1299 |
Wegen G\<turnstile>Super accessible_in (pid C) folgt:
|
|
1300 |
G\<turnstile>m declared_in C; G\<turnstile>m member_of D; accmodi m = Package (G\<turnstile>D \<preceq>\<^sub>C C)
|
|
1301 |
\<Longrightarrow> pid C = pid D
|
|
1302 |
|
|
1303 |
C package
|
|
1304 |
m public in C
|
|
1305 |
für alle anderen D: G\<turnstile>m undeclared_in C
|
|
1306 |
m wird in alle subklassen vererbt, auch aus dem Package heraus!
|
|
1307 |
|
|
1308 |
G\<turnstile>m member_of C \<Longrightarrow> \<exists> D. G\<turnstile>C \<preceq>\<^sub>C D \<and> G\<turnstile>m declared_in D
|
|
1309 |
*)
|
|
1310 |
|
|
1311 |
(* Begriff (C,m) overrides (D,m)
|
|
1312 |
3 Fälle: Direkt,
|
|
1313 |
Indirekt über eine Zwischenklasse (ohne weiteres override)
|
|
1314 |
Indirekt über override
|
|
1315 |
*)
|
|
1316 |
|
|
1317 |
(*
|
|
1318 |
"G\<turnstile>m member_of C \<equiv>
|
|
1319 |
constdefs declares_method:: "prog \<Rightarrow> sig \<Rightarrow> qtname \<Rightarrow> methd \<Rightarrow> bool"
|
|
1320 |
("_,_\<turnstile> _ declares'_method _" [61,61,61,61] 60)
|
|
1321 |
"G,sig\<turnstile>C declares_method m \<equiv> cdeclaredmethd G C sig = Some m"
|
|
1322 |
|
|
1323 |
constdefs is_declared:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> bool"
|
|
1324 |
"is_declared G sig em \<equiv> G,sig\<turnstile>declclass em declares_method mthd em"
|
|
1325 |
*)
|
|
1326 |
|
|
1327 |
lemma member_of_Private:
|
|
1328 |
"\<lbrakk>G\<turnstile>m member_of C; accmodi m = Private\<rbrakk> \<Longrightarrow> declclass m = C"
|
|
1329 |
by (induct set: members) (auto simp add: inheritable_in_def)
|
|
1330 |
|
|
1331 |
lemma member_of_subclseq_declC:
|
|
1332 |
"G\<turnstile>m member_of C \<Longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m"
|
|
1333 |
by (induct set: members) (auto dest: r_into_rtrancl intro: rtrancl_trans)
|
|
1334 |
|
|
1335 |
lemma member_of_inheritance:
|
|
1336 |
(assumes m: "G\<turnstile>m member_of D" and
|
|
1337 |
subclseq_D_C: "G\<turnstile>D \<preceq>\<^sub>C C" and
|
|
1338 |
subclseq_C_m: "G\<turnstile>C \<preceq>\<^sub>C declclass m" and
|
|
1339 |
ws: "ws_prog G"
|
|
1340 |
) "G\<turnstile>m member_of C"
|
|
1341 |
proof -
|
|
1342 |
from m subclseq_D_C subclseq_C_m
|
|
1343 |
show ?thesis
|
|
1344 |
proof (induct)
|
|
1345 |
case (Immediate D m)
|
|
1346 |
assume "declclass m = D" and
|
|
1347 |
"G\<turnstile>D\<preceq>\<^sub>C C" and "G\<turnstile>C\<preceq>\<^sub>C declclass m"
|
|
1348 |
with ws have "D=C"
|
|
1349 |
by (auto intro: subclseq_acyclic)
|
|
1350 |
with Immediate
|
|
1351 |
show "G\<turnstile>m member_of C"
|
|
1352 |
by (auto intro: members.Immediate)
|
|
1353 |
next
|
|
1354 |
case (Inherited D S m)
|
|
1355 |
assume member_of_D_props:
|
|
1356 |
"G \<turnstile> m inheritable_in pid D"
|
|
1357 |
"G\<turnstile> memberid m undeclared_in D"
|
|
1358 |
"G \<turnstile> Class S accessible_in pid D"
|
|
1359 |
"G \<turnstile> m member_of S"
|
|
1360 |
assume super: "G\<turnstile>D\<prec>\<^sub>C\<^sub>1S"
|
|
1361 |
assume hyp: "\<lbrakk>G\<turnstile>S\<preceq>\<^sub>C C; G\<turnstile>C\<preceq>\<^sub>C declclass m\<rbrakk> \<Longrightarrow> G \<turnstile> m member_of C"
|
|
1362 |
assume subclseq_C_m: "G\<turnstile>C\<preceq>\<^sub>C declclass m"
|
|
1363 |
assume "G\<turnstile>D\<preceq>\<^sub>C C"
|
|
1364 |
then show "G\<turnstile>m member_of C"
|
|
1365 |
proof (cases rule: subclseq_cases)
|
|
1366 |
case Eq
|
|
1367 |
assume "D=C"
|
|
1368 |
with super member_of_D_props
|
|
1369 |
show ?thesis
|
|
1370 |
by (auto intro: members.Inherited)
|
|
1371 |
next
|
|
1372 |
case Subcls
|
|
1373 |
assume "G\<turnstile>D\<prec>\<^sub>C C"
|
|
1374 |
with super
|
|
1375 |
have "G\<turnstile>S\<preceq>\<^sub>C C"
|
|
1376 |
by (auto dest: subcls1D subcls_superD)
|
|
1377 |
with subclseq_C_m hyp show ?thesis
|
|
1378 |
by blast
|
|
1379 |
qed
|
|
1380 |
qed
|
|
1381 |
qed
|
|
1382 |
|
|
1383 |
lemma member_of_subcls:
|
|
1384 |
(assumes old: "G\<turnstile>old member_of C" and
|
|
1385 |
new: "G\<turnstile>new member_of D" and
|
|
1386 |
eqid: "memberid new = memberid old" and
|
|
1387 |
subclseq_D_C: "G\<turnstile>D \<preceq>\<^sub>C C" and
|
|
1388 |
subcls_new_old: "G\<turnstile>declclass new \<prec>\<^sub>C declclass old" and
|
|
1389 |
ws: "ws_prog G"
|
|
1390 |
) "G\<turnstile>D \<prec>\<^sub>C C"
|
|
1391 |
proof -
|
|
1392 |
from old
|
|
1393 |
have subclseq_C_old: "G\<turnstile>C \<preceq>\<^sub>C declclass old"
|
|
1394 |
by (auto dest: member_of_subclseq_declC)
|
|
1395 |
from new
|
|
1396 |
have subclseq_D_new: "G\<turnstile>D \<preceq>\<^sub>C declclass new"
|
|
1397 |
by (auto dest: member_of_subclseq_declC)
|
|
1398 |
from subcls_new_old ws
|
|
1399 |
have neq_new_old: "new\<noteq>old"
|
|
1400 |
by (cases new,cases old) (auto dest: subcls_irrefl)
|
|
1401 |
from subclseq_D_new subclseq_D_C
|
|
1402 |
have "G\<turnstile>(declclass new) \<preceq>\<^sub>C C \<or> G\<turnstile>C \<preceq>\<^sub>C (declclass new)"
|
|
1403 |
by (rule subcls_compareable)
|
|
1404 |
then have "G\<turnstile>(declclass new) \<preceq>\<^sub>C C"
|
|
1405 |
proof
|
|
1406 |
assume "G\<turnstile>declclass new\<preceq>\<^sub>C C" then show ?thesis .
|
|
1407 |
next
|
|
1408 |
assume "G\<turnstile>C \<preceq>\<^sub>C (declclass new)"
|
|
1409 |
with new subclseq_D_C ws
|
|
1410 |
have "G\<turnstile>new member_of C"
|
|
1411 |
by (blast intro: member_of_inheritance)
|
|
1412 |
with eqid old
|
|
1413 |
have "new=old"
|
|
1414 |
by (blast intro: unique_member_of)
|
|
1415 |
with neq_new_old
|
|
1416 |
show ?thesis
|
|
1417 |
by contradiction
|
|
1418 |
qed
|
|
1419 |
then show ?thesis
|
|
1420 |
proof (cases rule: subclseq_cases)
|
|
1421 |
case Eq
|
|
1422 |
assume "declclass new = C"
|
|
1423 |
with new have "G\<turnstile>new member_of C"
|
|
1424 |
by (auto dest: member_of_member_of_declC)
|
|
1425 |
with eqid old
|
|
1426 |
have "new=old"
|
|
1427 |
by (blast intro: unique_member_of)
|
|
1428 |
with neq_new_old
|
|
1429 |
show ?thesis
|
|
1430 |
by contradiction
|
|
1431 |
next
|
|
1432 |
case Subcls
|
|
1433 |
assume "G\<turnstile>declclass new\<prec>\<^sub>C C"
|
|
1434 |
with subclseq_D_new
|
|
1435 |
show "G\<turnstile>D\<prec>\<^sub>C C"
|
|
1436 |
by (rule rtrancl_trancl_trancl)
|
|
1437 |
qed
|
|
1438 |
qed
|
|
1439 |
|
|
1440 |
corollary member_of_overrides_subcls:
|
|
1441 |
"\<lbrakk>G\<turnstile>Methd sig old member_of C; G\<turnstile>Methd sig new member_of D;G\<turnstile>D \<preceq>\<^sub>C C;
|
|
1442 |
G,sig\<turnstile>new overrides old; ws_prog G\<rbrakk>
|
|
1443 |
\<Longrightarrow> G\<turnstile>D \<prec>\<^sub>C C"
|
|
1444 |
by (drule overrides_commonD) (auto intro: member_of_subcls)
|
|
1445 |
|
|
1446 |
corollary member_of_stat_overrides_subcls:
|
|
1447 |
"\<lbrakk>G\<turnstile>Methd sig old member_of C; G\<turnstile>Methd sig new member_of D;G\<turnstile>D \<preceq>\<^sub>C C;
|
|
1448 |
G,sig\<turnstile>new overrides\<^sub>S old; ws_prog G\<rbrakk>
|
|
1449 |
\<Longrightarrow> G\<turnstile>D \<prec>\<^sub>C C"
|
|
1450 |
by (drule stat_overrides_commonD) (auto intro: member_of_subcls)
|
|
1451 |
|
|
1452 |
|
|
1453 |
|
|
1454 |
lemma inherited_field_access:
|
|
1455 |
(assumes stat_acc: "G\<turnstile>membr of statC accessible_from accC" and
|
|
1456 |
is_field: "is_field membr" and
|
|
1457 |
subclseq: "G \<turnstile> dynC \<preceq>\<^sub>C statC"
|
|
1458 |
) "G\<turnstile>membr in dynC dyn_accessible_from accC"
|
|
1459 |
proof -
|
|
1460 |
from stat_acc is_field subclseq
|
|
1461 |
show ?thesis
|
|
1462 |
by (auto dest: accessible_fieldD
|
|
1463 |
intro: dyn_accessible_fromR.immediate
|
|
1464 |
member_inI
|
|
1465 |
permits_acc_inheritance)
|
|
1466 |
qed
|
|
1467 |
|
|
1468 |
lemma accessible_inheritance:
|
|
1469 |
(assumes stat_acc: "G\<turnstile>m of statC accessible_from accC" and
|
|
1470 |
subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
|
|
1471 |
member_dynC: "G\<turnstile>m member_of dynC" and
|
|
1472 |
dynC_acc: "G\<turnstile>(Class dynC) accessible_in (pid accC)"
|
|
1473 |
) "G\<turnstile>m of dynC accessible_from accC"
|
|
1474 |
proof -
|
|
1475 |
from stat_acc
|
|
1476 |
have member_statC: "G\<turnstile>m member_of statC"
|
|
1477 |
by (auto dest: accessible_from_commonD)
|
|
1478 |
from stat_acc
|
|
1479 |
show ?thesis
|
|
1480 |
proof (cases)
|
|
1481 |
case immediate
|
|
1482 |
with member_dynC member_statC subclseq dynC_acc
|
|
1483 |
show ?thesis
|
|
1484 |
by (auto intro: accessible_fromR.immediate permits_acc_inheritance)
|
|
1485 |
next
|
|
1486 |
case overriding
|
|
1487 |
with member_dynC subclseq dynC_acc
|
|
1488 |
show ?thesis
|
|
1489 |
by (auto intro: accessible_fromR.overriding rtrancl_trancl_trancl)
|
|
1490 |
qed
|
|
1491 |
qed
|
|
1492 |
|
|
1493 |
section "fields and methods"
|
|
1494 |
|
|
1495 |
|
|
1496 |
types
|
|
1497 |
fspec = "vname \<times> qtname"
|
|
1498 |
|
|
1499 |
translations
|
|
1500 |
"fspec" <= (type) "vname \<times> qtname"
|
|
1501 |
|
|
1502 |
constdefs
|
|
1503 |
imethds:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
|
|
1504 |
"imethds G I
|
|
1505 |
\<equiv> iface_rec (G,I)
|
|
1506 |
(\<lambda>I i ts. (Un_tables ts) \<oplus>\<oplus>
|
|
1507 |
(o2s \<circ> table_of (map (\<lambda>(s,m). (s,I,m)) (imethods i))))"
|
|
1508 |
text {* methods of an interface, with overriding and inheritance, cf. 9.2 *}
|
|
1509 |
|
|
1510 |
constdefs
|
|
1511 |
accimethds:: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
|
|
1512 |
"accimethds G pack I
|
|
1513 |
\<equiv> if G\<turnstile>Iface I accessible_in pack
|
|
1514 |
then imethds G I
|
|
1515 |
else \<lambda> k. {}"
|
|
1516 |
text {* only returns imethds if the interface is accessible *}
|
|
1517 |
|
|
1518 |
constdefs
|
|
1519 |
methd:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
|
|
1520 |
|
|
1521 |
"methd G C
|
|
1522 |
\<equiv> class_rec (G,C) empty
|
|
1523 |
(\<lambda>C c subcls_mthds.
|
|
1524 |
filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m)
|
|
1525 |
subcls_mthds
|
|
1526 |
++
|
|
1527 |
table_of (map (\<lambda>(s,m). (s,C,m)) (methods c)))"
|
|
1528 |
text {* @{term "methd G C"}: methods of a class C (statically visible from C),
|
|
1529 |
with inheritance and hiding cf. 8.4.6;
|
|
1530 |
Overriding is captured by @{text dynmethd}.
|
|
1531 |
Every new method with the same signature coalesces the
|
|
1532 |
method of a superclass. *}
|
|
1533 |
|
|
1534 |
constdefs
|
|
1535 |
accmethd:: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
|
|
1536 |
"accmethd G S C
|
|
1537 |
\<equiv> filter_tab (\<lambda>sig m. G\<turnstile>method sig m of C accessible_from S)
|
|
1538 |
(methd G C)"
|
|
1539 |
text {* @{term "accmethd G S C"}: only those methods of @{term "methd G C"},
|
|
1540 |
accessible from S *}
|
|
1541 |
|
|
1542 |
text {* Note the class component in the accessibility filter. The class where
|
|
1543 |
method @{term m} is declared (@{term declC}) isn't necessarily accessible
|
|
1544 |
from the current scope @{term S}. The method can be made accessible
|
|
1545 |
through inheritance, too.
|
|
1546 |
So we must test accessibility of method @{term m} of class @{term C}
|
|
1547 |
(not @{term "declclass m"}) *}
|
|
1548 |
|
|
1549 |
constdefs
|
|
1550 |
dynmethd:: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
|
|
1551 |
"dynmethd G statC dynC
|
|
1552 |
\<equiv> \<lambda> sig.
|
|
1553 |
(if G\<turnstile>dynC \<preceq>\<^sub>C statC
|
|
1554 |
then (case methd G statC sig of
|
|
1555 |
None \<Rightarrow> None
|
|
1556 |
| Some statM
|
|
1557 |
\<Rightarrow> (class_rec (G,dynC) empty
|
|
1558 |
(\<lambda>C c subcls_mthds.
|
|
1559 |
subcls_mthds
|
|
1560 |
++
|
|
1561 |
(filter_tab
|
|
1562 |
(\<lambda> _ dynM. G,sig\<turnstile>dynM overrides statM \<or> dynM=statM)
|
|
1563 |
(methd G C) ))
|
|
1564 |
) sig
|
|
1565 |
)
|
|
1566 |
else None)"
|
|
1567 |
(*
|
|
1568 |
"dynmethd G statC dynC
|
|
1569 |
\<equiv> \<lambda> sig.
|
|
1570 |
(if G\<turnstile>dynC \<preceq>\<^sub>C statC
|
|
1571 |
then (case methd G statC sig of
|
|
1572 |
None \<Rightarrow> None
|
|
1573 |
| Some statM
|
|
1574 |
\<Rightarrow> (class_rec (G,statC) empty
|
|
1575 |
(\<lambda>C c subcls_mthds.
|
|
1576 |
subcls_mthds
|
|
1577 |
++
|
|
1578 |
(filter_tab
|
|
1579 |
(\<lambda> _ dynM. G,sig\<turnstile>dynM overrides statM)
|
|
1580 |
(table_of (map (\<lambda>(s,m). (s,C,m)) (methods c)))))
|
|
1581 |
) sig
|
|
1582 |
)
|
|
1583 |
else None)"*)
|
|
1584 |
text {* @{term "dynmethd G statC dynC"}: dynamic method lookup of a reference
|
|
1585 |
with dynamic class @{term dynC} and static class @{term statC} *}
|
|
1586 |
text {* Note some kind of duality between @{term methd} and @{term dynmethd}
|
|
1587 |
in the @{term class_rec} arguments. Whereas @{term methd} filters the
|
|
1588 |
subclass methods (to get only the inherited ones), @{term dynmethd}
|
|
1589 |
filters the new methods (to get only those methods which actually
|
|
1590 |
override the methods of the static class) *}
|
|
1591 |
|
|
1592 |
constdefs
|
|
1593 |
dynimethd:: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
|
|
1594 |
"dynimethd G I dynC
|
|
1595 |
\<equiv> \<lambda> sig. if imethds G I sig \<noteq> {}
|
|
1596 |
then methd G dynC sig
|
|
1597 |
else dynmethd G Object dynC sig"
|
|
1598 |
text {* @{term "dynimethd G I dynC"}: dynamic method lookup of a reference with
|
|
1599 |
dynamic class dynC and static interface type I *}
|
|
1600 |
text {*
|
|
1601 |
When calling an interface method, we must distinguish if the method signature
|
|
1602 |
was defined in the interface or if it must be an Object method in the other
|
|
1603 |
case. If it was an interface method we search the class hierarchy
|
|
1604 |
starting at the dynamic class of the object up to Object to find the
|
|
1605 |
first matching method (@{term methd}). Since all interface methods have
|
|
1606 |
public access the method can't be coalesced due to some odd visibility
|
|
1607 |
effects like in case of dynmethd. The method will be inherited or
|
|
1608 |
overridden in all classes from the first class implementing the interface
|
|
1609 |
down to the actual dynamic class.
|
|
1610 |
*}
|
|
1611 |
|
|
1612 |
constdefs
|
|
1613 |
dynlookup::"prog \<Rightarrow> ref_ty \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
|
|
1614 |
"dynlookup G statT dynC
|
|
1615 |
\<equiv> (case statT of
|
|
1616 |
NullT \<Rightarrow> empty
|
|
1617 |
| IfaceT I \<Rightarrow> dynimethd G I dynC
|
|
1618 |
| ClassT statC \<Rightarrow> dynmethd G statC dynC
|
|
1619 |
| ArrayT ty \<Rightarrow> dynmethd G Object dynC)"
|
|
1620 |
text {* @{term "dynlookup G statT dynC"}: dynamic lookup of a method within the
|
|
1621 |
static reference type statT and the dynamic class dynC.
|
|
1622 |
In a wellformd context statT will not be NullT and in case
|
|
1623 |
statT is an array type, dynC=Object *}
|
|
1624 |
|
|
1625 |
constdefs
|
|
1626 |
fields:: "prog \<Rightarrow> qtname \<Rightarrow> ((vname \<times> qtname) \<times> field) list"
|
|
1627 |
"fields G C
|
|
1628 |
\<equiv> class_rec (G,C) [] (\<lambda>C c ts. map (\<lambda>(n,t). ((n,C),t)) (cfields c) @ ts)"
|
|
1629 |
text {* @{term "fields G C"}
|
|
1630 |
list of fields of a class, including all the fields of the superclasses
|
|
1631 |
(private, inherited and hidden ones) not only the accessible ones
|
|
1632 |
(an instance of a object allocates all these fields *}
|
|
1633 |
|
|
1634 |
constdefs
|
|
1635 |
accfield:: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (vname, qtname \<times> field) table"
|
|
1636 |
"accfield G S C
|
|
1637 |
\<equiv> let field_tab = table_of((map (\<lambda>((n,d),f).(n,(d,f)))) (fields G C))
|
|
1638 |
in filter_tab (\<lambda>n (declC,f). G\<turnstile> (declC,fdecl (n,f)) of C accessible_from S)
|
|
1639 |
field_tab"
|
|
1640 |
text {* @{term "accfield G C S"}: fields of a class @{term C} which are
|
|
1641 |
accessible from scope of class
|
|
1642 |
@{term S} with inheritance and hiding, cf. 8.3 *}
|
|
1643 |
text {* note the class component in the accessibility filter (see also
|
|
1644 |
@{term methd}).
|
|
1645 |
The class declaring field @{term f} (@{term declC}) isn't necessarily
|
|
1646 |
accessible from scope @{term S}. The field can be made visible through
|
|
1647 |
inheritance, too. So we must test accessibility of field @{term f} of class
|
|
1648 |
@{term C} (not @{term "declclass f"}) *}
|
|
1649 |
|
|
1650 |
constdefs
|
|
1651 |
|
|
1652 |
is_methd :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> bool"
|
|
1653 |
"is_methd G \<equiv> \<lambda>C sig. is_class G C \<and> methd G C sig \<noteq> None"
|
|
1654 |
|
|
1655 |
constdefs efname:: "((vname \<times> qtname) \<times> field) \<Rightarrow> (vname \<times> qtname)"
|
|
1656 |
"efname \<equiv> fst"
|
|
1657 |
|
|
1658 |
lemma efname_simp[simp]:"efname (n,f) = n"
|
|
1659 |
by (simp add: efname_def)
|
|
1660 |
|
|
1661 |
|
|
1662 |
subsection "imethds"
|
|
1663 |
|
|
1664 |
lemma imethds_rec: "\<lbrakk>iface G I = Some i; ws_prog G\<rbrakk> \<Longrightarrow>
|
|
1665 |
imethds G I = Un_tables ((\<lambda>J. imethds G J)`set (isuperIfs i)) \<oplus>\<oplus>
|
|
1666 |
(o2s \<circ> table_of (map (\<lambda>(s,mh). (s,I,mh)) (imethods i)))"
|
|
1667 |
apply (unfold imethds_def)
|
|
1668 |
apply (rule iface_rec [THEN trans])
|
|
1669 |
apply auto
|
|
1670 |
done
|
|
1671 |
|
|
1672 |
|
|
1673 |
(* local lemma *)
|
|
1674 |
lemma imethds_norec:
|
|
1675 |
"\<lbrakk>iface G md = Some i; ws_prog G; table_of (imethods i) sig = Some mh\<rbrakk> \<Longrightarrow>
|
|
1676 |
(md, mh) \<in> imethds G md sig"
|
|
1677 |
apply (subst imethds_rec)
|
|
1678 |
apply assumption+
|
|
1679 |
apply (rule iffD2)
|
|
1680 |
apply (rule overrides_t_Some_iff)
|
|
1681 |
apply (rule disjI1)
|
|
1682 |
apply (auto elim: table_of_map_SomeI)
|
|
1683 |
done
|
|
1684 |
|
|
1685 |
lemma imethds_declI: "\<lbrakk>m \<in> imethds G I sig; ws_prog G; is_iface G I\<rbrakk> \<Longrightarrow>
|
|
1686 |
(\<exists>i. iface G (decliface m) = Some i \<and>
|
|
1687 |
table_of (imethods i) sig = Some (mthd m)) \<and>
|
|
1688 |
(I,decliface m) \<in> (subint1 G)^* \<and> m \<in> imethds G (decliface m) sig"
|
|
1689 |
apply (erule make_imp)
|
|
1690 |
apply (rule ws_subint1_induct, assumption, assumption)
|
|
1691 |
apply (subst imethds_rec, erule conjunct1, assumption)
|
|
1692 |
apply (force elim: imethds_norec intro: rtrancl_into_rtrancl2)
|
|
1693 |
done
|
|
1694 |
|
|
1695 |
lemma imethds_cases [consumes 3, case_names NewMethod InheritedMethod]:
|
|
1696 |
(assumes im: "im \<in> imethds G I sig" and
|
|
1697 |
ifI: "iface G I = Some i" and
|
|
1698 |
ws: "ws_prog G" and
|
|
1699 |
hyp_new: "table_of (map (\<lambda>(s, mh). (s, I, mh)) (imethods i)) sig
|
|
1700 |
= Some im \<Longrightarrow> P" and
|
|
1701 |
hyp_inh: "\<And> J. \<lbrakk>J \<in> set (isuperIfs i); im \<in> imethds G J sig\<rbrakk> \<Longrightarrow> P"
|
|
1702 |
) "P"
|
|
1703 |
proof -
|
|
1704 |
from ifI ws im hyp_new hyp_inh
|
|
1705 |
show "P"
|
|
1706 |
by (auto simp add: imethds_rec)
|
|
1707 |
qed
|
|
1708 |
|
|
1709 |
subsection "accimethd"
|
|
1710 |
|
|
1711 |
lemma accimethds_simp [simp]:
|
|
1712 |
"G\<turnstile>Iface I accessible_in pack \<Longrightarrow> accimethds G pack I = imethds G I"
|
|
1713 |
by (simp add: accimethds_def)
|
|
1714 |
|
|
1715 |
lemma accimethdsD:
|
|
1716 |
"im \<in> accimethds G pack I sig
|
|
1717 |
\<Longrightarrow> im \<in> imethds G I sig \<and> G\<turnstile>Iface I accessible_in pack"
|
|
1718 |
by (auto simp add: accimethds_def)
|
|
1719 |
|
|
1720 |
lemma accimethdsI:
|
|
1721 |
"\<lbrakk>im \<in> imethds G I sig;G\<turnstile>Iface I accessible_in pack\<rbrakk>
|
|
1722 |
\<Longrightarrow> im \<in> accimethds G pack I sig"
|
|
1723 |
by simp
|
|
1724 |
|
|
1725 |
subsection "methd"
|
|
1726 |
|
|
1727 |
lemma methd_rec: "\<lbrakk>class G C = Some c; ws_prog G\<rbrakk> \<Longrightarrow>
|
|
1728 |
methd G C
|
|
1729 |
= (if C = Object
|
|
1730 |
then empty
|
|
1731 |
else filter_tab (\<lambda>sig m. G\<turnstile>C inherits method sig m)
|
|
1732 |
(methd G (super c)))
|
|
1733 |
++ table_of (map (\<lambda>(s,m). (s,C,m)) (methods c))"
|
|
1734 |
apply (unfold methd_def)
|
|
1735 |
apply (erule class_rec [THEN trans], assumption)
|
|
1736 |
apply (simp)
|
|
1737 |
done
|
|
1738 |
|
|
1739 |
(* local lemma *)
|
|
1740 |
lemma methd_norec:
|
|
1741 |
"\<lbrakk>class G declC = Some c; ws_prog G;table_of (methods c) sig = Some m\<rbrakk>
|
|
1742 |
\<Longrightarrow> methd G declC sig = Some (declC, m)"
|
|
1743 |
apply (simp only: methd_rec)
|
|
1744 |
apply (rule disjI1 [THEN override_Some_iff [THEN iffD2]])
|
|
1745 |
apply (auto elim: table_of_map_SomeI)
|
|
1746 |
done
|
|
1747 |
|
|
1748 |
|
|
1749 |
lemma methd_declC:
|
|
1750 |
"\<lbrakk>methd G C sig = Some m; ws_prog G;is_class G C\<rbrakk> \<Longrightarrow>
|
|
1751 |
(\<exists>d. class G (declclass m)=Some d \<and> table_of (methods d) sig=Some (mthd m)) \<and>
|
|
1752 |
G\<turnstile>C \<preceq>\<^sub>C (declclass m) \<and> methd G (declclass m) sig = Some m"
|
|
1753 |
apply (erule make_imp)
|
|
1754 |
apply (rule ws_subcls1_induct, assumption, assumption)
|
|
1755 |
apply (subst methd_rec, assumption)
|
|
1756 |
apply (case_tac "Ca=Object")
|
|
1757 |
apply (force elim: methd_norec )
|
|
1758 |
|
|
1759 |
apply simp
|
|
1760 |
apply (case_tac "table_of (map (\<lambda>(s, m). (s, Ca, m)) (methods c)) sig")
|
|
1761 |
apply (force intro: rtrancl_into_rtrancl2)
|
|
1762 |
|
|
1763 |
apply (auto intro: methd_norec)
|
|
1764 |
done
|
|
1765 |
|
|
1766 |
lemma methd_inheritedD:
|
|
1767 |
"\<lbrakk>class G C = Some c; ws_prog G;methd G C sig = Some m\<rbrakk>
|
|
1768 |
\<Longrightarrow> (declclass m \<noteq> C \<longrightarrow> G \<turnstile>C inherits method sig m)"
|
|
1769 |
by (auto simp add: methd_rec)
|
|
1770 |
|
|
1771 |
lemma methd_diff_cls:
|
|
1772 |
"\<lbrakk>ws_prog G; is_class G C; is_class G D;
|
|
1773 |
methd G C sig = m; methd G D sig = n; m\<noteq>n
|
|
1774 |
\<rbrakk> \<Longrightarrow> C\<noteq>D"
|
|
1775 |
by (auto simp add: methd_rec)
|
|
1776 |
|
|
1777 |
lemma method_declared_inI:
|
|
1778 |
"\<lbrakk>table_of (methods c) sig = Some m; class G C = Some c\<rbrakk>
|
|
1779 |
\<Longrightarrow> G\<turnstile>mdecl (sig,m) declared_in C"
|
|
1780 |
by (auto simp add: cdeclaredmethd_def declared_in_def)
|
|
1781 |
|
|
1782 |
lemma methd_declared_in_declclass:
|
|
1783 |
"\<lbrakk>methd G C sig = Some m; ws_prog G;is_class G C\<rbrakk>
|
|
1784 |
\<Longrightarrow> G\<turnstile>Methd sig m declared_in (declclass m)"
|
|
1785 |
by (auto dest: methd_declC method_declared_inI)
|
|
1786 |
|
|
1787 |
lemma member_methd:
|
|
1788 |
(assumes member_of: "G\<turnstile>Methd sig m member_of C" and
|
|
1789 |
ws: "ws_prog G"
|
|
1790 |
) "methd G C sig = Some m"
|
|
1791 |
proof -
|
|
1792 |
from member_of
|
|
1793 |
have iscls_C: "is_class G C"
|
|
1794 |
by (rule member_of_is_classD)
|
|
1795 |
from iscls_C ws member_of
|
|
1796 |
show ?thesis (is "?Methd C")
|
|
1797 |
proof (induct rule: ws_class_induct')
|
|
1798 |
case (Object co)
|
|
1799 |
assume "G \<turnstile>Methd sig m member_of Object"
|
|
1800 |
then have "G\<turnstile>Methd sig m declared_in Object \<and> declclass m = Object"
|
|
1801 |
by (cases set: members) (cases m, auto dest: subcls1D)
|
|
1802 |
with ws Object
|
|
1803 |
show "?Methd Object"
|
|
1804 |
by (cases m)
|
|
1805 |
(auto simp add: declared_in_def cdeclaredmethd_def methd_rec
|
|
1806 |
intro: table_of_mapconst_SomeI)
|
|
1807 |
next
|
|
1808 |
case (Subcls C c)
|
|
1809 |
assume clsC: "class G C = Some c" and
|
|
1810 |
neq_C_Obj: "C \<noteq> Object" and
|
|
1811 |
hyp: "G \<turnstile>Methd sig m member_of super c \<Longrightarrow> ?Methd (super c)" and
|
|
1812 |
member_of: "G \<turnstile>Methd sig m member_of C"
|
|
1813 |
from member_of
|
|
1814 |
show "?Methd C"
|
|
1815 |
proof (cases)
|
|
1816 |
case (Immediate Ca membr)
|
|
1817 |
then have "Ca=C" "membr = method sig m" and
|
|
1818 |
"G\<turnstile>Methd sig m declared_in C" "declclass m = C"
|
|
1819 |
by (cases m,auto)
|
|
1820 |
with clsC
|
|
1821 |
have "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig = Some m"
|
|
1822 |
by (cases m)
|
|
1823 |
(auto simp add: declared_in_def cdeclaredmethd_def
|
|
1824 |
intro: table_of_mapconst_SomeI)
|
|
1825 |
with clsC neq_C_Obj ws
|
|
1826 |
show ?thesis
|
|
1827 |
by (simp add: methd_rec)
|
|
1828 |
next
|
|
1829 |
case (Inherited Ca S membr)
|
|
1830 |
with clsC
|
|
1831 |
have eq_Ca_C: "Ca=C" and
|
|
1832 |
undecl: "G\<turnstile>mid sig undeclared_in C" and
|
|
1833 |
super: "G \<turnstile>Methd sig m member_of (super c)"
|
|
1834 |
by (auto dest: subcls1D)
|
|
1835 |
from eq_Ca_C clsC undecl
|
|
1836 |
have "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig = None"
|
|
1837 |
by (auto simp add: undeclared_in_def cdeclaredmethd_def
|
|
1838 |
intro: table_of_mapconst_NoneI)
|
|
1839 |
moreover
|
|
1840 |
from Inherited have "G \<turnstile> C inherits (method sig m)"
|
|
1841 |
by (auto simp add: inherits_def)
|
|
1842 |
moreover
|
|
1843 |
note clsC neq_C_Obj ws super hyp
|
|
1844 |
ultimately
|
|
1845 |
show ?thesis
|
|
1846 |
by (auto simp add: methd_rec intro: filter_tab_SomeI)
|
|
1847 |
qed
|
|
1848 |
qed
|
|
1849 |
qed
|
|
1850 |
|
|
1851 |
(*unused*)
|
|
1852 |
lemma finite_methd:"ws_prog G \<Longrightarrow> finite {methd G C sig |sig C. is_class G C}"
|
|
1853 |
apply (rule finite_is_class [THEN finite_SetCompr2])
|
|
1854 |
apply (intro strip)
|
|
1855 |
apply (erule_tac ws_subcls1_induct, assumption)
|
|
1856 |
apply (subst methd_rec)
|
|
1857 |
apply (assumption)
|
|
1858 |
apply (auto intro!: finite_range_map_of finite_range_filter_tab finite_range_map_of_override)
|
|
1859 |
done
|
|
1860 |
|
|
1861 |
lemma finite_dom_methd:
|
|
1862 |
"\<lbrakk>ws_prog G; is_class G C\<rbrakk> \<Longrightarrow> finite (dom (methd G C))"
|
|
1863 |
apply (erule_tac ws_subcls1_induct)
|
|
1864 |
apply assumption
|
|
1865 |
apply (subst methd_rec)
|
|
1866 |
apply (assumption)
|
|
1867 |
apply (auto intro!: finite_dom_map_of finite_dom_filter_tab)
|
|
1868 |
done
|
|
1869 |
|
|
1870 |
|
|
1871 |
subsection "accmethd"
|
|
1872 |
|
|
1873 |
lemma accmethd_SomeD:
|
|
1874 |
"accmethd G S C sig = Some m
|
|
1875 |
\<Longrightarrow> methd G C sig = Some m \<and> G\<turnstile>method sig m of C accessible_from S"
|
|
1876 |
by (auto simp add: accmethd_def dest: filter_tab_SomeD)
|
|
1877 |
|
|
1878 |
lemma accmethd_SomeI:
|
|
1879 |
"\<lbrakk>methd G C sig = Some m; G\<turnstile>method sig m of C accessible_from S\<rbrakk>
|
|
1880 |
\<Longrightarrow> accmethd G S C sig = Some m"
|
|
1881 |
by (auto simp add: accmethd_def intro: filter_tab_SomeI)
|
|
1882 |
|
|
1883 |
lemma accmethd_declC:
|
|
1884 |
"\<lbrakk>accmethd G S C sig = Some m; ws_prog G; is_class G C\<rbrakk> \<Longrightarrow>
|
|
1885 |
(\<exists>d. class G (declclass m)=Some d \<and>
|
|
1886 |
table_of (methods d) sig=Some (mthd m)) \<and>
|
|
1887 |
G\<turnstile>C \<preceq>\<^sub>C (declclass m) \<and> methd G (declclass m) sig = Some m \<and>
|
|
1888 |
G\<turnstile>method sig m of C accessible_from S"
|
|
1889 |
by (auto dest: accmethd_SomeD methd_declC accmethd_SomeI)
|
|
1890 |
|
|
1891 |
|
|
1892 |
lemma finite_dom_accmethd:
|
|
1893 |
"\<lbrakk>ws_prog G; is_class G C\<rbrakk> \<Longrightarrow> finite (dom (accmethd G S C))"
|
|
1894 |
by (auto simp add: accmethd_def intro: finite_dom_filter_tab finite_dom_methd)
|
|
1895 |
|
|
1896 |
|
|
1897 |
subsection "dynmethd"
|
|
1898 |
|
|
1899 |
lemma dynmethd_rec:
|
|
1900 |
"\<lbrakk>class G dynC = Some c; ws_prog G\<rbrakk> \<Longrightarrow>
|
|
1901 |
dynmethd G statC dynC sig
|
|
1902 |
= (if G\<turnstile>dynC \<preceq>\<^sub>C statC
|
|
1903 |
then (case methd G statC sig of
|
|
1904 |
None \<Rightarrow> None
|
|
1905 |
| Some statM
|
|
1906 |
\<Rightarrow> (case methd G dynC sig of
|
|
1907 |
None \<Rightarrow> dynmethd G statC (super c) sig
|
|
1908 |
| Some dynM \<Rightarrow>
|
|
1909 |
(if G,sig\<turnstile> dynM overrides statM \<or> dynM = statM
|
|
1910 |
then Some dynM
|
|
1911 |
else (dynmethd G statC (super c) sig)
|
|
1912 |
)))
|
|
1913 |
else None)"
|
|
1914 |
(is "_ \<Longrightarrow> _ \<Longrightarrow> ?Dynmethd_def dynC sig = ?Dynmethd_rec dynC c sig")
|
|
1915 |
proof -
|
|
1916 |
assume clsDynC: "class G dynC = Some c" and
|
|
1917 |
ws: "ws_prog G"
|
|
1918 |
then show "?Dynmethd_def dynC sig = ?Dynmethd_rec dynC c sig"
|
|
1919 |
proof (induct rule: ws_class_induct'')
|
|
1920 |
case (Object co)
|
|
1921 |
show "?Dynmethd_def Object sig = ?Dynmethd_rec Object co sig"
|
|
1922 |
proof (cases "G\<turnstile>Object \<preceq>\<^sub>C statC")
|
|
1923 |
case False
|
|
1924 |
then show ?thesis by (simp add: dynmethd_def)
|
|
1925 |
next
|
|
1926 |
case True
|
|
1927 |
then have eq_statC_Obj: "statC = Object" ..
|
|
1928 |
show ?thesis
|
|
1929 |
proof (cases "methd G statC sig")
|
|
1930 |
case None then show ?thesis by (simp add: dynmethd_def)
|
|
1931 |
next
|
|
1932 |
case Some
|
|
1933 |
with True Object ws eq_statC_Obj
|
|
1934 |
show ?thesis
|
|
1935 |
by (auto simp add: dynmethd_def class_rec
|
|
1936 |
intro: filter_tab_SomeI)
|
|
1937 |
qed
|
|
1938 |
qed
|
|
1939 |
next
|
|
1940 |
case (Subcls dynC c sc)
|
|
1941 |
show "?Dynmethd_def dynC sig = ?Dynmethd_rec dynC c sig"
|
|
1942 |
proof (cases "G\<turnstile>dynC \<preceq>\<^sub>C statC")
|
|
1943 |
case False
|
|
1944 |
then show ?thesis by (simp add: dynmethd_def)
|
|
1945 |
next
|
|
1946 |
case True
|
|
1947 |
note subclseq_dynC_statC = True
|
|
1948 |
show ?thesis
|
|
1949 |
proof (cases "methd G statC sig")
|
|
1950 |
case None then show ?thesis by (simp add: dynmethd_def)
|
|
1951 |
next
|
|
1952 |
case (Some statM)
|
|
1953 |
note statM = Some
|
|
1954 |
let "?filter C" =
|
|
1955 |
"filter_tab
|
|
1956 |
(\<lambda>_ dynM. G,sig \<turnstile> dynM overrides statM \<or> dynM = statM)
|
|
1957 |
(methd G C)"
|
|
1958 |
let "?class_rec C" =
|
|
1959 |
"(class_rec (G, C) empty
|
|
1960 |
(\<lambda>C c subcls_mthds. subcls_mthds ++ (?filter C)))"
|
|
1961 |
from statM Subcls ws subclseq_dynC_statC
|
|
1962 |
have dynmethd_dynC_def:
|
|
1963 |
"?Dynmethd_def dynC sig =
|
|
1964 |
((?class_rec (super c))
|
|
1965 |
++
|
|
1966 |
(?filter dynC)) sig"
|
|
1967 |
by (simp (no_asm_simp) only: dynmethd_def class_rec)
|
|
1968 |
auto
|
|
1969 |
show ?thesis
|
|
1970 |
proof (cases "dynC = statC")
|
|
1971 |
case True
|
|
1972 |
with subclseq_dynC_statC statM dynmethd_dynC_def
|
|
1973 |
have "?Dynmethd_def dynC sig = Some statM"
|
|
1974 |
by (auto intro: override_find_right filter_tab_SomeI)
|
|
1975 |
with subclseq_dynC_statC True Some
|
|
1976 |
show ?thesis
|
|
1977 |
by auto
|
|
1978 |
next
|
|
1979 |
case False
|
|
1980 |
with subclseq_dynC_statC Subcls
|
|
1981 |
have subclseq_super_statC: "G\<turnstile>(super c) \<preceq>\<^sub>C statC"
|
|
1982 |
by (blast dest: subclseq_superD)
|
|
1983 |
show ?thesis
|
|
1984 |
proof (cases "methd G dynC sig")
|
|
1985 |
case None
|
|
1986 |
then have "?filter dynC sig = None"
|
|
1987 |
by (rule filter_tab_None)
|
|
1988 |
then have "?Dynmethd_def dynC sig=?class_rec (super c) sig"
|
|
1989 |
by (simp add: dynmethd_dynC_def)
|
|
1990 |
with subclseq_super_statC statM None
|
|
1991 |
have "?Dynmethd_def dynC sig = ?Dynmethd_def (super c) sig"
|
|
1992 |
by (auto simp add: empty_def dynmethd_def)
|
|
1993 |
with None subclseq_dynC_statC statM
|
|
1994 |
show ?thesis
|
|
1995 |
by simp
|
|
1996 |
next
|
|
1997 |
case (Some dynM)
|
|
1998 |
note dynM = Some
|
|
1999 |
let ?Termination = "G \<turnstile> qmdecl sig dynM overrides qmdecl sig statM \<or>
|
|
2000 |
dynM = statM"
|
|
2001 |
show ?thesis
|
|
2002 |
proof (cases "?filter dynC sig")
|
|
2003 |
case None
|
|
2004 |
with dynM
|
|
2005 |
have no_termination: "\<not> ?Termination"
|
|
2006 |
by (simp add: filter_tab_def)
|
|
2007 |
from None
|
|
2008 |
have "?Dynmethd_def dynC sig=?class_rec (super c) sig"
|
|
2009 |
by (simp add: dynmethd_dynC_def)
|
|
2010 |
with subclseq_super_statC statM dynM no_termination
|
|
2011 |
show ?thesis
|
|
2012 |
by (auto simp add: empty_def dynmethd_def)
|
|
2013 |
next
|
|
2014 |
case Some
|
|
2015 |
with dynM
|
|
2016 |
have termination: "?Termination"
|
|
2017 |
by (auto)
|
|
2018 |
with Some dynM
|
|
2019 |
have "?Dynmethd_def dynC sig=Some dynM"
|
|
2020 |
by (auto simp add: dynmethd_dynC_def)
|
|
2021 |
with subclseq_super_statC statM dynM termination
|
|
2022 |
show ?thesis
|
|
2023 |
by (auto simp add: dynmethd_def)
|
|
2024 |
qed
|
|
2025 |
qed
|
|
2026 |
qed
|
|
2027 |
qed
|
|
2028 |
qed
|
|
2029 |
qed
|
|
2030 |
qed
|
|
2031 |
|
|
2032 |
lemma dynmethd_C_C:"\<lbrakk>is_class G C; ws_prog G\<rbrakk>
|
|
2033 |
\<Longrightarrow> dynmethd G C C sig = methd G C sig"
|
|
2034 |
apply (auto simp add: dynmethd_rec)
|
|
2035 |
done
|
|
2036 |
|
|
2037 |
lemma dynmethdSomeD:
|
|
2038 |
"\<lbrakk>dynmethd G statC dynC sig = Some dynM; is_class G dynC; ws_prog G\<rbrakk>
|
|
2039 |
\<Longrightarrow> G\<turnstile>dynC \<preceq>\<^sub>C statC \<and> (\<exists> statM. methd G statC sig = Some statM)"
|
|
2040 |
apply clarify
|
|
2041 |
apply rotate_tac
|
|
2042 |
by (auto simp add: dynmethd_rec)
|
|
2043 |
|
|
2044 |
lemma dynmethd_Some_cases [consumes 3, case_names Static Overrides]:
|
|
2045 |
(assumes dynM: "dynmethd G statC dynC sig = Some dynM" and
|
|
2046 |
is_cls_dynC: "is_class G dynC" and
|
|
2047 |
ws: "ws_prog G" and
|
|
2048 |
hyp_static: "methd G statC sig = Some dynM \<Longrightarrow> P" and
|
|
2049 |
hyp_override: "\<And> statM. \<lbrakk>methd G statC sig = Some statM;dynM\<noteq>statM;
|
|
2050 |
G,sig\<turnstile>dynM overrides statM\<rbrakk> \<Longrightarrow> P"
|
|
2051 |
) "P"
|
|
2052 |
proof -
|
|
2053 |
from is_cls_dynC obtain dc where clsDynC: "class G dynC = Some dc" by blast
|
|
2054 |
from clsDynC ws dynM hyp_static hyp_override
|
|
2055 |
show "P"
|
|
2056 |
proof (induct rule: ws_class_induct)
|
|
2057 |
case (Object co)
|
|
2058 |
with ws have "statC = Object"
|
|
2059 |
by (auto simp add: dynmethd_rec)
|
|
2060 |
with ws Object show ?thesis by (auto simp add: dynmethd_C_C)
|
|
2061 |
next
|
|
2062 |
case (Subcls C c)
|
|
2063 |
with ws show ?thesis
|
|
2064 |
by (auto simp add: dynmethd_rec)
|
|
2065 |
qed
|
|
2066 |
qed
|
|
2067 |
|
|
2068 |
lemma no_override_in_Object:
|
|
2069 |
(assumes dynM: "dynmethd G statC dynC sig = Some dynM" and
|
|
2070 |
is_cls_dynC: "is_class G dynC" and
|
|
2071 |
ws: "ws_prog G" and
|
|
2072 |
statM: "methd G statC sig = Some statM" and
|
|
2073 |
neq_dynM_statM: "dynM\<noteq>statM"
|
|
2074 |
)
|
|
2075 |
"dynC \<noteq> Object"
|
|
2076 |
proof -
|
|
2077 |
from is_cls_dynC obtain dc where clsDynC: "class G dynC = Some dc" by blast
|
|
2078 |
from clsDynC ws dynM statM neq_dynM_statM
|
|
2079 |
show ?thesis (is "?P dynC")
|
|
2080 |
proof (induct rule: ws_class_induct)
|
|
2081 |
case (Object co)
|
|
2082 |
with ws have "statC = Object"
|
|
2083 |
by (auto simp add: dynmethd_rec)
|
|
2084 |
with ws Object show "?P Object" by (auto simp add: dynmethd_C_C)
|
|
2085 |
next
|
|
2086 |
case (Subcls dynC c)
|
|
2087 |
with ws show "?P dynC"
|
|
2088 |
by (auto simp add: dynmethd_rec)
|
|
2089 |
qed
|
|
2090 |
qed
|
|
2091 |
|
|
2092 |
|
|
2093 |
lemma dynmethd_Some_rec_cases [consumes 3,
|
|
2094 |
case_names Static Override Recursion]:
|
|
2095 |
(assumes dynM: "dynmethd G statC dynC sig = Some dynM" and
|
|
2096 |
clsDynC: "class G dynC = Some c" and
|
|
2097 |
ws: "ws_prog G" and
|
|
2098 |
hyp_static: "methd G statC sig = Some dynM \<Longrightarrow> P" and
|
|
2099 |
hyp_override: "\<And> statM. \<lbrakk>methd G statC sig = Some statM;
|
|
2100 |
methd G dynC sig = Some dynM; statM\<noteq>dynM;
|
|
2101 |
G,sig\<turnstile> dynM overrides statM\<rbrakk> \<Longrightarrow> P" and
|
|
2102 |
|
|
2103 |
hyp_recursion: "\<lbrakk>dynC\<noteq>Object;
|
|
2104 |
dynmethd G statC (super c) sig = Some dynM\<rbrakk> \<Longrightarrow> P"
|
|
2105 |
) "P"
|
|
2106 |
proof -
|
|
2107 |
from clsDynC have "is_class G dynC" by simp
|
|
2108 |
note no_override_in_Object' = no_override_in_Object [OF dynM this ws]
|
|
2109 |
from ws clsDynC dynM hyp_static hyp_override hyp_recursion
|
|
2110 |
show ?thesis
|
|
2111 |
by (auto simp add: dynmethd_rec dest: no_override_in_Object')
|
|
2112 |
qed
|
|
2113 |
|
|
2114 |
lemma dynmethd_declC:
|
|
2115 |
"\<lbrakk>dynmethd G statC dynC sig = Some m;
|
|
2116 |
is_class G statC;ws_prog G
|
|
2117 |
\<rbrakk> \<Longrightarrow>
|
|
2118 |
(\<exists>d. class G (declclass m)=Some d \<and> table_of (methods d) sig=Some (mthd m)) \<and>
|
|
2119 |
G\<turnstile>dynC \<preceq>\<^sub>C (declclass m) \<and> methd G (declclass m) sig = Some m"
|
|
2120 |
proof -
|
|
2121 |
assume is_cls_statC: "is_class G statC"
|
|
2122 |
assume ws: "ws_prog G"
|
|
2123 |
assume m: "dynmethd G statC dynC sig = Some m"
|
|
2124 |
from m
|
|
2125 |
have "G\<turnstile>dynC \<preceq>\<^sub>C statC" by (auto simp add: dynmethd_def)
|
|
2126 |
from this is_cls_statC
|
|
2127 |
have is_cls_dynC: "is_class G dynC" by (rule subcls_is_class2)
|
|
2128 |
from is_cls_dynC ws m
|
|
2129 |
show ?thesis (is "?P dynC")
|
|
2130 |
proof (induct rule: ws_class_induct')
|
|
2131 |
case (Object co)
|
|
2132 |
with ws have "statC=Object" by (auto simp add: dynmethd_rec)
|
|
2133 |
with ws Object
|
|
2134 |
show "?P Object"
|
|
2135 |
by (auto simp add: dynmethd_C_C dest: methd_declC)
|
|
2136 |
next
|
|
2137 |
case (Subcls dynC c)
|
|
2138 |
assume hyp: "dynmethd G statC (super c) sig = Some m \<Longrightarrow> ?P (super c)" and
|
|
2139 |
clsDynC: "class G dynC = Some c" and
|
|
2140 |
m': "dynmethd G statC dynC sig = Some m" and
|
|
2141 |
neq_dynC_Obj: "dynC \<noteq> Object"
|
|
2142 |
from ws this obtain statM where
|
|
2143 |
subclseq_dynC_statC: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
|
|
2144 |
statM: "methd G statC sig = Some statM"
|
|
2145 |
by (blast dest: dynmethdSomeD)
|
|
2146 |
from clsDynC neq_dynC_Obj
|
|
2147 |
have subclseq_dynC_super: "G\<turnstile>dynC \<preceq>\<^sub>C (super c)"
|
|
2148 |
by (auto intro: subcls1I)
|
|
2149 |
from m' clsDynC ws
|
|
2150 |
show "?P dynC"
|
|
2151 |
proof (cases rule: dynmethd_Some_rec_cases)
|
|
2152 |
case Static
|
|
2153 |
with is_cls_statC ws subclseq_dynC_statC
|
|
2154 |
show ?thesis
|
|
2155 |
by (auto intro: rtrancl_trans dest: methd_declC)
|
|
2156 |
next
|
|
2157 |
case Override
|
|
2158 |
with clsDynC ws
|
|
2159 |
show ?thesis
|
|
2160 |
by (auto dest: methd_declC)
|
|
2161 |
next
|
|
2162 |
case Recursion
|
|
2163 |
with hyp subclseq_dynC_super
|
|
2164 |
show ?thesis
|
|
2165 |
by (auto intro: rtrancl_trans)
|
|
2166 |
qed
|
|
2167 |
qed
|
|
2168 |
qed
|
|
2169 |
|
|
2170 |
lemma methd_Some_dynmethd_Some:
|
|
2171 |
(assumes statM: "methd G statC sig = Some statM" and
|
|
2172 |
subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
|
|
2173 |
is_cls_statC: "is_class G statC" and
|
|
2174 |
ws: "ws_prog G"
|
|
2175 |
) "\<exists> dynM. dynmethd G statC dynC sig = Some dynM"
|
|
2176 |
(is "?P dynC")
|
|
2177 |
proof -
|
|
2178 |
from subclseq is_cls_statC
|
|
2179 |
have is_cls_dynC: "is_class G dynC" by (rule subcls_is_class2)
|
|
2180 |
then obtain dc where
|
|
2181 |
clsDynC: "class G dynC = Some dc" by blast
|
|
2182 |
from clsDynC ws subclseq
|
|
2183 |
show "?thesis"
|
|
2184 |
proof (induct rule: ws_class_induct)
|
|
2185 |
case (Object co)
|
|
2186 |
with ws have "statC = Object"
|
|
2187 |
by (auto)
|
|
2188 |
with ws Object statM
|
|
2189 |
show "?P Object"
|
|
2190 |
by (auto simp add: dynmethd_C_C)
|
|
2191 |
next
|
|
2192 |
case (Subcls dynC dc)
|
|
2193 |
assume clsDynC': "class G dynC = Some dc"
|
|
2194 |
assume neq_dynC_Obj: "dynC \<noteq> Object"
|
|
2195 |
assume hyp: "G\<turnstile>super dc\<preceq>\<^sub>C statC \<Longrightarrow> ?P (super dc)"
|
|
2196 |
assume subclseq': "G\<turnstile>dynC\<preceq>\<^sub>C statC"
|
|
2197 |
then
|
|
2198 |
show "?P dynC"
|
|
2199 |
proof (cases rule: subclseq_cases)
|
|
2200 |
case Eq
|
|
2201 |
with ws statM clsDynC'
|
|
2202 |
show ?thesis
|
|
2203 |
by (auto simp add: dynmethd_rec)
|
|
2204 |
next
|
|
2205 |
case Subcls
|
|
2206 |
assume "G\<turnstile>dynC\<prec>\<^sub>C statC"
|
|
2207 |
from this clsDynC'
|
|
2208 |
have "G\<turnstile>super dc\<preceq>\<^sub>C statC" by (rule subcls_superD)
|
|
2209 |
with hyp ws clsDynC' subclseq' statM
|
|
2210 |
show ?thesis
|
|
2211 |
by (auto simp add: dynmethd_rec)
|
|
2212 |
qed
|
|
2213 |
qed
|
|
2214 |
qed
|
|
2215 |
|
|
2216 |
lemma dynmethd_cases [consumes 4, case_names Static Overrides]:
|
|
2217 |
(assumes statM: "methd G statC sig = Some statM" and
|
|
2218 |
subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
|
|
2219 |
is_cls_statC: "is_class G statC" and
|
|
2220 |
ws: "ws_prog G" and
|
|
2221 |
hyp_static: "dynmethd G statC dynC sig = Some statM \<Longrightarrow> P" and
|
|
2222 |
hyp_override: "\<And> dynM. \<lbrakk>dynmethd G statC dynC sig = Some dynM;
|
|
2223 |
dynM\<noteq>statM;
|
|
2224 |
G,sig\<turnstile>dynM overrides statM\<rbrakk> \<Longrightarrow> P"
|
|
2225 |
) "P"
|
|
2226 |
proof -
|
|
2227 |
from subclseq is_cls_statC
|
|
2228 |
have is_cls_dynC: "is_class G dynC" by (rule subcls_is_class2)
|
|
2229 |
then obtain dc where
|
|
2230 |
clsDynC: "class G dynC = Some dc" by blast
|
|
2231 |
from statM subclseq is_cls_statC ws
|
|
2232 |
obtain dynM
|
|
2233 |
where dynM: "dynmethd G statC dynC sig = Some dynM"
|
|
2234 |
by (blast dest: methd_Some_dynmethd_Some)
|
|
2235 |
from dynM is_cls_dynC ws
|
|
2236 |
show ?thesis
|
|
2237 |
proof (cases rule: dynmethd_Some_cases)
|
|
2238 |
case Static
|
|
2239 |
with hyp_static dynM statM show ?thesis by simp
|
|
2240 |
next
|
|
2241 |
case Overrides
|
|
2242 |
with hyp_override dynM statM show ?thesis by simp
|
|
2243 |
qed
|
|
2244 |
qed
|
|
2245 |
|
|
2246 |
lemma ws_dynmethd:
|
|
2247 |
(assumes statM: "methd G statC sig = Some statM" and
|
|
2248 |
subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
|
|
2249 |
is_cls_statC: "is_class G statC" and
|
|
2250 |
ws: "ws_prog G"
|
|
2251 |
)
|
|
2252 |
"\<exists> dynM. dynmethd G statC dynC sig = Some dynM \<and>
|
|
2253 |
is_static dynM = is_static statM \<and> G\<turnstile>resTy dynM\<preceq>resTy statM"
|
|
2254 |
proof -
|
|
2255 |
from statM subclseq is_cls_statC ws
|
|
2256 |
show ?thesis
|
|
2257 |
proof (cases rule: dynmethd_cases)
|
|
2258 |
case Static
|
|
2259 |
with statM
|
|
2260 |
show ?thesis
|
|
2261 |
by simp
|
|
2262 |
next
|
|
2263 |
case Overrides
|
|
2264 |
with ws
|
|
2265 |
show ?thesis
|
|
2266 |
by (auto dest: ws_overrides_commonD)
|
|
2267 |
qed
|
|
2268 |
qed
|
|
2269 |
|
|
2270 |
(*
|
|
2271 |
lemma dom_dynmethd:
|
|
2272 |
"dom (dynmethd G statC dynC) \<subseteq> dom (methd G statC) \<union> dom (methd G dynC)"
|
|
2273 |
by (auto simp add: dynmethd_def dom_def)
|
|
2274 |
|
|
2275 |
lemma finite_dom_dynmethd:
|
|
2276 |
"\<lbrakk>ws_prog G; is_class G statC; is_class G dynC\<rbrakk>
|
|
2277 |
\<Longrightarrow> finite (dom (dynmethd G statC dynC))"
|
|
2278 |
apply (rule_tac B="dom (methd G statC) \<union> dom (methd G dynC)" in finite_subset)
|
|
2279 |
apply (rule dom_dynmethd)
|
|
2280 |
apply (rule finite_UnI)
|
|
2281 |
apply (drule (2) finite_dom_methd)+
|
|
2282 |
done
|
|
2283 |
*)
|
|
2284 |
(*
|
|
2285 |
lemma dynmethd_SomeD:
|
|
2286 |
"\<lbrakk>ws_prog G; is_class G statC; is_class G dynC;
|
|
2287 |
methd G statC sig = Some sm; dynmethd G statC dynC sig = Some dm; sm \<noteq> dm
|
|
2288 |
\<rbrakk> \<Longrightarrow> G\<turnstile>dynC \<prec>\<^sub>C statC \<and>
|
|
2289 |
(declclass dm \<noteq> dynC \<longrightarrow> G \<turnstile> dm accessible_through_inheritance_in dynC)"
|
|
2290 |
by (auto simp add: dynmethd_def
|
|
2291 |
dest: methd_inheritedD methd_diff_cls
|
|
2292 |
intro: rtrancl_into_trancl3)
|
|
2293 |
*)
|
|
2294 |
|
|
2295 |
subsection "dynlookup"
|
|
2296 |
|
|
2297 |
lemma dynlookup_cases [consumes 1, case_names NullT IfaceT ClassT ArrayT]:
|
|
2298 |
"\<lbrakk>dynlookup G statT dynC sig = x;
|
|
2299 |
\<lbrakk>statT = NullT ; empty sig = x \<rbrakk> \<Longrightarrow> P;
|
|
2300 |
\<And> I. \<lbrakk>statT = IfaceT I ; dynimethd G I dynC sig = x\<rbrakk> \<Longrightarrow> P;
|
|
2301 |
\<And> statC.\<lbrakk>statT = ClassT statC; dynmethd G statC dynC sig = x\<rbrakk> \<Longrightarrow> P;
|
|
2302 |
\<And> ty. \<lbrakk>statT = ArrayT ty ; dynmethd G Object dynC sig = x\<rbrakk> \<Longrightarrow> P
|
|
2303 |
\<rbrakk> \<Longrightarrow> P"
|
|
2304 |
by (cases statT) (auto simp add: dynlookup_def)
|
|
2305 |
|
|
2306 |
subsection "fields"
|
|
2307 |
|
|
2308 |
lemma fields_rec: "\<lbrakk>class G C = Some c; ws_prog G\<rbrakk> \<Longrightarrow>
|
|
2309 |
fields G C = map (\<lambda>(fn,ft). ((fn,C),ft)) (cfields c) @
|
|
2310 |
(if C = Object then [] else fields G (super c))"
|
|
2311 |
apply (simp only: fields_def)
|
|
2312 |
apply (erule class_rec [THEN trans])
|
|
2313 |
apply assumption
|
|
2314 |
apply clarsimp
|
|
2315 |
done
|
|
2316 |
|
|
2317 |
(* local lemma *)
|
|
2318 |
lemma fields_norec:
|
|
2319 |
"\<lbrakk>class G fd = Some c; ws_prog G; table_of (cfields c) fn = Some f\<rbrakk>
|
|
2320 |
\<Longrightarrow> table_of (fields G fd) (fn,fd) = Some f"
|
|
2321 |
apply (subst fields_rec)
|
|
2322 |
apply assumption+
|
|
2323 |
apply (subst map_of_override [symmetric])
|
|
2324 |
apply (rule disjI1 [THEN override_Some_iff [THEN iffD2]])
|
|
2325 |
apply (auto elim: table_of_map2_SomeI)
|
|
2326 |
done
|
|
2327 |
|
|
2328 |
(* local lemma *)
|
|
2329 |
lemma table_of_fieldsD:
|
|
2330 |
"table_of (map (\<lambda>(fn,ft). ((fn,C),ft)) (cfields c)) efn = Some f
|
|
2331 |
\<Longrightarrow> (declclassf efn) = C \<and> table_of (cfields c) (fname efn) = Some f"
|
|
2332 |
apply (case_tac "efn")
|
|
2333 |
by auto
|
|
2334 |
|
|
2335 |
lemma fields_declC:
|
|
2336 |
"\<lbrakk>table_of (fields G C) efn = Some f; ws_prog G; is_class G C\<rbrakk> \<Longrightarrow>
|
|
2337 |
(\<exists>d. class G (declclassf efn) = Some d \<and>
|
|
2338 |
table_of (cfields d) (fname efn)=Some f) \<and>
|
|
2339 |
G\<turnstile>C \<preceq>\<^sub>C (declclassf efn) \<and> table_of (fields G (declclassf efn)) efn = Some f"
|
|
2340 |
apply (erule make_imp)
|
|
2341 |
apply (rule ws_subcls1_induct, assumption, assumption)
|
|
2342 |
apply (subst fields_rec, assumption)
|
|
2343 |
apply clarify
|
|
2344 |
apply (simp only: map_of_override [symmetric] del: map_of_override)
|
|
2345 |
apply (case_tac "table_of (map (split (\<lambda>fn. Pair (fn, Ca))) (cfields c)) efn")
|
|
2346 |
apply (force intro:rtrancl_into_rtrancl2 simp add: override_def)
|
|
2347 |
|
|
2348 |
apply (frule_tac fd="Ca" in fields_norec)
|
|
2349 |
apply assumption
|
|
2350 |
apply blast
|
|
2351 |
apply (frule table_of_fieldsD)
|
|
2352 |
apply (frule_tac n="table_of (map (split (\<lambda>fn. Pair (fn, Ca))) (cfields c))"
|
|
2353 |
and m="table_of (if Ca = Object then [] else fields G (super c))"
|
|
2354 |
in override_find_right)
|
|
2355 |
apply (case_tac "efn")
|
|
2356 |
apply (simp)
|
|
2357 |
done
|
|
2358 |
|
|
2359 |
lemma fields_emptyI: "\<And>y. \<lbrakk>ws_prog G; class G C = Some c;cfields c = [];
|
|
2360 |
C \<noteq> Object \<longrightarrow> class G (super c) = Some y \<and> fields G (super c) = []\<rbrakk> \<Longrightarrow>
|
|
2361 |
fields G C = []"
|
|
2362 |
apply (subst fields_rec)
|
|
2363 |
apply assumption
|
|
2364 |
apply auto
|
|
2365 |
done
|
|
2366 |
|
|
2367 |
(* easier than with table_of *)
|
|
2368 |
lemma fields_mono_lemma:
|
|
2369 |
"\<lbrakk>x \<in> set (fields G C); G\<turnstile>D \<preceq>\<^sub>C C; ws_prog G\<rbrakk>
|
|
2370 |
\<Longrightarrow> x \<in> set (fields G D)"
|
|
2371 |
apply (erule make_imp)
|
|
2372 |
apply (erule converse_rtrancl_induct)
|
|
2373 |
apply fast
|
|
2374 |
apply (drule subcls1D)
|
|
2375 |
apply clarsimp
|
|
2376 |
apply (subst fields_rec)
|
|
2377 |
apply auto
|
|
2378 |
done
|
|
2379 |
|
|
2380 |
|
|
2381 |
lemma ws_unique_fields_lemma:
|
|
2382 |
"\<lbrakk>(efn,fd) \<in> set (fields G (super c)); fc \<in> set (cfields c); ws_prog G;
|
|
2383 |
fname efn = fname fc; declclassf efn = C;
|
|
2384 |
class G C = Some c; C \<noteq> Object; class G (super c) = Some d\<rbrakk> \<Longrightarrow> R"
|
|
2385 |
apply (frule_tac ws_prog_cdeclD [THEN conjunct2], assumption, assumption)
|
|
2386 |
apply (drule_tac weak_map_of_SomeI)
|
|
2387 |
apply (frule_tac subcls1I [THEN subcls1_irrefl], assumption, assumption)
|
|
2388 |
apply (auto dest: fields_declC [THEN conjunct2 [THEN conjunct1[THEN rtranclD]]])
|
|
2389 |
done
|
|
2390 |
|
|
2391 |
lemma ws_unique_fields: "\<lbrakk>is_class G C; ws_prog G;
|
|
2392 |
\<And>C c. \<lbrakk>class G C = Some c\<rbrakk> \<Longrightarrow> unique (cfields c) \<rbrakk> \<Longrightarrow>
|
|
2393 |
unique (fields G C)"
|
|
2394 |
apply (rule ws_subcls1_induct, assumption, assumption)
|
|
2395 |
apply (subst fields_rec, assumption)
|
|
2396 |
apply (auto intro!: unique_map_inj injI
|
|
2397 |
elim!: unique_append ws_unique_fields_lemma fields_norec
|
|
2398 |
)
|
|
2399 |
done
|
|
2400 |
|
|
2401 |
|
|
2402 |
subsection "accfield"
|
|
2403 |
|
|
2404 |
lemma accfield_fields:
|
|
2405 |
"accfield G S C fn = Some f
|
|
2406 |
\<Longrightarrow> table_of (fields G C) (fn, declclass f) = Some (fld f)"
|
|
2407 |
apply (simp only: accfield_def Let_def)
|
|
2408 |
apply (rule table_of_remap_SomeD)
|
|
2409 |
apply (auto dest: filter_tab_SomeD)
|
|
2410 |
done
|
|
2411 |
|
|
2412 |
|
|
2413 |
lemma accfield_declC_is_class:
|
|
2414 |
"\<lbrakk>is_class G C; accfield G S C en = Some (fd, f); ws_prog G\<rbrakk> \<Longrightarrow>
|
|
2415 |
is_class G fd"
|
|
2416 |
apply (drule accfield_fields)
|
|
2417 |
apply (drule fields_declC [THEN conjunct1], assumption)
|
|
2418 |
apply auto
|
|
2419 |
done
|
|
2420 |
|
|
2421 |
lemma accfield_accessibleD:
|
|
2422 |
"accfield G S C fn = Some f \<Longrightarrow> G\<turnstile>Field fn f of C accessible_from S"
|
|
2423 |
by (auto simp add: accfield_def Let_def)
|
|
2424 |
|
|
2425 |
subsection "is methd"
|
|
2426 |
|
|
2427 |
lemma is_methdI:
|
|
2428 |
"\<lbrakk>class G C = Some y; methd G C sig = Some b\<rbrakk> \<Longrightarrow> is_methd G C sig"
|
|
2429 |
apply (unfold is_methd_def)
|
|
2430 |
apply auto
|
|
2431 |
done
|
|
2432 |
|
|
2433 |
lemma is_methdD:
|
|
2434 |
"is_methd G C sig \<Longrightarrow> class G C \<noteq> None \<and> methd G C sig \<noteq> None"
|
|
2435 |
apply (unfold is_methd_def)
|
|
2436 |
apply auto
|
|
2437 |
done
|
|
2438 |
|
|
2439 |
lemma finite_is_methd:
|
|
2440 |
"ws_prog G \<Longrightarrow> finite (Collect (split (is_methd G)))"
|
|
2441 |
apply (unfold is_methd_def)
|
|
2442 |
apply (subst SetCompr_Sigma_eq)
|
|
2443 |
apply (rule finite_is_class [THEN finite_SigmaI])
|
|
2444 |
apply (simp only: mem_Collect_eq)
|
|
2445 |
apply (fold dom_def)
|
|
2446 |
apply (erule finite_dom_methd)
|
|
2447 |
apply assumption
|
|
2448 |
done
|
|
2449 |
|
|
2450 |
section "calculation of the superclasses of a class"
|
|
2451 |
|
|
2452 |
constdefs
|
|
2453 |
superclasses:: "prog \<Rightarrow> qtname \<Rightarrow> qtname set"
|
|
2454 |
"superclasses G C \<equiv> class_rec (G,C) {}
|
|
2455 |
(\<lambda> C c superclss. (if C=Object
|
|
2456 |
then {}
|
|
2457 |
else insert (super c) superclss))"
|
|
2458 |
|
|
2459 |
lemma superclasses_rec: "\<lbrakk>class G C = Some c; ws_prog G\<rbrakk> \<Longrightarrow>
|
|
2460 |
superclasses G C
|
|
2461 |
= (if (C=Object)
|
|
2462 |
then {}
|
|
2463 |
else insert (super c) (superclasses G (super c)))"
|
|
2464 |
apply (unfold superclasses_def)
|
|
2465 |
apply (erule class_rec [THEN trans], assumption)
|
|
2466 |
apply (simp)
|
|
2467 |
done
|
|
2468 |
|
|
2469 |
lemma superclasses_mono:
|
|
2470 |
"\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D;ws_prog G; class G C = Some c;
|
|
2471 |
\<And> C c. \<lbrakk>class G C = Some c;C\<noteq>Object\<rbrakk> \<Longrightarrow> \<exists> sc. class G (super c) = Some sc;
|
|
2472 |
x\<in>superclasses G D
|
|
2473 |
\<rbrakk> \<Longrightarrow> x\<in>superclasses G C"
|
|
2474 |
proof -
|
|
2475 |
|
|
2476 |
assume ws: "ws_prog G" and
|
|
2477 |
cls_C: "class G C = Some c" and
|
|
2478 |
wf: "\<And>C c. \<lbrakk>class G C = Some c; C \<noteq> Object\<rbrakk>
|
|
2479 |
\<Longrightarrow> \<exists>sc. class G (super c) = Some sc"
|
|
2480 |
assume clsrel: "G\<turnstile>C\<prec>\<^sub>C D"
|
|
2481 |
thus "\<And> c. \<lbrakk>class G C = Some c; x\<in>superclasses G D\<rbrakk>\<Longrightarrow>
|
|
2482 |
x\<in>superclasses G C" (is "PROP ?P C"
|
|
2483 |
is "\<And> c. ?CLS C c \<Longrightarrow> ?SUP D \<Longrightarrow> ?SUP C")
|
|
2484 |
proof (induct ?P C rule: converse_trancl_induct)
|
|
2485 |
fix C c
|
|
2486 |
assume "G\<turnstile>C\<prec>\<^sub>C\<^sub>1D" "class G C = Some c" "x \<in> superclasses G D"
|
|
2487 |
with wf ws show "?SUP C"
|
|
2488 |
by (auto intro: no_subcls1_Object
|
|
2489 |
simp add: superclasses_rec subcls1_def)
|
|
2490 |
next
|
|
2491 |
fix C S c
|
|
2492 |
assume clsrel': "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S" "G\<turnstile>S \<prec>\<^sub>C D"
|
|
2493 |
and hyp : "\<And> s. \<lbrakk>class G S = Some s; x \<in> superclasses G D\<rbrakk>
|
|
2494 |
\<Longrightarrow> x \<in> superclasses G S"
|
|
2495 |
and cls_C': "class G C = Some c"
|
|
2496 |
and x: "x \<in> superclasses G D"
|
|
2497 |
moreover note wf ws
|
|
2498 |
moreover from calculation
|
|
2499 |
have "?SUP S"
|
|
2500 |
by (force intro: no_subcls1_Object simp add: subcls1_def)
|
|
2501 |
moreover from calculation
|
|
2502 |
have "super c = S"
|
|
2503 |
by (auto intro: no_subcls1_Object simp add: subcls1_def)
|
|
2504 |
ultimately show "?SUP C"
|
|
2505 |
by (auto intro: no_subcls1_Object simp add: superclasses_rec)
|
|
2506 |
qed
|
|
2507 |
qed
|
|
2508 |
|
|
2509 |
lemma subclsEval:
|
|
2510 |
"\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D;ws_prog G; class G C = Some c;
|
|
2511 |
\<And> C c. \<lbrakk>class G C = Some c;C\<noteq>Object\<rbrakk> \<Longrightarrow> \<exists> sc. class G (super c) = Some sc
|
|
2512 |
\<rbrakk> \<Longrightarrow> D\<in>superclasses G C"
|
|
2513 |
proof -
|
|
2514 |
note converse_trancl_induct
|
|
2515 |
= converse_trancl_induct [consumes 1,case_names Single Step]
|
|
2516 |
assume
|
|
2517 |
ws: "ws_prog G" and
|
|
2518 |
cls_C: "class G C = Some c" and
|
|
2519 |
wf: "\<And>C c. \<lbrakk>class G C = Some c; C \<noteq> Object\<rbrakk>
|
|
2520 |
\<Longrightarrow> \<exists>sc. class G (super c) = Some sc"
|
|
2521 |
assume clsrel: "G\<turnstile>C\<prec>\<^sub>C D"
|
|
2522 |
thus "\<And> c. class G C = Some c\<Longrightarrow> D\<in>superclasses G C"
|
|
2523 |
(is "PROP ?P C" is "\<And> c. ?CLS C c \<Longrightarrow> ?SUP C")
|
|
2524 |
proof (induct ?P C rule: converse_trancl_induct)
|
|
2525 |
fix C c
|
|
2526 |
assume "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" "class G C = Some c"
|
|
2527 |
with ws wf show "?SUP C"
|
|
2528 |
by (auto intro: no_subcls1_Object simp add: superclasses_rec subcls1_def)
|
|
2529 |
next
|
|
2530 |
fix C S c
|
|
2531 |
assume "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S" "G\<turnstile>S\<prec>\<^sub>C D"
|
|
2532 |
"\<And>s. class G S = Some s \<Longrightarrow> D \<in> superclasses G S"
|
|
2533 |
"class G C = Some c"
|
|
2534 |
with ws wf show "?SUP C"
|
|
2535 |
by - (rule superclasses_mono,
|
|
2536 |
auto dest: no_subcls1_Object simp add: subcls1_def )
|
|
2537 |
qed
|
|
2538 |
qed
|
|
2539 |
|
|
2540 |
end |