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