| author | wenzelm | 
| Fri, 20 Aug 2010 11:47:33 +0200 | |
| changeset 38566 | 8176107637ce | 
| parent 37595 | 9591362629e3 | 
| child 42150 | b0c0638c4aad | 
| permissions | -rw-r--r-- | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 1 | (* Title: HOL/MicroJava/BV/LBVComplete.thy | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 2 | Author: Gerwin Klein | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 3 | Copyright 2000 Technische Universitaet Muenchen | 
| 
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 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 6 | header {* \isaheader{Completeness of the LBV} *}
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 7 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 8 | theory LBVComplete | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 9 | imports LBVSpec Typing_Framework | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 10 | begin | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 11 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
33954diff
changeset | 12 | definition is_target :: "['s step_type, 's list, nat] \<Rightarrow> bool" where | 
| 37595 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
35416diff
changeset | 13 | "is_target step phi pc' \<longleftrightarrow> | 
| 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
35416diff
changeset | 14 | (\<exists>pc s'. pc' \<noteq> pc+1 \<and> pc < length phi \<and> (pc',s') \<in> set (step pc (phi!pc)))" | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 15 | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
33954diff
changeset | 16 | definition make_cert :: "['s step_type, 's list, 's] \<Rightarrow> 's certificate" where | 
| 37595 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
35416diff
changeset | 17 | "make_cert step phi B = | 
| 33954 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 18 | map (\<lambda>pc. if is_target step phi pc then phi!pc else B) [0..<length phi] @ [B]" | 
| 
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 | lemma [code]: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 21 | "is_target step phi pc' = | 
| 37595 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
35416diff
changeset | 22 | list_ex (\<lambda>pc. pc' \<noteq> pc+1 \<and> List.member (map fst (step pc (phi!pc))) pc') [0..<length phi]" | 
| 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
35416diff
changeset | 23 | by (force simp: list_ex_iff member_def is_target_def) | 
| 33954 
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 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 26 | locale lbvc = lbv + | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 27 |   fixes phi :: "'a list" ("\<phi>")
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 28 | fixes c :: "'a list" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 29 | defines cert_def: "c \<equiv> make_cert step \<phi> \<bottom>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 30 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 31 | assumes mono: "mono r step (length \<phi>) A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 32 | assumes pres: "pres_type step (length \<phi>) A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 33 | assumes phi: "\<forall>pc < length \<phi>. \<phi>!pc \<in> A \<and> \<phi>!pc \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 34 | assumes bounded: "bounded step (length \<phi>)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 35 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 36 | assumes B_neq_T: "\<bottom> \<noteq> \<top>" | 
| 
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 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 39 | lemma (in lbvc) cert: "cert_ok c (length \<phi>) \<top> \<bottom> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 40 | proof (unfold cert_ok_def, intro strip conjI) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 41 | note [simp] = make_cert_def cert_def nth_append | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 42 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 43 | show "c!length \<phi> = \<bottom>" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 44 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 45 | fix pc assume pc: "pc < length \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 46 | from pc phi B_A show "c!pc \<in> A" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 47 | from pc phi B_neq_T show "c!pc \<noteq> \<top>" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 48 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 49 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 50 | lemmas [simp del] = split_paired_Ex | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 51 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 52 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 53 | lemma (in lbvc) cert_target [intro?]: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 54 | "\<lbrakk> (pc',s') \<in> set (step pc (\<phi>!pc)); | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 55 | pc' \<noteq> pc+1; pc < length \<phi>; pc' < length \<phi> \<rbrakk> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 56 | \<Longrightarrow> c!pc' = \<phi>!pc'" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 57 | by (auto simp add: cert_def make_cert_def nth_append is_target_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 58 | |
| 
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 | lemma (in lbvc) cert_approx [intro?]: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 61 | "\<lbrakk> pc < length \<phi>; c!pc \<noteq> \<bottom> \<rbrakk> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 62 | \<Longrightarrow> c!pc = \<phi>!pc" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 63 | by (auto simp add: cert_def make_cert_def nth_append) | 
| 
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 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 66 | lemma (in lbv) le_top [simp, intro]: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 67 | "x <=_r \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 68 | by (insert top) simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 69 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 70 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 71 | lemma (in lbv) merge_mono: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 72 | assumes less: "ss2 <=|r| ss1" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 73 | assumes x: "x \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 74 | assumes ss1: "snd`set ss1 \<subseteq> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 75 | assumes ss2: "snd`set ss2 \<subseteq> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 76 | shows "merge c pc ss2 x <=_r merge c pc ss1 x" (is "?s2 <=_r ?s1") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 77 | proof- | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 78 | have "?s1 = \<top> \<Longrightarrow> ?thesis" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 79 |   moreover {
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 80 | assume merge: "?s1 \<noteq> T" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 81 | from x ss1 have "?s1 = | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 82 | (if \<forall>(pc', s')\<in>set ss1. pc' \<noteq> pc + 1 \<longrightarrow> s' <=_r c!pc' | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 83 | then (map snd [(p', t') \<leftarrow> ss1 . p'=pc+1]) ++_f x | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 84 | else \<top>)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 85 | by (rule merge_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 86 | with merge obtain | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 87 | app: "\<forall>(pc',s')\<in>set ss1. pc' \<noteq> pc+1 \<longrightarrow> s' <=_r c!pc'" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 88 | (is "?app ss1") and | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 89 | sum: "(map snd [(p',t') \<leftarrow> ss1 . p' = pc+1] ++_f x) = ?s1" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 90 | (is "?map ss1 ++_f x = _" is "?sum ss1 = _") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 91 | by (simp split: split_if_asm) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 92 | from app less | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 93 | have "?app ss2" by (blast dest: trans_r lesub_step_typeD) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 94 |     moreover {
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 95 | from ss1 have map1: "set (?map ss1) \<subseteq> A" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 96 | with x have "?sum ss1 \<in> A" by (auto intro!: plusplus_closed semilat) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 97 | with sum have "?s1 \<in> A" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 98 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 99 | have mapD: "\<And>x ss. x \<in> set (?map ss) \<Longrightarrow> \<exists>p. (p,x) \<in> set ss \<and> p=pc+1" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 100 | from x map1 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 101 | have "\<forall>x \<in> set (?map ss1). x <=_r ?sum ss1" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 102 | by clarify (rule pp_ub1) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 103 | with sum have "\<forall>x \<in> set (?map ss1). x <=_r ?s1" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 104 | with less have "\<forall>x \<in> set (?map ss2). x <=_r ?s1" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 105 | by (fastsimp dest!: mapD lesub_step_typeD intro: trans_r) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 106 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 107 | from map1 x have "x <=_r (?sum ss1)" by (rule pp_ub2) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 108 | with sum have "x <=_r ?s1" by simp | 
| 
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 ss2 have "set (?map ss2) \<subseteq> A" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 111 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 112 | have "?sum ss2 <=_r ?s1" using x by - (rule pp_lub) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 113 | } | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 114 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 115 | from x ss2 have | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 116 | "?s2 = | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 117 | (if \<forall>(pc', s')\<in>set ss2. pc' \<noteq> pc + 1 \<longrightarrow> s' <=_r c!pc' | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 118 | then map snd [(p', t') \<leftarrow> ss2 . p' = pc + 1] ++_f x | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 119 | else \<top>)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 120 | by (rule merge_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 121 | ultimately have ?thesis by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 122 | } | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 123 | ultimately show ?thesis by (cases "?s1 = \<top>") auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 124 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 125 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 126 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 127 | lemma (in lbvc) wti_mono: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 128 | assumes less: "s2 <=_r s1" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 129 | assumes pc: "pc < length \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 130 | assumes s1: "s1 \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 131 | assumes s2: "s2 \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 132 | shows "wti c pc s2 <=_r wti c pc s1" (is "?s2' <=_r ?s1'") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 133 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 134 | from mono pc s2 less have "step pc s2 <=|r| step pc s1" by (rule monoD) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 135 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 136 | from cert B_A pc have "c!Suc pc \<in> A" by (rule cert_okD3) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 137 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 138 | from pres s1 pc | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 139 | have "snd`set (step pc s1) \<subseteq> A" by (rule pres_typeD2) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 140 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 141 | from pres s2 pc | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 142 | have "snd`set (step pc s2) \<subseteq> A" by (rule pres_typeD2) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 143 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 144 | show ?thesis by (simp add: wti merge_mono) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 145 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 146 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 147 | lemma (in lbvc) wtc_mono: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 148 | assumes less: "s2 <=_r s1" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 149 | assumes pc: "pc < length \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 150 | assumes s1: "s1 \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 151 | assumes s2: "s2 \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 152 | shows "wtc c pc s2 <=_r wtc c pc s1" (is "?s2' <=_r ?s1'") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 153 | proof (cases "c!pc = \<bottom>") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 154 | case True | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 155 | moreover from less pc s1 s2 have "wti c pc s2 <=_r wti c pc s1" by (rule wti_mono) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 156 | ultimately show ?thesis by (simp add: wtc) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 157 | next | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 158 | case False | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 159 | have "?s1' = \<top> \<Longrightarrow> ?thesis" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 160 |   moreover {
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 161 | assume "?s1' \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 162 | with False have c: "s1 <=_r c!pc" by (simp add: wtc split: split_if_asm) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 163 | with less have "s2 <=_r c!pc" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 164 | with False c have ?thesis by (simp add: wtc) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 165 | } | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 166 | ultimately show ?thesis by (cases "?s1' = \<top>") auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 167 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 168 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 169 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 170 | lemma (in lbv) top_le_conv [simp]: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 171 | "\<top> <=_r x = (x = \<top>)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 172 | by (insert semilat) (simp add: top top_le_conv) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 173 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 174 | lemma (in lbv) neq_top [simp, elim]: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 175 | "\<lbrakk> x <=_r y; y \<noteq> \<top> \<rbrakk> \<Longrightarrow> x \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 176 | by (cases "x = T") auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 177 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 178 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 179 | lemma (in lbvc) stable_wti: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 180 | assumes stable: "stable r step \<phi> pc" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 181 | assumes pc: "pc < length \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 182 | shows "wti c pc (\<phi>!pc) \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 183 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 184 | let ?step = "step pc (\<phi>!pc)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 185 | from stable | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 186 | have less: "\<forall>(q,s')\<in>set ?step. s' <=_r \<phi>!q" by (simp add: stable_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 187 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 188 | from cert B_A pc | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 189 | have cert_suc: "c!Suc pc \<in> A" by (rule cert_okD3) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 190 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 191 | from phi pc have "\<phi>!pc \<in> A" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 192 | from pres this pc | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 193 | have stepA: "snd`set ?step \<subseteq> A" by (rule pres_typeD2) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 194 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 195 | have "merge c pc ?step (c!Suc pc) = | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 196 | (if \<forall>(pc',s')\<in>set ?step. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc' | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 197 | then map snd [(p',t') \<leftarrow> ?step.p'=pc+1] ++_f c!Suc pc | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 198 | else \<top>)" unfolding mrg_def by (rule lbv.merge_def [OF lbvc.axioms(1), OF lbvc_axioms]) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 199 |   moreover {
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 200 | fix pc' s' assume s': "(pc', s') \<in> set ?step" and suc_pc: "pc' \<noteq> pc+1" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 201 | with less have "s' <=_r \<phi>!pc'" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 202 | also | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 203 | from bounded pc s' have "pc' < length \<phi>" by (rule boundedD) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 204 | with s' suc_pc pc have "c!pc' = \<phi>!pc'" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 205 | hence "\<phi>!pc' = c!pc'" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 206 | finally have "s' <=_r c!pc'" . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 207 | } hence "\<forall>(pc',s')\<in>set ?step. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 208 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 209 | from pc have "Suc pc = length \<phi> \<or> Suc pc < length \<phi>" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 210 | hence "map snd [(p',t') \<leftarrow> ?step.p'=pc+1] ++_f c!Suc pc \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 211 | (is "?map ++_f _ \<noteq> _") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 212 | proof (rule disjE) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 213 | assume pc': "Suc pc = length \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 214 | with cert have "c!Suc pc = \<bottom>" by (simp add: cert_okD2) | 
| 
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 pc' bounded pc | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 217 | have "\<forall>(p',t')\<in>set ?step. p'\<noteq>pc+1" by clarify (drule boundedD, auto) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 218 | hence "[(p',t') \<leftarrow> ?step.p'=pc+1] = []" by (blast intro: filter_False) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 219 | hence "?map = []" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 220 | ultimately show ?thesis by (simp add: B_neq_T) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 221 | next | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 222 | assume pc': "Suc pc < length \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 223 | from pc' phi have "\<phi>!Suc pc \<in> A" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 224 | moreover note cert_suc | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 225 | moreover from stepA | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 226 | have "set ?map \<subseteq> A" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 227 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 228 | have "\<And>s. s \<in> set ?map \<Longrightarrow> \<exists>t. (Suc pc, t) \<in> set ?step" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 229 | with less have "\<forall>s' \<in> set ?map. s' <=_r \<phi>!Suc pc" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 230 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 231 | from pc' have "c!Suc pc <=_r \<phi>!Suc pc" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 232 | by (cases "c!Suc pc = \<bottom>") (auto dest: cert_approx) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 233 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 234 | have "?map ++_f c!Suc pc <=_r \<phi>!Suc pc" by (rule pp_lub) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 235 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 236 | from pc' phi have "\<phi>!Suc pc \<noteq> \<top>" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 237 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 238 | show ?thesis by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 239 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 240 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 241 | have "merge c pc ?step (c!Suc pc) \<noteq> \<top>" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 242 | thus ?thesis by (simp add: wti) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 243 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 244 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 245 | lemma (in lbvc) wti_less: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 246 | assumes stable: "stable r step \<phi> pc" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 247 | assumes suc_pc: "Suc pc < length \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 248 | shows "wti c pc (\<phi>!pc) <=_r \<phi>!Suc pc" (is "?wti <=_r _") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 249 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 250 | let ?step = "step pc (\<phi>!pc)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 251 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 252 | from stable | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 253 | have less: "\<forall>(q,s')\<in>set ?step. s' <=_r \<phi>!q" by (simp add: stable_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 254 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 255 | from suc_pc have pc: "pc < length \<phi>" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 256 | with cert B_A have cert_suc: "c!Suc pc \<in> A" by (rule cert_okD3) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 257 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 258 | from phi pc have "\<phi>!pc \<in> A" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 259 | with pres pc have stepA: "snd`set ?step \<subseteq> A" by - (rule pres_typeD2) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 260 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 261 | from stable pc have "?wti \<noteq> \<top>" by (rule stable_wti) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 262 | hence "merge c pc ?step (c!Suc pc) \<noteq> \<top>" by (simp add: wti) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 263 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 264 | have "merge c pc ?step (c!Suc pc) = | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 265 | map snd [(p',t')\<leftarrow> ?step.p'=pc+1] ++_f c!Suc pc" by (rule merge_not_top_s) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 266 | hence "?wti = \<dots>" (is "_ = (?map ++_f _)" is "_ = ?sum") by (simp add: wti) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 267 |   also {
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 268 | from suc_pc phi have "\<phi>!Suc pc \<in> A" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 269 | moreover note cert_suc | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 270 | moreover from stepA have "set ?map \<subseteq> A" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 271 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 272 | have "\<And>s. s \<in> set ?map \<Longrightarrow> \<exists>t. (Suc pc, t) \<in> set ?step" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 273 | with less have "\<forall>s' \<in> set ?map. s' <=_r \<phi>!Suc pc" by auto | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 274 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 275 | from suc_pc have "c!Suc pc <=_r \<phi>!Suc pc" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 276 | by (cases "c!Suc pc = \<bottom>") (auto dest: cert_approx) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 277 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 278 | have "?sum <=_r \<phi>!Suc pc" by (rule pp_lub) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 279 | } | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 280 | finally show ?thesis . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 281 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 282 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 283 | lemma (in lbvc) stable_wtc: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 284 | assumes stable: "stable r step phi pc" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 285 | assumes pc: "pc < length \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 286 | shows "wtc c pc (\<phi>!pc) \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 287 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 288 | from stable pc have wti: "wti c pc (\<phi>!pc) \<noteq> \<top>" by (rule stable_wti) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 289 | show ?thesis | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 290 | proof (cases "c!pc = \<bottom>") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 291 | case True with wti show ?thesis by (simp add: wtc) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 292 | next | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 293 | case False | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 294 | with pc have "c!pc = \<phi>!pc" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 295 | with False wti show ?thesis by (simp add: wtc) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 296 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 297 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 298 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 299 | lemma (in lbvc) wtc_less: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 300 | assumes stable: "stable r step \<phi> pc" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 301 | assumes suc_pc: "Suc pc < length \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 302 | shows "wtc c pc (\<phi>!pc) <=_r \<phi>!Suc pc" (is "?wtc <=_r _") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 303 | proof (cases "c!pc = \<bottom>") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 304 | case True | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 305 | moreover from stable suc_pc have "wti c pc (\<phi>!pc) <=_r \<phi>!Suc pc" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 306 | by (rule wti_less) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 307 | ultimately show ?thesis by (simp add: wtc) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 308 | next | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 309 | case False | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 310 | from suc_pc have pc: "pc < length \<phi>" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 311 | with stable have "?wtc \<noteq> \<top>" by (rule stable_wtc) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 312 | with False have "?wtc = wti c pc (c!pc)" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 313 | by (unfold wtc) (simp split: split_if_asm) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 314 | also from pc False have "c!pc = \<phi>!pc" .. | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 315 | finally have "?wtc = wti c pc (\<phi>!pc)" . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 316 | also from stable suc_pc have "wti c pc (\<phi>!pc) <=_r \<phi>!Suc pc" by (rule wti_less) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 317 | finally show ?thesis . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 318 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 319 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 320 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 321 | lemma (in lbvc) wt_step_wtl_lemma: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 322 | assumes wt_step: "wt_step r \<top> step \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 323 | shows "\<And>pc s. pc+length ls = length \<phi> \<Longrightarrow> s <=_r \<phi>!pc \<Longrightarrow> s \<in> A \<Longrightarrow> s\<noteq>\<top> \<Longrightarrow> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 324 | wtl ls c pc s \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 325 | (is "\<And>pc s. _ \<Longrightarrow> _ \<Longrightarrow> _ \<Longrightarrow> _ \<Longrightarrow> ?wtl ls pc s \<noteq> _") | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 326 | proof (induct ls) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 327 | fix pc s assume "s\<noteq>\<top>" thus "?wtl [] pc s \<noteq> \<top>" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 328 | next | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 329 | fix pc s i ls | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 330 | assume "\<And>pc s. pc+length ls=length \<phi> \<Longrightarrow> s <=_r \<phi>!pc \<Longrightarrow> s \<in> A \<Longrightarrow> s\<noteq>\<top> \<Longrightarrow> | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 331 | ?wtl ls pc s \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 332 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 333 | assume pc_l: "pc + length (i#ls) = length \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 334 | hence suc_pc_l: "Suc pc + length ls = length \<phi>" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 335 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 336 | have IH: "\<And>s. s <=_r \<phi>!Suc pc \<Longrightarrow> s \<in> A \<Longrightarrow> s \<noteq> \<top> \<Longrightarrow> ?wtl ls (Suc pc) s \<noteq> \<top>" . | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 337 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 338 | from pc_l obtain pc: "pc < length \<phi>" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 339 | with wt_step have stable: "stable r step \<phi> pc" by (simp add: wt_step_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 340 | from this pc have wt_phi: "wtc c pc (\<phi>!pc) \<noteq> \<top>" by (rule stable_wtc) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 341 | assume s_phi: "s <=_r \<phi>!pc" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 342 | from phi pc have phi_pc: "\<phi>!pc \<in> A" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 343 | assume s: "s \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 344 | with s_phi pc phi_pc have wt_s_phi: "wtc c pc s <=_r wtc c pc (\<phi>!pc)" by (rule wtc_mono) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 345 | with wt_phi have wt_s: "wtc c pc s \<noteq> \<top>" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 346 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 347 | assume s': "s \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 348 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 349 | have "ls = [] \<Longrightarrow> ?wtl (i#ls) pc s \<noteq> \<top>" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 350 |   moreover {
 | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 351 | assume "ls \<noteq> []" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 352 | with pc_l have suc_pc: "Suc pc < length \<phi>" by (auto simp add: neq_Nil_conv) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 353 | with stable have "wtc c pc (phi!pc) <=_r \<phi>!Suc pc" by (rule wtc_less) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 354 | with wt_s_phi have "wtc c pc s <=_r \<phi>!Suc pc" by (rule trans_r) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 355 | moreover | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 356 | from cert suc_pc have "c!pc \<in> A" "c!(pc+1) \<in> A" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 357 | by (auto simp add: cert_ok_def) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 358 | from pres this s pc have "wtc c pc s \<in> A" by (rule wtc_pres) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 359 | ultimately | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 360 | have "?wtl ls (Suc pc) (wtc c pc s) \<noteq> \<top>" using IH wt_s by blast | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 361 | with s' wt_s have "?wtl (i#ls) pc s \<noteq> \<top>" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 362 | } | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 363 | ultimately show "?wtl (i#ls) pc s \<noteq> \<top>" by (cases ls) blast+ | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 364 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 365 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 366 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 367 | theorem (in lbvc) wtl_complete: | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 368 | assumes wt: "wt_step r \<top> step \<phi>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 369 | and s: "s <=_r \<phi>!0" "s \<in> A" "s \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 370 | and len: "length ins = length phi" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 371 | shows "wtl ins c 0 s \<noteq> \<top>" | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 372 | proof - | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 373 | from len have "0+length ins = length phi" by simp | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 374 | from wt this s show ?thesis by (rule wt_step_wtl_lemma) | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 375 | qed | 
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 376 | |
| 
1bc3b688548c
backported parts of abstract byte code verifier from AFP/Jinja
 haftmann parents: diff
changeset | 377 | end |