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