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