src/HOL/Bali/TypeSafe.thy
changeset 12854 00d4a435777f
child 12857 a4386cc9b1c3
equal deleted inserted replaced
12853:de505273c971 12854:00d4a435777f
       
     1 (*  Title:      isabelle/Bali/TypeSafe.thy
       
     2     ID:         $Id$
       
     3     Author:     David von Oheimb
       
     4     Copyright   1997 Technische Universitaet Muenchen
       
     5 *)
       
     6 header {* The type soundness proof for Java *}
       
     7 
       
     8 
       
     9 theory TypeSafe = Eval + WellForm + Conform:
       
    10 
       
    11 section "result conformance"
       
    12 
       
    13 constdefs
       
    14   assign_conforms :: "st \<Rightarrow> (val \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> ty \<Rightarrow> env_ \<Rightarrow> bool"
       
    15           ("_\<le>|_\<preceq>_\<Colon>\<preceq>_"                                        [71,71,71,71] 70)
       
    16  "s\<le>|f\<preceq>T\<Colon>\<preceq>E \<equiv>
       
    17   \<forall>s' w. Norm s'\<Colon>\<preceq>E \<longrightarrow> fst E,s'\<turnstile>w\<Colon>\<preceq>T \<longrightarrow> s\<le>|s' \<longrightarrow> assign f w (Norm s')\<Colon>\<preceq>E"
       
    18 
       
    19   rconf :: "prog \<Rightarrow> lenv \<Rightarrow> st \<Rightarrow> term \<Rightarrow> vals \<Rightarrow> tys \<Rightarrow> bool"
       
    20           ("_,_,_\<turnstile>_\<succ>_\<Colon>\<preceq>_"                               [71,71,71,71,71,71] 70)
       
    21   "G,L,s\<turnstile>t\<succ>v\<Colon>\<preceq>T 
       
    22     \<equiv> case T of
       
    23         Inl T  \<Rightarrow> if (\<exists>vf. t=In2 vf)
       
    24                   then G,s\<turnstile>fst (the_In2 v)\<Colon>\<preceq>T \<and> s\<le>|snd (the_In2 v)\<preceq>T\<Colon>\<preceq>(G,L)
       
    25                   else G,s\<turnstile>the_In1 v\<Colon>\<preceq>T
       
    26       | Inr Ts \<Rightarrow> list_all2 (conf G s) (the_In3 v) Ts"
       
    27 
       
    28 lemma rconf_In1 [simp]: 
       
    29  "G,L,s\<turnstile>In1 ec\<succ>In1 v \<Colon>\<preceq>Inl T  =  G,s\<turnstile>v\<Colon>\<preceq>T"
       
    30 apply (unfold rconf_def)
       
    31 apply (simp (no_asm))
       
    32 done
       
    33 
       
    34 lemma rconf_In2 [simp]: 
       
    35  "G,L,s\<turnstile>In2 va\<succ>In2 vf\<Colon>\<preceq>Inl T  = (G,s\<turnstile>fst vf\<Colon>\<preceq>T \<and> s\<le>|snd vf\<preceq>T\<Colon>\<preceq>(G,L))"
       
    36 apply (unfold rconf_def)
       
    37 apply (simp (no_asm))
       
    38 done
       
    39 
       
    40 lemma rconf_In3 [simp]: 
       
    41  "G,L,s\<turnstile>In3 es\<succ>In3 vs\<Colon>\<preceq>Inr Ts = list_all2 (\<lambda>v T. G,s\<turnstile>v\<Colon>\<preceq>T) vs Ts"
       
    42 apply (unfold rconf_def)
       
    43 apply (simp (no_asm))
       
    44 done
       
    45 
       
    46 section "fits and conf"
       
    47 
       
    48 (* unused *)
       
    49 lemma conf_fits: "G,s\<turnstile>v\<Colon>\<preceq>T \<Longrightarrow> G,s\<turnstile>v fits T"
       
    50 apply (unfold fits_def)
       
    51 apply clarify
       
    52 apply (erule swap, simp (no_asm_use))
       
    53 apply (drule conf_RefTD)
       
    54 apply auto
       
    55 done
       
    56 
       
    57 lemma fits_conf: 
       
    58   "\<lbrakk>G,s\<turnstile>v\<Colon>\<preceq>T; G\<turnstile>T\<preceq>? T'; G,s\<turnstile>v fits T'; ws_prog G\<rbrakk> \<Longrightarrow> G,s\<turnstile>v\<Colon>\<preceq>T'"
       
    59 apply (auto dest!: fitsD cast_PrimT2 cast_RefT2)
       
    60 apply (force dest: conf_RefTD intro: conf_AddrI)
       
    61 done
       
    62 
       
    63 lemma fits_Array: 
       
    64  "\<lbrakk>G,s\<turnstile>v\<Colon>\<preceq>T; G\<turnstile>T'.[]\<preceq>T.[]; G,s\<turnstile>v fits T'; ws_prog G\<rbrakk> \<Longrightarrow> G,s\<turnstile>v\<Colon>\<preceq>T'"
       
    65 apply (auto dest!: fitsD widen_ArrayPrimT widen_ArrayRefT)
       
    66 apply (force dest: conf_RefTD intro: conf_AddrI)
       
    67 done
       
    68 
       
    69 
       
    70 
       
    71 section "gext"
       
    72 
       
    73 lemma halloc_gext: "\<And>s1 s2. G\<turnstile>s1 \<midarrow>halloc oi\<succ>a\<rightarrow> s2 \<Longrightarrow> snd s1\<le>|snd s2"
       
    74 apply (simp (no_asm_simp) only: split_tupled_all)
       
    75 apply (erule halloc.induct)
       
    76 apply  (auto dest!: new_AddrD)
       
    77 done
       
    78 
       
    79 lemma sxalloc_gext: "\<And>s1 s2. G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2 \<Longrightarrow> snd s1\<le>|snd s2"
       
    80 apply (simp (no_asm_simp) only: split_tupled_all)
       
    81 apply (erule sxalloc.induct)
       
    82 apply   (auto dest!: halloc_gext)
       
    83 done
       
    84 
       
    85 lemma eval_gext_lemma [rule_format (no_asm)]: 
       
    86  "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w,s') \<Longrightarrow> snd s\<le>|snd s' \<and> (case w of  
       
    87     In1 v \<Rightarrow> True  
       
    88   | In2 vf \<Rightarrow> normal s \<longrightarrow> (\<forall>v x s. s\<le>|snd (assign (snd vf) v (x,s)))  
       
    89   | In3 vs \<Rightarrow> True)"
       
    90 apply (erule eval_induct)
       
    91 prefer 24 
       
    92   apply (case_tac "inited C (globs s0)", clarsimp, erule thin_rl) (* Init *)
       
    93 apply (auto del: conjI  dest!: not_initedD gext_new sxalloc_gext halloc_gext
       
    94  simp  add: lvar_def fvar_def2 avar_def2 init_lvars_def2
       
    95  split del: split_if_asm split add: sum3.split)
       
    96 (* 6 subgoals *)
       
    97 apply force+
       
    98 done
       
    99 
       
   100 lemma evar_gext_f: 
       
   101   "G\<turnstile>Norm s1 \<midarrow>e=\<succ>vf \<rightarrow> s2 \<Longrightarrow> s\<le>|snd (assign (snd vf) v (x,s))"
       
   102 apply (drule eval_gext_lemma [THEN conjunct2])
       
   103 apply auto
       
   104 done
       
   105 
       
   106 lemmas eval_gext = eval_gext_lemma [THEN conjunct1]
       
   107 
       
   108 lemma eval_gext': "G\<turnstile>(x1,s1) \<midarrow>t\<succ>\<rightarrow> (w,x2,s2) \<Longrightarrow> s1\<le>|s2"
       
   109 apply (drule eval_gext)
       
   110 apply auto
       
   111 done
       
   112 
       
   113 lemma init_yields_initd: "G\<turnstile>Norm s1 \<midarrow>Init C\<rightarrow> s2 \<Longrightarrow> initd C s2"
       
   114 apply (erule eval_cases , auto split del: split_if_asm)
       
   115 apply (case_tac "inited C (globs s1)")
       
   116 apply  (clarsimp split del: split_if_asm)+
       
   117 apply (drule eval_gext')+
       
   118 apply (drule init_class_obj_inited)
       
   119 apply (erule inited_gext)
       
   120 apply (simp (no_asm_use))
       
   121 done
       
   122 
       
   123 
       
   124 section "Lemmas"
       
   125 
       
   126 lemma obj_ty_obj_class1: 
       
   127  "\<lbrakk>wf_prog G; is_type G (obj_ty obj)\<rbrakk> \<Longrightarrow> is_class G (obj_class obj)"
       
   128 apply (case_tac "tag obj")
       
   129 apply (auto simp add: obj_ty_def obj_class_def)
       
   130 done
       
   131 
       
   132 lemma oconf_init_obj: 
       
   133  "\<lbrakk>wf_prog G;  
       
   134  (case r of Heap a \<Rightarrow> is_type G (obj_ty obj) | Stat C \<Rightarrow> is_class G C)
       
   135 \<rbrakk> \<Longrightarrow> G,s\<turnstile>obj \<lparr>values:=init_vals (var_tys G (tag obj) r)\<rparr>\<Colon>\<preceq>\<surd>r"
       
   136 apply (auto intro!: oconf_init_obj_lemma unique_fields)
       
   137 done
       
   138 
       
   139 (*
       
   140 lemma obj_split: "P obj = (\<forall> oi vs. obj = \<lparr>tag=oi,values=vs\<rparr> \<longrightarrow> ?P \<lparr>tag=oi,values=vs\<rparr>)"
       
   141 apply auto
       
   142 apply (case_tac "obj")
       
   143 apply auto
       
   144 *)
       
   145 
       
   146 lemma conforms_newG: "\<lbrakk>globs s oref = None; (x, s)\<Colon>\<preceq>(G,L);   
       
   147   wf_prog G; case oref of Heap a \<Rightarrow> is_type G (obj_ty \<lparr>tag=oi,values=vs\<rparr>)  
       
   148                         | Stat C \<Rightarrow> is_class G C\<rbrakk> \<Longrightarrow>  
       
   149   (x, init_obj G oi oref s)\<Colon>\<preceq>(G, L)"
       
   150 apply (unfold init_obj_def)
       
   151 apply (auto elim!: conforms_gupd dest!: oconf_init_obj 
       
   152             simp add: obj.update_defs)
       
   153 done
       
   154 
       
   155 lemma conforms_init_class_obj: 
       
   156  "\<lbrakk>(x,s)\<Colon>\<preceq>(G, L); wf_prog G; class G C=Some y; \<not> inited C (globs s)\<rbrakk> \<Longrightarrow> 
       
   157   (x,init_class_obj G C s)\<Colon>\<preceq>(G, L)"
       
   158 apply (rule not_initedD [THEN conforms_newG])
       
   159 apply    (auto)
       
   160 done
       
   161 
       
   162 
       
   163 lemma fst_init_lvars[simp]: 
       
   164  "fst (init_lvars G C sig (invmode m e) a' pvs (x,s)) = 
       
   165   (if static m then x else (np a') x)"
       
   166 apply (simp (no_asm) add: init_lvars_def2)
       
   167 done
       
   168 
       
   169 
       
   170 lemma halloc_conforms: "\<And>s1. \<lbrakk>G\<turnstile>s1 \<midarrow>halloc oi\<succ>a\<rightarrow> s2; wf_prog G; s1\<Colon>\<preceq>(G, L); 
       
   171   is_type G (obj_ty \<lparr>tag=oi,values=fs\<rparr>)\<rbrakk> \<Longrightarrow> s2\<Colon>\<preceq>(G, L)"
       
   172 apply (simp (no_asm_simp) only: split_tupled_all)
       
   173 apply (case_tac "aa")
       
   174 apply  (auto elim!: halloc_elim_cases dest!: new_AddrD 
       
   175        intro!: conforms_newG [THEN conforms_xconf] conf_AddrI)
       
   176 done
       
   177 
       
   178 lemma halloc_type_sound: "\<And>s1. \<lbrakk>G\<turnstile>s1 \<midarrow>halloc oi\<succ>a\<rightarrow> (x,s); wf_prog G; s1\<Colon>\<preceq>(G, L);
       
   179   T = obj_ty \<lparr>tag=oi,values=fs\<rparr>; is_type G T\<rbrakk> \<Longrightarrow>  
       
   180   (x,s)\<Colon>\<preceq>(G, L) \<and> (x = None \<longrightarrow> G,s\<turnstile>Addr a\<Colon>\<preceq>T)"
       
   181 apply (auto elim!: halloc_conforms)
       
   182 apply (case_tac "aa")
       
   183 apply (subst obj_ty_eq)
       
   184 apply  (auto elim!: halloc_elim_cases dest!: new_AddrD intro!: conf_AddrI)
       
   185 done
       
   186 
       
   187 lemma sxalloc_type_sound: 
       
   188  "\<And>s1 s2. \<lbrakk>G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2; wf_prog G\<rbrakk> \<Longrightarrow> case fst s1 of  
       
   189   None \<Rightarrow> s2 = s1 | Some x \<Rightarrow>  
       
   190   (\<exists>a. fst s2 = Some(Xcpt (Loc a)) \<and> (\<forall>L. s1\<Colon>\<preceq>(G,L) \<longrightarrow> s2\<Colon>\<preceq>(G,L)))"
       
   191 apply (simp (no_asm_simp) only: split_tupled_all)
       
   192 apply (erule sxalloc.induct)
       
   193 apply   auto
       
   194 apply (rule halloc_conforms [THEN conforms_xconf])
       
   195 apply     (auto elim!: halloc_elim_cases dest!: new_AddrD intro!: conf_AddrI)
       
   196 done
       
   197 
       
   198 lemma wt_init_comp_ty: 
       
   199 "is_acc_type G (pid C) T \<Longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>init_comp_ty T\<Colon>\<surd>"
       
   200 apply (unfold init_comp_ty_def)
       
   201 apply (clarsimp simp add: accessible_in_RefT_simp 
       
   202                           is_acc_type_def is_acc_class_def)
       
   203 done
       
   204 
       
   205 
       
   206 declare fun_upd_same [simp]
       
   207 
       
   208 declare fun_upd_apply [simp del]
       
   209 
       
   210 
       
   211 constdefs
       
   212   DynT_prop::"[prog,inv_mode,qtname,ref_ty] \<Rightarrow> bool" 
       
   213                                               ("_\<turnstile>_\<rightarrow>_\<preceq>_"[71,71,71,71]70)
       
   214  "G\<turnstile>mode\<rightarrow>D\<preceq>t \<equiv> mode = IntVir \<longrightarrow> is_class G D \<and> 
       
   215                      (if (\<exists>T. t=ArrayT T) then D=Object else G\<turnstile>Class D\<preceq>RefT t)"
       
   216 
       
   217 lemma DynT_propI: 
       
   218  "\<lbrakk>(x,s)\<Colon>\<preceq>(G, L); G,s\<turnstile>a'\<Colon>\<preceq>RefT statT; wf_prog G; mode = IntVir \<longrightarrow> a' \<noteq> Null\<rbrakk> 
       
   219   \<Longrightarrow>  G\<turnstile>mode\<rightarrow>invocation_class mode s a' statT\<preceq>statT"
       
   220 proof (unfold DynT_prop_def)
       
   221   assume state_conform: "(x,s)\<Colon>\<preceq>(G, L)"
       
   222      and      statT_a': "G,s\<turnstile>a'\<Colon>\<preceq>RefT statT"
       
   223      and            wf: "wf_prog G"
       
   224      and          mode: "mode = IntVir \<longrightarrow> a' \<noteq> Null"
       
   225   let ?invCls = "(invocation_class mode s a' statT)"
       
   226   let ?IntVir = "mode = IntVir"
       
   227   let ?Concl  = "\<lambda>invCls. is_class G invCls \<and>
       
   228                           (if \<exists>T. statT = ArrayT T
       
   229                                   then invCls = Object
       
   230                                   else G\<turnstile>Class invCls\<preceq>RefT statT)"
       
   231   show "?IntVir \<longrightarrow> ?Concl ?invCls"
       
   232   proof  
       
   233     assume modeIntVir: ?IntVir 
       
   234     with mode have not_Null: "a' \<noteq> Null" ..
       
   235     from statT_a' not_Null state_conform 
       
   236     obtain a obj 
       
   237       where obj_props:  "a' = Addr a" "globs s (Inl a) = Some obj"   
       
   238                         "G\<turnstile>obj_ty obj\<preceq>RefT statT" "is_type G (obj_ty obj)"
       
   239       by (blast dest: conforms_RefTD)
       
   240     show "?Concl ?invCls"
       
   241     proof (cases "tag obj")
       
   242       case CInst
       
   243       with modeIntVir obj_props
       
   244       show ?thesis 
       
   245 	by (auto dest!: widen_Array2 split add: split_if)
       
   246     next
       
   247       case Arr
       
   248       from Arr obtain T where "obj_ty obj = T.[]" by (blast dest: obj_ty_Arr1)
       
   249       moreover from Arr have "obj_class obj = Object" 
       
   250 	by (blast dest: obj_class_Arr1)
       
   251       moreover note modeIntVir obj_props wf 
       
   252       ultimately show ?thesis by (auto dest!: widen_Array )
       
   253     qed
       
   254   qed
       
   255 qed
       
   256 
       
   257 lemma invocation_methd:
       
   258 "\<lbrakk>wf_prog G; statT \<noteq> NullT; 
       
   259  (\<forall> statC. statT = ClassT statC \<longrightarrow> is_class G statC);
       
   260  (\<forall>     I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> mode \<noteq> SuperM);
       
   261  (\<forall>     T. statT = ArrayT T \<longrightarrow> mode \<noteq> SuperM);
       
   262  G\<turnstile>mode\<rightarrow>invocation_class mode s a' statT\<preceq>statT;  
       
   263  dynlookup G statT (invocation_class mode s a' statT) sig = Some m \<rbrakk> 
       
   264 \<Longrightarrow> methd G (invocation_declclass G mode s a' statT sig) sig = Some m"
       
   265 proof -
       
   266   assume         wf: "wf_prog G"
       
   267      and  not_NullT: "statT \<noteq> NullT"
       
   268      and statC_prop: "(\<forall> statC. statT = ClassT statC \<longrightarrow> is_class G statC)"
       
   269      and statI_prop: "(\<forall> I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> mode \<noteq> SuperM)"
       
   270      and statA_prop: "(\<forall>     T. statT = ArrayT T \<longrightarrow> mode \<noteq> SuperM)"
       
   271      and  invC_prop: "G\<turnstile>mode\<rightarrow>invocation_class mode s a' statT\<preceq>statT"
       
   272      and  dynlookup: "dynlookup G statT (invocation_class mode s a' statT) sig 
       
   273                       = Some m"
       
   274   show ?thesis
       
   275   proof (cases statT)
       
   276     case NullT
       
   277     with not_NullT show ?thesis by simp
       
   278   next
       
   279     case IfaceT
       
   280     with statI_prop obtain I 
       
   281       where    statI: "statT = IfaceT I" and 
       
   282             is_iface: "is_iface G I"     and
       
   283           not_SuperM: "mode \<noteq> SuperM" by blast            
       
   284     
       
   285     show ?thesis 
       
   286     proof (cases mode)
       
   287       case Static
       
   288       with wf dynlookup statI is_iface 
       
   289       show ?thesis
       
   290 	by (auto simp add: invocation_declclass_def dynlookup_def 
       
   291                            dynimethd_def dynmethd_C_C 
       
   292 	            intro: dynmethd_declclass
       
   293 	            dest!: wf_imethdsD
       
   294                      dest: table_of_map_SomeI
       
   295                     split: split_if_asm)
       
   296     next	
       
   297       case SuperM
       
   298       with not_SuperM show ?thesis ..
       
   299     next
       
   300       case IntVir
       
   301       with wf dynlookup IfaceT invC_prop show ?thesis 
       
   302 	by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
       
   303                            DynT_prop_def
       
   304 	            intro: methd_declclass dynmethd_declclass
       
   305 	            split: split_if_asm)
       
   306     qed
       
   307   next
       
   308     case ClassT
       
   309     show ?thesis
       
   310     proof (cases mode)
       
   311       case Static
       
   312       with wf ClassT dynlookup statC_prop 
       
   313       show ?thesis by (auto simp add: invocation_declclass_def dynlookup_def
       
   314                                intro: dynmethd_declclass)
       
   315     next
       
   316       case SuperM
       
   317       with wf ClassT dynlookup statC_prop 
       
   318       show ?thesis by (auto simp add: invocation_declclass_def dynlookup_def
       
   319                                intro: dynmethd_declclass)
       
   320     next
       
   321       case IntVir
       
   322       with wf ClassT dynlookup statC_prop invC_prop 
       
   323       show ?thesis
       
   324 	by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
       
   325                            DynT_prop_def
       
   326 	            intro: dynmethd_declclass)
       
   327     qed
       
   328   next
       
   329     case ArrayT
       
   330     show ?thesis
       
   331     proof (cases mode)
       
   332       case Static
       
   333       with wf ArrayT dynlookup show ?thesis
       
   334 	by (auto simp add: invocation_declclass_def dynlookup_def 
       
   335                            dynimethd_def dynmethd_C_C
       
   336                     intro: dynmethd_declclass
       
   337                      dest: table_of_map_SomeI)
       
   338     next
       
   339       case SuperM
       
   340       with ArrayT statA_prop show ?thesis by blast
       
   341     next
       
   342       case IntVir
       
   343       with wf ArrayT dynlookup invC_prop show ?thesis
       
   344 	by (auto simp add: invocation_declclass_def dynlookup_def dynimethd_def
       
   345                            DynT_prop_def dynmethd_C_C
       
   346                     intro: dynmethd_declclass
       
   347                      dest: table_of_map_SomeI)
       
   348     qed
       
   349   qed
       
   350 qed
       
   351    
       
   352 declare split_paired_All [simp del] split_paired_Ex [simp del] 
       
   353 ML_setup {*
       
   354 simpset_ref() := simpset() delloop "split_all_tac";
       
   355 claset_ref () := claset () delSWrapper "split_all_tac"
       
   356 *}
       
   357 lemma DynT_mheadsD: 
       
   358 "\<lbrakk>G\<turnstile>invmode (mhd sm) e\<rightarrow>invC\<preceq>statT; 
       
   359   wf_prog G; \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT; 
       
   360   sm \<in> mheads G C statT sig; 
       
   361   invC = invocation_class (invmode (mhd sm) e) s a' statT;
       
   362   declC =invocation_declclass G (invmode (mhd sm) e) s a' statT sig
       
   363  \<rbrakk> \<Longrightarrow> 
       
   364   \<exists> dm. 
       
   365   methd G declC sig = Some dm  \<and> G\<turnstile>resTy (mthd dm)\<preceq>resTy (mhd sm) \<and> 
       
   366   wf_mdecl G declC (sig, mthd dm) \<and>
       
   367   declC = declclass dm \<and>
       
   368   is_static dm = is_static sm \<and>  
       
   369   is_class G invC \<and> is_class G declC  \<and> G\<turnstile>invC\<preceq>\<^sub>C declC \<and>  
       
   370   (if invmode (mhd sm) e = IntVir
       
   371       then (\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>invC \<preceq>\<^sub>C statC)
       
   372       else (  (\<exists> statC. statT=ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C declC)
       
   373             \<or> (\<forall> statC. statT\<noteq>ClassT statC \<and> declC=Object)) \<and> 
       
   374            (declrefT sm) = ClassT (declclass dm))"
       
   375 proof -
       
   376   assume invC_prop: "G\<turnstile>invmode (mhd sm) e\<rightarrow>invC\<preceq>statT" 
       
   377      and        wf: "wf_prog G" 
       
   378      and      wt_e: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT"
       
   379      and        sm: "sm \<in> mheads G C statT sig" 
       
   380      and      invC: "invC = invocation_class (invmode (mhd sm) e) s a' statT"
       
   381      and     declC: "declC = 
       
   382                     invocation_declclass G (invmode (mhd sm) e) s a' statT sig"
       
   383   from wt_e wf have type_statT: "is_type G (RefT statT)"
       
   384     by (auto dest: ty_expr_is_type)
       
   385   from sm have not_Null: "statT \<noteq> NullT" by auto
       
   386   from type_statT 
       
   387   have wf_C: "(\<forall> statC. statT = ClassT statC \<longrightarrow> is_class G statC)"
       
   388     by (auto)
       
   389   from type_statT wt_e 
       
   390   have wf_I: "(\<forall>I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> 
       
   391                                         invmode (mhd sm) e \<noteq> SuperM)"
       
   392     by (auto dest: invocationTypeExpr_noClassD)
       
   393   from wt_e
       
   394   have wf_A: "(\<forall>     T. statT = ArrayT T \<longrightarrow> invmode (mhd sm) e \<noteq> SuperM)"
       
   395     by (auto dest: invocationTypeExpr_noClassD)
       
   396   show ?thesis
       
   397   proof (cases "invmode (mhd sm) e = IntVir")
       
   398     case True
       
   399     with invC_prop not_Null
       
   400     have invC_prop': " is_class G invC \<and> 
       
   401                       (if (\<exists>T. statT=ArrayT T) then invC=Object 
       
   402                                               else G\<turnstile>Class invC\<preceq>RefT statT)"
       
   403       by (auto simp add: DynT_prop_def) 
       
   404     from True 
       
   405     have "\<not> is_static sm"
       
   406       by (simp add: invmode_IntVir_eq)
       
   407     with invC_prop' not_Null
       
   408     have "G,statT \<turnstile> invC valid_lookup_cls_for (is_static sm)"
       
   409       by (cases statT) auto
       
   410     with sm wf type_statT obtain dm where
       
   411            dm: "dynlookup G statT invC sig = Some dm" and
       
   412       resT_dm: "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mhd sm)"      and
       
   413        static: "is_static dm = is_static sm"
       
   414       by - (drule dynamic_mheadsD,auto)
       
   415     with declC invC not_Null 
       
   416     have declC': "declC = (declclass dm)" 
       
   417       by (auto simp add: invocation_declclass_def)
       
   418     with wf invC declC not_Null wf_C wf_I wf_A invC_prop dm 
       
   419     have dm': "methd G declC sig = Some dm"
       
   420       by - (drule invocation_methd,auto)
       
   421     from wf dm invC_prop' declC' type_statT 
       
   422     have declC_prop: "G\<turnstile>invC \<preceq>\<^sub>C declC \<and> is_class G declC"
       
   423       by (auto dest: dynlookup_declC)
       
   424     from wf dm' declC_prop declC' 
       
   425     have wf_dm: "wf_mdecl G declC (sig,(mthd dm))"
       
   426       by (auto dest: methd_wf_mdecl)
       
   427     from invC_prop' 
       
   428     have statC_prop: "(\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>invC \<preceq>\<^sub>C statC)"
       
   429       by auto
       
   430     from True dm' resT_dm wf_dm invC_prop' declC_prop statC_prop declC' static
       
   431     show ?thesis by auto
       
   432   next
       
   433     case False
       
   434     with type_statT wf invC not_Null wf_I wf_A
       
   435     have invC_prop': "is_class G invC \<and>  
       
   436                      ((\<exists> statC. statT=ClassT statC \<and> invC=statC) \<or>
       
   437                       (\<forall> statC. statT\<noteq>ClassT statC \<and> invC=Object)) "
       
   438         by (case_tac "statT") (auto simp add: invocation_class_def 
       
   439                                        split: inv_mode.splits)
       
   440     with not_Null wf
       
   441     have dynlookup_static: "dynlookup G statT invC sig = methd G invC sig"
       
   442       by (case_tac "statT") (auto simp add: dynlookup_def dynmethd_C_C
       
   443                                             dynimethd_def)
       
   444     from sm wf wt_e not_Null False invC_prop' obtain "dm" where
       
   445                     dm: "methd G invC sig = Some dm"          and
       
   446 	eq_declC_sm_dm:"declrefT sm = ClassT (declclass dm)"  and
       
   447 	     eq_mheads:"mhd sm=mhead (mthd dm) "
       
   448       by - (drule static_mheadsD, auto dest: accmethd_SomeD)
       
   449     then have static: "is_static dm = is_static sm" by - (case_tac "sm",auto)
       
   450     with declC invC dynlookup_static dm
       
   451     have declC': "declC = (declclass dm)"  
       
   452       by (auto simp add: invocation_declclass_def)
       
   453     from invC_prop' wf declC' dm 
       
   454     have dm': "methd G declC sig = Some dm"
       
   455       by (auto intro: methd_declclass)
       
   456     from wf dm invC_prop' declC' type_statT 
       
   457     have declC_prop: "G\<turnstile>invC \<preceq>\<^sub>C declC \<and> is_class G declC"
       
   458       by (auto dest: methd_declC )
       
   459     then have declC_prop1: "invC=Object \<longrightarrow> declC=Object"  by auto
       
   460     from wf dm' declC_prop declC' 
       
   461     have wf_dm: "wf_mdecl G declC (sig,(mthd dm))"
       
   462       by (auto dest: methd_wf_mdecl)
       
   463     from invC_prop' declC_prop declC_prop1
       
   464     have statC_prop: "(   (\<exists> statC. statT=ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C declC)
       
   465                        \<or>  (\<forall> statC. statT\<noteq>ClassT statC \<and> declC=Object))" 
       
   466       by auto
       
   467     from False dm' wf_dm invC_prop' declC_prop statC_prop declC' 
       
   468          eq_declC_sm_dm eq_mheads static
       
   469     show ?thesis by auto
       
   470   qed
       
   471 qed
       
   472 
       
   473 (*
       
   474 lemma DynT_mheadsD: 
       
   475 "\<lbrakk>G\<turnstile>invmode (mhd sm) e\<rightarrow>invC\<preceq>statT; 
       
   476   wf_prog G; \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT; 
       
   477   sm \<in> mheads G C statT sig; 
       
   478   invC = invocation_class (invmode (mhd sm) e) s a' statT;
       
   479   declC =invocation_declclass G (invmode (mhd sm) e) s a' statT sig
       
   480  \<rbrakk> \<Longrightarrow> 
       
   481   \<exists> dm. 
       
   482   methd G declC sig = Some dm  \<and> G\<turnstile>resTy (mthd dm)\<preceq>resTy (mhd sm) \<and> 
       
   483   wf_mdecl G declC (sig, mthd dm) \<and>  
       
   484   is_class G invC \<and> is_class G declC  \<and> G\<turnstile>invC\<preceq>\<^sub>C declC \<and>  
       
   485   (if invmode (mhd sm) e = IntVir
       
   486       then (\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>invC \<preceq>\<^sub>C statC)
       
   487       else (\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>statC \<preceq>\<^sub>C declC) \<and> 
       
   488            (declrefT sm) = ClassT (declclass dm))"
       
   489 proof -
       
   490   assume invC_prop: "G\<turnstile>invmode (mhd sm) e\<rightarrow>invC\<preceq>statT" 
       
   491      and        wf: "wf_prog G" 
       
   492      and      wt_e: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT"
       
   493      and        sm: "sm \<in> mheads G C statT sig" 
       
   494      and      invC: "invC = invocation_class (invmode (mhd sm) e) s a' statT"
       
   495      and     declC: "declC = 
       
   496                     invocation_declclass G (invmode (mhd sm) e) s a' statT sig"
       
   497   from wt_e wf have type_statT: "is_type G (RefT statT)"
       
   498     by (auto dest: ty_expr_is_type)
       
   499   from sm have not_Null: "statT \<noteq> NullT" by auto
       
   500   from type_statT 
       
   501   have wf_C: "(\<forall> statC. statT = ClassT statC \<longrightarrow> is_class G statC)"
       
   502     by (auto)
       
   503   from type_statT wt_e 
       
   504   have wf_I: "(\<forall>I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> 
       
   505                                         invmode (mhd sm) e \<noteq> SuperM)"
       
   506     by (auto dest: invocationTypeExpr_noClassD)
       
   507   from wt_e
       
   508   have wf_A: "(\<forall>     T. statT = ArrayT T \<longrightarrow> invmode (mhd sm) e \<noteq> SuperM)"
       
   509     by (auto dest: invocationTypeExpr_noClassD)
       
   510   show ?thesis
       
   511   proof (cases "invmode (mhd sm) e = IntVir")
       
   512     case True
       
   513     with invC_prop not_Null
       
   514     have invC_prop': "is_class G invC \<and>  
       
   515                       (if (\<exists>T. statT=ArrayT T) then invC=Object 
       
   516                                               else G\<turnstile>Class invC\<preceq>RefT statT)"
       
   517       by (auto simp add: DynT_prop_def) 
       
   518     with sm wf type_statT not_Null obtain dm where
       
   519            dm: "dynlookup G statT invC sig = Some dm" and
       
   520       resT_dm: "G\<turnstile>resTy (mthd dm)\<preceq>resTy (mhd sm)"
       
   521       by - (clarify,drule dynamic_mheadsD,auto)
       
   522     with declC invC not_Null 
       
   523     have declC': "declC = (declclass dm)" 
       
   524       by (auto simp add: invocation_declclass_def)
       
   525     with wf invC declC not_Null wf_C wf_I wf_A invC_prop dm 
       
   526     have dm': "methd G declC sig = Some dm"
       
   527       by - (drule invocation_methd,auto)
       
   528     from wf dm invC_prop' declC' type_statT 
       
   529     have declC_prop: "G\<turnstile>invC \<preceq>\<^sub>C declC \<and> is_class G declC"
       
   530       by (auto dest: dynlookup_declC)
       
   531     from wf dm' declC_prop declC' 
       
   532     have wf_dm: "wf_mdecl G declC (sig,(mthd dm))"
       
   533       by (auto dest: methd_wf_mdecl)
       
   534     from invC_prop' 
       
   535     have statC_prop: "(\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>invC \<preceq>\<^sub>C statC)"
       
   536       by auto
       
   537     from True dm' resT_dm wf_dm invC_prop' declC_prop statC_prop
       
   538     show ?thesis by auto
       
   539   next
       
   540     case False
       
   541     
       
   542     with type_statT wf invC not_Null wf_I wf_A
       
   543     have invC_prop': "is_class G invC \<and>  
       
   544                      ((\<exists> statC. statT=ClassT statC \<and> invC=statC) \<or>
       
   545                       (\<forall> statC. statT\<noteq>ClassT statC \<and> invC=Object)) "
       
   546         
       
   547         by (case_tac "statT") (auto simp add: invocation_class_def 
       
   548                                        split: inv_mode.splits)
       
   549     with not_Null 
       
   550     have dynlookup_static: "dynlookup G statT invC sig = methd G invC sig"
       
   551       by (case_tac "statT") (auto simp add: dynlookup_def dynmethd_def 
       
   552                                             dynimethd_def)
       
   553     from sm wf wt_e not_Null False invC_prop' obtain "dm" where
       
   554                     dm: "methd G invC sig = Some dm"          and
       
   555 	eq_declC_sm_dm:"declrefT sm = ClassT (declclass dm)"  and
       
   556 	     eq_mheads:"mhd sm=mhead (mthd dm) "
       
   557       by - (drule static_mheadsD, auto dest: accmethd_SomeD)
       
   558     with declC invC dynlookup_static dm
       
   559     have declC': "declC = (declclass dm)"  
       
   560       by (auto simp add: invocation_declclass_def)
       
   561     from invC_prop' wf declC' dm 
       
   562     have dm': "methd G declC sig = Some dm"
       
   563       by (auto intro: methd_declclass)
       
   564     from wf dm invC_prop' declC' type_statT 
       
   565     have declC_prop: "G\<turnstile>invC \<preceq>\<^sub>C declC \<and> is_class G declC"
       
   566       by (auto dest: methd_declC )   
       
   567     from wf dm' declC_prop declC' 
       
   568     have wf_dm: "wf_mdecl G declC (sig,(mthd dm))"
       
   569       by (auto dest: methd_wf_mdecl)
       
   570     from invC_prop' declC_prop
       
   571     have statC_prop: "(\<forall> statC. statT=ClassT statC \<longrightarrow> G\<turnstile>statC \<preceq>\<^sub>C declC)" 
       
   572       by auto
       
   573     from False dm' wf_dm invC_prop' declC_prop statC_prop 
       
   574          eq_declC_sm_dm eq_mheads
       
   575     show ?thesis by auto
       
   576   qed
       
   577 qed	
       
   578 *)
       
   579 
       
   580 declare split_paired_All [simp del] split_paired_Ex [simp del] 
       
   581 declare split_if     [split del] split_if_asm     [split del] 
       
   582         option.split [split del] option.split_asm [split del]
       
   583 ML_setup {*
       
   584 simpset_ref() := simpset() delloop "split_all_tac";
       
   585 claset_ref () := claset () delSWrapper "split_all_tac"
       
   586 *}
       
   587 
       
   588 lemma DynT_conf: "\<lbrakk>G\<turnstile>invocation_class mode s a' statT \<preceq>\<^sub>C declC; wf_prog G;
       
   589  isrtype G (statT);
       
   590  G,s\<turnstile>a'\<Colon>\<preceq>RefT statT; mode = IntVir \<longrightarrow> a' \<noteq> Null;  
       
   591   mode \<noteq> IntVir \<longrightarrow>    (\<exists> statC. statT=ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C declC)
       
   592                     \<or>  (\<forall> statC. statT\<noteq>ClassT statC \<and> declC=Object)\<rbrakk> 
       
   593  \<Longrightarrow>G,s\<turnstile>a'\<Colon>\<preceq> Class declC"
       
   594 apply (case_tac "mode = IntVir")
       
   595 apply (drule conf_RefTD)
       
   596 apply (force intro!: conf_AddrI 
       
   597             simp add: obj_class_def split add: obj_tag.split_asm)
       
   598 apply  clarsimp
       
   599 apply  safe
       
   600 apply    (erule (1) widen.subcls [THEN conf_widen])
       
   601 apply    (erule wf_ws_prog)
       
   602 
       
   603 apply    (frule widen_Object) apply (erule wf_ws_prog)
       
   604 apply    (erule (1) conf_widen) apply (erule wf_ws_prog)
       
   605 done
       
   606 
       
   607 
       
   608 lemma Ass_lemma: 
       
   609  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w, f)\<rightarrow> Norm s1; G\<turnstile>Norm s1 \<midarrow>e-\<succ>v\<rightarrow> Norm s2; G,s2\<turnstile>v\<Colon>\<preceq>T'; 
       
   610    s1\<le>|s2 \<longrightarrow> assign f v (Norm s2)\<Colon>\<preceq>(G, L)
       
   611   \<rbrakk> \<Longrightarrow> assign f v (Norm s2)\<Colon>\<preceq>(G, L) \<and> 
       
   612         (\<lambda>(x',s'). x' = None \<longrightarrow> G,s'\<turnstile>v\<Colon>\<preceq>T') (assign f v (Norm s2))"
       
   613 apply (drule_tac x = "None" and s = "s2" and v = "v" in evar_gext_f)
       
   614 apply (drule eval_gext', clarsimp)
       
   615 apply (erule conf_gext)
       
   616 apply simp
       
   617 done
       
   618 
       
   619 lemma Throw_lemma: "\<lbrakk>G\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable; wf_prog G; (x1,s1)\<Colon>\<preceq>(G, L);  
       
   620     x1 = None \<longrightarrow> G,s1\<turnstile>a'\<Colon>\<preceq> Class tn\<rbrakk> \<Longrightarrow> (throw a' x1, s1)\<Colon>\<preceq>(G, L)"
       
   621 apply (auto split add: split_abrupt_if simp add: throw_def2)
       
   622 apply (erule conforms_xconf)
       
   623 apply (frule conf_RefTD)
       
   624 apply (auto elim: widen.subcls [THEN conf_widen])
       
   625 done
       
   626 
       
   627 lemma Try_lemma: "\<lbrakk>G\<turnstile>obj_ty (the (globs s1' (Heap a)))\<preceq> Class tn; 
       
   628  (Some (Xcpt (Loc a)), s1')\<Colon>\<preceq>(G, L); wf_prog G\<rbrakk> 
       
   629  \<Longrightarrow> Norm (lupd(vn\<mapsto>Addr a) s1')\<Colon>\<preceq>(G, L(vn\<mapsto>Class tn))"
       
   630 apply (rule conforms_allocL)
       
   631 apply  (erule conforms_NormI)
       
   632 apply (drule conforms_XcptLocD [THEN conf_RefTD],rule HOL.refl)
       
   633 apply (auto intro!: conf_AddrI)
       
   634 done
       
   635 
       
   636 lemma Fin_lemma: 
       
   637 "\<lbrakk>G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> (x2,s2); wf_prog G; (Some a, s1)\<Colon>\<preceq>(G, L); (x2,s2)\<Colon>\<preceq>(G, L)\<rbrakk> 
       
   638 \<Longrightarrow>  (abrupt_if True (Some a) x2, s2)\<Colon>\<preceq>(G, L)"
       
   639 apply (force elim: eval_gext' conforms_xgext split add: split_abrupt_if)
       
   640 done
       
   641 
       
   642 lemma FVar_lemma1: "\<lbrakk>table_of (DeclConcepts.fields G Ca) (fn, C) = Some f ; 
       
   643   x2 = None \<longrightarrow> G,s2\<turnstile>a\<Colon>\<preceq> Class Ca; wf_prog G; G\<turnstile>Ca\<preceq>\<^sub>C C; C \<noteq> Object; 
       
   644   class G C = Some y; (x2,s2)\<Colon>\<preceq>(G, L); s1\<le>|s2; inited C (globs s1); 
       
   645   (if static f then id else np a) x2 = None\<rbrakk> 
       
   646  \<Longrightarrow>  
       
   647   \<exists>obj. globs s2 (if static f then Inr C else Inl (the_Addr a)) = Some obj \<and> 
       
   648   var_tys G (tag obj)  (if static f then Inr C else Inl(the_Addr a)) 
       
   649           (Inl(fn,C)) = Some (type f)"
       
   650 apply (drule initedD)
       
   651 apply (frule subcls_is_class2, simp (no_asm_simp))
       
   652 apply (case_tac "static f")
       
   653 apply  clarsimp
       
   654 apply  (drule (1) rev_gext_objD, clarsimp)
       
   655 apply  (frule fields_declC, erule wf_ws_prog, simp (no_asm_simp))
       
   656 apply  (rule var_tys_Some_eq [THEN iffD2])
       
   657 apply  clarsimp
       
   658 apply  (erule fields_table_SomeI, simp (no_asm))
       
   659 apply clarsimp
       
   660 apply (drule conf_RefTD, clarsimp, rule var_tys_Some_eq [THEN iffD2])
       
   661 apply (auto dest!: widen_Array split add: obj_tag.split)
       
   662 apply (rotate_tac -1) (* to front: tag obja = CInst pid_field_type to enable
       
   663                          conditional rewrite *)
       
   664 apply (rule fields_table_SomeI)
       
   665 apply (auto elim!: fields_mono subcls_is_class2)
       
   666 done
       
   667 
       
   668 lemma FVar_lemma: 
       
   669 "\<lbrakk>((v, f), Norm s2') = fvar C (static field) fn a (x2, s2); G\<turnstile>Ca\<preceq>\<^sub>C C;  
       
   670   table_of (DeclConcepts.fields G Ca) (fn, C) = Some field; wf_prog G;   
       
   671   x2 = None \<longrightarrow> G,s2\<turnstile>a\<Colon>\<preceq>Class Ca; C \<noteq> Object; class G C = Some y;   
       
   672   (x2, s2)\<Colon>\<preceq>(G, L); s1\<le>|s2; inited C (globs s1)\<rbrakk> \<Longrightarrow>  
       
   673   G,s2'\<turnstile>v\<Colon>\<preceq>type field \<and> s2'\<le>|f\<preceq>type field\<Colon>\<preceq>(G, L)"
       
   674 apply (unfold assign_conforms_def)
       
   675 apply (drule sym)
       
   676 apply (clarsimp simp add: fvar_def2)
       
   677 apply (drule (9) FVar_lemma1)
       
   678 apply (clarsimp)
       
   679 apply (drule (2) conforms_globsD [THEN oconf_lconf, THEN lconfD])
       
   680 apply clarsimp
       
   681 apply (drule (1) rev_gext_objD)
       
   682 apply (auto elim!: conforms_upd_gobj)
       
   683 done
       
   684 
       
   685 
       
   686 lemma AVar_lemma1: "\<lbrakk>globs s (Inl a) = Some obj;tag obj=Arr ty i; 
       
   687  the_Intg i' in_bounds i; wf_prog G; G\<turnstile>ty.[]\<preceq>Tb.[]; Norm s\<Colon>\<preceq>(G, L)
       
   688 \<rbrakk> \<Longrightarrow> G,s\<turnstile>the ((values obj) (Inr (the_Intg i')))\<Colon>\<preceq>Tb"
       
   689 apply (erule widen_Array_Array [THEN conf_widen])
       
   690 apply  (erule_tac [2] wf_ws_prog)
       
   691 apply (drule (1) conforms_globsD [THEN oconf_lconf, THEN lconfD])
       
   692 defer apply assumption
       
   693 apply (force intro: var_tys_Some_eq [THEN iffD2])
       
   694 done
       
   695 
       
   696 lemma obj_split: "\<And> obj. \<exists> t vs. obj = \<lparr>tag=t,values=vs\<rparr>"
       
   697 proof record_split
       
   698   fix tag values more
       
   699   show "\<exists>t vs. \<lparr>tag = tag, values = values, \<dots> = more\<rparr> = \<lparr>tag = t, values = vs\<rparr>"
       
   700     by auto
       
   701 qed
       
   702  
       
   703 lemma AVar_lemma: "\<lbrakk>wf_prog G; G\<turnstile>(x1, s1) \<midarrow>e2-\<succ>i\<rightarrow> (x2, s2);  
       
   704   ((v,f), Norm s2') = avar G i a (x2, s2); x1 = None \<longrightarrow> G,s1\<turnstile>a\<Colon>\<preceq>Ta.[];  
       
   705   (x2, s2)\<Colon>\<preceq>(G, L); s1\<le>|s2\<rbrakk> \<Longrightarrow> G,s2'\<turnstile>v\<Colon>\<preceq>Ta \<and> s2'\<le>|f\<preceq>Ta\<Colon>\<preceq>(G, L)"
       
   706 apply (unfold assign_conforms_def)
       
   707 apply (drule sym)
       
   708 apply (clarsimp simp add: avar_def2)
       
   709 apply (drule (1) conf_gext)
       
   710 apply (drule conf_RefTD, clarsimp)
       
   711 apply (subgoal_tac "\<exists> t vs. obj = \<lparr>tag=t,values=vs\<rparr>")
       
   712 defer
       
   713 apply   (rule obj_split)
       
   714 apply clarify
       
   715 apply (frule obj_ty_widenD)
       
   716 apply (auto dest!: widen_Class)
       
   717 apply  (force dest: AVar_lemma1)
       
   718 apply (auto split add: split_if)
       
   719 apply (force elim!: fits_Array dest: gext_objD 
       
   720        intro: var_tys_Some_eq [THEN iffD2] conforms_upd_gobj)
       
   721 done
       
   722 
       
   723 
       
   724 section "Call"
       
   725 lemma conforms_init_lvars_lemma: "\<lbrakk>wf_prog G;  
       
   726   wf_mhead G P sig mh; 
       
   727   Ball (set lvars) (split (\<lambda>vn. is_type G)); 
       
   728   list_all2 (conf G s) pvs pTsa; G\<turnstile>pTsa[\<preceq>](parTs sig)\<rbrakk> \<Longrightarrow>  
       
   729   G,s\<turnstile>init_vals (table_of lvars)(pars mh[\<mapsto>]pvs)
       
   730       [\<Colon>\<preceq>]table_of lvars(pars mh[\<mapsto>]parTs sig)"
       
   731 apply (unfold wf_mhead_def)
       
   732 apply clarify
       
   733 apply (erule (2) Ball_set_table [THEN lconf_init_vals, THEN lconf_ext_list])
       
   734 apply (drule wf_ws_prog)
       
   735 apply (erule (2) conf_list_widen)
       
   736 done
       
   737 
       
   738 
       
   739 lemma lconf_map_lname [simp]: 
       
   740   "G,s\<turnstile>(lname_case l1 l2)[\<Colon>\<preceq>](lname_case L1 L2)
       
   741    =
       
   742   (G,s\<turnstile>l1[\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>(\<lambda>x::unit . l2)[\<Colon>\<preceq>](\<lambda>x::unit. L2))"
       
   743 apply (unfold lconf_def)
       
   744 apply safe
       
   745 apply (case_tac "n")
       
   746 apply (force split add: lname.split)+
       
   747 done
       
   748 
       
   749 lemma lconf_map_ename [simp]:
       
   750   "G,s\<turnstile>(ename_case l1 l2)[\<Colon>\<preceq>](ename_case L1 L2)
       
   751    =
       
   752   (G,s\<turnstile>l1[\<Colon>\<preceq>]L1 \<and> G,s\<turnstile>(\<lambda>x::unit. l2)[\<Colon>\<preceq>](\<lambda>x::unit. L2))"
       
   753 apply (unfold lconf_def)
       
   754 apply safe
       
   755 apply (force split add: ename.split)+
       
   756 done
       
   757 
       
   758 
       
   759 lemma defval_conf1 [rule_format (no_asm), elim]: 
       
   760   "is_type G T \<longrightarrow> (\<exists>v\<in>Some (default_val T): G,s\<turnstile>v\<Colon>\<preceq>T)"
       
   761 apply (unfold conf_def)
       
   762 apply (induct "T")
       
   763 apply (auto intro: prim_ty.induct)
       
   764 done
       
   765 
       
   766 
       
   767 lemma conforms_init_lvars: 
       
   768 "\<lbrakk>wf_mhead G (pid declC) sig (mhead (mthd dm)); wf_prog G;  
       
   769   list_all2 (conf G s) pvs pTsa; G\<turnstile>pTsa[\<preceq>](parTs sig);  
       
   770   (x, s)\<Colon>\<preceq>(G, L); 
       
   771   methd G declC sig = Some dm;  
       
   772   isrtype G statT;
       
   773   G\<turnstile>invC\<preceq>\<^sub>C declC; 
       
   774   G,s\<turnstile>a'\<Colon>\<preceq>RefT statT;  
       
   775   invmode (mhd sm) e = IntVir \<longrightarrow> a' \<noteq> Null; 
       
   776   invmode (mhd sm) e \<noteq> IntVir \<longrightarrow>  
       
   777        (\<exists> statC. statT=ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C declC)
       
   778     \<or>  (\<forall> statC. statT\<noteq>ClassT statC \<and> declC=Object);
       
   779   invC  = invocation_class (invmode (mhd sm) e) s a' statT;
       
   780   declC = invocation_declclass G (invmode (mhd sm) e) s a' statT sig;
       
   781   Ball (set (lcls (mbody (mthd dm)))) 
       
   782        (split (\<lambda>vn. is_type G)) 
       
   783  \<rbrakk> \<Longrightarrow>  
       
   784   init_lvars G declC sig (invmode (mhd sm) e) a'  
       
   785   pvs (x,s)\<Colon>\<preceq>(G,\<lambda> k. 
       
   786                 (case k of
       
   787                    EName e \<Rightarrow> (case e of 
       
   788                                  VNam v 
       
   789                                   \<Rightarrow> (table_of (lcls (mbody (mthd dm)))
       
   790                                         (pars (mthd dm)[\<mapsto>]parTs sig)) v
       
   791                                | Res \<Rightarrow> Some (resTy (mthd dm)))
       
   792                  | This \<Rightarrow> if (static (mthd sm)) 
       
   793                               then None else Some (Class declC)))"
       
   794 apply (simp add: init_lvars_def2)
       
   795 apply (rule conforms_set_locals)
       
   796 apply  (simp (no_asm_simp) split add: split_if)
       
   797 apply (drule  (4) DynT_conf)
       
   798 apply clarsimp
       
   799 (* apply intro *)
       
   800 apply (drule (4) conforms_init_lvars_lemma)
       
   801 apply (case_tac "dm",simp)
       
   802 apply (rule conjI)
       
   803 apply (unfold lconf_def, clarify)
       
   804 apply (rule defval_conf1)
       
   805 apply (clarsimp simp add: wf_mhead_def is_acc_type_def)
       
   806 apply (case_tac "is_static sm")
       
   807 apply simp_all
       
   808 done
       
   809 
       
   810 
       
   811 lemma Call_type_sound: "\<lbrakk>wf_prog G; G\<turnstile>(x1, s1) \<midarrow>ps\<doteq>\<succ>pvs\<rightarrow> (x2, s2);  
       
   812  declC 
       
   813  = invocation_declclass G (invmode (mhd esm) e) s2 a' statT \<lparr>name=mn,parTs=pTs\<rparr>;
       
   814 s2'=init_lvars G declC \<lparr>name=mn,parTs=pTs\<rparr> (invmode (mhd esm) e) a' pvs (x2,s2);
       
   815  G\<turnstile>s2' \<midarrow>Methd declC \<lparr>name=mn,parTs=pTs\<rparr>-\<succ>v\<rightarrow> (x3, s3);    
       
   816  \<forall>L. s2'\<Colon>\<preceq>(G, L) 
       
   817      \<longrightarrow> (\<forall>T. \<lparr>prg=G,cls=declC,lcl=L\<rparr>\<turnstile> Methd declC \<lparr>name=mn,parTs=pTs\<rparr>\<Colon>-T 
       
   818      \<longrightarrow> (x3, s3)\<Colon>\<preceq>(G, L) \<and> (x3 = None \<longrightarrow> G,s3\<turnstile>v\<Colon>\<preceq>T));  
       
   819  Norm s0\<Colon>\<preceq>(G, L); 
       
   820  \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT; \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>ps\<Colon>\<doteq>pTsa;  
       
   821  max_spec G C statT \<lparr>name=mn,parTs=pTsa\<rparr> = {(esm, pTs)}; 
       
   822  (x1, s1)\<Colon>\<preceq>(G, L); 
       
   823  x1 = None \<longrightarrow> G,s1\<turnstile>a'\<Colon>\<preceq>RefT statT; (x2, s2)\<Colon>\<preceq>(G, L);  
       
   824  x2 = None \<longrightarrow> list_all2 (conf G s2) pvs pTsa;
       
   825  sm=(mhd esm)\<rbrakk> \<Longrightarrow>     
       
   826  (x3, set_locals (locals s2) s3)\<Colon>\<preceq>(G, L) \<and> 
       
   827  (x3 = None \<longrightarrow> G,s3\<turnstile>v\<Colon>\<preceq>resTy sm)"
       
   828 apply clarify
       
   829 apply (case_tac "x2")
       
   830 defer
       
   831 apply  (clarsimp split add: split_if_asm simp add: init_lvars_def2)
       
   832 apply (case_tac "a' = Null \<and> \<not> (static (mhd esm)) \<and> e \<noteq> Super")
       
   833 apply  (clarsimp simp add: init_lvars_def2)
       
   834 apply clarsimp
       
   835 apply (drule eval_gext')
       
   836 apply (frule (1) conf_gext)
       
   837 apply (drule max_spec2mheads, clarsimp)
       
   838 apply (subgoal_tac "invmode (mhd esm) e = IntVir \<longrightarrow> a' \<noteq> Null")
       
   839 defer  
       
   840 apply  (clarsimp simp add: invmode_IntVir_eq)
       
   841 apply (frule (6) DynT_mheadsD [OF DynT_propI,of _ "s2"],(rule HOL.refl)+)
       
   842 apply clarify
       
   843 apply (drule wf_mdeclD1, clarsimp) 
       
   844 apply (frule  ty_expr_is_type) apply simp
       
   845 apply (frule (2) conforms_init_lvars)
       
   846 apply   simp
       
   847 apply   assumption+
       
   848 apply   simp
       
   849 apply   assumption+
       
   850 apply   clarsimp
       
   851 apply   (rule HOL.refl)
       
   852 apply   simp
       
   853 apply   (rule Ball_weaken)
       
   854 apply     assumption
       
   855 apply     (force simp add: is_acc_type_def)
       
   856 apply (tactic "smp_tac 1 1")
       
   857 apply (frule (2) wt_MethdI, clarsimp)
       
   858 apply (subgoal_tac "is_static dm = (static (mthd esm))") 
       
   859 apply   (simp only:)
       
   860 apply   (tactic "smp_tac 1 1")
       
   861 apply   (rule conjI)
       
   862 apply     (erule  conforms_return)
       
   863 apply     blast
       
   864 
       
   865 apply     (force dest!: eval_gext del: impCE simp add: init_lvars_def2)
       
   866 apply     clarsimp
       
   867 apply     (drule (2) widen_trans, erule (1) conf_widen)
       
   868 apply     (erule wf_ws_prog)
       
   869 
       
   870 apply   auto
       
   871 done
       
   872 
       
   873 
       
   874 subsection "accessibility"
       
   875 
       
   876 theorem dynamic_field_access_ok:
       
   877   (assumes wf: "wf_prog G" and
       
   878        eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2" and
       
   879      not_Null: "a\<noteq>Null" and
       
   880     conform_a: "G,(store s2)\<turnstile>a\<Colon>\<preceq> Class statC" and
       
   881    conform_s2: "s2\<Colon>\<preceq>(G, L)" and 
       
   882     normal_s2: "normal s2" and
       
   883          wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>,dt\<Turnstile>e\<Colon>-Class statC" and
       
   884             f: "accfield G accC statC fn = Some f" and
       
   885          dynC: "if stat then dynC=statC  
       
   886                         else dynC=obj_class (lookup_obj (store s2) a)"
       
   887   ) "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f) \<and> 
       
   888      G\<turnstile>Field fn f in dynC dyn_accessible_from accC"
       
   889 proof (cases "stat")
       
   890   case True
       
   891   with dynC 
       
   892   have dynC': "dynC=statC" by simp
       
   893   with f
       
   894   have "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f)"
       
   895     by (auto simp add: accfield_def Let_def intro!: table_of_remap_SomeD)
       
   896   with dynC' f
       
   897   show ?thesis
       
   898     by (auto intro!: static_to_dynamic_accessible_from
       
   899          dest: accfield_accessibleD accessible_from_commonD)
       
   900 next
       
   901   case False
       
   902   with wf conform_a not_Null conform_s2 dynC
       
   903   obtain subclseq: "G\<turnstile>dynC \<preceq>\<^sub>C statC" and
       
   904     "is_class G dynC"
       
   905     by (auto dest!: conforms_RefTD [of _ _ _ _ "(fst s2)" L]
       
   906               dest: obj_ty_obj_class1
       
   907           simp add: obj_ty_obj_class )
       
   908   with wf f
       
   909   have "table_of (DeclConcepts.fields G dynC) (fn,declclass f) = Some (fld f)"
       
   910     by (auto simp add: accfield_def Let_def
       
   911                  dest: fields_mono
       
   912                 dest!: table_of_remap_SomeD)
       
   913   moreover
       
   914   from f subclseq
       
   915   have "G\<turnstile>Field fn f in dynC dyn_accessible_from accC"
       
   916     by (auto intro!: static_to_dynamic_accessible_from 
       
   917                dest: accfield_accessibleD)
       
   918   ultimately show ?thesis
       
   919     by blast
       
   920 qed
       
   921 
       
   922 lemma call_access_ok: 
       
   923 (assumes invC_prop: "G\<turnstile>invmode (mhd statM) e\<rightarrow>invC\<preceq>statT" 
       
   924      and        wf: "wf_prog G" 
       
   925      and      wt_e: "\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-RefT statT"
       
   926      and     statM: "statM \<in> mheads G accC statT sig" 
       
   927      and      invC: "invC = invocation_class (invmode (mhd statM) e) s a statT"
       
   928 )"\<exists> dynM. dynlookup G statT invC sig = Some dynM \<and>
       
   929   G\<turnstile>Methd sig dynM in invC dyn_accessible_from accC"
       
   930 proof -
       
   931   from wt_e wf have type_statT: "is_type G (RefT statT)"
       
   932     by (auto dest: ty_expr_is_type)
       
   933   from statM have not_Null: "statT \<noteq> NullT" by auto
       
   934   from type_statT wt_e 
       
   935   have wf_I: "(\<forall>I. statT = IfaceT I \<longrightarrow> is_iface G I \<and> 
       
   936                                         invmode (mhd statM) e \<noteq> SuperM)"
       
   937     by (auto dest: invocationTypeExpr_noClassD)
       
   938   from wt_e
       
   939   have wf_A: "(\<forall>     T. statT = ArrayT T \<longrightarrow> invmode (mhd statM) e \<noteq> SuperM)"
       
   940     by (auto dest: invocationTypeExpr_noClassD)
       
   941   show ?thesis
       
   942   proof (cases "invmode (mhd statM) e = IntVir")
       
   943     case True
       
   944     with invC_prop not_Null
       
   945     have invC_prop': "is_class G invC \<and>  
       
   946                       (if (\<exists>T. statT=ArrayT T) then invC=Object 
       
   947                                               else G\<turnstile>Class invC\<preceq>RefT statT)"
       
   948       by (auto simp add: DynT_prop_def)
       
   949     with True not_Null
       
   950     have "G,statT \<turnstile> invC valid_lookup_cls_for is_static statM"
       
   951      by (cases statT) (auto simp add: invmode_def 
       
   952                          split: split_if split_if_asm) (*  was deleted above *)
       
   953     with statM type_statT wf 
       
   954     show ?thesis
       
   955       by - (rule dynlookup_access,auto)
       
   956   next
       
   957     case False
       
   958     with type_statT wf invC not_Null wf_I wf_A
       
   959     have invC_prop': " is_class G invC \<and>
       
   960                       ((\<exists> statC. statT=ClassT statC \<and> invC=statC) \<or>
       
   961                       (\<forall> statC. statT\<noteq>ClassT statC \<and> invC=Object)) "
       
   962         by (case_tac "statT") (auto simp add: invocation_class_def 
       
   963                                        split: inv_mode.splits)
       
   964     with not_Null wf
       
   965     have dynlookup_static: "dynlookup G statT invC sig = methd G invC sig"
       
   966       by (case_tac "statT") (auto simp add: dynlookup_def dynmethd_C_C
       
   967                                             dynimethd_def)
       
   968    from statM wf wt_e not_Null False invC_prop' obtain dynM where
       
   969                 "accmethd G accC invC sig = Some dynM" 
       
   970      by (auto dest!: static_mheadsD)
       
   971    from invC_prop' False not_Null wf_I
       
   972    have "G,statT \<turnstile> invC valid_lookup_cls_for is_static statM"
       
   973      by (cases statT) (auto simp add: invmode_def
       
   974                         split: split_if split_if_asm) (*  was deleted above *)
       
   975    with statM type_statT wf 
       
   976     show ?thesis
       
   977       by - (rule dynlookup_access,auto)
       
   978   qed
       
   979 qed
       
   980 
       
   981 section "main proof of type safety"
       
   982 
       
   983 ML {*
       
   984 val forward_hyp_tac = EVERY' [smp_tac 1,
       
   985 	FIRST'[mp_tac,etac exI,smp_tac 2,smp_tac 1,EVERY'[etac impE,etac exI]],
       
   986 	REPEAT o (etac conjE)];
       
   987 val typD_tac = eresolve_tac (thms "wt_elim_cases") THEN_ALL_NEW 
       
   988 	EVERY' [full_simp_tac (simpset() setloop (K no_tac)), 
       
   989          clarify_tac(claset() addSEs[])]
       
   990 *}
       
   991 
       
   992 lemma conforms_locals [rule_format]: 
       
   993   "(a,b)\<Colon>\<preceq>(G, L) \<longrightarrow> L x = Some T \<longrightarrow> G,b\<turnstile>the (locals b x)\<Colon>\<preceq>T"
       
   994 apply (force simp: conforms_def Let_def lconf_def)
       
   995 done
       
   996 
       
   997 lemma eval_type_sound [rule_format (no_asm)]: 
       
   998  "wf_prog G \<Longrightarrow> G\<turnstile>s0 \<midarrow>t\<succ>\<rightarrow> (v,s1) \<Longrightarrow> (\<forall>L. s0\<Colon>\<preceq>(G,L) \<longrightarrow>    
       
   999   (\<forall>C T. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T \<longrightarrow> s1\<Colon>\<preceq>(G,L) \<and>  
       
  1000   (let (x,s) = s1 in x = None \<longrightarrow> G,L,s\<turnstile>t\<succ>v\<Colon>\<preceq>T)))"
       
  1001 apply (erule eval_induct)
       
  1002 
       
  1003 (* 29 subgoals *)
       
  1004 (* Xcpt, Inst, Methd, Nil, Skip, Expr, Comp *)
       
  1005 apply         (simp_all (no_asm_use) add: Let_def body_def)
       
  1006 apply       (tactic "ALLGOALS (EVERY'[Clarify_tac, TRY o typD_tac, 
       
  1007                      TRY o forward_hyp_tac])")
       
  1008 apply      (tactic"ALLGOALS(EVERY'[asm_simp_tac(simpset()),TRY o Clarify_tac])")
       
  1009 
       
  1010 (* 20 subgoals *)
       
  1011 
       
  1012 (* Break *)
       
  1013 apply (erule conforms_absorb)
       
  1014 
       
  1015 (* Cons *)
       
  1016 apply (erule_tac V = "G\<turnstile>Norm s0 \<midarrow>?ea\<succ>\<rightarrow> ?vs1" in thin_rl)
       
  1017 apply (frule eval_gext')
       
  1018 apply force
       
  1019 
       
  1020 (* LVar *)
       
  1021 apply (force elim: conforms_localD [THEN lconfD] conforms_lupd 
       
  1022        simp add: assign_conforms_def lvar_def)
       
  1023 
       
  1024 (* Cast *)
       
  1025 apply (force dest: fits_conf)
       
  1026 
       
  1027 (* Lit *)
       
  1028 apply (rule conf_litval)
       
  1029 apply (simp add: empty_dt_def)
       
  1030 
       
  1031 (* Super *)
       
  1032 apply (rule conf_widen)
       
  1033 apply   (erule (1) subcls_direct [THEN widen.subcls])
       
  1034 apply  (erule (1) conforms_localD [THEN lconfD])
       
  1035 apply (erule wf_ws_prog)
       
  1036 
       
  1037 (* Acc *)
       
  1038 apply fast
       
  1039 
       
  1040 (* Body *)
       
  1041 apply (rule conjI)
       
  1042 apply (rule conforms_absorb)
       
  1043 apply (fast)
       
  1044 apply (fast intro: conforms_locals)
       
  1045 
       
  1046 (* Cond *)
       
  1047 apply (simp split: split_if_asm)
       
  1048 apply  (tactic "forward_hyp_tac 1", force)
       
  1049 apply (tactic "forward_hyp_tac 1", force)
       
  1050 
       
  1051 (* If *)
       
  1052 apply (force split add: split_if_asm)
       
  1053 
       
  1054 (* Loop *)
       
  1055 apply (drule (1) wt.Loop)
       
  1056 apply (clarsimp split: split_if_asm)
       
  1057 apply (fast intro: conforms_absorb)
       
  1058 
       
  1059 (* Fin *)
       
  1060 apply (case_tac "x1", force)
       
  1061 apply (drule spec, erule impE, erule conforms_NormI)
       
  1062 apply (erule impE)
       
  1063 apply   blast
       
  1064 apply (clarsimp)
       
  1065 apply (erule (3) Fin_lemma)
       
  1066 
       
  1067 (* Throw *)
       
  1068 apply (erule (3) Throw_lemma)
       
  1069 
       
  1070 (* NewC *)
       
  1071 apply (clarsimp simp add: is_acc_class_def)
       
  1072 apply (drule (1) halloc_type_sound,blast, rule HOL.refl, simp, simp)
       
  1073 
       
  1074 (* NewA *)
       
  1075 apply (tactic "smp_tac 1 1",frule wt_init_comp_ty,erule impE,blast)
       
  1076 apply (tactic "forward_hyp_tac 1")
       
  1077 apply (case_tac "check_neg i' ab")
       
  1078 apply  (clarsimp simp add: is_acc_type_def)
       
  1079 apply  (drule (2) halloc_type_sound, rule HOL.refl, simp, simp)
       
  1080 apply force
       
  1081 
       
  1082 (* Level 34, 6 subgoals *)
       
  1083 
       
  1084 (* Init *)
       
  1085 apply (case_tac "inited C (globs s0)")
       
  1086 apply  (clarsimp)
       
  1087 apply (clarsimp)
       
  1088 apply (frule (1) wf_prog_cdecl)
       
  1089 apply (drule spec, erule impE, erule (3) conforms_init_class_obj)
       
  1090 apply (drule_tac "psi" = "class G C = ?x" in asm_rl,erule impE,
       
  1091       force dest!: wf_cdecl_supD split add: split_if simp add: is_acc_class_def)
       
  1092 apply (drule spec, erule impE, erule conforms_set_locals, rule lconf_empty)
       
  1093 apply (erule impE) apply (rule exI) apply (erule wf_cdecl_wt_init)
       
  1094 apply (drule (1) conforms_return, force dest: eval_gext', assumption)
       
  1095 
       
  1096 
       
  1097 (* Ass *)
       
  1098 apply (tactic "forward_hyp_tac 1")
       
  1099 apply (rename_tac x1 s1 x2 s2 v va w L C Ta T', case_tac x1)
       
  1100 prefer 2 apply force
       
  1101 apply (case_tac x2)
       
  1102 prefer 2 apply force
       
  1103 apply (simp, drule conjunct2)
       
  1104 apply (drule (1) conf_widen)
       
  1105 apply  (erule wf_ws_prog)
       
  1106 apply (erule (2) Ass_lemma)
       
  1107 apply (clarsimp simp add: assign_conforms_def)
       
  1108 
       
  1109 (* Try *)
       
  1110 apply (drule (1) sxalloc_type_sound, simp (no_asm_use))
       
  1111 apply (case_tac a)
       
  1112 apply  clarsimp
       
  1113 apply clarsimp
       
  1114 apply (tactic "smp_tac 1 1")
       
  1115 apply (simp split add: split_if_asm)
       
  1116 apply (fast dest: conforms_deallocL Try_lemma)
       
  1117 
       
  1118 (* FVar *)
       
  1119 
       
  1120 apply (frule accfield_fields)
       
  1121 apply (frule ty_expr_is_type [THEN type_is_class],simp)
       
  1122 apply simp
       
  1123 apply (frule wf_ws_prog)
       
  1124 apply (frule (1) fields_declC,simp)
       
  1125 apply clarsimp 
       
  1126 (*b y EVERY'[datac cfield_defpl_is_class 2, Clarsimp_tac] 1; not useful here*)
       
  1127 apply (tactic "smp_tac 1 1")
       
  1128 apply (tactic "forward_hyp_tac 1")
       
  1129 apply (rule conjI, force split add: split_if simp add: fvar_def2)
       
  1130 apply (drule init_yields_initd, frule eval_gext')
       
  1131 apply clarsimp
       
  1132 apply (case_tac "C=Object")
       
  1133 apply  clarsimp
       
  1134 apply (erule (9) FVar_lemma)
       
  1135 
       
  1136 (* AVar *)
       
  1137 apply (tactic "forward_hyp_tac 1")
       
  1138 apply (erule_tac V = "G\<turnstile>Norm s0 \<midarrow>?e1-\<succ>?a'\<rightarrow> (?x1 1, ?s1)" in thin_rl, 
       
  1139          frule eval_gext')
       
  1140 apply (rule conjI)
       
  1141 apply  (clarsimp simp add: avar_def2)
       
  1142 apply clarsimp
       
  1143 apply (erule (5) AVar_lemma)
       
  1144 
       
  1145 (* Call *)
       
  1146 apply (tactic "forward_hyp_tac 1")
       
  1147 apply (rule Call_type_sound)
       
  1148 apply auto
       
  1149 done
       
  1150 
       
  1151 
       
  1152 declare fun_upd_apply [simp]
       
  1153 declare split_paired_All [simp] split_paired_Ex [simp]
       
  1154 declare split_if     [split] split_if_asm     [split] 
       
  1155         option.split [split] option.split_asm [split]
       
  1156 ML_setup {* 
       
  1157 simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac);
       
  1158 claset_ref()  := claset () addSbefore ("split_all_tac", split_all_tac)
       
  1159 *}
       
  1160 
       
  1161 theorem eval_ts: 
       
  1162  "\<lbrakk>G\<turnstile>s \<midarrow>e-\<succ>v \<rightarrow> (x',s'); wf_prog G; s\<Colon>\<preceq>(G,L); \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>e\<Colon>-T\<rbrakk> 
       
  1163 \<Longrightarrow>  (x',s')\<Colon>\<preceq>(G,L) \<and> (x'=None \<longrightarrow> G,s'\<turnstile>v\<Colon>\<preceq>T)"
       
  1164 apply (drule (3) eval_type_sound)
       
  1165 apply (unfold Let_def)
       
  1166 apply clarsimp
       
  1167 done
       
  1168 
       
  1169 theorem evals_ts: 
       
  1170 "\<lbrakk>G\<turnstile>s \<midarrow>es\<doteq>\<succ>vs\<rightarrow> (x',s'); wf_prog G; s\<Colon>\<preceq>(G,L); \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>es\<Colon>\<doteq>Ts\<rbrakk> 
       
  1171 \<Longrightarrow>  (x',s')\<Colon>\<preceq>(G,L) \<and> (x'=None \<longrightarrow> list_all2 (conf G s') vs Ts)"
       
  1172 apply (drule (3) eval_type_sound)
       
  1173 apply (unfold Let_def)
       
  1174 apply clarsimp
       
  1175 done
       
  1176 
       
  1177 theorem evar_ts: 
       
  1178 "\<lbrakk>G\<turnstile>s \<midarrow>v=\<succ>vf\<rightarrow> (x',s'); wf_prog G; s\<Colon>\<preceq>(G,L); \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>v\<Colon>=T\<rbrakk> \<Longrightarrow>  
       
  1179   (x',s')\<Colon>\<preceq>(G,L) \<and> (x'=None \<longrightarrow> G,L,s'\<turnstile>In2 v\<succ>In2 vf\<Colon>\<preceq>Inl T)"
       
  1180 apply (drule (3) eval_type_sound)
       
  1181 apply (unfold Let_def)
       
  1182 apply clarsimp
       
  1183 done
       
  1184 
       
  1185 theorem exec_ts: 
       
  1186 "\<lbrakk>G\<turnstile>s \<midarrow>s0\<rightarrow> s'; wf_prog G; s\<Colon>\<preceq>(G,L); \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>s0\<Colon>\<surd>\<rbrakk> \<Longrightarrow> s'\<Colon>\<preceq>(G,L)"
       
  1187 apply (drule (3) eval_type_sound)
       
  1188 apply (unfold Let_def)
       
  1189 apply clarsimp
       
  1190 done
       
  1191 
       
  1192 (*
       
  1193 theorem dyn_methods_understood: 
       
  1194  "\<And>s. \<lbrakk>wf_prog G; \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>{t,md,IntVir}e..mn({pTs'}ps)\<Colon>-rT;  
       
  1195   s\<Colon>\<preceq>(G,L); G\<turnstile>s \<midarrow>e-\<succ>a'\<rightarrow> Norm s'; a' \<noteq> Null\<rbrakk> \<Longrightarrow>  
       
  1196   \<exists>a obj. a'=Addr a \<and> heap s' a = Some obj \<and> 
       
  1197   cmethd G (obj_class obj) (mn, pTs') \<noteq> None"
       
  1198 apply (erule wt_elim_cases)
       
  1199 apply (drule max_spec2mheads)
       
  1200 apply (drule (3) eval_ts)
       
  1201 apply (clarsimp split del: split_if split_if_asm)
       
  1202 apply (drule (2) DynT_propI)
       
  1203 apply  (simp (no_asm_simp))
       
  1204 apply (tactic *) (* {* exhaust_cmethd_tac "the (cmethd G (target (invmode m e) s' a' md) (mn, pTs'))" 1 *} *)(*)
       
  1205 apply (drule (4) DynT_mheadsD [THEN conjunct1], rule HOL.refl)
       
  1206 apply (drule conf_RefTD)
       
  1207 apply clarsimp
       
  1208 done 
       
  1209 *)
       
  1210 
       
  1211 end