src/HOL/IMP/Compiler.thy
changeset 44296 a666b8d11252
parent 43999 686fa0a0696e
child 44871 ab4d8499815c
     1.1 --- a/src/HOL/IMP/Compiler.thy	Fri Jun 17 14:35:24 2011 +0200
     1.2 +++ b/src/HOL/IMP/Compiler.thy	Fri Jun 17 20:38:43 2011 +0200
     1.3 @@ -2,56 +2,92 @@
     1.4  
     1.5  header "A Compiler for IMP"
     1.6  
     1.7 -theory Compiler imports Big_Step
     1.8 +theory Compiler imports Big_Step 
     1.9  begin
    1.10  
    1.11 +subsection "List setup"
    1.12 +
    1.13 +text {*
    1.14 +  We are going to define a small machine language where programs are
    1.15 +  lists of instructions. For nicer algebraic properties in our lemmas
    1.16 +  later, we prefer @{typ int} to @{term nat} as program counter.
    1.17 +  
    1.18 +  Therefore, we define notation for size and indexing for lists 
    1.19 +  on @{typ int}:
    1.20 +*}
    1.21 +abbreviation "isize xs == int (length xs)" 
    1.22 +
    1.23 +primrec
    1.24 +  inth :: "'a list => int => 'a" (infixl "!!" 100) where
    1.25 +  inth_Cons: "(x # xs) !! n = (if n = 0 then x else xs !! (n - 1))"
    1.26 +
    1.27 +text {*
    1.28 +  The only additional lemma we need is indexing over append:
    1.29 +*}
    1.30 +lemma inth_append [simp]:
    1.31 +  "0 \<le> n \<Longrightarrow>
    1.32 +  (xs @ ys) !! n = (if n < isize xs then xs !! n else ys !! (n - isize xs))"
    1.33 +  by (induct xs arbitrary: n) (auto simp: algebra_simps)
    1.34 +
    1.35  subsection "Instructions and Stack Machine"
    1.36  
    1.37  datatype instr = 
    1.38 -  LOADI int | LOAD string | ADD |
    1.39 +  LOADI int | 
    1.40 +  LOAD string | 
    1.41 +  ADD |
    1.42    STORE string |
    1.43 -  JMPF nat |
    1.44 -  JMPB nat |
    1.45 -  JMPFLESS nat |
    1.46 -  JMPFGE nat
    1.47 +  JMP int |
    1.48 +  JMPFLESS int |
    1.49 +  JMPFGE int
    1.50  
    1.51 -type_synonym stack = "int list"
    1.52 -type_synonym config = "nat\<times>state\<times>stack"
    1.53 +(* reads slightly nicer *)
    1.54 +abbreviation
    1.55 +  "JMPB i == JMP (-i)"
    1.56 +
    1.57 +type_synonym stack = "val list"
    1.58 +type_synonym config = "int\<times>state\<times>stack"
    1.59  
    1.60  abbreviation "hd2 xs == hd(tl xs)"
    1.61  abbreviation "tl2 xs == tl(tl xs)"
    1.62  
    1.63 -inductive exec1 :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool"
    1.64 -    ("(_/ \<turnstile> (_ \<rightarrow>/ _))" [50,0,0] 50)
    1.65 -  for P :: "instr list"
    1.66 +inductive iexec1 :: "instr \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool"
    1.67 +    ("(_/ \<turnstile>i (_ \<rightarrow>/ _))" [50,0,0] 50)
    1.68  where
    1.69 -"\<lbrakk> i < size P;  P!i = LOADI n \<rbrakk> \<Longrightarrow>
    1.70 - P \<turnstile> (i,s,stk) \<rightarrow> (i+1,s, n#stk)" |
    1.71 -"\<lbrakk> i < size P;  P!i = LOAD x \<rbrakk> \<Longrightarrow> 
    1.72 - P \<turnstile> (i,s,stk) \<rightarrow> (i+1,s, s x # stk)" |
    1.73 -"\<lbrakk> i < size P;  P!i = ADD \<rbrakk> \<Longrightarrow> 
    1.74 - P \<turnstile> (i,s,stk) \<rightarrow> (i+1,s, (hd2 stk + hd stk) # tl2 stk)" |
    1.75 -"\<lbrakk> i < size P;  P!i = STORE n \<rbrakk> \<Longrightarrow>
    1.76 - P \<turnstile> (i,s,stk) \<rightarrow> (i+1,s(n := hd stk),tl stk)" |
    1.77 -"\<lbrakk> i < size P;  P!i = JMPF n \<rbrakk> \<Longrightarrow>
    1.78 - P \<turnstile> (i,s,stk) \<rightarrow> (i+1+n,s,stk)" |
    1.79 -"\<lbrakk> i < size P;  P!i = JMPB n;  n \<le> i+1 \<rbrakk> \<Longrightarrow>
    1.80 - P \<turnstile> (i,s,stk) \<rightarrow> (i+1-n,s,stk)" |
    1.81 -"\<lbrakk> i < size P;  P!i = JMPFLESS n \<rbrakk> \<Longrightarrow>
    1.82 - P \<turnstile> (i,s,stk) \<rightarrow> (if hd2 stk < hd stk then i+1+n else i+1,s,tl2 stk)" |
    1.83 -"\<lbrakk> i < size P;  P!i = JMPFGE n \<rbrakk> \<Longrightarrow>
    1.84 - P \<turnstile> (i,s,stk) \<rightarrow> (if hd2 stk >= hd stk then i+1+n else i+1,s,tl2 stk)"
    1.85 +"LOADI n \<turnstile>i (i,s,stk) \<rightarrow> (i+1,s, n#stk)" |
    1.86 +"LOAD x  \<turnstile>i (i,s,stk) \<rightarrow> (i+1,s, s x # stk)" |
    1.87 +"ADD     \<turnstile>i (i,s,stk) \<rightarrow> (i+1,s, (hd2 stk + hd stk) # tl2 stk)" |
    1.88 +"STORE n \<turnstile>i (i,s,stk) \<rightarrow> (i+1,s(n := hd stk),tl stk)" |
    1.89 +"JMP n   \<turnstile>i (i,s,stk) \<rightarrow> (i+1+n,s,stk)" |
    1.90 +"JMPFLESS n \<turnstile>i (i,s,stk) \<rightarrow> (if hd2 stk < hd stk then i+1+n else i+1,s,tl2 stk)" |
    1.91 +"JMPFGE n \<turnstile>i (i,s,stk) \<rightarrow> (if hd2 stk >= hd stk then i+1+n else i+1,s,tl2 stk)"
    1.92 +
    1.93 +code_pred iexec1 .
    1.94 +
    1.95 +declare iexec1.intros
    1.96 +
    1.97 +(* FIXME: why does code gen not work with fun? *)
    1.98 +inductive
    1.99 +  exec1 :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool"
   1.100 +    ("(_/ \<turnstile> (_ \<rightarrow>/ _))" [50,0,0] 50) where
   1.101 + "\<lbrakk> P!!i \<turnstile>i (i,s,stk) \<rightarrow> c'; 0 \<le> i; i < isize P \<rbrakk> \<Longrightarrow> P \<turnstile> (i,s,stk) \<rightarrow> c'"
   1.102  
   1.103  code_pred exec1 .
   1.104  
   1.105 -declare exec1.intros[intro]
   1.106 +declare exec1.intros [intro!]
   1.107 +
   1.108 +inductive_cases exec1_elim [elim!]: "P \<turnstile> c \<rightarrow> c'"
   1.109 +
   1.110 +lemma exec1_simp [simp]:
   1.111 +  "P \<turnstile> c \<rightarrow> c' = 
   1.112 +   (\<exists>i s stk. c = (i,s,stk) \<and> P!!i \<turnstile>i (i,s,stk) \<rightarrow> c' \<and> 0 \<le> i \<and> i < isize P)"
   1.113 +  by auto
   1.114  
   1.115  inductive exec :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool" ("_/ \<turnstile> (_ \<rightarrow>*/ _)" 50)
   1.116  where
   1.117  refl: "P \<turnstile> c \<rightarrow>* c" |
   1.118  step: "P \<turnstile> c \<rightarrow> c' \<Longrightarrow> P \<turnstile> c' \<rightarrow>* c'' \<Longrightarrow> P \<turnstile> c \<rightarrow>* c''"
   1.119  
   1.120 -declare exec.intros[intro]
   1.121 +declare refl[intro] step[intro]
   1.122  
   1.123  lemmas exec_induct = exec.induct[split_format(complete)]
   1.124  
   1.125 @@ -66,45 +102,68 @@
   1.126  subsection{* Verification infrastructure *}
   1.127  
   1.128  lemma exec_trans: "P \<turnstile> c \<rightarrow>* c' \<Longrightarrow> P \<turnstile> c' \<rightarrow>* c'' \<Longrightarrow> P \<turnstile> c \<rightarrow>* c''"
   1.129 -apply(induct rule: exec.induct)
   1.130 - apply blast
   1.131 -by (metis exec.step)
   1.132 +  by (induct rule: exec.induct) fastsimp+
   1.133  
   1.134 -lemma exec1_subst: "P \<turnstile> c \<rightarrow> c' \<Longrightarrow> c' = c'' \<Longrightarrow> P \<turnstile> c \<rightarrow> c''"
   1.135 -by auto
   1.136 +inductive_cases iexec1_cases [elim!]:
   1.137 +  "LOADI n \<turnstile>i c \<rightarrow> c'" 
   1.138 +  "LOAD x  \<turnstile>i c \<rightarrow> c'"
   1.139 +  "ADD     \<turnstile>i c \<rightarrow> c'"
   1.140 +  "STORE n \<turnstile>i c \<rightarrow> c'" 
   1.141 +  "JMP n   \<turnstile>i c \<rightarrow> c'"
   1.142 +  "JMPFLESS n \<turnstile>i c \<rightarrow> c'"
   1.143 +  "JMPFGE n \<turnstile>i c \<rightarrow> c'"
   1.144  
   1.145 -lemmas exec1_simps = exec1.intros[THEN exec1_subst]
   1.146 +text {* Simplification rules for @{const iexec1}. *}
   1.147 +lemma iexec1_simps [simp]:
   1.148 +  "LOADI n \<turnstile>i c \<rightarrow> c' = (\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1, s, n # stk))"
   1.149 +  "LOAD x \<turnstile>i c \<rightarrow> c' = (\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1, s, s x # stk))"
   1.150 +  "ADD \<turnstile>i c \<rightarrow> c' = 
   1.151 +  (\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1, s, (hd2 stk + hd stk) # tl2 stk))"
   1.152 +  "STORE x \<turnstile>i c \<rightarrow> c' = 
   1.153 +  (\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1, s(x \<rightarrow> hd stk), tl stk))"
   1.154 +  "JMP n \<turnstile>i c \<rightarrow> c' = (\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1 + n, s, stk))"
   1.155 +   "JMPFLESS n \<turnstile>i c \<rightarrow> c' = 
   1.156 +  (\<exists>i s stk. c = (i, s, stk) \<and> 
   1.157 +             c' = (if hd2 stk < hd stk then i + 1 + n else i + 1, s, tl2 stk))"  
   1.158 +  "JMPFGE n \<turnstile>i c \<rightarrow> c' = 
   1.159 +  (\<exists>i s stk. c = (i, s, stk) \<and> 
   1.160 +             c' = (if hd stk \<le> hd2 stk then i + 1 + n else i + 1, s, tl2 stk))"
   1.161 +  by (auto split del: split_if intro!: iexec1.intros)
   1.162 +
   1.163  
   1.164  text{* Below we need to argue about the execution of code that is embedded in
   1.165  larger programs. For this purpose we show that execution is preserved by
   1.166  appending code to the left or right of a program. *}
   1.167  
   1.168 -lemma exec1_appendR: assumes "P \<turnstile> c \<rightarrow> c'" shows "P@P' \<turnstile> c \<rightarrow> c'"
   1.169 -proof-
   1.170 -  from assms show ?thesis
   1.171 -  by cases (simp_all add: exec1_simps nth_append)
   1.172 -  -- "All cases proved with the final simp-all"
   1.173 -qed
   1.174 +lemma exec1_appendR: "P \<turnstile> c \<rightarrow> c' \<Longrightarrow> P@P' \<turnstile> c \<rightarrow> c'"
   1.175 +  by auto
   1.176  
   1.177  lemma exec_appendR: "P \<turnstile> c \<rightarrow>* c' \<Longrightarrow> P@P' \<turnstile> c \<rightarrow>* c'"
   1.178 -apply(induct rule: exec.induct)
   1.179 - apply blast
   1.180 -by (metis exec1_appendR exec.step)
   1.181 +  by (induct rule: exec.induct) (fastsimp intro: exec1_appendR)+
   1.182  
   1.183 +lemma iexec1_shiftI:
   1.184 +  assumes "X \<turnstile>i (i,s,stk) \<rightarrow> (i',s',stk')"
   1.185 +  shows "X \<turnstile>i (n+i,s,stk) \<rightarrow> (n+i',s',stk')"
   1.186 +  using assms by cases auto
   1.187 +
   1.188 +lemma iexec1_shiftD:
   1.189 +  assumes "X \<turnstile>i (n+i,s,stk) \<rightarrow> (n+i',s',stk')"
   1.190 +  shows "X \<turnstile>i (i,s,stk) \<rightarrow> (i',s',stk')"
   1.191 +  using assms by cases auto
   1.192 +
   1.193 +lemma iexec_shift [simp]: 
   1.194 +  "(X \<turnstile>i (n+i,s,stk) \<rightarrow> (n+i',s',stk')) = (X \<turnstile>i (i,s,stk) \<rightarrow> (i',s',stk'))"
   1.195 +  by (blast intro: iexec1_shiftI dest: iexec1_shiftD)
   1.196 +  
   1.197  lemma exec1_appendL:
   1.198 -assumes "P \<turnstile> (i,s,stk) \<rightarrow> (i',s',stk')"
   1.199 -shows "P' @ P \<turnstile> (size(P')+i,s,stk) \<rightarrow> (size(P')+i',s',stk')"
   1.200 -proof-
   1.201 -  from assms show ?thesis
   1.202 -  by cases (simp_all add: exec1_simps)
   1.203 -qed
   1.204 +  "P \<turnstile> (i,s,stk) \<rightarrow> (i',s',stk') \<Longrightarrow>
   1.205 +   P' @ P \<turnstile> (isize(P')+i,s,stk) \<rightarrow> (isize(P')+i',s',stk')"
   1.206 +  by simp
   1.207  
   1.208  lemma exec_appendL:
   1.209   "P \<turnstile> (i,s,stk) \<rightarrow>* (i',s',stk')  \<Longrightarrow>
   1.210 -  P' @ P \<turnstile> (size(P')+i,s,stk) \<rightarrow>* (size(P')+i',s',stk')"
   1.211 -apply(induct rule: exec_induct)
   1.212 - apply blast
   1.213 -by (blast intro: exec1_appendL exec.step)
   1.214 +  P' @ P \<turnstile> (isize(P')+i,s,stk) \<rightarrow>* (isize(P')+i',s',stk')"
   1.215 +  by (induct rule: exec_induct) (blast intro!: exec1_appendL)+
   1.216  
   1.217  text{* Now we specialise the above lemmas to enable automatic proofs of
   1.218  @{prop "P \<turnstile> c \<rightarrow>* c'"} where @{text P} is a mixture of concrete instructions and
   1.219 @@ -112,37 +171,33 @@
   1.220  by @{text "@"} and @{text "#"}. Backward jumps are not supported.
   1.221  The details should be skipped on a first reading.
   1.222  
   1.223 -If the pc points beyond the first instruction or part of the program, drop it: *}
   1.224 +If we have just executed the first instruction of the program, drop it: *}
   1.225  
   1.226 -lemma exec_Cons_Suc[intro]:
   1.227 -  "P \<turnstile> (i,s,stk) \<rightarrow>* (j,t,stk') \<Longrightarrow>
   1.228 -  instr#P \<turnstile> (Suc i,s,stk) \<rightarrow>* (Suc j,t,stk')"
   1.229 -apply(drule exec_appendL[where P'="[instr]"])
   1.230 -apply simp
   1.231 -done
   1.232 +lemma exec_Cons_1 [intro]:
   1.233 +  "P \<turnstile> (0,s,stk) \<rightarrow>* (j,t,stk') \<Longrightarrow>
   1.234 +  instr#P \<turnstile> (1,s,stk) \<rightarrow>* (1+j,t,stk')"
   1.235 +  by (drule exec_appendL[where P'="[instr]"]) simp
   1.236  
   1.237  lemma exec_appendL_if[intro]:
   1.238 - "size P' <= i
   1.239 -  \<Longrightarrow> P \<turnstile> (i - size P',s,stk) \<rightarrow>* (i',s',stk')
   1.240 -  \<Longrightarrow> P' @ P \<turnstile> (i,s,stk) \<rightarrow>* (size P' + i',s',stk')"
   1.241 -apply(drule exec_appendL[where P'=P'])
   1.242 -apply simp
   1.243 -done
   1.244 + "isize P' <= i
   1.245 +  \<Longrightarrow> P \<turnstile> (i - isize P',s,stk) \<rightarrow>* (i',s',stk')
   1.246 +  \<Longrightarrow> P' @ P \<turnstile> (i,s,stk) \<rightarrow>* (isize P' + i',s',stk')"
   1.247 +  by (drule exec_appendL[where P'=P']) simp
   1.248  
   1.249  text{* Split the execution of a compound program up into the excution of its
   1.250  parts: *}
   1.251  
   1.252  lemma exec_append_trans[intro]:
   1.253  "P \<turnstile> (0,s,stk) \<rightarrow>* (i',s',stk') \<Longrightarrow>
   1.254 - size P \<le> i' \<Longrightarrow>
   1.255 - P' \<turnstile>  (i' - size P,s',stk') \<rightarrow>* (i'',s'',stk'') \<Longrightarrow>
   1.256 - j'' = size P + i''
   1.257 + isize P \<le> i' \<Longrightarrow>
   1.258 + P' \<turnstile>  (i' - isize P,s',stk') \<rightarrow>* (i'',s'',stk'') \<Longrightarrow>
   1.259 + j'' = isize P + i''
   1.260   \<Longrightarrow>
   1.261   P @ P' \<turnstile> (0,s,stk) \<rightarrow>* (j'',s'',stk'')"
   1.262 -by(metis exec_trans[OF  exec_appendR exec_appendL_if])
   1.263 +  by(metis exec_trans[OF exec_appendR exec_appendL_if])
   1.264  
   1.265  
   1.266 -declare Let_def[simp] eval_nat_numeral[simp]
   1.267 +declare Let_def[simp] 
   1.268  
   1.269  
   1.270  subsection "Compilation"
   1.271 @@ -153,17 +208,15 @@
   1.272  "acomp (Plus a1 a2) = acomp a1 @ acomp a2 @ [ADD]"
   1.273  
   1.274  lemma acomp_correct[intro]:
   1.275 -  "acomp a \<turnstile> (0,s,stk) \<rightarrow>* (size(acomp a),s,aval a s#stk)"
   1.276 -apply(induct a arbitrary: stk)
   1.277 -apply(fastsimp)+
   1.278 -done
   1.279 +  "acomp a \<turnstile> (0,s,stk) \<rightarrow>* (isize(acomp a),s,aval a s#stk)"
   1.280 +  by (induct a arbitrary: stk) fastsimp+
   1.281  
   1.282 -fun bcomp :: "bexp \<Rightarrow> bool \<Rightarrow> nat \<Rightarrow> instr list" where
   1.283 -"bcomp (B bv) c n = (if bv=c then [JMPF n] else [])" |
   1.284 +fun bcomp :: "bexp \<Rightarrow> bool \<Rightarrow> int \<Rightarrow> instr list" where
   1.285 +"bcomp (B bv) c n = (if bv=c then [JMP n] else [])" |
   1.286  "bcomp (Not b) c n = bcomp b (\<not>c) n" |
   1.287  "bcomp (And b1 b2) c n =
   1.288   (let cb2 = bcomp b2 c n;
   1.289 -      m = (if c then size cb2 else size cb2+n);
   1.290 +        m = (if c then isize cb2 else isize cb2+n);
   1.291        cb1 = bcomp b1 False m
   1.292    in cb1 @ cb2)" |
   1.293  "bcomp (Less a1 a2) c n =
   1.294 @@ -174,27 +227,30 @@
   1.295       False 3"
   1.296  
   1.297  lemma bcomp_correct[intro]:
   1.298 - "bcomp b c n \<turnstile>
   1.299 - (0,s,stk)  \<rightarrow>*  (size(bcomp b c n) + (if c = bval b s then n else 0),s,stk)"
   1.300 +  "0 \<le> n \<Longrightarrow>
   1.301 +  bcomp b c n \<turnstile>
   1.302 + (0,s,stk)  \<rightarrow>*  (isize(bcomp b c n) + (if c = bval b s then n else 0),s,stk)"
   1.303  proof(induct b arbitrary: c n m)
   1.304    case Not
   1.305 -  from Not[of "~c"] show ?case by fastsimp
   1.306 +  from Not(1)[where c="~c"] Not(2) show ?case by fastsimp
   1.307  next
   1.308    case (And b1 b2)
   1.309 -  from And(1)[of "False"] And(2)[of "c"] show ?case by fastsimp
   1.310 +  from And(1)[of "if c then isize (bcomp b2 c n) else isize (bcomp b2 c n) + n" 
   1.311 +                 "False"] 
   1.312 +       And(2)[of n  "c"] And(3) 
   1.313 +  show ?case by fastsimp
   1.314  qed fastsimp+
   1.315  
   1.316 -
   1.317  fun ccomp :: "com \<Rightarrow> instr list" where
   1.318  "ccomp SKIP = []" |
   1.319  "ccomp (x ::= a) = acomp a @ [STORE x]" |
   1.320  "ccomp (c\<^isub>1;c\<^isub>2) = ccomp c\<^isub>1 @ ccomp c\<^isub>2" |
   1.321  "ccomp (IF b THEN c\<^isub>1 ELSE c\<^isub>2) =
   1.322 -  (let cc\<^isub>1 = ccomp c\<^isub>1; cc\<^isub>2 = ccomp c\<^isub>2; cb = bcomp b False (size cc\<^isub>1 + 1)
   1.323 -   in cb @ cc\<^isub>1 @ JMPF(size cc\<^isub>2) # cc\<^isub>2)" |
   1.324 +  (let cc\<^isub>1 = ccomp c\<^isub>1; cc\<^isub>2 = ccomp c\<^isub>2; cb = bcomp b False (isize cc\<^isub>1 + 1)
   1.325 +   in cb @ cc\<^isub>1 @ JMP (isize cc\<^isub>2) # cc\<^isub>2)" |
   1.326  "ccomp (WHILE b DO c) =
   1.327 - (let cc = ccomp c; cb = bcomp b False (size cc + 1)
   1.328 -  in cb @ cc @ [JMPB (size cb + size cc + 1)])"
   1.329 + (let cc = ccomp c; cb = bcomp b False (isize cc + 1)
   1.330 +  in cb @ cc @ [JMPB (isize cb + isize cc + 1)])"
   1.331  
   1.332  value "ccomp
   1.333   (IF Less (V ''u'') (N 1) THEN ''u'' ::= Plus (V ''u'') (N 1)
   1.334 @@ -205,32 +261,32 @@
   1.335  
   1.336  subsection "Preservation of sematics"
   1.337  
   1.338 -lemma ccomp_correct:
   1.339 -  "(c,s) \<Rightarrow> t \<Longrightarrow> ccomp c \<turnstile> (0,s,stk) \<rightarrow>* (size(ccomp c),t,stk)"
   1.340 +lemma ccomp_bigstep:
   1.341 +  "(c,s) \<Rightarrow> t \<Longrightarrow> ccomp c \<turnstile> (0,s,stk) \<rightarrow>* (isize(ccomp c),t,stk)"
   1.342  proof(induct arbitrary: stk rule: big_step_induct)
   1.343    case (Assign x a s)
   1.344 -  show ?case by (fastsimp simp:fun_upd_def)
   1.345 +  show ?case by (fastsimp simp:fun_upd_def cong: if_cong)
   1.346  next
   1.347    case (Semi c1 s1 s2 c2 s3)
   1.348    let ?cc1 = "ccomp c1"  let ?cc2 = "ccomp c2"
   1.349 -  have "?cc1 @ ?cc2 \<turnstile> (0,s1,stk) \<rightarrow>* (size ?cc1,s2,stk)"
   1.350 -    using Semi.hyps(2) by (fastsimp)
   1.351 +  have "?cc1 @ ?cc2 \<turnstile> (0,s1,stk) \<rightarrow>* (isize ?cc1,s2,stk)"
   1.352 +    using Semi.hyps(2) by fastsimp
   1.353    moreover
   1.354 -  have "?cc1 @ ?cc2 \<turnstile> (size ?cc1,s2,stk) \<rightarrow>* (size(?cc1 @ ?cc2),s3,stk)"
   1.355 -    using Semi.hyps(4) by (fastsimp)
   1.356 +  have "?cc1 @ ?cc2 \<turnstile> (isize ?cc1,s2,stk) \<rightarrow>* (isize(?cc1 @ ?cc2),s3,stk)"
   1.357 +    using Semi.hyps(4) by fastsimp
   1.358    ultimately show ?case by simp (blast intro: exec_trans)
   1.359  next
   1.360    case (WhileTrue b s1 c s2 s3)
   1.361    let ?cc = "ccomp c"
   1.362 -  let ?cb = "bcomp b False (size ?cc + 1)"
   1.363 +  let ?cb = "bcomp b False (isize ?cc + 1)"
   1.364    let ?cw = "ccomp(WHILE b DO c)"
   1.365 -  have "?cw \<turnstile> (0,s1,stk) \<rightarrow>* (size ?cb + size ?cc,s2,stk)"
   1.366 +  have "?cw \<turnstile> (0,s1,stk) \<rightarrow>* (isize ?cb + isize ?cc,s2,stk)"
   1.367      using WhileTrue(1,3) by fastsimp
   1.368    moreover
   1.369 -  have "?cw \<turnstile> (size ?cb + size ?cc,s2,stk) \<rightarrow>* (0,s2,stk)"
   1.370 -    by (fastsimp)
   1.371 +  have "?cw \<turnstile> (isize ?cb + isize ?cc,s2,stk) \<rightarrow>* (0,s2,stk)"
   1.372 +    by fastsimp
   1.373    moreover
   1.374 -  have "?cw \<turnstile> (0,s2,stk) \<rightarrow>* (size ?cw,s3,stk)" by(rule WhileTrue(5))
   1.375 +  have "?cw \<turnstile> (0,s2,stk) \<rightarrow>* (isize ?cw,s3,stk)" by(rule WhileTrue(5))
   1.376    ultimately show ?case by(blast intro: exec_trans)
   1.377  qed fastsimp+
   1.378