src/HOL/MicroJava/BV/StepMono.thy
changeset 12516 d09d0f160888
parent 12515 3fb416265ba9
child 12517 360e3215f029
equal deleted inserted replaced
12515:3fb416265ba9 12516:d09d0f160888
     1 (*  Title:      HOL/MicroJava/BV/StepMono.thy
       
     2     ID:         $Id$
       
     3     Author:     Gerwin Klein
       
     4     Copyright   2000 Technische Universitaet Muenchen
       
     5 *)
       
     6 
       
     7 header {* Monotonicity of step and app *}
       
     8 
       
     9 theory StepMono = Step:
       
    10 
       
    11 
       
    12 lemma PrimT_PrimT: "(G \<turnstile> xb \<preceq> PrimT p) = (xb = PrimT p)"
       
    13   by (auto elim: widen.elims)
       
    14 
       
    15 
       
    16 lemma sup_loc_some [rule_format]:
       
    17 "\<forall> y n. (G \<turnstile> b <=l y) --> n < length y --> y!n = OK t -->
       
    18   (\<exists>t. b!n = OK t \<and> (G \<turnstile> (b!n) <=o (y!n)))" (is "?P b")
       
    19 proof (induct (open) ?P b)
       
    20   show "?P []" by simp
       
    21 
       
    22   case Cons
       
    23   show "?P (a#list)"
       
    24   proof (clarsimp simp add: list_all2_Cons1 sup_loc_def Listn.le_def lesub_def)
       
    25     fix z zs n
       
    26     assume * :
       
    27       "G \<turnstile> a <=o z" "list_all2 (sup_ty_opt G) list zs"
       
    28       "n < Suc (length list)" "(z # zs) ! n = OK t"
       
    29 
       
    30     show "(\<exists>t. (a # list) ! n = OK t) \<and> G \<turnstile>(a # list) ! n <=o OK t"
       
    31     proof (cases n)
       
    32       case 0
       
    33       with * show ?thesis by (simp add: sup_ty_opt_OK)
       
    34     next
       
    35       case Suc
       
    36       with Cons *
       
    37       show ?thesis by (simp add: sup_loc_def Listn.le_def lesub_def)
       
    38     qed
       
    39   qed
       
    40 qed
       
    41 
       
    42 
       
    43 lemma all_widen_is_sup_loc:
       
    44 "\<forall>b. length a = length b -->
       
    45      (\<forall>x\<in>set (zip a b). x \<in> widen G) = (G \<turnstile> (map OK a) <=l (map OK b))"
       
    46  (is "\<forall>b. length a = length b --> ?Q a b" is "?P a")
       
    47 proof (induct "a")
       
    48   show "?P []" by simp
       
    49 
       
    50   fix l ls assume Cons: "?P ls"
       
    51 
       
    52   show "?P (l#ls)"
       
    53   proof (intro allI impI)
       
    54     fix b
       
    55     assume "length (l # ls) = length (b::ty list)"
       
    56     with Cons
       
    57     show "?Q (l # ls) b" by - (cases b, auto)
       
    58   qed
       
    59 qed
       
    60 
       
    61 
       
    62 lemma append_length_n [rule_format]:
       
    63 "\<forall>n. n \<le> length x --> (\<exists>a b. x = a@b \<and> length a = n)" (is "?P x")
       
    64 proof (induct (open) ?P x)
       
    65   show "?P []" by simp
       
    66 
       
    67   fix l ls assume Cons: "?P ls"
       
    68 
       
    69   show "?P (l#ls)"
       
    70   proof (intro allI impI)
       
    71     fix n
       
    72     assume l: "n \<le> length (l # ls)"
       
    73 
       
    74     show "\<exists>a b. l # ls = a @ b \<and> length a = n"
       
    75     proof (cases n)
       
    76       assume "n=0" thus ?thesis by simp
       
    77     next
       
    78       fix "n'" assume s: "n = Suc n'"
       
    79       with l
       
    80       have  "n' \<le> length ls" by simp
       
    81       hence "\<exists>a b. ls = a @ b \<and> length a = n'" by (rule Cons [rule_format])
       
    82       then obtain a b where "ls = a @ b" "length a = n'" by rules
       
    83       with s have "l # ls = (l#a) @ b \<and> length (l#a) = n" by simp
       
    84       thus ?thesis by blast
       
    85     qed
       
    86   qed
       
    87 qed
       
    88 
       
    89 
       
    90 lemma rev_append_cons:
       
    91 "[|n < length x|] ==> \<exists>a b c. x = (rev a) @ b # c \<and> length a = n"
       
    92 proof -
       
    93   assume n: "n < length x"
       
    94   hence "n \<le> length x" by simp
       
    95   hence "\<exists>a b. x = a @ b \<and> length a = n" by (rule append_length_n)
       
    96   then obtain r d where x: "x = r@d" "length r = n" by rules
       
    97   with n have "\<exists>b c. d = b#c" by (simp add: neq_Nil_conv)
       
    98   then obtain b c where "d = b#c" by rules
       
    99   with x have "x = (rev (rev r)) @ b # c \<and> length (rev r) = n" by simp
       
   100   thus ?thesis by blast
       
   101 qed
       
   102 
       
   103 lemmas [iff] = not_Err_eq
       
   104 
       
   105 lemma app_mono:
       
   106 "[|G \<turnstile> s <=' s'; app i G m rT s'|] ==> app i G m rT s"
       
   107 proof -
       
   108 
       
   109   { fix s1 s2
       
   110     assume G:   "G \<turnstile> s2 <=s s1"
       
   111     assume app: "app i G m rT (Some s1)"
       
   112 
       
   113     have "app i G m rT (Some s2)"
       
   114     proof (cases (open) i)
       
   115       case Load
       
   116 
       
   117       from G Load app
       
   118       have "G \<turnstile> snd s2 <=l snd s1" by (auto simp add: sup_state_conv)
       
   119 
       
   120       with G Load app
       
   121       show ?thesis
       
   122         by (auto dest: sup_loc_some)
       
   123     next
       
   124       case Store
       
   125       with G app
       
   126       show ?thesis
       
   127         by (cases s2) (auto simp add: map_eq_Cons sup_loc_Cons2 sup_loc_length sup_state_conv)
       
   128     next
       
   129       case LitPush
       
   130       with G app
       
   131       show ?thesis by simp
       
   132     next
       
   133       case New
       
   134       with G app
       
   135       show ?thesis by simp
       
   136     next
       
   137       case Getfield
       
   138       with app G
       
   139       show ?thesis
       
   140         by (cases s2) (auto simp add: sup_state_Cons2, rule widen_trans)
       
   141     next
       
   142       case Putfield
       
   143 
       
   144       with app
       
   145       obtain vT oT ST LT b
       
   146         where s1: "s1 = (vT # oT # ST, LT)" and
       
   147                   "field (G, cname) vname = Some (cname, b)"
       
   148                   "is_class G cname" and
       
   149               oT: "G\<turnstile> oT\<preceq> (Class cname)" and
       
   150               vT: "G\<turnstile> vT\<preceq> b"
       
   151         by force
       
   152       moreover
       
   153       from s1 G
       
   154       obtain vT' oT' ST' LT'
       
   155         where s2:  "s2 = (vT' # oT' # ST', LT')" and
       
   156               oT': "G\<turnstile> oT' \<preceq> oT" and
       
   157               vT': "G\<turnstile> vT' \<preceq> vT"
       
   158         by (cases s2) (auto simp add: sup_state_Cons2)
       
   159       moreover
       
   160       from vT' vT
       
   161       have "G \<turnstile> vT' \<preceq> b" by (rule widen_trans)
       
   162       moreover
       
   163       from oT' oT
       
   164       have "G\<turnstile> oT' \<preceq> (Class cname)" by (rule widen_trans)
       
   165       ultimately
       
   166       show ?thesis
       
   167         by (auto simp add: Putfield)
       
   168     next
       
   169       case Checkcast
       
   170       with app G
       
   171       show ?thesis
       
   172         by (cases s2) (auto intro!: widen_RefT2 simp add: sup_state_Cons2)
       
   173     next
       
   174       case Return
       
   175       with app G
       
   176       show ?thesis
       
   177         by (cases s2) (auto simp add: sup_state_Cons2, rule widen_trans)
       
   178     next
       
   179       case Pop
       
   180       with app G
       
   181       show ?thesis
       
   182         by (cases s2) (auto simp add: sup_state_Cons2)
       
   183     next
       
   184       case Dup
       
   185       with app G
       
   186       show ?thesis
       
   187         by (cases s2) (clarsimp simp add: sup_state_Cons2,
       
   188               auto dest: sup_state_length)
       
   189     next
       
   190       case Dup_x1
       
   191       with app G
       
   192       show ?thesis
       
   193         by (cases s2) (clarsimp simp add: sup_state_Cons2,
       
   194               auto dest: sup_state_length)
       
   195     next
       
   196       case Dup_x2
       
   197       with app G
       
   198       show ?thesis
       
   199         by (cases s2) (clarsimp simp add: sup_state_Cons2,
       
   200               auto dest: sup_state_length)
       
   201     next
       
   202       case Swap
       
   203       with app G
       
   204       show ?thesis
       
   205         by (cases s2) (clarsimp simp add: sup_state_Cons2)
       
   206     next
       
   207       case IAdd
       
   208       with app G
       
   209       show ?thesis
       
   210         by (cases s2) (auto simp add: sup_state_Cons2 PrimT_PrimT)
       
   211     next
       
   212       case Goto
       
   213       with app
       
   214       show ?thesis by simp
       
   215     next
       
   216       case Ifcmpeq
       
   217       with app G
       
   218       show ?thesis
       
   219         by (cases s2) (auto simp add: sup_state_Cons2 PrimT_PrimT widen_RefT2)
       
   220     next
       
   221       case Invoke
       
   222 
       
   223       with app
       
   224       obtain apTs X ST LT mD' rT' b' where
       
   225         s1: "s1 = (rev apTs @ X # ST, LT)" and
       
   226         l:  "length apTs = length list" and
       
   227         c:  "is_class G cname" and
       
   228         C:  "G \<turnstile> X \<preceq> Class cname" and
       
   229         w:  "\<forall>x \<in> set (zip apTs list). x \<in> widen G" and
       
   230         m:  "method (G, cname) (mname, list) = Some (mD', rT', b')"
       
   231         by (simp del: not_None_eq) blast+
       
   232 
       
   233       obtain apTs' X' ST' LT' where
       
   234         s2: "s2 = (rev apTs' @ X' # ST', LT')" and
       
   235         l': "length apTs' = length list"
       
   236       proof -
       
   237         from l s1 G
       
   238         have "length list < length (fst s2)"
       
   239           by simp
       
   240         hence "\<exists>a b c. (fst s2) = rev a @ b # c \<and> length a = length list"
       
   241           by (rule rev_append_cons [rule_format])
       
   242         thus ?thesis
       
   243           by - (cases s2, elim exE conjE, simp, rule that)
       
   244       qed
       
   245 
       
   246       from l l'
       
   247       have "length (rev apTs') = length (rev apTs)" by simp
       
   248 
       
   249       from this s1 s2 G
       
   250       obtain
       
   251         G': "G \<turnstile> (apTs',LT') <=s (apTs,LT)" and
       
   252         X : "G \<turnstile>  X' \<preceq> X" and "G \<turnstile> (ST',LT') <=s (ST,LT)"
       
   253         by (simp add: sup_state_rev_fst sup_state_append_fst sup_state_Cons1)
       
   254 
       
   255       with C
       
   256       have C': "G \<turnstile> X' \<preceq> Class cname"
       
   257         by - (rule widen_trans, auto)
       
   258 
       
   259       from G'
       
   260       have "G \<turnstile> map OK apTs' <=l map OK apTs"
       
   261         by (simp add: sup_state_conv)
       
   262       also
       
   263       from l w
       
   264       have "G \<turnstile> map OK apTs <=l map OK list"
       
   265         by (simp add: all_widen_is_sup_loc)
       
   266       finally
       
   267       have "G \<turnstile> map OK apTs' <=l map OK list" .
       
   268 
       
   269       with l'
       
   270       have w': "\<forall>x \<in> set (zip apTs' list). x \<in> widen G"
       
   271         by (simp add: all_widen_is_sup_loc)
       
   272 
       
   273       from Invoke s2 l' w' C' m c
       
   274       show ?thesis
       
   275         by (simp del: split_paired_Ex) blast
       
   276     qed
       
   277   } note this [simp]
       
   278 
       
   279   assume "G \<turnstile> s <=' s'" "app i G m rT s'"
       
   280 
       
   281   thus ?thesis
       
   282     by - (cases s, cases s', auto)
       
   283 qed
       
   284 
       
   285 lemmas [simp del] = split_paired_Ex
       
   286 lemmas [simp] = step_def
       
   287 
       
   288 lemma step_mono_Some:
       
   289 "[| app i G m rT (Some s2); G \<turnstile> s1 <=s s2 |] ==>
       
   290   G \<turnstile> the (step i G (Some s1)) <=s the (step i G (Some s2))"
       
   291 proof (cases s1, cases s2)
       
   292   fix a1 b1 a2 b2
       
   293   assume s: "s1 = (a1,b1)" "s2 = (a2,b2)"
       
   294   assume app2: "app i G m rT (Some s2)"
       
   295   assume G: "G \<turnstile> s1 <=s s2"
       
   296 
       
   297   hence "G \<turnstile> Some s1 <=' Some s2"
       
   298     by simp
       
   299   from this app2
       
   300   have app1: "app i G m rT (Some s1)" by (rule app_mono)
       
   301 
       
   302   have "step i G (Some s1) \<noteq> None \<and> step i G (Some s2) \<noteq> None"
       
   303     by simp
       
   304   then
       
   305   obtain a1' b1' a2' b2'
       
   306     where step: "step i G (Some s1) = Some (a1',b1')"
       
   307                 "step i G (Some s2) = Some (a2',b2')"
       
   308     by (auto simp del: step_def simp add: s)
       
   309 
       
   310   have "G \<turnstile> (a1',b1') <=s (a2',b2')"
       
   311   proof (cases (open) i)
       
   312     case Load
       
   313 
       
   314     with s app1
       
   315     obtain y where
       
   316        y:  "nat < length b1" "b1 ! nat = OK y" by auto
       
   317 
       
   318     from Load s app2
       
   319     obtain y' where
       
   320        y': "nat < length b2" "b2 ! nat = OK y'" by auto
       
   321 
       
   322     from G s
       
   323     have "G \<turnstile> b1 <=l b2" by (simp add: sup_state_conv)
       
   324 
       
   325     with y y'
       
   326     have "G \<turnstile> y \<preceq> y'"
       
   327       by (auto dest: sup_loc_some)
       
   328 
       
   329     with Load G y y' s step app1 app2
       
   330     show ?thesis by (auto simp add: sup_state_conv)
       
   331   next
       
   332     case Store
       
   333     with G s step app1 app2
       
   334     show ?thesis
       
   335       by (auto simp add: sup_state_conv sup_loc_update)
       
   336   next
       
   337     case LitPush
       
   338     with G s step app1 app2
       
   339     show ?thesis
       
   340       by (auto simp add: sup_state_Cons1)
       
   341   next
       
   342     case New
       
   343     with G s step app1 app2
       
   344     show ?thesis
       
   345       by (auto simp add: sup_state_Cons1)
       
   346   next
       
   347     case Getfield
       
   348     with G s step app1 app2
       
   349     show ?thesis
       
   350       by (auto simp add: sup_state_Cons1)
       
   351   next
       
   352     case Putfield
       
   353     with G s step app1 app2
       
   354     show ?thesis
       
   355       by (auto simp add: sup_state_Cons1)
       
   356   next
       
   357     case Checkcast
       
   358     with G s step app1 app2
       
   359     show ?thesis
       
   360       by (auto simp add: sup_state_Cons1)
       
   361   next
       
   362     case Invoke
       
   363 
       
   364     with s app1
       
   365     obtain a X ST where
       
   366       s1: "s1 = (a @ X # ST, b1)" and
       
   367       l:  "length a = length list"
       
   368       by auto
       
   369 
       
   370     from Invoke s app2
       
   371     obtain a' X' ST' where
       
   372       s2: "s2 = (a' @ X' # ST', b2)" and
       
   373       l': "length a' = length list"
       
   374       by auto
       
   375 
       
   376     from l l'
       
   377     have lr: "length a = length a'" by simp
       
   378 
       
   379     from lr G s s1 s2
       
   380     have "G \<turnstile> (ST, b1) <=s (ST', b2)"
       
   381       by (simp add: sup_state_append_fst sup_state_Cons1)
       
   382 
       
   383     moreover
       
   384 
       
   385     from Invoke G s step app1 app2
       
   386     have "b1 = b1' \<and> b2 = b2'" by simp
       
   387 
       
   388     ultimately
       
   389 
       
   390     have "G \<turnstile> (ST, b1') <=s (ST', b2')" by simp
       
   391 
       
   392     with Invoke G s step app1 app2 s1 s2 l l'
       
   393     show ?thesis
       
   394       by (auto simp add: sup_state_conv)
       
   395   next
       
   396     case Return
       
   397     with G step
       
   398     show ?thesis
       
   399       by simp
       
   400   next
       
   401     case Pop
       
   402     with G s step app1 app2
       
   403     show ?thesis
       
   404       by (auto simp add: sup_state_Cons1)
       
   405   next
       
   406     case Dup
       
   407     with G s step app1 app2
       
   408     show ?thesis
       
   409       by (auto simp add: sup_state_Cons1)
       
   410   next
       
   411     case Dup_x1
       
   412     with G s step app1 app2
       
   413     show ?thesis
       
   414       by (auto simp add: sup_state_Cons1)
       
   415   next
       
   416     case Dup_x2
       
   417     with G s step app1 app2
       
   418     show ?thesis
       
   419       by (auto simp add: sup_state_Cons1)
       
   420   next
       
   421     case Swap
       
   422     with G s step app1 app2
       
   423     show ?thesis
       
   424       by (auto simp add: sup_state_Cons1)
       
   425   next
       
   426     case IAdd
       
   427     with G s step app1 app2
       
   428     show ?thesis
       
   429       by (auto simp add: sup_state_Cons1)
       
   430   next
       
   431     case Goto
       
   432     with G s step app1 app2
       
   433     show ?thesis by simp
       
   434   next
       
   435     case Ifcmpeq
       
   436     with G s step app1 app2
       
   437     show ?thesis
       
   438       by (auto simp add: sup_state_Cons1)
       
   439   qed
       
   440 
       
   441   with step
       
   442   show ?thesis by auto
       
   443 qed
       
   444 
       
   445 lemma step_mono:
       
   446   "[| app i G m rT s2; G \<turnstile> s1 <=' s2 |] ==>
       
   447   G \<turnstile> step i G s1 <=' step i G s2"
       
   448   by (cases s1, cases s2, auto dest: step_mono_Some)
       
   449 
       
   450 lemmas [simp del] = step_def
       
   451 lemmas [iff del] = not_Err_eq
       
   452 
       
   453 end
       
   454