src/HOL/MicroJava/J/JTypeSafe.ML
changeset 8011 d14c4e9e9c8e
child 8034 6fc37b5c5e98
equal deleted inserted replaced
8010:69032a618aa9 8011:d14c4e9e9c8e
       
     1 (*  Title:      HOL/MicroJava/J/JTypeSafe.ML
       
     2     ID:         $Id$
       
     3     Author:     David von Oheimb
       
     4     Copyright   1999 Technische Universitaet Muenchen
       
     5 
       
     6 Type safety proof
       
     7 *)
       
     8 
       
     9 Addsimps [split_beta];
       
    10 
       
    11 Goal "\\<lbrakk>h a = None; (h, l)\\<Colon>\\<preceq>(G, lT); wf_prog wtm G; is_class G C\\<rbrakk> \\<Longrightarrow> \
       
    12 \ (h(a\\<mapsto>(C,(init_vars (fields (G,C))))), l)\\<Colon>\\<preceq>(G, lT)";
       
    13 by( etac conforms_upd_obj 1);
       
    14 by(  rewtac oconf_def);
       
    15 by(  auto_tac (claset() addSDs [is_type_fields, map_of_SomeD], simpset()));
       
    16 qed "NewC_conforms";
       
    17 
       
    18 Goalw [cast_ok_def]
       
    19  "\\<lbrakk> wf_prog wtm G; G,h\\<turnstile>v\\<Colon>\\<preceq>Tb; G\\<turnstile>Tb\\<Rightarrow>? T'; cast_ok G T' h v\\<rbrakk> \
       
    20 \ \\<Longrightarrow> G,h\\<turnstile>v\\<Colon>\\<preceq>T'";
       
    21 by( case_tac1 "v = Null" 1);
       
    22 by(  Asm_full_simp_tac 1);
       
    23 by(  dtac widen_RefT 1);
       
    24 by(  Clarify_tac 1);
       
    25 by(  forward_tac [cast_RefT] 1);
       
    26 by(  Clarify_tac 1);
       
    27 by(  rtac widen.null 1);
       
    28 by( case_tac1 "\\<exists>pt. T' = PrimT pt" 1);
       
    29 by(  strip_tac1 1);
       
    30 by(  dtac cast_PrimT2 1);
       
    31 by(  etac conf_widen 1);
       
    32 by(   atac 1);
       
    33 by(  atac 1);
       
    34 by( Asm_full_simp_tac 1);
       
    35 by( dtac (non_PrimT RS iffD1) 1);
       
    36 by( strip_tac1 1);
       
    37 by( forward_tac [cast_RefT2] 1);
       
    38 by( strip_tac1 1);
       
    39 by( dtac non_npD 1);
       
    40 by(  atac 1);
       
    41 by( rewrite_goals_tac [the_Addr_def,obj_ty_def]);
       
    42 by Auto_tac ;
       
    43 by(  ALLGOALS (rtac conf_AddrI));
       
    44 by     Auto_tac;
       
    45 qed "Cast_conf";
       
    46 
       
    47 Goal "\\<lbrakk> wf_prog wtm G; cfield (G,C) fn = Some (fd, ft); (h,l)\\<Colon>\\<preceq>(G,lT); \
       
    48 \    x' = None \\<longrightarrow> G,h\\<turnstile>a'\\<Colon>\\<preceq>Class C; np a' x' = None \\<rbrakk> \\<Longrightarrow> \
       
    49 \ G,h\\<turnstile>the (snd (the (h (the_Addr a'))) (fn, fd))\\<Colon>\\<preceq>ft";
       
    50 by( dtac np_NoneD 1);
       
    51 by( etac conjE 1);
       
    52 by( mp_tac 1);
       
    53 by( dtac non_np_objD 1);
       
    54 by   Auto_tac;
       
    55 by( dtac (conforms_heapD RS hconfD) 1);
       
    56 by(  atac 1);
       
    57 by( dtac widen_cfs_fields 1);
       
    58 by(   atac 1);
       
    59 by(  atac 1);
       
    60 by( dtac oconf_objD 1);
       
    61 by(  atac 1);
       
    62 by Auto_tac;
       
    63 val FAcc_type_sound = result();
       
    64 
       
    65 Goal
       
    66  "\\<lbrakk> wf_prog wtm G; a = the_Addr a'; (c, fs) = the (h a); \
       
    67 \   (G, lT)\\<turnstile>v\\<Colon>T'; G\\<turnstile>T'\\<preceq>ft; \
       
    68 \   (G, lT)\\<turnstile>aa\\<Colon>Class C; \
       
    69 \   cfield (G,C) fn = Some (fd, ft); h''\\<le>|h'; \
       
    70 \   x' = None \\<longrightarrow> G,h'\\<turnstile>a'\\<Colon>\\<preceq>Class C; h'\\<le>|h; \
       
    71 \   (h, l)\\<Colon>\\<preceq>(G, lT); G,h\\<turnstile>x\\<Colon>\\<preceq>T'; np a' x' = None\\<rbrakk> \\<Longrightarrow> \
       
    72 \ h''\\<le>|h(a\\<mapsto>(c,(fs((fn,fd)\\<mapsto>x)))) \\<and>  \
       
    73 \ (h(a\\<mapsto>(c,(fs((fn,fd)\\<mapsto>x)))), l)\\<Colon>\\<preceq>(G, lT) \\<and>  \
       
    74 \ G,h(a\\<mapsto>(c,(fs((fn,fd)\\<mapsto>x))))\\<turnstile>x\\<Colon>\\<preceq>T'";
       
    75 by( dtac np_NoneD 1);
       
    76 by( etac conjE 1);
       
    77 by( Asm_full_simp_tac 1);
       
    78 by( dtac non_np_objD 1);
       
    79 by(   atac 1);
       
    80 by(  SELECT_GOAL Auto_tac 1);
       
    81 by( strip_tac1 1);
       
    82 by( Full_simp_tac 1);
       
    83 by( EVERY [forward_tac [hext_objD] 1, atac 1]);
       
    84 by( etac exE 1);
       
    85 by( Asm_full_simp_tac 1);
       
    86 by( strip_tac1 1);
       
    87 by( rtac conjI 1);
       
    88 by(  fast_tac (HOL_cs addEs [hext_trans, hext_upd_obj]) 1);
       
    89 by( rtac conjI 1);
       
    90 by(  fast_tac (HOL_cs addEs [conf_upd_obj RS iffD2]) 2);
       
    91 
       
    92 by( rtac conforms_upd_obj 1);
       
    93 by   Auto_tac;
       
    94 by(  rtac hextI 2);
       
    95 by(  Force_tac 2);
       
    96 by( rtac oconf_hext 1);
       
    97 by(  etac hext_upd_obj 2);
       
    98 by( dtac widen_cfs_fields 1);
       
    99 by(   atac 1);
       
   100 by(  atac 1);
       
   101 by( rtac (oconf_obj RS iffD2) 1);
       
   102 by( Simp_tac 1);
       
   103 by( strip_tac 1);
       
   104 by( case_tac1 "(aaa, b) = (fn, fd)" 1);
       
   105 by(  Asm_full_simp_tac 1);
       
   106 by(  fast_tac (HOL_cs addIs [conf_widen]) 1);
       
   107 by( fast_tac (HOL_cs addDs [conforms_heapD RS hconfD, oconf_objD]) 1);
       
   108 val FAss_type_sound = result();
       
   109 
       
   110 Goalw [wf_mhead_def] "\\<lbrakk> wf_prog wtm G; list_all2 (conf G h) pvs pTs; \
       
   111  \ list_all2 (\\<lambda>T T'. G\\<turnstile>T\\<preceq>T') pTs pTs'; wf_mhead G (mn,pTs') rT; \
       
   112 \ length pTs' = length pns; nodups pns; \
       
   113 \ Ball (set lvars) (split (\\<lambda>vn. is_type G)) \
       
   114 \ \\<rbrakk> \\<Longrightarrow> G,h\\<turnstile>init_vars lvars(pns[\\<mapsto>]pvs)[\\<Colon>\\<preceq>]map_of lvars(pns[\\<mapsto>]pTs')";
       
   115 by( Clarsimp_tac 1);
       
   116 by( rtac lconf_ext_list 1);
       
   117 by(    rtac (Ball_set_table RS lconf_init_vars) 1);
       
   118 by(    Force_tac 1);
       
   119 by(   atac 1);
       
   120 by(  atac 1);
       
   121 by( (etac conf_list_gext_widen THEN_ALL_NEW atac) 1);
       
   122 qed "Call_lemma2";
       
   123 
       
   124 Goalw [wf_java_prog_def]
       
   125  "\\<lbrakk> wf_java_prog G; a' \\<noteq> Null; (h, l)\\<Colon>\\<preceq>(G, lT); \
       
   126 \    max_spec G T (mn,pTsa) = {((mda,rTa),pTs')}; xc\\<le>|xh; xh\\<le>|h; \
       
   127 \    list_all2 (conf G h) pvs pTsa;\
       
   128 \    (md, rT, pns, lvars, blk, res) = \
       
   129 \              the (cmethd (G,fst (the (h (the_Addr a')))) (mn, pTs'));\
       
   130 \ \\<forall>lT. (h, init_vars lvars(pns[\\<mapsto>]pvs)(This\\<mapsto>a'))\\<Colon>\\<preceq>(G, lT) \\<longrightarrow> \
       
   131 \ (G, lT)\\<turnstile>blk\\<surd> \\<longrightarrow>  h\\<le>|xi \\<and>  (xi, xl)\\<Colon>\\<preceq>(G, lT); \
       
   132 \ \\<forall>lT. (xi, xl)\\<Colon>\\<preceq>(G, lT) \\<longrightarrow> (\\<forall>T. (G, lT)\\<turnstile>res\\<Colon>T \\<longrightarrow> \
       
   133 \         xi\\<le>|h' \\<and> (h', xj)\\<Colon>\\<preceq>(G, lT) \\<and> (x' = None \\<longrightarrow> G,h'\\<turnstile>v\\<Colon>\\<preceq>T)); \
       
   134 \ G,xh\\<turnstile>a'\\<Colon>\\<preceq>RefT T \\<rbrakk> \\<Longrightarrow> \
       
   135 \ xc\\<le>|h' \\<and> (h', l)\\<Colon>\\<preceq>(G, lT) \\<and>  (x' = None \\<longrightarrow> G,h'\\<turnstile>v\\<Colon>\\<preceq>rTa)";
       
   136 by( dtac (insertI1 RSN (2,(equalityD2 RS subsetD))) 1);
       
   137 by( dtac (max_spec2appl_meths RS appl_methsD) 1);
       
   138 by( etac conjE 1);
       
   139 by( dtac non_np_objD' 1);
       
   140 by(    atac 1);
       
   141 by(   atac 1);
       
   142 by(  strip_tac1 1);
       
   143 by(  Asm_full_simp_tac 1);
       
   144 by( Clarsimp_tac 1);
       
   145 by( EVERY'[forward_tac [hext_objD], atac] 1);
       
   146 by( Clarsimp_tac 1);
       
   147 by( EVERY'[dtac Call_lemma, atac, atac] 1);
       
   148 by( clarsimp_tac (claset(),simpset() addsimps [wt_java_mdecl_def])1);
       
   149 by( thin_tac "cmethd ?sig ?x = ?y" 1);
       
   150 by( EVERY'[dtac spec, etac impE] 1);
       
   151 by(  mp_tac 2);
       
   152 by(  rtac conformsI 1);
       
   153 by(   etac conforms_heapD 1);
       
   154 by(  rtac lconf_ext 1);
       
   155 by(   force_tac (claset() addSEs [Call_lemma2],simpset()) 1);
       
   156 by(  EVERY'[etac conf_hext, etac conf_obj_AddrI, atac] 1);
       
   157 by( thin_tac "?E\\<turnstile>?blk\\<surd>" 1);
       
   158 by( etac conjE 1);
       
   159 by( EVERY'[dtac spec, mp_tac] 1);
       
   160 (*by( thin_tac "?E\\<Colon>\\<preceq>(G, pT')" 1);*)
       
   161 by( EVERY'[dtac spec, mp_tac] 1);
       
   162 by( thin_tac "?E\\<turnstile>res\\<Colon>?rT" 1);
       
   163 by( strip_tac1 1);
       
   164 by( rtac conjI 1);
       
   165 by(  fast_tac (HOL_cs addIs [hext_trans]) 1);
       
   166 by( rtac conjI 1);
       
   167 by(  rtac impI 2);
       
   168 by(  mp_tac 2);
       
   169 by(  forward_tac [conf_widen] 2);
       
   170 by(    atac 4);
       
   171 by(   atac 2);
       
   172 by(  fast_tac (HOL_cs addSEs [widen_trans]) 2);
       
   173 by( etac conforms_hext 1);
       
   174 by(  etac hext_trans 1);
       
   175 by(  atac 1);
       
   176 by( etac conforms_heapD 1);
       
   177 qed "Call_type_sound";
       
   178 
       
   179 
       
   180 
       
   181 Unify.search_bound := 40;
       
   182 Unify.trace_bound  := 40;
       
   183 Delsplits[split_if];
       
   184 Delsimps[fun_upd_apply];(*###*)
       
   185 Goal
       
   186 "wf_java_prog G \\<Longrightarrow> \
       
   187 \ (G\\<turnstile>(x,(h,l)) -e  \\<succ>v  \\<rightarrow> (x', (h',l')) \\<longrightarrow> \
       
   188 \     (\\<forall>lT.    (h ,l )\\<Colon>\\<preceq>(G,lT) \\<longrightarrow> (\\<forall>T . (G,lT)\\<turnstile>e  \\<Colon> T \\<longrightarrow> \
       
   189 \     h\\<le>|h' \\<and> (h',l')\\<Colon>\\<preceq>(G,lT) \\<and> (x'=None \\<longrightarrow> G,h'\\<turnstile>v  \\<Colon>\\<preceq> T )))) \\<and> \
       
   190 \ (G\\<turnstile>(x,(h,l)) -es[\\<succ>]vs\\<rightarrow> (x', (h',l')) \\<longrightarrow> \
       
   191 \     (\\<forall>lT.    (h ,l )\\<Colon>\\<preceq>(G,lT) \\<longrightarrow> (\\<forall>Ts. (G,lT)\\<turnstile>es[\\<Colon>]Ts \\<longrightarrow> \
       
   192 \     h\\<le>|h' \\<and> (h',l')\\<Colon>\\<preceq>(G,lT) \\<and> (x'=None \\<longrightarrow> list_all2 (\\<lambda>v T. G,h'\\<turnstile>v\\<Colon>\\<preceq>T) vs Ts)))) \\<and> \
       
   193 \ (G\\<turnstile>(x,(h,l)) -c       \\<rightarrow> (x', (h',l')) \\<longrightarrow> \
       
   194 \     (\\<forall>lT.    (h ,l )\\<Colon>\\<preceq>(G,lT) \\<longrightarrow>       (G,lT)\\<turnstile>c  \\<surd> \\<longrightarrow> \
       
   195 \     h\\<le>|h' \\<and> (h',l')\\<Colon>\\<preceq>(G,lT)))";
       
   196 by( rtac eval_evals_exec.induct 1);
       
   197 by( rewtac c_hupd_def);
       
   198 
       
   199 (* several simplifications, XcptE, XcptEs, XcptS, Skip *)
       
   200 by( ALLGOALS Asm_full_simp_tac);
       
   201 by( ALLGOALS strip_tac);
       
   202 by( ALLGOALS (eresolve_tac ty_expr_ty_exprs_wt_stmt.elims 
       
   203 		THEN_ALL_NEW Full_simp_tac));
       
   204 by( ALLGOALS (EVERY' [REPEAT o (etac conjE), REPEAT o hyp_subst_tac]));
       
   205 by( rewtac wf_java_prog_def);
       
   206 
       
   207 (* Level 7 *)
       
   208 
       
   209 (* 14 NewC *)
       
   210 by( dtac new_AddrD 1);
       
   211 by( etac disjE 1);
       
   212 by(  Asm_simp_tac 2);
       
   213 by( etac conjE 1);
       
   214 by( Asm_simp_tac 1);
       
   215 by( rtac conjI 1);
       
   216 by(  fast_tac (HOL_cs addSEs [NewC_conforms]) 1);
       
   217 by( rtac conf_obj_AddrI 1);
       
   218 by(  rtac widen.refl 2);
       
   219 by(  Asm_simp_tac 2);
       
   220 by( Simp_tac 1);
       
   221 
       
   222 (* for Cast *)
       
   223 by( defer_tac 1);
       
   224 
       
   225 (* 13 Lit *)
       
   226 by( etac conf_litval 1);
       
   227 
       
   228 (* 12 LAcc *)
       
   229 by( fast_tac (claset() addEs [conforms_localD RS lconfD]) 1);
       
   230 
       
   231 (* 11 Nil *)
       
   232 by( Simp_tac 5);
       
   233 
       
   234 (* for FAss *)
       
   235 by( EVERY'[eresolve_tac ty_expr_ty_exprs_wt_stmt.elims THEN_ALL_NEW Full_simp_tac, 
       
   236 			REPEAT o (etac conjE), hyp_subst_tac] 3);
       
   237 
       
   238 (* for if *)
       
   239 by( (case_tac1 "the_Bool v" THEN_ALL_NEW Asm_full_simp_tac) 8);
       
   240 
       
   241 val forward_hyp_tac = ALLGOALS (TRY o (EVERY' [dtac spec, mp_tac,
       
   242 	(mp_tac ORELSE' (dtac spec THEN' mp_tac)), REPEAT o (etac conjE)]));
       
   243 by forward_hyp_tac;
       
   244 
       
   245 (* 10+1 if *)
       
   246 by(  fast_tac (HOL_cs addIs [hext_trans]) 8);
       
   247 by( fast_tac (HOL_cs addIs [hext_trans]) 8);
       
   248 
       
   249 (* 9 Expr *)
       
   250 by( Fast_tac 6);
       
   251 
       
   252 by( ALLGOALS Asm_full_simp_tac);
       
   253 
       
   254 (* 8 Cast *)
       
   255 by( EVERY'[rtac impI, dtac raise_if_NoneD, Clarsimp_tac, 
       
   256            fast_tac (claset() addEs [Cast_conf])] 8);
       
   257 
       
   258 (* 7 LAss *)
       
   259 by( asm_simp_tac (simpset() addsplits [expand_if]) 1);
       
   260 by( EVERY'[eresolve_tac ty_expr_ty_exprs_wt_stmt.elims THEN_ALL_NEW Full_simp_tac] 1);
       
   261 by( blast_tac (claset() addIs [conforms_upd_local, conf_widen]) 1);
       
   262 
       
   263 (* 6 FAcc *)
       
   264 by( fast_tac (claset() addSEs [FAcc_type_sound]) 1);
       
   265 
       
   266 (* 5 While *)
       
   267 by( fast_tac (claset() addIs [ty_expr_ty_exprs_wt_stmt.Cond, ty_expr_ty_exprs_wt_stmt.Comp, ty_expr_ty_exprs_wt_stmt.Skip]
       
   268 		       addEs [ty_expr_ty_exprs_wt_stmt.Loop]) 5);
       
   269 
       
   270 by forward_hyp_tac;
       
   271 
       
   272 (* 4 Cons *)
       
   273 by( fast_tac (claset() addDs [evals_no_xcpt] addIs [conf_hext,hext_trans]) 3);
       
   274 
       
   275 (* 3 ;; *)
       
   276 by( fast_tac (claset() addIs [hext_trans]) 3);
       
   277 
       
   278 (* 2 FAss *)
       
   279 by( case_tac1 "x2 = None" 1);
       
   280 by(  Asm_simp_tac 2);
       
   281 by(  fast_tac (claset() addIs [hext_trans]) 2);
       
   282 by( Asm_full_simp_tac 1);
       
   283 by( dtac eval_no_xcpt 1);
       
   284 by( SELECT_GOAL (etac FAss_type_sound 1 THEN rtac refl 1 THEN ALLGOALS atac) 1);
       
   285 
       
   286 by prune_params_tac;
       
   287 (* Level 45 *)
       
   288 
       
   289 (* 1 Call *)
       
   290 by( case_tac1 "x = None" 1);
       
   291 by(  dtac (not_None_eq RS iffD1) 2);
       
   292 by(  Clarsimp_tac 2);
       
   293 by(  dtac exec_xcpt 2);
       
   294 by(  Asm_full_simp_tac 2);
       
   295 by(  dtac eval_xcpt 2);
       
   296 by(  Asm_full_simp_tac 2);
       
   297 by(  fast_tac (HOL_cs addEs [hext_trans]) 2);
       
   298 by( Clarify_tac 1);
       
   299 by( dtac evals_no_xcpt 1);
       
   300 by( Asm_full_simp_tac 1);
       
   301 by( case_tac1 "a' = Null" 1);
       
   302 by(  Asm_full_simp_tac 1);
       
   303 by(  dtac exec_xcpt 1);
       
   304 by(  Asm_full_simp_tac 1);
       
   305 by(  dtac eval_xcpt 1);
       
   306 by(  Asm_full_simp_tac 1);
       
   307 by(  fast_tac (HOL_cs addEs [hext_trans]) 1);
       
   308 by( (rtac (rewrite_rule[wf_java_prog_def]Call_type_sound) THEN_ALL_NEW Asm_simp_tac) 1);
       
   309 qed "eval_evals_exec_type_sound";
       
   310 
       
   311 Goal "\\<And>E s s'. \
       
   312 \ \\<lbrakk> G=prg E; wf_java_prog G; G\\<turnstile>(x,s) -e\\<succ>v \\<rightarrow> (x',s'); s\\<Colon>\\<preceq>E; E\\<turnstile>e\\<Colon>T \\<rbrakk> \
       
   313 \ \\<Longrightarrow> s'\\<Colon>\\<preceq>E \\<and> (x'=None \\<longrightarrow> G,heap s'\\<turnstile>v\\<Colon>\\<preceq>T)";
       
   314 by( split_all_tac 1);
       
   315 bd (eval_evals_exec_type_sound RS conjunct1 RS mp RS spec RS mp) 1;
       
   316 by Auto_tac;
       
   317 qed "eval_type_sound";
       
   318 
       
   319 Goal "\\<And>E s s'. \
       
   320 \ \\<lbrakk> G=prg E; wf_java_prog G; G\\<turnstile>(x,s) -s0\\<rightarrow> (x',s'); s\\<Colon>\\<preceq>E; E\\<turnstile>s0\\<surd> \\<rbrakk> \
       
   321 \ \\<Longrightarrow> s'\\<Colon>\\<preceq>E";
       
   322 by( split_all_tac 1);
       
   323 bd (eval_evals_exec_type_sound RS conjunct2 RS conjunct2 RS mp RS spec RS mp) 1;
       
   324 by   Auto_tac;
       
   325 qed "exec_type_sound";
       
   326 
       
   327 Goal "\\<lbrakk>G=prg E; wf_java_prog G; G\\<turnstile>(x,s) -e\\<succ>a'\\<rightarrow> Norm s'; a' \\<noteq> Null;\
       
   328 \         s\\<Colon>\\<preceq>E; E\\<turnstile>e\\<Colon>RefT T; m_head G T sig \\<noteq> None\\<rbrakk> \\<Longrightarrow> \
       
   329 \ cmethd (G,fst (the (heap s' (the_Addr a')))) sig \\<noteq> None";
       
   330 by( dtac eval_type_sound 1);
       
   331 by(     atac 1);
       
   332 by(    atac 1);
       
   333 by(   atac 1);
       
   334 by(  atac 1);
       
   335 by( not_None_tac 1);
       
   336 by( split_all_tac 1);
       
   337 by(rewtac wf_java_prog_def);
       
   338 by( forward_tac [widen_methd] 1);
       
   339 by(   atac 1);
       
   340 by(  rtac (not_None_eq RS iffD1) 2);
       
   341 by(  Fast_tac 2);
       
   342 by( etac conjE 1);
       
   343 by( dtac non_npD 1);
       
   344 by Auto_tac;
       
   345 qed "all_methods_understood";
       
   346 
       
   347 Delsimps [split_beta];
       
   348 Addsimps[fun_upd_apply];(*###*)
       
   349