author | paulson |
Wed, 26 Jun 2002 10:26:46 +0200 | |
changeset 13248 | ae66c22ed52e |
parent 13098 | e0644528e21e |
child 13612 | 55d32e76ef4e |
permissions | -rw-r--r-- |
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
1 |
theory Machines = Natural: |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
2 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
3 |
lemma rtrancl_eq: "R^* = Id \<union> (R O R^*)" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
4 |
by(fast intro:rtrancl.intros elim:rtranclE) |
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 converse_rtrancl_eq: "R^* = Id \<union> (R^* O R)" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
7 |
by(subst r_comp_rtrancl_eq[symmetric], rule rtrancl_eq) |
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 |
lemmas converse_rel_powE = rel_pow_E2 |
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 |
lemma R_O_Rn_commute: "R O R^n = R^n O R" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
12 |
by(induct_tac n, simp, simp add: O_assoc[symmetric]) |
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 converse_in_rel_pow_eq: |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
15 |
"((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))" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
16 |
apply(rule iffI) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
17 |
apply(blast elim:converse_rel_powE) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
18 |
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
|
19 |
done |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
20 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
21 |
lemma rel_pow_plus: "R^(m+n) = R^n O R^m" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
22 |
by(induct n, simp, simp add:O_assoc) |
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_plusI: "\<lbrakk> (x,y) \<in> R^m; (y,z) \<in> R^n \<rbrakk> \<Longrightarrow> (x,z) \<in> R^(m+n)" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
25 |
by(simp add:rel_pow_plus rel_compI) |
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 |
subsection "Instructions" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
28 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
29 |
text {* There are only three instructions: *} |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
30 |
datatype instr = ASIN loc aexp | JMPF bexp nat | JMPB nat |
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 |
types instrs = "instr list" |
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 |
subsection "M0 with PC" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
35 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
36 |
consts exec01 :: "instr list \<Rightarrow> ((nat\<times>state) \<times> (nat\<times>state))set" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
37 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
38 |
syntax |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
39 |
"_exec01" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
40 |
("(_/ |- (1<_,/_>)/ -1-> (1<_,/_>))" [50,0,0,0,0] 50) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
41 |
"_exec0s" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
42 |
("(_/ |- (1<_,/_>)/ -*-> (1<_,/_>))" [50,0,0,0,0] 50) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
43 |
"_exec0n" :: "[instrs, nat,state, nat, 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 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
46 |
syntax (xsymbols) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
47 |
"_exec01" :: "[instrs, nat,state, nat,state] \<Rightarrow> bool" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
48 |
("(_/ \<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
|
49 |
"_exec0s" :: "[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>)/ -*\<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 |
"_exec0n" :: "[instrs, nat,state, nat, 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 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
54 |
translations |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
55 |
"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
|
56 |
"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
|
57 |
"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
|
58 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
59 |
inductive "exec01 P" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
60 |
intros |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
61 |
ASIN: "\<lbrakk> n<size P; P!n = ASIN x a \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>Suc n,s[x\<mapsto> a s]\<rangle>" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
62 |
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
|
63 |
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
|
64 |
\<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
|
65 |
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
|
66 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
67 |
subsection "M0 with lists" |
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 |
text {* We describe execution of programs in the machine by |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
70 |
an operational (small step) semantics: |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
71 |
*} |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
72 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
73 |
types config = "instrs \<times> instrs \<times> state" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
74 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
75 |
consts stepa1 :: "(config \<times> config)set" |
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 |
syntax |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
78 |
"_stepa1" :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
79 |
("((1<_,/_,/_>)/ -1-> (1<_,/_,/_>))" 50) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
80 |
"_stepa" :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
81 |
("((1<_,/_,/_>)/ -*-> (1<_,/_,/_>))" 50) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
82 |
"_stepan" :: "[state,instrs,instrs, nat, instrs,instrs,state] \<Rightarrow> bool" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
83 |
("((1<_,/_,/_>)/ -_-> (1<_,/_,/_>))" 50) |
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 |
syntax (xsymbols) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
86 |
"_stepa1" :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
87 |
("((1\<langle>_,/_,/_\<rangle>)/ -1\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
88 |
"_stepa" :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
89 |
("((1\<langle>_,/_,/_\<rangle>)/ -*\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
90 |
"_stepan" :: "[instrs,instrs,state, nat, instrs,instrs,state] \<Rightarrow> bool" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
91 |
("((1\<langle>_,/_,/_\<rangle>)/ -_\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
92 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
93 |
translations |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
94 |
"\<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
|
95 |
"\<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
|
96 |
"\<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
|
97 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
98 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
99 |
inductive "stepa1" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
100 |
intros |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
101 |
"\<langle>ASIN x a#p,q,s\<rangle> -1\<rightarrow> \<langle>p,ASIN x a#q,s[x\<mapsto> a s]\<rangle>" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
102 |
"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
|
103 |
"\<lbrakk> \<not> b s; i \<le> size p \<rbrakk> |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
104 |
\<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
|
105 |
"i \<le> size q |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
106 |
\<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
|
107 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
108 |
inductive_cases execE: "((i#is,p,s),next) : stepa1" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
109 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
110 |
lemma exec_simp[simp]: |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
111 |
"(\<langle>i#p,q,s\<rangle> -1\<rightarrow> \<langle>p',q',t\<rangle>) = (case i of |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
112 |
ASIN x a \<Rightarrow> t = s[x\<mapsto> a s] \<and> p' = p \<and> q' = i#q | |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
113 |
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
|
114 |
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
|
115 |
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
|
116 |
apply(rule iffI) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
117 |
defer |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
118 |
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
|
119 |
apply(erule execE) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
120 |
apply(simp_all) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
121 |
done |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
122 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
123 |
lemma execn_simp[simp]: |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
124 |
"(\<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
|
125 |
(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
|
126 |
((\<exists>m p' q' t. n = Suc m \<and> |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
127 |
\<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
|
128 |
by(subst converse_in_rel_pow_eq, simp) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
129 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
130 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
131 |
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
|
132 |
(p'' = i#p & q''=q & u=s | |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
133 |
(\<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
|
134 |
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
|
135 |
apply(blast) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
136 |
done |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
137 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
138 |
declare nth_append[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 |
lemma rev_revD: "rev xs = rev ys \<Longrightarrow> xs = ys" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
141 |
by simp |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
142 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
143 |
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
|
144 |
apply(rule iffI) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
145 |
apply(rule rev_revD, simp) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
146 |
apply fastsimp |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
147 |
done |
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 |
lemma direction1: |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
150 |
"\<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
|
151 |
rev p' @ q' = rev p @ q \<and> rev p @ q \<turnstile> \<langle>size p,s\<rangle> -1\<rightarrow> \<langle>size p',t\<rangle>" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
152 |
apply(erule stepa1.induct) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
153 |
apply(simp add:exec01.ASIN) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
154 |
apply(fastsimp intro:exec01.JMPFT) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
155 |
apply simp |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
156 |
apply(rule exec01.JMPFF) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
157 |
apply simp |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
158 |
apply fastsimp |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
159 |
apply simp |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
160 |
apply simp |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
161 |
apply arith |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
162 |
apply simp |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
163 |
apply arith |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
164 |
apply(fastsimp simp add:exec01.JMPB) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
165 |
done |
13098 | 166 |
(* |
167 |
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
|
168 |
apply(induct xs) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
169 |
apply simp_all |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
170 |
apply(case_tac i) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
171 |
apply simp_all |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
172 |
done |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
173 |
|
13098 | 174 |
lemma rev_drop: "\<And>i. rev (drop i xs) = take (length xs - i) (rev xs)" |
175 |
apply(induct xs) |
|
176 |
apply simp_all |
|
177 |
apply(case_tac i) |
|
178 |
apply simp_all |
|
179 |
done |
|
180 |
*) |
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
181 |
lemma direction2: |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
182 |
"rpq \<turnstile> \<langle>sp,s\<rangle> -1\<rightarrow> \<langle>sp',t\<rangle> \<Longrightarrow> |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
183 |
\<forall>p q p' q'. rpq = rev p @ q & sp = size p & sp' = size p' \<longrightarrow> |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
184 |
rev p' @ q' = rev p @ q \<longrightarrow> \<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
|
185 |
apply(erule exec01.induct) |
13098 | 186 |
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
|
187 |
apply(drule sym) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
188 |
apply simp |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
189 |
apply(rule rev_revD) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
190 |
apply simp |
13098 | 191 |
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
|
192 |
apply(drule sym) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
193 |
apply simp |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
194 |
apply(rule rev_revD) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
195 |
apply simp |
13098 | 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 | 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) |
13098 | 203 |
apply(simp add:rev_take) |
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
204 |
apply(rule rev_revD) |
13098 | 205 |
apply(simp add:rev_drop) |
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
206 |
done |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
207 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
208 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
209 |
theorem M_eqiv: |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
210 |
"(\<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
|
211 |
(rev p' @ q' = rev p @ q \<and> rev p @ q \<turnstile> \<langle>size p,s\<rangle> -1\<rightarrow> \<langle>size p',t\<rangle>)" |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
212 |
by(fast dest:direction1 direction2) |
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
213 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
214 |
end |