| author | haftmann |
| Sun, 26 Apr 2009 08:34:53 +0200 | |
| changeset 30988 | b53800e3ee47 |
| parent 30952 | 7ab2716dd93b |
| child 31969 | 09524788a6b9 |
| permissions | -rw-r--r-- |
|
30952
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents:
23746
diff
changeset
|
1 |
theory Machines |
|
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents:
23746
diff
changeset
|
2 |
imports Natural |
|
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents:
23746
diff
changeset
|
3 |
begin |
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
4 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
5 |
lemma rtrancl_eq: "R^* = Id \<union> (R O R^*)" |
|
22267
ea31e6ea0e2e
Adapted to changes in Transitive_Closure theory.
berghofe
parents:
20503
diff
changeset
|
6 |
by (fast intro: rtrancl_into_rtrancl elim: rtranclE) |
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
7 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
8 |
lemma converse_rtrancl_eq: "R^* = Id \<union> (R^* O R)" |
| 18372 | 9 |
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
|
10 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
11 |
lemmas converse_rel_powE = rel_pow_E2 |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
12 |
|
|
30952
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents:
23746
diff
changeset
|
13 |
lemma R_O_Rn_commute: "R O R ^^ n = R ^^ n O R" |
| 18372 | 14 |
by (induct n) (simp, simp add: O_assoc [symmetric]) |
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
15 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
16 |
lemma converse_in_rel_pow_eq: |
|
30952
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents:
23746
diff
changeset
|
17 |
"((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
|
18 |
apply(rule iffI) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
19 |
apply(blast elim:converse_rel_powE) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
20 |
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
|
21 |
done |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
22 |
|
|
30952
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents:
23746
diff
changeset
|
23 |
lemma rel_pow_plus: |
|
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents:
23746
diff
changeset
|
24 |
"R ^^ (m+n) = R ^^ n O R ^^ m" |
| 18372 | 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 |
|
|
30952
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents:
23746
diff
changeset
|
27 |
lemma rel_pow_plusI: |
|
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents:
23746
diff
changeset
|
28 |
"\<lbrakk> (x,y) \<in> R ^^ m; (y,z) \<in> R ^^ n \<rbrakk> \<Longrightarrow> (x,z) \<in> R ^^ (m+n)" |
| 18372 | 29 |
by (simp add: rel_pow_plus rel_compI) |
|
13095
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 |
subsection "Instructions" |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
32 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
33 |
text {* There are only three instructions: *}
|
| 13675 | 34 |
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
|
35 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
36 |
types instrs = "instr list" |
|
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 |
subsection "M0 with PC" |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
39 |
|
| 23746 | 40 |
inductive_set |
41 |
exec01 :: "instr list \<Rightarrow> ((nat\<times>state) \<times> (nat\<times>state))set" |
|
42 |
and exec01' :: "[instrs, nat,state, nat,state] \<Rightarrow> bool" |
|
43 |
("(_/ \<turnstile> (1\<langle>_,/_\<rangle>)/ -1\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
|
|
44 |
for P :: "instr list" |
|
45 |
where |
|
46 |
"p \<turnstile> \<langle>i,s\<rangle> -1\<rightarrow> \<langle>j,t\<rangle> == ((i,s),j,t) : (exec01 p)" |
|
47 |
| 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>" |
|
48 |
| 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>" |
|
49 |
| JMPFF: "\<lbrakk> n<size P; P!n = JMPF b i; \<not>b s; m=n+i+1; m \<le> size P \<rbrakk> |
|
50 |
\<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>m,s\<rangle>" |
|
51 |
| 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>" |
|
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
52 |
|
| 23746 | 53 |
abbreviation |
54 |
exec0s :: "[instrs, nat,state, nat,state] \<Rightarrow> bool" |
|
55 |
("(_/ \<turnstile> (1\<langle>_,/_\<rangle>)/ -*\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50) where
|
|
56 |
"p \<turnstile> \<langle>i,s\<rangle> -*\<rightarrow> \<langle>j,t\<rangle> == ((i,s),j,t) : (exec01 p)^*" |
|
| 14565 | 57 |
|
| 23746 | 58 |
abbreviation |
59 |
exec0n :: "[instrs, nat,state, nat, nat,state] \<Rightarrow> bool" |
|
60 |
("(_/ \<turnstile> (1\<langle>_,/_\<rangle>)/ -_\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50) where
|
|
|
30952
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents:
23746
diff
changeset
|
61 |
"p \<turnstile> \<langle>i,s\<rangle> -n\<rightarrow> \<langle>j,t\<rangle> == ((i,s),j,t) : (exec01 p)^^n" |
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
62 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
63 |
subsection "M0 with lists" |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
64 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
65 |
text {* We describe execution of programs in the machine by
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
66 |
an operational (small step) semantics: |
|
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 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
69 |
types config = "instrs \<times> instrs \<times> state" |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
70 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
71 |
|
| 23746 | 72 |
inductive_set |
73 |
stepa1 :: "(config \<times> config)set" |
|
74 |
and stepa1' :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool" |
|
75 |
("((1\<langle>_,/_,/_\<rangle>)/ -1\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50)
|
|
76 |
where |
|
77 |
"\<langle>p,q,s\<rangle> -1\<rightarrow> \<langle>p',q',t\<rangle> == ((p,q,s),p',q',t) : stepa1" |
|
78 |
| "\<langle>SET x a#p,q,s\<rangle> -1\<rightarrow> \<langle>p,SET x a#q,s[x\<mapsto> a s]\<rangle>" |
|
79 |
| "b s \<Longrightarrow> \<langle>JMPF b i#p,q,s\<rangle> -1\<rightarrow> \<langle>p,JMPF b i#q,s\<rangle>" |
|
80 |
| "\<lbrakk> \<not> b s; i \<le> size p \<rbrakk> |
|
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
81 |
\<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>" |
| 23746 | 82 |
| "i \<le> size q |
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
83 |
\<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
|
84 |
|
| 23746 | 85 |
abbreviation |
86 |
stepa :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool" |
|
87 |
("((1\<langle>_,/_,/_\<rangle>)/ -*\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50) where
|
|
88 |
"\<langle>p,q,s\<rangle> -*\<rightarrow> \<langle>p',q',t\<rangle> == ((p,q,s),p',q',t) : (stepa1^*)" |
|
89 |
||
90 |
abbreviation |
|
91 |
stepan :: "[instrs,instrs,state, nat, instrs,instrs,state] \<Rightarrow> bool" |
|
92 |
("((1\<langle>_,/_,/_\<rangle>)/ -_\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50) where
|
|
|
30952
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents:
23746
diff
changeset
|
93 |
"\<langle>p,q,s\<rangle> -i\<rightarrow> \<langle>p',q',t\<rangle> == ((p,q,s),p',q',t) : (stepa1^^i)" |
| 23746 | 94 |
|
95 |
inductive_cases execE: "((i#is,p,s), (is',p',s')) : stepa1" |
|
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
96 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
97 |
lemma exec_simp[simp]: |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
98 |
"(\<langle>i#p,q,s\<rangle> -1\<rightarrow> \<langle>p',q',t\<rangle>) = (case i of |
| 13675 | 99 |
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
|
100 |
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
|
101 |
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
|
102 |
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
|
103 |
apply(rule iffI) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
104 |
defer |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
105 |
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
|
106 |
apply(erule execE) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
107 |
apply(simp_all) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
108 |
done |
|
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 execn_simp[simp]: |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
111 |
"(\<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
|
112 |
(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
|
113 |
((\<exists>m p' q' t. n = Suc m \<and> |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
114 |
\<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
|
115 |
by(subst converse_in_rel_pow_eq, simp) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
116 |
|
|
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 |
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
|
119 |
(p'' = i#p & q''=q & u=s | |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
120 |
(\<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
|
121 |
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
|
122 |
apply(blast) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
123 |
done |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
124 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
125 |
declare nth_append[simp] |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
126 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
127 |
lemma rev_revD: "rev xs = rev ys \<Longrightarrow> xs = ys" |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
128 |
by 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 |
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
|
131 |
apply(rule iffI) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
132 |
apply(rule rev_revD, simp) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
133 |
apply fastsimp |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
134 |
done |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
135 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
136 |
lemma direction1: |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
137 |
"\<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
|
138 |
rev p' @ q' = rev p @ q \<and> rev p @ q \<turnstile> \<langle>size p,s\<rangle> -1\<rightarrow> \<langle>size p',t\<rangle>" |
| 18372 | 139 |
apply(induct set: stepa1) |
| 13675 | 140 |
apply(simp add:exec01.SET) |
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
141 |
apply(fastsimp intro:exec01.JMPFT) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
142 |
apply simp |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
143 |
apply(rule exec01.JMPFF) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
144 |
apply simp |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
145 |
apply fastsimp |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
146 |
apply simp |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
147 |
apply simp |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
148 |
apply simp |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
149 |
apply(fastsimp simp add:exec01.JMPB) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
150 |
done |
|
20217
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents:
18372
diff
changeset
|
151 |
|
| 13098 | 152 |
(* |
153 |
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
|
154 |
apply(induct xs) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
155 |
apply simp_all |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
156 |
apply(case_tac i) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
157 |
apply simp_all |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
158 |
done |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
159 |
|
| 13098 | 160 |
lemma rev_drop: "\<And>i. rev (drop i xs) = take (length xs - i) (rev xs)" |
161 |
apply(induct xs) |
|
162 |
apply simp_all |
|
163 |
apply(case_tac i) |
|
164 |
apply simp_all |
|
165 |
done |
|
166 |
*) |
|
|
20217
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents:
18372
diff
changeset
|
167 |
|
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
168 |
lemma direction2: |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
169 |
"rpq \<turnstile> \<langle>sp,s\<rangle> -1\<rightarrow> \<langle>sp',t\<rangle> \<Longrightarrow> |
| 18372 | 170 |
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
|
171 |
rev p' @ q' = rev p @ q \<longrightarrow> \<langle>q,p,s\<rangle> -1\<rightarrow> \<langle>q',p',t\<rangle>" |
| 20503 | 172 |
apply(induct arbitrary: p q p' q' set: exec01) |
| 13098 | 173 |
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
|
174 |
apply(drule sym) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
175 |
apply simp |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
176 |
apply(rule rev_revD) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
177 |
apply simp |
| 13098 | 178 |
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
|
179 |
apply(drule sym) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
180 |
apply simp |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
181 |
apply(rule rev_revD) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
182 |
apply simp |
| 13612 | 183 |
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
|
184 |
apply(drule sym) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
185 |
apply simp |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
186 |
apply(rule rev_revD) |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
187 |
apply simp |
| 13098 | 188 |
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
|
189 |
apply(drule sym) |
| 13098 | 190 |
apply(simp add:rev_take) |
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
191 |
apply(rule rev_revD) |
| 13098 | 192 |
apply(simp add:rev_drop) |
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
193 |
done |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
194 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
195 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
196 |
theorem M_eqiv: |
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
197 |
"(\<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
|
198 |
(rev p' @ q' = rev p @ q \<and> rev p @ q \<turnstile> \<langle>size p,s\<rangle> -1\<rightarrow> \<langle>size p',t\<rangle>)" |
| 18372 | 199 |
by (blast dest: direction1 direction2) |
|
13095
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
200 |
|
|
8ed413a57bdc
New machine architecture and other direction of compiler proof.
nipkow
parents:
diff
changeset
|
201 |
end |