src/HOL/IMP/Compiler0.thy
author haftmann
Fri, 17 Jun 2005 16:12:49 +0200
changeset 16417 9bc16273c2d4
parent 14565 c6dc17aab88a
child 18372 2bffdf62fe7f
permissions -rw-r--r--
migrated theory headers to new format
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/IMP/Compiler.thy
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow, TUM
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
     4
    Copyright   1996 TUM
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
     5
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
     6
This is an early version of the compiler, where the abstract machine
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
     7
has an explicit pc. This turned out to be awkward, and a second
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
     8
development was started. See Machines.thy and Compiler.thy.
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
     9
*)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    10
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    11
header "A Simple Compiler"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    12
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14565
diff changeset
    13
theory Compiler0 imports Natural begin
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    14
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    15
subsection "An abstract, simplistic machine"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    16
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    17
text {* There are only three instructions: *}
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    18
datatype instr = ASIN loc aexp | JMPF bexp nat | JMPB nat
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    19
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    20
text {* We describe execution of programs in the machine by
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    21
  an operational (small step) semantics:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    22
*}
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    23
consts  stepa1 :: "instr list \<Rightarrow> ((state\<times>nat) \<times> (state\<times>nat))set"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    24
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    25
syntax
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    26
  "_stepa1" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    27
               ("_ |- (3<_,_>/ -1-> <_,_>)" [50,0,0,0,0] 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    28
  "_stepa" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    29
               ("_ |-/ (3<_,_>/ -*-> <_,_>)" [50,0,0,0,0] 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    30
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    31
  "_stepan" :: "[instr list,state,nat,nat,state,nat] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    32
               ("_ |-/ (3<_,_>/ -(_)-> <_,_>)" [50,0,0,0,0,0] 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    33
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    34
syntax (xsymbols)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    35
  "_stepa1" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    36
               ("_ \<turnstile> (3\<langle>_,_\<rangle>/ -1\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    37
  "_stepa" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    38
               ("_ \<turnstile>/ (3\<langle>_,_\<rangle>/ -*\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    39
  "_stepan" :: "[instr list,state,nat,nat,state,nat] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    40
               ("_ \<turnstile>/ (3\<langle>_,_\<rangle>/ -(_)\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0,0] 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    41
14565
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13130
diff changeset
    42
syntax (HTML output)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13130
diff changeset
    43
  "_stepa1" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool"
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13130
diff changeset
    44
               ("_ |- (3\<langle>_,_\<rangle>/ -1\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13130
diff changeset
    45
  "_stepa" :: "[instr list,state,nat,state,nat] \<Rightarrow> bool"
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13130
diff changeset
    46
               ("_ |-/ (3\<langle>_,_\<rangle>/ -*\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0] 50)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13130
diff changeset
    47
  "_stepan" :: "[instr list,state,nat,nat,state,nat] \<Rightarrow> bool"
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13130
diff changeset
    48
               ("_ |-/ (3\<langle>_,_\<rangle>/ -(_)\<rightarrow> \<langle>_,_\<rangle>)" [50,0,0,0,0,0] 50)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13130
diff changeset
    49
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    50
translations  
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    51
  "P \<turnstile> \<langle>s,m\<rangle> -1\<rightarrow> \<langle>t,n\<rangle>" == "((s,m),t,n) : stepa1 P"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    52
  "P \<turnstile> \<langle>s,m\<rangle> -*\<rightarrow> \<langle>t,n\<rangle>" == "((s,m),t,n) : ((stepa1 P)^*)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    53
  "P \<turnstile> \<langle>s,m\<rangle> -(i)\<rightarrow> \<langle>t,n\<rangle>" == "((s,m),t,n) : ((stepa1 P)^i)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    54
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    55
inductive "stepa1 P"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    56
intros
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    57
ASIN[simp]:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    58
  "\<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>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    59
JMPFT[simp,intro]:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    60
  "\<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>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    61
JMPFF[simp,intro]:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    62
  "\<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>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    63
JMPB[simp]:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    64
  "\<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>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    65
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    66
subsection "The compiler"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    67
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    68
consts compile :: "com \<Rightarrow> instr list"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    69
primrec
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    70
"compile \<SKIP> = []"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    71
"compile (x:==a) = [ASIN x a]"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    72
"compile (c1;c2) = compile c1 @ compile c2"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    73
"compile (\<IF> b \<THEN> c1 \<ELSE> c2) =
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    74
 [JMPF b (length(compile c1) + 2)] @ compile c1 @
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    75
 [JMPF (%x. False) (length(compile c2)+1)] @ compile c2"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    76
"compile (\<WHILE> b \<DO> c) = [JMPF b (length(compile c) + 2)] @ compile c @
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    77
 [JMPB (length(compile c)+1)]"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    78
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    79
declare nth_append[simp]
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    80
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    81
subsection "Context lifting lemmas"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    82
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    83
text {* 
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    84
  Some lemmas for lifting an execution into a prefix and suffix
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    85
  of instructions; only needed for the first proof.
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    86
*}
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    87
lemma app_right_1:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    88
  assumes A: "is1 \<turnstile> \<langle>s1,i1\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    89
  shows "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    90
proof -
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    91
 from A show ?thesis
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    92
 by induct force+
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    93
qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    94
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    95
lemma app_left_1:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    96
  assumes A: "is2 \<turnstile> \<langle>s1,i1\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    97
  shows "is1 @ is2 \<turnstile> \<langle>s1,size is1+i1\<rangle> -1\<rightarrow> \<langle>s2,size is1+i2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    98
proof -
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    99
 from A show ?thesis
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   100
 by induct force+
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   101
qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   102
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   103
declare rtrancl_induct2 [induct set: rtrancl]
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   104
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   105
lemma app_right:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   106
assumes A: "is1 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   107
shows "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   108
proof -
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   109
 from A show ?thesis
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   110
 proof induct
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   111
   show "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s1,i1\<rangle>" by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   112
 next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   113
   fix s1' i1' s2 i2
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   114
   assume "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s1',i1'\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   115
          "is1 \<turnstile> \<langle>s1',i1'\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   116
   thus "is1 @ is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   117
     by(blast intro:app_right_1 rtrancl_trans)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   118
 qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   119
qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   120
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   121
lemma app_left:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   122
assumes A: "is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   123
shows "is1 @ is2 \<turnstile> \<langle>s1,size is1+i1\<rangle> -*\<rightarrow> \<langle>s2,size is1+i2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   124
proof -
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   125
  from A show ?thesis
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   126
  proof induct
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   127
    show "is1 @ is2 \<turnstile> \<langle>s1,length is1 + i1\<rangle> -*\<rightarrow> \<langle>s1,length is1 + i1\<rangle>" by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   128
  next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   129
    fix s1' i1' s2 i2
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   130
    assume "is1 @ is2 \<turnstile> \<langle>s1,length is1 + i1\<rangle> -*\<rightarrow> \<langle>s1',length is1 + i1'\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   131
           "is2 \<turnstile> \<langle>s1',i1'\<rangle> -1\<rightarrow> \<langle>s2,i2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   132
    thus "is1 @ is2 \<turnstile> \<langle>s1,length is1 + i1\<rangle> -*\<rightarrow> \<langle>s2,length is1 + i2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   133
      by(blast intro:app_left_1 rtrancl_trans)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   134
 qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   135
qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   136
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   137
lemma app_left2:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   138
  "\<lbrakk> is2 \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle>; j1 = size is1+i1; j2 = size is1+i2 \<rbrakk> \<Longrightarrow>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   139
   is1 @ is2 \<turnstile> \<langle>s1,j1\<rangle> -*\<rightarrow> \<langle>s2,j2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   140
  by (simp add:app_left)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   141
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   142
lemma app1_left:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   143
  "is \<turnstile> \<langle>s1,i1\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle> \<Longrightarrow>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   144
   instr # is \<turnstile> \<langle>s1,Suc i1\<rangle> -*\<rightarrow> \<langle>s2,Suc i2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   145
  by(erule app_left[of _ _ _ _ _ "[instr]",simplified])
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   146
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   147
subsection "Compiler correctness"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   148
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   149
declare rtrancl_into_rtrancl[trans]
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   150
        converse_rtrancl_into_rtrancl[trans]
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   151
        rtrancl_trans[trans]
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   152
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   153
text {*
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   154
  The first proof; The statement is very intuitive,
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   155
  but application of induction hypothesis requires the above lifting lemmas
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   156
*}
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   157
theorem assumes A: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   158
shows "compile c \<turnstile> \<langle>s,0\<rangle> -*\<rightarrow> \<langle>t,length(compile c)\<rangle>" (is "?P c s t")
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   159
proof -
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   160
  from A show ?thesis
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   161
  proof induct
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   162
    show "\<And>s. ?P \<SKIP> s s" by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   163
  next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   164
    show "\<And>a s x. ?P (x :== a) s (s[x\<mapsto> a s])" by force
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   165
  next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   166
    fix c0 c1 s0 s1 s2
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   167
    assume "?P c0 s0 s1"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   168
    hence "compile c0 @ compile c1 \<turnstile> \<langle>s0,0\<rangle> -*\<rightarrow> \<langle>s1,length(compile c0)\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   169
      by(rule app_right)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   170
    moreover assume "?P c1 s1 s2"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   171
    hence "compile c0 @ compile c1 \<turnstile> \<langle>s1,length(compile c0)\<rangle> -*\<rightarrow>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   172
           \<langle>s2,length(compile c0)+length(compile c1)\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   173
    proof -
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   174
      show "\<And>is1 is2 s1 s2 i2.
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   175
	      is2 \<turnstile> \<langle>s1,0\<rangle> -*\<rightarrow> \<langle>s2,i2\<rangle> \<Longrightarrow>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   176
	      is1 @ is2 \<turnstile> \<langle>s1,size is1\<rangle> -*\<rightarrow> \<langle>s2,size is1+i2\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   177
	using app_left[of _ 0] by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   178
    qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   179
    ultimately have "compile c0 @ compile c1 \<turnstile> \<langle>s0,0\<rangle> -*\<rightarrow>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   180
                       \<langle>s2,length(compile c0)+length(compile c1)\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   181
      by (rule rtrancl_trans)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   182
    thus "?P (c0; c1) s0 s2" by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   183
  next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   184
    fix b c0 c1 s0 s1
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   185
    let ?comp = "compile(\<IF> b \<THEN> c0 \<ELSE> c1)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   186
    assume "b s0" and IH: "?P c0 s0 s1"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   187
    hence "?comp \<turnstile> \<langle>s0,0\<rangle> -1\<rightarrow> \<langle>s0,1\<rangle>" by auto
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   188
    also from IH
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   189
    have "?comp \<turnstile> \<langle>s0,1\<rangle> -*\<rightarrow> \<langle>s1,length(compile c0)+1\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   190
      by(auto intro:app1_left app_right)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   191
    also have "?comp \<turnstile> \<langle>s1,length(compile c0)+1\<rangle> -1\<rightarrow> \<langle>s1,length ?comp\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   192
      by(auto)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   193
    finally show "?P (\<IF> b \<THEN> c0 \<ELSE> c1) s0 s1" .
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   194
  next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   195
    fix b c0 c1 s0 s1
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   196
    let ?comp = "compile(\<IF> b \<THEN> c0 \<ELSE> c1)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   197
    assume "\<not>b s0" and IH: "?P c1 s0 s1"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   198
    hence "?comp \<turnstile> \<langle>s0,0\<rangle> -1\<rightarrow> \<langle>s0,length(compile c0) + 2\<rangle>" by auto
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   199
    also from IH
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   200
    have "?comp \<turnstile> \<langle>s0,length(compile c0)+2\<rangle> -*\<rightarrow> \<langle>s1,length ?comp\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   201
      by(force intro!:app_left2 app1_left)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   202
    finally show "?P (\<IF> b \<THEN> c0 \<ELSE> c1) s0 s1" .
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   203
  next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   204
    fix b c and s::state
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   205
    assume "\<not>b s"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   206
    thus "?P (\<WHILE> b \<DO> c) s s" by force
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   207
  next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   208
    fix b c and s0::state and s1 s2
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   209
    let ?comp = "compile(\<WHILE> b \<DO> c)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   210
    assume "b s0" and
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   211
      IHc: "?P c s0 s1" and IHw: "?P (\<WHILE> b \<DO> c) s1 s2"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   212
    hence "?comp \<turnstile> \<langle>s0,0\<rangle> -1\<rightarrow> \<langle>s0,1\<rangle>" by auto
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   213
    also from IHc
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   214
    have "?comp \<turnstile> \<langle>s0,1\<rangle> -*\<rightarrow> \<langle>s1,length(compile c)+1\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   215
      by(auto intro:app1_left app_right)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   216
    also have "?comp \<turnstile> \<langle>s1,length(compile c)+1\<rangle> -1\<rightarrow> \<langle>s1,0\<rangle>" by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   217
    also note IHw
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   218
    finally show "?P (\<WHILE> b \<DO> c) s0 s2".
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   219
  qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   220
qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   221
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   222
text {*
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   223
  Second proof; statement is generalized to cater for prefixes and suffixes;
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   224
  needs none of the lifting lemmas, but instantiations of pre/suffix.
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   225
  *}
13130
423ce375bf65 commented out half converted proof
nipkow
parents: 13112
diff changeset
   226
(*
13112
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   227
theorem assumes A: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t"
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   228
shows "\<And>a z. a@compile c@z \<turnstile> \<langle>s,size a\<rangle> -*\<rightarrow> \<langle>t,size a + size(compile c)\<rangle>"
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   229
      (is "\<And>a z. ?P c s t a z")
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   230
proof -
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   231
  from A show "\<And>a z. ?thesis a z"
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   232
  proof induct
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   233
    case Skip thus ?case by simp
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   234
  next
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   235
    case Assign thus ?case by (force intro!: ASIN)
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   236
  next
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   237
    fix c1 c2 s s' s'' a z
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   238
    assume IH1: "\<And>a z. ?P c1 s s' a z" and IH2: "\<And>a z. ?P c2 s' s'' a z"
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   239
    from IH1 IH2[of "a@compile c1"]
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   240
    show "?P (c1;c2) s s'' a z"
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   241
      by(simp add:add_assoc[THEN sym])(blast intro:rtrancl_trans)
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   242
  next
7275750553b7 a bit of conversion to structured proofs
nipkow
parents: 13095
diff changeset
   243
(* at this point I gave up converting to structured proofs *)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   244
(* \<IF> b \<THEN> c0 \<ELSE> c1; case b is true *)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   245
   apply(intro strip)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   246
   (* instantiate assumption sufficiently for later: *)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   247
   apply(erule_tac x = "a@[?I]" in allE)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   248
   apply(simp)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   249
   (* execute JMPF: *)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   250
   apply(rule converse_rtrancl_into_rtrancl)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   251
    apply(force intro!: JMPFT)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   252
   (* execute compile c0: *)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   253
   apply(rule rtrancl_trans)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   254
    apply(erule allE)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   255
    apply assumption
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   256
   (* execute JMPF: *)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   257
   apply(rule r_into_rtrancl)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   258
   apply(force intro!: JMPFF)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   259
(* end of case b is true *)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   260
  apply(intro strip)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   261
  apply(erule_tac x = "a@[?I]@compile c0@[?J]" in allE)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   262
  apply(simp add:add_assoc)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   263
  apply(rule converse_rtrancl_into_rtrancl)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   264
   apply(force intro!: JMPFF)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   265
  apply(blast)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   266
 apply(force intro: JMPFF)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   267
apply(intro strip)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   268
apply(erule_tac x = "a@[?I]" in allE)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   269
apply(erule_tac x = a in allE)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   270
apply(simp)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   271
apply(rule converse_rtrancl_into_rtrancl)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   272
 apply(force intro!: JMPFT)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   273
apply(rule rtrancl_trans)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   274
 apply(erule allE)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   275
 apply assumption
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   276
apply(rule converse_rtrancl_into_rtrancl)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   277
 apply(force intro!: JMPB)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   278
apply(simp)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   279
done
13130
423ce375bf65 commented out half converted proof
nipkow
parents: 13112
diff changeset
   280
*)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   281
text {* Missing: the other direction! I did much of it, and although
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   282
the main lemma is very similar to the one in the new development, the
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   283
lemmas surrounding it seemed much more complicated. In the end I gave
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   284
up. *}
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   285
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   286
end