diff --git a/src/HOL/IMP/Compiler.thy b/src/HOL/IMP/Compiler.thy --- a/src/HOL/IMP/Compiler.thy +++ b/src/HOL/IMP/Compiler.thy @@ -1,261 +1,261 @@ (* Author: Tobias Nipkow and Gerwin Klein *) section "Compiler for IMP" theory Compiler imports Big_Step Star begin subsection "List setup" text \ In the following, we use the length of lists as integers instead of natural numbers. Instead of converting \<^typ>\nat\ to \<^typ>\int\ explicitly, we tell Isabelle to coerce \<^typ>\nat\ automatically when necessary. \ declare [[coercion_enabled]] declare [[coercion "int :: nat \ int"]] text \ Similarly, we will want to access the ith element of a list, where \<^term>\i\ is an \<^typ>\int\. \ fun inth :: "'a list \ int \ 'a" (infixl "!!" 100) where "(x # xs) !! i = (if i = 0 then x else xs !! (i - 1))" text \ The only additional lemma we need about this function is indexing over append: \ lemma inth_append [simp]: "0 \ i \ (xs @ ys) !! i = (if i < size xs then xs !! i else ys !! (i - size xs))" by (induction xs arbitrary: i) (auto simp: algebra_simps) text\We hide coercion \<^const>\int\ applied to \<^const>\length\:\ abbreviation (output) "isize xs == int (length xs)" notation isize ("size") subsection "Instructions and Stack Machine" text_raw\\snip{instrdef}{0}{1}{%\ datatype instr = LOADI int | LOAD vname | ADD | STORE vname | JMP int | JMPLESS int | JMPGE int text_raw\}%endsnip\ 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) | JMPLESS n \ (if hd2 stk < hd stk then i+1+n else i+1,s,tl2 stk) | JMPGE n \ (if hd2 stk >= hd stk then i+1+n else i+1,s,tl2 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 < size P)" lemma exec1I [intro, code_pred_intro]: "c' = iexec (P!!i) (i,s,stk) \ 0 \ i \ i < size P \ P \ (i,s,stk) \ c'" by (simp add: exec1_def) abbreviation exec :: "instr list \ config \ config \ bool" ("(_/ \ (_ \*/ _))" 50) where "exec P \ star (exec1 P)" lemmas exec_induct = star.induct [of "exec1 P", split_format(complete)] code_pred exec1 by (metis exec1_def) 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\ 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 (auto simp: exec1_def) lemma exec_appendR: "P \ c \* c' \ P@P' \ c \* c'" by (induction rule: star.induct) (fastforce intro: star.step exec1_appendR)+ lemma exec1_appendL: fixes i i' :: int shows "P \ (i,s,stk) \ (i',s',stk') \ P' @ P \ (size(P')+i,s,stk) \ (size(P')+i',s',stk')" unfolding exec1_def by (auto simp del: iexec.simps) lemma exec_appendL: fixes i i' :: int shows "P \ (i,s,stk) \* (i',s',stk') \ P' @ P \ (size(P')+i,s,stk) \* (size(P')+i',s',stk')" by (induction rule: exec_induct) (blast intro: star.step exec1_appendL)+ text\Now we specialise the above lemmas to enable automatic proofs of \<^prop>\P \ c \* c'\ where \P\ is a mixture of concrete instructions and pieces of code that we already know how they execute (by induction), combined by \@\ and \#\. 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]: fixes i i' j :: int shows "size P' <= i \ P \ (i - size P',s,stk) \* (j,s',stk') \ i' = size P' + j \ P' @ P \ (i,s,stk) \* (i',s',stk')" by (drule exec_appendL[where P'=P']) simp text\Split the execution of a compound program up into the execution of its parts:\ lemma exec_append_trans[intro]: fixes i' i'' j'' :: int shows "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 star_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) \* (size(acomp a),s,aval a s#stk)" by (induction a arbitrary: stk) fastforce+ fun bcomp :: "bexp \ bool \ int \ instr list" where "bcomp (Bc v) f n = (if v=f then [JMP n] else [])" | "bcomp (Not b) f n = bcomp b (\f) n" | "bcomp (And b1 b2) f n = (let cb2 = bcomp b2 f n; - m = if f then size cb2 else (size cb2::int)+n; + m = if f then size cb2 else (size cb2)+n; cb1 = bcomp b1 False m in cb1 @ cb2)" | "bcomp (Less a1 a2) f n = acomp a1 @ acomp a2 @ (if f then [JMPLESS n] else [JMPGE n])" value "bcomp (And (Less (V ''x'') (V ''y'')) (Not(Less (V ''u'') (V ''v'')))) False 3" lemma bcomp_correct[intro]: fixes n :: int shows "0 \ n \ bcomp b f n \ (0,s,stk) \* (size(bcomp b f n) + (if f = bval b s then n else 0),s,stk)" proof(induction b arbitrary: f n) case Not from Not(1)[where f="~f"] Not(2) show ?case by fastforce next case (And b1 b2) from And(1)[of "if f then size(bcomp b2 f n) else size(bcomp b2 f n) + n" "False"] And(2)[of n f] And(3) show ?case by fastforce qed fastforce+ fun ccomp :: "com \ instr list" where "ccomp SKIP = []" | "ccomp (x ::= a) = acomp a @ [STORE x]" | "ccomp (c\<^sub>1;;c\<^sub>2) = ccomp c\<^sub>1 @ ccomp c\<^sub>2" | "ccomp (IF b THEN c\<^sub>1 ELSE c\<^sub>2) = (let cc\<^sub>1 = ccomp c\<^sub>1; cc\<^sub>2 = ccomp c\<^sub>2; cb = bcomp b False (size cc\<^sub>1 + 1) in cb @ cc\<^sub>1 @ JMP (size cc\<^sub>2) # cc\<^sub>2)" | "ccomp (WHILE b DO c) = (let cc = ccomp c; cb = bcomp b False (size cc + 1) in cb @ cc @ [JMP (-(size cb + size cc + 1))])" 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) \* (size(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) next case (Seq c1 s1 s2 c2 s3) let ?cc1 = "ccomp c1" let ?cc2 = "ccomp c2" have "?cc1 @ ?cc2 \ (0,s1,stk) \* (size ?cc1,s2,stk)" using Seq.IH(1) by fastforce moreover have "?cc1 @ ?cc2 \ (size ?cc1,s2,stk) \* (size(?cc1 @ ?cc2),s3,stk)" using Seq.IH(2) by fastforce ultimately show ?case by simp (blast intro: star_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,s1,stk)" using \bval b s1\ by fastforce moreover have "?cw \ (size ?cb,s1,stk) \* (size ?cb + size ?cc,s2,stk)" using WhileTrue.IH(1) by fastforce moreover have "?cw \ (size ?cb + size ?cc,s2,stk) \* (0,s2,stk)" by fastforce moreover have "?cw \ (0,s2,stk) \* (size ?cw,s3,stk)" by(rule WhileTrue.IH(2)) ultimately show ?case by(blast intro: star_trans) qed fastforce+ end