src/HOL/IMP/Compiler.thy
author kleing
Thu, 28 Jul 2011 15:15:26 +0200
changeset 44871 ab4d8499815c
parent 44296 a666b8d11252
child 44875 a9fcbafdf208
permissions -rw-r--r--
resolved code_pred FIXME in IMP; clearer notation for exec_n
nipkow@43982
     1
(* Author: Tobias Nipkow *)
nipkow@10343
     2
nipkow@43982
     3
header "A Compiler for IMP"
nipkow@10342
     4
kleing@44296
     5
theory Compiler imports Big_Step 
nipkow@43982
     6
begin
kleing@12429
     7
kleing@44296
     8
subsection "List setup"
kleing@44296
     9
kleing@44296
    10
text {*
kleing@44296
    11
  We are going to define a small machine language where programs are
kleing@44296
    12
  lists of instructions. For nicer algebraic properties in our lemmas
kleing@44296
    13
  later, we prefer @{typ int} to @{term nat} as program counter.
kleing@44296
    14
  
kleing@44296
    15
  Therefore, we define notation for size and indexing for lists 
kleing@44296
    16
  on @{typ int}:
kleing@44296
    17
*}
kleing@44296
    18
abbreviation "isize xs == int (length xs)" 
kleing@44296
    19
kleing@44296
    20
primrec
kleing@44296
    21
  inth :: "'a list => int => 'a" (infixl "!!" 100) where
kleing@44296
    22
  inth_Cons: "(x # xs) !! n = (if n = 0 then x else xs !! (n - 1))"
kleing@44296
    23
kleing@44296
    24
text {*
kleing@44296
    25
  The only additional lemma we need is indexing over append:
kleing@44296
    26
*}
kleing@44296
    27
lemma inth_append [simp]:
kleing@44296
    28
  "0 \<le> n \<Longrightarrow>
kleing@44296
    29
  (xs @ ys) !! n = (if n < isize xs then xs !! n else ys !! (n - isize xs))"
kleing@44296
    30
  by (induct xs arbitrary: n) (auto simp: algebra_simps)
kleing@44296
    31
nipkow@43982
    32
subsection "Instructions and Stack Machine"
nipkow@43982
    33
nipkow@43982
    34
datatype instr = 
kleing@44296
    35
  LOADI int | 
kleing@44296
    36
  LOAD string | 
kleing@44296
    37
  ADD |
nipkow@43982
    38
  STORE string |
kleing@44296
    39
  JMP int |
kleing@44296
    40
  JMPFLESS int |
kleing@44296
    41
  JMPFGE int
nipkow@43982
    42
kleing@44296
    43
(* reads slightly nicer *)
kleing@44296
    44
abbreviation
kleing@44296
    45
  "JMPB i == JMP (-i)"
kleing@44296
    46
kleing@44296
    47
type_synonym stack = "val list"
kleing@44296
    48
type_synonym config = "int\<times>state\<times>stack"
nipkow@43982
    49
nipkow@43982
    50
abbreviation "hd2 xs == hd(tl xs)"
nipkow@43982
    51
abbreviation "tl2 xs == tl(tl xs)"
nipkow@43982
    52
kleing@44296
    53
inductive iexec1 :: "instr \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool"
kleing@44296
    54
    ("(_/ \<turnstile>i (_ \<rightarrow>/ _))" [50,0,0] 50)
wenzelm@27362
    55
where
kleing@44296
    56
"LOADI n \<turnstile>i (i,s,stk) \<rightarrow> (i+1,s, n#stk)" |
kleing@44296
    57
"LOAD x  \<turnstile>i (i,s,stk) \<rightarrow> (i+1,s, s x # stk)" |
kleing@44296
    58
"ADD     \<turnstile>i (i,s,stk) \<rightarrow> (i+1,s, (hd2 stk + hd stk) # tl2 stk)" |
kleing@44296
    59
"STORE n \<turnstile>i (i,s,stk) \<rightarrow> (i+1,s(n := hd stk),tl stk)" |
kleing@44296
    60
"JMP n   \<turnstile>i (i,s,stk) \<rightarrow> (i+1+n,s,stk)" |
kleing@44296
    61
"JMPFLESS n \<turnstile>i (i,s,stk) \<rightarrow> (if hd2 stk < hd stk then i+1+n else i+1,s,tl2 stk)" |
kleing@44296
    62
"JMPFGE n \<turnstile>i (i,s,stk) \<rightarrow> (if hd2 stk >= hd stk then i+1+n else i+1,s,tl2 stk)"
kleing@44296
    63
kleing@44296
    64
code_pred iexec1 .
kleing@44296
    65
kleing@44296
    66
declare iexec1.intros
kleing@44296
    67
kleing@44871
    68
definition
kleing@44871
    69
  exec1 :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool"  ("(_/ \<turnstile> (_ \<rightarrow>/ _))" [50,0,0] 50) 
kleing@44871
    70
where
kleing@44296
    71
  "P \<turnstile> c \<rightarrow> c' = 
kleing@44296
    72
   (\<exists>i s stk. c = (i,s,stk) \<and> P!!i \<turnstile>i (i,s,stk) \<rightarrow> c' \<and> 0 \<le> i \<and> i < isize P)"
kleing@44871
    73
kleing@44871
    74
declare exec1_def [simp] 
kleing@44871
    75
kleing@44871
    76
lemma exec1I [intro, code_pred_intro]:
kleing@44871
    77
  "\<lbrakk> P!!i \<turnstile>i (i,s,stk) \<rightarrow> c'; 0 \<le> i; i < isize P \<rbrakk> \<Longrightarrow> P \<turnstile> (i,s,stk) \<rightarrow> c'"
kleing@44871
    78
  by simp
nipkow@43982
    79
nipkow@43982
    80
inductive exec :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool" ("_/ \<turnstile> (_ \<rightarrow>*/ _)" 50)
nipkow@43982
    81
where
nipkow@43982
    82
refl: "P \<turnstile> c \<rightarrow>* c" |
nipkow@43982
    83
step: "P \<turnstile> c \<rightarrow> c' \<Longrightarrow> P \<turnstile> c' \<rightarrow>* c'' \<Longrightarrow> P \<turnstile> c \<rightarrow>* c''"
nipkow@43982
    84
kleing@44296
    85
declare refl[intro] step[intro]
nipkow@43982
    86
nipkow@43982
    87
lemmas exec_induct = exec.induct[split_format(complete)]
nipkow@43982
    88
kleing@44871
    89
code_pred exec by force
nipkow@43982
    90
nipkow@43982
    91
values
nipkow@43982
    92
  "{(i,map t [''x'',''y''],stk) | i t stk.
nipkow@43982
    93
    [LOAD ''y'', STORE ''x''] \<turnstile>
kleing@43999
    94
    (0, [''x'' \<rightarrow> 3, ''y'' \<rightarrow> 4], []) \<rightarrow>* (i,t,stk)}"
nipkow@43982
    95
nipkow@43982
    96
nipkow@43982
    97
subsection{* Verification infrastructure *}
nipkow@43982
    98
nipkow@43982
    99
lemma exec_trans: "P \<turnstile> c \<rightarrow>* c' \<Longrightarrow> P \<turnstile> c' \<rightarrow>* c'' \<Longrightarrow> P \<turnstile> c \<rightarrow>* c''"
kleing@44296
   100
  by (induct rule: exec.induct) fastsimp+
nipkow@43982
   101
kleing@44296
   102
inductive_cases iexec1_cases [elim!]:
kleing@44296
   103
  "LOADI n \<turnstile>i c \<rightarrow> c'" 
kleing@44296
   104
  "LOAD x  \<turnstile>i c \<rightarrow> c'"
kleing@44296
   105
  "ADD     \<turnstile>i c \<rightarrow> c'"
kleing@44296
   106
  "STORE n \<turnstile>i c \<rightarrow> c'" 
kleing@44296
   107
  "JMP n   \<turnstile>i c \<rightarrow> c'"
kleing@44296
   108
  "JMPFLESS n \<turnstile>i c \<rightarrow> c'"
kleing@44296
   109
  "JMPFGE n \<turnstile>i c \<rightarrow> c'"
nipkow@43982
   110
kleing@44296
   111
text {* Simplification rules for @{const iexec1}. *}
kleing@44296
   112
lemma iexec1_simps [simp]:
kleing@44296
   113
  "LOADI n \<turnstile>i c \<rightarrow> c' = (\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1, s, n # stk))"
kleing@44296
   114
  "LOAD x \<turnstile>i c \<rightarrow> c' = (\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1, s, s x # stk))"
kleing@44296
   115
  "ADD \<turnstile>i c \<rightarrow> c' = 
kleing@44296
   116
  (\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1, s, (hd2 stk + hd stk) # tl2 stk))"
kleing@44296
   117
  "STORE x \<turnstile>i c \<rightarrow> c' = 
kleing@44296
   118
  (\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1, s(x \<rightarrow> hd stk), tl stk))"
kleing@44296
   119
  "JMP n \<turnstile>i c \<rightarrow> c' = (\<exists>i s stk. c = (i, s, stk) \<and> c' = (i + 1 + n, s, stk))"
kleing@44296
   120
   "JMPFLESS n \<turnstile>i c \<rightarrow> c' = 
kleing@44296
   121
  (\<exists>i s stk. c = (i, s, stk) \<and> 
kleing@44296
   122
             c' = (if hd2 stk < hd stk then i + 1 + n else i + 1, s, tl2 stk))"  
kleing@44296
   123
  "JMPFGE n \<turnstile>i c \<rightarrow> c' = 
kleing@44296
   124
  (\<exists>i s stk. c = (i, s, stk) \<and> 
kleing@44296
   125
             c' = (if hd stk \<le> hd2 stk then i + 1 + n else i + 1, s, tl2 stk))"
kleing@44296
   126
  by (auto split del: split_if intro!: iexec1.intros)
kleing@44296
   127
nipkow@43982
   128
nipkow@43982
   129
text{* Below we need to argue about the execution of code that is embedded in
nipkow@43982
   130
larger programs. For this purpose we show that execution is preserved by
nipkow@43982
   131
appending code to the left or right of a program. *}
nipkow@43982
   132
kleing@44296
   133
lemma exec1_appendR: "P \<turnstile> c \<rightarrow> c' \<Longrightarrow> P@P' \<turnstile> c \<rightarrow> c'"
kleing@44296
   134
  by auto
nipkow@11275
   135
nipkow@43982
   136
lemma exec_appendR: "P \<turnstile> c \<rightarrow>* c' \<Longrightarrow> P@P' \<turnstile> c \<rightarrow>* c'"
kleing@44296
   137
  by (induct rule: exec.induct) (fastsimp intro: exec1_appendR)+
nipkow@13095
   138
kleing@44296
   139
lemma iexec1_shiftI:
kleing@44296
   140
  assumes "X \<turnstile>i (i,s,stk) \<rightarrow> (i',s',stk')"
kleing@44296
   141
  shows "X \<turnstile>i (n+i,s,stk) \<rightarrow> (n+i',s',stk')"
kleing@44296
   142
  using assms by cases auto
kleing@44296
   143
kleing@44296
   144
lemma iexec1_shiftD:
kleing@44296
   145
  assumes "X \<turnstile>i (n+i,s,stk) \<rightarrow> (n+i',s',stk')"
kleing@44296
   146
  shows "X \<turnstile>i (i,s,stk) \<rightarrow> (i',s',stk')"
kleing@44296
   147
  using assms by cases auto
kleing@44296
   148
kleing@44296
   149
lemma iexec_shift [simp]: 
kleing@44296
   150
  "(X \<turnstile>i (n+i,s,stk) \<rightarrow> (n+i',s',stk')) = (X \<turnstile>i (i,s,stk) \<rightarrow> (i',s',stk'))"
kleing@44296
   151
  by (blast intro: iexec1_shiftI dest: iexec1_shiftD)
kleing@44296
   152
  
nipkow@43982
   153
lemma exec1_appendL:
kleing@44296
   154
  "P \<turnstile> (i,s,stk) \<rightarrow> (i',s',stk') \<Longrightarrow>
kleing@44296
   155
   P' @ P \<turnstile> (isize(P')+i,s,stk) \<rightarrow> (isize(P')+i',s',stk')"
kleing@44296
   156
  by simp
nipkow@13095
   157
nipkow@43982
   158
lemma exec_appendL:
nipkow@43982
   159
 "P \<turnstile> (i,s,stk) \<rightarrow>* (i',s',stk')  \<Longrightarrow>
kleing@44296
   160
  P' @ P \<turnstile> (isize(P')+i,s,stk) \<rightarrow>* (isize(P')+i',s',stk')"
kleing@44296
   161
  by (induct rule: exec_induct) (blast intro!: exec1_appendL)+
nipkow@43982
   162
nipkow@43982
   163
text{* Now we specialise the above lemmas to enable automatic proofs of
nipkow@43982
   164
@{prop "P \<turnstile> c \<rightarrow>* c'"} where @{text P} is a mixture of concrete instructions and
nipkow@43982
   165
pieces of code that we already know how they execute (by induction), combined
nipkow@43982
   166
by @{text "@"} and @{text "#"}. Backward jumps are not supported.
nipkow@43982
   167
The details should be skipped on a first reading.
nipkow@43982
   168
kleing@44296
   169
If we have just executed the first instruction of the program, drop it: *}
nipkow@43982
   170
kleing@44296
   171
lemma exec_Cons_1 [intro]:
kleing@44296
   172
  "P \<turnstile> (0,s,stk) \<rightarrow>* (j,t,stk') \<Longrightarrow>
kleing@44296
   173
  instr#P \<turnstile> (1,s,stk) \<rightarrow>* (1+j,t,stk')"
kleing@44296
   174
  by (drule exec_appendL[where P'="[instr]"]) simp
nipkow@10342
   175
nipkow@43982
   176
lemma exec_appendL_if[intro]:
kleing@44296
   177
 "isize P' <= i
kleing@44296
   178
  \<Longrightarrow> P \<turnstile> (i - isize P',s,stk) \<rightarrow>* (i',s',stk')
kleing@44296
   179
  \<Longrightarrow> P' @ P \<turnstile> (i,s,stk) \<rightarrow>* (isize P' + i',s',stk')"
kleing@44296
   180
  by (drule exec_appendL[where P'=P']) simp
nipkow@13095
   181
nipkow@43982
   182
text{* Split the execution of a compound program up into the excution of its
nipkow@43982
   183
parts: *}
nipkow@13095
   184
nipkow@43982
   185
lemma exec_append_trans[intro]:
nipkow@43982
   186
"P \<turnstile> (0,s,stk) \<rightarrow>* (i',s',stk') \<Longrightarrow>
kleing@44296
   187
 isize P \<le> i' \<Longrightarrow>
kleing@44296
   188
 P' \<turnstile>  (i' - isize P,s',stk') \<rightarrow>* (i'',s'',stk'') \<Longrightarrow>
kleing@44296
   189
 j'' = isize P + i''
nipkow@43982
   190
 \<Longrightarrow>
nipkow@43982
   191
 P @ P' \<turnstile> (0,s,stk) \<rightarrow>* (j'',s'',stk'')"
kleing@44296
   192
  by(metis exec_trans[OF exec_appendR exec_appendL_if])
nipkow@13095
   193
nipkow@43982
   194
kleing@44296
   195
declare Let_def[simp] 
nipkow@43982
   196
nipkow@43982
   197
nipkow@43982
   198
subsection "Compilation"
nipkow@43982
   199
nipkow@43982
   200
fun acomp :: "aexp \<Rightarrow> instr list" where
nipkow@43982
   201
"acomp (N n) = [LOADI n]" |
nipkow@43982
   202
"acomp (V x) = [LOAD x]" |
nipkow@43982
   203
"acomp (Plus a1 a2) = acomp a1 @ acomp a2 @ [ADD]"
nipkow@43982
   204
nipkow@43982
   205
lemma acomp_correct[intro]:
kleing@44296
   206
  "acomp a \<turnstile> (0,s,stk) \<rightarrow>* (isize(acomp a),s,aval a s#stk)"
kleing@44296
   207
  by (induct a arbitrary: stk) fastsimp+
nipkow@13095
   208
kleing@44296
   209
fun bcomp :: "bexp \<Rightarrow> bool \<Rightarrow> int \<Rightarrow> instr list" where
kleing@44296
   210
"bcomp (B bv) c n = (if bv=c then [JMP n] else [])" |
nipkow@43982
   211
"bcomp (Not b) c n = bcomp b (\<not>c) n" |
nipkow@43982
   212
"bcomp (And b1 b2) c n =
nipkow@43982
   213
 (let cb2 = bcomp b2 c n;
kleing@44296
   214
        m = (if c then isize cb2 else isize cb2+n);
nipkow@43982
   215
      cb1 = bcomp b1 False m
nipkow@43982
   216
  in cb1 @ cb2)" |
nipkow@43982
   217
"bcomp (Less a1 a2) c n =
nipkow@43982
   218
 acomp a1 @ acomp a2 @ (if c then [JMPFLESS n] else [JMPFGE n])"
nipkow@43982
   219
nipkow@43982
   220
value
nipkow@43982
   221
  "bcomp (And (Less (V ''x'') (V ''y'')) (Not(Less (V ''u'') (V ''v''))))
nipkow@43982
   222
     False 3"
nipkow@43982
   223
nipkow@43982
   224
lemma bcomp_correct[intro]:
kleing@44296
   225
  "0 \<le> n \<Longrightarrow>
kleing@44296
   226
  bcomp b c n \<turnstile>
kleing@44296
   227
 (0,s,stk)  \<rightarrow>*  (isize(bcomp b c n) + (if c = bval b s then n else 0),s,stk)"
nipkow@43982
   228
proof(induct b arbitrary: c n m)
nipkow@43982
   229
  case Not
kleing@44296
   230
  from Not(1)[where c="~c"] Not(2) show ?case by fastsimp
nipkow@13095
   231
next
nipkow@43982
   232
  case (And b1 b2)
kleing@44296
   233
  from And(1)[of "if c then isize (bcomp b2 c n) else isize (bcomp b2 c n) + n" 
kleing@44296
   234
                 "False"] 
kleing@44296
   235
       And(2)[of n  "c"] And(3) 
kleing@44296
   236
  show ?case by fastsimp
nipkow@43982
   237
qed fastsimp+
nipkow@13095
   238
nipkow@43982
   239
fun ccomp :: "com \<Rightarrow> instr list" where
nipkow@43982
   240
"ccomp SKIP = []" |
nipkow@43982
   241
"ccomp (x ::= a) = acomp a @ [STORE x]" |
nipkow@43982
   242
"ccomp (c\<^isub>1;c\<^isub>2) = ccomp c\<^isub>1 @ ccomp c\<^isub>2" |
nipkow@43982
   243
"ccomp (IF b THEN c\<^isub>1 ELSE c\<^isub>2) =
kleing@44296
   244
  (let cc\<^isub>1 = ccomp c\<^isub>1; cc\<^isub>2 = ccomp c\<^isub>2; cb = bcomp b False (isize cc\<^isub>1 + 1)
kleing@44296
   245
   in cb @ cc\<^isub>1 @ JMP (isize cc\<^isub>2) # cc\<^isub>2)" |
nipkow@43982
   246
"ccomp (WHILE b DO c) =
kleing@44296
   247
 (let cc = ccomp c; cb = bcomp b False (isize cc + 1)
kleing@44296
   248
  in cb @ cc @ [JMPB (isize cb + isize cc + 1)])"
nipkow@13095
   249
nipkow@43982
   250
value "ccomp
nipkow@43982
   251
 (IF Less (V ''u'') (N 1) THEN ''u'' ::= Plus (V ''u'') (N 1)
nipkow@43982
   252
  ELSE ''v'' ::= V ''u'')"
nipkow@13095
   253
nipkow@43982
   254
value "ccomp (WHILE Less (V ''u'') (N 1) DO (''u'' ::= Plus (V ''u'') (N 1)))"
nipkow@13095
   255
nipkow@13095
   256
nipkow@43982
   257
subsection "Preservation of sematics"
nipkow@43982
   258
kleing@44296
   259
lemma ccomp_bigstep:
kleing@44296
   260
  "(c,s) \<Rightarrow> t \<Longrightarrow> ccomp c \<turnstile> (0,s,stk) \<rightarrow>* (isize(ccomp c),t,stk)"
nipkow@43982
   261
proof(induct arbitrary: stk rule: big_step_induct)
nipkow@43982
   262
  case (Assign x a s)
kleing@44296
   263
  show ?case by (fastsimp simp:fun_upd_def cong: if_cong)
nipkow@13095
   264
next
nipkow@43982
   265
  case (Semi c1 s1 s2 c2 s3)
nipkow@43982
   266
  let ?cc1 = "ccomp c1"  let ?cc2 = "ccomp c2"
kleing@44296
   267
  have "?cc1 @ ?cc2 \<turnstile> (0,s1,stk) \<rightarrow>* (isize ?cc1,s2,stk)"
kleing@44296
   268
    using Semi.hyps(2) by fastsimp
nipkow@43982
   269
  moreover
kleing@44296
   270
  have "?cc1 @ ?cc2 \<turnstile> (isize ?cc1,s2,stk) \<rightarrow>* (isize(?cc1 @ ?cc2),s3,stk)"
kleing@44296
   271
    using Semi.hyps(4) by fastsimp
nipkow@43982
   272
  ultimately show ?case by simp (blast intro: exec_trans)
nipkow@13095
   273
next
nipkow@43982
   274
  case (WhileTrue b s1 c s2 s3)
nipkow@43982
   275
  let ?cc = "ccomp c"
kleing@44296
   276
  let ?cb = "bcomp b False (isize ?cc + 1)"
nipkow@43982
   277
  let ?cw = "ccomp(WHILE b DO c)"
kleing@44296
   278
  have "?cw \<turnstile> (0,s1,stk) \<rightarrow>* (isize ?cb + isize ?cc,s2,stk)"
nipkow@43982
   279
    using WhileTrue(1,3) by fastsimp
nipkow@43982
   280
  moreover
kleing@44296
   281
  have "?cw \<turnstile> (isize ?cb + isize ?cc,s2,stk) \<rightarrow>* (0,s2,stk)"
kleing@44296
   282
    by fastsimp
nipkow@43982
   283
  moreover
kleing@44296
   284
  have "?cw \<turnstile> (0,s2,stk) \<rightarrow>* (isize ?cw,s3,stk)" by(rule WhileTrue(5))
nipkow@43982
   285
  ultimately show ?case by(blast intro: exec_trans)
nipkow@43982
   286
qed fastsimp+
nipkow@13095
   287
webertj@20217
   288
end