| author | blanchet | 
| Tue, 03 May 2011 08:52:16 +0200 | |
| changeset 42648 | 6099b85ae48e | 
| parent 42463 | f270e3e18be5 | 
| child 42793 | 88bee9f6eec7 | 
| permissions | -rw-r--r-- | 
| 8011 | 1  | 
(* Title: HOL/MicroJava/J/WellForm.thy  | 
2  | 
Author: David von Oheimb  | 
|
3  | 
Copyright 1999 Technische Universitaet Muenchen  | 
|
| 11070 | 4  | 
*)  | 
| 8011 | 5  | 
|
| 12911 | 6  | 
header {* \isaheader{Well-formedness of Java programs} *}
 | 
| 8011 | 7  | 
|
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
8  | 
theory WellForm  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
9  | 
imports TypeRel SystemClasses  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
10  | 
begin  | 
| 8011 | 11  | 
|
| 11070 | 12  | 
text {*
 | 
13  | 
for static checks on expressions and statements, see WellType.  | 
|
14  | 
||
15  | 
\begin{description}
 | 
|
16  | 
\item[improvements over Java Specification 1.0 (cf. 8.4.6.3, 8.4.6.4, 9.4.1):]\ \\  | 
|
17  | 
\begin{itemize}
 | 
|
18  | 
\item a method implementing or overwriting another method may have a result type  | 
|
19  | 
that widens to the result type of the other method (instead of identical type)  | 
|
20  | 
\end{itemize}
 | 
|
21  | 
||
22  | 
\item[simplifications:]\ \\  | 
|
23  | 
\begin{itemize}
 | 
|
24  | 
\item for uniformity, Object is assumed to be declared like any other class  | 
|
25  | 
\end{itemize}
 | 
|
26  | 
\end{description}
 | 
|
27  | 
*}  | 
|
| 42463 | 28  | 
type_synonym 'c wf_mb = "'c prog => cname => 'c mdecl => bool"  | 
| 8011 | 29  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
30  | 
definition wf_syscls :: "'c prog => bool" where  | 
| 14045 | 31  | 
"wf_syscls G == let cs = set G in Object \<in> fst ` cs \<and> (\<forall>x. Xcpt x \<in> fst ` cs)"  | 
32  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
33  | 
definition wf_fdecl :: "'c prog => fdecl => bool" where  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
34  | 
"wf_fdecl G == \<lambda>(fn,ft). is_type G ft"  | 
| 8011 | 35  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
36  | 
definition wf_mhead :: "'c prog => sig => ty => bool" where  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
37  | 
"wf_mhead G == \<lambda>(mn,pTs) rT. (\<forall>T\<in>set pTs. is_type G T) \<and> is_type G rT"  | 
| 8011 | 38  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
39  | 
definition ws_cdecl :: "'c prog => 'c cdecl => bool" where  | 
| 14045 | 40  | 
"ws_cdecl G ==  | 
41  | 
\<lambda>(C,(D,fs,ms)).  | 
|
42  | 
(\<forall>f\<in>set fs. wf_fdecl G f) \<and> unique fs \<and>  | 
|
43  | 
(\<forall>(sig,rT,mb)\<in>set ms. wf_mhead G sig rT) \<and> unique ms \<and>  | 
|
44  | 
(C \<noteq> Object \<longrightarrow> is_class G D \<and> \<not>G\<turnstile>D\<preceq>C C)"  | 
|
45  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
46  | 
definition ws_prog :: "'c prog => bool" where  | 
| 14045 | 47  | 
"ws_prog G ==  | 
48  | 
wf_syscls G \<and> (\<forall>c\<in>set G. ws_cdecl G c) \<and> unique G"  | 
|
49  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
50  | 
definition wf_mrT :: "'c prog => 'c cdecl => bool" where  | 
| 14045 | 51  | 
"wf_mrT G ==  | 
52  | 
\<lambda>(C,(D,fs,ms)).  | 
|
53  | 
(C \<noteq> Object \<longrightarrow> (\<forall>(sig,rT,b)\<in>set ms. \<forall>D' rT' b'.  | 
|
54  | 
method(G,D) sig = Some(D',rT',b') --> G\<turnstile>rT\<preceq>rT'))"  | 
|
55  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
56  | 
definition wf_cdecl_mdecl :: "'c wf_mb => 'c prog => 'c cdecl => bool" where  | 
| 14045 | 57  | 
"wf_cdecl_mdecl wf_mb G ==  | 
58  | 
\<lambda>(C,(D,fs,ms)). (\<forall>m\<in>set ms. wf_mb G C m)"  | 
|
59  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
60  | 
definition wf_prog :: "'c wf_mb => 'c prog => bool" where  | 
| 14045 | 61  | 
"wf_prog wf_mb G ==  | 
62  | 
ws_prog G \<and> (\<forall>c\<in> set G. wf_mrT G c \<and> wf_cdecl_mdecl wf_mb G c)"  | 
|
63  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
64  | 
definition wf_mdecl :: "'c wf_mb => 'c wf_mb" where  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
65  | 
"wf_mdecl wf_mb G C == \<lambda>(sig,rT,mb). wf_mhead G sig rT \<and> wf_mb G C (sig,rT,mb)"  | 
| 8011 | 66  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
33954 
diff
changeset
 | 
67  | 
definition wf_cdecl :: "'c wf_mb => 'c prog => 'c cdecl => bool" where  | 
| 10042 | 68  | 
"wf_cdecl wf_mb G ==  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
69  | 
\<lambda>(C,(D,fs,ms)).  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
70  | 
(\<forall>f\<in>set fs. wf_fdecl G f) \<and> unique fs \<and>  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
71  | 
(\<forall>m\<in>set ms. wf_mdecl wf_mb G C m) \<and> unique ms \<and>  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
72  | 
(C \<noteq> Object \<longrightarrow> is_class G D \<and> \<not>G\<turnstile>D\<preceq>C C \<and>  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
73  | 
(\<forall>(sig,rT,b)\<in>set ms. \<forall>D' rT' b'.  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
74  | 
method(G,D) sig = Some(D',rT',b') --> G\<turnstile>rT\<preceq>rT'))"  | 
| 8011 | 75  | 
|
| 14045 | 76  | 
lemma wf_cdecl_mrT_cdecl_mdecl:  | 
77  | 
"(wf_cdecl wf_mb G c) = (ws_cdecl G c \<and> wf_mrT G c \<and> wf_cdecl_mdecl wf_mb G c)"  | 
|
78  | 
apply (rule iffI)  | 
|
79  | 
apply (simp add: wf_cdecl_def ws_cdecl_def wf_mrT_def wf_cdecl_mdecl_def  | 
|
80  | 
wf_mdecl_def wf_mhead_def split_beta)+  | 
|
81  | 
done  | 
|
| 12951 | 82  | 
|
| 14045 | 83  | 
lemma wf_cdecl_ws_cdecl [intro]: "wf_cdecl wf_mb G cd \<Longrightarrow> ws_cdecl G cd"  | 
84  | 
by (simp add: wf_cdecl_mrT_cdecl_mdecl)  | 
|
85  | 
||
86  | 
lemma wf_prog_ws_prog [intro]: "wf_prog wf_mb G \<Longrightarrow> ws_prog G"  | 
|
87  | 
by (simp add: wf_prog_def ws_prog_def)  | 
|
88  | 
||
89  | 
lemma wf_prog_wf_mdecl:  | 
|
90  | 
"\<lbrakk> wf_prog wf_mb G; (C,S,fs,mdecls) \<in> set G; ((mn,pTs),rT,code) \<in> set mdecls\<rbrakk>  | 
|
91  | 
\<Longrightarrow> wf_mdecl wf_mb G C ((mn,pTs),rT,code)"  | 
|
92  | 
by (auto simp add: wf_prog_def ws_prog_def wf_mdecl_def  | 
|
93  | 
wf_cdecl_mdecl_def ws_cdecl_def)  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
94  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
95  | 
lemma class_wf:  | 
| 14045 | 96  | 
"[|class G C = Some c; wf_prog wf_mb G|]  | 
97  | 
==> wf_cdecl wf_mb G (C,c) \<and> wf_mrT G (C,c)"  | 
|
98  | 
apply (unfold wf_prog_def ws_prog_def wf_cdecl_def class_def)  | 
|
99  | 
apply clarify  | 
|
100  | 
apply (drule_tac x="(C,c)" in bspec, fast dest: map_of_SomeD)  | 
|
101  | 
apply (drule_tac x="(C,c)" in bspec, fast dest: map_of_SomeD)  | 
|
102  | 
apply (simp add: wf_cdecl_def ws_cdecl_def wf_mdecl_def  | 
|
103  | 
wf_cdecl_mdecl_def wf_mrT_def split_beta)  | 
|
104  | 
done  | 
|
105  | 
||
106  | 
lemma class_wf_struct:  | 
|
107  | 
"[|class G C = Some c; ws_prog G|]  | 
|
108  | 
==> ws_cdecl G (C,c)"  | 
|
109  | 
apply (unfold ws_prog_def class_def)  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
110  | 
apply (fast dest: map_of_SomeD)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
111  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
112  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
113  | 
lemma class_Object [simp]:  | 
| 14045 | 114  | 
"ws_prog G ==> \<exists>X fs ms. class G Object = Some (X,fs,ms)"  | 
115  | 
apply (unfold ws_prog_def wf_syscls_def class_def)  | 
|
| 12951 | 116  | 
apply (auto simp: map_of_SomeI)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
117  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
118  | 
|
| 14045 | 119  | 
lemma class_Object_syscls [simp]:  | 
120  | 
"wf_syscls G ==> unique G \<Longrightarrow> \<exists>X fs ms. class G Object = Some (X,fs,ms)"  | 
|
121  | 
apply (unfold wf_syscls_def class_def)  | 
|
122  | 
apply (auto simp: map_of_SomeI)  | 
|
123  | 
done  | 
|
124  | 
||
125  | 
lemma is_class_Object [simp]: "ws_prog G ==> is_class G Object"  | 
|
| 18576 | 126  | 
by (simp add: is_class_def)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
127  | 
|
| 14045 | 128  | 
lemma is_class_xcpt [simp]: "ws_prog G \<Longrightarrow> is_class G (Xcpt x)"  | 
129  | 
apply (simp add: ws_prog_def wf_syscls_def)  | 
|
| 13051 | 130  | 
apply (simp add: is_class_def class_def)  | 
131  | 
apply clarify  | 
|
132  | 
apply (erule_tac x = x in allE)  | 
|
133  | 
apply clarify  | 
|
134  | 
apply (auto intro!: map_of_SomeI)  | 
|
135  | 
done  | 
|
136  | 
||
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
137  | 
lemma subcls1_wfD: "[|G\<turnstile>C\<prec>C1D; ws_prog G|] ==> D \<noteq> C \<and> (D, C) \<notin> (subcls1 G)^+"  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
138  | 
apply( frule trancl.r_into_trancl [where r="subcls1 G"])  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
139  | 
apply( drule subcls1D)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
140  | 
apply(clarify)  | 
| 14045 | 141  | 
apply( drule (1) class_wf_struct)  | 
142  | 
apply( unfold ws_cdecl_def)  | 
|
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
143  | 
apply(force simp add: reflcl_trancl [THEN sym] simp del: reflcl_trancl)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
144  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
145  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
146  | 
lemma wf_cdecl_supD:  | 
| 14045 | 147  | 
"!!r. \<lbrakk>ws_cdecl G (C,D,r); C \<noteq> Object\<rbrakk> \<Longrightarrow> is_class G D"  | 
148  | 
apply (unfold ws_cdecl_def)  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
149  | 
apply (auto split add: option.split_asm)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
150  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
151  | 
|
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
152  | 
lemma subcls_asym: "[|ws_prog G; (C, D) \<in> (subcls1 G)^+|] ==> (D, C) \<notin> (subcls1 G)^+"  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
153  | 
apply(erule trancl.cases)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
154  | 
apply(fast dest!: subcls1_wfD )  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
155  | 
apply(fast dest!: subcls1_wfD intro: trancl_trans)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
156  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
157  | 
|
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
158  | 
lemma subcls_irrefl: "[|ws_prog G; (C, D) \<in> (subcls1 G)^+|] ==> C \<noteq> D"  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
159  | 
apply (erule trancl_trans_induct)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
160  | 
apply (auto dest: subcls1_wfD subcls_asym)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
161  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
162  | 
|
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
163  | 
lemma acyclic_subcls1: "ws_prog G ==> acyclic (subcls1 G)"  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
164  | 
apply (simp add: acyclic_def)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
165  | 
apply (fast dest: subcls_irrefl)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
166  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
167  | 
|
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
168  | 
lemma wf_subcls1: "ws_prog G ==> wf ((subcls1 G)^-1)"  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
169  | 
apply (rule finite_acyclic_wf)  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
170  | 
apply ( subst finite_converse)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
171  | 
apply ( rule finite_subcls1)  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
172  | 
apply (subst acyclic_converse)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
173  | 
apply (erule acyclic_subcls1)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
174  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
175  | 
|
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
176  | 
lemma subcls_induct_struct:  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
177  | 
"[|ws_prog G; !!C. \<forall>D. (C, D) \<in> (subcls1 G)^+ --> P D ==> P C|] ==> P C"  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
178  | 
(is "?A \<Longrightarrow> PROP ?P \<Longrightarrow> _")  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
179  | 
proof -  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
180  | 
assume p: "PROP ?P"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
181  | 
assume ?A thus ?thesis apply -  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
182  | 
apply(drule wf_subcls1)  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
183  | 
apply(drule wf_trancl)  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
184  | 
apply(simp only: trancl_converse)  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
185  | 
apply(erule_tac a = C in wf_induct)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
186  | 
apply(rule p)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
187  | 
apply(auto)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
188  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
189  | 
qed  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
190  | 
|
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
191  | 
lemma subcls_induct:  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
192  | 
"[|wf_prog wf_mb G; !!C. \<forall>D. (C, D) \<in> (subcls1 G)^+ --> P D ==> P C|] ==> P C"  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
193  | 
(is "?A \<Longrightarrow> PROP ?P \<Longrightarrow> _")  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
194  | 
by (fact subcls_induct_struct [OF wf_prog_ws_prog])  | 
| 
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
195  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
196  | 
lemma subcls1_induct:  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
197  | 
"[|is_class G C; wf_prog wf_mb G; P Object;  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
198  | 
!!C D fs ms. [|C \<noteq> Object; is_class G C; class G C = Some (D,fs,ms) \<and>  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
199  | 
wf_cdecl wf_mb G (C,D,fs,ms) \<and> G\<turnstile>C\<prec>C1D \<and> is_class G D \<and> P D|] ==> P C  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
200  | 
|] ==> P C"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
201  | 
(is "?A \<Longrightarrow> ?B \<Longrightarrow> ?C \<Longrightarrow> PROP ?P \<Longrightarrow> _")  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
202  | 
proof -  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
203  | 
assume p: "PROP ?P"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
204  | 
assume ?A ?B ?C thus ?thesis apply -  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
205  | 
apply(unfold is_class_def)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
206  | 
apply( rule impE)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
207  | 
prefer 2  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
208  | 
apply( assumption)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
209  | 
prefer 2  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
210  | 
apply( assumption)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
211  | 
apply( erule thin_rl)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
212  | 
apply( rule subcls_induct)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
213  | 
apply( assumption)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
214  | 
apply( rule impI)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
215  | 
apply( case_tac "C = Object")  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
216  | 
apply( fast)  | 
| 18447 | 217  | 
apply auto  | 
| 14045 | 218  | 
apply( frule (1) class_wf) apply (erule conjE)+  | 
219  | 
apply (frule wf_cdecl_ws_cdecl)  | 
|
220  | 
apply( frule (1) wf_cdecl_supD)  | 
|
221  | 
||
222  | 
apply( subgoal_tac "G\<turnstile>C\<prec>C1a")  | 
|
223  | 
apply( erule_tac [2] subcls1I)  | 
|
224  | 
apply( rule p)  | 
|
225  | 
apply (unfold is_class_def)  | 
|
226  | 
apply auto  | 
|
227  | 
done  | 
|
228  | 
qed  | 
|
229  | 
||
230  | 
lemma subcls1_induct_struct:  | 
|
231  | 
"[|is_class G C; ws_prog G; P Object;  | 
|
232  | 
!!C D fs ms. [|C \<noteq> Object; is_class G C; class G C = Some (D,fs,ms) \<and>  | 
|
233  | 
ws_cdecl G (C,D,fs,ms) \<and> G\<turnstile>C\<prec>C1D \<and> is_class G D \<and> P D|] ==> P C  | 
|
234  | 
|] ==> P C"  | 
|
235  | 
(is "?A \<Longrightarrow> ?B \<Longrightarrow> ?C \<Longrightarrow> PROP ?P \<Longrightarrow> _")  | 
|
236  | 
proof -  | 
|
237  | 
assume p: "PROP ?P"  | 
|
238  | 
assume ?A ?B ?C thus ?thesis apply -  | 
|
239  | 
apply(unfold is_class_def)  | 
|
240  | 
apply( rule impE)  | 
|
241  | 
prefer 2  | 
|
242  | 
apply( assumption)  | 
|
243  | 
prefer 2  | 
|
244  | 
apply( assumption)  | 
|
245  | 
apply( erule thin_rl)  | 
|
246  | 
apply( rule subcls_induct_struct)  | 
|
247  | 
apply( assumption)  | 
|
248  | 
apply( rule impI)  | 
|
249  | 
apply( case_tac "C = Object")  | 
|
250  | 
apply( fast)  | 
|
| 18447 | 251  | 
apply auto  | 
| 14045 | 252  | 
apply( frule (1) class_wf_struct)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
253  | 
apply( frule (1) wf_cdecl_supD)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
254  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
255  | 
apply( subgoal_tac "G\<turnstile>C\<prec>C1a")  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
256  | 
apply( erule_tac [2] subcls1I)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
257  | 
apply( rule p)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
258  | 
apply (unfold is_class_def)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
259  | 
apply auto  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
260  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
261  | 
qed  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
262  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
263  | 
lemmas method_rec = wf_subcls1 [THEN [2] method_rec_lemma];  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
264  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
265  | 
lemmas fields_rec = wf_subcls1 [THEN [2] fields_rec_lemma];  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
266  | 
|
| 14045 | 267  | 
lemma field_rec: "\<lbrakk>class G C = Some (D, fs, ms); ws_prog G\<rbrakk>  | 
| 13672 | 268  | 
\<Longrightarrow> field (G, C) =  | 
269  | 
(if C = Object then empty else field (G, D)) ++  | 
|
270  | 
map_of (map (\<lambda>(s, f). (s, C, f)) fs)"  | 
|
271  | 
apply (simp only: field_def)  | 
|
272  | 
apply (frule fields_rec, assumption)  | 
|
273  | 
apply (rule HOL.trans)  | 
|
274  | 
apply (simp add: o_def)  | 
|
| 33640 | 275  | 
apply (simp (no_asm_use) add: split_beta split_def o_def)  | 
| 13672 | 276  | 
done  | 
277  | 
||
| 12951 | 278  | 
lemma method_Object [simp]:  | 
| 14045 | 279  | 
"method (G, Object) sig = Some (D, mh, code) \<Longrightarrow> ws_prog G \<Longrightarrow> D = Object"  | 
| 12951 | 280  | 
apply (frule class_Object, clarify)  | 
281  | 
apply (drule method_rec, assumption)  | 
|
282  | 
apply (auto dest: map_of_SomeD)  | 
|
283  | 
done  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
284  | 
|
| 13672 | 285  | 
|
| 14045 | 286  | 
lemma fields_Object [simp]: "\<lbrakk> ((vn, C), T) \<in> set (fields (G, Object)); ws_prog G \<rbrakk>  | 
| 13672 | 287  | 
\<Longrightarrow> C = Object"  | 
288  | 
apply (frule class_Object)  | 
|
289  | 
apply clarify  | 
|
290  | 
apply (subgoal_tac "fields (G, Object) = map (\<lambda>(fn,ft). ((fn,Object),ft)) fs")  | 
|
291  | 
apply (simp add: image_iff split_beta)  | 
|
292  | 
apply auto  | 
|
293  | 
apply (rule trans)  | 
|
294  | 
apply (rule fields_rec, assumption+)  | 
|
295  | 
apply simp  | 
|
296  | 
done  | 
|
297  | 
||
| 14045 | 298  | 
lemma subcls_C_Object: "[|is_class G C; ws_prog G|] ==> G\<turnstile>C\<preceq>C Object"  | 
299  | 
apply(erule subcls1_induct_struct)  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
300  | 
apply( assumption)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
301  | 
apply( fast)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
302  | 
apply(auto dest!: wf_cdecl_supD)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
303  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
304  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
305  | 
lemma is_type_rTI: "wf_mhead G sig rT ==> is_type G rT"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
306  | 
apply (unfold wf_mhead_def)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
307  | 
apply auto  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
308  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
309  | 
|
| 14045 | 310  | 
lemma widen_fields_defpl': "[|is_class G C; ws_prog G|] ==>  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
311  | 
\<forall>((fn,fd),fT)\<in>set (fields (G,C)). G\<turnstile>C\<preceq>C fd"  | 
| 14045 | 312  | 
apply( erule subcls1_induct_struct)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
313  | 
apply( assumption)  | 
| 12951 | 314  | 
apply( frule class_Object)  | 
315  | 
apply( clarify)  | 
|
316  | 
apply( frule fields_rec, assumption)  | 
|
317  | 
apply( fastsimp)  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
318  | 
apply( tactic "safe_tac HOL_cs")  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
319  | 
apply( subst fields_rec)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
320  | 
apply( assumption)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
321  | 
apply( assumption)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
322  | 
apply( simp (no_asm) split del: split_if)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
323  | 
apply( rule ballI)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
324  | 
apply( simp (no_asm_simp) only: split_tupled_all)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
325  | 
apply( simp (no_asm))  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
326  | 
apply( erule UnE)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
327  | 
apply( force)  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
328  | 
apply( erule r_into_rtrancl [THEN rtrancl_trans])  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
329  | 
apply auto  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
330  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
331  | 
|
| 12517 | 332  | 
lemma widen_fields_defpl:  | 
| 14045 | 333  | 
"[|((fn,fd),fT) \<in> set (fields (G,C)); ws_prog G; is_class G C|] ==>  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
334  | 
G\<turnstile>C\<preceq>C fd"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
335  | 
apply( drule (1) widen_fields_defpl')  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
336  | 
apply (fast)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
337  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
338  | 
|
| 12517 | 339  | 
lemma unique_fields:  | 
| 14045 | 340  | 
"[|is_class G C; ws_prog G|] ==> unique (fields (G,C))"  | 
341  | 
apply( erule subcls1_induct_struct)  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
342  | 
apply( assumption)  | 
| 12951 | 343  | 
apply( frule class_Object)  | 
344  | 
apply( clarify)  | 
|
345  | 
apply( frule fields_rec, assumption)  | 
|
| 14045 | 346  | 
apply( drule class_wf_struct, assumption)  | 
347  | 
apply( simp add: ws_cdecl_def)  | 
|
| 12951 | 348  | 
apply( rule unique_map_inj)  | 
349  | 
apply( simp)  | 
|
| 13585 | 350  | 
apply( rule inj_onI)  | 
| 12951 | 351  | 
apply( simp)  | 
352  | 
apply( safe dest!: wf_cdecl_supD)  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
353  | 
apply( drule subcls1_wfD)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
354  | 
apply( assumption)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
355  | 
apply( subst fields_rec)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
356  | 
apply auto  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
357  | 
apply( rotate_tac -1)  | 
| 14045 | 358  | 
apply( frule class_wf_struct)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
359  | 
apply auto  | 
| 14045 | 360  | 
apply( simp add: ws_cdecl_def)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
361  | 
apply( erule unique_append)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
362  | 
apply( rule unique_map_inj)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
363  | 
apply( clarsimp)  | 
| 13585 | 364  | 
apply (rule inj_onI)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
365  | 
apply( simp)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
366  | 
apply(auto dest!: widen_fields_defpl)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
367  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
368  | 
|
| 12517 | 369  | 
lemma fields_mono_lemma [rule_format (no_asm)]:  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
370  | 
"[|ws_prog G; (C', C) \<in> (subcls1 G)^*|] ==>  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
371  | 
x \<in> set (fields (G,C)) --> x \<in> set (fields (G,C'))"  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
372  | 
apply(erule converse_rtrancl_induct)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
373  | 
apply( safe dest!: subcls1D)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
374  | 
apply(subst fields_rec)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
375  | 
apply( auto)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
376  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
377  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
378  | 
lemma fields_mono:  | 
| 14045 | 379  | 
"\<lbrakk>map_of (fields (G,C)) fn = Some f; G\<turnstile>D\<preceq>C C; is_class G D; ws_prog G\<rbrakk>  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
380  | 
\<Longrightarrow> map_of (fields (G,D)) fn = Some f"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
381  | 
apply (rule map_of_SomeI)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
382  | 
apply (erule (1) unique_fields)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
383  | 
apply (erule (1) fields_mono_lemma)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
384  | 
apply (erule map_of_SomeD)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
385  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
386  | 
|
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
387  | 
lemma widen_cfs_fields:  | 
| 14045 | 388  | 
"[|field (G,C) fn = Some (fd, fT); G\<turnstile>D\<preceq>C C; ws_prog G|]==>  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
389  | 
map_of (fields (G,D)) (fn, fd) = Some fT"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
390  | 
apply (drule field_fields)  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
391  | 
apply (drule rtranclD)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
392  | 
apply safe  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
393  | 
apply (frule subcls_is_class)  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
394  | 
apply (drule trancl_into_rtrancl)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
395  | 
apply (fast dest: fields_mono)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
396  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
397  | 
|
| 12517 | 398  | 
lemma method_wf_mdecl [rule_format (no_asm)]:  | 
399  | 
"wf_prog wf_mb G ==> is_class G C \<Longrightarrow>  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
400  | 
method (G,C) sig = Some (md,mh,m)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
401  | 
--> G\<turnstile>C\<preceq>C md \<and> wf_mdecl wf_mb G md (sig,(mh,m))"  | 
| 14045 | 402  | 
apply (frule wf_prog_ws_prog)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
403  | 
apply( erule subcls1_induct)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
404  | 
apply( assumption)  | 
| 12951 | 405  | 
apply( clarify)  | 
406  | 
apply( frule class_Object)  | 
|
407  | 
apply( clarify)  | 
|
408  | 
apply( frule method_rec, assumption)  | 
|
409  | 
apply( drule class_wf, assumption)  | 
|
410  | 
apply( simp add: wf_cdecl_def)  | 
|
411  | 
apply( drule map_of_SomeD)  | 
|
412  | 
apply( subgoal_tac "md = Object")  | 
|
| 14045 | 413  | 
apply( fastsimp)  | 
| 12951 | 414  | 
apply( fastsimp)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
415  | 
apply( clarify)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
416  | 
apply( frule_tac C = C in method_rec)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
417  | 
apply( assumption)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
418  | 
apply( rotate_tac -1)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
419  | 
apply( simp)  | 
| 14025 | 420  | 
apply( drule map_add_SomeD)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
421  | 
apply( erule disjE)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
422  | 
apply( erule_tac V = "?P --> ?Q" in thin_rl)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
423  | 
apply (frule map_of_SomeD)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
424  | 
apply (clarsimp simp add: wf_cdecl_def)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
425  | 
apply( clarify)  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
426  | 
apply( rule rtrancl_trans)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
427  | 
prefer 2  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
428  | 
apply( assumption)  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
429  | 
apply( rule r_into_rtrancl)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
430  | 
apply( fast intro: subcls1I)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
431  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
432  | 
|
| 13672 | 433  | 
|
| 14045 | 434  | 
lemma method_wf_mhead [rule_format (no_asm)]:  | 
435  | 
"ws_prog G ==> is_class G C \<Longrightarrow>  | 
|
436  | 
method (G,C) sig = Some (md,rT,mb)  | 
|
437  | 
--> G\<turnstile>C\<preceq>C md \<and> wf_mhead G sig rT"  | 
|
438  | 
apply( erule subcls1_induct_struct)  | 
|
439  | 
apply( assumption)  | 
|
440  | 
apply( clarify)  | 
|
441  | 
apply( frule class_Object)  | 
|
442  | 
apply( clarify)  | 
|
443  | 
apply( frule method_rec, assumption)  | 
|
444  | 
apply( drule class_wf_struct, assumption)  | 
|
445  | 
apply( simp add: ws_cdecl_def)  | 
|
446  | 
apply( drule map_of_SomeD)  | 
|
447  | 
apply( subgoal_tac "md = Object")  | 
|
448  | 
apply( fastsimp)  | 
|
449  | 
apply( fastsimp)  | 
|
450  | 
apply( clarify)  | 
|
451  | 
apply( frule_tac C = C in method_rec)  | 
|
452  | 
apply( assumption)  | 
|
453  | 
apply( rotate_tac -1)  | 
|
454  | 
apply( simp)  | 
|
455  | 
apply( drule map_add_SomeD)  | 
|
456  | 
apply( erule disjE)  | 
|
457  | 
apply( erule_tac V = "?P --> ?Q" in thin_rl)  | 
|
458  | 
apply (frule map_of_SomeD)  | 
|
459  | 
apply (clarsimp simp add: ws_cdecl_def)  | 
|
460  | 
apply blast  | 
|
461  | 
apply clarify  | 
|
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
462  | 
apply( rule rtrancl_trans)  | 
| 14045 | 463  | 
prefer 2  | 
464  | 
apply( assumption)  | 
|
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
465  | 
apply( rule r_into_rtrancl)  | 
| 14045 | 466  | 
apply( fast intro: subcls1I)  | 
| 13672 | 467  | 
done  | 
468  | 
||
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
469  | 
lemma subcls_widen_methd [rule_format (no_asm)]:  | 
| 14045 | 470  | 
"[|G\<turnstile>T'\<preceq>C T; wf_prog wf_mb G|] ==>  | 
471  | 
\<forall>D rT b. method (G,T) sig = Some (D,rT ,b) -->  | 
|
472  | 
(\<exists>D' rT' b'. method (G,T') sig = Some (D',rT',b') \<and> G\<turnstile>D'\<preceq>C D \<and> G\<turnstile>rT'\<preceq>rT)"  | 
|
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
473  | 
apply( drule rtranclD)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
474  | 
apply( erule disjE)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
475  | 
apply( fast)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
476  | 
apply( erule conjE)  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
477  | 
apply( erule trancl_trans_induct)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
478  | 
prefer 2  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
479  | 
apply( clarify)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
480  | 
apply( drule spec, drule spec, drule spec, erule (1) impE)  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
481  | 
apply( fast elim: widen_trans rtrancl_trans)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
482  | 
apply( clarify)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
483  | 
apply( drule subcls1D)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
484  | 
apply( clarify)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
485  | 
apply( subst method_rec)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
486  | 
apply( assumption)  | 
| 14025 | 487  | 
apply( unfold map_add_def)  | 
| 14045 | 488  | 
apply( simp (no_asm_simp) add: wf_prog_ws_prog del: split_paired_Ex)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
489  | 
apply( case_tac "\<exists>z. map_of(map (\<lambda>(s,m). (s, ?C, m)) ms) sig = Some z")  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
490  | 
apply( erule exE)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
491  | 
apply( rotate_tac -1, frule ssubst, erule_tac [2] asm_rl)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
492  | 
prefer 2  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
493  | 
apply( rotate_tac -1, frule ssubst, erule_tac [2] asm_rl)  | 
| 22633 | 494  | 
apply(  tactic "asm_full_simp_tac (HOL_ss addsimps [@{thm not_None_eq} RS sym]) 1")
 | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
495  | 
apply( simp_all (no_asm_simp) del: split_paired_Ex)  | 
| 14045 | 496  | 
apply( frule (1) class_wf)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
497  | 
apply( simp (no_asm_simp) only: split_tupled_all)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
498  | 
apply( unfold wf_cdecl_def)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
499  | 
apply( drule map_of_SomeD)  | 
| 14045 | 500  | 
apply (auto simp add: wf_mrT_def)  | 
| 
33954
 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 
haftmann 
parents: 
33640 
diff
changeset
 | 
501  | 
apply (rule rtrancl_trans)  | 
| 14045 | 502  | 
defer  | 
503  | 
apply (rule method_wf_mhead [THEN conjunct1])  | 
|
504  | 
apply (simp only: wf_prog_def)  | 
|
505  | 
apply (simp add: is_class_def)  | 
|
506  | 
apply assumption  | 
|
507  | 
apply (auto intro: subcls1I)  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
508  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
509  | 
|
| 14045 | 510  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
511  | 
lemma subtype_widen_methd:  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
512  | 
"[| G\<turnstile> C\<preceq>C D; wf_prog wf_mb G;  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
513  | 
method (G,D) sig = Some (md, rT, b) |]  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
514  | 
==> \<exists>mD' rT' b'. method (G,C) sig= Some(mD',rT',b') \<and> G\<turnstile>rT'\<preceq>rT"  | 
| 14045 | 515  | 
apply(auto dest: subcls_widen_methd  | 
| 12517 | 516  | 
simp add: wf_mdecl_def wf_mhead_def split_def)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
517  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
518  | 
|
| 14045 | 519  | 
|
| 12517 | 520  | 
lemma method_in_md [rule_format (no_asm)]:  | 
| 14045 | 521  | 
"ws_prog G ==> is_class G C \<Longrightarrow> \<forall>D. method (G,C) sig = Some(D,mh,code)  | 
| 12517 | 522  | 
--> is_class G D \<and> method (G,D) sig = Some(D,mh,code)"  | 
| 14045 | 523  | 
apply (erule (1) subcls1_induct_struct)  | 
| 12951 | 524  | 
apply clarify  | 
525  | 
apply (frule method_Object, assumption)  | 
|
526  | 
apply hypsubst  | 
|
527  | 
apply simp  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
528  | 
apply (erule conjE)  | 
| 15481 | 529  | 
apply (simplesubst method_rec, assumption+)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
530  | 
apply (clarify)  | 
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14045 
diff
changeset
 | 
531  | 
apply (erule_tac x = "Da" in allE)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
532  | 
apply (clarsimp)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
533  | 
apply (simp add: map_of_map)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
534  | 
apply (clarify)  | 
| 15481 | 535  | 
apply (subst method_rec, assumption+)  | 
| 14025 | 536  | 
apply (simp add: map_add_def map_of_map split add: option.split)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
537  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
538  | 
|
| 13672 | 539  | 
|
| 14045 | 540  | 
lemma method_in_md_struct [rule_format (no_asm)]:  | 
541  | 
"ws_prog G ==> is_class G C \<Longrightarrow> \<forall>D. method (G,C) sig = Some(D,mh,code)  | 
|
542  | 
--> is_class G D \<and> method (G,D) sig = Some(D,mh,code)"  | 
|
543  | 
apply (erule (1) subcls1_induct_struct)  | 
|
544  | 
apply clarify  | 
|
545  | 
apply (frule method_Object, assumption)  | 
|
546  | 
apply hypsubst  | 
|
547  | 
apply simp  | 
|
548  | 
apply (erule conjE)  | 
|
| 15481 | 549  | 
apply (simplesubst method_rec, assumption+)  | 
| 14045 | 550  | 
apply (clarify)  | 
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
14045 
diff
changeset
 | 
551  | 
apply (erule_tac x = "Da" in allE)  | 
| 14045 | 552  | 
apply (clarsimp)  | 
553  | 
apply (simp add: map_of_map)  | 
|
554  | 
apply (clarify)  | 
|
| 15481 | 555  | 
apply (subst method_rec, assumption+)  | 
| 14045 | 556  | 
apply (simp add: map_add_def map_of_map split add: option.split)  | 
557  | 
done  | 
|
558  | 
||
| 13672 | 559  | 
lemma fields_in_fd [rule_format (no_asm)]: "\<lbrakk> wf_prog wf_mb G; is_class G C\<rbrakk>  | 
560  | 
\<Longrightarrow> \<forall> vn D T. (((vn,D),T) \<in> set (fields (G,C))  | 
|
561  | 
\<longrightarrow> (is_class G D \<and> ((vn,D),T) \<in> set (fields (G,D))))"  | 
|
562  | 
apply (erule (1) subcls1_induct)  | 
|
563  | 
||
564  | 
apply clarify  | 
|
| 14045 | 565  | 
apply (frule wf_prog_ws_prog)  | 
| 13672 | 566  | 
apply (frule fields_Object, assumption+)  | 
567  | 
apply (simp only: is_class_Object) apply simp  | 
|
568  | 
||
569  | 
apply clarify  | 
|
570  | 
apply (frule fields_rec)  | 
|
| 14045 | 571  | 
apply (simp (no_asm_simp) add: wf_prog_ws_prog)  | 
| 13672 | 572  | 
|
573  | 
apply (case_tac "Da=C")  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
23757 
diff
changeset
 | 
574  | 
apply blast (* case Da=C *)  | 
| 13672 | 575  | 
|
576  | 
apply (subgoal_tac "((vn, Da), T) \<in> set (fields (G, D))") apply blast  | 
|
577  | 
apply (erule thin_rl)  | 
|
578  | 
apply (rotate_tac 1)  | 
|
579  | 
apply (erule thin_rl, erule thin_rl, erule thin_rl,  | 
|
580  | 
erule thin_rl, erule thin_rl, erule thin_rl)  | 
|
581  | 
apply auto  | 
|
582  | 
done  | 
|
583  | 
||
584  | 
lemma field_in_fd [rule_format (no_asm)]: "\<lbrakk> wf_prog wf_mb G; is_class G C\<rbrakk>  | 
|
585  | 
\<Longrightarrow> \<forall> vn D T. (field (G,C) vn = Some(D,T)  | 
|
586  | 
\<longrightarrow> is_class G D \<and> field (G,D) vn = Some(D,T))"  | 
|
587  | 
apply (erule (1) subcls1_induct)  | 
|
588  | 
||
589  | 
apply clarify  | 
|
590  | 
apply (frule field_fields)  | 
|
591  | 
apply (drule map_of_SomeD)  | 
|
| 14045 | 592  | 
apply (frule wf_prog_ws_prog)  | 
| 13672 | 593  | 
apply (drule fields_Object, assumption+)  | 
594  | 
apply simp  | 
|
595  | 
||
596  | 
apply clarify  | 
|
597  | 
apply (subgoal_tac "((field (G, D)) ++ map_of (map (\<lambda>(s, f). (s, C, f)) fs)) vn = Some (Da, T)")  | 
|
| 14025 | 598  | 
apply (simp (no_asm_use) only: map_add_Some_iff)  | 
| 13672 | 599  | 
apply (erule disjE)  | 
600  | 
apply (simp (no_asm_use) add: map_of_map) apply blast  | 
|
601  | 
apply blast  | 
|
602  | 
apply (rule trans [THEN sym], rule sym, assumption)  | 
|
603  | 
apply (rule_tac x=vn in fun_cong)  | 
|
604  | 
apply (rule trans, rule field_rec, assumption+)  | 
|
| 14045 | 605  | 
apply (simp (no_asm_simp) add: wf_prog_ws_prog)  | 
| 13672 | 606  | 
apply (simp (no_asm_use)) apply blast  | 
607  | 
done  | 
|
608  | 
||
609  | 
||
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
610  | 
lemma widen_methd:  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
611  | 
"[| method (G,C) sig = Some (md,rT,b); wf_prog wf_mb G; G\<turnstile>T''\<preceq>C C|]  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
612  | 
==> \<exists>md' rT' b'. method (G,T'') sig = Some (md',rT',b') \<and> G\<turnstile>rT'\<preceq>rT"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
613  | 
apply( drule subcls_widen_methd)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
614  | 
apply auto  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
615  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
616  | 
|
| 13672 | 617  | 
lemma widen_field: "\<lbrakk> (field (G,C) fn) = Some (fd, fT); wf_prog wf_mb G; is_class G C \<rbrakk>  | 
618  | 
\<Longrightarrow> G\<turnstile>C\<preceq>C fd"  | 
|
619  | 
apply (rule widen_fields_defpl)  | 
|
620  | 
apply (simp add: field_def)  | 
|
621  | 
apply (rule map_of_SomeD)  | 
|
622  | 
apply (rule table_of_remap_SomeD)  | 
|
623  | 
apply assumption+  | 
|
| 14045 | 624  | 
apply (simp (no_asm_simp) add: wf_prog_ws_prog)+  | 
| 13672 | 625  | 
done  | 
626  | 
||
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
627  | 
lemma Call_lemma:  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
628  | 
"[|method (G,C) sig = Some (md,rT,b); G\<turnstile>T''\<preceq>C C; wf_prog wf_mb G;  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
629  | 
class G C = Some y|] ==> \<exists>T' rT' b. method (G,T'') sig = Some (T',rT',b) \<and>  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
630  | 
G\<turnstile>rT'\<preceq>rT \<and> G\<turnstile>T''\<preceq>C T' \<and> wf_mhead G sig rT' \<and> wf_mb G T' (sig,rT',b)"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
631  | 
apply( drule (2) widen_methd)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
632  | 
apply( clarify)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
633  | 
apply( frule subcls_is_class2)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
634  | 
apply (unfold is_class_def)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
635  | 
apply (simp (no_asm_simp))  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
636  | 
apply( drule method_wf_mdecl)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
637  | 
apply( unfold wf_mdecl_def)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
638  | 
apply( unfold is_class_def)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
639  | 
apply auto  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
640  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
641  | 
|
| 12517 | 642  | 
lemma fields_is_type_lemma [rule_format (no_asm)]:  | 
| 14045 | 643  | 
"[|is_class G C; ws_prog G|] ==>  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
644  | 
\<forall>f\<in>set (fields (G,C)). is_type G (snd f)"  | 
| 14045 | 645  | 
apply( erule (1) subcls1_induct_struct)  | 
| 12951 | 646  | 
apply( frule class_Object)  | 
647  | 
apply( clarify)  | 
|
648  | 
apply( frule fields_rec, assumption)  | 
|
| 14045 | 649  | 
apply( drule class_wf_struct, assumption)  | 
650  | 
apply( simp add: ws_cdecl_def wf_fdecl_def)  | 
|
| 12951 | 651  | 
apply( fastsimp)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
652  | 
apply( subst fields_rec)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
653  | 
apply( fast)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
654  | 
apply( assumption)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
655  | 
apply( clarsimp)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
656  | 
apply( safe)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
657  | 
prefer 2  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
658  | 
apply( force)  | 
| 14045 | 659  | 
apply( drule (1) class_wf_struct)  | 
660  | 
apply( unfold ws_cdecl_def)  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
661  | 
apply( clarsimp)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
662  | 
apply( drule (1) bspec)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
663  | 
apply( unfold wf_fdecl_def)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
664  | 
apply auto  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
665  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
666  | 
|
| 14045 | 667  | 
|
| 12517 | 668  | 
lemma fields_is_type:  | 
| 14045 | 669  | 
"[|map_of (fields (G,C)) fn = Some f; ws_prog G; is_class G C|] ==>  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
670  | 
is_type G f"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
671  | 
apply(drule map_of_SomeD)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
672  | 
apply(drule (2) fields_is_type_lemma)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
673  | 
apply(auto)  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
674  | 
done  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
675  | 
|
| 14045 | 676  | 
|
677  | 
lemma field_is_type: "\<lbrakk> ws_prog G; is_class G C; field (G, C) fn = Some (fd, fT) \<rbrakk>  | 
|
678  | 
\<Longrightarrow> is_type G fT"  | 
|
679  | 
apply (frule_tac f="((fn, fd), fT)" in fields_is_type_lemma)  | 
|
680  | 
apply (auto simp add: field_def dest: map_of_SomeD)  | 
|
681  | 
done  | 
|
682  | 
||
683  | 
||
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
684  | 
lemma methd:  | 
| 14045 | 685  | 
"[| ws_prog G; (C,S,fs,mdecls) \<in> set G; (sig,rT,code) \<in> set mdecls |]  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
686  | 
==> method (G,C) sig = Some(C,rT,code) \<and> is_class G C"  | 
| 
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
687  | 
proof -  | 
| 14045 | 688  | 
assume wf: "ws_prog G" and C: "(C,S,fs,mdecls) \<in> set G" and  | 
| 12951 | 689  | 
m: "(sig,rT,code) \<in> set mdecls"  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
690  | 
moreover  | 
| 12951 | 691  | 
from wf C have "class G C = Some (S,fs,mdecls)"  | 
| 14045 | 692  | 
by (auto simp add: ws_prog_def class_def is_class_def intro: map_of_SomeI)  | 
| 12951 | 693  | 
moreover  | 
694  | 
from wf C  | 
|
| 14045 | 695  | 
have "unique mdecls" by (unfold ws_prog_def ws_cdecl_def) auto  | 
| 12951 | 696  | 
hence "unique (map (\<lambda>(s,m). (s,C,m)) mdecls)" by (induct mdecls, auto)  | 
697  | 
with m  | 
|
698  | 
have "map_of (map (\<lambda>(s,m). (s,C,m)) mdecls) sig = Some (C,rT,code)"  | 
|
699  | 
by (force intro: map_of_SomeI)  | 
|
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
700  | 
ultimately  | 
| 12517 | 701  | 
show ?thesis by (auto simp add: is_class_def dest: method_rec)  | 
| 
11026
 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 
oheimb 
parents: 
10613 
diff
changeset
 | 
702  | 
qed  | 
| 8011 | 703  | 
|
| 12951 | 704  | 
|
705  | 
lemma wf_mb'E:  | 
|
706  | 
"\<lbrakk> wf_prog wf_mb G; \<And>C S fs ms m.\<lbrakk>(C,S,fs,ms) \<in> set G; m \<in> set ms\<rbrakk> \<Longrightarrow> wf_mb' G C m \<rbrakk>  | 
|
707  | 
\<Longrightarrow> wf_prog wf_mb' G"  | 
|
| 14045 | 708  | 
apply (simp only: wf_prog_def)  | 
| 12951 | 709  | 
apply auto  | 
| 14045 | 710  | 
apply (simp add: wf_cdecl_mdecl_def)  | 
| 12951 | 711  | 
apply safe  | 
712  | 
apply (drule bspec, assumption) apply simp  | 
|
713  | 
done  | 
|
714  | 
||
715  | 
||
716  | 
lemma fst_mono: "A \<subseteq> B \<Longrightarrow> fst ` A \<subseteq> fst ` B" by fast  | 
|
717  | 
||
718  | 
lemma wf_syscls:  | 
|
719  | 
"set SystemClasses \<subseteq> set G \<Longrightarrow> wf_syscls G"  | 
|
720  | 
apply (drule fst_mono)  | 
|
721  | 
apply (simp add: SystemClasses_def wf_syscls_def)  | 
|
722  | 
apply (simp add: ObjectC_def)  | 
|
723  | 
apply (rule allI, case_tac x)  | 
|
724  | 
apply (auto simp add: NullPointerC_def ClassCastC_def OutOfMemoryC_def)  | 
|
725  | 
done  | 
|
726  | 
||
| 8011 | 727  | 
end  |