src/HOL/IMP/Compiler.thy
author blanchet
Mon, 06 Jun 2011 20:36:35 +0200
changeset 43181 cd3b7798ecc2
parent 43158 686fa0a0696e
child 43438 a666b8d11252
permissions -rw-r--r--
don't stumble on Skolem names
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
     1
(* Author: Tobias Nipkow *)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
     2
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
     3
header "A Compiler for IMP"
10343
24c87e1366d8 *** empty log message ***
nipkow
parents: 10342
diff changeset
     4
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
     5
theory Compiler imports Big_Step
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
     6
begin
12429
15c13bdc94c8 tuned for latex output
kleing
parents: 12275
diff changeset
     7
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
     8
subsection "Instructions and Stack Machine"
10342
b124d59f7b61 *** empty log message ***
nipkow
parents:
diff changeset
     9
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    10
datatype instr = 
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    11
  LOADI int | LOAD string | ADD |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    12
  STORE string |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    13
  JMPF nat |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    14
  JMPB nat |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    15
  JMPFLESS nat |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    16
  JMPFGE nat
10342
b124d59f7b61 *** empty log message ***
nipkow
parents:
diff changeset
    17
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    18
type_synonym stack = "int list"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    19
type_synonym config = "nat\<times>state\<times>stack"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    20
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    21
abbreviation "hd2 xs == hd(tl xs)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    22
abbreviation "tl2 xs == tl(tl xs)"
11275
71498de45241 new proof
nipkow
parents: 10425
diff changeset
    23
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    24
inductive exec1 :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    25
    ("(_/ \<turnstile> (_ \<rightarrow>/ _))" [50,0,0] 50)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    26
  for P :: "instr list"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    27
where
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    28
"\<lbrakk> i < size P;  P!i = LOADI n \<rbrakk> \<Longrightarrow>
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    29
 P \<turnstile> (i,s,stk) \<rightarrow> (i+1,s, n#stk)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    30
"\<lbrakk> i < size P;  P!i = LOAD x \<rbrakk> \<Longrightarrow> 
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    31
 P \<turnstile> (i,s,stk) \<rightarrow> (i+1,s, s x # stk)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    32
"\<lbrakk> i < size P;  P!i = ADD \<rbrakk> \<Longrightarrow> 
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    33
 P \<turnstile> (i,s,stk) \<rightarrow> (i+1,s, (hd2 stk + hd stk) # tl2 stk)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    34
"\<lbrakk> i < size P;  P!i = STORE n \<rbrakk> \<Longrightarrow>
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    35
 P \<turnstile> (i,s,stk) \<rightarrow> (i+1,s(n := hd stk),tl stk)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    36
"\<lbrakk> i < size P;  P!i = JMPF n \<rbrakk> \<Longrightarrow>
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    37
 P \<turnstile> (i,s,stk) \<rightarrow> (i+1+n,s,stk)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    38
"\<lbrakk> i < size P;  P!i = JMPB n;  n \<le> i+1 \<rbrakk> \<Longrightarrow>
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    39
 P \<turnstile> (i,s,stk) \<rightarrow> (i+1-n,s,stk)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    40
"\<lbrakk> i < size P;  P!i = JMPFLESS n \<rbrakk> \<Longrightarrow>
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    41
 P \<turnstile> (i,s,stk) \<rightarrow> (if hd2 stk < hd stk then i+1+n else i+1,s,tl2 stk)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    42
"\<lbrakk> i < size P;  P!i = JMPFGE n \<rbrakk> \<Longrightarrow>
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    43
 P \<turnstile> (i,s,stk) \<rightarrow> (if hd2 stk >= hd stk then i+1+n else i+1,s,tl2 stk)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    44
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    45
code_pred exec1 .
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    46
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    47
declare exec1.intros[intro]
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    48
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    49
inductive exec :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool" ("_/ \<turnstile> (_ \<rightarrow>*/ _)" 50)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    50
where
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    51
refl: "P \<turnstile> c \<rightarrow>* c" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    52
step: "P \<turnstile> c \<rightarrow> c' \<Longrightarrow> P \<turnstile> c' \<rightarrow>* c'' \<Longrightarrow> P \<turnstile> c \<rightarrow>* c''"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    53
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    54
declare exec.intros[intro]
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    55
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    56
lemmas exec_induct = exec.induct[split_format(complete)]
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    57
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    58
code_pred exec .
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    59
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    60
values
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    61
  "{(i,map t [''x'',''y''],stk) | i t stk.
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    62
    [LOAD ''y'', STORE ''x''] \<turnstile>
43158
686fa0a0696e imported rest of new IMP
kleing
parents: 43141
diff changeset
    63
    (0, [''x'' \<rightarrow> 3, ''y'' \<rightarrow> 4], []) \<rightarrow>* (i,t,stk)}"
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    64
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    65
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    66
subsection{* Verification infrastructure *}
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    67
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    68
lemma exec_trans: "P \<turnstile> c \<rightarrow>* c' \<Longrightarrow> P \<turnstile> c' \<rightarrow>* c'' \<Longrightarrow> P \<turnstile> c \<rightarrow>* c''"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    69
apply(induct rule: exec.induct)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    70
 apply blast
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    71
by (metis exec.step)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    72
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    73
lemma exec1_subst: "P \<turnstile> c \<rightarrow> c' \<Longrightarrow> c' = c'' \<Longrightarrow> P \<turnstile> c \<rightarrow> c''"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    74
by auto
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    75
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    76
lemmas exec1_simps = exec1.intros[THEN exec1_subst]
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    77
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    78
text{* Below we need to argue about the execution of code that is embedded in
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    79
larger programs. For this purpose we show that execution is preserved by
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    80
appending code to the left or right of a program. *}
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    81
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    82
lemma exec1_appendR: assumes "P \<turnstile> c \<rightarrow> c'" shows "P@P' \<turnstile> c \<rightarrow> c'"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    83
proof-
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    84
  from assms show ?thesis
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    85
  by cases (simp_all add: exec1_simps nth_append)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    86
  -- "All cases proved with the final simp-all"
12275
aa2b7b475a94 Isar conversion
nipkow
parents: 11704
diff changeset
    87
qed
11275
71498de45241 new proof
nipkow
parents: 10425
diff changeset
    88
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    89
lemma exec_appendR: "P \<turnstile> c \<rightarrow>* c' \<Longrightarrow> P@P' \<turnstile> c \<rightarrow>* c'"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    90
apply(induct rule: exec.induct)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    91
 apply blast
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    92
by (metis exec1_appendR exec.step)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    93
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    94
lemma exec1_appendL:
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    95
assumes "P \<turnstile> (i,s,stk) \<rightarrow> (i',s',stk')"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    96
shows "P' @ P \<turnstile> (size(P')+i,s,stk) \<rightarrow> (size(P')+i',s',stk')"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    97
proof-
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    98
  from assms show ?thesis
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
    99
  by cases (simp_all add: exec1_simps)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   100
qed
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   101
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   102
lemma exec_appendL:
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   103
 "P \<turnstile> (i,s,stk) \<rightarrow>* (i',s',stk')  \<Longrightarrow>
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   104
  P' @ P \<turnstile> (size(P')+i,s,stk) \<rightarrow>* (size(P')+i',s',stk')"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   105
apply(induct rule: exec_induct)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   106
 apply blast
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   107
by (blast intro: exec1_appendL exec.step)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   108
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   109
text{* Now we specialise the above lemmas to enable automatic proofs of
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   110
@{prop "P \<turnstile> c \<rightarrow>* c'"} where @{text P} is a mixture of concrete instructions and
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   111
pieces of code that we already know how they execute (by induction), combined
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   112
by @{text "@"} and @{text "#"}. Backward jumps are not supported.
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   113
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
   114
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   115
If the pc points beyond the first instruction or part of the program, drop it: *}
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   116
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   117
lemma exec_Cons_Suc[intro]:
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   118
  "P \<turnstile> (i,s,stk) \<rightarrow>* (j,t,stk') \<Longrightarrow>
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   119
  instr#P \<turnstile> (Suc i,s,stk) \<rightarrow>* (Suc j,t,stk')"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   120
apply(drule exec_appendL[where P'="[instr]"])
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   121
apply simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   122
done
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   123
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   124
lemma exec_appendL_if[intro]:
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   125
 "size P' <= i
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   126
  \<Longrightarrow> P \<turnstile> (i - size P',s,stk) \<rightarrow>* (i',s',stk')
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   127
  \<Longrightarrow> P' @ P \<turnstile> (i,s,stk) \<rightarrow>* (size P' + i',s',stk')"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   128
apply(drule exec_appendL[where P'=P'])
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   129
apply simp
10342
b124d59f7b61 *** empty log message ***
nipkow
parents:
diff changeset
   130
done
b124d59f7b61 *** empty log message ***
nipkow
parents:
diff changeset
   131
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   132
text{* Split the execution of a compound program up into the excution of its
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   133
parts: *}
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   134
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   135
lemma exec_append_trans[intro]:
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   136
"P \<turnstile> (0,s,stk) \<rightarrow>* (i',s',stk') \<Longrightarrow>
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   137
 size P \<le> i' \<Longrightarrow>
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   138
 P' \<turnstile>  (i' - size P,s',stk') \<rightarrow>* (i'',s'',stk'') \<Longrightarrow>
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   139
 j'' = size P + i''
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   140
 \<Longrightarrow>
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   141
 P @ P' \<turnstile> (0,s,stk) \<rightarrow>* (j'',s'',stk'')"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   142
by(metis exec_trans[OF  exec_appendR exec_appendL_if])
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   143
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   144
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   145
declare Let_def[simp] eval_nat_numeral[simp]
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   146
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   147
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   148
subsection "Compilation"
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   149
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   150
fun acomp :: "aexp \<Rightarrow> instr list" where
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   151
"acomp (N n) = [LOADI n]" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   152
"acomp (V x) = [LOAD x]" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   153
"acomp (Plus a1 a2) = acomp a1 @ acomp a2 @ [ADD]"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   154
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   155
lemma acomp_correct[intro]:
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   156
  "acomp a \<turnstile> (0,s,stk) \<rightarrow>* (size(acomp a),s,aval a s#stk)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   157
apply(induct a arbitrary: stk)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   158
apply(fastsimp)+
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   159
done
10342
b124d59f7b61 *** empty log message ***
nipkow
parents:
diff changeset
   160
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   161
fun bcomp :: "bexp \<Rightarrow> bool \<Rightarrow> nat \<Rightarrow> instr list" where
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   162
"bcomp (B bv) c n = (if bv=c then [JMPF n] else [])" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   163
"bcomp (Not b) c n = bcomp b (\<not>c) n" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   164
"bcomp (And b1 b2) c n =
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   165
 (let cb2 = bcomp b2 c n;
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   166
      m = (if c then size cb2 else size cb2+n);
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   167
      cb1 = bcomp b1 False m
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   168
  in cb1 @ cb2)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   169
"bcomp (Less a1 a2) c n =
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   170
 acomp a1 @ acomp a2 @ (if c then [JMPFLESS n] else [JMPFGE n])"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   171
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   172
value
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   173
  "bcomp (And (Less (V ''x'') (V ''y'')) (Not(Less (V ''u'') (V ''v''))))
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   174
     False 3"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   175
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   176
lemma bcomp_correct[intro]:
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   177
 "bcomp b c n \<turnstile>
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   178
 (0,s,stk)  \<rightarrow>*  (size(bcomp b c n) + (if c = bval b s then n else 0),s,stk)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   179
proof(induct b arbitrary: c n m)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   180
  case Not
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   181
  from Not[of "~c"] show ?case by fastsimp
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   182
next
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   183
  case (And b1 b2)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   184
  from And(1)[of "False"] And(2)[of "c"] show ?case by fastsimp
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   185
qed fastsimp+
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   186
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   187
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   188
fun ccomp :: "com \<Rightarrow> instr list" where
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   189
"ccomp SKIP = []" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   190
"ccomp (x ::= a) = acomp a @ [STORE x]" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   191
"ccomp (c\<^isub>1;c\<^isub>2) = ccomp c\<^isub>1 @ ccomp c\<^isub>2" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   192
"ccomp (IF b THEN c\<^isub>1 ELSE c\<^isub>2) =
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   193
  (let cc\<^isub>1 = ccomp c\<^isub>1; cc\<^isub>2 = ccomp c\<^isub>2; cb = bcomp b False (size cc\<^isub>1 + 1)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   194
   in cb @ cc\<^isub>1 @ JMPF(size cc\<^isub>2) # cc\<^isub>2)" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   195
"ccomp (WHILE b DO c) =
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   196
 (let cc = ccomp c; cb = bcomp b False (size cc + 1)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   197
  in cb @ cc @ [JMPB (size cb + size cc + 1)])"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   198
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   199
value "ccomp
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   200
 (IF Less (V ''u'') (N 1) THEN ''u'' ::= Plus (V ''u'') (N 1)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   201
  ELSE ''v'' ::= V ''u'')"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   202
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   203
value "ccomp (WHILE Less (V ''u'') (N 1) DO (''u'' ::= Plus (V ''u'') (N 1)))"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   204
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   205
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   206
subsection "Preservation of sematics"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   207
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   208
lemma ccomp_correct:
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   209
  "(c,s) \<Rightarrow> t \<Longrightarrow> ccomp c \<turnstile> (0,s,stk) \<rightarrow>* (size(ccomp c),t,stk)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   210
proof(induct arbitrary: stk rule: big_step_induct)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   211
  case (Assign x a s)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   212
  show ?case by (fastsimp simp:fun_upd_def)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   213
next
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   214
  case (Semi c1 s1 s2 c2 s3)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   215
  let ?cc1 = "ccomp c1"  let ?cc2 = "ccomp c2"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   216
  have "?cc1 @ ?cc2 \<turnstile> (0,s1,stk) \<rightarrow>* (size ?cc1,s2,stk)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   217
    using Semi.hyps(2) by (fastsimp)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   218
  moreover
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   219
  have "?cc1 @ ?cc2 \<turnstile> (size ?cc1,s2,stk) \<rightarrow>* (size(?cc1 @ ?cc2),s3,stk)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   220
    using Semi.hyps(4) by (fastsimp)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   221
  ultimately show ?case by simp (blast intro: exec_trans)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   222
next
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   223
  case (WhileTrue b s1 c s2 s3)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   224
  let ?cc = "ccomp c"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   225
  let ?cb = "bcomp b False (size ?cc + 1)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   226
  let ?cw = "ccomp(WHILE b DO c)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   227
  have "?cw \<turnstile> (0,s1,stk) \<rightarrow>* (size ?cb + size ?cc,s2,stk)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   228
    using WhileTrue(1,3) by fastsimp
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   229
  moreover
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   230
  have "?cw \<turnstile> (size ?cb + size ?cc,s2,stk) \<rightarrow>* (0,s2,stk)"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   231
    by (fastsimp)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   232
  moreover
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   233
  have "?cw \<turnstile> (0,s2,stk) \<rightarrow>* (size ?cw,s3,stk)" by(rule WhileTrue(5))
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   234
  ultimately show ?case by(blast intro: exec_trans)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents: 32960
diff changeset
   235
qed fastsimp+
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   236
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 16417
diff changeset
   237
end