(* Title: HOL/Bali/WellForm.thy
Author: David von Oheimb and Norbert Schirmer
*)
subsection \<open>Well-formedness of Java programs\<close>
theory WellForm imports DefiniteAssignment begin
text \<open>
For static checks on expressions and statements, see WellType.thy
improvements over Java Specification 1.0 (cf. 8.4.6.3, 8.4.6.4, 9.4.1):
\begin{itemize}
\item a method implementing or overwriting another method may have a result
type that widens to the result type of the other method
(instead of identical type)
\item if a method hides another method (both methods have to be static!)
there are no restrictions to the result type
since the methods have to be static and there is no dynamic binding of
static methods
\item if an interface inherits more than one method with the same signature, the
methods need not have identical return types
\end{itemize}
simplifications:
\begin{itemize}
\item Object and standard exceptions are assumed to be declared like normal
classes
\end{itemize}
\<close>
subsubsection "well-formed field declarations"
text \<open>well-formed field declaration (common part for classes and interfaces),
cf. 8.3 and (9.3)\<close>
definition
wf_fdecl :: "prog \ pname \ fdecl \ bool"
where "wf_fdecl G P = (\(fn,f). is_acc_type G P (type f))"
lemma wf_fdecl_def2: "\fd. wf_fdecl G P fd = is_acc_type G P (type (snd fd))"
apply (unfold wf_fdecl_def)
apply simp
done
subsubsection "well-formed method declarations"
(*well-formed method declaration,cf. 8.4, 8.4.1, 8.4.3, 8.4.5, 14.3.2, (9.4)*)
(* cf. 14.15, 15.7.2, for scope issues cf. 8.4.1 and 14.3.2 *)
text \<open>
A method head is wellformed if:
\begin{itemize}
\item the signature and the method head agree in the number of parameters
\item all types of the parameters are visible
\item the result type is visible
\item the parameter names are unique
\end{itemize}
\<close>
definition
wf_mhead :: "prog \ pname \ sig \ mhead \ bool" where
"wf_mhead G P = (\ sig mh. length (parTs sig) = length (pars mh) \
( \<forall>T\<in>set (parTs sig). is_acc_type G P T) \<and>
is_acc_type G P (resTy mh) \<and>
distinct (pars mh))"
text \<open>
A method declaration is wellformed if:
\begin{itemize}
\item the method head is wellformed
\item the names of the local variables are unique
\item the types of the local variables must be accessible
\item the local variables don't shadow the parameters
\item the class of the method is defined
\item the body statement is welltyped with respect to the
modified environment of local names, were the local variables,
the parameters the special result variable (Res) and This are assoziated
with there types.
\end{itemize}
\<close>
definition
callee_lcl :: "qtname \ sig \ methd \ lenv" where
"callee_lcl C sig m =
(\<lambda>k. (case k of
EName e
\<Rightarrow> (case e of
VNam v
\<Rightarrow>(table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v
| Res \<Rightarrow> Some (resTy m))
| This \<Rightarrow> if is_static m then None else Some (Class C)))"
definition
parameters :: "methd \ lname set" where
"parameters m = set (map (EName \ VNam) (pars m)) \ (if (static m) then {} else {This})"
definition
wf_mdecl :: "prog \ qtname \ mdecl \ bool" where
"wf_mdecl G C =
(\<lambda>(sig,m).
wf_mhead G (pid C) sig (mhead m) \<and>
unique (lcls (mbody m)) \<and>
(\<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T) \<and>
(\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None) \<and>
jumpNestingOkS {Ret} (stmt (mbody m)) \<and>
is_class G C \<and>
\<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd> \<and>
(\<exists> A. \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>
\<turnstile> parameters m \<guillemotright>\<langle>stmt (mbody m)\<rangle>\<guillemotright> A
\<and> Result \<in> nrm A))"
lemma callee_lcl_VNam_simp [simp]:
"callee_lcl C sig m (EName (VNam v))
= (table_of (lcls (mbody m))((pars m)[\<mapsto>](parTs sig))) v"
by (simp add: callee_lcl_def)
lemma callee_lcl_Res_simp [simp]:
"callee_lcl C sig m (EName Res) = Some (resTy m)"
by (simp add: callee_lcl_def)
lemma callee_lcl_This_simp [simp]:
"callee_lcl C sig m (This) = (if is_static m then None else Some (Class C))"
by (simp add: callee_lcl_def)
lemma callee_lcl_This_static_simp:
"is_static m \ callee_lcl C sig m (This) = None"
by simp
lemma callee_lcl_This_not_static_simp:
"\ is_static m \ callee_lcl C sig m (This) = Some (Class C)"
by simp
lemma wf_mheadI:
"\length (parTs sig) = length (pars m); \T\set (parTs sig). is_acc_type G P T;
is_acc_type G P (resTy m); distinct (pars m)\<rbrakk> \<Longrightarrow>
wf_mhead G P sig m"
apply (unfold wf_mhead_def)
apply (simp (no_asm_simp))
done
lemma wf_mdeclI: "\
wf_mhead G (pid C) sig (mhead m); unique (lcls (mbody m));
(\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None);
\<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T;
jumpNestingOkS {Ret} (stmt (mbody m));
is_class G C;
\<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd>;
(\<exists> A. \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr> \<turnstile> parameters m \<guillemotright>\<langle>stmt (mbody m)\<rangle>\<guillemotright> A
\<and> Result \<in> nrm A)
\<rbrakk> \<Longrightarrow>
wf_mdecl G C (sig,m)"
apply (unfold wf_mdecl_def)
apply simp
done
lemma wf_mdeclE [consumes 1]:
"\wf_mdecl G C (sig,m);
\<lbrakk>wf_mhead G (pid C) sig (mhead m); unique (lcls (mbody m));
\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None;
\<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T;
jumpNestingOkS {Ret} (stmt (mbody m));
is_class G C;
\<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile>(stmt (mbody m))\<Colon>\<surd>;
(\<exists> A. \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile> parameters m \<guillemotright>\<langle>stmt (mbody m)\<rangle>\<guillemotright> A
\<and> Result \<in> nrm A)
\<rbrakk> \<Longrightarrow> P
\<rbrakk> \<Longrightarrow> P"
by (unfold wf_mdecl_def) simp
lemma wf_mdeclD1:
"wf_mdecl G C (sig,m) \
wf_mhead G (pid C) sig (mhead m) \<and> unique (lcls (mbody m)) \<and>
(\<forall>pn\<in>set (pars m). table_of (lcls (mbody m)) pn = None) \<and>
(\<forall>(vn,T)\<in>set (lcls (mbody m)). is_acc_type G (pid C) T)"
apply (unfold wf_mdecl_def)
apply simp
done
lemma wf_mdecl_bodyD:
"wf_mdecl G C (sig,m) \
(\<exists>T. \<lparr>prg=G,cls=C,lcl=callee_lcl C sig m\<rparr>\<turnstile>Body C (stmt (mbody m))\<Colon>-T \<and>
G\<turnstile>T\<preceq>(resTy m))"
apply (unfold wf_mdecl_def)
apply clarify
apply (rule_tac x="(resTy m)" in exI)
apply (unfold wf_mhead_def)
apply (auto simp add: wf_mhead_def is_acc_type_def intro: wt.Body )
done
(*
lemma static_Object_methodsE [elim!]:
"\<lbrakk>wf_mdecl G Object (sig, m);static m\<rbrakk> \<Longrightarrow> R"
apply (unfold wf_mdecl_def)
apply auto
done
*)
lemma rT_is_acc_type:
"wf_mhead G P sig m \ is_acc_type G P (resTy m)"
apply (unfold wf_mhead_def)
apply auto
done
subsubsection "well-formed interface declarations"
(* well-formed interface declaration, cf. 9.1, 9.1.2.1, 9.1.3, 9.4 *)
text \<open>
A interface declaration is wellformed if:
\begin{itemize}
\item the interface hierarchy is wellstructured
\item there is no class with the same name
\item the method heads are wellformed and not static and have Public access
\item the methods are uniquely named
\item all superinterfaces are accessible
\item the result type of a method overriding a method of Object widens to the
result type of the overridden method.
Shadowing static methods is forbidden.
\item the result type of a method overriding a set of methods defined in the
superinterfaces widens to each of the corresponding result types
\end{itemize}
\<close>
definition
wf_idecl :: "prog \ idecl \ bool" where
"wf_idecl G =
(\<lambda>(I,i).
ws_idecl G I (isuperIfs i) \<and>
\<not>is_class G I \<and>
(\<forall>(sig,mh)\<in>set (imethods i). wf_mhead G (pid I) sig mh \<and>
\<not>is_static mh \<and>
accmodi mh = Public) \<and>
unique (imethods i) \<and>
(\<forall> J\<in>set (isuperIfs i). is_acc_iface G (pid I) J) \<and>
(table_of (imethods i)
hiding (methd G Object)
under (\<lambda> new old. accmodi old \<noteq> Private)
entails (\<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old \<and>
is_static new = is_static old)) \<and>
(set_option \<circ> table_of (imethods i)
hidings Un_tables((\<lambda>J.(imethds G J))`set (isuperIfs i))
entails (\<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old)))"
lemma wf_idecl_mhead: "\wf_idecl G (I,i); (sig,mh)\set (imethods i)\ \
wf_mhead G (pid I) sig mh \<and> \<not>is_static mh \<and> accmodi mh = Public"
apply (unfold wf_idecl_def)
apply auto
done
lemma wf_idecl_hidings:
"wf_idecl G (I, i) \
(\<lambda>s. set_option (table_of (imethods i) s))
hidings Un_tables ((\<lambda>J. imethds G J) ` set (isuperIfs i))
entails \<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old"
apply (unfold wf_idecl_def o_def)
apply simp
done
lemma wf_idecl_hiding:
"wf_idecl G (I, i) \
(table_of (imethods i)
hiding (methd G Object)
under (\<lambda> new old. accmodi old \<noteq> Private)
entails (\<lambda>new old. G\<turnstile>resTy new\<preceq>resTy old \<and>
is_static new = is_static old))"
apply (unfold wf_idecl_def)
apply simp
done
lemma wf_idecl_supD:
"\wf_idecl G (I,i); J \ set (isuperIfs i)\
\<Longrightarrow> is_acc_iface G (pid I) J \<and> (J, I) \<notin> (subint1 G)\<^sup>+"
apply (unfold wf_idecl_def ws_idecl_def)
apply auto
done
subsubsection "well-formed class declarations"
(* well-formed class declaration, cf. 8.1, 8.1.2.1, 8.1.2.2, 8.1.3, 8.1.4 and
class method declaration, cf. 8.4.3.3, 8.4.6.1, 8.4.6.2, 8.4.6.3, 8.4.6.4 *)
text \<open>
A class declaration is wellformed if:
\begin{itemize}
\item there is no interface with the same name
\item all superinterfaces are accessible and for all methods implementing
an interface method the result type widens to the result type of
the interface method, the method is not static and offers at least
as much access
(this actually means that the method has Public access, since all
interface methods have public access)
\item all field declarations are wellformed and the field names are unique
\item all method declarations are wellformed and the method names are unique
\item the initialization statement is welltyped
\item the classhierarchy is wellstructured
\item Unless the class is Object:
\begin{itemize}
\item the superclass is accessible
\item for all methods overriding another method (of a superclass )the
result type widens to the result type of the overridden method,
the access modifier of the new method provides at least as much
access as the overwritten one.
\item for all methods hiding a method (of a superclass) the hidden
method must be static and offer at least as much access rights.
Remark: In contrast to the Java Language Specification we don't
restrict the result types of the method
(as in case of overriding), because there seems to be no reason,
since there is no dynamic binding of static methods.
(cf. 8.4.6.3 vs. 15.12.1).
Stricly speaking the restrictions on the access rights aren't
necessary to, since the static type and the access rights
together determine which method is to be called statically.
But if a class gains more then one static method with the
same signature due to inheritance, it is confusing when the
method selection depends on the access rights only:
e.g.
Class C declares static public method foo().
Class D is subclass of C and declares static method foo()
with default package access.
D.foo() ? if this call is in the same package as D then
foo of class D is called, otherwise foo of class C.
\end{itemize}
\end{itemize}
\<close>
(* to Table *)
definition
entails :: "('a,'b) table \ ('b \ bool) \ bool" ("_ entails _" 20)
where "(t entails P) = (\k. \ x \ t k: P x)"
lemma entailsD:
"\t entails P; t k = Some x\ \ P x"
by (simp add: entails_def)
lemma empty_entails[simp]: "Map.empty entails P"
by (simp add: entails_def)
definition
wf_cdecl :: "prog \ cdecl \ bool" where
"wf_cdecl G =
(\<lambda>(C,c).
\<not>is_iface G C \<and>
(\<forall>I\<in>set (superIfs c). is_acc_iface G (pid C) I \<and>
(\<forall>s. \<forall> im \<in> imethds G I s.
(\<exists> cm \<in> methd G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and>
\<not> is_static cm \<and>
accmodi im \<le> accmodi cm))) \<and>
(\<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f) \<and> unique (cfields c) \<and>
(\<forall>m\<in>set (methods c). wf_mdecl G C m) \<and> unique (methods c) \<and>
jumpNestingOkS {} (init c) \<and>
(\<exists> A. \<lparr>prg=G,cls=C,lcl=Map.empty\<rparr>\<turnstile> {} \<guillemotright>\<langle>init c\<rangle>\<guillemotright> A) \<and>
\<lparr>prg=G,cls=C,lcl=Map.empty\<rparr>\<turnstile>(init c)\<Colon>\<surd> \<and> ws_cdecl G C (super c) \<and>
(C \<noteq> Object \<longrightarrow>
(is_acc_class G (pid C) (super c) \<and>
(table_of (map (\<lambda> (s,m). (s,C,m)) (methods c))
entails (\<lambda> new. \<forall> old sig.
(G,sig\<turnstile>new overrides\<^sub>S old
\<longrightarrow> (G\<turnstile>resTy new\<preceq>resTy old \<and>
accmodi old \<le> accmodi new \<and>
\<not>is_static old)) \<and>
(G,sig\<turnstile>new hides old
\<longrightarrow> (accmodi old \<le> accmodi new \<and>
is_static old))))
)))"
(*
definition wf_cdecl :: "prog \<Rightarrow> cdecl \<Rightarrow> bool" where
"wf_cdecl G \<equiv>
\<lambda>(C,c).
\<not>is_iface G C \<and>
(\<forall>I\<in>set (superIfs c). is_acc_iface G (pid C) I \<and>
(\<forall>s. \<forall> im \<in> imethds G I s.
(\<exists> cm \<in> methd G C s: G\<turnstile>resTy (mthd cm)\<preceq>resTy (mthd im) \<and>
\<not> is_static cm \<and>
accmodi im \<le> accmodi cm))) \<and>
(\<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f) \<and> unique (cfields c) \<and>
(\<forall>m\<in>set (methods c). wf_mdecl G C m) \<and> unique (methods c) \<and>
\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd> \<and> ws_cdecl G C (super c) \<and>
(C \<noteq> Object \<longrightarrow>
(is_acc_class G (pid C) (super c) \<and>
(table_of (map (\<lambda> (s,m). (s,C,m)) (methods c))
hiding methd G (super c)
under (\<lambda> new old. G\<turnstile>new overrides old)
entails (\<lambda> new old.
(G\<turnstile>resTy (mthd new)\<preceq>resTy (mthd old) \<and>
accmodi old \<le> accmodi new \<and>
\<not> is_static old))) \<and>
(table_of (map (\<lambda> (s,m). (s,C,m)) (methods c))
hiding methd G (super c)
under (\<lambda> new old. G\<turnstile>new hides old)
entails (\<lambda> new old. is_static old \<and>
accmodi old \<le> accmodi new)) \<and>
(table_of (cfields c) hiding accfield G C (super c)
entails (\<lambda> newF oldF. accmodi oldF \<le> access newF))))"
*)
lemma wf_cdeclE [consumes 1]:
"\wf_cdecl G (C,c);
\<lbrakk>\<not>is_iface G C;
(\<forall>I\<in>set (superIfs c). is_acc_iface G (pid C) I \<and>
(\<forall>s. \<forall> im \<in> imethds G I s.
(\<exists> cm \<in> methd G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and>
\<not> is_static cm \<and>
accmodi im \<le> accmodi cm)));
\<forall>f\<in>set (cfields c). wf_fdecl G (pid C) f; unique (cfields c);
\<forall>m\<in>set (methods c). wf_mdecl G C m; unique (methods c);
jumpNestingOkS {} (init c);
\<exists> A. \<lparr>prg=G,cls=C,lcl=Map.empty\<rparr>\<turnstile> {} \<guillemotright>\<langle>init c\<rangle>\<guillemotright> A;
\<lparr>prg=G,cls=C,lcl=Map.empty\<rparr>\<turnstile>(init c)\<Colon>\<surd>;
ws_cdecl G C (super c);
(C \<noteq> Object \<longrightarrow>
(is_acc_class G (pid C) (super c) \<and>
(table_of (map (\<lambda> (s,m). (s,C,m)) (methods c))
entails (\<lambda> new. \<forall> old sig.
(G,sig\<turnstile>new overrides\<^sub>S old
\<longrightarrow> (G\<turnstile>resTy new\<preceq>resTy old \<and>
accmodi old \<le> accmodi new \<and>
\<not>is_static old)) \<and>
(G,sig\<turnstile>new hides old
\<longrightarrow> (accmodi old \<le> accmodi new \<and>
is_static old))))
))\<rbrakk> \<Longrightarrow> P
\<rbrakk> \<Longrightarrow> P"
by (unfold wf_cdecl_def) simp
lemma wf_cdecl_unique:
"wf_cdecl G (C,c) \ unique (cfields c) \ unique (methods c)"
apply (unfold wf_cdecl_def)
apply auto
done
lemma wf_cdecl_fdecl:
"\wf_cdecl G (C,c); f\set (cfields c)\ \ wf_fdecl G (pid C) f"
apply (unfold wf_cdecl_def)
apply auto
done
lemma wf_cdecl_mdecl:
"\wf_cdecl G (C,c); m\set (methods c)\ \ wf_mdecl G C m"
apply (unfold wf_cdecl_def)
apply auto
done
lemma wf_cdecl_impD:
"\wf_cdecl G (C,c); I\set (superIfs c)\
\<Longrightarrow> is_acc_iface G (pid C) I \<and>
(\<forall>s. \<forall>im \<in> imethds G I s.
(\<exists>cm \<in> methd G C s: G\<turnstile>resTy cm\<preceq>resTy im \<and> \<not>is_static cm \<and>
accmodi im \<le> accmodi cm))"
apply (unfold wf_cdecl_def)
apply auto
done
lemma wf_cdecl_supD:
"\wf_cdecl G (C,c); C \ Object\ \
is_acc_class G (pid C) (super c) \<and> (super c,C) \<notin> (subcls1 G)\<^sup>+ \<and>
(table_of (map (\<lambda> (s,m). (s,C,m)) (methods c))
entails (\<lambda> new. \<forall> old sig.
(G,sig\<turnstile>new overrides\<^sub>S old
\<longrightarrow> (G\<turnstile>resTy new\<preceq>resTy old \<and>
accmodi old \<le> accmodi new \<and>
\<not>is_static old)) \<and>
(G,sig\<turnstile>new hides old
\<longrightarrow> (accmodi old \<le> accmodi new \<and>
is_static old))))"
apply (unfold wf_cdecl_def ws_cdecl_def)
apply auto
done
lemma wf_cdecl_overrides_SomeD:
"\wf_cdecl G (C,c); C \ Object; table_of (methods c) sig = Some newM;
G,sig\<turnstile>(C,newM) overrides\<^sub>S old
\<rbrakk> \<Longrightarrow> G\<turnstile>resTy newM\<preceq>resTy old \<and>
accmodi old \<le> accmodi newM \<and>
\<not> is_static old"
apply (drule (1) wf_cdecl_supD)
apply (clarify)
apply (drule entailsD)
apply (blast intro: table_of_map_SomeI)
apply (drule_tac x="old" in spec)
apply (auto dest: overrides_eq_sigD simp add: msig_def)
done
lemma wf_cdecl_hides_SomeD:
"\wf_cdecl G (C,c); C \ Object; table_of (methods c) sig = Some newM;
G,sig\<turnstile>(C,newM) hides old
\<rbrakk> \<Longrightarrow> accmodi old \<le> access newM \<and>
is_static old"
apply (drule (1) wf_cdecl_supD)
apply (clarify)
apply (drule entailsD)
apply (blast intro: table_of_map_SomeI)
apply (drule_tac x="old" in spec)
apply (auto dest: hides_eq_sigD simp add: msig_def)
done
lemma wf_cdecl_wt_init:
"wf_cdecl G (C, c) \ \prg=G,cls=C,lcl=Map.empty\\init c\\"
apply (unfold wf_cdecl_def)
apply auto
done
subsubsection "well-formed programs"
(* well-formed program, cf. 8.1, 9.1 *)
text \<open>
A program declaration is wellformed if:
\begin{itemize}
\item the class ObjectC of Object is defined
\item every method of Object has an access modifier distinct from Package.
This is
necessary since every interface automatically inherits from Object.
We must know, that every time a Object method is "overriden" by an
interface method this is also overriden by the class implementing the
the interface (see \<open>implement_dynmethd and class_mheadsD\<close>)
\item all standard Exceptions are defined
\item all defined interfaces are wellformed
\item all defined classes are wellformed
\end{itemize}
\<close>
definition
wf_prog :: "prog \ bool" where
"wf_prog G = (let is = ifaces G; cs = classes G in
ObjectC \<in> set cs \<and>
(\<forall> m\<in>set Object_mdecls. accmodi m \<noteq> Package) \<and>
(\<forall>xn. SXcptC xn \<in> set cs) \<and>
(\<forall>i\<in>set is. wf_idecl G i) \<and> unique is \<and>
(\<forall>c\<in>set cs. wf_cdecl G c) \<and> unique cs)"
lemma wf_prog_idecl: "\iface G I = Some i; wf_prog G\ \ wf_idecl G (I,i)"
apply (unfold wf_prog_def Let_def)
apply simp
apply (fast dest: map_of_SomeD)
done
lemma wf_prog_cdecl: "\class G C = Some c; wf_prog G\ \ wf_cdecl G (C,c)"
apply (unfold wf_prog_def Let_def)
apply simp
apply (fast dest: map_of_SomeD)
done
lemma wf_prog_Object_mdecls:
"wf_prog G \ (\ m\set Object_mdecls. accmodi m \ Package)"
apply (unfold wf_prog_def Let_def)
apply simp
done
lemma wf_prog_acc_superD:
"\wf_prog G; class G C = Some c; C \ Object \
\<Longrightarrow> is_acc_class G (pid C) (super c)"
by (auto dest: wf_prog_cdecl wf_cdecl_supD)
lemma wf_ws_prog [elim!,simp]: "wf_prog G \ ws_prog G"
apply (unfold wf_prog_def Let_def)
apply (rule ws_progI)
apply (simp_all (no_asm))
apply (auto simp add: is_acc_class_def is_acc_iface_def
dest!: wf_idecl_supD wf_cdecl_supD )+
done
lemma class_Object [simp]:
"wf_prog G \
class G Object = Some \<lparr>access=Public,cfields=[],methods=Object_mdecls,
init=Skip,super=undefined,superIfs=[]\<rparr>"
apply (unfold wf_prog_def Let_def ObjectC_def)
apply (fast dest!: map_of_SomeI)
done
lemma methd_Object[simp]: "wf_prog G \ methd G Object =
table_of (map (\<lambda>(s,m). (s, Object, m)) Object_mdecls)"
apply (subst methd_rec)
apply (auto simp add: Let_def)
done
lemma wf_prog_Object_methd:
"\wf_prog G; methd G Object sig = Some m\ \ accmodi m \ Package"
by (auto dest!: wf_prog_Object_mdecls) (auto dest!: map_of_SomeD)
lemma wf_prog_Object_is_public[intro]:
"wf_prog G \ is_public G Object"
by (auto simp add: is_public_def dest: class_Object)
lemma class_SXcpt [simp]:
"wf_prog G \
class G (SXcpt xn) = Some \<lparr>access=Public,cfields=[],methods=SXcpt_mdecls,
init=Skip,
super=if xn = Throwable then Object
else SXcpt Throwable,
superIfs=[]\<rparr>"
apply (unfold wf_prog_def Let_def SXcptC_def)
apply (fast dest!: map_of_SomeI)
done
lemma wf_ObjectC [simp]:
"wf_cdecl G ObjectC = (\is_iface G Object \ Ball (set Object_mdecls)
(wf_mdecl G Object) \<and> unique Object_mdecls)"
apply (unfold wf_cdecl_def ws_cdecl_def ObjectC_def)
apply (auto intro: da.Skip)
done
lemma Object_is_class [simp,elim!]: "wf_prog G \ is_class G Object"
apply (simp (no_asm_simp))
done
lemma Object_is_acc_class [simp,elim!]: "wf_prog G \ is_acc_class G S Object"
apply (simp (no_asm_simp) add: is_acc_class_def is_public_def
accessible_in_RefT_simp)
done
lemma SXcpt_is_class [simp,elim!]: "wf_prog G \ is_class G (SXcpt xn)"
apply (simp (no_asm_simp))
done
lemma SXcpt_is_acc_class [simp,elim!]:
"wf_prog G \ is_acc_class G S (SXcpt xn)"
apply (simp (no_asm_simp) add: is_acc_class_def is_public_def
accessible_in_RefT_simp)
done
lemma fields_Object [simp]: "wf_prog G \ DeclConcepts.fields G Object = []"
by (force intro: fields_emptyI)
lemma accfield_Object [simp]:
"wf_prog G \ accfield G S Object = Map.empty"
apply (unfold accfield_def)
apply (simp (no_asm_simp) add: Let_def)
done
lemma fields_Throwable [simp]:
"wf_prog G \ DeclConcepts.fields G (SXcpt Throwable) = []"
by (force intro: fields_emptyI)
lemma fields_SXcpt [simp]: "wf_prog G \ DeclConcepts.fields G (SXcpt xn) = []"
apply (case_tac "xn = Throwable")
apply (simp (no_asm_simp))
by (force intro: fields_emptyI)
lemmas widen_trans = ws_widen_trans [OF _ _ wf_ws_prog, elim]
lemma widen_trans2 [elim]: "\G\U\T; G\S\U; wf_prog G\ \ G\S\T"
apply (erule (2) widen_trans)
done
lemma Xcpt_subcls_Throwable [simp]:
"wf_prog G \ G\SXcpt xn\\<^sub>C SXcpt Throwable"
apply (rule SXcpt_subcls_Throwable_lemma)
apply auto
done
lemma unique_fields:
"\is_class G C; wf_prog G\ \ unique (DeclConcepts.fields G C)"
apply (erule ws_unique_fields)
apply (erule wf_ws_prog)
apply (erule (1) wf_prog_cdecl [THEN wf_cdecl_unique [THEN conjunct1]])
done
lemma fields_mono:
"\table_of (DeclConcepts.fields G C) fn = Some f; G\D\\<^sub>C C;
is_class G D; wf_prog G\<rbrakk>
\<Longrightarrow> table_of (DeclConcepts.fields G D) fn = Some f"
apply (rule map_of_SomeI)
apply (erule (1) unique_fields)
apply (erule (1) map_of_SomeD [THEN fields_mono_lemma])
apply (erule wf_ws_prog)
done
lemma fields_is_type [elim]:
"\table_of (DeclConcepts.fields G C) m = Some f; wf_prog G; is_class G C\ \
is_type G (type f)"
apply (frule wf_ws_prog)
apply (force dest: fields_declC [THEN conjunct1]
wf_prog_cdecl [THEN wf_cdecl_fdecl]
simp add: wf_fdecl_def2 is_acc_type_def)
done
lemma imethds_wf_mhead [rule_format (no_asm)]:
"\m \ imethds G I sig; wf_prog G; is_iface G I\ \
wf_mhead G (pid (decliface m)) sig (mthd m) \<and>
\<not> is_static m \<and> accmodi m = Public"
apply (frule wf_ws_prog)
apply (drule (2) imethds_declI [THEN conjunct1])
apply clarify
apply (frule_tac I="(decliface m)" in wf_prog_idecl,assumption)
apply (drule wf_idecl_mhead)
apply (erule map_of_SomeD)
apply (cases m, simp)
done
lemma methd_wf_mdecl:
"\methd G C sig = Some m; wf_prog G; class G C = Some y\ \
G\<turnstile>C\<preceq>\<^sub>C (declclass m) \<and> is_class G (declclass m) \<and>
wf_mdecl G (declclass m) (sig,(mthd m))"
apply (frule wf_ws_prog)
apply (drule (1) methd_declC)
apply fast
apply clarsimp
apply (frule (1) wf_prog_cdecl, erule wf_cdecl_mdecl, erule map_of_SomeD)
done
(*
This lemma doesn't hold!
lemma methd_rT_is_acc_type:
"\<lbrakk>wf_prog G;methd G C C sig = Some (D,m);
class G C = Some y\<rbrakk>
\<Longrightarrow> is_acc_type G (pid C) (resTy m)"
The result Type is only visible in the scope of defining class D
"is_vis_type G (pid D) (resTy m)" but not necessarily in scope of class C!
(The same is true for the type of pramaters of a method)
*)
lemma methd_rT_is_type:
"\wf_prog G;methd G C sig = Some m;
class G C = Some y\<rbrakk>
\<Longrightarrow> is_type G (resTy m)"
apply (drule (2) methd_wf_mdecl)
apply clarify
apply (drule wf_mdeclD1)
apply clarify
apply (drule rT_is_acc_type)
apply (cases m, simp add: is_acc_type_def)
done
lemma accmethd_rT_is_type:
"\wf_prog G;accmethd G S C sig = Some m;
class G C = Some y\<rbrakk>
\<Longrightarrow> is_type G (resTy m)"
by (auto simp add: accmethd_def
intro: methd_rT_is_type)
lemma methd_Object_SomeD:
"\wf_prog G;methd G Object sig = Some m\
\<Longrightarrow> declclass m = Object"
by (auto dest: class_Object simp add: methd_rec )
lemmas iface_rec_induct' = iface_rec.induct [of "%x y z. P x y"] for P
lemma wf_imethdsD:
"\im \ imethds G I sig;wf_prog G; is_iface G I\
\<Longrightarrow> \<not>is_static im \<and> accmodi im = Public"
proof -
assume asm: "wf_prog G" "is_iface G I" "im \ imethds G I sig"
have "wf_prog G \
(\<forall> i im. iface G I = Some i \<longrightarrow> im \<in> imethds G I sig
\<longrightarrow> \<not>is_static im \<and> accmodi im = Public)" (is "?P G I")
proof (induct G I rule: iface_rec_induct', intro allI impI)
fix G I i im
assume hyp: "\ i J. iface G I = Some i \ ws_prog G \ J \ set (isuperIfs i)
\<Longrightarrow> ?P G J"
assume wf: "wf_prog G" and if_I: "iface G I = Some i" and
im: "im \ imethds G I sig"
show "\is_static im \ accmodi im = Public"
proof -
let ?inherited = "Un_tables (imethds G ` set (isuperIfs i))"
let ?new = "(set_option \ table_of (map (\(s, mh). (s, I, mh)) (imethods i)))"
from if_I wf im have imethds:"im \ (?inherited \\ ?new) sig"
by (simp add: imethds_rec)
from wf if_I have
wf_supI: "\ J. J \ set (isuperIfs i) \ (\ j. iface G J = Some j)"
by (blast dest: wf_prog_idecl wf_idecl_supD is_acc_ifaceD)
from wf if_I have
"\ im \ set (imethods i). \ is_static im \ accmodi im = Public"
by (auto dest!: wf_prog_idecl wf_idecl_mhead)
then have new_ok: "\ im. table_of (imethods i) sig = Some im
\<longrightarrow> \<not> is_static im \<and> accmodi im = Public"
by (auto dest!: table_of_Some_in_set)
show ?thesis
proof (cases "?new sig = {}")
case True
from True wf wf_supI if_I imethds hyp
show ?thesis by (auto simp del: split_paired_All)
next
case False
from False wf wf_supI if_I imethds new_ok hyp
show ?thesis by (auto dest: wf_idecl_hidings hidings_entailsD)
qed
qed
qed
with asm show ?thesis by (auto simp del: split_paired_All)
qed
lemma wf_prog_hidesD:
assumes hides: "G \new hides old" and wf: "wf_prog G"
shows
"accmodi old \ accmodi new \
is_static old"
proof -
from hides
obtain c where
clsNew: "class G (declclass new) = Some c" and
neqObj: "declclass new \ Object"
by (auto dest: hidesD declared_in_classD)
with hides obtain newM oldM where
newM: "table_of (methods c) (msig new) = Some newM" and
new: "new = (declclass new,(msig new),newM)" and
old: "old = (declclass old,(msig old),oldM)" and
"msig new = msig old"
by (cases new,cases old)
(auto dest: hidesD
simp add: cdeclaredmethd_def declared_in_def)
with hides
have hides':
"G,(msig new)\(declclass new,newM) hides (declclass old,oldM)"
by auto
from clsNew wf
have "wf_cdecl G (declclass new,c)" by (blast intro: wf_prog_cdecl)
note wf_cdecl_hides_SomeD [OF this neqObj newM hides']
with new old
show ?thesis
by (cases new, cases old) auto
qed
text \<open>Compare this lemma about static
overriding \<^term>\<open>G \<turnstile>new overrides\<^sub>S old\<close> with the definition of
dynamic overriding \<^term>\<open>G \<turnstile>new overrides old\<close>.
Conforming result types and restrictions on the access modifiers of the old
and the new method are not part of the predicate for static overriding. But
they are enshured in a wellfromed program. Dynamic overriding has
no restrictions on the access modifiers but enforces confrom result types
as precondition. But with some efford we can guarantee the access modifier
restriction for dynamic overriding, too. See lemma
\<open>wf_prog_dyn_override_prop\<close>.
\<close>
lemma wf_prog_stat_overridesD:
assumes stat_override: "G \new overrides\<^sub>S old" and wf: "wf_prog G"
shows
"G\resTy new\resTy old \
accmodi old \<le> accmodi new \<and>
\<not> is_static old"
proof -
from stat_override
obtain c where
clsNew: "class G (declclass new) = Some c" and
neqObj: "declclass new \ Object"
by (auto dest: stat_overrides_commonD declared_in_classD)
with stat_override obtain newM oldM where
newM: "table_of (methods c) (msig new) = Some newM" and
new: "new = (declclass new,(msig new),newM)" and
old: "old = (declclass old,(msig old),oldM)" and
"msig new = msig old"
by (cases new,cases old)
(auto dest: stat_overrides_commonD
simp add: cdeclaredmethd_def declared_in_def)
with stat_override
have stat_override':
"G,(msig new)\(declclass new,newM) overrides\<^sub>S (declclass old,oldM)"
by auto
from clsNew wf
have "wf_cdecl G (declclass new,c)" by (blast intro: wf_prog_cdecl)
note wf_cdecl_overrides_SomeD [OF this neqObj newM stat_override']
with new old
show ?thesis
by (cases new, cases old) auto
qed
lemma static_to_dynamic_overriding:
assumes stat_override: "G\new overrides\<^sub>S old" and wf : "wf_prog G"
shows "G\new overrides old"
proof -
from stat_override
show ?thesis (is "?Overrides new old")
proof (induct)
case (Direct new old superNew)
then have stat_override:"G\new overrides\<^sub>S old"
by (rule stat_overridesR.Direct)
from stat_override wf
have resTy_widen: "G\resTy new\resTy old" and
not_static_old: "\ is_static old"
by (auto dest: wf_prog_stat_overridesD)
have not_private_new: "accmodi new \ Private"
proof -
from stat_override
have "accmodi old \ Private"
by (rule no_Private_stat_override)
moreover
from stat_override wf
have "accmodi old \ accmodi new"
by (auto dest: wf_prog_stat_overridesD)
ultimately
show ?thesis
by (auto dest: acc_modi_bottom)
qed
with Direct resTy_widen not_static_old
show "?Overrides new old"
by (auto intro: overridesR.Direct stat_override_declclasses_relation)
next
case (Indirect new inter old)
then show "?Overrides new old"
by (blast intro: overridesR.Indirect)
qed
qed
lemma non_Package_instance_method_inheritance:
assumes old_inheritable: "G\Method old inheritable_in (pid C)" and
accmodi_old: "accmodi old \ Package" and
instance_method: "\ is_static old" and
subcls: "G\C \\<^sub>C declclass old" and
old_declared: "G\Method old declared_in (declclass old)" and
wf: "wf_prog G"
shows "G\Method old member_of C \
(\<exists> new. G\<turnstile> new overrides\<^sub>S old \<and> G\<turnstile>Method new member_of C)"
proof -
from wf have ws: "ws_prog G" by auto
from old_declared have iscls_declC_old: "is_class G (declclass old)"
by (auto simp add: declared_in_def cdeclaredmethd_def)
from subcls have iscls_C: "is_class G C"
by (blast dest: subcls_is_class)
from iscls_C ws old_inheritable subcls
show ?thesis (is "?P C old")
proof (induct rule: ws_class_induct')
case Object
assume "G\Object\\<^sub>C declclass old"
then show "?P Object old"
by blast
next
case (Subcls C c)
assume cls_C: "class G C = Some c" and
neq_C_Obj: "C \ Object" and
hyp: "\G \Method old inheritable_in pid (super c);
G\<turnstile>super c\<prec>\<^sub>C declclass old\<rbrakk> \<Longrightarrow> ?P (super c) old" and
inheritable: "G \Method old inheritable_in pid C" and
subclsC: "G\C\\<^sub>C declclass old"
from cls_C neq_C_Obj
have super: "G\C \\<^sub>C1 super c"
by (rule subcls1I)
from wf cls_C neq_C_Obj
have accessible_super: "G\(Class (super c)) accessible_in (pid C)"
by (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD)
{
fix old
assume member_super: "G\Method old member_of (super c)"
assume inheritable: "G \Method old inheritable_in pid C"
assume instance_method: "\ is_static old"
from member_super
have old_declared: "G\Method old declared_in (declclass old)"
by (cases old) (auto dest: member_of_declC)
have "?P C old"
proof (cases "G\mid (msig old) undeclared_in C")
case True
with inheritable super accessible_super member_super
have "G\Method old member_of C"
by (cases old) (auto intro: members.Inherited)
then show ?thesis
by auto
next
case False
then obtain new_member where
"G\new_member declared_in C" and
"mid (msig old) = memberid new_member"
by (auto dest: not_undeclared_declared)
then obtain new where
new: "G\Method new declared_in C" and
eq_sig: "msig old = msig new" and
declC_new: "declclass new = C"
by (cases new_member) auto
then have member_new: "G\Method new member_of C"
by (cases new) (auto intro: members.Immediate)
from declC_new super member_super
have subcls_new_old: "G\declclass new \\<^sub>C declclass old"
by (auto dest!: member_of_subclseq_declC
dest: r_into_trancl intro: trancl_rtrancl_trancl)
show ?thesis
proof (cases "is_static new")
case False
with eq_sig declC_new new old_declared inheritable
super member_super subcls_new_old
have "G\new overrides\<^sub>S old"
by (auto intro!: stat_overridesR.Direct)
with member_new show ?thesis
by blast
next
case True
with eq_sig declC_new subcls_new_old new old_declared inheritable
have "G\new hides old"
by (auto intro: hidesI)
with wf
have "is_static old"
by (blast dest: wf_prog_hidesD)
with instance_method
show ?thesis
by (contradiction)
qed
qed
} note hyp_member_super = this
from subclsC cls_C
have "G\(super c)\\<^sub>C declclass old"
by (rule subcls_superD)
then
show "?P C old"
proof (cases rule: subclseq_cases)
case Eq
assume "super c = declclass old"
with old_declared
have "G\Method old member_of (super c)"
by (cases old) (auto intro: members.Immediate)
with inheritable instance_method
show ?thesis
by (blast dest: hyp_member_super)
next
case Subcls
assume "G\super c\\<^sub>C declclass old"
moreover
from inheritable accmodi_old
have "G \Method old inheritable_in pid (super c)"
by (cases "accmodi old") (auto simp add: inheritable_in_def)
ultimately
have "?P (super c) old"
by (blast dest: hyp)
then show ?thesis
proof
assume "G \Method old member_of super c"
with inheritable instance_method
show ?thesis
by (blast dest: hyp_member_super)
next
assume "\new. G \ new overrides\<^sub>S old \ G \Method new member_of super c"
then obtain super_new where
super_new_override: "G \ super_new overrides\<^sub>S old" and
super_new_member: "G \Method super_new member_of super c"
by blast
from super_new_override wf
have "accmodi old \ accmodi super_new"
by (auto dest: wf_prog_stat_overridesD)
with inheritable accmodi_old
have "G \Method super_new inheritable_in pid C"
by (auto simp add: inheritable_in_def
split: acc_modi.splits
dest: acc_modi_le_Dests)
moreover
from super_new_override
have "\ is_static super_new"
by (auto dest: stat_overrides_commonD)
moreover
note super_new_member
ultimately have "?P C super_new"
by (auto dest: hyp_member_super)
then show ?thesis
proof
assume "G \Method super_new member_of C"
with super_new_override
show ?thesis
by blast
next
assume "\new. G \ new overrides\<^sub>S super_new \
G \<turnstile>Method new member_of C"
with super_new_override show ?thesis
by (blast intro: stat_overridesR.Indirect)
qed
qed
qed
qed
qed
lemma non_Package_instance_method_inheritance_cases:
assumes old_inheritable: "G\Method old inheritable_in (pid C)" and
accmodi_old: "accmodi old \ Package" and
instance_method: "\ is_static old" and
subcls: "G\C \\<^sub>C declclass old" and
old_declared: "G\Method old declared_in (declclass old)" and
wf: "wf_prog G"
obtains (Inheritance) "G\Method old member_of C"
| (Overriding) new where "G\ new overrides\<^sub>S old" and "G\Method new member_of C"
proof -
from old_inheritable accmodi_old instance_method subcls old_declared wf
Inheritance Overriding
show thesis
by (auto dest: non_Package_instance_method_inheritance)
qed
lemma dynamic_to_static_overriding:
assumes dyn_override: "G\ new overrides old" and
accmodi_old: "accmodi old \ Package" and
wf: "wf_prog G"
shows "G\ new overrides\<^sub>S old"
proof -
from dyn_override accmodi_old
show ?thesis (is "?Overrides new old")
proof (induct rule: overridesR.induct)
case (Direct new old)
assume new_declared: "G\Method new declared_in declclass new"
assume eq_sig_new_old: "msig new = msig old"
assume subcls_new_old: "G\declclass new \\<^sub>C declclass old"
assume "G \Method old inheritable_in pid (declclass new)" and
"accmodi old \ Package" and
"\ is_static old" and
"G\declclass new\\<^sub>C declclass old" and
"G\Method old declared_in declclass old"
from this wf
show "?Overrides new old"
proof (cases rule: non_Package_instance_method_inheritance_cases)
case Inheritance
assume "G \Method old member_of declclass new"
then have "G\mid (msig old) undeclared_in declclass new"
proof cases
case Immediate
with subcls_new_old wf show ?thesis
by (auto dest: subcls_irrefl)
next
case Inherited
then show ?thesis
by (cases old) auto
qed
with eq_sig_new_old new_declared
show ?thesis
by (cases old,cases new) (auto dest!: declared_not_undeclared)
next
case (Overriding new')
assume stat_override_new': "G \ new' overrides\<^sub>S old"
then have "msig new' = msig old"
by (auto dest: stat_overrides_commonD)
with eq_sig_new_old have eq_sig_new_new': "msig new=msig new'"
by simp
assume "G \Method new' member_of declclass new"
then show ?thesis
proof (cases)
case Immediate
then have declC_new: "declclass new' = declclass new"
by auto
from Immediate
have "G\Method new' declared_in declclass new"
by (cases new') auto
with new_declared eq_sig_new_new' declC_new
have "new=new'"
by (cases new, cases new') (auto dest: unique_declared_in)
with stat_override_new'
show ?thesis
by simp
next
case Inherited
then have "G\mid (msig new') undeclared_in declclass new"
by (cases new') (auto)
with eq_sig_new_new' new_declared
show ?thesis
by (cases new,cases new') (auto dest!: declared_not_undeclared)
qed
qed
next
case (Indirect new inter old)
assume accmodi_old: "accmodi old \ Package"
assume "accmodi old \ Package \ G \ inter overrides\<^sub>S old"
with accmodi_old
have stat_override_inter_old: "G \ inter overrides\<^sub>S old"
by blast
moreover
assume hyp_inter: "accmodi inter \ Package \ G \ new overrides\<^sub>S inter"
moreover
have "accmodi inter \ Package"
proof -
from stat_override_inter_old wf
have "accmodi old \ accmodi inter"
by (auto dest: wf_prog_stat_overridesD)
with stat_override_inter_old accmodi_old
show ?thesis
by (auto dest!: no_Private_stat_override
split: acc_modi.splits
dest: acc_modi_le_Dests)
qed
ultimately show "?Overrides new old"
by (blast intro: stat_overridesR.Indirect)
qed
qed
lemma wf_prog_dyn_override_prop:
assumes dyn_override: "G \ new overrides old" and
wf: "wf_prog G"
shows "accmodi old \ accmodi new"
proof (cases "accmodi old = Package")
case True
note old_Package = this
show ?thesis
proof (cases "accmodi old \ accmodi new")
case True then show ?thesis .
next
case False
with old_Package
have "accmodi new = Private"
by (cases "accmodi new") (auto simp add: le_acc_def less_acc_def)
with dyn_override
show ?thesis
by (auto dest: overrides_commonD)
qed
next
case False
with dyn_override wf
have "G \ new overrides\<^sub>S old"
by (blast intro: dynamic_to_static_overriding)
with wf
show ?thesis
by (blast dest: wf_prog_stat_overridesD)
qed
lemma overrides_Package_old:
assumes dyn_override: "G \ new overrides old" and
accmodi_new: "accmodi new = Package" and
wf: "wf_prog G "
shows "accmodi old = Package"
proof (cases "accmodi old")
case Private
with dyn_override show ?thesis
by (simp add: no_Private_override)
next
case Package
then show ?thesis .
next
case Protected
with dyn_override wf
have "G \ new overrides\<^sub>S old"
by (auto intro: dynamic_to_static_overriding)
with wf
have "accmodi old \ accmodi new"
by (auto dest: wf_prog_stat_overridesD)
with Protected accmodi_new
show ?thesis
by (simp add: less_acc_def le_acc_def)
next
case Public
with dyn_override wf
have "G \ new overrides\<^sub>S old"
by (auto intro: dynamic_to_static_overriding)
with wf
have "accmodi old \ accmodi new"
by (auto dest: wf_prog_stat_overridesD)
with Public accmodi_new
show ?thesis
by (simp add: less_acc_def le_acc_def)
qed
lemma dyn_override_Package:
assumes dyn_override: "G \ new overrides old" and
accmodi_old: "accmodi old = Package" and
accmodi_new: "accmodi new = Package" and
wf: "wf_prog G"
shows "pid (declclass old) = pid (declclass new)"
proof -
from dyn_override accmodi_old accmodi_new
show ?thesis (is "?EqPid old new")
proof (induct rule: overridesR.induct)
case (Direct new old)
assume "accmodi old = Package"
"G \Method old inheritable_in pid (declclass new)"
then show "pid (declclass old) = pid (declclass new)"
by (auto simp add: inheritable_in_def)
next
case (Indirect new inter old)
assume accmodi_old: "accmodi old = Package" and
accmodi_new: "accmodi new = Package"
assume "G \ new overrides inter"
with accmodi_new wf
have "accmodi inter = Package"
by (auto intro: overrides_Package_old)
with Indirect
show "pid (declclass old) = pid (declclass new)"
by auto
qed
qed
lemma dyn_override_Package_escape:
assumes dyn_override: "G \ new overrides old" and
accmodi_old: "accmodi old = Package" and
outside_pack: "pid (declclass old) \ pid (declclass new)" and
wf: "wf_prog G"
shows "\ inter. G \ new overrides inter \ G \ inter overrides old \
pid (declclass old) = pid (declclass inter) \<and>
Protected \<le> accmodi inter"
proof -
from dyn_override accmodi_old outside_pack
show ?thesis (is "?P new old")
proof (induct rule: overridesR.induct)
case (Direct new old)
assume accmodi_old: "accmodi old = Package"
assume outside_pack: "pid (declclass old) \ pid (declclass new)"
assume "G \Method old inheritable_in pid (declclass new)"
with accmodi_old
have "pid (declclass old) = pid (declclass new)"
by (simp add: inheritable_in_def)
with outside_pack
show "?P new old"
by (contradiction)
next
case (Indirect new inter old)
assume accmodi_old: "accmodi old = Package"
assume outside_pack: "pid (declclass old) \ pid (declclass new)"
assume override_new_inter: "G \ new overrides inter"
assume override_inter_old: "G \ inter overrides old"
assume hyp_new_inter: "\accmodi inter = Package;
pid (declclass inter) \<noteq> pid (declclass new)\<rbrakk>
\<Longrightarrow> ?P new inter"
assume hyp_inter_old: "\accmodi old = Package;
pid (declclass old) \<noteq> pid (declclass inter)\<rbrakk>
\<Longrightarrow> ?P inter old"
show "?P new old"
proof (cases "pid (declclass old) = pid (declclass inter)")
case True
note same_pack_old_inter = this
show ?thesis
proof (cases "pid (declclass inter) = pid (declclass new)")
case True
with same_pack_old_inter outside_pack
show ?thesis
by auto
next
case False
note diff_pack_inter_new = this
show ?thesis
proof (cases "accmodi inter = Package")
case True
with diff_pack_inter_new hyp_new_inter
obtain newinter where
over_new_newinter: "G \ new overrides newinter" and
over_newinter_inter: "G \ newinter overrides inter" and
eq_pid: "pid (declclass inter) = pid (declclass newinter)" and
accmodi_newinter: "Protected \ accmodi newinter"
by auto
from over_newinter_inter override_inter_old
have "G\newinter overrides old"
by (rule overridesR.Indirect)
moreover
from eq_pid same_pack_old_inter
have "pid (declclass old) = pid (declclass newinter)"
by simp
moreover
note over_new_newinter accmodi_newinter
ultimately show ?thesis
by blast
next
case False
with override_new_inter
have "Protected \ accmodi inter"
by (cases "accmodi inter") (auto dest: no_Private_override)
with override_new_inter override_inter_old same_pack_old_inter
show ?thesis
by blast
qed
qed
next
case False
with accmodi_old hyp_inter_old
obtain newinter where
over_inter_newinter: "G \ inter overrides newinter" and
over_newinter_old: "G \ newinter overrides old" and
eq_pid: "pid (declclass old) = pid (declclass newinter)" and
accmodi_newinter: "Protected \ accmodi newinter"
by auto
from override_new_inter over_inter_newinter
have "G \ new overrides newinter"
by (rule overridesR.Indirect)
with eq_pid over_newinter_old accmodi_newinter
show ?thesis
by blast
qed
qed
qed
lemmas class_rec_induct' = class_rec.induct [of "%x y z w. P x y"] for P
lemma declclass_widen[rule_format]:
"wf_prog G
\<longrightarrow> (\<forall>c m. class G C = Some c \<longrightarrow> methd G C sig = Some m
\<longrightarrow> G\<turnstile>C \<preceq>\<^sub>C declclass m)" (is "?P G C")
proof (induct G C rule: class_rec_induct', intro allI impI)
fix G C c m
assume Hyp: "\c. class G C = Some c \ ws_prog G \ C \ Object
\<Longrightarrow> ?P G (super c)"
assume wf: "wf_prog G" and cls_C: "class G C = Some c" and
m: "methd G C sig = Some m"
show "G\C\\<^sub>C declclass m"
proof (cases "C=Object")
case True
with wf m show ?thesis by (simp add: methd_Object_SomeD)
next
let ?filter="filter_tab (\sig m. G\C inherits method sig m)"
let ?table = "table_of (map (\(s, m). (s, C, m)) (methods c))"
case False
with cls_C wf m
have methd_C: "(?filter (methd G (super c)) ++ ?table) sig = Some m "
by (simp add: methd_rec)
show ?thesis
proof (cases "?table sig")
case None
from this methd_C have "?filter (methd G (super c)) sig = Some m"
by simp
moreover
from wf cls_C False obtain sup where "class G (super c) = Some sup"
by (blast dest: wf_prog_cdecl wf_cdecl_supD is_acc_class_is_class)
moreover note wf False cls_C
ultimately have "G\super c \\<^sub>C declclass m"
by (auto intro: Hyp [rule_format])
moreover from cls_C False have "G\C \\<^sub>C1 super c" by (rule subcls1I)
ultimately show ?thesis by - (rule rtrancl_into_rtrancl2)
next
case Some
from this wf False cls_C methd_C show ?thesis by auto
qed
qed
qed
lemma declclass_methd_Object:
"\wf_prog G; methd G Object sig = Some m\ \ declclass m = Object"
by auto
lemma methd_declaredD:
"\wf_prog G; is_class G C;methd G C sig = Some m\
\<Longrightarrow> G\<turnstile>(mdecl (sig,mthd m)) declared_in (declclass m)"
proof -
assume wf: "wf_prog G"
then have ws: "ws_prog G" ..
assume clsC: "is_class G C"
from clsC ws
show "methd G C sig = Some m
\<Longrightarrow> G\<turnstile>(mdecl (sig,mthd m)) declared_in (declclass m)"
proof (induct C rule: ws_class_induct')
case Object
assume "methd G Object sig = Some m"
with wf show ?thesis
by - (rule method_declared_inI, auto)
next
case Subcls
fix C c
assume clsC: "class G C = Some c"
and m: "methd G C sig = Some m"
and hyp: "methd G (super c) sig = Some m \ ?thesis"
let ?newMethods = "table_of (map (\(s, m). (s, C, m)) (methods c))"
show ?thesis
proof (cases "?newMethods sig")
case None
from None ws clsC m hyp
show ?thesis by (auto intro: method_declared_inI simp add: methd_rec)
next
case Some
from Some ws clsC m
show ?thesis by (auto intro: method_declared_inI simp add: methd_rec)
qed
qed
qed
lemma methd_rec_Some_cases:
assumes methd_C: "methd G C sig = Some m" and
ws: "ws_prog G" and
clsC: "class G C = Some c" and
neq_C_Obj: "C\Object"
obtains (NewMethod) "table_of (map (\(s, m). (s, C, m)) (methods c)) sig = Some m"
| (InheritedMethod) "G\C inherits (method sig m)" and "methd G (super c) sig = Some m"
proof -
let ?inherited = "filter_tab (\sig m. G\C inherits method sig m)
(methd G (super c))"
let ?new = "table_of (map (\(s, m). (s, C, m)) (methods c))"
from ws clsC neq_C_Obj methd_C
have methd_unfold: "(?inherited ++ ?new) sig = Some m"
by (simp add: methd_rec)
show thesis
proof (cases "?new sig")
case None
with methd_unfold have "?inherited sig = Some m"
by (auto)
with InheritedMethod show ?thesis by blast
next
case Some
with methd_unfold have "?new sig = Some m"
by auto
with NewMethod show ?thesis by blast
qed
qed
lemma methd_member_of:
assumes wf: "wf_prog G"
shows
"\is_class G C; methd G C sig = Some m\ \ G\Methd sig m member_of C"
(is "?Class C \ ?Method C \ ?MemberOf C")
proof -
from wf have ws: "ws_prog G" ..
assume defC: "is_class G C"
from defC ws
show "?Class C \ ?Method C \ ?MemberOf C"
proof (induct rule: ws_class_induct')
case Object
with wf have declC: "Object = declclass m"
by (simp add: declclass_methd_Object)
from Object wf have "G\Methd sig m declared_in Object"
by (auto intro: methd_declaredD simp add: declC)
with declC
show "?MemberOf Object"
by (auto intro!: members.Immediate
simp del: methd_Object)
next
case (Subcls C c)
assume clsC: "class G C = Some c" and
neq_C_Obj: "C \ Object"
assume methd: "?Method C"
from methd ws clsC neq_C_Obj
show "?MemberOf C"
proof (cases rule: methd_rec_Some_cases)
case NewMethod
with clsC show ?thesis
by (auto dest: method_declared_inI intro!: members.Immediate)
next
case InheritedMethod
then show "?thesis"
by (blast dest: inherits_member)
qed
qed
qed
lemma current_methd:
"\table_of (methods c) sig = Some new;
ws_prog G; class G C = Some c; C \<noteq> Object;
methd G (super c) sig = Some old\<rbrakk>
\<Longrightarrow> methd G C sig = Some (C,new)"
by (auto simp add: methd_rec
intro: filter_tab_SomeI map_add_find_right table_of_map_SomeI)
lemma wf_prog_staticD:
assumes wf: "wf_prog G" and
clsC: "class G C = Some c" and
neq_C_Obj: "C \ Object" and
old: "methd G (super c) sig = Some old" and
accmodi_old: "Protected \ accmodi old" and
new: "table_of (methods c) sig = Some new"
shows "is_static new = is_static old"
proof -
from clsC wf
have wf_cdecl: "wf_cdecl G (C,c)" by (rule wf_prog_cdecl)
from wf clsC neq_C_Obj
have is_cls_super: "is_class G (super c)"
by (blast dest: wf_prog_acc_superD is_acc_classD)
from wf is_cls_super old
have old_member_of: "G\Methd sig old member_of (super c)"
by (rule methd_member_of)
from old wf is_cls_super
have old_declared: "G\Methd sig old declared_in (declclass old)"
by (auto dest: methd_declared_in_declclass)
from new clsC
have new_declared: "G\Methd sig (C,new) declared_in C"
by (auto intro: method_declared_inI)
note trancl_rtrancl_tranc = trancl_rtrancl_trancl [trans] (* ### in Basis *)
from clsC neq_C_Obj
have subcls1_C_super: "G\C \\<^sub>C1 super c"
by (rule subcls1I)
then have "G\C \\<^sub>C super c" ..
also from old wf is_cls_super
have "G\super c \\<^sub>C (declclass old)" by (auto dest: methd_declC)
finally have subcls_C_old: "G\C \\<^sub>C (declclass old)" .
from accmodi_old
have inheritable: "G\Methd sig old inheritable_in pid C"
by (auto simp add: inheritable_in_def
dest: acc_modi_le_Dests)
show ?thesis
proof (cases "is_static new")
case True
with subcls_C_old new_declared old_declared inheritable
have "G,sig\(C,new) hides old"
by (auto intro: hidesI)
with True wf_cdecl neq_C_Obj new
show ?thesis
by (auto dest: wf_cdecl_hides_SomeD)
next
case False
with subcls_C_old new_declared old_declared inheritable subcls1_C_super
old_member_of
have "G,sig\(C,new) overrides\<^sub>S old"
by (auto intro: stat_overridesR.Direct)
with False wf_cdecl neq_C_Obj new
show ?thesis
by (auto dest: wf_cdecl_overrides_SomeD)
qed
qed
lemma inheritable_instance_methd:
assumes subclseq_C_D: "G\C \\<^sub>C D" and
is_cls_D: "is_class G D" and
wf: "wf_prog G" and
old: "methd G D sig = Some old" and
accmodi_old: "Protected \ accmodi old" and
not_static_old: "\ is_static old"
shows
"\new. methd G C sig = Some new \
(new = old \<or> G,sig\<turnstile>new overrides\<^sub>S old)"
(is "(\new. (?Constraint C new old))")
proof -
from subclseq_C_D is_cls_D
have is_cls_C: "is_class G C" by (rule subcls_is_class2)
from wf
have ws: "ws_prog G" ..
from is_cls_C ws subclseq_C_D
show "\new. ?Constraint C new old"
proof (induct rule: ws_class_induct')
case (Object co)
then have eq_D_Obj: "D=Object" by auto
with old
have "?Constraint Object old old"
by auto
with eq_D_Obj
show "\ new. ?Constraint Object new old" by auto
next
case (Subcls C c)
assume hyp: "G\super c\\<^sub>C D \ \new. ?Constraint (super c) new old"
assume clsC: "class G C = Some c"
assume neq_C_Obj: "C\Object"
from clsC wf
have wf_cdecl: "wf_cdecl G (C,c)"
by (rule wf_prog_cdecl)
from ws clsC neq_C_Obj
have is_cls_super: "is_class G (super c)"
by (auto dest: ws_prog_cdeclD)
from clsC wf neq_C_Obj
have superAccessible: "G\(Class (super c)) accessible_in (pid C)" and
subcls1_C_super: "G\C \\<^sub>C1 super c"
by (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD
intro: subcls1I)
show "\new. ?Constraint C new old"
proof (cases "G\super c\\<^sub>C D")
case False
from False Subcls
have eq_C_D: "C=D"
by (auto dest: subclseq_superD)
with old
have "?Constraint C old old"
by auto
with eq_C_D
show "\ new. ?Constraint C new old" by auto
next
case True
with hyp obtain super_method
where super: "?Constraint (super c) super_method old" by blast
from super not_static_old
have not_static_super: "\ is_static super_method"
by (auto dest!: stat_overrides_commonD)
from super old wf accmodi_old
have accmodi_super_method: "Protected \ accmodi super_method"
by (auto dest!: wf_prog_stat_overridesD)
from super accmodi_old wf
have inheritable: "G\Methd sig super_method inheritable_in (pid C)"
by (auto dest!: wf_prog_stat_overridesD
acc_modi_le_Dests
simp add: inheritable_in_def)
from super wf is_cls_super
--> --------------------
--> maximum size reached
--> --------------------
¤ Dauer der Verarbeitung: 0.158 Sekunden
(vorverarbeitet)
¤
|
Haftungshinweis
Die Informationen auf dieser Webseite wurden
nach bestem Wissen sorgfältig zusammengestellt. Es wird jedoch weder Vollständigkeit, noch Richtigkeit,
noch Qualität der bereit gestellten Informationen zugesichert.
Bemerkung:
Die farbliche Syntaxdarstellung ist noch experimentell.
|