(* Author: Tobias Nipkow *) header "A Compiler for IMP" theory Compiler imports Big_Step begin subsection "Instructions and Stack Machine" datatype instr = PUSH_N nat | PUSH_V nat | ADD | STORE nat | JMPF nat | JMPB nat | JMPFLESS nat | JMPFGE nat types stack = "nat list" config = "nat\state\stack" abbreviation "hd2 xs == hd(tl xs)" abbreviation "tl2 xs == tl(tl xs)" inductive exec1 :: "instr list \ config \ config \ bool" ("(_/ \ (_ \/ _))" [50,0,0] 50) for P :: "instr list" where "\ i < size P; P!i = PUSH_N n \ \ P \ (i,s,stk) \ (i+1,s, n#stk)" | "\ i < size P; P!i = PUSH_V x \ \ P \ (i,s,stk) \ (i+1,s, s x # stk)" | "\ i < size P; P!i = ADD \ \ P \ (i,s,stk) \ (i+1,s, (hd2 stk + hd stk) # tl2 stk)" | "\ i < size P; P!i = STORE n \ \ P \ (i,s,stk) \ (i+1,s(n := hd stk),tl stk)" | "\ i < size P; P!i = JMPF n \ \ P \ (i,s,stk) \ (i+1+n,s,stk)" | "\ i < size P; P!i = JMPB n; n \ i+1 \ \ P \ (i,s,stk) \ (i+1-n,s,stk)" | "\ i < size P; P!i = JMPFLESS n \ \ P \ (i,s,stk) \ (if hd2 stk < hd stk then i+1+n else i+1,s,tl2 stk)" | "\ i < size P; P!i = JMPFGE n \ \ P \ (i,s,stk) \ (if hd2 stk >= hd stk then i+1+n else i+1,s,tl2 stk)" code_pred exec1 . declare exec1.intros[intro] inductive exec :: "instr list \ config \ config \ bool" ("_/ \ (_ \*/ _)" 50) where refl: "P \ c \* c" | step: "P \ c \ c' \ P \ c' \* c'' \ P \ c \* c''" declare exec.intros[intro] lemmas exec_induct = exec.induct[split_format(complete)] code_pred exec . text{* Integrating the state to list transformation: *} inductive execl :: "instr list \ nat \ nat list \ stack \ nat \ nat list \ stack \ bool" where "P \ (i,nth ns,stk) \* (i',s',stk') \ execl P i ns stk i' (list s' (size ns)) stk'" code_pred execl . values "{(i,ns,stk). execl [PUSH_V 1, STORE 0] 0 [3,4] [] i ns stk}" subsection{* Verification infrastructure *} lemma exec_trans: "P \ c \* c' \ P \ c' \* c'' \ P \ c \* c''" apply(induct rule: exec.induct) apply blast by (metis exec.step) lemma exec1_subst: "P \ c \ c' \ c' = c'' \ P \ c \ c''" by auto lemmas exec1_simps = exec1.intros[THEN exec1_subst] 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 exec1_appendR: assumes "P \ c \ c'" shows "P@P' \ c \ c'" proof- from assms show ?thesis by cases (simp_all add: exec1_simps nth_append) -- "All cases proved with the final simp-all" qed lemma exec_appendR: "P \ c \* c' \ P@P' \ c \* c'" apply(induct rule: exec.induct) apply blast by (metis exec1_appendR exec.step) lemma exec1_appendL: assumes "P \ (i,s,stk) \ (i',s',stk')" shows "P' @ P \ (size(P')+i,s,stk) \ (size(P')+i',s',stk')" proof- from assms show ?thesis by cases (simp_all add: exec1_simps) qed lemma exec_appendL: "P \ (i,s,stk) \* (i',s',stk') \ P' @ P \ (size(P')+i,s,stk) \* (size(P')+i',s',stk')" apply(induct rule: exec_induct) apply blast by (blast intro: exec1_appendL exec.step) 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 the pc points beyond the first instruction or part of the program, drop it: *} lemma exec_Cons_Suc[intro]: "P \ (i,s,stk) \* (j,t,stk') \ instr#P \ (Suc i,s,stk) \* (Suc j,t,stk')" apply(drule exec_appendL[where P'="[instr]"]) apply simp done lemma exec_appendL_if[intro]: "size P' <= i \ P \ (i - size P',s,stk) \* (i',s',stk') \ P' @ P \ (i,s,stk) \* (size P' + i',s',stk')" apply(drule exec_appendL[where P'=P']) apply simp done 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') \ size P \ i' \ P' \ (i' - size P,s',stk') \* (i'',s'',stk'') \ j'' = size P + i'' \ P @ P' \ (0,s,stk) \* (j'',s'',stk'')" by(metis exec_trans[OF exec_appendR exec_appendL_if]) declare Let_def[simp] nat_number[simp] subsection "Compilation" fun acomp :: "aexp \ instr list" where "acomp (N n) = [PUSH_N n]" | "acomp (V n) = [PUSH_V n]" | "acomp (Plus a\<^isub>1 a\<^isub>2) = acomp a\<^isub>1 @ acomp a\<^isub>2 @ [ADD]" lemma acomp_correct[intro]: "acomp a \ (0,s,stk) \* (size(acomp a),s,aval a s#stk)" apply(induct a arbitrary: stk) apply(fastsimp)+ done fun bcomp :: "bexp \ bool \ nat \ instr list" where "bcomp (B v) c n = (if v=c then [JMPF n] else [])" | "bcomp (Not b) c n = bcomp b (\c) n" | "bcomp (And b\<^isub>1 b\<^isub>2) c n = (let cb\<^isub>2 = bcomp b\<^isub>2 c n; m = (if c then size cb\<^isub>2 else size cb\<^isub>2+n); cb\<^isub>1 = bcomp b\<^isub>1 False m in cb\<^isub>1 @ cb\<^isub>2)" | "bcomp (Less a\<^isub>1 a\<^isub>2) c n = acomp a\<^isub>1 @ acomp a\<^isub>2 @ (if c then [JMPFLESS n] else [JMPFGE n])" value "bcomp (And (Less (V 0) (V 1)) (Not(Less (V 2) (V 3)))) False 3" lemma bcomp_correct[intro]: "bcomp b c n \ (0,s,stk) \* (size(bcomp b c n) + (if c = bval b s then n else 0),s,stk)" proof(induct b arbitrary: c n m) case Not from Not[of "~c"] show ?case by fastsimp next case (And b1 b2) from And(1)[of "False"] And(2)[of "c"] show ?case by fastsimp qed fastsimp+ 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 False (size cc\<^isub>1 + 1) in cb @ cc\<^isub>1 @ JMPF(size cc\<^isub>2) # cc\<^isub>2)" | "ccomp (WHILE b DO c) = (let cc = ccomp c; cb = bcomp b False (size cc + 1) in cb @ cc @ [JMPB (size cb + size cc + 1)])" value "ccomp (IF Less (V 4) (N 1) THEN 4 ::= Plus (V 4) (N 1) ELSE 3 ::= V 4)" value "ccomp (WHILE Less (V 4) (N 1) DO (4 ::= Plus (V 4) (N 1)))" subsection "Preservation of sematics" lemma ccomp_correct: "(c,s) \ t \ ccomp c \ (0,s,stk) \* (size(ccomp c),t,stk)" proof(induct arbitrary: stk rule: big_step_induct) case (Assign x a s) show ?case by (fastsimp simp:fun_upd_def) next case (Semi c1 s1 s2 c2 s3) let ?cc1 = "ccomp c1" let ?cc2 = "ccomp c2" have "?cc1 @ ?cc2 \ (0,s1,stk) \* (size ?cc1,s2,stk)" using Semi.hyps(2) by (fastsimp) moreover have "?cc1 @ ?cc2 \ (size ?cc1,s2,stk) \* (size(?cc1 @ ?cc2),s3,stk)" using Semi.hyps(4) by (fastsimp) ultimately show ?case by simp (blast intro: exec_trans) next case (WhileTrue b s1 c s2 s3) let ?cc = "ccomp c" let ?cb = "bcomp b False (size ?cc + 1)" let ?cw = "ccomp(WHILE b DO c)" have "?cw \ (0,s1,stk) \* (size ?cb + size ?cc,s2,stk)" using WhileTrue(1,3) by fastsimp moreover have "?cw \ (size ?cb + size ?cc,s2,stk) \* (0,s2,stk)" by (fastsimp) moreover have "?cw \ (0,s2,stk) \* (size ?cw,s3,stk)" by(rule WhileTrue(5)) ultimately show ?case by(blast intro: exec_trans) qed fastsimp+ end