| author | huffman | 
| Sun, 11 Sep 2011 09:40:18 -0700 | |
| changeset 44884 | 02efd5a6b6e5 | 
| parent 33954 | 1bc3b688548c | 
| child 58886 | 8a6cac7c7247 | 
| permissions | -rw-r--r-- | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 1 | (* Author: Gerwin Klein | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 2 | Copyright 1999 Technische Universitaet Muenchen | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 3 | *) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 4 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 5 | header {* \isaheader{Correctness of the LBV} *}
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 6 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 7 | theory LBVCorrect | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 8 | imports LBVSpec Typing_Framework | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 9 | begin | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 10 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 11 | locale lbvs = lbv + | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 12 |   fixes s0  :: 'a ("s\<^sub>0")
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 13 | fixes c :: "'a list" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 14 | fixes ins :: "'b list" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 15 |   fixes phi :: "'a list" ("\<phi>")
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 16 | defines phi_def: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 17 | "\<phi> \<equiv> map (\<lambda>pc. if c!pc = \<bottom> then wtl (take pc ins) c 0 s0 else c!pc) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 18 | [0..<length ins]" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 19 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 20 | assumes bounded: "bounded step (length ins)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 21 | assumes cert: "cert_ok c (length ins) \<top> \<bottom> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 22 | assumes pres: "pres_type step (length ins) A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 23 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 24 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 25 | lemma (in lbvs) phi_None [intro?]: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 26 | "\<lbrakk> pc < length ins; c!pc = \<bottom> \<rbrakk> \<Longrightarrow> \<phi> ! pc = wtl (take pc ins) c 0 s0" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 27 | by (simp add: phi_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 28 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 29 | lemma (in lbvs) phi_Some [intro?]: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 30 | "\<lbrakk> pc < length ins; c!pc \<noteq> \<bottom> \<rbrakk> \<Longrightarrow> \<phi> ! pc = c ! pc" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 31 | by (simp add: phi_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 32 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 33 | lemma (in lbvs) phi_len [simp]: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 34 | "length \<phi> = length ins" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 35 | by (simp add: phi_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 36 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 37 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 38 | lemma (in lbvs) wtl_suc_pc: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 39 | assumes all: "wtl ins c 0 s\<^sub>0 \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 40 | assumes pc: "pc+1 < length ins" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 41 | shows "wtl (take (pc+1) ins) c 0 s0 \<sqsubseteq>\<^sub>r \<phi>!(pc+1)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 42 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 43 | from all pc | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 44 | have "wtc c (pc+1) (wtl (take (pc+1) ins) c 0 s0) \<noteq> T" by (rule wtl_all) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 45 | with pc show ?thesis by (simp add: phi_def wtc split: split_if_asm) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 46 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 47 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 48 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 49 | lemma (in lbvs) wtl_stable: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 50 | assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 51 | assumes s0: "s0 \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 52 | assumes pc: "pc < length ins" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 53 | shows "stable r step \<phi> pc" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 54 | proof (unfold stable_def, clarify) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 55 | fix pc' s' assume step: "(pc',s') \<in> set (step pc (\<phi> ! pc))" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 56 | (is "(pc',s') \<in> set (?step pc)") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 57 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 58 | from bounded pc step have pc': "pc' < length ins" by (rule boundedD) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 59 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 60 | from wtl have tkpc: "wtl (take pc ins) c 0 s0 \<noteq> \<top>" (is "?s1 \<noteq> _") by (rule wtl_take) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 61 | from wtl have s2: "wtl (take (pc+1) ins) c 0 s0 \<noteq> \<top>" (is "?s2 \<noteq> _") by (rule wtl_take) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 62 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 63 | from wtl pc have wt_s1: "wtc c pc ?s1 \<noteq> \<top>" by (rule wtl_all) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 64 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 65 | have c_Some: "\<forall>pc t. pc < length ins \<longrightarrow> c!pc \<noteq> \<bottom> \<longrightarrow> \<phi>!pc = c!pc" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 66 | by (simp add: phi_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 67 | from pc have c_None: "c!pc = \<bottom> \<Longrightarrow> \<phi>!pc = ?s1" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 68 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 69 | from wt_s1 pc c_None c_Some | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 70 | have inst: "wtc c pc ?s1 = wti c pc (\<phi>!pc)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 71 | by (simp add: wtc split: split_if_asm) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 72 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 73 | from pres cert s0 wtl pc have "?s1 \<in> A" by (rule wtl_pres) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 74 | with pc c_Some cert c_None | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 75 | have "\<phi>!pc \<in> A" by (cases "c!pc = \<bottom>") (auto dest: cert_okD1) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 76 | with pc pres | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 77 | have step_in_A: "snd`set (?step pc) \<subseteq> A" by (auto dest: pres_typeD2) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 78 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 79 | show "s' <=_r \<phi>!pc'" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 80 | proof (cases "pc' = pc+1") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 81 | case True | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 82 | with pc' cert | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 83 | have cert_in_A: "c!(pc+1) \<in> A" by (auto dest: cert_okD1) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 84 | from True pc' have pc1: "pc+1 < length ins" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 85 | with tkpc have "?s2 = wtc c pc ?s1" by - (rule wtl_Suc) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 86 | with inst | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 87 | have merge: "?s2 = merge c pc (?step pc) (c!(pc+1))" by (simp add: wti) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 88 | also | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 89 | from s2 merge have "\<dots> \<noteq> \<top>" (is "?merge \<noteq> _") by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 90 | with cert_in_A step_in_A | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 91 | have "?merge = (map snd [(p',t') \<leftarrow> ?step pc. p'=pc+1] ++_f (c!(pc+1)))" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 92 | by (rule merge_not_top_s) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 93 | finally | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 94 | have "s' <=_r ?s2" using step_in_A cert_in_A True step | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 95 | by (auto intro: pp_ub1') | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 96 | also | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 97 | from wtl pc1 have "?s2 <=_r \<phi>!(pc+1)" by (rule wtl_suc_pc) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 98 | also note True [symmetric] | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 99 | finally show ?thesis by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 100 | next | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 101 | case False | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 102 | from wt_s1 inst | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 103 | have "merge c pc (?step pc) (c!(pc+1)) \<noteq> \<top>" by (simp add: wti) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 104 | with step_in_A | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 105 | have "\<forall>(pc', s')\<in>set (?step pc). pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 106 | by - (rule merge_not_top) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 107 | with step False | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 108 | have ok: "s' <=_r c!pc'" by blast | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 109 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 110 | from ok | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 111 | have "c!pc' = \<bottom> \<Longrightarrow> s' = \<bottom>" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 112 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 113 | from c_Some pc' | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 114 | have "c!pc' \<noteq> \<bottom> \<Longrightarrow> \<phi>!pc' = c!pc'" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 115 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 116 | show ?thesis by (cases "c!pc' = \<bottom>") auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 117 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 118 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 119 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 120 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 121 | lemma (in lbvs) phi_not_top: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 122 | assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 123 | assumes pc: "pc < length ins" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 124 | shows "\<phi>!pc \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 125 | proof (cases "c!pc = \<bottom>") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 126 | case False with pc | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 127 | have "\<phi>!pc = c!pc" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 128 | also from cert pc have "\<dots> \<noteq> \<top>" by (rule cert_okD4) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 129 | finally show ?thesis . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 130 | next | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 131 | case True with pc | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 132 | have "\<phi>!pc = wtl (take pc ins) c 0 s0" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 133 | also from wtl have "\<dots> \<noteq> \<top>" by (rule wtl_take) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 134 | finally show ?thesis . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 135 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 136 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 137 | lemma (in lbvs) phi_in_A: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 138 | assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 139 | assumes s0: "s0 \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 140 | shows "\<phi> \<in> list (length ins) A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 141 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 142 |   { fix x assume "x \<in> set \<phi>"
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 143 | then obtain xs ys where "\<phi> = xs @ x # ys" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 144 | by (auto simp add: in_set_conv_decomp) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 145 | then obtain pc where pc: "pc < length \<phi>" and x: "\<phi>!pc = x" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 146 | by (simp add: that [of "length xs"] nth_append) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 147 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 148 | from pres cert wtl s0 pc | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 149 | have "wtl (take pc ins) c 0 s0 \<in> A" by (auto intro!: wtl_pres) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 150 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 151 | from pc have "pc < length ins" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 152 | with cert have "c!pc \<in> A" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 153 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 154 | have "\<phi>!pc \<in> A" using pc by (simp add: phi_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 155 | hence "x \<in> A" using x by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 156 | } | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 157 | hence "set \<phi> \<subseteq> A" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 158 | thus ?thesis by (unfold list_def) simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 159 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 160 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 161 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 162 | lemma (in lbvs) phi0: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 163 | assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 164 | assumes 0: "0 < length ins" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 165 | shows "s0 <=_r \<phi>!0" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 166 | proof (cases "c!0 = \<bottom>") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 167 | case True | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 168 | with 0 have "\<phi>!0 = wtl (take 0 ins) c 0 s0" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 169 | moreover have "wtl (take 0 ins) c 0 s0 = s0" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 170 | ultimately have "\<phi>!0 = s0" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 171 | thus ?thesis by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 172 | next | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 173 | case False | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 174 | with 0 have "phi!0 = c!0" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 175 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 176 | from wtl have "wtl (take 1 ins) c 0 s0 \<noteq> \<top>" by (rule wtl_take) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 177 | with 0 False | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 178 | have "s0 <=_r c!0" by (auto simp add: neq_Nil_conv wtc split: split_if_asm) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 179 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 180 | show ?thesis by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 181 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 182 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 183 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 184 | theorem (in lbvs) wtl_sound: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 185 | assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 186 | assumes s0: "s0 \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 187 | shows "\<exists>ts. wt_step r \<top> step ts" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 188 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 189 | have "wt_step r \<top> step \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 190 | proof (unfold wt_step_def, intro strip conjI) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 191 | fix pc assume "pc < length \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 192 | then have pc: "pc < length ins" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 193 | with wtl show "\<phi>!pc \<noteq> \<top>" by (rule phi_not_top) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 194 | from wtl s0 pc show "stable r step \<phi> pc" by (rule wtl_stable) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 195 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 196 | thus ?thesis .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 197 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 198 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 199 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 200 | theorem (in lbvs) wtl_sound_strong: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 201 | assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 202 | assumes s0: "s0 \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 203 | assumes nz: "0 < length ins" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 204 | shows "\<exists>ts \<in> list (length ins) A. wt_step r \<top> step ts \<and> s0 <=_r ts!0" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 205 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 206 | from wtl s0 have "\<phi> \<in> list (length ins) A" by (rule phi_in_A) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 207 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 208 | have "wt_step r \<top> step \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 209 | proof (unfold wt_step_def, intro strip conjI) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 210 | fix pc assume "pc < length \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 211 | then have pc: "pc < length ins" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 212 | with wtl show "\<phi>!pc \<noteq> \<top>" by (rule phi_not_top) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 213 | from wtl s0 pc show "stable r step \<phi> pc" by (rule wtl_stable) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 214 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 215 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 216 | from wtl nz have "s0 <=_r \<phi>!0" by (rule phi0) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 217 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 218 | show ?thesis by fast | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 219 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 220 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 221 | end |