| author | huffman | 
| Sun, 22 Feb 2009 08:52:44 -0800 | |
| changeset 30066 | 9223483cc927 | 
| parent 27363 | 6d93bbe5633e | 
| child 30952 | 7ab2716dd93b | 
| permissions | -rw-r--r-- | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 1 | (* Title: HOL/IMP/Compiler.thy | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 2 | ID: $Id$ | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 3 | Author: Tobias Nipkow, TUM | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 4 | Copyright 1996 TUM | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 5 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 6 | This is an early version of the compiler, where the abstract machine | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 7 | has an explicit pc. This turned out to be awkward, and a second | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 8 | development was started. See Machines.thy and Compiler.thy. | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 9 | *) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 10 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 11 | header "A Simple Compiler" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 12 | |
| 16417 | 13 | theory Compiler0 imports Natural begin | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 14 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 15 | subsection "An abstract, simplistic machine" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 16 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 17 | text {* There are only three instructions: *}
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 18 | datatype instr = ASIN loc aexp | JMPF bexp nat | JMPB nat | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 19 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 20 | text {* We describe execution of programs in the machine by
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 21 | an operational (small step) semantics: | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 22 | *} | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 23 | |
| 23746 | 24 | inductive_set | 
| 25 | stepa1 :: "instr list \<Rightarrow> ((state\<times>nat) \<times> (state\<times>nat))set" | |
| 26 | and stepa1' :: "[instr list,state,nat,state,nat] \<Rightarrow> bool" | |
| 27 |     ("_ \<turnstile> (3\<langle>_,_\<rangle>/ -1\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)
 | |
| 28 | for P :: "instr list" | |
| 29 | where | |
| 30 | "P \<turnstile> \<langle>s,m\<rangle> -1\<rightarrow> \<langle>t,n\<rangle> == ((s,m),t,n) : stepa1 P" | |
| 31 | | ASIN[simp]: | |
| 32 | "\<lbrakk> n<size P; P!n = ASIN x a \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s[x\<mapsto> a s],Suc n\<rangle>" | |
| 33 | | JMPFT[simp,intro]: | |
| 34 | "\<lbrakk> n<size P; P!n = JMPF b i; b s \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s,Suc n\<rangle>" | |
| 35 | | JMPFF[simp,intro]: | |
| 36 | "\<lbrakk> n<size P; P!n = JMPF b i; ~b s; m=n+i \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s,m\<rangle>" | |
| 37 | | JMPB[simp]: | |
| 38 | "\<lbrakk> n<size P; P!n = JMPB i; i <= n; j = n-i \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s,j\<rangle>" | |
| 14565 | 39 | |
| 23746 | 40 | abbreviation | 
| 41 | stepa :: "[instr list,state,nat,state,nat] \<Rightarrow> bool" | |
| 42 |     ("_ \<turnstile>/ (3\<langle>_,_\<rangle>/ -*\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)  where
 | |
| 43 | "P \<turnstile> \<langle>s,m\<rangle> -*\<rightarrow> \<langle>t,n\<rangle> == ((s,m),t,n) : ((stepa1 P)^*)" | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 44 | |
| 23746 | 45 | abbreviation | 
| 46 | stepan :: "[instr list,state,nat,nat,state,nat] \<Rightarrow> bool" | |
| 47 |     ("_ \<turnstile>/ (3\<langle>_,_\<rangle>/ -(_)\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0,0] 50)  where
 | |
| 48 | "P \<turnstile> \<langle>s,m\<rangle> -(i)\<rightarrow> \<langle>t,n\<rangle> == ((s,m),t,n) : ((stepa1 P)^i)" | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 49 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 50 | subsection "The compiler" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 51 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 52 | consts compile :: "com \<Rightarrow> instr list" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 53 | primrec | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 54 | "compile \<SKIP> = []" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 55 | "compile (x:==a) = [ASIN x a]" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 56 | "compile (c1;c2) = compile c1 @ compile c2" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 57 | "compile (\<IF> b \<THEN> c1 \<ELSE> c2) = | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 58 | [JMPF b (length(compile c1) + 2)] @ compile c1 @ | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 59 | [JMPF (%x. False) (length(compile c2)+1)] @ compile c2" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 60 | "compile (\<WHILE> b \<DO> c) = [JMPF b (length(compile c) + 2)] @ compile c @ | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 61 | [JMPB (length(compile c)+1)]" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 62 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 63 | declare nth_append[simp] | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 64 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 65 | subsection "Context lifting lemmas" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 66 | |
| 18372 | 67 | text {*
 | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 68 | Some lemmas for lifting an execution into a prefix and suffix | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 69 | of instructions; only needed for the first proof. | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 70 | *} | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 71 | lemma app_right_1: | 
| 18372 | 72 | assumes "is1 \<turnstile> \<langle>s1,i1\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>" | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 73 | shows "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>" | 
| 27363 | 74 | using assms | 
| 18372 | 75 | by induct auto | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 76 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 77 | lemma app_left_1: | 
| 18372 | 78 | assumes "is2 \<turnstile> \<langle>s1,i1\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>" | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 79 | shows "is1 @ is2 \<turnstile> \<langle>s1,size is1+i1\<rangle> -1\<rightarrow> \<langle>s2,size is1+i2\<rangle>" | 
| 27363 | 80 | using assms | 
| 18372 | 81 | by induct auto | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 82 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 83 | declare rtrancl_induct2 [induct set: rtrancl] | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 84 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 85 | lemma app_right: | 
| 18372 | 86 | assumes "is1 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>" | 
| 87 | shows "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>" | |
| 27363 | 88 | using assms | 
| 18372 | 89 | proof induct | 
| 90 | show "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s1,i1\<rangle>" by simp | |
| 91 | next | |
| 92 | fix s1' i1' s2 i2 | |
| 93 | assume "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s1',i1'\<rangle>" | |
| 94 | and "is1 \<turnstile> \<langle>s1',i1'\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>" | |
| 95 | thus "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>" | |
| 96 | by (blast intro: app_right_1 rtrancl_trans) | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 97 | qed | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 98 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 99 | lemma app_left: | 
| 18372 | 100 | assumes "is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>" | 
| 101 | shows "is1 @ is2 \<turnstile> \<langle>s1,size is1+i1\<rangle> -*\<rightarrow> \<langle>s2,size is1+i2\<rangle>" | |
| 27363 | 102 | using assms | 
| 18372 | 103 | proof induct | 
| 104 | show "is1 @ is2 \<turnstile> \<langle>s1,length is1 + i1\<rangle> -*\<rightarrow> \<langle>s1,length is1 + i1\<rangle>" by simp | |
| 105 | next | |
| 106 | fix s1' i1' s2 i2 | |
| 107 | assume "is1 @ is2 \<turnstile> \<langle>s1,length is1 + i1\<rangle> -*\<rightarrow> \<langle>s1',length is1 + i1'\<rangle>" | |
| 108 | and "is2 \<turnstile> \<langle>s1',i1'\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>" | |
| 109 | thus "is1 @ is2 \<turnstile> \<langle>s1,length is1 + i1\<rangle> -*\<rightarrow> \<langle>s2,length is1 + i2\<rangle>" | |
| 110 | by (blast intro: app_left_1 rtrancl_trans) | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 111 | qed | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 112 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 113 | lemma app_left2: | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 114 | "\<lbrakk> is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>; j1 = size is1+i1; j2 = size is1+i2 \<rbrakk> \<Longrightarrow> | 
| 18372 | 115 | is1 @ is2 \<turnstile> \<langle>s1,j1\<rangle> -*\<rightarrow> \<langle>s2,j2\<rangle>" | 
| 116 | by (simp add: app_left) | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 117 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 118 | lemma app1_left: | 
| 18372 | 119 | assumes "is \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>" | 
| 120 | shows "instr # is \<turnstile> \<langle>s1,Suc i1\<rangle> -*\<rightarrow> \<langle>s2,Suc i2\<rangle>" | |
| 121 | proof - | |
| 27363 | 122 | from app_left [OF assms, of "[instr]"] | 
| 18372 | 123 | show ?thesis by simp | 
| 124 | qed | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 125 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 126 | subsection "Compiler correctness" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 127 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 128 | declare rtrancl_into_rtrancl[trans] | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 129 | converse_rtrancl_into_rtrancl[trans] | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 130 | rtrancl_trans[trans] | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 131 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 132 | text {*
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 133 | The first proof; The statement is very intuitive, | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 134 | but application of induction hypothesis requires the above lifting lemmas | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 135 | *} | 
| 18372 | 136 | theorem | 
| 137 | assumes "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t" | |
| 138 | shows "compile c \<turnstile> \<langle>s,0\<rangle> -*\<rightarrow> \<langle>t,length(compile c)\<rangle>" (is "?P c s t") | |
| 27363 | 139 | using assms | 
| 18372 | 140 | proof induct | 
| 141 | show "\<And>s. ?P \<SKIP> s s" by simp | |
| 142 | next | |
| 143 | show "\<And>a s x. ?P (x :== a) s (s[x\<mapsto> a s])" by force | |
| 144 | next | |
| 145 | fix c0 c1 s0 s1 s2 | |
| 146 | assume "?P c0 s0 s1" | |
| 147 | hence "compile c0 @ compile c1 \<turnstile> \<langle>s0,0\<rangle> -*\<rightarrow> \<langle>s1,length(compile c0)\<rangle>" | |
| 148 | by (rule app_right) | |
| 149 | moreover assume "?P c1 s1 s2" | |
| 150 | hence "compile c0 @ compile c1 \<turnstile> \<langle>s1,length(compile c0)\<rangle> -*\<rightarrow> | |
| 151 | \<langle>s2,length(compile c0)+length(compile c1)\<rangle>" | |
| 152 | proof - | |
| 153 | show "\<And>is1 is2 s1 s2 i2. | |
| 154 | is2 \<turnstile> \<langle>s1,0\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle> \<Longrightarrow> | |
| 155 | is1 @ is2 \<turnstile> \<langle>s1,size is1\<rangle> -*\<rightarrow> \<langle>s2,size is1+i2\<rangle>" | |
| 156 | using app_left[of _ 0] by simp | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 157 | qed | 
| 18372 | 158 | ultimately have "compile c0 @ compile c1 \<turnstile> \<langle>s0,0\<rangle> -*\<rightarrow> | 
| 159 | \<langle>s2,length(compile c0)+length(compile c1)\<rangle>" | |
| 160 | by (rule rtrancl_trans) | |
| 161 | thus "?P (c0; c1) s0 s2" by simp | |
| 162 | next | |
| 163 | fix b c0 c1 s0 s1 | |
| 164 | let ?comp = "compile(\<IF> b \<THEN> c0 \<ELSE> c1)" | |
| 165 | assume "b s0" and IH: "?P c0 s0 s1" | |
| 166 | hence "?comp \<turnstile> \<langle>s0,0\<rangle> -1\<rightarrow> \<langle>s0,1\<rangle>" by auto | |
| 167 | also from IH | |
| 168 | have "?comp \<turnstile> \<langle>s0,1\<rangle> -*\<rightarrow> \<langle>s1,length(compile c0)+1\<rangle>" | |
| 169 | by(auto intro:app1_left app_right) | |
| 170 | also have "?comp \<turnstile> \<langle>s1,length(compile c0)+1\<rangle> -1\<rightarrow> \<langle>s1,length ?comp\<rangle>" | |
| 171 | by(auto) | |
| 172 | finally show "?P (\<IF> b \<THEN> c0 \<ELSE> c1) s0 s1" . | |
| 173 | next | |
| 174 | fix b c0 c1 s0 s1 | |
| 175 | let ?comp = "compile(\<IF> b \<THEN> c0 \<ELSE> c1)" | |
| 176 | assume "\<not>b s0" and IH: "?P c1 s0 s1" | |
| 177 | hence "?comp \<turnstile> \<langle>s0,0\<rangle> -1\<rightarrow> \<langle>s0,length(compile c0) + 2\<rangle>" by auto | |
| 178 | also from IH | |
| 179 | have "?comp \<turnstile> \<langle>s0,length(compile c0)+2\<rangle> -*\<rightarrow> \<langle>s1,length ?comp\<rangle>" | |
| 180 | by (force intro!: app_left2 app1_left) | |
| 181 | finally show "?P (\<IF> b \<THEN> c0 \<ELSE> c1) s0 s1" . | |
| 182 | next | |
| 183 | fix b c and s::state | |
| 184 | assume "\<not>b s" | |
| 185 | thus "?P (\<WHILE> b \<DO> c) s s" by force | |
| 186 | next | |
| 187 | fix b c and s0::state and s1 s2 | |
| 188 | let ?comp = "compile(\<WHILE> b \<DO> c)" | |
| 189 | assume "b s0" and | |
| 190 | IHc: "?P c s0 s1" and IHw: "?P (\<WHILE> b \<DO> c) s1 s2" | |
| 191 | hence "?comp \<turnstile> \<langle>s0,0\<rangle> -1\<rightarrow> \<langle>s0,1\<rangle>" by auto | |
| 192 | also from IHc | |
| 193 | have "?comp \<turnstile> \<langle>s0,1\<rangle> -*\<rightarrow> \<langle>s1,length(compile c)+1\<rangle>" | |
| 194 | by (auto intro: app1_left app_right) | |
| 195 | also have "?comp \<turnstile> \<langle>s1,length(compile c)+1\<rangle> -1\<rightarrow> \<langle>s1,0\<rangle>" by simp | |
| 196 | also note IHw | |
| 197 | finally show "?P (\<WHILE> b \<DO> c) s0 s2". | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 198 | qed | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 199 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 200 | text {*
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 201 | Second proof; statement is generalized to cater for prefixes and suffixes; | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 202 | needs none of the lifting lemmas, but instantiations of pre/suffix. | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 203 | *} | 
| 13130 | 204 | (* | 
| 13112 | 205 | theorem assumes A: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t" | 
| 206 | shows "\<And>a z. a@compile c@z \<turnstile> \<langle>s,size a\<rangle> -*\<rightarrow> \<langle>t,size a + size(compile c)\<rangle>" | |
| 207 | (is "\<And>a z. ?P c s t a z") | |
| 208 | proof - | |
| 209 | from A show "\<And>a z. ?thesis a z" | |
| 210 | proof induct | |
| 211 | case Skip thus ?case by simp | |
| 212 | next | |
| 213 | case Assign thus ?case by (force intro!: ASIN) | |
| 214 | next | |
| 215 | fix c1 c2 s s' s'' a z | |
| 216 | assume IH1: "\<And>a z. ?P c1 s s' a z" and IH2: "\<And>a z. ?P c2 s' s'' a z" | |
| 217 | from IH1 IH2[of "a@compile c1"] | |
| 218 | show "?P (c1;c2) s s'' a z" | |
| 219 | by(simp add:add_assoc[THEN sym])(blast intro:rtrancl_trans) | |
| 220 | next | |
| 221 | (* at this point I gave up converting to structured proofs *) | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 222 | (* \<IF> b \<THEN> c0 \<ELSE> c1; case b is true *) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 223 | apply(intro strip) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 224 | (* instantiate assumption sufficiently for later: *) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 225 | apply(erule_tac x = "a@[?I]" in allE) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 226 | apply(simp) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 227 | (* execute JMPF: *) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 228 | apply(rule converse_rtrancl_into_rtrancl) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 229 | apply(force intro!: JMPFT) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 230 | (* execute compile c0: *) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 231 | apply(rule rtrancl_trans) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 232 | apply(erule allE) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 233 | apply assumption | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 234 | (* execute JMPF: *) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 235 | apply(rule r_into_rtrancl) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 236 | apply(force intro!: JMPFF) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 237 | (* end of case b is true *) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 238 | apply(intro strip) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 239 | apply(erule_tac x = "a@[?I]@compile c0@[?J]" in allE) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 240 | apply(simp add:add_assoc) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 241 | apply(rule converse_rtrancl_into_rtrancl) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 242 | apply(force intro!: JMPFF) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 243 | apply(blast) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 244 | apply(force intro: JMPFF) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 245 | apply(intro strip) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 246 | apply(erule_tac x = "a@[?I]" in allE) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 247 | apply(erule_tac x = a in allE) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 248 | apply(simp) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 249 | apply(rule converse_rtrancl_into_rtrancl) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 250 | apply(force intro!: JMPFT) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 251 | apply(rule rtrancl_trans) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 252 | apply(erule allE) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 253 | apply assumption | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 254 | apply(rule converse_rtrancl_into_rtrancl) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 255 | apply(force intro!: JMPB) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 256 | apply(simp) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 257 | done | 
| 13130 | 258 | *) | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 259 | text {* Missing: the other direction! I did much of it, and although
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 260 | the main lemma is very similar to the one in the new development, the | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 261 | lemmas surrounding it seemed much more complicated. In the end I gave | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 262 | up. *} | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 263 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 264 | end |