src/HOL/IMP/Compiler.thy
author nipkow
Wed, 14 Nov 2012 14:11:47 +0100
changeset 51075 43753eca324a
parent 48689 151d137f1095
child 51076 7110422d4cb3
permissions -rw-r--r--
replaced relation by function - simplifies development
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
nipkow@46501
    20
fun inth :: "'a list \<Rightarrow> int \<Rightarrow> 'a" (infixl "!!" 100) where
nipkow@46501
    21
"(x # xs) !! n = (if n = 0 then x else xs !! (n - 1))"
kleing@44296
    22
kleing@44296
    23
text {*
kleing@44296
    24
  The only additional lemma we need is indexing over append:
kleing@44296
    25
*}
kleing@44296
    26
lemma inth_append [simp]:
kleing@44296
    27
  "0 \<le> n \<Longrightarrow>
kleing@44296
    28
  (xs @ ys) !! n = (if n < isize xs then xs !! n else ys !! (n - isize xs))"
nipkow@51075
    29
by (induction xs arbitrary: n) (auto simp: algebra_simps)
kleing@44296
    30
nipkow@43982
    31
subsection "Instructions and Stack Machine"
nipkow@43982
    32
nipkow@43982
    33
datatype instr = 
kleing@46194
    34
  LOADI int |
kleing@46194
    35
  LOAD vname |
kleing@44296
    36
  ADD |
kleing@46194
    37
  STORE vname |
kleing@44296
    38
  JMP int |
kleing@46193
    39
  JMPLESS int |
kleing@46193
    40
  JMPGE int
nipkow@43982
    41
kleing@44296
    42
type_synonym stack = "val list"
nipkow@46501
    43
type_synonym config = "int \<times> state \<times> stack"
nipkow@43982
    44
nipkow@43982
    45
abbreviation "hd2 xs == hd(tl xs)"
nipkow@43982
    46
abbreviation "tl2 xs == tl(tl xs)"
nipkow@43982
    47
nipkow@51075
    48
fun iexec :: "instr \<Rightarrow> config \<Rightarrow> config" where
nipkow@51075
    49
"iexec instr (i,s,stk) = (case instr of
nipkow@51075
    50
  LOADI n \<Rightarrow> (i+1,s, n#stk) |
nipkow@51075
    51
  LOAD x \<Rightarrow> (i+1,s, s x # stk) |
nipkow@51075
    52
  ADD \<Rightarrow> (i+1,s, (hd2 stk + hd stk) # tl2 stk) |
nipkow@51075
    53
  STORE x \<Rightarrow> (i+1,s(x := hd stk),tl stk) |
nipkow@51075
    54
  JMP n \<Rightarrow>  (i+1+n,s,stk) |
nipkow@51075
    55
  JMPLESS n \<Rightarrow> (if hd2 stk < hd stk then i+1+n else i+1,s,tl2 stk) |
nipkow@51075
    56
  JMPGE n \<Rightarrow> (if hd2 stk >= hd stk then i+1+n else i+1,s,tl2 stk))"
kleing@44296
    57
kleing@44871
    58
definition
nipkow@51075
    59
  exec1 :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool"
nipkow@51075
    60
     ("(_/ \<turnstile> (_ \<rightarrow>/ _))" [59,0,59] 60) 
kleing@44871
    61
where
kleing@44296
    62
  "P \<turnstile> c \<rightarrow> c' = 
nipkow@51075
    63
  (\<exists>i s stk. c = (i,s,stk) \<and> c' = iexec(P!!i) (i,s,stk) \<and> 0 \<le> i \<and> i < isize P)"
kleing@44871
    64
nipkow@51075
    65
declare exec1_def [simp]
kleing@44871
    66
kleing@44871
    67
lemma exec1I [intro, code_pred_intro]:
nipkow@51075
    68
  "c' = iexec (P!!i) (i,s,stk) \<Longrightarrow> 0 \<le> i \<Longrightarrow> i < isize P
nipkow@51075
    69
  \<Longrightarrow> P \<turnstile> (i,s,stk) \<rightarrow> c'"
nipkow@51075
    70
by simp
nipkow@43982
    71
nipkow@51075
    72
inductive exec :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool"
nipkow@51075
    73
   ("(_/ \<turnstile> (_ \<rightarrow>*/ _))" 50)
nipkow@43982
    74
where
nipkow@43982
    75
refl: "P \<turnstile> c \<rightarrow>* c" |
nipkow@43982
    76
step: "P \<turnstile> c \<rightarrow> c' \<Longrightarrow> P \<turnstile> c' \<rightarrow>* c'' \<Longrightarrow> P \<turnstile> c \<rightarrow>* c''"
nipkow@43982
    77
kleing@44296
    78
declare refl[intro] step[intro]
nipkow@43982
    79
nipkow@43982
    80
lemmas exec_induct = exec.induct[split_format(complete)]
nipkow@43982
    81
kleing@44871
    82
code_pred exec by force
nipkow@43982
    83
nipkow@43982
    84
values
nipkow@43982
    85
  "{(i,map t [''x'',''y''],stk) | i t stk.
nipkow@43982
    86
    [LOAD ''y'', STORE ''x''] \<turnstile>
kleing@44907
    87
    (0, <''x'' := 3, ''y'' := 4>, []) \<rightarrow>* (i,t,stk)}"
nipkow@43982
    88
nipkow@43982
    89
nipkow@43982
    90
subsection{* Verification infrastructure *}
nipkow@43982
    91
nipkow@43982
    92
lemma exec_trans: "P \<turnstile> c \<rightarrow>* c' \<Longrightarrow> P \<turnstile> c' \<rightarrow>* c'' \<Longrightarrow> P \<turnstile> c \<rightarrow>* c''"
nipkow@51075
    93
by (induction rule: exec.induct) fastforce+
nipkow@43982
    94
nipkow@43982
    95
text{* Below we need to argue about the execution of code that is embedded in
nipkow@43982
    96
larger programs. For this purpose we show that execution is preserved by
nipkow@43982
    97
appending code to the left or right of a program. *}
nipkow@43982
    98
nipkow@51075
    99
lemma iexec_shift [simp]: 
nipkow@51075
   100
  "((n+i',s',stk') = iexec x (n+i,s,stk)) = ((i',s',stk') = iexec x (i,s,stk))"
nipkow@51075
   101
by(auto split:instr.split)
nipkow@51075
   102
kleing@44296
   103
lemma exec1_appendR: "P \<turnstile> c \<rightarrow> c' \<Longrightarrow> P@P' \<turnstile> c \<rightarrow> c'"
nipkow@51075
   104
by auto
nipkow@11275
   105
nipkow@43982
   106
lemma exec_appendR: "P \<turnstile> c \<rightarrow>* c' \<Longrightarrow> P@P' \<turnstile> c \<rightarrow>* c'"
nipkow@51075
   107
by (induction rule: exec.induct) (fastforce intro: exec1_appendR)+
nipkow@13095
   108
nipkow@43982
   109
lemma exec1_appendL:
kleing@44296
   110
  "P \<turnstile> (i,s,stk) \<rightarrow> (i',s',stk') \<Longrightarrow>
kleing@44296
   111
   P' @ P \<turnstile> (isize(P')+i,s,stk) \<rightarrow> (isize(P')+i',s',stk')"
nipkow@51075
   112
by (auto split: instr.split)
nipkow@13095
   113
nipkow@43982
   114
lemma exec_appendL:
nipkow@43982
   115
 "P \<turnstile> (i,s,stk) \<rightarrow>* (i',s',stk')  \<Longrightarrow>
kleing@44296
   116
  P' @ P \<turnstile> (isize(P')+i,s,stk) \<rightarrow>* (isize(P')+i',s',stk')"
nipkow@51075
   117
by (induction rule: exec_induct) (blast intro!: exec1_appendL)+
nipkow@43982
   118
nipkow@43982
   119
text{* Now we specialise the above lemmas to enable automatic proofs of
nipkow@43982
   120
@{prop "P \<turnstile> c \<rightarrow>* c'"} where @{text P} is a mixture of concrete instructions and
nipkow@43982
   121
pieces of code that we already know how they execute (by induction), combined
nipkow@43982
   122
by @{text "@"} and @{text "#"}. Backward jumps are not supported.
nipkow@43982
   123
The details should be skipped on a first reading.
nipkow@43982
   124
kleing@44296
   125
If we have just executed the first instruction of the program, drop it: *}
nipkow@43982
   126
kleing@44296
   127
lemma exec_Cons_1 [intro]:
kleing@44296
   128
  "P \<turnstile> (0,s,stk) \<rightarrow>* (j,t,stk') \<Longrightarrow>
kleing@44296
   129
  instr#P \<turnstile> (1,s,stk) \<rightarrow>* (1+j,t,stk')"
nipkow@51075
   130
by (drule exec_appendL[where P'="[instr]"]) simp
nipkow@10342
   131
nipkow@43982
   132
lemma exec_appendL_if[intro]:
kleing@44296
   133
 "isize P' <= i
kleing@44296
   134
  \<Longrightarrow> P \<turnstile> (i - isize P',s,stk) \<rightarrow>* (i',s',stk')
kleing@44296
   135
  \<Longrightarrow> P' @ P \<turnstile> (i,s,stk) \<rightarrow>* (isize P' + i',s',stk')"
nipkow@51075
   136
by (drule exec_appendL[where P'=P']) simp
nipkow@13095
   137
nipkow@43982
   138
text{* Split the execution of a compound program up into the excution of its
nipkow@43982
   139
parts: *}
nipkow@13095
   140
nipkow@43982
   141
lemma exec_append_trans[intro]:
nipkow@43982
   142
"P \<turnstile> (0,s,stk) \<rightarrow>* (i',s',stk') \<Longrightarrow>
kleing@44296
   143
 isize P \<le> i' \<Longrightarrow>
kleing@44296
   144
 P' \<turnstile>  (i' - isize P,s',stk') \<rightarrow>* (i'',s'',stk'') \<Longrightarrow>
kleing@44296
   145
 j'' = isize P + i''
nipkow@43982
   146
 \<Longrightarrow>
nipkow@43982
   147
 P @ P' \<turnstile> (0,s,stk) \<rightarrow>* (j'',s'',stk'')"
nipkow@51075
   148
by(metis exec_trans[OF exec_appendR exec_appendL_if])
nipkow@13095
   149
nipkow@43982
   150
nipkow@51075
   151
declare Let_def[simp]
nipkow@43982
   152
nipkow@43982
   153
nipkow@43982
   154
subsection "Compilation"
nipkow@43982
   155
nipkow@43982
   156
fun acomp :: "aexp \<Rightarrow> instr list" where
nipkow@43982
   157
"acomp (N n) = [LOADI n]" |
nipkow@43982
   158
"acomp (V x) = [LOAD x]" |
nipkow@43982
   159
"acomp (Plus a1 a2) = acomp a1 @ acomp a2 @ [ADD]"
nipkow@43982
   160
nipkow@43982
   161
lemma acomp_correct[intro]:
kleing@44296
   162
  "acomp a \<turnstile> (0,s,stk) \<rightarrow>* (isize(acomp a),s,aval a s#stk)"
nipkow@51075
   163
by (induction a arbitrary: stk) fastforce+
nipkow@13095
   164
kleing@44296
   165
fun bcomp :: "bexp \<Rightarrow> bool \<Rightarrow> int \<Rightarrow> instr list" where
nipkow@46068
   166
"bcomp (Bc v) c n = (if v=c then [JMP n] else [])" |
nipkow@43982
   167
"bcomp (Not b) c n = bcomp b (\<not>c) n" |
nipkow@43982
   168
"bcomp (And b1 b2) c n =
nipkow@43982
   169
 (let cb2 = bcomp b2 c n;
kleing@44296
   170
        m = (if c then isize cb2 else isize cb2+n);
nipkow@43982
   171
      cb1 = bcomp b1 False m
nipkow@43982
   172
  in cb1 @ cb2)" |
nipkow@43982
   173
"bcomp (Less a1 a2) c n =
kleing@46193
   174
 acomp a1 @ acomp a2 @ (if c then [JMPLESS n] else [JMPGE n])"
nipkow@43982
   175
nipkow@43982
   176
value
nipkow@43982
   177
  "bcomp (And (Less (V ''x'') (V ''y'')) (Not(Less (V ''u'') (V ''v''))))
nipkow@43982
   178
     False 3"
nipkow@43982
   179
nipkow@43982
   180
lemma bcomp_correct[intro]:
kleing@44296
   181
  "0 \<le> n \<Longrightarrow>
kleing@44296
   182
  bcomp b c n \<turnstile>
kleing@44296
   183
 (0,s,stk)  \<rightarrow>*  (isize(bcomp b c n) + (if c = bval b s then n else 0),s,stk)"
wenzelm@46000
   184
proof(induction b arbitrary: c n)
nipkow@43982
   185
  case Not
nipkow@45761
   186
  from Not(1)[where c="~c"] Not(2) show ?case by fastforce
nipkow@13095
   187
next
nipkow@43982
   188
  case (And b1 b2)
nipkow@51075
   189
  from And(1)[of "if c then isize(bcomp b2 c n) else isize(bcomp b2 c n) + n" 
kleing@44296
   190
                 "False"] 
kleing@44296
   191
       And(2)[of n  "c"] And(3) 
nipkow@45761
   192
  show ?case by fastforce
nipkow@45761
   193
qed fastforce+
nipkow@13095
   194
nipkow@43982
   195
fun ccomp :: "com \<Rightarrow> instr list" where
nipkow@43982
   196
"ccomp SKIP = []" |
nipkow@43982
   197
"ccomp (x ::= a) = acomp a @ [STORE x]" |
nipkow@43982
   198
"ccomp (c\<^isub>1;c\<^isub>2) = ccomp c\<^isub>1 @ ccomp c\<^isub>2" |
nipkow@43982
   199
"ccomp (IF b THEN c\<^isub>1 ELSE c\<^isub>2) =
kleing@44296
   200
  (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
   201
   in cb @ cc\<^isub>1 @ JMP (isize cc\<^isub>2) # cc\<^isub>2)" |
nipkow@43982
   202
"ccomp (WHILE b DO c) =
kleing@44296
   203
 (let cc = ccomp c; cb = bcomp b False (isize cc + 1)
kleing@44875
   204
  in cb @ cc @ [JMP (-(isize cb + isize cc + 1))])"
kleing@44875
   205
nipkow@13095
   206
nipkow@43982
   207
value "ccomp
nipkow@43982
   208
 (IF Less (V ''u'') (N 1) THEN ''u'' ::= Plus (V ''u'') (N 1)
nipkow@43982
   209
  ELSE ''v'' ::= V ''u'')"
nipkow@13095
   210
nipkow@43982
   211
value "ccomp (WHILE Less (V ''u'') (N 1) DO (''u'' ::= Plus (V ''u'') (N 1)))"
nipkow@13095
   212
nipkow@13095
   213
Jean@45966
   214
subsection "Preservation of semantics"
nipkow@43982
   215
kleing@44296
   216
lemma ccomp_bigstep:
kleing@44296
   217
  "(c,s) \<Rightarrow> t \<Longrightarrow> ccomp c \<turnstile> (0,s,stk) \<rightarrow>* (isize(ccomp c),t,stk)"
nipkow@45886
   218
proof(induction arbitrary: stk rule: big_step_induct)
nipkow@43982
   219
  case (Assign x a s)
nipkow@45761
   220
  show ?case by (fastforce simp:fun_upd_def cong: if_cong)
nipkow@13095
   221
next
nipkow@48689
   222
  case (Seq c1 s1 s2 c2 s3)
nipkow@43982
   223
  let ?cc1 = "ccomp c1"  let ?cc2 = "ccomp c2"
kleing@44296
   224
  have "?cc1 @ ?cc2 \<turnstile> (0,s1,stk) \<rightarrow>* (isize ?cc1,s2,stk)"
nipkow@48689
   225
    using Seq.IH(1) by fastforce
nipkow@43982
   226
  moreover
kleing@44296
   227
  have "?cc1 @ ?cc2 \<turnstile> (isize ?cc1,s2,stk) \<rightarrow>* (isize(?cc1 @ ?cc2),s3,stk)"
nipkow@48689
   228
    using Seq.IH(2) by fastforce
nipkow@43982
   229
  ultimately show ?case by simp (blast intro: exec_trans)
nipkow@13095
   230
next
nipkow@43982
   231
  case (WhileTrue b s1 c s2 s3)
nipkow@43982
   232
  let ?cc = "ccomp c"
kleing@44296
   233
  let ?cb = "bcomp b False (isize ?cc + 1)"
nipkow@43982
   234
  let ?cw = "ccomp(WHILE b DO c)"
kleing@44296
   235
  have "?cw \<turnstile> (0,s1,stk) \<rightarrow>* (isize ?cb + isize ?cc,s2,stk)"
nipkow@45886
   236
    using WhileTrue.IH(1) WhileTrue.hyps(1) by fastforce
nipkow@43982
   237
  moreover
kleing@44296
   238
  have "?cw \<turnstile> (isize ?cb + isize ?cc,s2,stk) \<rightarrow>* (0,s2,stk)"
nipkow@45761
   239
    by fastforce
nipkow@43982
   240
  moreover
nipkow@45886
   241
  have "?cw \<turnstile> (0,s2,stk) \<rightarrow>* (isize ?cw,s3,stk)" by(rule WhileTrue.IH(2))
nipkow@43982
   242
  ultimately show ?case by(blast intro: exec_trans)
nipkow@45761
   243
qed fastforce+
nipkow@13095
   244
webertj@20217
   245
end