| author | wenzelm | 
| Fri, 17 Dec 2010 18:10:37 +0100 | |
| changeset 41249 | 26f12f98f50a | 
| parent 32235 | 8f9b8d14fc9f | 
| child 42174 | d0be2722ce9f | 
| 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 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
 | 
6  | 
"((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
 | 
7  | 
apply(rule iffI)  | 
| 31969 | 8  | 
apply(blast elim:rel_pow_E2)  | 
| 
32235
 
8f9b8d14fc9f
"more standard" argument order of relation composition (op O)
 
krauss 
parents: 
31970 
diff
changeset
 | 
9  | 
apply (auto simp: rel_pow_commute[symmetric])  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
10  | 
done  | 
| 
 
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  | 
subsection "Instructions"  | 
| 
 
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  | 
text {* There are only three instructions: *}
 | 
| 13675 | 15  | 
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
 | 
16  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
17  | 
types instrs = "instr list"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
19  | 
subsection "M0 with PC"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
20  | 
|
| 23746 | 21  | 
inductive_set  | 
22  | 
exec01 :: "instr list \<Rightarrow> ((nat\<times>state) \<times> (nat\<times>state))set"  | 
|
23  | 
and exec01' :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"  | 
|
24  | 
    ("(_/ \<turnstile> (1\<langle>_,/_\<rangle>)/ -1\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)
 | 
|
25  | 
for P :: "instr list"  | 
|
26  | 
where  | 
|
27  | 
"p \<turnstile> \<langle>i,s\<rangle> -1\<rightarrow> \<langle>j,t\<rangle> == ((i,s),j,t) : (exec01 p)"  | 
|
28  | 
| 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>"  | 
|
29  | 
| 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>"  | 
|
30  | 
| JMPFF: "\<lbrakk> n<size P; P!n = JMPF b i; \<not>b s; m=n+i+1; m \<le> size P \<rbrakk>  | 
|
31  | 
\<Longrightarrow> P \<turnstile> \<langle>n,s\<rangle> -1\<rightarrow> \<langle>m,s\<rangle>"  | 
|
32  | 
| 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
 | 
33  | 
|
| 23746 | 34  | 
abbreviation  | 
35  | 
exec0s :: "[instrs, nat,state, nat,state] \<Rightarrow> bool"  | 
|
36  | 
    ("(_/ \<turnstile> (1\<langle>_,/_\<rangle>)/ -*\<rightarrow> (1\<langle>_,/_\<rangle>))" [50,0,0,0,0] 50)  where
 | 
|
37  | 
"p \<turnstile> \<langle>i,s\<rangle> -*\<rightarrow> \<langle>j,t\<rangle> == ((i,s),j,t) : (exec01 p)^*"  | 
|
| 14565 | 38  | 
|
| 23746 | 39  | 
abbreviation  | 
40  | 
exec0n :: "[instrs, nat,state, nat, nat,state] \<Rightarrow> bool"  | 
|
41  | 
    ("(_/ \<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
 | 
42  | 
"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
 | 
43  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
44  | 
subsection "M0 with lists"  | 
| 
 
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  | 
text {* We describe execution of programs in the machine by
 | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
47  | 
an operational (small step) semantics:  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
48  | 
*}  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
49  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
50  | 
types config = "instrs \<times> instrs \<times> state"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
52  | 
|
| 23746 | 53  | 
inductive_set  | 
54  | 
stepa1 :: "(config \<times> config)set"  | 
|
55  | 
and stepa1' :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool"  | 
|
56  | 
    ("((1\<langle>_,/_,/_\<rangle>)/ -1\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50)
 | 
|
57  | 
where  | 
|
58  | 
"\<langle>p,q,s\<rangle> -1\<rightarrow> \<langle>p',q',t\<rangle> == ((p,q,s),p',q',t) : stepa1"  | 
|
59  | 
| "\<langle>SET x a#p,q,s\<rangle> -1\<rightarrow> \<langle>p,SET x a#q,s[x\<mapsto> a s]\<rangle>"  | 
|
60  | 
| "b s \<Longrightarrow> \<langle>JMPF b i#p,q,s\<rangle> -1\<rightarrow> \<langle>p,JMPF b i#q,s\<rangle>"  | 
|
61  | 
| "\<lbrakk> \<not> b s; i \<le> size p \<rbrakk>  | 
|
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
62  | 
\<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 | 63  | 
| "i \<le> size q  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
64  | 
\<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
 | 
65  | 
|
| 23746 | 66  | 
abbreviation  | 
67  | 
stepa :: "[instrs,instrs,state, instrs,instrs,state] \<Rightarrow> bool"  | 
|
68  | 
    ("((1\<langle>_,/_,/_\<rangle>)/ -*\<rightarrow> (1\<langle>_,/_,/_\<rangle>))" 50)  where
 | 
|
69  | 
"\<langle>p,q,s\<rangle> -*\<rightarrow> \<langle>p',q',t\<rangle> == ((p,q,s),p',q',t) : (stepa1^*)"  | 
|
70  | 
||
71  | 
abbreviation  | 
|
72  | 
stepan :: "[instrs,instrs,state, nat, instrs,instrs,state] \<Rightarrow> bool"  | 
|
73  | 
    ("((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
 | 
74  | 
"\<langle>p,q,s\<rangle> -i\<rightarrow> \<langle>p',q',t\<rangle> == ((p,q,s),p',q',t) : (stepa1^^i)"  | 
| 23746 | 75  | 
|
76  | 
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
 | 
77  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
78  | 
lemma exec_simp[simp]:  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
79  | 
"(\<langle>i#p,q,s\<rangle> -1\<rightarrow> \<langle>p',q',t\<rangle>) = (case i of  | 
| 13675 | 80  | 
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
 | 
81  | 
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
 | 
82  | 
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
 | 
83  | 
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
 | 
84  | 
apply(rule iffI)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
85  | 
defer  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
86  | 
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
 | 
87  | 
apply(erule execE)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
88  | 
apply(simp_all)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
89  | 
done  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
90  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
91  | 
lemma execn_simp[simp]:  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
92  | 
"(\<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
 | 
93  | 
(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
 | 
94  | 
((\<exists>m p' q' t. n = Suc m \<and>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
95  | 
\<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
 | 
96  | 
by(subst converse_in_rel_pow_eq, simp)  | 
| 
 
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  | 
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
 | 
100  | 
(p'' = i#p & q''=q & u=s |  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
101  | 
(\<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
 | 
102  | 
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
 | 
103  | 
apply(blast)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
104  | 
done  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
105  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
106  | 
declare nth_append[simp]  | 
| 
 
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  | 
lemma rev_revD: "rev xs = rev ys \<Longrightarrow> xs = ys"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
109  | 
by simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
110  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
111  | 
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
 | 
112  | 
apply(rule iffI)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
113  | 
apply(rule rev_revD, simp)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
114  | 
apply fastsimp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
115  | 
done  | 
| 
 
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  | 
lemma direction1:  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
118  | 
"\<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
 | 
119  | 
rev p' @ q' = rev p @ q \<and> rev p @ q \<turnstile> \<langle>size p,s\<rangle> -1\<rightarrow> \<langle>size p',t\<rangle>"  | 
| 18372 | 120  | 
apply(induct set: stepa1)  | 
| 13675 | 121  | 
apply(simp add:exec01.SET)  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
122  | 
apply(fastsimp intro:exec01.JMPFT)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
123  | 
apply simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
124  | 
apply(rule exec01.JMPFF)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
125  | 
apply simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
126  | 
apply fastsimp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
127  | 
apply simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
128  | 
apply simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
129  | 
apply simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
130  | 
apply(fastsimp simp add:exec01.JMPB)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
131  | 
done  | 
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18372 
diff
changeset
 | 
132  | 
|
| 13098 | 133  | 
(*  | 
134  | 
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
 | 
135  | 
apply(induct xs)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
136  | 
apply simp_all  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
137  | 
apply(case_tac i)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
138  | 
apply simp_all  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
139  | 
done  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
140  | 
|
| 13098 | 141  | 
lemma rev_drop: "\<And>i. rev (drop i xs) = take (length xs - i) (rev xs)"  | 
142  | 
apply(induct xs)  | 
|
143  | 
apply simp_all  | 
|
144  | 
apply(case_tac i)  | 
|
145  | 
apply simp_all  | 
|
146  | 
done  | 
|
147  | 
*)  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
18372 
diff
changeset
 | 
148  | 
|
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
149  | 
lemma direction2:  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
150  | 
"rpq \<turnstile> \<langle>sp,s\<rangle> -1\<rightarrow> \<langle>sp',t\<rangle> \<Longrightarrow>  | 
| 18372 | 151  | 
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
 | 
152  | 
rev p' @ q' = rev p @ q \<longrightarrow> \<langle>q,p,s\<rangle> -1\<rightarrow> \<langle>q',p',t\<rangle>"  | 
| 20503 | 153  | 
apply(induct arbitrary: p q p' q' set: exec01)  | 
| 13098 | 154  | 
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
 | 
155  | 
apply(drule sym)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
156  | 
apply simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
157  | 
apply(rule rev_revD)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
158  | 
apply simp  | 
| 13098 | 159  | 
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
 | 
160  | 
apply(drule sym)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
161  | 
apply simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
162  | 
apply(rule rev_revD)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
163  | 
apply simp  | 
| 13612 | 164  | 
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
 | 
165  | 
apply(drule sym)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
166  | 
apply simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
167  | 
apply(rule rev_revD)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
168  | 
apply simp  | 
| 13098 | 169  | 
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
 | 
170  | 
apply(drule sym)  | 
| 13098 | 171  | 
apply(simp add:rev_take)  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
172  | 
apply(rule rev_revD)  | 
| 13098 | 173  | 
apply(simp add:rev_drop)  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
174  | 
done  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
175  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
176  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
177  | 
theorem M_eqiv:  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
178  | 
"(\<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
 | 
179  | 
(rev p' @ q' = rev p @ q \<and> rev p @ q \<turnstile> \<langle>size p,s\<rangle> -1\<rightarrow> \<langle>size p',t\<rangle>)"  | 
| 18372 | 180  | 
by (blast dest: direction1 direction2)  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
181  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents:  
diff
changeset
 | 
182  | 
end  |