header "Compiler for IMP" theory solution06_2 imports Big_Step begin subsection "List setup" text {* We are going to define a small machine language where programs are lists of instructions. For nicer algebraic properties in our lemmas later, we prefer @{typ int} to @{term nat} as program counter. Therefore, we define notation for size and indexing for lists on @{typ int}: *} abbreviation "isize xs == int (length xs)" fun inth :: "'a list \ int \ 'a" (infixl "!!" 100) where "(x # xs) !! n = (if n = 0 then x else xs !! (n - 1))" text {* The only additional lemma we need is indexing over append: *} lemma inth_append [simp]: "0 \ n \ (xs @ ys) !! n = (if n < isize xs then xs !! n else ys !! (n - isize xs))" by (induction xs arbitrary: n) (auto simp: algebra_simps) subsection "Instructions and Stack Machine" datatype instr = LOADI int | LOAD vname | ADD | STORE vname | JMP int | LOADB bool | NOT | AND | LESS | JMP0 int type_synonym stack = "val list" type_synonym config = "int \ state \ stack" abbreviation "hd2 xs == hd(tl xs)" abbreviation "tl2 xs == tl(tl xs)" fun iexec :: "instr \ config \ config" where "iexec instr (i,s,stk) = (case instr of LOADI n \ (i+1,s, n#stk) | LOAD x \ (i+1,s, s x # stk) | ADD \ (i+1,s, (hd2 stk + hd stk) # tl2 stk) | STORE x \ (i+1,s(x := hd stk),tl stk) | JMP n \ (i+1+n,s,stk) | LOADB c \ (i+1,s, (if c then 1 else 0) # stk) | NOT \ (i+1,s, (if hd stk \ 0 then 0 else 1) # tl stk) | AND \ (i+1,s, (if hd2 stk \ 0 then if hd stk \ 0 then 1 else 0 else 0) # tl2 stk) | LESS \ (i+1,s, (if hd2 stk < hd stk then 1 else 0) # tl2 stk) | JMP0 n \ (if hd stk > 0 then i+1 else i+1+n,s,tl stk))" definition exec1 :: "instr list \ config \ config \ bool" ("(_/ \ (_ \/ _))" [59,0,59] 60) where "P \ c \ c' = (\i s stk. c = (i,s,stk) \ c' = iexec(P!!i) (i,s,stk) \ 0 \ i \ i < isize P)" declare exec1_def [simp] lemma exec1I [intro, code_pred_intro]: "c' = iexec (P!!i) (i,s,stk) \ 0 \ i \ i < isize P \ P \ (i,s,stk) \ c'" by simp inductive exec :: "instr list \ config \ config \ bool" ("(_/ \ (_ \*/ _))" 50) where refl: "P \ c \* c" | step: "P \ c \ c' \ P \ c' \* c'' \ P \ c \* c''" declare refl[intro] step[intro] lemmas exec_induct = exec.induct[split_format(complete)] code_pred exec by force values "{(i,map t [''x'',''y''],stk) | i t stk. [LOAD ''y'', STORE ''x''] \ (0, <''x'' := 3, ''y'' := 4>, []) \* (i,t,stk)}" subsection{* Verification infrastructure *} lemma exec_trans: "P \ c \* c' \ P \ c' \* c'' \ P \ c \* c''" by (induction rule: exec.induct) fastforce+ text{* Below we need to argue about the execution of code that is embedded in larger programs. For this purpose we show that execution is preserved by appending code to the left or right of a program. *} lemma iexec_shift [simp]: "((n+i',s',stk') = iexec x (n+i,s,stk)) = ((i',s',stk') = iexec x (i,s,stk))" by(auto split:instr.split) lemma exec1_appendR: "P \ c \ c' \ P@P' \ c \ c'" by fastforce+ lemma exec_appendR: "P \ c \* c' \ P@P' \ c \* c'" by (induction rule: exec.induct) (fastforce intro: exec1_appendR)+ lemma exec1_appendL: "P \ (i,s,stk) \ (i',s',stk') \ P' @ P \ (isize(P')+i,s,stk) \ (isize(P')+i',s',stk')" by (auto split: instr.split) lemma exec_appendL: "P \ (i,s,stk) \* (i',s',stk') \ P' @ P \ (isize(P')+i,s,stk) \* (isize(P')+i',s',stk')" by (induction rule: exec_induct) (blast intro!: exec1_appendL)+ text{* Now we specialise the above lemmas to enable automatic proofs of @{prop "P \ c \* c'"} where @{text P} is a mixture of concrete instructions and pieces of code that we already know how they execute (by induction), combined by @{text "@"} and @{text "#"}. Backward jumps are not supported. The details should be skipped on a first reading. If we have just executed the first instruction of the program, drop it: *} lemma exec_Cons_1 [intro]: "P \ (0,s,stk) \* (j,t,stk') \ instr#P \ (1,s,stk) \* (1+j,t,stk')" by (drule exec_appendL[where P'="[instr]"]) simp lemma exec_appendL_if[intro]: "isize P' <= i \ P \ (i - isize P',s,stk) \* (i',s',stk') \ P' @ P \ (i,s,stk) \* (isize P' + i',s',stk')" by (drule exec_appendL[where P'=P']) simp text{* Split the execution of a compound program up into the excution of its parts: *} lemma exec_append_trans[intro]: "P \ (0,s,stk) \* (i',s',stk') \ isize P \ i' \ P' \ (i' - isize P,s',stk') \* (i'',s'',stk'') \ j'' = isize P + i'' \ P @ P' \ (0,s,stk) \* (j'',s'',stk'')" by(metis exec_trans[OF exec_appendR exec_appendL_if]) declare Let_def[simp] subsection "Compilation" fun acomp :: "aexp \ instr list" where "acomp (N n) = [LOADI n]" | "acomp (V x) = [LOAD x]" | "acomp (Plus a1 a2) = acomp a1 @ acomp a2 @ [ADD]" lemma acomp_correct[intro]: "acomp a \ (0,s,stk) \* (isize(acomp a),s,aval a s#stk)" apply (induction a arbitrary: stk) by (fastforce intro!: exec_append_trans)+ fun bcomp :: "bexp \ instr list" where "bcomp (Bc v) = [LOADB v]" | "bcomp (Not b) = bcomp b @ [NOT]" | "bcomp (And b1 b2) = bcomp b1 @ bcomp b2 @ [AND]"| "bcomp (Less a1 a2) = acomp a1 @ acomp a2 @ [LESS]" lemma bcomp_correct[intro]: "bcomp b \ (0,s,stk) \* (isize(bcomp b),s,(if bval b s then 1 else 0) #stk)" apply (induction b arbitrary: stk) by (fastforce intro!: exec_append_trans)+ fun ccomp :: "com \ instr list" where "ccomp SKIP = []" | "ccomp (x ::= a) = acomp a @ [STORE x]" | "ccomp (c\<^isub>1;c\<^isub>2) = ccomp c\<^isub>1 @ ccomp c\<^isub>2" | "ccomp (IF b THEN c\<^isub>1 ELSE c\<^isub>2) = (let cc\<^isub>1 = ccomp c\<^isub>1; cc\<^isub>2 = ccomp c\<^isub>2; cb = bcomp b in cb @ JMP0 (isize cc\<^isub>1 + 1) # cc\<^isub>1 @ JMP (isize cc\<^isub>2) # cc\<^isub>2)" | "ccomp (WHILE b DO c) = (let cc = ccomp c; cb = bcomp b in cb @ JMP0 (isize cc + 1) # cc @ [JMP (-(isize cb + isize cc + 2))])" value "ccomp (IF Less (V ''u'') (N 1) THEN ''u'' ::= Plus (V ''u'') (N 1) ELSE ''v'' ::= V ''u'')" value "ccomp (WHILE Less (V ''u'') (N 1) DO (''u'' ::= Plus (V ''u'') (N 1)))" subsection "Preservation of semantics" lemma ccomp_bigstep: "(c,s) \ t \ ccomp c \ (0,s,stk) \* (isize(ccomp c),t,stk)" proof(induction arbitrary: stk rule: big_step_induct) case (Assign x a s) show ?case by (fastforce simp:fun_upd_def cong: if_cong intro!: exec_append_trans) next case (Seq c1 s1 s2 c2 s3) let ?cc1 = "ccomp c1" let ?cc2 = "ccomp c2" have "?cc1 @ ?cc2 \ (0,s1,stk) \* (isize ?cc1,s2,stk)" using Seq.IH(1) by fastforce moreover have "?cc1 @ ?cc2 \ (isize ?cc1,s2,stk) \* (isize(?cc1 @ ?cc2),s3,stk)" using Seq.IH(2) by fastforce ultimately show ?case by simp (blast intro: exec_trans) next case IfTrue thus ?case by (fastforce intro!: exec_append_trans) next case (IfFalse b s c2 t c1 stk) let ?cc1 = "ccomp c1" let ?cc2 = "ccomp c2" let ?cb = "bcomp b" let ?ci = "ccomp (IF b THEN c1 ELSE c2)" let ?cii = "?cb @ JMP0 (isize ?cc1 + 1) # ?cc1 @ [JMP (isize ?cc2)]" have "?ci \ (0,s,stk) \* (2 + isize ?cb + isize ?cc1,s,stk)" using IfFalse by (fastforce intro!: exec_append_trans) moreover have "?ci \ (2 + isize ?cb + isize ?cc1,s,stk) \* (isize ?ci, t, stk)" using exec_appendL[OF IfFalse(3), of ?cii] by simp smt ultimately show ?case by(blast intro!: exec_trans) next case (WhileTrue b s1 c s2 s3) let ?cc = "ccomp c" let ?cb = "bcomp b" let ?cw = "ccomp(WHILE b DO c)" have "?cw \ (0,s1,stk) \* (isize ?cb + isize ?cc + 1,s2,stk)" using WhileTrue.IH(1) WhileTrue.hyps(1) by (fastforce intro!: exec_append_trans) moreover have "?cw \ (isize ?cb + isize ?cc + 1,s2,stk) \* (0,s2,stk)" by (fastforce simp add: nth_append) moreover have "?cw \ (0,s2,stk) \* (isize ?cw,s3,stk)" by(rule WhileTrue.IH(2)) ultimately show ?case by(blast intro!: exec_trans) next case WhileFalse thus ?case by (fastforce intro!: exec_append_trans) qed fastforce end