src/HOL/MicroJava/JVM/JVMDefensive.thy
author wenzelm
Sun, 02 Nov 2014 17:58:35 +0100
changeset 58886 8a6cac7c7247
parent 58310 91ea607a34d8
child 61361 8b5f00202e1a
permissions -rw-r--r--
modernized header;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/JVM/JVMDefensive.thy
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
     2
    Author:     Gerwin Klein
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
     3
*)
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
     4
58886
8a6cac7c7247 modernized header;
wenzelm
parents: 58310
diff changeset
     5
section {* A Defensive JVM *}
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
     6
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
     7
theory JVMDefensive
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
     8
imports JVMExec
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
     9
begin
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    10
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    11
text {*
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    12
  Extend the state space by one element indicating a type error (or
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    13
  other abnormal termination) *}
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    14
datatype 'a type_error = TypeError | Normal 'a
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    15
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    16
35102
cc7a0b9f938c modernized translations;
wenzelm
parents: 16417
diff changeset
    17
abbreviation
cc7a0b9f938c modernized translations;
wenzelm
parents: 16417
diff changeset
    18
  fifth :: "'a \<times> 'b \<times> 'c \<times> 'd \<times> 'e \<times> 'f \<Rightarrow> 'e"
cc7a0b9f938c modernized translations;
wenzelm
parents: 16417
diff changeset
    19
  where "fifth x == fst(snd(snd(snd(snd x))))"
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    20
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    21
fun isAddr :: "val \<Rightarrow> bool" where
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    22
  "isAddr (Addr loc) = True"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    23
| "isAddr v          = False"
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    24
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    25
fun isIntg :: "val \<Rightarrow> bool" where
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    26
  "isIntg (Intg i) = True"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    27
| "isIntg v        = False"
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    28
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    29
definition isRef :: "val \<Rightarrow> bool" where
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    30
  "isRef v \<equiv> v = Null \<or> isAddr v"
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    31
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    32
primrec check_instr :: "[instr, jvm_prog, aheap, opstack, locvars, 
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    33
                  cname, sig, p_count, nat, frame list] \<Rightarrow> bool" where
13822
bb5eda7416e5 check maxs in defensive machine
kleing
parents: 13677
diff changeset
    34
  "check_instr (Load idx) G hp stk vars C sig pc mxs frs = 
bb5eda7416e5 check maxs in defensive machine
kleing
parents: 13677
diff changeset
    35
  (idx < length vars \<and> size stk < mxs)"
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    36
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    37
| "check_instr (Store idx) G hp stk vars Cl sig pc mxs frs = 
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    38
  (0 < length stk \<and> idx < length vars)"
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    39
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    40
| "check_instr (LitPush v) G hp stk vars Cl sig pc mxs frs = 
13822
bb5eda7416e5 check maxs in defensive machine
kleing
parents: 13677
diff changeset
    41
  (\<not>isAddr v \<and> size stk < mxs)"
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    42
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    43
| "check_instr (New C) G hp stk vars Cl sig pc mxs frs = 
13822
bb5eda7416e5 check maxs in defensive machine
kleing
parents: 13677
diff changeset
    44
  (is_class G C \<and> size stk < mxs)"
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    45
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    46
| "check_instr (Getfield F C) G hp stk vars Cl sig pc mxs frs = 
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    47
  (0 < length stk \<and> is_class G C \<and> field (G,C) F \<noteq> None \<and> 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    48
  (let (C', T) = the (field (G,C) F); ref = hd stk in 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    49
    C' = C \<and> isRef ref \<and> (ref \<noteq> Null \<longrightarrow> 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    50
      hp (the_Addr ref) \<noteq> None \<and> 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    51
      (let (D,vs) = the (hp (the_Addr ref)) in 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    52
        G \<turnstile> D \<preceq>C C \<and> vs (F,C) \<noteq> None \<and> G,hp \<turnstile> the (vs (F,C)) ::\<preceq> T))))" 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    53
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    54
| "check_instr (Putfield F C) G hp stk vars Cl sig pc mxs frs = 
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    55
  (1 < length stk \<and> is_class G C \<and> field (G,C) F \<noteq> None \<and> 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    56
  (let (C', T) = the (field (G,C) F); v = hd stk; ref = hd (tl stk) in 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    57
    C' = C \<and> isRef ref \<and> (ref \<noteq> Null \<longrightarrow> 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    58
      hp (the_Addr ref) \<noteq> None \<and> 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    59
      (let (D,vs) = the (hp (the_Addr ref)) in 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    60
        G \<turnstile> D \<preceq>C C \<and> G,hp \<turnstile> v ::\<preceq> T))))" 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    61
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    62
| "check_instr (Checkcast C) G hp stk vars Cl sig pc mxs frs =
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    63
  (0 < length stk \<and> is_class G C \<and> isRef (hd stk))"
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    64
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    65
| "check_instr (Invoke C mn ps) G hp stk vars Cl sig pc mxs frs =
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    66
  (length ps < length stk \<and> 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    67
  (let n = length ps; v = stk!n in
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    68
  isRef v \<and> (v \<noteq> Null \<longrightarrow> 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    69
    hp (the_Addr v) \<noteq> None \<and>
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    70
    method (G,cname_of hp v) (mn,ps) \<noteq> None \<and>
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    71
    list_all2 (\<lambda>v T. G,hp \<turnstile> v ::\<preceq> T) (rev (take n stk)) ps)))"
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    72
  
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    73
| "check_instr Return G hp stk0 vars Cl sig0 pc mxs frs =
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    74
  (0 < length stk0 \<and> (0 < length frs \<longrightarrow> 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    75
    method (G,Cl) sig0 \<noteq> None \<and>    
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    76
    (let v = hd stk0;  (C, rT, body) = the (method (G,Cl) sig0) in
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    77
    Cl = C \<and> G,hp \<turnstile> v ::\<preceq> rT)))"
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    78
 
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    79
| "check_instr Pop G hp stk vars Cl sig pc mxs frs = 
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    80
  (0 < length stk)"
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    81
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    82
| "check_instr Dup G hp stk vars Cl sig pc mxs frs = 
13822
bb5eda7416e5 check maxs in defensive machine
kleing
parents: 13677
diff changeset
    83
  (0 < length stk \<and> size stk < mxs)"
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    84
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    85
| "check_instr Dup_x1 G hp stk vars Cl sig pc mxs frs = 
13822
bb5eda7416e5 check maxs in defensive machine
kleing
parents: 13677
diff changeset
    86
  (1 < length stk \<and> size stk < mxs)"
bb5eda7416e5 check maxs in defensive machine
kleing
parents: 13677
diff changeset
    87
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    88
| "check_instr Dup_x2 G hp stk vars Cl sig pc mxs frs = 
13822
bb5eda7416e5 check maxs in defensive machine
kleing
parents: 13677
diff changeset
    89
  (2 < length stk \<and> size stk < mxs)"
bb5eda7416e5 check maxs in defensive machine
kleing
parents: 13677
diff changeset
    90
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    91
| "check_instr Swap G hp stk vars Cl sig pc mxs frs =
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    92
  (1 < length stk)"
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    93
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    94
| "check_instr IAdd G hp stk vars Cl sig pc mxs frs =
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    95
  (1 < length stk \<and> isIntg (hd stk) \<and> isIntg (hd (tl stk)))"
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    96
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
    97
| "check_instr (Ifcmpeq b) G hp stk vars Cl sig pc mxs frs =
13677
5fad004bd9df cleanup, beautified
kleing
parents: 13631
diff changeset
    98
  (1 < length stk \<and> 0 \<le> int pc+b)"
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
    99
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
   100
| "check_instr (Goto b) G hp stk vars Cl sig pc mxs frs =
13677
5fad004bd9df cleanup, beautified
kleing
parents: 13631
diff changeset
   101
  (0 \<le> int pc+b)"
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   102
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
   103
| "check_instr Throw G hp stk vars Cl sig pc mxs frs =
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   104
  (0 < length stk \<and> isRef (hd stk))"
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   105
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
   106
definition check :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> bool" where
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   107
  "check G s \<equiv> let (xcpt, hp, frs) = s in
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   108
               (case frs of [] \<Rightarrow> True | (stk,loc,C,sig,pc)#frs' \<Rightarrow> 
13822
bb5eda7416e5 check maxs in defensive machine
kleing
parents: 13677
diff changeset
   109
                (let  (C',rt,mxs,mxl,ins,et) = the (method (G,C) sig); i = ins!pc in
bb5eda7416e5 check maxs in defensive machine
kleing
parents: 13677
diff changeset
   110
                 pc < size ins \<and> 
bb5eda7416e5 check maxs in defensive machine
kleing
parents: 13677
diff changeset
   111
                 check_instr i G hp stk loc C sig pc mxs frs'))"
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   112
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   113
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
   114
definition exec_d :: "jvm_prog \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state option type_error" where
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   115
  "exec_d G s \<equiv> case s of 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   116
      TypeError \<Rightarrow> TypeError 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   117
    | Normal s' \<Rightarrow> if check G s' then Normal (exec (G, s')) else TypeError"
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   118
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   119
37406
982f3e02f3c4 modernized specifications
haftmann
parents: 35417
diff changeset
   120
definition
35355
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35102
diff changeset
   121
  exec_all_d :: "jvm_prog \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state type_error \<Rightarrow> bool" 
53024
e0968e1f6fe9 more symbolic notation;
wenzelm
parents: 37406
diff changeset
   122
                   ("_ \<turnstile> _ \<midarrow>jvmd\<rightarrow> _" [61,61,61]60) where
e0968e1f6fe9 more symbolic notation;
wenzelm
parents: 37406
diff changeset
   123
  "G \<turnstile> s \<midarrow>jvmd\<rightarrow> t \<longleftrightarrow>
35355
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35102
diff changeset
   124
         (s,t) \<in> ({(s,t). exec_d G s = TypeError \<and> t = TypeError} \<union>
613e133966ea modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents: 35102
diff changeset
   125
                  {(s,t). \<exists>t'. exec_d G s = Normal (Some t') \<and> t = Normal t'})\<^sup>*"
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   126
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   127
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   128
declare split_paired_All [simp del]
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   129
declare split_paired_Ex [simp del]
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   130
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   131
lemma [dest!]:
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   132
  "(if P then A else B) \<noteq> B \<Longrightarrow> P"
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   133
  by (cases P, auto)
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   134
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   135
lemma exec_d_no_errorI [intro]:
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   136
  "check G s \<Longrightarrow> exec_d G (Normal s) \<noteq> TypeError"
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   137
  by (unfold exec_d_def) simp
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   138
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   139
theorem no_type_error_commutes:
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   140
  "exec_d G (Normal s) \<noteq> TypeError \<Longrightarrow> 
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   141
  exec_d G (Normal s) = Normal (exec (G, s))"
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   142
  by (unfold exec_d_def, auto)
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   143
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   144
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   145
lemma defensive_imp_aggressive:
53024
e0968e1f6fe9 more symbolic notation;
wenzelm
parents: 37406
diff changeset
   146
  "G \<turnstile> (Normal s) \<midarrow>jvmd\<rightarrow> (Normal t) \<Longrightarrow> G \<turnstile> s \<midarrow>jvm\<rightarrow> t"
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   147
proof -
53024
e0968e1f6fe9 more symbolic notation;
wenzelm
parents: 37406
diff changeset
   148
  have "\<And>x y. G \<turnstile> x \<midarrow>jvmd\<rightarrow> y \<Longrightarrow> \<forall>s t. x = Normal s \<longrightarrow> y = Normal t \<longrightarrow>  G \<turnstile> s \<midarrow>jvm\<rightarrow> t"
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   149
    apply (unfold exec_all_d_def)
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   150
    apply (erule rtrancl_induct)
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   151
     apply (simp add: exec_all_def)
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   152
    apply (fold exec_all_d_def)
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   153
    apply simp
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   154
    apply (intro allI impI)
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   155
    apply (erule disjE, simp)
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   156
    apply (elim exE conjE)
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   157
    apply (erule allE, erule impE, assumption)
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   158
    apply (simp add: exec_all_def exec_d_def split: type_error.splits split_if_asm)
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   159
    apply (rule rtrancl_trans, assumption)
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   160
    apply blast
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   161
    done
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   162
  moreover
53024
e0968e1f6fe9 more symbolic notation;
wenzelm
parents: 37406
diff changeset
   163
  assume "G \<turnstile> (Normal s) \<midarrow>jvmd\<rightarrow> (Normal t)" 
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   164
  ultimately
53024
e0968e1f6fe9 more symbolic notation;
wenzelm
parents: 37406
diff changeset
   165
  show "G \<turnstile> s \<midarrow>jvm\<rightarrow> t" by blast
13631
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   166
qed
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   167
23ab136db946 defensive machine without obj init and jsr
kleing
parents:
diff changeset
   168
end