| author | haftmann | 
| Tue, 06 Jun 2006 14:56:42 +0200 | |
| changeset 19787 | b949911ecff5 | 
| parent 18372 | 2bffdf62fe7f | 
| child 23746 | a455e69c31cc | 
| permissions | -rw-r--r-- | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 1 | (* Title: HOL/IMP/Compiler.thy | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 2 | ID: $Id$ | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 3 | Author: Tobias Nipkow, TUM | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 4 | Copyright 1996 TUM | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 5 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 6 | This is an early version of the compiler, where the abstract machine | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 7 | has an explicit pc. This turned out to be awkward, and a second | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 8 | development was started. See Machines.thy and Compiler.thy. | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 9 | *) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 10 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 11 | header "A Simple Compiler" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 12 | |
| 16417 | 13 | theory Compiler0 imports Natural begin | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 14 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 15 | subsection "An abstract, simplistic machine" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 16 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 17 | text {* There are only three instructions: *}
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 18 | datatype instr = ASIN loc aexp | JMPF bexp nat | JMPB nat | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 19 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 20 | text {* We describe execution of programs in the machine by
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 21 | an operational (small step) semantics: | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 22 | *} | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 23 | consts stepa1 :: "instr list \<Rightarrow> ((state\<times>nat) \<times> (state\<times>nat))set" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 24 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 25 | syntax | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 26 | "_stepa1" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 27 |                ("_ |- (3<_,_>/ -1-> <_,_>)" [50,0,0,0,0] 50)
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 28 | "_stepa" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 29 |                ("_ |-/ (3<_,_>/ -*-> <_,_>)" [50,0,0,0,0] 50)
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 30 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 31 | "_stepan" :: "[instr list,state,nat,nat,state,nat] \<Rightarrow> bool" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 32 |                ("_ |-/ (3<_,_>/ -(_)-> <_,_>)" [50,0,0,0,0,0] 50)
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 33 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 34 | syntax (xsymbols) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 35 | "_stepa1" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 36 |                ("_ \<turnstile> (3\<langle>_,_\<rangle>/ -1\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 37 | "_stepa" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 38 |                ("_ \<turnstile>/ (3\<langle>_,_\<rangle>/ -*\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 39 | "_stepan" :: "[instr list,state,nat,nat,state,nat] \<Rightarrow> bool" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 40 |                ("_ \<turnstile>/ (3\<langle>_,_\<rangle>/ -(_)\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0,0] 50)
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 41 | |
| 14565 | 42 | syntax (HTML output) | 
| 43 | "_stepa1" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool" | |
| 44 |                ("_ |- (3\<langle>_,_\<rangle>/ -1\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)
 | |
| 45 | "_stepa" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool" | |
| 46 |                ("_ |-/ (3\<langle>_,_\<rangle>/ -*\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)
 | |
| 47 | "_stepan" :: "[instr list,state,nat,nat,state,nat] \<Rightarrow> bool" | |
| 48 |                ("_ |-/ (3\<langle>_,_\<rangle>/ -(_)\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0,0] 50)
 | |
| 49 | ||
| 18372 | 50 | translations | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 51 | "P \<turnstile> \<langle>s,m\<rangle> -1\<rightarrow> \<langle>t,n\<rangle>" == "((s,m),t,n) : stepa1 P" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 52 | "P \<turnstile> \<langle>s,m\<rangle> -*\<rightarrow> \<langle>t,n\<rangle>" == "((s,m),t,n) : ((stepa1 P)^*)" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 53 | "P \<turnstile> \<langle>s,m\<rangle> -(i)\<rightarrow> \<langle>t,n\<rangle>" == "((s,m),t,n) : ((stepa1 P)^i)" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 54 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 55 | inductive "stepa1 P" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 56 | intros | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 57 | ASIN[simp]: | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 58 | "\<lbrakk> n<size P; P!n = ASIN x a \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s[x\<mapsto> a s],Suc n\<rangle>" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 59 | JMPFT[simp,intro]: | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 60 | "\<lbrakk> n<size P; P!n = JMPF b i; b s \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s,Suc n\<rangle>" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 61 | JMPFF[simp,intro]: | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 62 | "\<lbrakk> n<size P; P!n = JMPF b i; ~b s; m=n+i \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s,m\<rangle>" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 63 | JMPB[simp]: | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 64 | "\<lbrakk> n<size P; P!n = JMPB i; i <= n; j = n-i \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>s,n\<rangle> -1\<rightarrow> \<langle>s,j\<rangle>" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 65 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 66 | subsection "The compiler" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 67 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 68 | consts compile :: "com \<Rightarrow> instr list" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 69 | primrec | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 70 | "compile \<SKIP> = []" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 71 | "compile (x:==a) = [ASIN x a]" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 72 | "compile (c1;c2) = compile c1 @ compile c2" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 73 | "compile (\<IF> b \<THEN> c1 \<ELSE> c2) = | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 74 | [JMPF b (length(compile c1) + 2)] @ compile c1 @ | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 75 | [JMPF (%x. False) (length(compile c2)+1)] @ compile c2" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 76 | "compile (\<WHILE> b \<DO> c) = [JMPF b (length(compile c) + 2)] @ compile c @ | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 77 | [JMPB (length(compile c)+1)]" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 78 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 79 | declare nth_append[simp] | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 80 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 81 | subsection "Context lifting lemmas" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 82 | |
| 18372 | 83 | text {*
 | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 84 | Some lemmas for lifting an execution into a prefix and suffix | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 85 | of instructions; only needed for the first proof. | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 86 | *} | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 87 | lemma app_right_1: | 
| 18372 | 88 | assumes "is1 \<turnstile> \<langle>s1,i1\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>" | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 89 | shows "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>" | 
| 18372 | 90 | using prems | 
| 91 | by induct auto | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 92 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 93 | lemma app_left_1: | 
| 18372 | 94 | assumes "is2 \<turnstile> \<langle>s1,i1\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>" | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 95 | shows "is1 @ is2 \<turnstile> \<langle>s1,size is1+i1\<rangle> -1\<rightarrow> \<langle>s2,size is1+i2\<rangle>" | 
| 18372 | 96 | using prems | 
| 97 | by induct auto | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 98 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 99 | declare rtrancl_induct2 [induct set: rtrancl] | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 100 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 101 | lemma app_right: | 
| 18372 | 102 | assumes "is1 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>" | 
| 103 | shows "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>" | |
| 104 | using prems | |
| 105 | proof induct | |
| 106 | show "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s1,i1\<rangle>" by simp | |
| 107 | next | |
| 108 | fix s1' i1' s2 i2 | |
| 109 | assume "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s1',i1'\<rangle>" | |
| 110 | and "is1 \<turnstile> \<langle>s1',i1'\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>" | |
| 111 | thus "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>" | |
| 112 | by (blast intro: app_right_1 rtrancl_trans) | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 113 | qed | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 114 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 115 | lemma app_left: | 
| 18372 | 116 | assumes "is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>" | 
| 117 | shows "is1 @ is2 \<turnstile> \<langle>s1,size is1+i1\<rangle> -*\<rightarrow> \<langle>s2,size is1+i2\<rangle>" | |
| 118 | using prems | |
| 119 | proof induct | |
| 120 | show "is1 @ is2 \<turnstile> \<langle>s1,length is1 + i1\<rangle> -*\<rightarrow> \<langle>s1,length is1 + i1\<rangle>" by simp | |
| 121 | next | |
| 122 | fix s1' i1' s2 i2 | |
| 123 | assume "is1 @ is2 \<turnstile> \<langle>s1,length is1 + i1\<rangle> -*\<rightarrow> \<langle>s1',length is1 + i1'\<rangle>" | |
| 124 | and "is2 \<turnstile> \<langle>s1',i1'\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>" | |
| 125 | thus "is1 @ is2 \<turnstile> \<langle>s1,length is1 + i1\<rangle> -*\<rightarrow> \<langle>s2,length is1 + i2\<rangle>" | |
| 126 | by (blast intro: app_left_1 rtrancl_trans) | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 127 | qed | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 128 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 129 | lemma app_left2: | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 130 | "\<lbrakk> is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>; j1 = size is1+i1; j2 = size is1+i2 \<rbrakk> \<Longrightarrow> | 
| 18372 | 131 | is1 @ is2 \<turnstile> \<langle>s1,j1\<rangle> -*\<rightarrow> \<langle>s2,j2\<rangle>" | 
| 132 | by (simp add: app_left) | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 133 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 134 | lemma app1_left: | 
| 18372 | 135 | assumes "is \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>" | 
| 136 | shows "instr # is \<turnstile> \<langle>s1,Suc i1\<rangle> -*\<rightarrow> \<langle>s2,Suc i2\<rangle>" | |
| 137 | proof - | |
| 138 | from app_left [OF prems, of "[instr]"] | |
| 139 | show ?thesis by simp | |
| 140 | qed | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 141 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 142 | subsection "Compiler correctness" | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 143 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 144 | declare rtrancl_into_rtrancl[trans] | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 145 | converse_rtrancl_into_rtrancl[trans] | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 146 | rtrancl_trans[trans] | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 147 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 148 | text {*
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 149 | The first proof; The statement is very intuitive, | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 150 | but application of induction hypothesis requires the above lifting lemmas | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 151 | *} | 
| 18372 | 152 | theorem | 
| 153 | assumes "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t" | |
| 154 | shows "compile c \<turnstile> \<langle>s,0\<rangle> -*\<rightarrow> \<langle>t,length(compile c)\<rangle>" (is "?P c s t") | |
| 155 | using prems | |
| 156 | proof induct | |
| 157 | show "\<And>s. ?P \<SKIP> s s" by simp | |
| 158 | next | |
| 159 | show "\<And>a s x. ?P (x :== a) s (s[x\<mapsto> a s])" by force | |
| 160 | next | |
| 161 | fix c0 c1 s0 s1 s2 | |
| 162 | assume "?P c0 s0 s1" | |
| 163 | hence "compile c0 @ compile c1 \<turnstile> \<langle>s0,0\<rangle> -*\<rightarrow> \<langle>s1,length(compile c0)\<rangle>" | |
| 164 | by (rule app_right) | |
| 165 | moreover assume "?P c1 s1 s2" | |
| 166 | hence "compile c0 @ compile c1 \<turnstile> \<langle>s1,length(compile c0)\<rangle> -*\<rightarrow> | |
| 167 | \<langle>s2,length(compile c0)+length(compile c1)\<rangle>" | |
| 168 | proof - | |
| 169 | show "\<And>is1 is2 s1 s2 i2. | |
| 170 | is2 \<turnstile> \<langle>s1,0\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle> \<Longrightarrow> | |
| 171 | is1 @ is2 \<turnstile> \<langle>s1,size is1\<rangle> -*\<rightarrow> \<langle>s2,size is1+i2\<rangle>" | |
| 172 | using app_left[of _ 0] by simp | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 173 | qed | 
| 18372 | 174 | ultimately have "compile c0 @ compile c1 \<turnstile> \<langle>s0,0\<rangle> -*\<rightarrow> | 
| 175 | \<langle>s2,length(compile c0)+length(compile c1)\<rangle>" | |
| 176 | by (rule rtrancl_trans) | |
| 177 | thus "?P (c0; c1) s0 s2" by simp | |
| 178 | next | |
| 179 | fix b c0 c1 s0 s1 | |
| 180 | let ?comp = "compile(\<IF> b \<THEN> c0 \<ELSE> c1)" | |
| 181 | assume "b s0" and IH: "?P c0 s0 s1" | |
| 182 | hence "?comp \<turnstile> \<langle>s0,0\<rangle> -1\<rightarrow> \<langle>s0,1\<rangle>" by auto | |
| 183 | also from IH | |
| 184 | have "?comp \<turnstile> \<langle>s0,1\<rangle> -*\<rightarrow> \<langle>s1,length(compile c0)+1\<rangle>" | |
| 185 | by(auto intro:app1_left app_right) | |
| 186 | also have "?comp \<turnstile> \<langle>s1,length(compile c0)+1\<rangle> -1\<rightarrow> \<langle>s1,length ?comp\<rangle>" | |
| 187 | by(auto) | |
| 188 | finally show "?P (\<IF> b \<THEN> c0 \<ELSE> c1) s0 s1" . | |
| 189 | next | |
| 190 | fix b c0 c1 s0 s1 | |
| 191 | let ?comp = "compile(\<IF> b \<THEN> c0 \<ELSE> c1)" | |
| 192 | assume "\<not>b s0" and IH: "?P c1 s0 s1" | |
| 193 | hence "?comp \<turnstile> \<langle>s0,0\<rangle> -1\<rightarrow> \<langle>s0,length(compile c0) + 2\<rangle>" by auto | |
| 194 | also from IH | |
| 195 | have "?comp \<turnstile> \<langle>s0,length(compile c0)+2\<rangle> -*\<rightarrow> \<langle>s1,length ?comp\<rangle>" | |
| 196 | by (force intro!: app_left2 app1_left) | |
| 197 | finally show "?P (\<IF> b \<THEN> c0 \<ELSE> c1) s0 s1" . | |
| 198 | next | |
| 199 | fix b c and s::state | |
| 200 | assume "\<not>b s" | |
| 201 | thus "?P (\<WHILE> b \<DO> c) s s" by force | |
| 202 | next | |
| 203 | fix b c and s0::state and s1 s2 | |
| 204 | let ?comp = "compile(\<WHILE> b \<DO> c)" | |
| 205 | assume "b s0" and | |
| 206 | IHc: "?P c s0 s1" and IHw: "?P (\<WHILE> b \<DO> c) s1 s2" | |
| 207 | hence "?comp \<turnstile> \<langle>s0,0\<rangle> -1\<rightarrow> \<langle>s0,1\<rangle>" by auto | |
| 208 | also from IHc | |
| 209 | have "?comp \<turnstile> \<langle>s0,1\<rangle> -*\<rightarrow> \<langle>s1,length(compile c)+1\<rangle>" | |
| 210 | by (auto intro: app1_left app_right) | |
| 211 | also have "?comp \<turnstile> \<langle>s1,length(compile c)+1\<rangle> -1\<rightarrow> \<langle>s1,0\<rangle>" by simp | |
| 212 | also note IHw | |
| 213 | finally show "?P (\<WHILE> b \<DO> c) s0 s2". | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 214 | qed | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 215 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 216 | text {*
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 217 | Second proof; statement is generalized to cater for prefixes and suffixes; | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 218 | needs none of the lifting lemmas, but instantiations of pre/suffix. | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 219 | *} | 
| 13130 | 220 | (* | 
| 13112 | 221 | theorem assumes A: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t" | 
| 222 | shows "\<And>a z. a@compile c@z \<turnstile> \<langle>s,size a\<rangle> -*\<rightarrow> \<langle>t,size a + size(compile c)\<rangle>" | |
| 223 | (is "\<And>a z. ?P c s t a z") | |
| 224 | proof - | |
| 225 | from A show "\<And>a z. ?thesis a z" | |
| 226 | proof induct | |
| 227 | case Skip thus ?case by simp | |
| 228 | next | |
| 229 | case Assign thus ?case by (force intro!: ASIN) | |
| 230 | next | |
| 231 | fix c1 c2 s s' s'' a z | |
| 232 | assume IH1: "\<And>a z. ?P c1 s s' a z" and IH2: "\<And>a z. ?P c2 s' s'' a z" | |
| 233 | from IH1 IH2[of "a@compile c1"] | |
| 234 | show "?P (c1;c2) s s'' a z" | |
| 235 | by(simp add:add_assoc[THEN sym])(blast intro:rtrancl_trans) | |
| 236 | next | |
| 237 | (* at this point I gave up converting to structured proofs *) | |
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 238 | (* \<IF> b \<THEN> c0 \<ELSE> c1; case b is true *) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 239 | apply(intro strip) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 240 | (* instantiate assumption sufficiently for later: *) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 241 | apply(erule_tac x = "a@[?I]" in allE) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 242 | apply(simp) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 243 | (* execute JMPF: *) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 244 | apply(rule converse_rtrancl_into_rtrancl) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 245 | apply(force intro!: JMPFT) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 246 | (* execute compile c0: *) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 247 | apply(rule rtrancl_trans) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 248 | apply(erule allE) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 249 | apply assumption | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 250 | (* execute JMPF: *) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 251 | apply(rule r_into_rtrancl) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 252 | apply(force intro!: JMPFF) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 253 | (* end of case b is true *) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 254 | apply(intro strip) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 255 | apply(erule_tac x = "a@[?I]@compile c0@[?J]" in allE) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 256 | apply(simp add:add_assoc) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 257 | apply(rule converse_rtrancl_into_rtrancl) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 258 | apply(force intro!: JMPFF) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 259 | apply(blast) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 260 | apply(force intro: JMPFF) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 261 | apply(intro strip) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 262 | apply(erule_tac x = "a@[?I]" in allE) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 263 | apply(erule_tac x = a in allE) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 264 | apply(simp) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 265 | apply(rule converse_rtrancl_into_rtrancl) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 266 | apply(force intro!: JMPFT) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 267 | apply(rule rtrancl_trans) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 268 | apply(erule allE) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 269 | apply assumption | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 270 | apply(rule converse_rtrancl_into_rtrancl) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 271 | apply(force intro!: JMPB) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 272 | apply(simp) | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 273 | done | 
| 13130 | 274 | *) | 
| 13095 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 275 | text {* Missing: the other direction! I did much of it, and although
 | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 276 | the main lemma is very similar to the one in the new development, the | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 277 | lemmas surrounding it seemed much more complicated. In the end I gave | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 278 | up. *} | 
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 279 | |
| 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 nipkow parents: diff
changeset | 280 | end |