diff -r 643fce75f6cd -r 8ed413a57bdc src/HOL/IMP/Compiler.thy --- a/src/HOL/IMP/Compiler.thy Thu Apr 25 17:36:29 2002 +0200 +++ b/src/HOL/IMP/Compiler.thy Fri Apr 26 11:47:01 2002 +0200 @@ -4,46 +4,7 @@ Copyright 1996 TUM *) -header "A Simple Compiler" - -theory Compiler = Natural: - -subsection "An abstract, simplistic machine" - -text {* There are only three instructions: *} -datatype instr = ASIN loc aexp | JMPF bexp nat | JMPB nat - -text {* We describe execution of programs in the machine by - an operational (small step) semantics: -*} -consts stepa1 :: "instr list \ ((state\nat) \ (state\nat))set" - -syntax - "_stepa1" :: "[instr list,state,nat,state,nat] \ bool" - ("_ |- <_,_>/ -1-> <_,_>" [50,0,0,0,0] 50) - "_stepa" :: "[instr list,state,nat,state,nat] \ bool" - ("_ |-/ <_,_>/ -*-> <_,_>" [50,0,0,0,0] 50) - -syntax (xsymbols) - "_stepa1" :: "[instr list,state,nat,state,nat] \ bool" - ("_ \ \_,_\/ -1\ \_,_\" [50,0,0,0,0] 50) - "_stepa" :: "[instr list,state,nat,state,nat] \ bool" - ("_ \/ \_,_\/ -*\ \_,_\" [50,0,0,0,0] 50) - -translations - "P \ \s,m\ -1\ \t,n\" == "((s,m),t,n) : stepa1 P" - "P \ \s,m\ -*\ \t,n\" == "((s,m),t,n) : ((stepa1 P)^*)" - -inductive "stepa1 P" -intros -ASIN[simp]: - "\ n \ P \ \s,n\ -1\ \s[x\ a s],Suc n\" -JMPFT[simp,intro]: - "\ n \ P \ \s,n\ -1\ \s,Suc n\" -JMPFF[simp,intro]: - "\ n \ P \ \s,n\ -1\ \s,m\" -JMPB[simp]: - "\ n \ P \ \s,n\ -1\ \s,j\" +theory Compiler = Machines: subsection "The compiler" @@ -53,216 +14,287 @@ "compile (x:==a) = [ASIN x a]" "compile (c1;c2) = compile c1 @ compile c2" "compile (\ b \ c1 \ c2) = - [JMPF b (length(compile c1) + 2)] @ compile c1 @ - [JMPF (%x. False) (length(compile c2)+1)] @ compile c2" -"compile (\ b \ c) = [JMPF b (length(compile c) + 2)] @ compile c @ + [JMPF b (length(compile c1) + 1)] @ compile c1 @ + [JMPF (%x. False) (length(compile c2))] @ compile c2" +"compile (\ b \ c) = [JMPF b (length(compile c) + 1)] @ compile c @ [JMPB (length(compile c)+1)]" -declare nth_append[simp] - -subsection "Context lifting lemmas" - -text {* - Some lemmas for lifting an execution into a prefix and suffix - of instructions; only needed for the first proof. -*} -lemma app_right_1: - "is1 \ \s1,i1\ -1\ \s2,i2\ \ is1 @ is2 \ \s1,i1\ -1\ \s2,i2\" - (is "?P \ _") -proof - - assume ?P - then show ?thesis - by induct force+ -qed - -lemma app_left_1: - "is2 \ \s1,i1\ -1\ \s2,i2\ \ - is1 @ is2 \ \s1,size is1+i1\ -1\ \s2,size is1+i2\" - (is "?P \ _") -proof - - assume ?P - then show ?thesis - by induct force+ -qed - -declare rtrancl_induct2 [induct set: rtrancl] - -lemma app_right: - "is1 \ \s1,i1\ -*\ \s2,i2\ \ is1 @ is2 \ \s1,i1\ -*\ \s2,i2\" - (is "?P \ _") -proof - - assume ?P - then show ?thesis - proof induct - show "is1 @ is2 \ \s1,i1\ -*\ \s1,i1\" by simp - next - fix s1' i1' s2 i2 - assume "is1 @ is2 \ \s1,i1\ -*\ \s1',i1'\" - "is1 \ \s1',i1'\ -1\ \s2,i2\" - thus "is1 @ is2 \ \s1,i1\ -*\ \s2,i2\" - by(blast intro:app_right_1 rtrancl_trans) - qed -qed - -lemma app_left: - "is2 \ \s1,i1\ -*\ \s2,i2\ \ - is1 @ is2 \ \s1,size is1+i1\ -*\ \s2,size is1+i2\" - (is "?P \ _") -proof - - assume ?P - then show ?thesis - proof induct - show "is1 @ is2 \ \s1,length is1 + i1\ -*\ \s1,length is1 + i1\" by simp - next - fix s1' i1' s2 i2 - assume "is1 @ is2 \ \s1,length is1 + i1\ -*\ \s1',length is1 + i1'\" - "is2 \ \s1',i1'\ -1\ \s2,i2\" - thus "is1 @ is2 \ \s1,length is1 + i1\ -*\ \s2,length is1 + i2\" - by(blast intro:app_left_1 rtrancl_trans) - qed -qed - -lemma app_left2: - "\ is2 \ \s1,i1\ -*\ \s2,i2\; j1 = size is1+i1; j2 = size is1+i2 \ \ - is1 @ is2 \ \s1,j1\ -*\ \s2,j2\" - by (simp add:app_left) - -lemma app1_left: - "is \ \s1,i1\ -*\ \s2,i2\ \ - instr # is \ \s1,Suc i1\ -*\ \s2,Suc i2\" - by(erule app_left[of _ _ _ _ _ "[instr]",simplified]) - subsection "Compiler correctness" -declare rtrancl_into_rtrancl[trans] - converse_rtrancl_into_rtrancl[trans] - rtrancl_trans[trans] - -text {* - The first proof; The statement is very intuitive, - but application of induction hypothesis requires the above lifting lemmas -*} -theorem "\c,s\ \\<^sub>c t \ compile c \ \s,0\ -*\ \t,length(compile c)\" - (is "?P \ ?Q c s t") +theorem assumes A: "\c,s\ \\<^sub>c t" +shows "\p q. \compile c @ p,q,s\ -*\ \p,rev(compile c)@q,t\" + (is "\p q. ?P c s t p q") proof - - assume ?P - then show ?thesis + from A show "\p q. ?thesis p q" proof induct - show "\s. ?Q \ s s" by simp + case Skip thus ?case by simp next - show "\a s x. ?Q (x :== a) s (s[x\ a s])" by force + case Assign thus ?case by force + next + case Semi thus ?case by simp (blast intro:rtrancl_trans) next - fix c0 c1 s0 s1 s2 - assume "?Q c0 s0 s1" - hence "compile c0 @ compile c1 \ \s0,0\ -*\ \s1,length(compile c0)\" - by(rule app_right) - moreover assume "?Q c1 s1 s2" - hence "compile c0 @ compile c1 \ \s1,length(compile c0)\ -*\ - \s2,length(compile c0)+length(compile c1)\" - proof - - note app_left[of _ 0] - thus - "\is1 is2 s1 s2 i2. - is2 \ \s1,0\ -*\ \s2,i2\ \ - is1 @ is2 \ \s1,size is1\ -*\ \s2,size is1+i2\" - by simp - qed - ultimately have "compile c0 @ compile c1 \ \s0,0\ -*\ - \s2,length(compile c0)+length(compile c1)\" - by (rule rtrancl_trans) - thus "?Q (c0; c1) s0 s2" by simp + fix b c0 c1 s0 s1 p q + assume IH: "\p q. ?P c0 s0 s1 p q" + assume "b s0" + thus "?P (\ b \ c0 \ c1) s0 s1 p q" + by(simp add: IH[THEN rtrancl_trans]) + next + case IfFalse thus ?case by(simp) next - fix b c0 c1 s0 s1 - let ?comp = "compile(\ b \ c0 \ c1)" - assume "b s0" and IH: "?Q c0 s0 s1" - hence "?comp \ \s0,0\ -1\ \s0,1\" by auto - also from IH - have "?comp \ \s0,1\ -*\ \s1,length(compile c0)+1\" - by(auto intro:app1_left app_right) - also have "?comp \ \s1,length(compile c0)+1\ -1\ \s1,length ?comp\" - by(auto) - finally show "?Q (\ b \ c0 \ c1) s0 s1" . + case WhileFalse thus ?case by simp next - fix b c0 c1 s0 s1 - let ?comp = "compile(\ b \ c0 \ c1)" - assume "\b s0" and IH: "?Q c1 s0 s1" - hence "?comp \ \s0,0\ -1\ \s0,length(compile c0) + 2\" by auto - also from IH - have "?comp \ \s0,length(compile c0)+2\ -*\ \s1,length ?comp\" - by(force intro!:app_left2 app1_left) - finally show "?Q (\ b \ c0 \ c1) s0 s1" . - next - fix b c and s::state - assume "\b s" - thus "?Q (\ b \ c) s s" by force - next - fix b c and s0::state and s1 s2 - let ?comp = "compile(\ b \ c)" - assume "b s0" and - IHc: "?Q c s0 s1" and IHw: "?Q (\ b \ c) s1 s2" - hence "?comp \ \s0,0\ -1\ \s0,1\" by auto - also from IHc - have "?comp \ \s0,1\ -*\ \s1,length(compile c)+1\" - by(auto intro:app1_left app_right) - also have "?comp \ \s1,length(compile c)+1\ -1\ \s1,0\" by simp - also note IHw - finally show "?Q (\ b \ c) s0 s2". + fix b c and s0::state and s1 s2 p q + assume b: "b s0" and + IHc: "\p q. ?P c s0 s1 p q" and + IHw: "\p q. ?P (\ b \ c) s1 s2 p q" + show "?P (\ b \ c) s0 s2 p q" + using b IHc[THEN rtrancl_trans] IHw by(simp) qed qed -text {* - Second proof; statement is generalized to cater for prefixes and suffixes; - needs none of the lifting lemmas, but instantiations of pre/suffix. - *} -theorem "\c,s\ \\<^sub>c t \ - !a z. a@compile c@z \ \s,length a\ -*\ \t,length a + length(compile c)\"; -apply(erule evalc.induct) - apply simp - apply(force intro!: ASIN) - apply(intro strip) - apply(erule_tac x = a in allE) - apply(erule_tac x = "a@compile c0" in allE) - apply(erule_tac x = "compile c1@z" in allE) - apply(erule_tac x = z in allE) - apply(simp add:add_assoc[THEN sym]) - apply(blast intro:rtrancl_trans) -(* \ b \ c0 \ c1; case b is true *) - apply(intro strip) - (* instantiate assumption sufficiently for later: *) - apply(erule_tac x = "a@[?I]" in allE) - apply(simp) - (* execute JMPF: *) - apply(rule converse_rtrancl_into_rtrancl) - apply(force intro!: JMPFT) - (* execute compile c0: *) - apply(rule rtrancl_trans) - apply(erule allE) - apply assumption - (* execute JMPF: *) - apply(rule r_into_rtrancl) - apply(force intro!: JMPFF) -(* end of case b is true *) - apply(intro strip) - apply(erule_tac x = "a@[?I]@compile c0@[?J]" in allE) - apply(simp add:add_assoc) - apply(rule converse_rtrancl_into_rtrancl) - apply(force intro!: JMPFF) - apply(blast) - apply(force intro: JMPFF) -apply(intro strip) -apply(erule_tac x = "a@[?I]" in allE) -apply(erule_tac x = a in allE) -apply(simp) -apply(rule converse_rtrancl_into_rtrancl) - apply(force intro!: JMPFT) -apply(rule rtrancl_trans) - apply(erule allE) - apply assumption -apply(rule converse_rtrancl_into_rtrancl) - apply(force intro!: JMPB) -apply(simp) +text {* The other direction! *} + +inductive_cases [elim!]: "(([],p,s),next) : stepa1" + +lemma [simp]: "(\[],q,s\ -n\ \p',q',t\) = (n=0 \ p' = [] \ q' = q \ t = s)" +apply(rule iffI) + apply(erule converse_rel_powE, simp, fast) +apply simp +done + +lemma [simp]: "(\[],q,s\ -*\ \p',q',t\) = (p' = [] \ q' = q \ t = s)" +by(simp add: rtrancl_is_UN_rel_pow) + +constdefs + forws :: "instr \ nat set" +"forws instr == case instr of + ASIN x a \ {0} | + JMPF b n \ {0,n} | + JMPB n \ {}" + backws :: "instr \ nat set" +"backws instr == case instr of + ASIN x a \ {} | + JMPF b n \ {} | + JMPB n \ {n}" + +consts closed :: "nat \ nat \ instr list \ bool" +primrec +"closed m n [] = True" +"closed m n (instr#is) = ((\j \ forws instr. j \ size is+n) \ + (\j \ backws instr. j \ m) \ closed (Suc m) n is)" + +lemma [simp]: + "\m n. closed m n (C1@C2) = + (closed m (n+size C2) C1 \ closed (m+size C1) n C2)" +by(induct C1, simp, simp add:plus_ac0) + +theorem [simp]: "\m n. closed m n (compile c)" +by(induct c, simp_all add:backws_def forws_def) + +lemma drop_lem: "n \ size(p1@p2) + \ (p1' @ p2 = drop n p1 @ drop (n - size p1) p2) = + (n \ size p1 & p1' = drop n p1)" +apply(rule iffI) + defer apply simp +apply(subgoal_tac "n \ size p1") + apply(rotate_tac -1) + apply simp +apply(rule ccontr) +apply(rotate_tac -1) +apply(drule_tac f = length in arg_cong) +apply simp +apply arith +done + +lemma reduce_exec1: + "\i # p1 @ p2,q1 @ q2,s\ -1\ \p1' @ p2,q1' @ q2,s'\ \ + \i # p1,q1,s\ -1\ \p1',q1',s'\" +by(clarsimp simp add: drop_lem split:instr.split_asm split_if_asm) + + +lemma closed_exec1: + "\ closed 0 0 (rev q1 @ instr # p1); + \instr # p1 @ p2, q1 @ q2,r\ -1\ \p',q',r'\ \ \ + \p1' q1'. p' = p1'@p2 \ q' = q1'@q2 \ rev q1' @ p1' = rev q1 @ instr # p1" +apply(clarsimp simp add:forws_def backws_def + split:instr.split_asm split_if_asm) done -text {* Missing: the other direction! *} +theorem closed_execn_decomp: "\C1 C2 r. + \ closed 0 0 (rev C1 @ C2); + \C2 @ p1 @ p2, C1 @ q,r\ -n\ \p2,rev p1 @ rev C2 @ C1 @ q,t\ \ + \ \s n1 n2. \C2,C1,r\ -n1\ \[],rev C2 @ C1,s\ \ + \p1@p2,rev C2 @ C1 @ q,s\ -n2\ \p2, rev p1 @ rev C2 @ C1 @ q,t\ \ + n = n1+n2" +(is "\C1 C2 r. \?CL C1 C2; ?H C1 C2 r n\ \ ?P C1 C2 r n") +proof(induct n) + fix C1 C2 r + assume "?H C1 C2 r 0" + thus "?P C1 C2 r 0" by simp +next + fix C1 C2 r n + assume IH: "\C1 C2 r. ?CL C1 C2 \ ?H C1 C2 r n \ ?P C1 C2 r n" + assume CL: "?CL C1 C2" and H: "?H C1 C2 r (Suc n)" + show "?P C1 C2 r (Suc n)" + proof (cases C2) + assume "C2 = []" with H show ?thesis by simp + next + fix instr tlC2 + assume C2: "C2 = instr # tlC2" + from H C2 obtain p' q' r' + where 1: "\instr # tlC2 @ p1 @ p2, C1 @ q,r\ -1\ \p',q',r'\" + and n: "\p',q',r'\ -n\ \p2,rev p1 @ rev C2 @ C1 @ q,t\" + by(fastsimp simp add:R_O_Rn_commute) + from CL closed_exec1[OF _ 1] C2 + obtain C2' C1' where pq': "p' = C2' @ p1 @ p2 \ q' = C1' @ q" + and same: "rev C1' @ C2' = rev C1 @ C2" + by fastsimp + have rev_same: "rev C2' @ C1' = rev C2 @ C1" + proof - + have "rev C2' @ C1' = rev(rev C1' @ C2')" by simp + also have "\ = rev(rev C1 @ C2)" by(simp only:same) + also have "\ = rev C2 @ C1" by simp + finally show ?thesis . + qed + hence rev_same': "\p. rev C2' @ C1' @ p = rev C2 @ C1 @ p" by simp + from n have n': "\C2' @ p1 @ p2,C1' @ q,r'\ -n\ + \p2,rev p1 @ rev C2' @ C1' @ q,t\" + by(simp add:pq' rev_same') + from IH[OF _ n'] CL + obtain s n1 n2 where n1: "\C2',C1',r'\ -n1\ \[],rev C2 @ C1,s\" and + "\p1 @ p2,rev C2 @ C1 @ q,s\ -n2\ \p2,rev p1 @ rev C2 @ C1 @ q,t\ \ + n = n1 + n2" + by(fastsimp simp add: same rev_same rev_same') + moreover + from 1 n1 pq' C2 have "\C2,C1,r\ -Suc n1\ \[],rev C2 @ C1,s\" + by (simp del:relpow.simps exec_simp) (fast dest:reduce_exec1) + ultimately show ?thesis by (fastsimp simp del:relpow.simps) + qed +qed + +lemma execn_decomp: +"\compile c @ p1 @ p2,q,r\ -n\ \p2,rev p1 @ rev(compile c) @ q,t\ + \ \s n1 n2. \compile c,[],r\ -n1\ \[],rev(compile c),s\ \ + \p1@p2,rev(compile c) @ q,s\ -n2\ \p2, rev p1 @ rev(compile c) @ q,t\ \ + n = n1+n2" +using closed_execn_decomp[of "[]",simplified] by simp + +lemma exec_star_decomp: +"\compile c @ p1 @ p2,q,r\ -*\ \p2,rev p1 @ rev(compile c) @ q,t\ + \ \s. \compile c,[],r\ -*\ \[],rev(compile c),s\ \ + \p1@p2,rev(compile c) @ q,s\ -*\ \p2, rev p1 @ rev(compile c) @ q,t\" +by(simp add:rtrancl_is_UN_rel_pow)(fast dest: execn_decomp) + + +(* Alternative: +lemma exec_comp_n: +"\p1 p2 q r t n. + \compile c @ p1 @ p2,q,r\ -n\ \p2,rev p1 @ rev(compile c) @ q,t\ + \ \s n1 n2. \compile c,[],r\ -n1\ \[],rev(compile c),s\ \ + \p1@p2,rev(compile c) @ q,s\ -n2\ \p2, rev p1 @ rev(compile c) @ q,t\ \ + n = n1+n2" + (is "\p1 p2 q r t n. ?H c p1 p2 q r t n \ ?P c p1 p2 q r t n") +proof (induct c) +*) + +text{*Warning: +@{prop"\compile c @ p,q,s\ -*\ \p,rev(compile c)@q,t\ \ \c,s\ \\<^sub>c t"} +is not true! *} -end +theorem "\s t. + \compile c,[],s\ -*\ \[],rev(compile c),t\ \ \c,s\ \\<^sub>c t" +proof (induct c) + fix s t + assume "\compile SKIP,[],s\ -*\ \[],rev(compile SKIP),t\" + thus "\SKIP,s\ \\<^sub>c t" by simp +next + fix s t v f + assume "\compile(v :== f),[],s\ -*\ \[],rev(compile(v :== f)),t\" + thus "\v :== f,s\ \\<^sub>c t" by simp +next + fix s1 s3 c1 c2 + let ?C1 = "compile c1" let ?C2 = "compile c2" + assume IH1: "\s t. \?C1,[],s\ -*\ \[],rev ?C1,t\ \ \c1,s\ \\<^sub>c t" + and IH2: "\s t. \?C2,[],s\ -*\ \[],rev ?C2,t\ \ \c2,s\ \\<^sub>c t" + assume "\compile(c1;c2),[],s1\ -*\ \[],rev(compile(c1;c2)),s3\" + then obtain s2 where exec1: "\?C1,[],s1\ -*\ \[],rev ?C1,s2\" and + exec2: "\?C2,rev ?C1,s2\ -*\ \[],rev(compile(c1;c2)),s3\" + by(fastsimp dest:exec_star_decomp[of _ _ "[]" "[]",simplified]) + from exec2 have exec2': "\?C2,[],s2\ -*\ \[],rev ?C2,s3\" + using exec_star_decomp[of _ "[]" "[]"] by fastsimp + have "\c1,s1\ \\<^sub>c s2" using IH1 exec1 by simp + moreover have "\c2,s2\ \\<^sub>c s3" using IH2 exec2' by fastsimp + ultimately show "\c1;c2,s1\ \\<^sub>c s3" .. +next + fix s t b c1 c2 + let ?if = "IF b THEN c1 ELSE c2" let ?C = "compile ?if" + let ?C1 = "compile c1" let ?C2 = "compile c2" + assume IH1: "\s t. \?C1,[],s\ -*\ \[],rev ?C1,t\ \ \c1,s\ \\<^sub>c t" + and IH2: "\s t. \?C2,[],s\ -*\ \[],rev ?C2,t\ \ \c2,s\ \\<^sub>c t" + and H: "\?C,[],s\ -*\ \[],rev ?C,t\" + show "\?if,s\ \\<^sub>c t" + proof cases + assume b: "b s" + with H have "\?C1,[],s\ -*\ \[],rev ?C1,t\" + by (fastsimp dest:exec_star_decomp + [of _ "[JMPF (\x. False) (size ?C2)]@?C2" "[]",simplified]) + hence "\c1,s\ \\<^sub>c t" by(rule IH1) + with b show ?thesis .. + next + assume b: "\ b s" + with H have "\?C2,[],s\ -*\ \[],rev ?C2,t\" + using exec_star_decomp[of _ "[]" "[]"] by simp + hence "\c2,s\ \\<^sub>c t" by(rule IH2) + with b show ?thesis .. + qed +next + fix b c s t + let ?w = "WHILE b DO c" let ?W = "compile ?w" let ?C = "compile c" + let ?j1 = "JMPF b (size ?C + 1)" let ?j2 = "JMPB (size ?C + 1)" + assume IHc: "\s t. \?C,[],s\ -*\ \[],rev ?C,t\ \ \c,s\ \\<^sub>c t" + and H: "\?W,[],s\ -*\ \[],rev ?W,t\" + from H obtain k where ob:"\?W,[],s\ -k\ \[],rev ?W,t\" + by(simp add:rtrancl_is_UN_rel_pow) blast + { fix n have "\s. \?W,[],s\ -n\ \[],rev ?W,t\ \ \?w,s\ \\<^sub>c t" + proof (induct n rule: less_induct) + fix n + assume IHm: "\m s. \m < n; \?W,[],s\ -m\ \[],rev ?W,t\ \ \ \?w,s\ \\<^sub>c t" + fix s + assume H: "\?W,[],s\ -n\ \[],rev ?W,t\" + show "\?w,s\ \\<^sub>c t" + proof cases + assume b: "b s" + then obtain m where m: "n = Suc m" + and 1: "\?C @ [?j2],[?j1],s\ -m\ \[],rev ?W,t\" + using H by fastsimp + then obtain r n1 n2 where n1: "\?C,[],s\ -n1\ \[],rev ?C,r\" + and n2: "\[?j2],rev ?C @ [?j1],r\ -n2\ \[],rev ?W,t\" + and n12: "m = n1+n2" + using execn_decomp[of _ "[?j2]"] + by(simp del: execn_simp) fast + have n2n: "n2 - 1 < n" using m n12 by arith + note b + moreover + { from n1 have "\?C,[],s\ -*\ \[],rev ?C,r\" + by (simp add:rtrancl_is_UN_rel_pow) fast + hence "\c,s\ \\<^sub>c r" by(rule IHc) + } + moreover + { have "n2 - 1 < n" using m n12 by arith + moreover from n2 have "\?W,[],r\ -n2- 1\ \[],rev ?W,t\" by fastsimp + ultimately have "\?w,r\ \\<^sub>c t" by(rule IHm) + } + ultimately show ?thesis .. + next + assume b: "\ b s" + hence "t = s" using H by simp + with b show ?thesis by simp + qed + qed + } + with ob show "\?w,s\ \\<^sub>c t" by fast +qed + +(* To Do: connect with Machine 0 using M_equiv *) + +end \ No newline at end of file