| author | hoelzl | 
| Tue, 05 Jul 2016 20:29:58 +0200 | |
| changeset 63393 | c22928719e19 | 
| parent 62390 | 842917225d56 | 
| child 63540 | f8652d0534fa | 
| permissions | -rw-r--r-- | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 1 | (* Author: Gerwin Klein *) | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 2 | |
| 52400 | 3 | theory Compiler2 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 4 | imports Compiler | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 5 | begin | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 6 | |
| 56224 | 7 | text {*
 | 
| 8 | The preservation of the source code semantics is already shown in the | |
| 9 | parent theory @{theory Compiler}. This here shows the second direction.
 | |
| 10 | *} | |
| 11 | ||
| 50061 | 12 | section {* Compiler Correctness, Reverse Direction *}
 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 13 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 14 | subsection {* Definitions *}
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 15 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 16 | text {* Execution in @{term n} steps for simpler induction *}
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 17 | primrec | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 18 | exec_n :: "instr list \<Rightarrow> config \<Rightarrow> nat \<Rightarrow> config \<Rightarrow> bool" | 
| 44000 
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
 kleing parents: 
43438diff
changeset | 19 |   ("_/ \<turnstile> (_ \<rightarrow>^_/ _)" [65,0,1000,55] 55)
 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 20 | where | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 21 | "P \<turnstile> c \<rightarrow>^0 c' = (c'=c)" | | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 22 | "P \<turnstile> c \<rightarrow>^(Suc n) c'' = (\<exists>c'. (P \<turnstile> c \<rightarrow> c') \<and> P \<turnstile> c' \<rightarrow>^n c'')" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 23 | |
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 24 | text {* The possible successor PCs of an instruction at position @{term n} *}
 | 
| 51705 | 25 | text_raw{*\snip{isuccsdef}{0}{1}{% *}
 | 
| 53911 | 26 | definition isuccs :: "instr \<Rightarrow> int \<Rightarrow> int set" where | 
| 27 | "isuccs i n = (case i of | |
| 28 |   JMP j \<Rightarrow> {n + 1 + j} |
 | |
| 29 |   JMPLESS j \<Rightarrow> {n + 1 + j, n + 1} |
 | |
| 30 |   JMPGE j \<Rightarrow> {n + 1 + j, n + 1} |
 | |
| 31 |   _ \<Rightarrow> {n +1})"
 | |
| 51705 | 32 | text_raw{*}%endsnip*}
 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 33 | |
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 34 | text {* The possible successors PCs of an instruction list *}
 | 
| 53911 | 35 | definition succs :: "instr list \<Rightarrow> int \<Rightarrow> int set" where | 
| 36 | "succs P n = {s. \<exists>i::int. 0 \<le> i \<and> i < size P \<and> s \<in> isuccs (P!!i) (n+i)}" 
 | |
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 37 | |
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 38 | text {* Possible exit PCs of a program *}
 | 
| 53911 | 39 | definition exits :: "instr list \<Rightarrow> int set" where | 
| 40 | "exits P = succs P 0 - {0..< size P}"
 | |
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 41 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 42 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 43 | subsection {* Basic properties of @{term exec_n} *}
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 44 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 45 | lemma exec_n_exec: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 46 | "P \<turnstile> c \<rightarrow>^n c' \<Longrightarrow> P \<turnstile> c \<rightarrow>* c'" | 
| 61147 | 47 | by (induct n arbitrary: c) (auto intro: star.step) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 48 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 49 | lemma exec_0 [intro!]: "P \<turnstile> c \<rightarrow>^0 c" by simp | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 50 | |
| 45218 | 51 | lemma exec_Suc: | 
| 44000 
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
 kleing parents: 
43438diff
changeset | 52 | "\<lbrakk> P \<turnstile> c \<rightarrow> c'; P \<turnstile> c' \<rightarrow>^n c'' \<rbrakk> \<Longrightarrow> P \<turnstile> c \<rightarrow>^(Suc n) c''" | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44070diff
changeset | 53 | by (fastforce simp del: split_paired_Ex) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 54 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 55 | lemma exec_exec_n: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 56 | "P \<turnstile> c \<rightarrow>* c' \<Longrightarrow> \<exists>n. P \<turnstile> c \<rightarrow>^n c'" | 
| 52915 | 57 | by (induct rule: star.induct) (auto intro: exec_Suc) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 58 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 59 | lemma exec_eq_exec_n: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 60 | "(P \<turnstile> c \<rightarrow>* c') = (\<exists>n. P \<turnstile> c \<rightarrow>^n c')" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 61 | by (blast intro: exec_exec_n exec_n_exec) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 62 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 63 | lemma exec_n_Nil [simp]: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 64 | "[] \<turnstile> c \<rightarrow>^k c' = (c' = c \<and> k = 0)" | 
| 52915 | 65 | by (induct k) (auto simp: exec1_def) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 66 | |
| 44004 | 67 | lemma exec1_exec_n [intro!]: | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 68 | "P \<turnstile> c \<rightarrow> c' \<Longrightarrow> P \<turnstile> c \<rightarrow>^1 c'" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 69 | by (cases c') simp | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 70 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 71 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 72 | subsection {* Concrete symbolic execution steps *}
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 73 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 74 | lemma exec_n_step: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 75 | "n \<noteq> n' \<Longrightarrow> | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 76 | P \<turnstile> (n,stk,s) \<rightarrow>^k (n',stk',s') = | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 77 | (\<exists>c. P \<turnstile> (n,stk,s) \<rightarrow> c \<and> P \<turnstile> c \<rightarrow>^(k - 1) (n',stk',s') \<and> 0 < k)" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 78 | by (cases k) auto | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 79 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 80 | lemma exec1_end: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 81 | "size P <= fst c \<Longrightarrow> \<not> P \<turnstile> c \<rightarrow> c'" | 
| 52915 | 82 | by (auto simp: exec1_def) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 83 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 84 | lemma exec_n_end: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 85 | "size P <= (n::int) \<Longrightarrow> | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 86 | P \<turnstile> (n,s,stk) \<rightarrow>^k (n',s',stk') = (n' = n \<and> stk'=stk \<and> s'=s \<and> k =0)" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 87 | by (cases k) (auto simp: exec1_end) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 88 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 89 | lemmas exec_n_simps = exec_n_step exec_n_end | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 90 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 91 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 92 | subsection {* Basic properties of @{term succs} *}
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 93 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 94 | lemma succs_simps [simp]: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 95 |   "succs [ADD] n = {n + 1}"
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 96 |   "succs [LOADI v] n = {n + 1}"
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 97 |   "succs [LOAD x] n = {n + 1}"
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 98 |   "succs [STORE x] n = {n + 1}"
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 99 |   "succs [JMP i] n = {n + 1 + i}"
 | 
| 45322 
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
 kleing parents: 
45218diff
changeset | 100 |   "succs [JMPGE i] n = {n + 1 + i, n + 1}"
 | 
| 
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
 kleing parents: 
45218diff
changeset | 101 |   "succs [JMPLESS i] n = {n + 1 + i, n + 1}"
 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 102 | by (auto simp: succs_def isuccs_def) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 103 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 104 | lemma succs_empty [iff]: "succs [] n = {}"
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 105 | by (simp add: succs_def) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 106 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 107 | lemma succs_Cons: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 108 | "succs (x#xs) n = isuccs x n \<union> succs xs (1+n)" (is "_ = ?x \<union> ?xs") | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 109 | proof | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 110 | let ?isuccs = "\<lambda>p P n i::int. 0 \<le> i \<and> i < size P \<and> p \<in> isuccs (P!!i) (n+i)" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 111 |   { fix p assume "p \<in> succs (x#xs) n"
 | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 112 | then obtain i::int where isuccs: "?isuccs p (x#xs) n i" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 113 | unfolding succs_def by auto | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 114 | have "p \<in> ?x \<union> ?xs" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 115 | proof cases | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 116 | assume "i = 0" with isuccs show ?thesis by simp | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 117 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 118 | assume "i \<noteq> 0" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 119 | with isuccs | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 120 | have "?isuccs p xs (1+n) (i - 1)" by auto | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 121 | hence "p \<in> ?xs" unfolding succs_def by blast | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 122 | thus ?thesis .. | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 123 | qed | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 124 | } | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 125 | thus "succs (x#xs) n \<subseteq> ?x \<union> ?xs" .. | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 126 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 127 |   { fix p assume "p \<in> ?x \<or> p \<in> ?xs"
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 128 | hence "p \<in> succs (x#xs) n" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 129 | proof | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44070diff
changeset | 130 | assume "p \<in> ?x" thus ?thesis by (fastforce simp: succs_def) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 131 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 132 | assume "p \<in> ?xs" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 133 | then obtain i where "?isuccs p xs (1+n) i" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 134 | unfolding succs_def by auto | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 135 | hence "?isuccs p (x#xs) n (1+i)" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 136 | by (simp add: algebra_simps) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 137 | thus ?thesis unfolding succs_def by blast | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 138 | qed | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 139 | } | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 140 | thus "?x \<union> ?xs \<subseteq> succs (x#xs) n" by blast | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 141 | qed | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 142 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 143 | lemma succs_iexec1: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 144 | assumes "c' = iexec (P!!i) (i,s,stk)" "0 \<le> i" "i < size P" | 
| 44004 | 145 | shows "fst c' \<in> succs P 0" | 
| 50060 
43753eca324a
replaced relation by function - simplifies development
 nipkow parents: 
47818diff
changeset | 146 | using assms by (auto simp: succs_def isuccs_def split: instr.split) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 147 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 148 | lemma succs_shift: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 149 | "(p - n \<in> succs P 0) = (p \<in> succs P n)" | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44070diff
changeset | 150 | by (fastforce simp: succs_def isuccs_def split: instr.split) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 151 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 152 | lemma inj_op_plus [simp]: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 153 | "inj (op + (i::int))" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 154 | by (metis add_minus_cancel inj_on_inverseI) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 155 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 156 | lemma succs_set_shift [simp]: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 157 | "op + i ` succs xs 0 = succs xs i" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 158 | by (force simp: succs_shift [where n=i, symmetric] intro: set_eqI) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 159 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 160 | lemma succs_append [simp]: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 161 | "succs (xs @ ys) n = succs xs n \<union> succs ys (n + size xs)" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 162 | by (induct xs arbitrary: n) (auto simp: succs_Cons algebra_simps) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 163 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 164 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 165 | lemma exits_append [simp]: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 166 | "exits (xs @ ys) = exits xs \<union> (op + (size xs)) ` exits ys - | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 167 |                      {0..<size xs + size ys}" 
 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 168 | by (auto simp: exits_def image_set_diff) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 169 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 170 | lemma exits_single: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 171 |   "exits [x] = isuccs x 0 - {0}"
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 172 | by (auto simp: exits_def succs_def) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 173 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 174 | lemma exits_Cons: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 175 |   "exits (x # xs) = (isuccs x 0 - {0}) \<union> (op + 1) ` exits xs - 
 | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 176 |                      {0..<1 + size xs}" 
 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 177 | using exits_append [of "[x]" xs] | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 178 | by (simp add: exits_single) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 179 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 180 | lemma exits_empty [iff]: "exits [] = {}" by (simp add: exits_def)
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 181 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 182 | lemma exits_simps [simp]: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 183 |   "exits [ADD] = {1}"
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 184 |   "exits [LOADI v] = {1}"
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 185 |   "exits [LOAD x] = {1}"
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 186 |   "exits [STORE x] = {1}"
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 187 |   "i \<noteq> -1 \<Longrightarrow> exits [JMP i] = {1 + i}"
 | 
| 45322 
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
 kleing parents: 
45218diff
changeset | 188 |   "i \<noteq> -1 \<Longrightarrow> exits [JMPGE i] = {1 + i, 1}"
 | 
| 
654cc47f6115
JMPF(LESS|GE) -> JMP(LESS|GE) because jumps are int now.
 kleing parents: 
45218diff
changeset | 189 |   "i \<noteq> -1 \<Longrightarrow> exits [JMPLESS i] = {1 + i, 1}"
 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 190 | by (auto simp: exits_def) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 191 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 192 | lemma acomp_succs [simp]: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 193 |   "succs (acomp a) n = {n + 1 .. n + size (acomp a)}"
 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 194 | by (induct a arbitrary: n) auto | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 195 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 196 | lemma acomp_size: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 197 | "(1::int) \<le> size (acomp a)" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 198 | by (induct a) auto | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 199 | |
| 44010 | 200 | lemma acomp_exits [simp]: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 201 |   "exits (acomp a) = {size (acomp a)}"
 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 202 | by (auto simp: exits_def acomp_size) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 203 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 204 | lemma bcomp_succs: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 205 | "0 \<le> i \<Longrightarrow> | 
| 53880 | 206 |   succs (bcomp b f i) n \<subseteq> {n .. n + size (bcomp b f i)}
 | 
| 207 |                            \<union> {n + i + size (bcomp b f i)}" 
 | |
| 208 | proof (induction b arbitrary: f i n) | |
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 209 | case (And b1 b2) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 210 | from And.prems | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 211 | show ?case | 
| 53880 | 212 | by (cases f) | 
| 45015 | 213 | (auto dest: And.IH(1) [THEN subsetD, rotated] | 
| 214 | And.IH(2) [THEN subsetD, rotated]) | |
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 215 | qed auto | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 216 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 217 | lemmas bcomp_succsD [dest!] = bcomp_succs [THEN subsetD, rotated] | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 218 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 219 | lemma bcomp_exits: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 220 | fixes i :: int | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 221 | shows | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 222 | "0 \<le> i \<Longrightarrow> | 
| 53880 | 223 |   exits (bcomp b f i) \<subseteq> {size (bcomp b f i), i + size (bcomp b f i)}" 
 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 224 | by (auto simp: exits_def) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 225 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 226 | lemma bcomp_exitsD [dest!]: | 
| 53880 | 227 | "p \<in> exits (bcomp b f i) \<Longrightarrow> 0 \<le> i \<Longrightarrow> | 
| 228 | p = size (bcomp b f i) \<or> p = i + size (bcomp b f i)" | |
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 229 | using bcomp_exits by auto | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 230 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 231 | lemma ccomp_succs: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 232 |   "succs (ccomp c) n \<subseteq> {n..n + size (ccomp c)}"
 | 
| 45015 | 233 | proof (induction c arbitrary: n) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 234 | case SKIP thus ?case by simp | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 235 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 236 | case Assign thus ?case by simp | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 237 | next | 
| 47818 | 238 | case (Seq c1 c2) | 
| 239 | from Seq.prems | |
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 240 | show ?case | 
| 47818 | 241 | by (fastforce dest: Seq.IH [THEN subsetD]) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 242 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 243 | case (If b c1 c2) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 244 | from If.prems | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 245 | show ?case | 
| 45015 | 246 | by (auto dest!: If.IH [THEN subsetD] simp: isuccs_def succs_Cons) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 247 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 248 | case (While b c) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 249 | from While.prems | 
| 45015 | 250 | show ?case by (auto dest!: While.IH [THEN subsetD]) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 251 | qed | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 252 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 253 | lemma ccomp_exits: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 254 |   "exits (ccomp c) \<subseteq> {size (ccomp c)}"
 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 255 | using ccomp_succs [of c 0] by (auto simp: exits_def) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 256 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 257 | lemma ccomp_exitsD [dest!]: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 258 | "p \<in> exits (ccomp c) \<Longrightarrow> p = size (ccomp c)" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 259 | using ccomp_exits by auto | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 260 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 261 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 262 | subsection {* Splitting up machine executions *}
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 263 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 264 | lemma exec1_split: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 265 | fixes i j :: int | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 266 | shows | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 267 | "P @ c @ P' \<turnstile> (size P + i, s) \<rightarrow> (j,s') \<Longrightarrow> 0 \<le> i \<Longrightarrow> i < size c \<Longrightarrow> | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 268 | c \<turnstile> (i,s) \<rightarrow> (j - size P, s')" | 
| 52915 | 269 | by (auto split: instr.splits simp: exec1_def) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 270 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 271 | lemma exec_n_split: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 272 | fixes i j :: int | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 273 | assumes "P @ c @ P' \<turnstile> (size P + i, s) \<rightarrow>^n (j, s')" | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 274 | "0 \<le> i" "i < size c" | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 275 |           "j \<notin> {size P ..< size P + size c}"
 | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 276 | shows "\<exists>s'' (i'::int) k m. | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 277 | c \<turnstile> (i, s) \<rightarrow>^k (i', s'') \<and> | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 278 | i' \<in> exits c \<and> | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 279 | P @ c @ P' \<turnstile> (size P + i', s'') \<rightarrow>^m (j, s') \<and> | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 280 | n = k + m" | 
| 45015 | 281 | using assms proof (induction n arbitrary: i j s) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 282 | case 0 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 283 | thus ?case by simp | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 284 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 285 | case (Suc n) | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 286 | have i: "0 \<le> i" "i < size c" by fact+ | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 287 | from Suc.prems | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 288 | have j: "\<not> (size P \<le> j \<and> j < size P + size c)" by simp | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 289 | from Suc.prems | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 290 | obtain i0 s0 where | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 291 | step: "P @ c @ P' \<turnstile> (size P + i, s) \<rightarrow> (i0,s0)" and | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 292 | rest: "P @ c @ P' \<turnstile> (i0,s0) \<rightarrow>^n (j, s')" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 293 | by clarsimp | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 294 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 295 | from step i | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 296 | have c: "c \<turnstile> (i,s) \<rightarrow> (i0 - size P, s0)" by (rule exec1_split) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 297 | |
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 298 | have "i0 = size P + (i0 - size P) " by simp | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 299 | then obtain j0::int where j0: "i0 = size P + j0" .. | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 300 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 301 | note split_paired_Ex [simp del] | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 302 | |
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 303 |   { assume "j0 \<in> {0 ..< size c}"
 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 304 | with j0 j rest c | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 305 | have ?case | 
| 45015 | 306 | by (fastforce dest!: Suc.IH intro!: exec_Suc) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 307 |   } moreover {
 | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 308 |     assume "j0 \<notin> {0 ..< size c}"
 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 309 | moreover | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 310 | from c j0 have "j0 \<in> succs c 0" | 
| 52915 | 311 | by (auto dest: succs_iexec1 simp: exec1_def simp del: iexec.simps) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 312 | ultimately | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 313 | have "j0 \<in> exits c" by (simp add: exits_def) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 314 | with c j0 rest | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44070diff
changeset | 315 | have ?case by fastforce | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 316 | } | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 317 | ultimately | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 318 | show ?case by cases | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 319 | qed | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 320 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 321 | lemma exec_n_drop_right: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 322 | fixes j :: int | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 323 |   assumes "c @ P' \<turnstile> (0, s) \<rightarrow>^n (j, s')" "j \<notin> {0..<size c}"
 | 
| 44004 | 324 | shows "\<exists>s'' i' k m. | 
| 325 | (if c = [] then s'' = s \<and> i' = 0 \<and> k = 0 | |
| 326 | else c \<turnstile> (0, s) \<rightarrow>^k (i', s'') \<and> | |
| 327 | i' \<in> exits c) \<and> | |
| 328 | c @ P' \<turnstile> (i', s'') \<rightarrow>^m (j, s') \<and> | |
| 329 | n = k + m" | |
| 330 | using assms | |
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 331 | by (cases "c = []") | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 332 | (auto dest: exec_n_split [where P="[]", simplified]) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 333 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 334 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 335 | text {*
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 336 |   Dropping the left context of a potentially incomplete execution of @{term c}.
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 337 | *} | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 338 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 339 | lemma exec1_drop_left: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 340 | fixes i n :: int | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 341 | assumes "P1 @ P2 \<turnstile> (i, s, stk) \<rightarrow> (n, s', stk')" and "size P1 \<le> i" | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 342 | shows "P2 \<turnstile> (i - size P1, s, stk) \<rightarrow> (n - size P1, s', stk')" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 343 | proof - | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 344 | have "i = size P1 + (i - size P1)" by simp | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 345 | then obtain i' :: int where "i = size P1 + i'" .. | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 346 | moreover | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 347 | have "n = size P1 + (n - size P1)" by simp | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 348 | then obtain n' :: int where "n = size P1 + n'" .. | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 349 | ultimately | 
| 52915 | 350 | show ?thesis using assms | 
| 351 | by (clarsimp simp: exec1_def simp del: iexec.simps) | |
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 352 | qed | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 353 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 354 | lemma exec_n_drop_left: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 355 | fixes i n :: int | 
| 44004 | 356 | assumes "P @ P' \<turnstile> (i, s, stk) \<rightarrow>^k (n, s', stk')" | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 357 |           "size P \<le> i" "exits P' \<subseteq> {0..}"
 | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 358 | shows "P' \<turnstile> (i - size P, s, stk) \<rightarrow>^k (n - size P, s', stk')" | 
| 45015 | 359 | using assms proof (induction k arbitrary: i s stk) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 360 | case 0 thus ?case by simp | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 361 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 362 | case (Suc k) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 363 | from Suc.prems | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 364 | obtain i' s'' stk'' where | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 365 | step: "P @ P' \<turnstile> (i, s, stk) \<rightarrow> (i', s'', stk'')" and | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 366 | rest: "P @ P' \<turnstile> (i', s'', stk'') \<rightarrow>^k (n, s', stk')" | 
| 53356 | 367 | by auto | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 368 | from step `size P \<le> i` | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 369 | have "P' \<turnstile> (i - size P, s, stk) \<rightarrow> (i' - size P, s'', stk'')" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 370 | by (rule exec1_drop_left) | 
| 45218 | 371 | moreover | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 372 | then have "i' - size P \<in> succs P' 0" | 
| 52915 | 373 | by (fastforce dest!: succs_iexec1 simp: exec1_def simp del: iexec.simps) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 374 |   with `exits P' \<subseteq> {0..}`
 | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 375 | have "size P \<le> i'" by (auto simp: exits_def) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 376 |   from rest this `exits P' \<subseteq> {0..}`     
 | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 377 | have "P' \<turnstile> (i' - size P, s'', stk'') \<rightarrow>^k (n - size P, s', stk')" | 
| 45015 | 378 | by (rule Suc.IH) | 
| 45218 | 379 | ultimately | 
| 380 | show ?case by auto | |
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 381 | qed | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 382 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 383 | lemmas exec_n_drop_Cons = | 
| 45605 | 384 | exec_n_drop_left [where P="[instr]", simplified] for instr | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 385 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 386 | definition | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 387 |   "closed P \<longleftrightarrow> exits P \<subseteq> {size P}" 
 | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 388 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 389 | lemma ccomp_closed [simp, intro!]: "closed (ccomp c)" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 390 | using ccomp_exits by (auto simp: closed_def) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 391 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 392 | lemma acomp_closed [simp, intro!]: "closed (acomp c)" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 393 | by (simp add: closed_def) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 394 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 395 | lemma exec_n_split_full: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 396 | fixes j :: int | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 397 | assumes exec: "P @ P' \<turnstile> (0,s,stk) \<rightarrow>^k (j, s', stk')" | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 398 | assumes P: "size P \<le> j" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 399 | assumes closed: "closed P" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 400 |   assumes exits: "exits P' \<subseteq> {0..}"
 | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 401 | shows "\<exists>k1 k2 s'' stk''. P \<turnstile> (0,s,stk) \<rightarrow>^k1 (size P, s'', stk'') \<and> | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 402 | P' \<turnstile> (0,s'',stk'') \<rightarrow>^k2 (j - size P, s', stk')" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 403 | proof (cases "P") | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 404 | case Nil with exec | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44070diff
changeset | 405 | show ?thesis by fastforce | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 406 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 407 | case Cons | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 408 | hence "0 < size P" by simp | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 409 | with exec P closed | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 410 | obtain k1 k2 s'' stk'' where | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 411 | 1: "P \<turnstile> (0,s,stk) \<rightarrow>^k1 (size P, s'', stk'')" and | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 412 | 2: "P @ P' \<turnstile> (size P,s'',stk'') \<rightarrow>^k2 (j, s', stk')" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 413 | by (auto dest!: exec_n_split [where P="[]" and i=0, simplified] | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 414 | simp: closed_def) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 415 | moreover | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 416 | have "j = size P + (j - size P)" by simp | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 417 | then obtain j0 :: int where "j = size P + j0" .. | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 418 | ultimately | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 419 | show ?thesis using exits | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44070diff
changeset | 420 | by (fastforce dest: exec_n_drop_left) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 421 | qed | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 422 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 423 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 424 | subsection {* Correctness theorem *}
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 425 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 426 | lemma acomp_neq_Nil [simp]: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 427 | "acomp a \<noteq> []" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 428 | by (induct a) auto | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 429 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 430 | lemma acomp_exec_n [dest!]: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 431 | "acomp a \<turnstile> (0,s,stk) \<rightarrow>^n (size (acomp a),s',stk') \<Longrightarrow> | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 432 | s' = s \<and> stk' = aval a s#stk" | 
| 45015 | 433 | proof (induction a arbitrary: n s' stk stk') | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 434 | case (Plus a1 a2) | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 435 | let ?sz = "size (acomp a1) + (size (acomp a2) + 1)" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 436 | from Plus.prems | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 437 | have "acomp a1 @ acomp a2 @ [ADD] \<turnstile> (0,s,stk) \<rightarrow>^n (?sz, s', stk')" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 438 | by (simp add: algebra_simps) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 439 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 440 | then obtain n1 s1 stk1 n2 s2 stk2 n3 where | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 441 | "acomp a1 \<turnstile> (0,s,stk) \<rightarrow>^n1 (size (acomp a1), s1, stk1)" | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 442 | "acomp a2 \<turnstile> (0,s1,stk1) \<rightarrow>^n2 (size (acomp a2), s2, stk2)" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 443 | "[ADD] \<turnstile> (0,s2,stk2) \<rightarrow>^n3 (1, s', stk')" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 444 | by (auto dest!: exec_n_split_full) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 445 | |
| 52915 | 446 | thus ?case by (fastforce dest: Plus.IH simp: exec_n_simps exec1_def) | 
| 447 | qed (auto simp: exec_n_simps exec1_def) | |
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 448 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 449 | lemma bcomp_split: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 450 | fixes i j :: int | 
| 53880 | 451 | assumes "bcomp b f i @ P' \<turnstile> (0, s, stk) \<rightarrow>^n (j, s', stk')" | 
| 452 |           "j \<notin> {0..<size (bcomp b f i)}" "0 \<le> i"
 | |
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 453 | shows "\<exists>s'' stk'' (i'::int) k m. | 
| 53880 | 454 | bcomp b f i \<turnstile> (0, s, stk) \<rightarrow>^k (i', s'', stk'') \<and> | 
| 455 | (i' = size (bcomp b f i) \<or> i' = i + size (bcomp b f i)) \<and> | |
| 456 | bcomp b f i @ P' \<turnstile> (i', s'', stk'') \<rightarrow>^m (j, s', stk') \<and> | |
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 457 | n = k + m" | 
| 53880 | 458 | using assms by (cases "bcomp b f i = []") (fastforce dest!: exec_n_drop_right)+ | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 459 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 460 | lemma bcomp_exec_n [dest]: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 461 | fixes i j :: int | 
| 53880 | 462 | assumes "bcomp b f j \<turnstile> (0, s, stk) \<rightarrow>^n (i, s', stk')" | 
| 463 | "size (bcomp b f j) \<le> i" "0 \<le> j" | |
| 464 | shows "i = size(bcomp b f j) + (if f = bval b s then j else 0) \<and> | |
| 44004 | 465 | s' = s \<and> stk' = stk" | 
| 53880 | 466 | using assms proof (induction b arbitrary: f j i n s' stk') | 
| 45200 | 467 | case Bc thus ?case | 
| 62390 | 468 | by (simp split: if_split_asm add: exec_n_simps exec1_def) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 469 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 470 | case (Not b) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 471 | from Not.prems show ?case | 
| 45015 | 472 | by (fastforce dest!: Not.IH) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 473 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 474 | case (And b1 b2) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 475 | |
| 53880 | 476 | let ?b2 = "bcomp b2 f j" | 
| 477 | let ?m = "if f then size ?b2 else size ?b2 + j" | |
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 478 | let ?b1 = "bcomp b1 False ?m" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 479 | |
| 53880 | 480 | have j: "size (bcomp (And b1 b2) f j) \<le> i" "0 \<le> j" by fact+ | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 481 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 482 | from And.prems | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 483 | obtain s'' stk'' and i'::int and k m where | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 484 | b1: "?b1 \<turnstile> (0, s, stk) \<rightarrow>^k (i', s'', stk'')" | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 485 | "i' = size ?b1 \<or> i' = ?m + size ?b1" and | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 486 | b2: "?b2 \<turnstile> (i' - size ?b1, s'', stk'') \<rightarrow>^m (i - size ?b1, s', stk')" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 487 | by (auto dest!: bcomp_split dest: exec_n_drop_left) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 488 | from b1 j | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 489 | have "i' = size ?b1 + (if \<not>bval b1 s then ?m else 0) \<and> s'' = s \<and> stk'' = stk" | 
| 45015 | 490 | by (auto dest!: And.IH) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 491 | with b2 j | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 492 | show ?case | 
| 62390 | 493 | by (fastforce dest!: And.IH simp: exec_n_end split: if_split_asm) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 494 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 495 | case Less | 
| 52915 | 496 | thus ?case by (auto dest!: exec_n_split_full simp: exec_n_simps exec1_def) (* takes time *) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 497 | qed | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 498 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 499 | lemma ccomp_empty [elim!]: | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 500 | "ccomp c = [] \<Longrightarrow> (c,s) \<Rightarrow> s" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 501 | by (induct c) auto | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 502 | |
| 44070 | 503 | declare assign_simp [simp] | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 504 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 505 | lemma ccomp_exec_n: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 506 | "ccomp c \<turnstile> (0,s,stk) \<rightarrow>^n (size(ccomp c),t,stk') | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 507 | \<Longrightarrow> (c,s) \<Rightarrow> t \<and> stk'=stk" | 
| 45015 | 508 | proof (induction c arbitrary: s t stk stk' n) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 509 | case SKIP | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 510 | thus ?case by auto | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 511 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 512 | case (Assign x a) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 513 | thus ?case | 
| 52915 | 514 | by simp (fastforce dest!: exec_n_split_full simp: exec_n_simps exec1_def) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 515 | next | 
| 47818 | 516 | case (Seq c1 c2) | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44070diff
changeset | 517 | thus ?case by (fastforce dest!: exec_n_split_full) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 518 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 519 | case (If b c1 c2) | 
| 45015 | 520 | note If.IH [dest!] | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 521 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 522 | let ?if = "IF b THEN c1 ELSE c2" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 523 | let ?cs = "ccomp ?if" | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 524 | let ?bcomp = "bcomp b False (size (ccomp c1) + 1)" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 525 | |
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 526 | from `?cs \<turnstile> (0,s,stk) \<rightarrow>^n (size ?cs,t,stk')` | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 527 | obtain i' :: int and k m s'' stk'' where | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 528 | cs: "?cs \<turnstile> (i',s'',stk'') \<rightarrow>^m (size ?cs,t,stk')" and | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 529 | "?bcomp \<turnstile> (0,s,stk) \<rightarrow>^k (i', s'', stk'')" | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 530 | "i' = size ?bcomp \<or> i' = size ?bcomp + size (ccomp c1) + 1" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 531 | by (auto dest!: bcomp_split) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 532 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 533 | hence i': | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 534 | "s''=s" "stk'' = stk" | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 535 | "i' = (if bval b s then size ?bcomp else size ?bcomp+size(ccomp c1)+1)" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 536 | by auto | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 537 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 538 | with cs have cs': | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 539 | "ccomp c1@JMP (size (ccomp c2))#ccomp c2 \<turnstile> | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 540 | (if bval b s then 0 else size (ccomp c1)+1, s, stk) \<rightarrow>^m | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 541 | (1 + size (ccomp c1) + size (ccomp c2), t, stk')" | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44070diff
changeset | 542 | by (fastforce dest: exec_n_drop_left simp: exits_Cons isuccs_def algebra_simps) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 543 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 544 | show ?case | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 545 | proof (cases "bval b s") | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 546 | case True with cs' | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 547 | show ?thesis | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 548 | by simp | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44070diff
changeset | 549 | (fastforce dest: exec_n_drop_right | 
| 62390 | 550 | split: if_split_asm | 
| 52915 | 551 | simp: exec_n_simps exec1_def) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 552 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 553 | case False with cs' | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 554 | show ?thesis | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 555 | by (auto dest!: exec_n_drop_Cons exec_n_drop_left | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 556 | simp: exits_Cons isuccs_def) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 557 | qed | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 558 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 559 | case (While b c) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 560 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 561 | from While.prems | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 562 | show ?case | 
| 45015 | 563 | proof (induction n arbitrary: s rule: nat_less_induct) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 564 | case (1 n) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 565 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 566 |     { assume "\<not> bval b s"
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 567 | with "1.prems" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 568 | have ?case | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 569 | by simp | 
| 53880 | 570 | (fastforce dest!: bcomp_exec_n bcomp_split simp: exec_n_simps) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 571 |     } moreover {
 | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 572 | assume b: "bval b s" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 573 | let ?c0 = "WHILE b DO c" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 574 | let ?cs = "ccomp ?c0" | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 575 | let ?bs = "bcomp b False (size (ccomp c) + 1)" | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 576 | let ?jmp = "[JMP (-((size ?bs + size (ccomp c) + 1)))]" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 577 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 578 | from "1.prems" b | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 579 | obtain k where | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 580 | cs: "?cs \<turnstile> (size ?bs, s, stk) \<rightarrow>^k (size ?cs, t, stk')" and | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 581 | k: "k \<le> n" | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44070diff
changeset | 582 | by (fastforce dest!: bcomp_split) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 583 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 584 | have ?case | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 585 | proof cases | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 586 | assume "ccomp c = []" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 587 | with cs k | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 588 | obtain m where | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 589 | "?cs \<turnstile> (0,s,stk) \<rightarrow>^m (size (ccomp ?c0), t, stk')" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 590 | "m < n" | 
| 52915 | 591 | by (auto simp: exec_n_step [where k=k] exec1_def) | 
| 45015 | 592 | with "1.IH" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 593 | show ?case by blast | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 594 | next | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 595 | assume "ccomp c \<noteq> []" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 596 | with cs | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 597 | obtain m m' s'' stk'' where | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 598 | c: "ccomp c \<turnstile> (0, s, stk) \<rightarrow>^m' (size (ccomp c), s'', stk'')" and | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 599 | rest: "?cs \<turnstile> (size ?bs + size (ccomp c), s'', stk'') \<rightarrow>^m | 
| 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 600 | (size ?cs, t, stk')" and | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 601 | m: "k = m + m'" | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 602 | by (auto dest: exec_n_split [where i=0, simplified]) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 603 | from c | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 604 | have "(c,s) \<Rightarrow> s''" and stk: "stk'' = stk" | 
| 45015 | 605 | by (auto dest!: While.IH) | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 606 | moreover | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 607 | from rest m k stk | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 608 | obtain k' where | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 609 | "?cs \<turnstile> (0, s'', stk) \<rightarrow>^k' (size ?cs, t, stk')" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 610 | "k' < n" | 
| 52915 | 611 | by (auto simp: exec_n_step [where k=m] exec1_def) | 
| 45015 | 612 | with "1.IH" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 613 | have "(?c0, s'') \<Rightarrow> t \<and> stk' = stk" by blast | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 614 | ultimately | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 615 | show ?case using b by blast | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 616 | qed | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 617 | } | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 618 | ultimately show ?case by cases | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 619 | qed | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 620 | qed | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 621 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 622 | theorem ccomp_exec: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 623 | "ccomp c \<turnstile> (0,s,stk) \<rightarrow>* (size(ccomp c),t,stk') \<Longrightarrow> (c,s) \<Rightarrow> t" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 624 | by (auto dest: exec_exec_n ccomp_exec_n) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 625 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 626 | corollary ccomp_sound: | 
| 51259 
1491459df114
eliminated isize in favour of size + type coercion
 kleing parents: 
50061diff
changeset | 627 | "ccomp c \<turnstile> (0,s,stk) \<rightarrow>* (size(ccomp c),t,stk) \<longleftrightarrow> (c,s) \<Rightarrow> t" | 
| 43438 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 628 | by (blast intro!: ccomp_exec ccomp_bigstep) | 
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 629 | |
| 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 kleing parents: diff
changeset | 630 | end |