src/HOL/NanoJava/TypeRel.thy
changeset 63167 0909deb8059b
parent 62390 842917225d56
child 67443 3abf6a722518
equal deleted inserted replaced
63166:143f58bb34f9 63167:0909deb8059b
     6 
     6 
     7 theory TypeRel
     7 theory TypeRel
     8 imports Decl
     8 imports Decl
     9 begin
     9 begin
    10 
    10 
    11 text{* Direct subclass relation *}
    11 text\<open>Direct subclass relation\<close>
    12 
    12 
    13 definition subcls1 :: "(cname \<times> cname) set"
    13 definition subcls1 :: "(cname \<times> cname) set"
    14 where
    14 where
    15   "subcls1 \<equiv> {(C,D). C\<noteq>Object \<and> (\<exists>c. class C = Some c \<and> super c=D)}"
    15   "subcls1 \<equiv> {(C,D). C\<noteq>Object \<and> (\<exists>c. class C = Some c \<and> super c=D)}"
    16 
    16 
    22   where "C \<preceq>C D == (C,D) \<in> subcls1^*"
    22   where "C \<preceq>C D == (C,D) \<in> subcls1^*"
    23 
    23 
    24 
    24 
    25 subsection "Declarations and properties not used in the meta theory"
    25 subsection "Declarations and properties not used in the meta theory"
    26 
    26 
    27 text{* Widening, viz. method invocation conversion *}
    27 text\<open>Widening, viz. method invocation conversion\<close>
    28 inductive
    28 inductive
    29   widen :: "ty => ty => bool"  ("_ \<preceq> _" [71,71] 70)
    29   widen :: "ty => ty => bool"  ("_ \<preceq> _" [71,71] 70)
    30 where
    30 where
    31   refl [intro!, simp]: "T \<preceq> T"
    31   refl [intro!, simp]: "T \<preceq> T"
    32 | subcls: "C\<preceq>C D \<Longrightarrow> Class C \<preceq> Class D"
    32 | subcls: "C\<preceq>C D \<Longrightarrow> Class C \<preceq> Class D"
   101 apply (drule wf_subcls1)
   101 apply (drule wf_subcls1)
   102 apply (subst def_wfrec[OF class_rec_def], auto)
   102 apply (subst def_wfrec[OF class_rec_def], auto)
   103 apply (subst cut_apply, auto intro: subcls1I)
   103 apply (subst cut_apply, auto intro: subcls1I)
   104 done
   104 done
   105 
   105 
   106 --{* Methods of a class, with inheritance and hiding *}
   106 \<comment>\<open>Methods of a class, with inheritance and hiding\<close>
   107 definition "method" :: "cname => (mname \<rightharpoonup> methd)" where
   107 definition "method" :: "cname => (mname \<rightharpoonup> methd)" where
   108   "method C \<equiv> class_rec C methods"
   108   "method C \<equiv> class_rec C methods"
   109 
   109 
   110 lemma method_rec: "\<lbrakk>class C = Some m; ws_prog\<rbrakk> \<Longrightarrow>
   110 lemma method_rec: "\<lbrakk>class C = Some m; ws_prog\<rbrakk> \<Longrightarrow>
   111 method C = (if C=Object then empty else method (super m)) ++ map_of (methods m)"
   111 method C = (if C=Object then empty else method (super m)) ++ map_of (methods m)"
   113 apply (erule (1) class_rec [THEN trans])
   113 apply (erule (1) class_rec [THEN trans])
   114 apply simp
   114 apply simp
   115 done
   115 done
   116 
   116 
   117 
   117 
   118 --{* Fields of a class, with inheritance and hiding *}
   118 \<comment>\<open>Fields of a class, with inheritance and hiding\<close>
   119 definition field  :: "cname => (fname \<rightharpoonup> ty)" where
   119 definition field  :: "cname => (fname \<rightharpoonup> ty)" where
   120   "field C \<equiv> class_rec C flds"
   120   "field C \<equiv> class_rec C flds"
   121 
   121 
   122 lemma flds_rec: "\<lbrakk>class C = Some m; ws_prog\<rbrakk> \<Longrightarrow>
   122 lemma flds_rec: "\<lbrakk>class C = Some m; ws_prog\<rbrakk> \<Longrightarrow>
   123 field C = (if C=Object then empty else field (super m)) ++ map_of (flds m)"
   123 field C = (if C=Object then empty else field (super m)) ++ map_of (flds m)"