src/HOL/MicroJava/BV/LBVCorrect.thy
author ballarin
Mon, 02 Aug 2004 10:12:02 +0200
changeset 15096 be1d3b8cfbd5
parent 13649 0f562a70c07d
child 15425 6356d2523f73
permissions -rw-r--r--
Documentation added; minor improvements.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
     1
(*
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     2
    ID:         $Id$
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     3
    Author:     Gerwin Klein
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     4
    Copyright   1999 Technische Universitaet Muenchen
9054
0e48e7d4d4f9 minor tuning for pdf documents
kleing
parents: 9012
diff changeset
     5
*)
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     6
12911
704713ca07ea new document
kleing
parents: 12516
diff changeset
     7
header {* \isaheader{Correctness of the LBV} *}
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     8
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
     9
theory LBVCorrect = LBVSpec + Typing_Framework:
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    10
13365
a2c4faad4d35 adapted to locale defs;
wenzelm
parents: 13215
diff changeset
    11
locale (open) lbvs = lbv +
13649
0f562a70c07d alternative syntax
kleing
parents: 13365
diff changeset
    12
  fixes s0  :: 'a ("s\<^sub>0")
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    13
  fixes c   :: "'a list"
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    14
  fixes ins :: "'b list"
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    15
  fixes phi :: "'a list" ("\<phi>")
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    16
  defines phi_def:
13080
kleing
parents: 13078
diff changeset
    17
  "\<phi> \<equiv> map (\<lambda>pc. if c!pc = \<bottom> then wtl (take pc ins) c 0 s0 else c!pc) 
kleing
parents: 13078
diff changeset
    18
       [0..length ins(]"
kleing
parents: 13078
diff changeset
    19
kleing
parents: 13078
diff changeset
    20
  assumes bounded: "bounded step (length ins)"
kleing
parents: 13078
diff changeset
    21
  assumes cert: "cert_ok c (length ins) \<top> \<bottom> A"
kleing
parents: 13078
diff changeset
    22
  assumes pres: "pres_type step (length ins) A"
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    23
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    24
13080
kleing
parents: 13078
diff changeset
    25
lemma (in lbvs) phi_None [intro?]:
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    26
  "\<lbrakk> pc < length ins; c!pc = \<bottom> \<rbrakk> \<Longrightarrow> \<phi> ! pc = wtl (take pc ins) c 0 s0"
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    27
  by (simp add: phi_def)
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    28
13080
kleing
parents: 13078
diff changeset
    29
lemma (in lbvs) phi_Some [intro?]:
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    30
  "\<lbrakk> pc < length ins; c!pc \<noteq> \<bottom> \<rbrakk> \<Longrightarrow> \<phi> ! pc = c ! pc"
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    31
  by (simp add: phi_def)
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    32
13080
kleing
parents: 13078
diff changeset
    33
lemma (in lbvs) phi_len [simp]:
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    34
  "length \<phi> = length ins"
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    35
  by (simp add: phi_def)
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    36
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    37
13080
kleing
parents: 13078
diff changeset
    38
lemma (in lbvs) wtl_suc_pc:
13649
0f562a70c07d alternative syntax
kleing
parents: 13365
diff changeset
    39
  assumes all: "wtl ins c 0 s\<^sub>0 \<noteq> \<top>" 
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    40
  assumes pc:  "pc+1 < length ins"
13649
0f562a70c07d alternative syntax
kleing
parents: 13365
diff changeset
    41
  shows "wtl (take (pc+1) ins) c 0 s0 \<le>\<^sub>r \<phi>!(pc+1)"
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    42
proof -
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    43
  from all pc
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    44
  have "wtc c (pc+1) (wtl (take (pc+1) ins) c 0 s0) \<noteq> T" by (rule wtl_all)
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    45
  with pc show ?thesis by (simp add: phi_def wtc split: split_if_asm)
9580
kleing
parents: 9549
diff changeset
    46
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    47
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    48
13080
kleing
parents: 13078
diff changeset
    49
lemma (in lbvs) wtl_stable:
kleing
parents: 13078
diff changeset
    50
  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>" 
kleing
parents: 13078
diff changeset
    51
  assumes s0:  "s0 \<in> A" 
kleing
parents: 13078
diff changeset
    52
  assumes pc:  "pc < length ins" 
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    53
  shows "stable r step \<phi> pc"
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    54
proof (unfold stable_def, clarify)
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    55
  fix pc' s' assume step: "(pc',s') \<in> set (step pc (\<phi> ! pc))" 
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    56
                      (is "(pc',s') \<in> set (?step pc)")
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    57
  
13080
kleing
parents: 13078
diff changeset
    58
  from bounded pc step have pc': "pc' < length ins" by (rule boundedD)
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    59
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    60
  have tkpc: "wtl (take pc ins) c 0 s0 \<noteq> \<top>" (is "?s1 \<noteq> _") by (rule wtl_take)
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    61
  have s2: "wtl (take (pc+1) ins) c 0 s0 \<noteq> \<top>" (is "?s2 \<noteq> _") by (rule wtl_take)
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    62
  
13080
kleing
parents: 13078
diff changeset
    63
  from wtl pc have wt_s1: "wtc c pc ?s1 \<noteq> \<top>" by (rule wtl_all)
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    64
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    65
  have c_Some: "\<forall>pc t. pc < length ins \<longrightarrow> c!pc \<noteq> \<bottom> \<longrightarrow> \<phi>!pc = c!pc" 
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    66
    by (simp add: phi_def)
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    67
  have c_None: "c!pc = \<bottom> \<Longrightarrow> \<phi>!pc = ?s1" ..
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    68
13080
kleing
parents: 13078
diff changeset
    69
  from wt_s1 pc c_None c_Some
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    70
  have inst: "wtc c pc ?s1  = wti c pc (\<phi>!pc)"
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    71
    by (simp add: wtc split: split_if_asm)
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    72
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    73
  have "?s1 \<in> A" by (rule wtl_pres) 
13080
kleing
parents: 13078
diff changeset
    74
  with pc c_Some cert c_None
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    75
  have "\<phi>!pc \<in> A" by (cases "c!pc = \<bottom>") (auto dest: cert_okD1)
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    76
  with pc pres
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    77
  have step_in_A: "snd`set (?step pc) \<subseteq> A" by (auto dest: pres_typeD2)
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    78
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    79
  show "s' <=_r \<phi>!pc'" 
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    80
  proof (cases "pc' = pc+1")
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    81
    case True
13080
kleing
parents: 13078
diff changeset
    82
    with pc' cert
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    83
    have cert_in_A: "c!(pc+1) \<in> A" by (auto dest: cert_okD1)
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    84
    from True pc' have pc1: "pc+1 < length ins" by simp
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    85
    with tkpc have "?s2 = wtc c pc ?s1" by - (rule wtl_Suc)
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    86
    with inst 
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    87
    have merge: "?s2 = merge c pc (?step pc) (c!(pc+1))" by (simp add: wti)
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    88
    also    
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    89
    from s2 merge have "\<dots> \<noteq> \<top>" (is "?merge \<noteq> _") by simp
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    90
    with cert_in_A step_in_A
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    91
    have "?merge = (map snd [(p',t')\<in>?step pc. p'=pc+1] ++_f (c!(pc+1)))"
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    92
      by (rule merge_not_top_s) 
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    93
    finally
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    94
    have "s' <=_r ?s2" using step_in_A cert_in_A True step 
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    95
      by (auto intro: pp_ub1')
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    96
    also 
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    97
    from wtl pc1 have "?s2 <=_r \<phi>!(pc+1)" by (rule wtl_suc_pc)
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    98
    also note True [symmetric]
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
    99
    finally show ?thesis by simp    
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   100
  next
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   101
    case False
13080
kleing
parents: 13078
diff changeset
   102
    from wt_s1 inst
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   103
    have "merge c pc (?step pc) (c!(pc+1)) \<noteq> \<top>" by (simp add: wti)
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   104
    with step_in_A
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   105
    have "\<forall>(pc', s')\<in>set (?step pc). pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'" 
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   106
      by - (rule merge_not_top)
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   107
    with step False 
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   108
    have ok: "s' <=_r c!pc'" by blast
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   109
    moreover
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   110
    from ok
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   111
    have "c!pc' = \<bottom> \<Longrightarrow> s' = \<bottom>" by simp
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   112
    moreover
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   113
    from c_Some pc'
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   114
    have "c!pc' \<noteq> \<bottom> \<Longrightarrow> \<phi>!pc' = c!pc'" by auto
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   115
    ultimately
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   116
    show ?thesis by (cases "c!pc' = \<bottom>") auto 
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   117
  qed
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   118
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   119
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   120
  
13080
kleing
parents: 13078
diff changeset
   121
lemma (in lbvs) phi_not_top:
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   122
  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>"
13080
kleing
parents: 13078
diff changeset
   123
  assumes pc:  "pc < length ins"
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   124
  shows "\<phi>!pc \<noteq> \<top>"
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   125
proof (cases "c!pc = \<bottom>")
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   126
  case False with pc
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   127
  have "\<phi>!pc = c!pc" ..
13080
kleing
parents: 13078
diff changeset
   128
  also from cert pc have "\<dots> \<noteq> \<top>" by (rule cert_okD4)
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   129
  finally show ?thesis .
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   130
next
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   131
  case True with pc
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   132
  have "\<phi>!pc = wtl (take pc ins) c 0 s0" ..
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   133
  also from wtl have "\<dots> \<noteq> \<top>" by (rule wtl_take)
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   134
  finally show ?thesis .
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   135
qed
13207
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   136
13215
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   137
lemma (in lbvs) phi_in_A:
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   138
  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>"
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   139
  assumes s0:  "s0 \<in> A"
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   140
  shows "\<phi> \<in> list (length ins) A"
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   141
proof -
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   142
  { fix x assume "x \<in> set \<phi>"
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   143
    then obtain xs ys where "\<phi> = xs @ x # ys" 
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   144
      by (auto simp add: in_set_conv_decomp)
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   145
    then obtain pc where pc: "pc < length \<phi>" and x: "\<phi>!pc = x"
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   146
      by (simp add: that [of "length xs"] nth_append)
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   147
    
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   148
    from wtl s0 pc 
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   149
    have "wtl (take pc ins) c 0 s0 \<in> A" by (auto intro!: wtl_pres)
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   150
    moreover
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   151
    from pc have "pc < length ins" by simp
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   152
    with cert have "c!pc \<in> A" ..
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   153
    ultimately
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   154
    have "\<phi>!pc \<in> A" using pc by (simp add: phi_def)
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   155
    hence "x \<in> A" using x by simp
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   156
  } 
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   157
  hence "set \<phi> \<subseteq> A" ..
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   158
  thus ?thesis by (unfold list_def) simp
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   159
qed
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   160
13207
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   161
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   162
lemma (in lbvs) phi0:
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   163
  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>"
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   164
  assumes 0:   "0 < length ins"
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   165
  shows "s0 <=_r \<phi>!0"
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   166
proof (cases "c!0 = \<bottom>")
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   167
  case True
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   168
  with 0 have "\<phi>!0 = wtl (take 0 ins) c 0 s0" ..
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   169
  moreover have "wtl (take 0 ins) c 0 s0 = s0" by simp
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   170
  ultimately have "\<phi>!0 = s0" by simp
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   171
  thus ?thesis by simp
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   172
next
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   173
  case False
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   174
  with 0 have "phi!0 = c!0" ..
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   175
  moreover 
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   176
  have "wtl (take 1 ins) c 0 s0 \<noteq> \<top>"  by (rule wtl_take)
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   177
  with 0 False 
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   178
  have "s0 <=_r c!0" by (auto simp add: neq_Nil_conv wtc split: split_if_asm)
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   179
  ultimately
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   180
  show ?thesis by simp
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   181
qed
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   182
9376
c32c5696ec2a flat instruction set
kleing
parents: 9260
diff changeset
   183
13080
kleing
parents: 13078
diff changeset
   184
theorem (in lbvs) wtl_sound:
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   185
  assumes "wtl ins c 0 s0 \<noteq> \<top>" 
13080
kleing
parents: 13078
diff changeset
   186
  assumes "s0 \<in> A" 
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   187
  shows "\<exists>ts. wt_step r \<top> step ts"
13207
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   188
proof -
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   189
  have "wt_step r \<top> step \<phi>"
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   190
  proof (unfold wt_step_def, intro strip conjI)
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   191
    fix pc assume "pc < length \<phi>"
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   192
    then obtain "pc < length ins" by simp
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   193
    show "\<phi>!pc \<noteq> \<top>" by (rule phi_not_top)
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   194
    show "stable r step \<phi> pc" by (rule wtl_stable)
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   195
  qed
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   196
  thus ?thesis ..
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   197
qed
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   198
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   199
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   200
theorem (in lbvs) wtl_sound_strong:
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   201
  assumes "wtl ins c 0 s0 \<noteq> \<top>" 
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   202
  assumes "s0 \<in> A" 
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   203
  assumes "0 < length ins"
13215
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   204
  shows "\<exists>ts \<in> list (length ins) A. wt_step r \<top> step ts \<and> s0 <=_r ts!0"
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   205
proof -
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   206
  have "\<phi> \<in> list (length ins) A" by (rule phi_in_A)
072a77989ce0 LBV instantiated for JVM
kleing
parents: 13209
diff changeset
   207
  moreover
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   208
  have "wt_step r \<top> step \<phi>"
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   209
  proof (unfold wt_step_def, intro strip conjI)
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   210
    fix pc assume "pc < length \<phi>"
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   211
    then obtain "pc < length ins" by simp
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   212
    show "\<phi>!pc \<noteq> \<top>" by (rule phi_not_top)
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   213
    show "stable r step \<phi> pc" by (rule wtl_stable)
13207
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   214
  qed
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   215
  moreover
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   216
  have "s0 <=_r \<phi>!0" by (rule phi0)
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   217
  ultimately
0d07e49dc9a5 included strong soundness (sound + s0 <= phi!0)
kleing
parents: 13080
diff changeset
   218
  show ?thesis by fast
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   219
qed
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   220
13078
1dd711c6b93c flattened, uses locales
kleing
parents: 13071
diff changeset
   221
end