author | kuncar |
Fri, 23 Mar 2012 14:20:09 +0100 | |
changeset 47094 | 1a7ad2601cb5 |
parent 46212 | d86ef6b96097 |
child 47176 | 568fdc70e565 |
permissions | -rw-r--r-- |
12857 | 1 |
(* Title: HOL/Bali/Decl.thy |
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12859
diff
changeset
|
2 |
Author: David von Oheimb and Norbert Schirmer |
12854 | 3 |
*) |
4 |
header {* Field, method, interface, and class declarations, whole Java programs |
|
5 |
*} |
|
6 |
||
26566
36a93808642c
instantiation replacing primitive instance plus overloaded defs
haftmann
parents:
24783
diff
changeset
|
7 |
theory Decl |
44013
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
krauss
parents:
41778
diff
changeset
|
8 |
imports Term Table |
5cfc1c36ae97
moved recdef package to HOL/Library/Old_Recdef.thy
krauss
parents:
41778
diff
changeset
|
9 |
(** order is significant, because of clash for "var" **) |
26566
36a93808642c
instantiation replacing primitive instance plus overloaded defs
haftmann
parents:
24783
diff
changeset
|
10 |
begin |
12854 | 11 |
|
12 |
text {* |
|
13 |
improvements: |
|
14 |
\begin{itemize} |
|
15 |
\item clarification and correction of some aspects of the package/access concept |
|
16 |
(Also submitted as bug report to the Java Bug Database: |
|
17 |
Bug Id: 4485402 and Bug Id: 4493343 |
|
18 |
http://developer.java.sun.com/developer/bugParade/index.jshtml |
|
19 |
) |
|
20 |
\end{itemize} |
|
21 |
simplifications: |
|
22 |
\begin{itemize} |
|
23 |
\item the only field and method modifiers are static and the access modifiers |
|
24 |
\item no constructors, which may be simulated by new + suitable methods |
|
25 |
\item there is just one global initializer per class, which can simulate all |
|
26 |
others |
|
27 |
||
28 |
\item no throws clause |
|
29 |
\item a void method is replaced by one that returns Unit (of dummy type Void) |
|
30 |
||
31 |
\item no interface fields |
|
32 |
||
33 |
\item every class has an explicit superclass (unused for Object) |
|
34 |
\item the (standard) methods of Object and of standard exceptions are not |
|
35 |
specified |
|
36 |
||
37 |
\item no main method |
|
38 |
\end{itemize} |
|
39 |
*} |
|
40 |
||
41 |
subsection {* Modifier*} |
|
42 |
||
43 |
subsubsection {* Access modifier *} |
|
44 |
||
45 |
datatype acc_modi (* access modifier *) |
|
46 |
= Private | Package | Protected | Public |
|
47 |
||
48 |
text {* |
|
49 |
We can define a linear order for the access modifiers. With Private yielding the |
|
50 |
most restrictive access and public the most liberal access policy: |
|
51 |
Private < Package < Protected < Public |
|
52 |
*} |
|
53 |
||
26566
36a93808642c
instantiation replacing primitive instance plus overloaded defs
haftmann
parents:
24783
diff
changeset
|
54 |
instantiation acc_modi :: linorder |
36a93808642c
instantiation replacing primitive instance plus overloaded defs
haftmann
parents:
24783
diff
changeset
|
55 |
begin |
12854 | 56 |
|
26566
36a93808642c
instantiation replacing primitive instance plus overloaded defs
haftmann
parents:
24783
diff
changeset
|
57 |
definition |
36a93808642c
instantiation replacing primitive instance plus overloaded defs
haftmann
parents:
24783
diff
changeset
|
58 |
less_acc_def: "a < b |
36a93808642c
instantiation replacing primitive instance plus overloaded defs
haftmann
parents:
24783
diff
changeset
|
59 |
\<longleftrightarrow> (case a of |
12854 | 60 |
Private \<Rightarrow> (b=Package \<or> b=Protected \<or> b=Public) |
61 |
| Package \<Rightarrow> (b=Protected \<or> b=Public) |
|
62 |
| Protected \<Rightarrow> (b=Public) |
|
63 |
| Public \<Rightarrow> False)" |
|
64 |
||
26566
36a93808642c
instantiation replacing primitive instance plus overloaded defs
haftmann
parents:
24783
diff
changeset
|
65 |
definition |
27682 | 66 |
le_acc_def: "(a :: acc_modi) \<le> b \<longleftrightarrow> a < b \<or> a = b" |
26566
36a93808642c
instantiation replacing primitive instance plus overloaded defs
haftmann
parents:
24783
diff
changeset
|
67 |
|
41525 | 68 |
instance |
69 |
proof |
|
12854 | 70 |
fix x y z::acc_modi |
27682 | 71 |
show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)" |
72 |
by (auto simp add: le_acc_def less_acc_def split add: acc_modi.split) |
|
46212 | 73 |
show "x \<le> x" -- reflexivity |
12854 | 74 |
by (auto simp add: le_acc_def) |
41525 | 75 |
{ |
76 |
assume "x \<le> y" "y \<le> z" -- transitivity |
|
77 |
then show "x \<le> z" |
|
78 |
by (auto simp add: le_acc_def less_acc_def split add: acc_modi.split) |
|
12854 | 79 |
next |
41525 | 80 |
assume "x \<le> y" "y \<le> x" -- antisymmetry |
81 |
moreover have "\<forall> x y. x < (y::acc_modi) \<and> y < x \<longrightarrow> False" |
|
12854 | 82 |
by (auto simp add: less_acc_def split add: acc_modi.split) |
41525 | 83 |
ultimately show "x = y" by (unfold le_acc_def) iprover |
12854 | 84 |
next |
41525 | 85 |
fix x y:: acc_modi |
86 |
show "x \<le> y \<or> y \<le> x" |
|
87 |
by (auto simp add: less_acc_def le_acc_def split add: acc_modi.split) |
|
27682 | 88 |
} |
12854 | 89 |
qed |
26566
36a93808642c
instantiation replacing primitive instance plus overloaded defs
haftmann
parents:
24783
diff
changeset
|
90 |
|
36a93808642c
instantiation replacing primitive instance plus overloaded defs
haftmann
parents:
24783
diff
changeset
|
91 |
end |
12854 | 92 |
|
93 |
lemma acc_modi_top [simp]: "Public \<le> a \<Longrightarrow> a = Public" |
|
94 |
by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) |
|
95 |
||
96 |
lemma acc_modi_top1 [simp, intro!]: "a \<le> Public" |
|
97 |
by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) |
|
98 |
||
99 |
lemma acc_modi_le_Public: |
|
100 |
"a \<le> Public \<Longrightarrow> a=Private \<or> a = Package \<or> a=Protected \<or> a=Public" |
|
101 |
by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) |
|
102 |
||
103 |
lemma acc_modi_bottom: "a \<le> Private \<Longrightarrow> a = Private" |
|
104 |
by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) |
|
105 |
||
106 |
lemma acc_modi_Private_le: |
|
107 |
"Private \<le> a \<Longrightarrow> a=Private \<or> a = Package \<or> a=Protected \<or> a=Public" |
|
108 |
by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) |
|
109 |
||
110 |
lemma acc_modi_Package_le: |
|
111 |
"Package \<le> a \<Longrightarrow> a = Package \<or> a=Protected \<or> a=Public" |
|
112 |
by (auto simp add: le_acc_def less_acc_def split: acc_modi.split) |
|
113 |
||
114 |
lemma acc_modi_le_Package: |
|
115 |
"a \<le> Package \<Longrightarrow> a=Private \<or> a = Package" |
|
116 |
by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) |
|
117 |
||
118 |
lemma acc_modi_Protected_le: |
|
119 |
"Protected \<le> a \<Longrightarrow> a=Protected \<or> a=Public" |
|
120 |
by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) |
|
121 |
||
122 |
lemma acc_modi_le_Protected: |
|
123 |
"a \<le> Protected \<Longrightarrow> a=Private \<or> a = Package \<or> a = Protected" |
|
124 |
by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) |
|
125 |
||
126 |
||
127 |
lemmas acc_modi_le_Dests = acc_modi_top acc_modi_le_Public |
|
128 |
acc_modi_Private_le acc_modi_bottom |
|
129 |
acc_modi_Package_le acc_modi_le_Package |
|
130 |
acc_modi_Protected_le acc_modi_le_Protected |
|
131 |
||
132 |
lemma acc_modi_Package_le_cases |
|
133 |
[consumes 1,case_names Package Protected Public]: |
|
134 |
"Package \<le> m \<Longrightarrow> ( m = Package \<Longrightarrow> P m) \<Longrightarrow> (m=Protected \<Longrightarrow> P m) \<Longrightarrow> |
|
135 |
(m=Public \<Longrightarrow> P m) \<Longrightarrow> P m" |
|
136 |
by (auto dest: acc_modi_Package_le) |
|
137 |
||
138 |
||
139 |
subsubsection {* Static Modifier *} |
|
41778 | 140 |
type_synonym stat_modi = bool (* modifier: static *) |
12854 | 141 |
|
142 |
subsection {* Declaration (base "class" for member,interface and class |
|
143 |
declarations *} |
|
144 |
||
145 |
record decl = |
|
146 |
access :: acc_modi |
|
147 |
||
148 |
translations |
|
35431 | 149 |
(type) "decl" <= (type) "\<lparr>access::acc_modi\<rparr>" |
150 |
(type) "decl" <= (type) "\<lparr>access::acc_modi,\<dots>::'a\<rparr>" |
|
12854 | 151 |
|
152 |
subsection {* Member (field or method)*} |
|
153 |
record member = decl + |
|
154 |
static :: stat_modi |
|
155 |
||
156 |
translations |
|
35431 | 157 |
(type) "member" <= (type) "\<lparr>access::acc_modi,static::bool\<rparr>" |
158 |
(type) "member" <= (type) "\<lparr>access::acc_modi,static::bool,\<dots>::'a\<rparr>" |
|
12854 | 159 |
|
160 |
subsection {* Field *} |
|
161 |
||
162 |
record field = member + |
|
163 |
type :: ty |
|
164 |
translations |
|
35431 | 165 |
(type) "field" <= (type) "\<lparr>access::acc_modi, static::bool, type::ty\<rparr>" |
166 |
(type) "field" <= (type) "\<lparr>access::acc_modi, static::bool, type::ty,\<dots>::'a\<rparr>" |
|
12854 | 167 |
|
41778 | 168 |
type_synonym fdecl (* field declaration, cf. 8.3 *) |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
169 |
= "vname \<times> field" |
12854 | 170 |
|
171 |
||
172 |
translations |
|
35431 | 173 |
(type) "fdecl" <= (type) "vname \<times> field" |
12854 | 174 |
|
175 |
subsection {* Method *} |
|
176 |
||
177 |
record mhead = member + (* method head (excluding signature) *) |
|
178 |
pars ::"vname list" (* parameter names *) |
|
179 |
resT ::ty (* result type *) |
|
180 |
||
181 |
record mbody = (* method body *) |
|
182 |
lcls:: "(vname \<times> ty) list" (* local variables *) |
|
183 |
stmt:: stmt (* the body statement *) |
|
184 |
||
185 |
record methd = mhead + (* method in a class *) |
|
186 |
mbody::mbody |
|
187 |
||
41778 | 188 |
type_synonym mdecl = "sig \<times> methd" (* method declaration in a class *) |
12854 | 189 |
|
190 |
||
191 |
translations |
|
35431 | 192 |
(type) "mhead" <= (type) "\<lparr>access::acc_modi, static::bool, |
12854 | 193 |
pars::vname list, resT::ty\<rparr>" |
35431 | 194 |
(type) "mhead" <= (type) "\<lparr>access::acc_modi, static::bool, |
12854 | 195 |
pars::vname list, resT::ty,\<dots>::'a\<rparr>" |
35431 | 196 |
(type) "mbody" <= (type) "\<lparr>lcls::(vname \<times> ty) list,stmt::stmt\<rparr>" |
197 |
(type) "mbody" <= (type) "\<lparr>lcls::(vname \<times> ty) list,stmt::stmt,\<dots>::'a\<rparr>" |
|
198 |
(type) "methd" <= (type) "\<lparr>access::acc_modi, static::bool, |
|
12854 | 199 |
pars::vname list, resT::ty,mbody::mbody\<rparr>" |
35431 | 200 |
(type) "methd" <= (type) "\<lparr>access::acc_modi, static::bool, |
12854 | 201 |
pars::vname list, resT::ty,mbody::mbody,\<dots>::'a\<rparr>" |
35431 | 202 |
(type) "mdecl" <= (type) "sig \<times> methd" |
12854 | 203 |
|
204 |
||
37956 | 205 |
definition |
206 |
mhead :: "methd \<Rightarrow> mhead" |
|
207 |
where "mhead m = \<lparr>access=access m, static=static m, pars=pars m, resT=resT m\<rparr>" |
|
12854 | 208 |
|
209 |
lemma access_mhead [simp]:"access (mhead m) = access m" |
|
210 |
by (simp add: mhead_def) |
|
211 |
||
212 |
lemma static_mhead [simp]:"static (mhead m) = static m" |
|
213 |
by (simp add: mhead_def) |
|
214 |
||
215 |
lemma pars_mhead [simp]:"pars (mhead m) = pars m" |
|
216 |
by (simp add: mhead_def) |
|
217 |
||
218 |
lemma resT_mhead [simp]:"resT (mhead m) = resT m" |
|
219 |
by (simp add: mhead_def) |
|
220 |
||
221 |
text {* To be able to talk uniformaly about field and method declarations we |
|
222 |
introduce the notion of a member declaration (e.g. useful to define |
|
223 |
accessiblity ) *} |
|
224 |
||
225 |
datatype memberdecl = fdecl fdecl | mdecl mdecl |
|
226 |
||
227 |
datatype memberid = fid vname | mid sig |
|
228 |
||
35315 | 229 |
class has_memberid = |
230 |
fixes memberid :: "'a \<Rightarrow> memberid" |
|
12854 | 231 |
|
35315 | 232 |
instantiation memberdecl :: has_memberid |
233 |
begin |
|
12854 | 234 |
|
35315 | 235 |
definition |
12854 | 236 |
memberdecl_memberid_def: |
37956 | 237 |
"memberid m = (case m of |
12854 | 238 |
fdecl (vn,f) \<Rightarrow> fid vn |
239 |
| mdecl (sig,m) \<Rightarrow> mid sig)" |
|
240 |
||
35315 | 241 |
instance .. |
242 |
||
243 |
end |
|
244 |
||
12854 | 245 |
lemma memberid_fdecl_simp[simp]: "memberid (fdecl (vn,f)) = fid vn" |
246 |
by (simp add: memberdecl_memberid_def) |
|
247 |
||
248 |
lemma memberid_fdecl_simp1: "memberid (fdecl f) = fid (fst f)" |
|
249 |
by (cases f) (simp add: memberdecl_memberid_def) |
|
250 |
||
251 |
lemma memberid_mdecl_simp[simp]: "memberid (mdecl (sig,m)) = mid sig" |
|
252 |
by (simp add: memberdecl_memberid_def) |
|
253 |
||
254 |
lemma memberid_mdecl_simp1: "memberid (mdecl m) = mid (fst m)" |
|
255 |
by (cases m) (simp add: memberdecl_memberid_def) |
|
256 |
||
37678
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
haftmann
parents:
35547
diff
changeset
|
257 |
instantiation prod :: (type, has_memberid) has_memberid |
35315 | 258 |
begin |
12854 | 259 |
|
35315 | 260 |
definition |
12854 | 261 |
pair_memberid_def: |
37956 | 262 |
"memberid p = memberid (snd p)" |
12854 | 263 |
|
35315 | 264 |
instance .. |
265 |
||
266 |
end |
|
267 |
||
12854 | 268 |
lemma memberid_pair_simp[simp]: "memberid (c,m) = memberid m" |
269 |
by (simp add: pair_memberid_def) |
|
270 |
||
271 |
lemma memberid_pair_simp1: "memberid p = memberid (snd p)" |
|
272 |
by (simp add: pair_memberid_def) |
|
273 |
||
37956 | 274 |
definition |
275 |
is_field :: "qtname \<times> memberdecl \<Rightarrow> bool" |
|
276 |
where "is_field m = (\<exists> declC f. m=(declC,fdecl f))" |
|
12854 | 277 |
|
278 |
lemma is_fieldD: "is_field m \<Longrightarrow> \<exists> declC f. m=(declC,fdecl f)" |
|
279 |
by (simp add: is_field_def) |
|
280 |
||
281 |
lemma is_fieldI: "is_field (C,fdecl f)" |
|
282 |
by (simp add: is_field_def) |
|
283 |
||
37956 | 284 |
definition |
285 |
is_method :: "qtname \<times> memberdecl \<Rightarrow> bool" |
|
286 |
where "is_method membr = (\<exists>declC m. membr=(declC,mdecl m))" |
|
12854 | 287 |
|
288 |
lemma is_methodD: "is_method membr \<Longrightarrow> \<exists> declC m. membr=(declC,mdecl m)" |
|
289 |
by (simp add: is_method_def) |
|
290 |
||
291 |
lemma is_methodI: "is_method (C,mdecl m)" |
|
292 |
by (simp add: is_method_def) |
|
293 |
||
294 |
||
295 |
subsection {* Interface *} |
|
296 |
||
297 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13601
diff
changeset
|
298 |
record ibody = decl + --{* interface body *} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13601
diff
changeset
|
299 |
imethods :: "(sig \<times> mhead) list" --{* method heads *} |
12854 | 300 |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13601
diff
changeset
|
301 |
record iface = ibody + --{* interface *} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13601
diff
changeset
|
302 |
isuperIfs:: "qtname list" --{* superinterface list *} |
41778 | 303 |
type_synonym |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
304 |
idecl --{* interface declaration, cf. 9.1 *} |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
305 |
= "qtname \<times> iface" |
12854 | 306 |
|
307 |
translations |
|
35431 | 308 |
(type) "ibody" <= (type) "\<lparr>access::acc_modi,imethods::(sig \<times> mhead) list\<rparr>" |
309 |
(type) "ibody" <= (type) "\<lparr>access::acc_modi,imethods::(sig \<times> mhead) list,\<dots>::'a\<rparr>" |
|
310 |
(type) "iface" <= (type) "\<lparr>access::acc_modi,imethods::(sig \<times> mhead) list, |
|
12854 | 311 |
isuperIfs::qtname list\<rparr>" |
35431 | 312 |
(type) "iface" <= (type) "\<lparr>access::acc_modi,imethods::(sig \<times> mhead) list, |
12854 | 313 |
isuperIfs::qtname list,\<dots>::'a\<rparr>" |
35431 | 314 |
(type) "idecl" <= (type) "qtname \<times> iface" |
12854 | 315 |
|
37956 | 316 |
definition |
317 |
ibody :: "iface \<Rightarrow> ibody" |
|
318 |
where "ibody i = \<lparr>access=access i,imethods=imethods i\<rparr>" |
|
12854 | 319 |
|
320 |
lemma access_ibody [simp]: "(access (ibody i)) = access i" |
|
321 |
by (simp add: ibody_def) |
|
322 |
||
323 |
lemma imethods_ibody [simp]: "(imethods (ibody i)) = imethods i" |
|
324 |
by (simp add: ibody_def) |
|
325 |
||
326 |
subsection {* Class *} |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13601
diff
changeset
|
327 |
record cbody = decl + --{* class body *} |
12854 | 328 |
cfields:: "fdecl list" |
329 |
methods:: "mdecl list" |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13601
diff
changeset
|
330 |
init :: "stmt" --{* initializer *} |
12854 | 331 |
|
18551 | 332 |
record "class" = cbody + --{* class *} |
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13601
diff
changeset
|
333 |
super :: "qtname" --{* superclass *} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13601
diff
changeset
|
334 |
superIfs:: "qtname list" --{* implemented interfaces *} |
41778 | 335 |
type_synonym |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
336 |
cdecl --{* class declaration, cf. 8.1 *} |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
337 |
= "qtname \<times> class" |
12854 | 338 |
|
339 |
translations |
|
35431 | 340 |
(type) "cbody" <= (type) "\<lparr>access::acc_modi,cfields::fdecl list, |
12854 | 341 |
methods::mdecl list,init::stmt\<rparr>" |
35431 | 342 |
(type) "cbody" <= (type) "\<lparr>access::acc_modi,cfields::fdecl list, |
12854 | 343 |
methods::mdecl list,init::stmt,\<dots>::'a\<rparr>" |
35431 | 344 |
(type) "class" <= (type) "\<lparr>access::acc_modi,cfields::fdecl list, |
12854 | 345 |
methods::mdecl list,init::stmt, |
346 |
super::qtname,superIfs::qtname list\<rparr>" |
|
35431 | 347 |
(type) "class" <= (type) "\<lparr>access::acc_modi,cfields::fdecl list, |
12854 | 348 |
methods::mdecl list,init::stmt, |
349 |
super::qtname,superIfs::qtname list,\<dots>::'a\<rparr>" |
|
35431 | 350 |
(type) "cdecl" <= (type) "qtname \<times> class" |
12854 | 351 |
|
37956 | 352 |
definition |
353 |
cbody :: "class \<Rightarrow> cbody" |
|
354 |
where "cbody c = \<lparr>access=access c, cfields=cfields c,methods=methods c,init=init c\<rparr>" |
|
12854 | 355 |
|
356 |
lemma access_cbody [simp]:"access (cbody c) = access c" |
|
357 |
by (simp add: cbody_def) |
|
358 |
||
359 |
lemma cfields_cbody [simp]:"cfields (cbody c) = cfields c" |
|
360 |
by (simp add: cbody_def) |
|
361 |
||
362 |
lemma methods_cbody [simp]:"methods (cbody c) = methods c" |
|
363 |
by (simp add: cbody_def) |
|
364 |
||
365 |
lemma init_cbody [simp]:"init (cbody c) = init c" |
|
366 |
by (simp add: cbody_def) |
|
367 |
||
368 |
||
369 |
section "standard classes" |
|
370 |
||
371 |
consts |
|
14674 | 372 |
Object_mdecls :: "mdecl list" --{* methods of Object *} |
373 |
SXcpt_mdecls :: "mdecl list" --{* methods of SXcpts *} |
|
12854 | 374 |
|
37956 | 375 |
definition |
376 |
ObjectC :: "cdecl" --{* declaration of root class *} where |
|
377 |
"ObjectC = (Object,\<lparr>access=Public,cfields=[],methods=Object_mdecls, |
|
378 |
init=Skip,super=undefined,superIfs=[]\<rparr>)" |
|
12854 | 379 |
|
37956 | 380 |
definition |
381 |
SXcptC ::"xname \<Rightarrow> cdecl" --{* declarations of throwable classes *} where |
|
382 |
"SXcptC xn = (SXcpt xn,\<lparr>access=Public,cfields=[],methods=SXcpt_mdecls, |
|
12854 | 383 |
init=Skip, |
384 |
super=if xn = Throwable then Object |
|
385 |
else SXcpt Throwable, |
|
386 |
superIfs=[]\<rparr>)" |
|
387 |
||
388 |
lemma ObjectC_neq_SXcptC [simp]: "ObjectC \<noteq> SXcptC xn" |
|
389 |
by (simp add: ObjectC_def SXcptC_def Object_def SXcpt_def) |
|
390 |
||
391 |
lemma SXcptC_inject [simp]: "(SXcptC xn = SXcptC xm) = (xn = xm)" |
|
17778 | 392 |
by (simp add: SXcptC_def) |
12854 | 393 |
|
37956 | 394 |
definition |
395 |
standard_classes :: "cdecl list" where |
|
396 |
"standard_classes = [ObjectC, SXcptC Throwable, |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
397 |
SXcptC NullPointer, SXcptC OutOfMemory, SXcptC ClassCast, |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
398 |
SXcptC NegArrSize , SXcptC IndOutBound, SXcptC ArrStore]" |
12854 | 399 |
|
400 |
||
401 |
section "programs" |
|
402 |
||
403 |
record prog = |
|
404 |
ifaces ::"idecl list" |
|
405 |
"classes"::"cdecl list" |
|
406 |
||
407 |
translations |
|
35431 | 408 |
(type) "prog" <= (type) "\<lparr>ifaces::idecl list,classes::cdecl list\<rparr>" |
409 |
(type) "prog" <= (type) "\<lparr>ifaces::idecl list,classes::cdecl list,\<dots>::'a\<rparr>" |
|
12854 | 410 |
|
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
411 |
abbreviation |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
412 |
iface :: "prog \<Rightarrow> (qtname, iface) table" |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
413 |
where "iface G I == table_of (ifaces G) I" |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
414 |
|
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
415 |
abbreviation |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
416 |
"class" :: "prog \<Rightarrow> (qtname, class) table" |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
417 |
where "class G C == table_of (classes G) C" |
12854 | 418 |
|
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
419 |
abbreviation |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
420 |
is_iface :: "prog \<Rightarrow> qtname \<Rightarrow> bool" |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
421 |
where "is_iface G I == iface G I \<noteq> None" |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
422 |
|
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
423 |
abbreviation |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
424 |
is_class :: "prog \<Rightarrow> qtname \<Rightarrow> bool" |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
425 |
where "is_class G C == class G C \<noteq> None" |
12854 | 426 |
|
427 |
||
428 |
section "is type" |
|
429 |
||
37956 | 430 |
primrec is_type :: "prog \<Rightarrow> ty \<Rightarrow> bool" |
431 |
and isrtype :: "prog \<Rightarrow> ref_ty \<Rightarrow> bool" |
|
432 |
where |
|
433 |
"is_type G (PrimT pt) = True" |
|
434 |
| "is_type G (RefT rt) = isrtype G rt" |
|
435 |
| "isrtype G (NullT) = True" |
|
436 |
| "isrtype G (IfaceT tn) = is_iface G tn" |
|
437 |
| "isrtype G (ClassT tn) = is_class G tn" |
|
438 |
| "isrtype G (ArrayT T ) = is_type G T" |
|
12854 | 439 |
|
440 |
lemma type_is_iface: "is_type G (Iface I) \<Longrightarrow> is_iface G I" |
|
441 |
by auto |
|
442 |
||
443 |
lemma type_is_class: "is_type G (Class C) \<Longrightarrow> is_class G C" |
|
444 |
by auto |
|
445 |
||
446 |
||
447 |
section "subinterface and subclass relation, in anticipation of TypeRel.thy" |
|
448 |
||
37956 | 449 |
definition |
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13601
diff
changeset
|
450 |
subint1 :: "prog \<Rightarrow> (qtname \<times> qtname) set" --{* direct subinterface *} |
37956 | 451 |
where "subint1 G = {(I,J). \<exists>i\<in>iface G I: J\<in>set (isuperIfs i)}" |
12854 | 452 |
|
37956 | 453 |
definition |
454 |
subcls1 :: "prog \<Rightarrow> (qtname \<times> qtname) set" --{* direct subclass *} |
|
455 |
where "subcls1 G = {(C,D). C\<noteq>Object \<and> (\<exists>c\<in>class G C: super c = D)}" |
|
12854 | 456 |
|
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
457 |
abbreviation |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
458 |
subcls1_syntax :: "prog => [qtname, qtname] => bool" ("_|-_<:C1_" [71,71,71] 70) |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
459 |
where "G|-C <:C1 D == (C,D) \<in> subcls1 G" |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
460 |
|
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
461 |
abbreviation |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
462 |
subclseq_syntax :: "prog => [qtname, qtname] => bool" ("_|-_<=:C _"[71,71,71] 70) |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
463 |
where "G|-C <=:C D == (C,D) \<in>(subcls1 G)^*" (* cf. 8.1.3 *) |
12854 | 464 |
|
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
465 |
abbreviation |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
466 |
subcls_syntax :: "prog => [qtname, qtname] => bool" ("_|-_<:C _"[71,71,71] 70) |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
467 |
where "G|-C <:C D == (C,D) \<in>(subcls1 G)^+" |
12854 | 468 |
|
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
469 |
notation (xsymbols) |
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35315
diff
changeset
|
470 |
subcls1_syntax ("_\<turnstile>_\<prec>\<^sub>C1_" [71,71,71] 70) and |
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
471 |
subclseq_syntax ("_\<turnstile>_\<preceq>\<^sub>C _" [71,71,71] 70) and |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
472 |
subcls_syntax ("_\<turnstile>_\<prec>\<^sub>C _" [71,71,71] 70) |
12854 | 473 |
|
474 |
lemma subint1I: "\<lbrakk>iface G I = Some i; J \<in> set (isuperIfs i)\<rbrakk> |
|
475 |
\<Longrightarrow> (I,J) \<in> subint1 G" |
|
476 |
apply (simp add: subint1_def) |
|
477 |
done |
|
478 |
||
479 |
lemma subcls1I:"\<lbrakk>class G C = Some c; C \<noteq> Object\<rbrakk> \<Longrightarrow> (C,(super c)) \<in> subcls1 G" |
|
480 |
apply (simp add: subcls1_def) |
|
481 |
done |
|
482 |
||
483 |
||
484 |
lemma subint1D: "(I,J)\<in>subint1 G\<Longrightarrow> \<exists>i\<in>iface G I: J\<in>set (isuperIfs i)" |
|
485 |
by (simp add: subint1_def) |
|
486 |
||
487 |
lemma subcls1D: |
|
488 |
"(C,D)\<in>subcls1 G \<Longrightarrow> C\<noteq>Object \<and> (\<exists>c. class G C = Some c \<and> (super c = D))" |
|
489 |
apply (simp add: subcls1_def) |
|
490 |
apply auto |
|
491 |
done |
|
492 |
||
493 |
lemma subint1_def2: |
|
14952
47455995693d
removal of x-symbol syntax <Sigma> for dependent products
paulson
parents:
14674
diff
changeset
|
494 |
"subint1 G = (SIGMA I: {I. is_iface G I}. set (isuperIfs (the (iface G I))))" |
12854 | 495 |
apply (unfold subint1_def) |
496 |
apply auto |
|
497 |
done |
|
498 |
||
499 |
lemma subcls1_def2: |
|
14952
47455995693d
removal of x-symbol syntax <Sigma> for dependent products
paulson
parents:
14674
diff
changeset
|
500 |
"subcls1 G = |
47455995693d
removal of x-symbol syntax <Sigma> for dependent products
paulson
parents:
14674
diff
changeset
|
501 |
(SIGMA C: {C. is_class G C}. {D. C\<noteq>Object \<and> super (the(class G C))=D})" |
12854 | 502 |
apply (unfold subcls1_def) |
503 |
apply auto |
|
504 |
done |
|
505 |
||
506 |
lemma subcls_is_class: |
|
507 |
"\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D\<rbrakk> \<Longrightarrow> \<exists> c. class G C = Some c" |
|
508 |
by (auto simp add: subcls1_def dest: tranclD) |
|
509 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35315
diff
changeset
|
510 |
lemma no_subcls1_Object:"G\<turnstile>Object\<prec>\<^sub>C1 D \<Longrightarrow> P" |
12854 | 511 |
by (auto simp add: subcls1_def) |
512 |
||
513 |
lemma no_subcls_Object: "G\<turnstile>Object\<prec>\<^sub>C D \<Longrightarrow> P" |
|
514 |
apply (erule trancl_induct) |
|
515 |
apply (auto intro: no_subcls1_Object) |
|
516 |
done |
|
517 |
||
518 |
section "well-structured programs" |
|
519 |
||
37956 | 520 |
definition |
521 |
ws_idecl :: "prog \<Rightarrow> qtname \<Rightarrow> qtname list \<Rightarrow> bool" |
|
522 |
where "ws_idecl G I si = (\<forall>J\<in>set si. is_iface G J \<and> (J,I)\<notin>(subint1 G)^+)" |
|
12854 | 523 |
|
37956 | 524 |
definition |
525 |
ws_cdecl :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool" |
|
526 |
where "ws_cdecl G C sc = (C\<noteq>Object \<longrightarrow> is_class G sc \<and> (sc,C)\<notin>(subcls1 G)^+)" |
|
12854 | 527 |
|
37956 | 528 |
definition |
529 |
ws_prog :: "prog \<Rightarrow> bool" where |
|
530 |
"ws_prog G = ((\<forall>(I,i)\<in>set (ifaces G). ws_idecl G I (isuperIfs i)) \<and> |
|
531 |
(\<forall>(C,c)\<in>set (classes G). ws_cdecl G C (super c)))" |
|
12854 | 532 |
|
533 |
||
534 |
lemma ws_progI: |
|
535 |
"\<lbrakk>\<forall>(I,i)\<in>set (ifaces G). \<forall>J\<in>set (isuperIfs i). is_iface G J \<and> |
|
536 |
(J,I) \<notin> (subint1 G)^+; |
|
537 |
\<forall>(C,c)\<in>set (classes G). C\<noteq>Object \<longrightarrow> is_class G (super c) \<and> |
|
538 |
((super c),C) \<notin> (subcls1 G)^+ |
|
539 |
\<rbrakk> \<Longrightarrow> ws_prog G" |
|
540 |
apply (unfold ws_prog_def ws_idecl_def ws_cdecl_def) |
|
541 |
apply (erule_tac conjI) |
|
542 |
apply blast |
|
543 |
done |
|
544 |
||
545 |
lemma ws_prog_ideclD: |
|
546 |
"\<lbrakk>iface G I = Some i; J\<in>set (isuperIfs i); ws_prog G\<rbrakk> \<Longrightarrow> |
|
547 |
is_iface G J \<and> (J,I)\<notin>(subint1 G)^+" |
|
548 |
apply (unfold ws_prog_def ws_idecl_def) |
|
549 |
apply clarify |
|
550 |
apply (drule_tac map_of_SomeD) |
|
551 |
apply auto |
|
552 |
done |
|
553 |
||
554 |
lemma ws_prog_cdeclD: |
|
555 |
"\<lbrakk>class G C = Some c; C\<noteq>Object; ws_prog G\<rbrakk> \<Longrightarrow> |
|
556 |
is_class G (super c) \<and> (super c,C)\<notin>(subcls1 G)^+" |
|
557 |
apply (unfold ws_prog_def ws_cdecl_def) |
|
558 |
apply clarify |
|
559 |
apply (drule_tac map_of_SomeD) |
|
560 |
apply auto |
|
561 |
done |
|
562 |
||
563 |
||
564 |
section "well-foundedness" |
|
565 |
||
566 |
lemma finite_is_iface: "finite {I. is_iface G I}" |
|
567 |
apply (fold dom_def) |
|
568 |
apply (rule_tac finite_dom_map_of) |
|
569 |
done |
|
570 |
||
571 |
lemma finite_is_class: "finite {C. is_class G C}" |
|
572 |
apply (fold dom_def) |
|
573 |
apply (rule_tac finite_dom_map_of) |
|
574 |
done |
|
575 |
||
576 |
lemma finite_subint1: "finite (subint1 G)" |
|
577 |
apply (subst subint1_def2) |
|
578 |
apply (rule finite_SigmaI) |
|
579 |
apply (rule finite_is_iface) |
|
580 |
apply (simp (no_asm)) |
|
581 |
done |
|
582 |
||
583 |
lemma finite_subcls1: "finite (subcls1 G)" |
|
584 |
apply (subst subcls1_def2) |
|
585 |
apply (rule finite_SigmaI) |
|
586 |
apply (rule finite_is_class) |
|
587 |
apply (rule_tac B = "{super (the (class G C))}" in finite_subset) |
|
588 |
apply auto |
|
589 |
done |
|
590 |
||
591 |
lemma subint1_irrefl_lemma1: |
|
592 |
"ws_prog G \<Longrightarrow> (subint1 G)^-1 \<inter> (subint1 G)^+ = {}" |
|
593 |
apply (force dest: subint1D ws_prog_ideclD conjunct2) |
|
594 |
done |
|
595 |
||
596 |
lemma subcls1_irrefl_lemma1: |
|
597 |
"ws_prog G \<Longrightarrow> (subcls1 G)^-1 \<inter> (subcls1 G)^+ = {}" |
|
598 |
apply (force dest: subcls1D ws_prog_cdeclD conjunct2) |
|
599 |
done |
|
600 |
||
601 |
lemmas subint1_irrefl_lemma2 = subint1_irrefl_lemma1 [THEN irrefl_tranclI'] |
|
602 |
lemmas subcls1_irrefl_lemma2 = subcls1_irrefl_lemma1 [THEN irrefl_tranclI'] |
|
603 |
||
604 |
lemma subint1_irrefl: "\<lbrakk>(x, y) \<in> subint1 G; ws_prog G\<rbrakk> \<Longrightarrow> x \<noteq> y" |
|
605 |
apply (rule irrefl_trancl_rD) |
|
606 |
apply (rule subint1_irrefl_lemma2) |
|
607 |
apply auto |
|
608 |
done |
|
609 |
||
610 |
lemma subcls1_irrefl: "\<lbrakk>(x, y) \<in> subcls1 G; ws_prog G\<rbrakk> \<Longrightarrow> x \<noteq> y" |
|
611 |
apply (rule irrefl_trancl_rD) |
|
612 |
apply (rule subcls1_irrefl_lemma2) |
|
613 |
apply auto |
|
614 |
done |
|
615 |
||
45605 | 616 |
lemmas subint1_acyclic = subint1_irrefl_lemma2 [THEN acyclicI] |
617 |
lemmas subcls1_acyclic = subcls1_irrefl_lemma2 [THEN acyclicI] |
|
12854 | 618 |
|
619 |
||
620 |
lemma wf_subint1: "ws_prog G \<Longrightarrow> wf ((subint1 G)\<inverse>)" |
|
621 |
by (auto intro: finite_acyclic_wf_converse finite_subint1 subint1_acyclic) |
|
622 |
||
623 |
lemma wf_subcls1: "ws_prog G \<Longrightarrow> wf ((subcls1 G)\<inverse>)" |
|
624 |
by (auto intro: finite_acyclic_wf_converse finite_subcls1 subcls1_acyclic) |
|
625 |
||
626 |
||
627 |
lemma subint1_induct: |
|
628 |
"\<lbrakk>ws_prog G; \<And>x. \<forall>y. (x, y) \<in> subint1 G \<longrightarrow> P y \<Longrightarrow> P x\<rbrakk> \<Longrightarrow> P a" |
|
629 |
apply (frule wf_subint1) |
|
630 |
apply (erule wf_induct) |
|
631 |
apply (simp (no_asm_use) only: converse_iff) |
|
632 |
apply blast |
|
633 |
done |
|
634 |
||
635 |
lemma subcls1_induct [consumes 1]: |
|
636 |
"\<lbrakk>ws_prog G; \<And>x. \<forall>y. (x, y) \<in> subcls1 G \<longrightarrow> P y \<Longrightarrow> P x\<rbrakk> \<Longrightarrow> P a" |
|
637 |
apply (frule wf_subcls1) |
|
638 |
apply (erule wf_induct) |
|
639 |
apply (simp (no_asm_use) only: converse_iff) |
|
640 |
apply blast |
|
641 |
done |
|
642 |
||
643 |
lemma ws_subint1_induct: |
|
644 |
"\<lbrakk>is_iface G I; ws_prog G; \<And>I i. \<lbrakk>iface G I = Some i \<and> |
|
645 |
(\<forall>J \<in> set (isuperIfs i). (I,J)\<in>subint1 G \<and> P J \<and> is_iface G J)\<rbrakk> \<Longrightarrow> P I |
|
646 |
\<rbrakk> \<Longrightarrow> P I" |
|
24038 | 647 |
apply (erule rev_mp) |
12854 | 648 |
apply (rule subint1_induct) |
649 |
apply assumption |
|
18447 | 650 |
apply (simp (no_asm)) |
12854 | 651 |
apply safe |
18447 | 652 |
apply (blast dest: subint1I ws_prog_ideclD) |
12854 | 653 |
done |
654 |
||
655 |
||
656 |
lemma ws_subcls1_induct: "\<lbrakk>is_class G C; ws_prog G; |
|
657 |
\<And>C c. \<lbrakk>class G C = Some c; |
|
658 |
(C \<noteq> Object \<longrightarrow> (C,(super c))\<in>subcls1 G \<and> |
|
659 |
P (super c) \<and> is_class G (super c))\<rbrakk> \<Longrightarrow> P C |
|
660 |
\<rbrakk> \<Longrightarrow> P C" |
|
24038 | 661 |
apply (erule rev_mp) |
12854 | 662 |
apply (rule subcls1_induct) |
663 |
apply assumption |
|
18447 | 664 |
apply (simp (no_asm)) |
12854 | 665 |
apply safe |
666 |
apply (fast dest: subcls1I ws_prog_cdeclD) |
|
667 |
done |
|
668 |
||
669 |
lemma ws_class_induct [consumes 2, case_names Object Subcls]: |
|
670 |
"\<lbrakk>class G C = Some c; ws_prog G; |
|
671 |
\<And> co. class G Object = Some co \<Longrightarrow> P Object; |
|
672 |
\<And> C c. \<lbrakk>class G C = Some c; C \<noteq> Object; P (super c)\<rbrakk> \<Longrightarrow> P C |
|
673 |
\<rbrakk> \<Longrightarrow> P C" |
|
674 |
proof - |
|
675 |
assume clsC: "class G C = Some c" |
|
676 |
and init: "\<And> co. class G Object = Some co \<Longrightarrow> P Object" |
|
677 |
and step: "\<And> C c. \<lbrakk>class G C = Some c; C \<noteq> Object; P (super c)\<rbrakk> \<Longrightarrow> P C" |
|
678 |
assume ws: "ws_prog G" |
|
679 |
then have "is_class G C \<Longrightarrow> P C" |
|
680 |
proof (induct rule: subcls1_induct) |
|
681 |
fix C |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35315
diff
changeset
|
682 |
assume hyp:"\<forall> S. G\<turnstile>C \<prec>\<^sub>C1 S \<longrightarrow> is_class G S \<longrightarrow> P S" |
12854 | 683 |
and iscls:"is_class G C" |
684 |
show "P C" |
|
685 |
proof (cases "C=Object") |
|
686 |
case True with iscls init show "P C" by auto |
|
687 |
next |
|
688 |
case False with ws step hyp iscls |
|
18576 | 689 |
show "P C" by (auto dest: subcls1I ws_prog_cdeclD) |
12854 | 690 |
qed |
691 |
qed |
|
692 |
with clsC show ?thesis by simp |
|
693 |
qed |
|
694 |
||
695 |
lemma ws_class_induct' [consumes 2, case_names Object Subcls]: |
|
696 |
"\<lbrakk>is_class G C; ws_prog G; |
|
697 |
\<And> co. class G Object = Some co \<Longrightarrow> P Object; |
|
698 |
\<And> C c. \<lbrakk>class G C = Some c; C \<noteq> Object; P (super c)\<rbrakk> \<Longrightarrow> P C |
|
699 |
\<rbrakk> \<Longrightarrow> P C" |
|
18447 | 700 |
by (auto intro: ws_class_induct) |
12854 | 701 |
|
702 |
lemma ws_class_induct'' [consumes 2, case_names Object Subcls]: |
|
703 |
"\<lbrakk>class G C = Some c; ws_prog G; |
|
704 |
\<And> co. class G Object = Some co \<Longrightarrow> P Object co; |
|
705 |
\<And> C c sc. \<lbrakk>class G C = Some c; class G (super c) = Some sc; |
|
706 |
C \<noteq> Object; P (super c) sc\<rbrakk> \<Longrightarrow> P C c |
|
707 |
\<rbrakk> \<Longrightarrow> P C c" |
|
708 |
proof - |
|
709 |
assume clsC: "class G C = Some c" |
|
710 |
and init: "\<And> co. class G Object = Some co \<Longrightarrow> P Object co" |
|
711 |
and step: "\<And> C c sc . \<lbrakk>class G C = Some c; class G (super c) = Some sc; |
|
712 |
C \<noteq> Object; P (super c) sc\<rbrakk> \<Longrightarrow> P C c" |
|
713 |
assume ws: "ws_prog G" |
|
714 |
then have "\<And> c. class G C = Some c\<Longrightarrow> P C c" |
|
715 |
proof (induct rule: subcls1_induct) |
|
716 |
fix C c |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35315
diff
changeset
|
717 |
assume hyp:"\<forall> S. G\<turnstile>C \<prec>\<^sub>C1 S \<longrightarrow> (\<forall> s. class G S = Some s \<longrightarrow> P S s)" |
12854 | 718 |
and iscls:"class G C = Some c" |
719 |
show "P C c" |
|
720 |
proof (cases "C=Object") |
|
721 |
case True with iscls init show "P C c" by auto |
|
722 |
next |
|
723 |
case False |
|
724 |
with ws iscls obtain sc where |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
725 |
sc: "class G (super c) = Some sc" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
726 |
by (auto dest: ws_prog_cdeclD) |
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35315
diff
changeset
|
727 |
from iscls False have "G\<turnstile>C \<prec>\<^sub>C1 (super c)" by (rule subcls1I) |
12854 | 728 |
with False ws step hyp iscls sc |
729 |
show "P C c" |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
730 |
by (auto) |
12854 | 731 |
qed |
732 |
qed |
|
733 |
with clsC show "P C c" by auto |
|
734 |
qed |
|
735 |
||
736 |
lemma ws_interface_induct [consumes 2, case_names Step]: |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12925
diff
changeset
|
737 |
assumes is_if_I: "is_iface G I" and |
12854 | 738 |
ws: "ws_prog G" and |
739 |
hyp_sub: "\<And>I i. \<lbrakk>iface G I = Some i; |
|
740 |
\<forall> J \<in> set (isuperIfs i). |
|
741 |
(I,J)\<in>subint1 G \<and> P J \<and> is_iface G J\<rbrakk> \<Longrightarrow> P I" |
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12925
diff
changeset
|
742 |
shows "P I" |
12854 | 743 |
proof - |
744 |
from is_if_I ws |
|
745 |
show "P I" |
|
746 |
proof (rule ws_subint1_induct) |
|
747 |
fix I i |
|
748 |
assume hyp: "iface G I = Some i \<and> |
|
749 |
(\<forall>J\<in>set (isuperIfs i). (I,J) \<in>subint1 G \<and> P J \<and> is_iface G J)" |
|
750 |
then have if_I: "iface G I = Some i" |
|
751 |
by blast |
|
752 |
show "P I" |
|
753 |
proof (cases "isuperIfs i") |
|
754 |
case Nil |
|
755 |
with if_I hyp_sub |
|
756 |
show "P I" |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
757 |
by auto |
12854 | 758 |
next |
759 |
case (Cons hd tl) |
|
760 |
with hyp if_I hyp_sub |
|
761 |
show "P I" |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
762 |
by auto |
12854 | 763 |
qed |
764 |
qed |
|
765 |
qed |
|
766 |
||
767 |
section "general recursion operators for the interface and class hiearchies" |
|
768 |
||
46212 | 769 |
function iface_rec :: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<Rightarrow> iface \<Rightarrow> 'a set \<Rightarrow> 'a) \<Rightarrow> 'a" |
35440 | 770 |
where |
771 |
[simp del]: "iface_rec G I f = |
|
772 |
(case iface G I of |
|
28524 | 773 |
None \<Rightarrow> undefined |
12854 | 774 |
| Some i \<Rightarrow> if ws_prog G |
775 |
then f I i |
|
35440 | 776 |
((\<lambda>J. iface_rec G J f)`set (isuperIfs i)) |
28524 | 777 |
else undefined)" |
35440 | 778 |
by auto |
779 |
termination |
|
780 |
by (relation "inv_image (same_fst ws_prog (\<lambda>G. (subint1 G)^-1)) (%(x,y,z). (x,y))") |
|
781 |
(auto simp: wf_subint1 subint1I wf_same_fst) |
|
12854 | 782 |
|
783 |
lemma iface_rec: |
|
784 |
"\<lbrakk>iface G I = Some i; ws_prog G\<rbrakk> \<Longrightarrow> |
|
35440 | 785 |
iface_rec G I f = f I i ((\<lambda>J. iface_rec G J f)`set (isuperIfs i))" |
12854 | 786 |
apply (subst iface_rec.simps) |
787 |
apply simp |
|
788 |
done |
|
789 |
||
35440 | 790 |
|
791 |
function |
|
792 |
class_rec :: "prog \<Rightarrow> qtname \<Rightarrow> 'a \<Rightarrow> (qtname \<Rightarrow> class \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a" |
|
793 |
where |
|
794 |
[simp del]: "class_rec G C t f = |
|
795 |
(case class G C of |
|
28524 | 796 |
None \<Rightarrow> undefined |
12854 | 797 |
| Some c \<Rightarrow> if ws_prog G |
798 |
then f C c |
|
799 |
(if C = Object then t |
|
35440 | 800 |
else class_rec G (super c) t f) |
28524 | 801 |
else undefined)" |
35440 | 802 |
|
803 |
by auto |
|
804 |
termination |
|
805 |
by (relation "inv_image (same_fst ws_prog (\<lambda>G. (subcls1 G)^-1)) (%(x,y,z,w). (x,y))") |
|
806 |
(auto simp: wf_subcls1 subcls1I wf_same_fst) |
|
12854 | 807 |
|
808 |
lemma class_rec: "\<lbrakk>class G C = Some c; ws_prog G\<rbrakk> \<Longrightarrow> |
|
35440 | 809 |
class_rec G C t f = |
810 |
f C c (if C = Object then t else class_rec G (super c) t f)" |
|
811 |
apply (subst class_rec.simps) |
|
12854 | 812 |
apply simp |
813 |
done |
|
814 |
||
37956 | 815 |
definition |
816 |
imethds :: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables" where |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13601
diff
changeset
|
817 |
--{* methods of an interface, with overriding and inheritance, cf. 9.2 *} |
37956 | 818 |
"imethds G I = iface_rec G I |
12854 | 819 |
(\<lambda>I i ts. (Un_tables ts) \<oplus>\<oplus> |
30235
58d147683393
Made Option a separate theory and renamed option_map to Option.map
nipkow
parents:
28524
diff
changeset
|
820 |
(Option.set \<circ> table_of (map (\<lambda>(s,m). (s,I,m)) (imethods i))))" |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30235
diff
changeset
|
821 |
|
12854 | 822 |
|
823 |
||
824 |
end |