src/HOL/IMP/Machines.thy
author ballarin
Wed, 19 Jul 2006 19:25:58 +0200
changeset 20168 ed7bced29e1b
parent 18372 2bffdf62fe7f
child 20217 25b068a99d2b
permissions -rw-r--r--
Reimplemented algebra method; now controlled by attribute.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
     1
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
     2
(* $Id$ *)
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
     3
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14565
diff changeset
     4
theory Machines imports Natural begin
13095
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
lemma rtrancl_eq: "R^* = Id \<union> (R O R^*)"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
     7
  by (fast intro: rtrancl.intros elim: rtranclE)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
     8
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
     9
lemma converse_rtrancl_eq: "R^* = Id \<union> (R^* O R)"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    10
  by (subst r_comp_rtrancl_eq[symmetric], rule rtrancl_eq)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    11
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    12
lemmas converse_rel_powE = rel_pow_E2
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    13
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    14
lemma R_O_Rn_commute: "R O R^n = R^n O R"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    15
  by (induct n) (simp, simp add: O_assoc [symmetric])
13095
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
lemma converse_in_rel_pow_eq:
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    18
  "((x,z) \<in> R^n) = (n=0 \<and> z=x \<or> (\<exists>m y. n = Suc m \<and> (x,y) \<in> R \<and> (y,z) \<in> R^m))"
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    19
apply(rule iffI)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    20
 apply(blast elim:converse_rel_powE)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    21
apply (fastsimp simp add:gr0_conv_Suc R_O_Rn_commute)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    22
done
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    23
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    24
lemma rel_pow_plus: "R^(m+n) = R^n O R^m"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    25
  by (induct n) (simp, simp add: O_assoc)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    26
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    27
lemma rel_pow_plusI: "\<lbrakk> (x,y) \<in> R^m; (y,z) \<in> R^n \<rbrakk> \<Longrightarrow> (x,z) \<in> R^(m+n)"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    28
  by (simp add: rel_pow_plus rel_compI)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    29
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    30
subsection "Instructions"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    31
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    32
text {* There are only three instructions: *}
13675
01fc1fc61384 ASIN -> SET
nipkow
parents: 13612
diff changeset
    33
datatype instr = SET loc aexp | JMPF bexp nat | JMPB nat
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    34
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    35
types instrs = "instr list"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    36
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    37
subsection "M0 with PC"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    38
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    39
consts  exec01 :: "instr list \<Rightarrow> ((nat\<times>state) \<times> (nat\<times>state))set"
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    40
syntax
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    41
  "_exec01" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    42
               ("(_/ |- (1<_,/_>)/ -1-> (1<_,/_>))" [50,0,0,0,0] 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    43
  "_exec0s" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    44
               ("(_/ |- (1<_,/_>)/ -*-> (1<_,/_>))" [50,0,0,0,0] 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    45
  "_exec0n" :: "[instrs, nat,state, nat, nat,state] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    46
               ("(_/ |- (1<_,/_>)/ -_-> (1<_,/_>))" [50,0,0,0,0] 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    47
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    48
syntax (xsymbols)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    49
  "_exec01" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    50
               ("(_/ \<turnstile> (1\<langle>_,/_\<rangle>)/ -1\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    51
  "_exec0s" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    52
               ("(_/ \<turnstile> (1\<langle>_,/_\<rangle>)/ -*\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    53
  "_exec0n" :: "[instrs, nat,state, nat, nat,state] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    54
               ("(_/ \<turnstile> (1\<langle>_,/_\<rangle>)/ -_\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    55
14565
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13675
diff changeset
    56
syntax (HTML output)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13675
diff changeset
    57
  "_exec01" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13675
diff changeset
    58
               ("(_/ |- (1\<langle>_,/_\<rangle>)/ -1\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13675
diff changeset
    59
  "_exec0s" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13675
diff changeset
    60
               ("(_/ |- (1\<langle>_,/_\<rangle>)/ -*\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13675
diff changeset
    61
  "_exec0n" :: "[instrs, nat,state, nat, nat,state] \<Rightarrow> bool"
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13675
diff changeset
    62
               ("(_/ |- (1\<langle>_,/_\<rangle>)/ -_\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
c6dc17aab88a use more symbols in HTML output
kleing
parents: 13675
diff changeset
    63
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    64
translations
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    65
  "p \<turnstile> \<langle>i,s\<rangle> -1\<rightarrow> \<langle>j,t\<rangle>" == "((i,s),j,t) : (exec01 p)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    66
  "p \<turnstile> \<langle>i,s\<rangle> -*\<rightarrow> \<langle>j,t\<rangle>" == "((i,s),j,t) : (exec01 p)^*"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    67
  "p \<turnstile> \<langle>i,s\<rangle> -n\<rightarrow> \<langle>j,t\<rangle>" == "((i,s),j,t) : (exec01 p)^n"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    68
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    69
inductive "exec01 P"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    70
intros
13675
01fc1fc61384 ASIN -> SET
nipkow
parents: 13612
diff changeset
    71
SET: "\<lbrakk> n<size P; P!n = SET x a \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>Suc n,s[x\<mapsto> a s]\<rangle>"
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    72
JMPFT: "\<lbrakk> n<size P; P!n = JMPF b i;  b s \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>Suc n,s\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    73
JMPFF: "\<lbrakk> n<size P; P!n = JMPF b i; \<not>b s; m=n+i+1; m \<le> size P \<rbrakk>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    74
        \<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>m,s\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    75
JMPB:  "\<lbrakk> n<size P; P!n = JMPB i; i \<le> n; j = n-i \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>j,s\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    76
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    77
subsection "M0 with lists"
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
text {* We describe execution of programs in the machine by
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    80
  an operational (small step) semantics:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    81
*}
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
types config = "instrs \<times> instrs \<times> state"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    84
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    85
consts  stepa1 :: "(config \<times> config)set"
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
syntax
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    88
  "_stepa1" :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    89
               ("((1<_,/_,/_>)/ -1-> (1<_,/_,/_>))" 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    90
  "_stepa" :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    91
               ("((1<_,/_,/_>)/ -*-> (1<_,/_,/_>))" 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    92
  "_stepan" :: "[state,instrs,instrs, nat, instrs,instrs,state] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    93
               ("((1<_,/_,/_>)/ -_-> (1<_,/_,/_>))" 50)
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
syntax (xsymbols)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    96
  "_stepa1" :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    97
               ("((1\<langle>_,/_,/_\<rangle>)/ -1\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    98
  "_stepa" :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
    99
               ("((1\<langle>_,/_,/_\<rangle>)/ -*\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   100
  "_stepan" :: "[instrs,instrs,state, nat, instrs,instrs,state] \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   101
               ("((1\<langle>_,/_,/_\<rangle>)/ -_\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   102
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
   103
translations
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   104
  "\<langle>p,q,s\<rangle> -1\<rightarrow> \<langle>p',q',t\<rangle>" == "((p,q,s),p',q',t) : stepa1"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   105
  "\<langle>p,q,s\<rangle> -*\<rightarrow> \<langle>p',q',t\<rangle>" == "((p,q,s),p',q',t) : (stepa1^*)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   106
  "\<langle>p,q,s\<rangle> -i\<rightarrow> \<langle>p',q',t\<rangle>" == "((p,q,s),p',q',t) : (stepa1^i)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   107
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   108
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   109
inductive "stepa1"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   110
intros
13675
01fc1fc61384 ASIN -> SET
nipkow
parents: 13612
diff changeset
   111
  "\<langle>SET x a#p,q,s\<rangle> -1\<rightarrow> \<langle>p,SET x a#q,s[x\<mapsto> a s]\<rangle>"
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   112
  "b s \<Longrightarrow> \<langle>JMPF b i#p,q,s\<rangle> -1\<rightarrow> \<langle>p,JMPF b i#q,s\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   113
  "\<lbrakk> \<not> b s; i \<le> size p \<rbrakk>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   114
   \<Longrightarrow> \<langle>JMPF b i # p, q, s\<rangle> -1\<rightarrow> \<langle>drop i p, rev(take i p) @ JMPF b i # q, s\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   115
  "i \<le> size q
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   116
   \<Longrightarrow> \<langle>JMPB i # p, q, s\<rangle> -1\<rightarrow> \<langle>rev(take i q) @ JMPB i # p, drop i q, s\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   117
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   118
inductive_cases execE: "((i#is,p,s),next) : stepa1"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   119
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   120
lemma exec_simp[simp]:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   121
 "(\<langle>i#p,q,s\<rangle> -1\<rightarrow> \<langle>p',q',t\<rangle>) = (case i of
13675
01fc1fc61384 ASIN -> SET
nipkow
parents: 13612
diff changeset
   122
 SET x a \<Rightarrow> t = s[x\<mapsto> a s] \<and> p' = p \<and> q' = i#q |
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   123
 JMPF b n \<Rightarrow> t=s \<and> (if b s then p' = p \<and> q' = i#q
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   124
            else n \<le> size p \<and> p' = drop n p \<and> q' = rev(take n p) @ i # q) |
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   125
 JMPB n \<Rightarrow> n \<le> size q \<and> t=s \<and> p' = rev(take n q) @ i # p \<and> q' = drop n q)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   126
apply(rule iffI)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   127
defer
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   128
apply(clarsimp simp add: stepa1.intros split: instr.split_asm split_if_asm)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   129
apply(erule execE)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   130
apply(simp_all)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   131
done
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   132
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   133
lemma execn_simp[simp]:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   134
"(\<langle>i#p,q,s\<rangle> -n\<rightarrow> \<langle>p'',q'',u\<rangle>) =
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   135
 (n=0 \<and> p'' = i#p \<and> q'' = q \<and> u = s \<or>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   136
  ((\<exists>m p' q' t. n = Suc m \<and>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   137
                \<langle>i#p,q,s\<rangle> -1\<rightarrow> \<langle>p',q',t\<rangle> \<and> \<langle>p',q',t\<rangle> -m\<rightarrow> \<langle>p'',q'',u\<rangle>)))"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   138
by(subst converse_in_rel_pow_eq, simp)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   139
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   140
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   141
lemma exec_star_simp[simp]: "(\<langle>i#p,q,s\<rangle> -*\<rightarrow> \<langle>p'',q'',u\<rangle>) =
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   142
 (p'' = i#p & q''=q & u=s |
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   143
 (\<exists>p' q' t. \<langle>i#p,q,s\<rangle> -1\<rightarrow> \<langle>p',q',t\<rangle> \<and> \<langle>p',q',t\<rangle> -*\<rightarrow> \<langle>p'',q'',u\<rangle>))"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   144
apply(simp add: rtrancl_is_UN_rel_pow del:exec_simp)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   145
apply(blast)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   146
done
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   147
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   148
declare nth_append[simp]
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   149
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   150
lemma rev_revD: "rev xs = rev ys \<Longrightarrow> xs = ys"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   151
by simp
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
lemma [simp]: "(rev xs @ rev ys = rev zs) = (ys @ xs = zs)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   154
apply(rule iffI)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   155
 apply(rule rev_revD, simp)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   156
apply fastsimp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   157
done
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   158
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   159
lemma direction1:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   160
 "\<langle>q,p,s\<rangle> -1\<rightarrow> \<langle>q',p',t\<rangle> \<Longrightarrow>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   161
  rev p' @ q' = rev p @ q \<and> rev p @ q \<turnstile> \<langle>size p,s\<rangle> -1\<rightarrow> \<langle>size p',t\<rangle>"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
   162
apply(induct set: stepa1)
13675
01fc1fc61384 ASIN -> SET
nipkow
parents: 13612
diff changeset
   163
   apply(simp add:exec01.SET)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   164
  apply(fastsimp intro:exec01.JMPFT)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   165
 apply simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   166
 apply(rule exec01.JMPFF)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   167
     apply simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   168
    apply fastsimp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   169
   apply simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   170
  apply simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   171
  apply arith
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   172
 apply simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   173
 apply arith
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   174
apply(fastsimp simp add:exec01.JMPB)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   175
done
13098
e0644528e21e Better compiler proof
nipkow
parents: 13095
diff changeset
   176
(*
e0644528e21e Better compiler proof
nipkow
parents: 13095
diff changeset
   177
lemma rev_take: "\<And>i. rev (take i xs) = drop (length xs - i) (rev xs)"
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   178
apply(induct xs)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   179
 apply simp_all
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   180
apply(case_tac i)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   181
apply simp_all
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   182
done
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   183
13098
e0644528e21e Better compiler proof
nipkow
parents: 13095
diff changeset
   184
lemma rev_drop: "\<And>i. rev (drop i xs) = take (length xs - i) (rev xs)"
e0644528e21e Better compiler proof
nipkow
parents: 13095
diff changeset
   185
apply(induct xs)
e0644528e21e Better compiler proof
nipkow
parents: 13095
diff changeset
   186
 apply simp_all
e0644528e21e Better compiler proof
nipkow
parents: 13095
diff changeset
   187
apply(case_tac i)
e0644528e21e Better compiler proof
nipkow
parents: 13095
diff changeset
   188
apply simp_all
e0644528e21e Better compiler proof
nipkow
parents: 13095
diff changeset
   189
done
e0644528e21e Better compiler proof
nipkow
parents: 13095
diff changeset
   190
*)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   191
lemma direction2:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   192
 "rpq \<turnstile> \<langle>sp,s\<rangle> -1\<rightarrow> \<langle>sp',t\<rangle> \<Longrightarrow>
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
   193
  rpq = rev p @ q & sp = size p & sp' = size p' \<longrightarrow>
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   194
          rev p' @ q' = rev p @ q \<longrightarrow> \<langle>q,p,s\<rangle> -1\<rightarrow> \<langle>q',p',t\<rangle>"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
   195
apply(induct fixing: p q p' q' set: exec01)
13098
e0644528e21e Better compiler proof
nipkow
parents: 13095
diff changeset
   196
   apply(clarsimp simp add: neq_Nil_conv append_eq_conv_conj)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   197
   apply(drule sym)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   198
   apply simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   199
   apply(rule rev_revD)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   200
   apply simp
13098
e0644528e21e Better compiler proof
nipkow
parents: 13095
diff changeset
   201
  apply(clarsimp simp add: neq_Nil_conv append_eq_conv_conj)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   202
  apply(drule sym)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   203
  apply simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   204
  apply(rule rev_revD)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   205
  apply simp
13612
55d32e76ef4e Adapted to new simplifier.
berghofe
parents: 13098
diff changeset
   206
 apply(simp (no_asm_use) add: neq_Nil_conv append_eq_conv_conj, clarify)+
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   207
 apply(drule sym)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   208
 apply simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   209
 apply(rule rev_revD)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   210
 apply simp
13098
e0644528e21e Better compiler proof
nipkow
parents: 13095
diff changeset
   211
apply(clarsimp simp add: neq_Nil_conv append_eq_conv_conj)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   212
apply(drule sym)
13098
e0644528e21e Better compiler proof
nipkow
parents: 13095
diff changeset
   213
apply(simp add:rev_take)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   214
apply(rule rev_revD)
13098
e0644528e21e Better compiler proof
nipkow
parents: 13095
diff changeset
   215
apply(simp add:rev_drop)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   216
done
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   217
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   218
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   219
theorem M_eqiv:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   220
"(\<langle>q,p,s\<rangle> -1\<rightarrow> \<langle>q',p',t\<rangle>) =
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   221
 (rev p' @ q' = rev p @ q \<and> rev p @ q \<turnstile> \<langle>size p,s\<rangle> -1\<rightarrow> \<langle>size p',t\<rangle>)"
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
   222
  by (blast dest: direction1 direction2)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   223
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents:
diff changeset
   224
end