| author | nipkow | 
| Mon, 18 Nov 2019 10:34:21 +0100 | |
| changeset 71139 | 87fd0374b3a0 | 
| parent 69597 | ff784d5a5bfb | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
1  | 
(* Author: Tobias Nipkow and Gerwin Klein *)  | 
| 43141 | 2  | 
|
| 58889 | 3  | 
section "Compiler for IMP"  | 
| 10343 | 4  | 
|
| 52915 | 5  | 
theory Compiler imports Big_Step Star  | 
| 43141 | 6  | 
begin  | 
| 12429 | 7  | 
|
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
8  | 
subsection "List setup"  | 
| 
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
9  | 
|
| 67406 | 10  | 
text \<open>  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
11  | 
In the following, we use the length of lists as integers  | 
| 69597 | 12  | 
instead of natural numbers. Instead of converting \<^typ>\<open>nat\<close>  | 
13  | 
to \<^typ>\<open>int\<close> explicitly, we tell Isabelle to coerce \<^typ>\<open>nat\<close>  | 
|
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
14  | 
automatically when necessary.  | 
| 67406 | 15  | 
\<close>  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
16  | 
declare [[coercion_enabled]]  | 
| 
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
17  | 
declare [[coercion "int :: nat \<Rightarrow> int"]]  | 
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
18  | 
|
| 67406 | 19  | 
text \<open>  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
20  | 
Similarly, we will want to access the ith element of a list,  | 
| 69597 | 21  | 
where \<^term>\<open>i\<close> is an \<^typ>\<open>int\<close>.  | 
| 67406 | 22  | 
\<close>  | 
| 45637 | 23  | 
fun inth :: "'a list \<Rightarrow> int \<Rightarrow> 'a" (infixl "!!" 100) where  | 
| 53869 | 24  | 
"(x # xs) !! i = (if i = 0 then x else xs !! (i - 1))"  | 
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
25  | 
|
| 67406 | 26  | 
text \<open>  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
27  | 
The only additional lemma we need about this function  | 
| 
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
28  | 
is indexing over append:  | 
| 67406 | 29  | 
\<close>  | 
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
30  | 
lemma inth_append [simp]:  | 
| 53869 | 31  | 
"0 \<le> i \<Longrightarrow>  | 
32  | 
(xs @ ys) !! i = (if i < size xs then xs !! i else ys !! (i - size xs))"  | 
|
33  | 
by (induction xs arbitrary: i) (auto simp: algebra_simps)  | 
|
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
34  | 
|
| 69597 | 35  | 
text\<open>We hide coercion \<^const>\<open>int\<close> applied to \<^const>\<open>length\<close>:\<close>  | 
| 53958 | 36  | 
|
37  | 
abbreviation (output)  | 
|
38  | 
"isize xs == int (length xs)"  | 
|
39  | 
||
40  | 
notation isize ("size")
 | 
|
41  | 
||
42  | 
||
| 43141 | 43  | 
subsection "Instructions and Stack Machine"  | 
| 10342 | 44  | 
|
| 67406 | 45  | 
text_raw\<open>\snip{instrdef}{0}{1}{%\<close>
 | 
| 58310 | 46  | 
datatype instr =  | 
| 52906 | 47  | 
LOADI int | LOAD vname | ADD | STORE vname |  | 
48  | 
JMP int | JMPLESS int | JMPGE int  | 
|
| 67406 | 49  | 
text_raw\<open>}%endsnip\<close>  | 
| 10342 | 50  | 
|
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
51  | 
type_synonym stack = "val list"  | 
| 45637 | 52  | 
type_synonym config = "int \<times> state \<times> stack"  | 
| 43141 | 53  | 
|
54  | 
abbreviation "hd2 xs == hd(tl xs)"  | 
|
55  | 
abbreviation "tl2 xs == tl(tl xs)"  | 
|
| 11275 | 56  | 
|
| 
50060
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
57  | 
fun iexec :: "instr \<Rightarrow> config \<Rightarrow> config" where  | 
| 
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
58  | 
"iexec instr (i,s,stk) = (case instr of  | 
| 
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
59  | 
LOADI n \<Rightarrow> (i+1,s, n#stk) |  | 
| 
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
60  | 
LOAD x \<Rightarrow> (i+1,s, s x # stk) |  | 
| 
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
61  | 
ADD \<Rightarrow> (i+1,s, (hd2 stk + hd stk) # tl2 stk) |  | 
| 
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
62  | 
STORE x \<Rightarrow> (i+1,s(x := hd stk),tl stk) |  | 
| 
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
63  | 
JMP n \<Rightarrow> (i+1+n,s,stk) |  | 
| 
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
64  | 
JMPLESS n \<Rightarrow> (if hd2 stk < hd stk then i+1+n else i+1,s,tl2 stk) |  | 
| 
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
65  | 
JMPGE n \<Rightarrow> (if hd2 stk >= hd stk then i+1+n else i+1,s,tl2 stk))"  | 
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
66  | 
|
| 
44000
 
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
 
kleing 
parents: 
43438 
diff
changeset
 | 
67  | 
definition  | 
| 
50060
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
68  | 
exec1 :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool"  | 
| 
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
69  | 
     ("(_/ \<turnstile> (_ \<rightarrow>/ _))" [59,0,59] 60) 
 | 
| 
44000
 
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
 
kleing 
parents: 
43438 
diff
changeset
 | 
70  | 
where  | 
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
71  | 
"P \<turnstile> c \<rightarrow> c' =  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
72  | 
(\<exists>i s stk. c = (i,s,stk) \<and> c' = iexec(P!!i) (i,s,stk) \<and> 0 \<le> i \<and> i < size P)"  | 
| 
44000
 
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
 
kleing 
parents: 
43438 
diff
changeset
 | 
73  | 
|
| 
 
ab4d8499815c
resolved code_pred FIXME in IMP; clearer notation for exec_n
 
kleing 
parents: 
43438 
diff
changeset
 | 
74  | 
lemma exec1I [intro, code_pred_intro]:  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
75  | 
"c' = iexec (P!!i) (i,s,stk) \<Longrightarrow> 0 \<le> i \<Longrightarrow> i < size P  | 
| 
50060
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
76  | 
\<Longrightarrow> P \<turnstile> (i,s,stk) \<rightarrow> c'"  | 
| 52915 | 77  | 
by (simp add: exec1_def)  | 
| 43141 | 78  | 
|
| 52915 | 79  | 
abbreviation  | 
80  | 
  exec :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool" ("(_/ \<turnstile> (_ \<rightarrow>*/ _))" 50)
 | 
|
| 43141 | 81  | 
where  | 
| 52915 | 82  | 
"exec P \<equiv> star (exec1 P)"  | 
83  | 
||
84  | 
lemmas exec_induct = star.induct [of "exec1 P", split_format(complete)]  | 
|
| 43141 | 85  | 
|
| 52915 | 86  | 
code_pred exec1 by (metis exec1_def)  | 
| 43141 | 87  | 
|
88  | 
values  | 
|
89  | 
  "{(i,map t [''x'',''y''],stk) | i t stk.
 | 
|
90  | 
[LOAD ''y'', STORE ''x''] \<turnstile>  | 
|
| 44036 | 91  | 
(0, <''x'' := 3, ''y'' := 4>, []) \<rightarrow>* (i,t,stk)}"  | 
| 43141 | 92  | 
|
93  | 
||
| 67406 | 94  | 
subsection\<open>Verification infrastructure\<close>  | 
| 43141 | 95  | 
|
| 67406 | 96  | 
text\<open>Below we need to argue about the execution of code that is embedded in  | 
| 43141 | 97  | 
larger programs. For this purpose we show that execution is preserved by  | 
| 67406 | 98  | 
appending code to the left or right of a program.\<close>  | 
| 43141 | 99  | 
|
| 
50060
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
100  | 
lemma iexec_shift [simp]:  | 
| 
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
101  | 
"((n+i',s',stk') = iexec x (n+i,s,stk)) = ((i',s',stk') = iexec x (i,s,stk))"  | 
| 
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
102  | 
by(auto split:instr.split)  | 
| 
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
103  | 
|
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
104  | 
lemma exec1_appendR: "P \<turnstile> c \<rightarrow> c' \<Longrightarrow> P@P' \<turnstile> c \<rightarrow> c'"  | 
| 52915 | 105  | 
by (auto simp: exec1_def)  | 
| 11275 | 106  | 
|
| 43141 | 107  | 
lemma exec_appendR: "P \<turnstile> c \<rightarrow>* c' \<Longrightarrow> P@P' \<turnstile> c \<rightarrow>* c'"  | 
| 61147 | 108  | 
by (induction rule: star.induct) (fastforce intro: star.step exec1_appendR)+  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
109  | 
|
| 43141 | 110  | 
lemma exec1_appendL:  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
111  | 
fixes i i' :: int  | 
| 
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
112  | 
shows  | 
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
113  | 
"P \<turnstile> (i,s,stk) \<rightarrow> (i',s',stk') \<Longrightarrow>  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
114  | 
P' @ P \<turnstile> (size(P')+i,s,stk) \<rightarrow> (size(P')+i',s',stk')"  | 
| 52915 | 115  | 
unfolding exec1_def  | 
| 52952 | 116  | 
by (auto simp del: iexec.simps)  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
117  | 
|
| 43141 | 118  | 
lemma exec_appendL:  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
119  | 
fixes i i' :: int  | 
| 
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
120  | 
shows  | 
| 43141 | 121  | 
"P \<turnstile> (i,s,stk) \<rightarrow>* (i',s',stk') \<Longrightarrow>  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
122  | 
P' @ P \<turnstile> (size(P')+i,s,stk) \<rightarrow>* (size(P')+i',s',stk')"  | 
| 61147 | 123  | 
by (induction rule: exec_induct) (blast intro: star.step exec1_appendL)+  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
124  | 
|
| 67406 | 125  | 
text\<open>Now we specialise the above lemmas to enable automatic proofs of  | 
| 69597 | 126  | 
\<^prop>\<open>P \<turnstile> c \<rightarrow>* c'\<close> where \<open>P\<close> is a mixture of concrete instructions and  | 
| 43141 | 127  | 
pieces of code that we already know how they execute (by induction), combined  | 
| 69505 | 128  | 
by \<open>@\<close> and \<open>#\<close>. Backward jumps are not supported.  | 
| 43141 | 129  | 
The details should be skipped on a first reading.  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
130  | 
|
| 67406 | 131  | 
If we have just executed the first instruction of the program, drop it:\<close>  | 
| 43141 | 132  | 
|
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
133  | 
lemma exec_Cons_1 [intro]:  | 
| 
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
134  | 
"P \<turnstile> (0,s,stk) \<rightarrow>* (j,t,stk') \<Longrightarrow>  | 
| 
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
135  | 
instr#P \<turnstile> (1,s,stk) \<rightarrow>* (1+j,t,stk')"  | 
| 
50060
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
136  | 
by (drule exec_appendL[where P'="[instr]"]) simp  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
137  | 
|
| 43141 | 138  | 
lemma exec_appendL_if[intro]:  | 
| 54428 | 139  | 
fixes i i' j :: int  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
140  | 
shows  | 
| 
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
141  | 
"size P' <= i  | 
| 54428 | 142  | 
\<Longrightarrow> P \<turnstile> (i - size P',s,stk) \<rightarrow>* (j,s',stk')  | 
143  | 
\<Longrightarrow> i' = size P' + j  | 
|
144  | 
\<Longrightarrow> P' @ P \<turnstile> (i,s,stk) \<rightarrow>* (i',s',stk')"  | 
|
| 
50060
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
145  | 
by (drule exec_appendL[where P'=P']) simp  | 
| 10342 | 146  | 
|
| 67406 | 147  | 
text\<open>Split the execution of a compound program up into the execution of its  | 
148  | 
parts:\<close>  | 
|
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
149  | 
|
| 43141 | 150  | 
lemma exec_append_trans[intro]:  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
151  | 
fixes i' i'' j'' :: int  | 
| 
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
152  | 
shows  | 
| 43141 | 153  | 
"P \<turnstile> (0,s,stk) \<rightarrow>* (i',s',stk') \<Longrightarrow>  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
154  | 
size P \<le> i' \<Longrightarrow>  | 
| 
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
155  | 
P' \<turnstile> (i' - size P,s',stk') \<rightarrow>* (i'',s'',stk'') \<Longrightarrow>  | 
| 
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
156  | 
j'' = size P + i''  | 
| 43141 | 157  | 
\<Longrightarrow>  | 
158  | 
P @ P' \<turnstile> (0,s,stk) \<rightarrow>* (j'',s'',stk'')"  | 
|
| 52915 | 159  | 
by(metis star_trans[OF exec_appendR exec_appendL_if])  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
160  | 
|
| 43141 | 161  | 
|
| 
50060
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
162  | 
declare Let_def[simp]  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
163  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
164  | 
|
| 43141 | 165  | 
subsection "Compilation"  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
166  | 
|
| 43141 | 167  | 
fun acomp :: "aexp \<Rightarrow> instr list" where  | 
168  | 
"acomp (N n) = [LOADI n]" |  | 
|
169  | 
"acomp (V x) = [LOAD x]" |  | 
|
170  | 
"acomp (Plus a1 a2) = acomp a1 @ acomp a2 @ [ADD]"  | 
|
171  | 
||
172  | 
lemma acomp_correct[intro]:  | 
|
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
173  | 
"acomp a \<turnstile> (0,s,stk) \<rightarrow>* (size(acomp a),s,aval a s#stk)"  | 
| 
50060
 
43753eca324a
replaced relation by function - simplifies development
 
nipkow 
parents: 
47818 
diff
changeset
 | 
174  | 
by (induction a arbitrary: stk) fastforce+  | 
| 10342 | 175  | 
|
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
176  | 
fun bcomp :: "bexp \<Rightarrow> bool \<Rightarrow> int \<Rightarrow> instr list" where  | 
| 53880 | 177  | 
"bcomp (Bc v) f n = (if v=f then [JMP n] else [])" |  | 
178  | 
"bcomp (Not b) f n = bcomp b (\<not>f) n" |  | 
|
179  | 
"bcomp (And b1 b2) f n =  | 
|
180  | 
(let cb2 = bcomp b2 f n;  | 
|
| 71139 | 181  | 
m = if f then size cb2 else (size cb2)+n;  | 
| 43141 | 182  | 
cb1 = bcomp b1 False m  | 
183  | 
in cb1 @ cb2)" |  | 
|
| 53880 | 184  | 
"bcomp (Less a1 a2) f n =  | 
185  | 
acomp a1 @ acomp a2 @ (if f then [JMPLESS n] else [JMPGE n])"  | 
|
| 43141 | 186  | 
|
187  | 
value  | 
|
188  | 
"bcomp (And (Less (V ''x'') (V ''y'')) (Not(Less (V ''u'') (V ''v''))))  | 
|
189  | 
False 3"  | 
|
190  | 
||
191  | 
lemma bcomp_correct[intro]:  | 
|
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
192  | 
fixes n :: int  | 
| 
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
193  | 
shows  | 
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
194  | 
"0 \<le> n \<Longrightarrow>  | 
| 53880 | 195  | 
bcomp b f n \<turnstile>  | 
196  | 
(0,s,stk) \<rightarrow>* (size(bcomp b f n) + (if f = bval b s then n else 0),s,stk)"  | 
|
197  | 
proof(induction b arbitrary: f n)  | 
|
| 43141 | 198  | 
case Not  | 
| 53880 | 199  | 
from Not(1)[where f="~f"] Not(2) show ?case by fastforce  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
200  | 
next  | 
| 43141 | 201  | 
case (And b1 b2)  | 
| 53880 | 202  | 
from And(1)[of "if f then size(bcomp b2 f n) else size(bcomp b2 f n) + n"  | 
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
203  | 
"False"]  | 
| 53880 | 204  | 
And(2)[of n f] And(3)  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44036 
diff
changeset
 | 
205  | 
show ?case by fastforce  | 
| 
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44036 
diff
changeset
 | 
206  | 
qed fastforce+  | 
| 43141 | 207  | 
|
208  | 
fun ccomp :: "com \<Rightarrow> instr list" where  | 
|
209  | 
"ccomp SKIP = []" |  | 
|
210  | 
"ccomp (x ::= a) = acomp a @ [STORE x]" |  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52952 
diff
changeset
 | 
211  | 
"ccomp (c\<^sub>1;;c\<^sub>2) = ccomp c\<^sub>1 @ ccomp c\<^sub>2" |  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52952 
diff
changeset
 | 
212  | 
"ccomp (IF b THEN c\<^sub>1 ELSE c\<^sub>2) =  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52952 
diff
changeset
 | 
213  | 
(let cc\<^sub>1 = ccomp c\<^sub>1; cc\<^sub>2 = ccomp c\<^sub>2; cb = bcomp b False (size cc\<^sub>1 + 1)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52952 
diff
changeset
 | 
214  | 
in cb @ cc\<^sub>1 @ JMP (size cc\<^sub>2) # cc\<^sub>2)" |  | 
| 43141 | 215  | 
"ccomp (WHILE b DO c) =  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
216  | 
(let cc = ccomp c; cb = bcomp b False (size cc + 1)  | 
| 
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
217  | 
in cb @ cc @ [JMP (-(size cb + size cc + 1))])"  | 
| 44004 | 218  | 
|
| 43141 | 219  | 
|
220  | 
value "ccomp  | 
|
221  | 
(IF Less (V ''u'') (N 1) THEN ''u'' ::= Plus (V ''u'') (N 1)  | 
|
222  | 
ELSE ''v'' ::= V ''u'')"  | 
|
223  | 
||
224  | 
value "ccomp (WHILE Less (V ''u'') (N 1) DO (''u'' ::= Plus (V ''u'') (N 1)))"
 | 
|
225  | 
||
226  | 
||
| 45114 | 227  | 
subsection "Preservation of semantics"  | 
| 43141 | 228  | 
|
| 
43438
 
a666b8d11252
IMP compiler with int, added reverse soundness direction
 
kleing 
parents: 
43158 
diff
changeset
 | 
229  | 
lemma ccomp_bigstep:  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
230  | 
"(c,s) \<Rightarrow> t \<Longrightarrow> ccomp c \<turnstile> (0,s,stk) \<rightarrow>* (size(ccomp c),t,stk)"  | 
| 45015 | 231  | 
proof(induction arbitrary: stk rule: big_step_induct)  | 
| 43141 | 232  | 
case (Assign x a s)  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44036 
diff
changeset
 | 
233  | 
show ?case by (fastforce simp:fun_upd_def cong: if_cong)  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
234  | 
next  | 
| 47818 | 235  | 
case (Seq c1 s1 s2 c2 s3)  | 
| 43141 | 236  | 
let ?cc1 = "ccomp c1" let ?cc2 = "ccomp c2"  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
237  | 
have "?cc1 @ ?cc2 \<turnstile> (0,s1,stk) \<rightarrow>* (size ?cc1,s2,stk)"  | 
| 47818 | 238  | 
using Seq.IH(1) by fastforce  | 
| 43141 | 239  | 
moreover  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
240  | 
have "?cc1 @ ?cc2 \<turnstile> (size ?cc1,s2,stk) \<rightarrow>* (size(?cc1 @ ?cc2),s3,stk)"  | 
| 47818 | 241  | 
using Seq.IH(2) by fastforce  | 
| 52915 | 242  | 
ultimately show ?case by simp (blast intro: star_trans)  | 
| 43141 | 243  | 
next  | 
244  | 
case (WhileTrue b s1 c s2 s3)  | 
|
245  | 
let ?cc = "ccomp c"  | 
|
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
246  | 
let ?cb = "bcomp b False (size ?cc + 1)"  | 
| 43141 | 247  | 
let ?cw = "ccomp(WHILE b DO c)"  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
248  | 
have "?cw \<turnstile> (0,s1,stk) \<rightarrow>* (size ?cb,s1,stk)"  | 
| 67406 | 249  | 
using \<open>bval b s1\<close> by fastforce  | 
| 50133 | 250  | 
moreover  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
251  | 
have "?cw \<turnstile> (size ?cb,s1,stk) \<rightarrow>* (size ?cb + size ?cc,s2,stk)"  | 
| 50133 | 252  | 
using WhileTrue.IH(1) by fastforce  | 
| 43141 | 253  | 
moreover  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
254  | 
have "?cw \<turnstile> (size ?cb + size ?cc,s2,stk) \<rightarrow>* (0,s2,stk)"  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44036 
diff
changeset
 | 
255  | 
by fastforce  | 
| 43141 | 256  | 
moreover  | 
| 
51259
 
1491459df114
eliminated isize in favour of size + type coercion
 
kleing 
parents: 
50133 
diff
changeset
 | 
257  | 
have "?cw \<turnstile> (0,s2,stk) \<rightarrow>* (size ?cw,s3,stk)" by(rule WhileTrue.IH(2))  | 
| 52915 | 258  | 
ultimately show ?case by(blast intro: star_trans)  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44036 
diff
changeset
 | 
259  | 
qed fastforce+  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
260  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
16417 
diff
changeset
 | 
261  | 
end  |