src/HOL/IMP/Compiler.thy
author wenzelm
Tue, 13 Aug 2013 16:25:47 +0200
changeset 54152 a1119cf551e8
parent 54089 07423b37bc22
child 55006 a6f6df7f01cf
permissions -rw-r--r--
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
kleing@52396
     1
(* Author: Tobias Nipkow and Gerwin Klein *)
nipkow@10343
     2
nipkow@51076
     3
header "Compiler for IMP"
nipkow@10342
     4
kleing@54052
     5
theory Compiler imports Big_Step Star
nipkow@43982
     6
begin
kleing@12429
     7
kleing@44296
     8
subsection "List setup"
kleing@44296
     9
kleing@52396
    10
text {* 
kleing@52396
    11
  In the following, we use the length of lists as integers 
kleing@52396
    12
  instead of natural numbers. Instead of converting @{typ nat}
kleing@52396
    13
  to @{typ int} explicitly, we tell Isabelle to coerce @{typ nat}
kleing@52396
    14
  automatically when necessary.
kleing@44296
    15
*}
kleing@52396
    16
declare [[coercion_enabled]] 
kleing@52396
    17
declare [[coercion "int :: nat \<Rightarrow> int"]]
kleing@44296
    18
kleing@52396
    19
text {* 
kleing@52396
    20
  Similarly, we will want to access the ith element of a list, 
kleing@52396
    21
  where @{term i} is an @{typ int}.
kleing@52396
    22
*}
nipkow@46501
    23
fun inth :: "'a list \<Rightarrow> int \<Rightarrow> 'a" (infixl "!!" 100) where
nipkow@46501
    24
"(x # xs) !! n = (if n = 0 then x else xs !! (n - 1))"
kleing@44296
    25
kleing@44296
    26
text {*
kleing@52396
    27
  The only additional lemma we need about this function 
kleing@52396
    28
  is indexing over append:
kleing@44296
    29
*}
kleing@44296
    30
lemma inth_append [simp]:
kleing@44296
    31
  "0 \<le> n \<Longrightarrow>
kleing@52396
    32
  (xs @ ys) !! n = (if n < size xs then xs !! n else ys !! (n - size xs))"
nipkow@51075
    33
by (induction xs arbitrary: n) (auto simp: algebra_simps)
kleing@44296
    34
nipkow@43982
    35
subsection "Instructions and Stack Machine"
nipkow@43982
    36
kleing@54043
    37
text_raw{*\snip{instrdef}{0}{1}{% *}
nipkow@43982
    38
datatype instr = 
kleing@54043
    39
  LOADI int | LOAD vname | ADD | STORE vname |
kleing@54043
    40
  JMP int | JMPLESS int | JMPGE int
kleing@54043
    41
text_raw{*}%endsnip*}
nipkow@43982
    42
kleing@44296
    43
type_synonym stack = "val list"
nipkow@46501
    44
type_synonym config = "int \<times> state \<times> stack"
nipkow@43982
    45
nipkow@43982
    46
abbreviation "hd2 xs == hd(tl xs)"
nipkow@43982
    47
abbreviation "tl2 xs == tl(tl xs)"
nipkow@43982
    48
nipkow@51075
    49
fun iexec :: "instr \<Rightarrow> config \<Rightarrow> config" where
nipkow@51075
    50
"iexec instr (i,s,stk) = (case instr of
nipkow@51075
    51
  LOADI n \<Rightarrow> (i+1,s, n#stk) |
nipkow@51075
    52
  LOAD x \<Rightarrow> (i+1,s, s x # stk) |
nipkow@51075
    53
  ADD \<Rightarrow> (i+1,s, (hd2 stk + hd stk) # tl2 stk) |
nipkow@51075
    54
  STORE x \<Rightarrow> (i+1,s(x := hd stk),tl stk) |
nipkow@51075
    55
  JMP n \<Rightarrow>  (i+1+n,s,stk) |
nipkow@51075
    56
  JMPLESS n \<Rightarrow> (if hd2 stk < hd stk then i+1+n else i+1,s,tl2 stk) |
nipkow@51075
    57
  JMPGE n \<Rightarrow> (if hd2 stk >= hd stk then i+1+n else i+1,s,tl2 stk))"
kleing@44296
    58
kleing@44871
    59
definition
nipkow@51075
    60
  exec1 :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool"
nipkow@51075
    61
     ("(_/ \<turnstile> (_ \<rightarrow>/ _))" [59,0,59] 60) 
kleing@44871
    62
where
kleing@44296
    63
  "P \<turnstile> c \<rightarrow> c' = 
kleing@52396
    64
  (\<exists>i s stk. c = (i,s,stk) \<and> c' = iexec(P!!i) (i,s,stk) \<and> 0 \<le> i \<and> i < size P)"
kleing@44871
    65
kleing@44871
    66
lemma exec1I [intro, code_pred_intro]:
kleing@52396
    67
  "c' = iexec (P!!i) (i,s,stk) \<Longrightarrow> 0 \<le> i \<Longrightarrow> i < size P
nipkow@51075
    68
  \<Longrightarrow> P \<turnstile> (i,s,stk) \<rightarrow> c'"
kleing@54052
    69
by (simp add: exec1_def)
nipkow@43982
    70
kleing@54052
    71
abbreviation 
kleing@54052
    72
  exec :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool" ("(_/ \<turnstile> (_ \<rightarrow>*/ _))" 50)
nipkow@43982
    73
where
kleing@54052
    74
  "exec P \<equiv> star (exec1 P)"
nipkow@43982
    75
kleing@54052
    76
declare star.step[intro]
nipkow@43982
    77
kleing@54052
    78
lemmas exec_induct = star.induct [of "exec1 P", split_format(complete)]
nipkow@43982
    79
kleing@54052
    80
code_pred exec1 by (metis exec1_def)
nipkow@43982
    81
nipkow@43982
    82
values
nipkow@43982
    83
  "{(i,map t [''x'',''y''],stk) | i t stk.
nipkow@43982
    84
    [LOAD ''y'', STORE ''x''] \<turnstile>
kleing@44907
    85
    (0, <''x'' := 3, ''y'' := 4>, []) \<rightarrow>* (i,t,stk)}"
nipkow@43982
    86
nipkow@43982
    87
nipkow@43982
    88
subsection{* Verification infrastructure *}
nipkow@43982
    89
nipkow@43982
    90
text{* Below we need to argue about the execution of code that is embedded in
nipkow@43982
    91
larger programs. For this purpose we show that execution is preserved by
nipkow@43982
    92
appending code to the left or right of a program. *}
nipkow@43982
    93
nipkow@51075
    94
lemma iexec_shift [simp]: 
nipkow@51075
    95
  "((n+i',s',stk') = iexec x (n+i,s,stk)) = ((i',s',stk') = iexec x (i,s,stk))"
nipkow@51075
    96
by(auto split:instr.split)
nipkow@51075
    97
kleing@44296
    98
lemma exec1_appendR: "P \<turnstile> c \<rightarrow> c' \<Longrightarrow> P@P' \<turnstile> c \<rightarrow> c'"
kleing@54052
    99
by (auto simp: exec1_def)
nipkow@11275
   100
nipkow@43982
   101
lemma exec_appendR: "P \<turnstile> c \<rightarrow>* c' \<Longrightarrow> P@P' \<turnstile> c \<rightarrow>* c'"
kleing@54052
   102
by (induction rule: star.induct) (fastforce intro: exec1_appendR)+
nipkow@13095
   103
nipkow@43982
   104
lemma exec1_appendL:
kleing@52396
   105
  fixes i i' :: int 
kleing@52396
   106
  shows
kleing@44296
   107
  "P \<turnstile> (i,s,stk) \<rightarrow> (i',s',stk') \<Longrightarrow>
kleing@52396
   108
   P' @ P \<turnstile> (size(P')+i,s,stk) \<rightarrow> (size(P')+i',s',stk')"
kleing@54052
   109
  unfolding exec1_def
kleing@54089
   110
  by (auto simp del: iexec.simps)
nipkow@13095
   111
nipkow@43982
   112
lemma exec_appendL:
kleing@52396
   113
  fixes i i' :: int 
kleing@52396
   114
  shows
nipkow@43982
   115
 "P \<turnstile> (i,s,stk) \<rightarrow>* (i',s',stk')  \<Longrightarrow>
kleing@52396
   116
  P' @ P \<turnstile> (size(P')+i,s,stk) \<rightarrow>* (size(P')+i',s',stk')"
kleing@52396
   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@52396
   133
  fixes i i' :: int
kleing@52396
   134
  shows
kleing@52396
   135
  "size P' <= i
kleing@52396
   136
   \<Longrightarrow> P \<turnstile> (i - size P',s,stk) \<rightarrow>* (i',s',stk')
kleing@52396
   137
   \<Longrightarrow> P' @ P \<turnstile> (i,s,stk) \<rightarrow>* (size P' + i',s',stk')"
nipkow@51075
   138
by (drule exec_appendL[where P'=P']) simp
nipkow@13095
   139
nipkow@43982
   140
text{* Split the execution of a compound program up into the excution of its
nipkow@43982
   141
parts: *}
nipkow@13095
   142
nipkow@43982
   143
lemma exec_append_trans[intro]:
kleing@52396
   144
  fixes i' i'' j'' :: int
kleing@52396
   145
  shows
nipkow@43982
   146
"P \<turnstile> (0,s,stk) \<rightarrow>* (i',s',stk') \<Longrightarrow>
kleing@52396
   147
 size P \<le> i' \<Longrightarrow>
kleing@52396
   148
 P' \<turnstile>  (i' - size P,s',stk') \<rightarrow>* (i'',s'',stk'') \<Longrightarrow>
kleing@52396
   149
 j'' = size P + i''
nipkow@43982
   150
 \<Longrightarrow>
nipkow@43982
   151
 P @ P' \<turnstile> (0,s,stk) \<rightarrow>* (j'',s'',stk'')"
kleing@54052
   152
by(metis star_trans[OF exec_appendR exec_appendL_if])
nipkow@13095
   153
nipkow@43982
   154
nipkow@51075
   155
declare Let_def[simp]
nipkow@43982
   156
nipkow@43982
   157
nipkow@43982
   158
subsection "Compilation"
nipkow@43982
   159
nipkow@43982
   160
fun acomp :: "aexp \<Rightarrow> instr list" where
nipkow@43982
   161
"acomp (N n) = [LOADI n]" |
nipkow@43982
   162
"acomp (V x) = [LOAD x]" |
nipkow@43982
   163
"acomp (Plus a1 a2) = acomp a1 @ acomp a2 @ [ADD]"
nipkow@43982
   164
nipkow@43982
   165
lemma acomp_correct[intro]:
kleing@52396
   166
  "acomp a \<turnstile> (0,s,stk) \<rightarrow>* (size(acomp a),s,aval a s#stk)"
nipkow@51075
   167
by (induction a arbitrary: stk) fastforce+
nipkow@13095
   168
kleing@44296
   169
fun bcomp :: "bexp \<Rightarrow> bool \<Rightarrow> int \<Rightarrow> instr list" where
nipkow@46068
   170
"bcomp (Bc v) c n = (if v=c then [JMP n] else [])" |
nipkow@43982
   171
"bcomp (Not b) c n = bcomp b (\<not>c) n" |
nipkow@43982
   172
"bcomp (And b1 b2) c n =
nipkow@43982
   173
 (let cb2 = bcomp b2 c n;
kleing@52396
   174
        m = (if c then size cb2 else (size cb2::int)+n);
nipkow@43982
   175
      cb1 = bcomp b1 False m
nipkow@43982
   176
  in cb1 @ cb2)" |
nipkow@43982
   177
"bcomp (Less a1 a2) c n =
kleing@46193
   178
 acomp a1 @ acomp a2 @ (if c then [JMPLESS n] else [JMPGE n])"
nipkow@43982
   179
nipkow@43982
   180
value
nipkow@43982
   181
  "bcomp (And (Less (V ''x'') (V ''y'')) (Not(Less (V ''u'') (V ''v''))))
nipkow@43982
   182
     False 3"
nipkow@43982
   183
nipkow@43982
   184
lemma bcomp_correct[intro]:
kleing@52396
   185
  fixes n :: int
kleing@52396
   186
  shows
kleing@44296
   187
  "0 \<le> n \<Longrightarrow>
kleing@44296
   188
  bcomp b c n \<turnstile>
kleing@52396
   189
 (0,s,stk)  \<rightarrow>*  (size(bcomp b c n) + (if c = bval b s then n else 0),s,stk)"
wenzelm@46000
   190
proof(induction b arbitrary: c n)
nipkow@43982
   191
  case Not
nipkow@45761
   192
  from Not(1)[where c="~c"] Not(2) show ?case by fastforce
nipkow@13095
   193
next
nipkow@43982
   194
  case (And b1 b2)
kleing@52396
   195
  from And(1)[of "if c then size(bcomp b2 c n) else size(bcomp b2 c n) + n" 
kleing@44296
   196
                 "False"] 
kleing@44296
   197
       And(2)[of n  "c"] And(3) 
nipkow@45761
   198
  show ?case by fastforce
nipkow@45761
   199
qed fastforce+
nipkow@13095
   200
nipkow@43982
   201
fun ccomp :: "com \<Rightarrow> instr list" where
nipkow@43982
   202
"ccomp SKIP = []" |
nipkow@43982
   203
"ccomp (x ::= a) = acomp a @ [STORE x]" |
wenzelm@54152
   204
"ccomp (c\<^sub>1;;c\<^sub>2) = ccomp c\<^sub>1 @ ccomp c\<^sub>2" |
wenzelm@54152
   205
"ccomp (IF b THEN c\<^sub>1 ELSE c\<^sub>2) =
wenzelm@54152
   206
  (let cc\<^sub>1 = ccomp c\<^sub>1; cc\<^sub>2 = ccomp c\<^sub>2; cb = bcomp b False (size cc\<^sub>1 + 1)
wenzelm@54152
   207
   in cb @ cc\<^sub>1 @ JMP (size cc\<^sub>2) # cc\<^sub>2)" |
nipkow@43982
   208
"ccomp (WHILE b DO c) =
kleing@52396
   209
 (let cc = ccomp c; cb = bcomp b False (size cc + 1)
kleing@52396
   210
  in cb @ cc @ [JMP (-(size cb + size cc + 1))])"
kleing@44875
   211
nipkow@13095
   212
nipkow@43982
   213
value "ccomp
nipkow@43982
   214
 (IF Less (V ''u'') (N 1) THEN ''u'' ::= Plus (V ''u'') (N 1)
nipkow@43982
   215
  ELSE ''v'' ::= V ''u'')"
nipkow@13095
   216
nipkow@43982
   217
value "ccomp (WHILE Less (V ''u'') (N 1) DO (''u'' ::= Plus (V ''u'') (N 1)))"
nipkow@13095
   218
nipkow@13095
   219
Jean@45966
   220
subsection "Preservation of semantics"
nipkow@43982
   221
kleing@44296
   222
lemma ccomp_bigstep:
kleing@52396
   223
  "(c,s) \<Rightarrow> t \<Longrightarrow> ccomp c \<turnstile> (0,s,stk) \<rightarrow>* (size(ccomp c),t,stk)"
nipkow@45886
   224
proof(induction arbitrary: stk rule: big_step_induct)
nipkow@43982
   225
  case (Assign x a s)
nipkow@45761
   226
  show ?case by (fastforce simp:fun_upd_def cong: if_cong)
nipkow@13095
   227
next
nipkow@48689
   228
  case (Seq c1 s1 s2 c2 s3)
nipkow@43982
   229
  let ?cc1 = "ccomp c1"  let ?cc2 = "ccomp c2"
kleing@52396
   230
  have "?cc1 @ ?cc2 \<turnstile> (0,s1,stk) \<rightarrow>* (size ?cc1,s2,stk)"
nipkow@48689
   231
    using Seq.IH(1) by fastforce
nipkow@43982
   232
  moreover
kleing@52396
   233
  have "?cc1 @ ?cc2 \<turnstile> (size ?cc1,s2,stk) \<rightarrow>* (size(?cc1 @ ?cc2),s3,stk)"
nipkow@48689
   234
    using Seq.IH(2) by fastforce
kleing@54052
   235
  ultimately show ?case by simp (blast intro: star_trans)
nipkow@13095
   236
next
nipkow@43982
   237
  case (WhileTrue b s1 c s2 s3)
nipkow@43982
   238
  let ?cc = "ccomp c"
kleing@52396
   239
  let ?cb = "bcomp b False (size ?cc + 1)"
nipkow@43982
   240
  let ?cw = "ccomp(WHILE b DO c)"
kleing@52396
   241
  have "?cw \<turnstile> (0,s1,stk) \<rightarrow>* (size ?cb,s1,stk)"
nipkow@51148
   242
    using `bval b s1` by fastforce
nipkow@51148
   243
  moreover
kleing@52396
   244
  have "?cw \<turnstile> (size ?cb,s1,stk) \<rightarrow>* (size ?cb + size ?cc,s2,stk)"
nipkow@51148
   245
    using WhileTrue.IH(1) by fastforce
nipkow@43982
   246
  moreover
kleing@52396
   247
  have "?cw \<turnstile> (size ?cb + size ?cc,s2,stk) \<rightarrow>* (0,s2,stk)"
nipkow@45761
   248
    by fastforce
nipkow@43982
   249
  moreover
kleing@52396
   250
  have "?cw \<turnstile> (0,s2,stk) \<rightarrow>* (size ?cw,s3,stk)" by(rule WhileTrue.IH(2))
kleing@54052
   251
  ultimately show ?case by(blast intro: star_trans)
nipkow@45761
   252
qed fastforce+
nipkow@13095
   253
webertj@20217
   254
end