| author | haftmann | 
| Tue, 11 May 2010 18:46:03 +0200 | |
| changeset 36832 | e6078ef937df | 
| parent 32960 | 69916a850301 | 
| child 43141 | 11fce8564415 | 
| permissions | -rw-r--r-- | 
| 10343 | 1  | 
(* Title: HOL/IMP/Compiler.thy  | 
2  | 
Author: Tobias Nipkow, TUM  | 
|
3  | 
Copyright 1996 TUM  | 
|
| 12429 | 4  | 
*)  | 
| 10343 | 5  | 
|
| 16417 | 6  | 
theory Compiler imports Machines begin  | 
| 12429 | 7  | 
|
8  | 
subsection "The compiler"  | 
|
| 10342 | 9  | 
|
| 27362 | 10  | 
primrec compile :: "com \<Rightarrow> instr list"  | 
11  | 
where  | 
|
12  | 
"compile \<SKIP> = []"  | 
|
13  | 
| "compile (x:==a) = [SET x a]"  | 
|
14  | 
| "compile (c1;c2) = compile c1 @ compile c2"  | 
|
15  | 
| "compile (\<IF> b \<THEN> c1 \<ELSE> c2) =  | 
|
16  | 
[JMPF b (length(compile c1) + 1)] @ compile c1 @  | 
|
17  | 
[JMPF (\<lambda>x. False) (length(compile c2))] @ compile c2"  | 
|
18  | 
| "compile (\<WHILE> b \<DO> c) = [JMPF b (length(compile c) + 1)] @ compile c @  | 
|
19  | 
[JMPB (length(compile c)+1)]"  | 
|
| 10342 | 20  | 
|
| 12429 | 21  | 
subsection "Compiler correctness"  | 
| 11275 | 22  | 
|
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
23  | 
theorem assumes A: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
24  | 
shows "\<And>p q. \<langle>compile c @ p,q,s\<rangle> -*\<rightarrow> \<langle>p,rev(compile c)@q,t\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
25  | 
(is "\<And>p q. ?P c s t p q")  | 
| 12275 | 26  | 
proof -  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
27  | 
from A show "\<And>p q. ?thesis p q"  | 
| 12275 | 28  | 
proof induct  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
29  | 
case Skip thus ?case by simp  | 
| 12275 | 30  | 
next  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
31  | 
case Assign thus ?case by force  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
32  | 
next  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
33  | 
case Semi thus ?case by simp (blast intro:rtrancl_trans)  | 
| 12275 | 34  | 
next  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
35  | 
fix b c0 c1 s0 s1 p q  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
36  | 
assume IH: "\<And>p q. ?P c0 s0 s1 p q"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
37  | 
assume "b s0"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
38  | 
thus "?P (\<IF> b \<THEN> c0 \<ELSE> c1) s0 s1 p q"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
39  | 
by(simp add: IH[THEN rtrancl_trans])  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
40  | 
next  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
41  | 
case IfFalse thus ?case by(simp)  | 
| 12275 | 42  | 
next  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
43  | 
case WhileFalse thus ?case by simp  | 
| 12275 | 44  | 
next  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
45  | 
fix b c and s0::state and s1 s2 p q  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
46  | 
assume b: "b s0" and  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
47  | 
IHc: "\<And>p q. ?P c s0 s1 p q" and  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
48  | 
IHw: "\<And>p q. ?P (\<WHILE> b \<DO> c) s1 s2 p q"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
49  | 
show "?P (\<WHILE> b \<DO> c) s0 s2 p q"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
50  | 
using b IHc[THEN rtrancl_trans] IHw by(simp)  | 
| 12275 | 51  | 
qed  | 
52  | 
qed  | 
|
| 11275 | 53  | 
|
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
54  | 
text {* The other direction! *}
 | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
55  | 
|
| 23746 | 56  | 
inductive_cases [elim!]: "(([],p,s),(is',p',s')) : stepa1"  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
57  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
58  | 
lemma [simp]: "(\<langle>[],q,s\<rangle> -n\<rightarrow> \<langle>p',q',t\<rangle>) = (n=0 \<and> p' = [] \<and> q' = q \<and> t = s)"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
59  | 
apply(rule iffI)  | 
| 31969 | 60  | 
apply(erule rel_pow_E2, simp, fast)  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
61  | 
apply simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
62  | 
done  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
63  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
64  | 
lemma [simp]: "(\<langle>[],q,s\<rangle> -*\<rightarrow> \<langle>p',q',t\<rangle>) = (p' = [] \<and> q' = q \<and> t = s)"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
65  | 
by(simp add: rtrancl_is_UN_rel_pow)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
66  | 
|
| 27362 | 67  | 
definition  | 
68  | 
forws :: "instr \<Rightarrow> nat set" where  | 
|
69  | 
"forws instr = (case instr of  | 
|
70  | 
     SET x a \<Rightarrow> {0} |
 | 
|
71  | 
     JMPF b n \<Rightarrow> {0,n} |
 | 
|
72  | 
     JMPB n \<Rightarrow> {})"
 | 
|
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
73  | 
|
| 27362 | 74  | 
definition  | 
75  | 
backws :: "instr \<Rightarrow> nat set" where  | 
|
76  | 
"backws instr = (case instr of  | 
|
77  | 
     SET x a \<Rightarrow> {} |
 | 
|
78  | 
     JMPF b n \<Rightarrow> {} |
 | 
|
79  | 
     JMPB n \<Rightarrow> {n})"
 | 
|
80  | 
||
81  | 
primrec closed :: "nat \<Rightarrow> nat \<Rightarrow> instr list \<Rightarrow> bool"  | 
|
82  | 
where  | 
|
83  | 
"closed m n [] = True"  | 
|
84  | 
| "closed m n (instr#is) = ((\<forall>j \<in> forws instr. j \<le> size is+n) \<and>  | 
|
85  | 
(\<forall>j \<in> backws instr. j \<le> m) \<and> closed (Suc m) n is)"  | 
|
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
86  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
87  | 
lemma [simp]:  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
88  | 
"\<And>m n. closed m n (C1@C2) =  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
89  | 
(closed m (n+size C2) C1 \<and> closed (m+size C1) n C2)"  | 
| 27362 | 90  | 
by(induct C1) (simp, simp add:add_ac)  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
91  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
92  | 
theorem [simp]: "\<And>m n. closed m n (compile c)"  | 
| 27362 | 93  | 
by(induct c) (simp_all add:backws_def forws_def)  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
94  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
95  | 
lemma drop_lem: "n \<le> size(p1@p2)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
96  | 
\<Longrightarrow> (p1' @ p2 = drop n p1 @ drop (n - size p1) p2) =  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
97  | 
(n \<le> size p1 & p1' = drop n p1)"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
98  | 
apply(rule iffI)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
99  | 
defer apply simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
100  | 
apply(subgoal_tac "n \<le> size p1")  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
101  | 
apply simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
102  | 
apply(rule ccontr)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
103  | 
apply(drule_tac f = length in arg_cong)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
104  | 
apply simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
105  | 
done  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
106  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
107  | 
lemma reduce_exec1:  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
108  | 
"\<langle>i # p1 @ p2,q1 @ q2,s\<rangle> -1\<rightarrow> \<langle>p1' @ p2,q1' @ q2,s'\<rangle> \<Longrightarrow>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
109  | 
\<langle>i # p1,q1,s\<rangle> -1\<rightarrow> \<langle>p1',q1',s'\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
110  | 
by(clarsimp simp add: drop_lem split:instr.split_asm split_if_asm)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
111  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
112  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
113  | 
lemma closed_exec1:  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
114  | 
"\<lbrakk> closed 0 0 (rev q1 @ instr # p1);  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
115  | 
\<langle>instr # p1 @ p2, q1 @ q2,r\<rangle> -1\<rightarrow> \<langle>p',q',r'\<rangle> \<rbrakk> \<Longrightarrow>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
116  | 
\<exists>p1' q1'. p' = p1'@p2 \<and> q' = q1'@q2 \<and> rev q1' @ p1' = rev q1 @ instr # p1"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
117  | 
apply(clarsimp simp add:forws_def backws_def  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
118  | 
split:instr.split_asm split_if_asm)  | 
| 10342 | 119  | 
done  | 
120  | 
||
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
121  | 
theorem closed_execn_decomp: "\<And>C1 C2 r.  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
122  | 
\<lbrakk> closed 0 0 (rev C1 @ C2);  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
123  | 
\<langle>C2 @ p1 @ p2, C1 @ q,r\<rangle> -n\<rightarrow> \<langle>p2,rev p1 @ rev C2 @ C1 @ q,t\<rangle> \<rbrakk>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
124  | 
\<Longrightarrow> \<exists>s n1 n2. \<langle>C2,C1,r\<rangle> -n1\<rightarrow> \<langle>[],rev C2 @ C1,s\<rangle> \<and>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
125  | 
\<langle>p1@p2,rev C2 @ C1 @ q,s\<rangle> -n2\<rightarrow> \<langle>p2, rev p1 @ rev C2 @ C1 @ q,t\<rangle> \<and>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
126  | 
n = n1+n2"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
127  | 
(is "\<And>C1 C2 r. \<lbrakk>?CL C1 C2; ?H C1 C2 r n\<rbrakk> \<Longrightarrow> ?P C1 C2 r n")  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
128  | 
proof(induct n)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
129  | 
fix C1 C2 r  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
130  | 
assume "?H C1 C2 r 0"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
131  | 
thus "?P C1 C2 r 0" by simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
132  | 
next  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
133  | 
fix C1 C2 r n  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
134  | 
assume IH: "\<And>C1 C2 r. ?CL C1 C2 \<Longrightarrow> ?H C1 C2 r n \<Longrightarrow> ?P C1 C2 r n"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
135  | 
assume CL: "?CL C1 C2" and H: "?H C1 C2 r (Suc n)"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
136  | 
show "?P C1 C2 r (Suc n)"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
137  | 
proof (cases C2)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
138  | 
assume "C2 = []" with H show ?thesis by simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
139  | 
next  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
140  | 
fix instr tlC2  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
141  | 
assume C2: "C2 = instr # tlC2"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
142  | 
from H C2 obtain p' q' r'  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
143  | 
where 1: "\<langle>instr # tlC2 @ p1 @ p2, C1 @ q,r\<rangle> -1\<rightarrow> \<langle>p',q',r'\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
144  | 
and n: "\<langle>p',q',r'\<rangle> -n\<rightarrow> \<langle>p2,rev p1 @ rev C2 @ C1 @ q,t\<rangle>"  | 
| 31969 | 145  | 
by(fastsimp simp add:rel_pow_commute)  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
146  | 
from CL closed_exec1[OF _ 1] C2  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
147  | 
obtain C2' C1' where pq': "p' = C2' @ p1 @ p2 \<and> q' = C1' @ q"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
148  | 
and same: "rev C1' @ C2' = rev C1 @ C2"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
149  | 
by fastsimp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
150  | 
have rev_same: "rev C2' @ C1' = rev C2 @ C1"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
151  | 
proof -  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
152  | 
have "rev C2' @ C1' = rev(rev C1' @ C2')" by simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
153  | 
also have "\<dots> = rev(rev C1 @ C2)" by(simp only:same)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
154  | 
also have "\<dots> = rev C2 @ C1" by simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
155  | 
finally show ?thesis .  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
156  | 
qed  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
157  | 
hence rev_same': "\<And>p. rev C2' @ C1' @ p = rev C2 @ C1 @ p" by simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
158  | 
from n have n': "\<langle>C2' @ p1 @ p2,C1' @ q,r'\<rangle> -n\<rightarrow>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
159  | 
\<langle>p2,rev p1 @ rev C2' @ C1' @ q,t\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
160  | 
by(simp add:pq' rev_same')  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
161  | 
from IH[OF _ n'] CL  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
162  | 
obtain s n1 n2 where n1: "\<langle>C2',C1',r'\<rangle> -n1\<rightarrow> \<langle>[],rev C2 @ C1,s\<rangle>" and  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
163  | 
"\<langle>p1 @ p2,rev C2 @ C1 @ q,s\<rangle> -n2\<rightarrow> \<langle>p2,rev p1 @ rev C2 @ C1 @ q,t\<rangle> \<and>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
164  | 
n = n1 + n2"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
165  | 
by(fastsimp simp add: same rev_same rev_same')  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
166  | 
moreover  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
167  | 
from 1 n1 pq' C2 have "\<langle>C2,C1,r\<rangle> -Suc n1\<rightarrow> \<langle>[],rev C2 @ C1,s\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
168  | 
by (simp del:relpow.simps exec_simp) (fast dest:reduce_exec1)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
169  | 
ultimately show ?thesis by (fastsimp simp del:relpow.simps)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
170  | 
qed  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
171  | 
qed  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
172  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
173  | 
lemma execn_decomp:  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
174  | 
"\<langle>compile c @ p1 @ p2,q,r\<rangle> -n\<rightarrow> \<langle>p2,rev p1 @ rev(compile c) @ q,t\<rangle>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
175  | 
\<Longrightarrow> \<exists>s n1 n2. \<langle>compile c,[],r\<rangle> -n1\<rightarrow> \<langle>[],rev(compile c),s\<rangle> \<and>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
176  | 
\<langle>p1@p2,rev(compile c) @ q,s\<rangle> -n2\<rightarrow> \<langle>p2, rev p1 @ rev(compile c) @ q,t\<rangle> \<and>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
177  | 
n = n1+n2"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
178  | 
using closed_execn_decomp[of "[]",simplified] by simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
179  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
180  | 
lemma exec_star_decomp:  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
181  | 
"\<langle>compile c @ p1 @ p2,q,r\<rangle> -*\<rightarrow> \<langle>p2,rev p1 @ rev(compile c) @ q,t\<rangle>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
182  | 
\<Longrightarrow> \<exists>s. \<langle>compile c,[],r\<rangle> -*\<rightarrow> \<langle>[],rev(compile c),s\<rangle> \<and>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
183  | 
\<langle>p1@p2,rev(compile c) @ q,s\<rangle> -*\<rightarrow> \<langle>p2, rev p1 @ rev(compile c) @ q,t\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
184  | 
by(simp add:rtrancl_is_UN_rel_pow)(fast dest: execn_decomp)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
185  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
186  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
187  | 
(* Alternative:  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
188  | 
lemma exec_comp_n:  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
189  | 
"\<And>p1 p2 q r t n.  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
190  | 
\<langle>compile c @ p1 @ p2,q,r\<rangle> -n\<rightarrow> \<langle>p2,rev p1 @ rev(compile c) @ q,t\<rangle>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
191  | 
\<Longrightarrow> \<exists>s n1 n2. \<langle>compile c,[],r\<rangle> -n1\<rightarrow> \<langle>[],rev(compile c),s\<rangle> \<and>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
192  | 
\<langle>p1@p2,rev(compile c) @ q,s\<rangle> -n2\<rightarrow> \<langle>p2, rev p1 @ rev(compile c) @ q,t\<rangle> \<and>  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
193  | 
n = n1+n2"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
194  | 
(is "\<And>p1 p2 q r t n. ?H c p1 p2 q r t n \<Longrightarrow> ?P c p1 p2 q r t n")  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
195  | 
proof (induct c)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
196  | 
*)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
197  | 
|
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
198  | 
text{*Warning: 
 | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
199  | 
@{prop"\<langle>compile c @ p,q,s\<rangle> -*\<rightarrow> \<langle>p,rev(compile c)@q,t\<rangle> \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t"}
 | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
200  | 
is not true! *}  | 
| 10342 | 201  | 
|
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
202  | 
theorem "\<And>s t.  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
203  | 
\<langle>compile c,[],s\<rangle> -*\<rightarrow> \<langle>[],rev(compile c),t\<rangle> \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
204  | 
proof (induct c)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
205  | 
fix s t  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
206  | 
assume "\<langle>compile SKIP,[],s\<rangle> -*\<rightarrow> \<langle>[],rev(compile SKIP),t\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
207  | 
thus "\<langle>SKIP,s\<rangle> \<longrightarrow>\<^sub>c t" by simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
208  | 
next  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
209  | 
fix s t v f  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
210  | 
assume "\<langle>compile(v :== f),[],s\<rangle> -*\<rightarrow> \<langle>[],rev(compile(v :== f)),t\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
211  | 
thus "\<langle>v :== f,s\<rangle> \<longrightarrow>\<^sub>c t" by simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
212  | 
next  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
213  | 
fix s1 s3 c1 c2  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
214  | 
let ?C1 = "compile c1" let ?C2 = "compile c2"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
215  | 
assume IH1: "\<And>s t. \<langle>?C1,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C1,t\<rangle> \<Longrightarrow> \<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c t"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
216  | 
and IH2: "\<And>s t. \<langle>?C2,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C2,t\<rangle> \<Longrightarrow> \<langle>c2,s\<rangle> \<longrightarrow>\<^sub>c t"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
217  | 
assume "\<langle>compile(c1;c2),[],s1\<rangle> -*\<rightarrow> \<langle>[],rev(compile(c1;c2)),s3\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
218  | 
then obtain s2 where exec1: "\<langle>?C1,[],s1\<rangle> -*\<rightarrow> \<langle>[],rev ?C1,s2\<rangle>" and  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
219  | 
exec2: "\<langle>?C2,rev ?C1,s2\<rangle> -*\<rightarrow> \<langle>[],rev(compile(c1;c2)),s3\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
220  | 
by(fastsimp dest:exec_star_decomp[of _ _ "[]" "[]",simplified])  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
221  | 
from exec2 have exec2': "\<langle>?C2,[],s2\<rangle> -*\<rightarrow> \<langle>[],rev ?C2,s3\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
222  | 
using exec_star_decomp[of _ "[]" "[]"] by fastsimp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
223  | 
have "\<langle>c1,s1\<rangle> \<longrightarrow>\<^sub>c s2" using IH1 exec1 by simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
224  | 
moreover have "\<langle>c2,s2\<rangle> \<longrightarrow>\<^sub>c s3" using IH2 exec2' by fastsimp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
225  | 
ultimately show "\<langle>c1;c2,s1\<rangle> \<longrightarrow>\<^sub>c s3" ..  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
226  | 
next  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
227  | 
fix s t b c1 c2  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
228  | 
let ?if = "IF b THEN c1 ELSE c2" let ?C = "compile ?if"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
229  | 
let ?C1 = "compile c1" let ?C2 = "compile c2"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
230  | 
assume IH1: "\<And>s t. \<langle>?C1,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C1,t\<rangle> \<Longrightarrow> \<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c t"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
231  | 
and IH2: "\<And>s t. \<langle>?C2,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C2,t\<rangle> \<Longrightarrow> \<langle>c2,s\<rangle> \<longrightarrow>\<^sub>c t"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
232  | 
and H: "\<langle>?C,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C,t\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
233  | 
show "\<langle>?if,s\<rangle> \<longrightarrow>\<^sub>c t"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
234  | 
proof cases  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
235  | 
assume b: "b s"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
236  | 
with H have "\<langle>?C1,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C1,t\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
237  | 
by (fastsimp dest:exec_star_decomp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
238  | 
[of _ "[JMPF (\<lambda>x. False) (size ?C2)]@?C2" "[]",simplified])  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
239  | 
hence "\<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c t" by(rule IH1)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
240  | 
with b show ?thesis ..  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
241  | 
next  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
242  | 
assume b: "\<not> b s"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
243  | 
with H have "\<langle>?C2,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C2,t\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
244  | 
using exec_star_decomp[of _ "[]" "[]"] by simp  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
245  | 
hence "\<langle>c2,s\<rangle> \<longrightarrow>\<^sub>c t" by(rule IH2)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
246  | 
with b show ?thesis ..  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
247  | 
qed  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
248  | 
next  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
249  | 
fix b c s t  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
250  | 
let ?w = "WHILE b DO c" let ?W = "compile ?w" let ?C = "compile c"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
251  | 
let ?j1 = "JMPF b (size ?C + 1)" let ?j2 = "JMPB (size ?C + 1)"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
252  | 
assume IHc: "\<And>s t. \<langle>?C,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C,t\<rangle> \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
253  | 
and H: "\<langle>?W,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?W,t\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
254  | 
from H obtain k where ob:"\<langle>?W,[],s\<rangle> -k\<rightarrow> \<langle>[],rev ?W,t\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
255  | 
by(simp add:rtrancl_is_UN_rel_pow) blast  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
256  | 
  { fix n have "\<And>s. \<langle>?W,[],s\<rangle> -n\<rightarrow> \<langle>[],rev ?W,t\<rangle> \<Longrightarrow> \<langle>?w,s\<rangle> \<longrightarrow>\<^sub>c t"
 | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
257  | 
proof (induct n rule: less_induct)  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
258  | 
fix n  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
259  | 
assume IHm: "\<And>m s. \<lbrakk>m < n; \<langle>?W,[],s\<rangle> -m\<rightarrow> \<langle>[],rev ?W,t\<rangle> \<rbrakk> \<Longrightarrow> \<langle>?w,s\<rangle> \<longrightarrow>\<^sub>c t"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
260  | 
fix s  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
261  | 
assume H: "\<langle>?W,[],s\<rangle> -n\<rightarrow> \<langle>[],rev ?W,t\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
262  | 
show "\<langle>?w,s\<rangle> \<longrightarrow>\<^sub>c t"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
263  | 
proof cases  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
264  | 
assume b: "b s"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
265  | 
then obtain m where m: "n = Suc m"  | 
| 13675 | 266  | 
and "\<langle>?C @ [?j2],[?j1],s\<rangle> -m\<rightarrow> \<langle>[],rev ?W,t\<rangle>"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
267  | 
using H by fastsimp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
268  | 
then obtain r n1 n2 where n1: "\<langle>?C,[],s\<rangle> -n1\<rightarrow> \<langle>[],rev ?C,r\<rangle>"  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
269  | 
and n2: "\<langle>[?j2],rev ?C @ [?j1],r\<rangle> -n2\<rightarrow> \<langle>[],rev ?W,t\<rangle>"  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
270  | 
and n12: "m = n1+n2"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
271  | 
using execn_decomp[of _ "[?j2]"]  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
272  | 
by(simp del: execn_simp) fast  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
273  | 
have n2n: "n2 - 1 < n" using m n12 by arith  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
274  | 
note b  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
275  | 
moreover  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
276  | 
        { from n1 have "\<langle>?C,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C,r\<rangle>"
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
277  | 
by (simp add:rtrancl_is_UN_rel_pow) fast  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
278  | 
hence "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c r" by(rule IHc)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
279  | 
}  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
280  | 
moreover  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
281  | 
        { have "n2 - 1 < n" using m n12 by arith
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
282  | 
moreover from n2 have "\<langle>?W,[],r\<rangle> -n2- 1\<rightarrow> \<langle>[],rev ?W,t\<rangle>" by fastsimp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
283  | 
ultimately have "\<langle>?w,r\<rangle> \<longrightarrow>\<^sub>c t" by(rule IHm)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
284  | 
}  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
285  | 
ultimately show ?thesis ..  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
286  | 
next  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
287  | 
assume b: "\<not> b s"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
288  | 
hence "t = s" using H by simp  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
31969 
diff
changeset
 | 
289  | 
with b show ?thesis by simp  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
290  | 
qed  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
291  | 
qed  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
292  | 
}  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
293  | 
with ob show "\<langle>?w,s\<rangle> \<longrightarrow>\<^sub>c t" by fast  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
294  | 
qed  | 
| 
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
295  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
16417 
diff
changeset
 | 
296  | 
(* TODO: connect with Machine 0 using M_equiv *)  | 
| 
13095
 
8ed413a57bdc
New machine architecture and other direction of compiler proof.
 
nipkow 
parents: 
12566 
diff
changeset
 | 
297  | 
|
| 
20217
 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 
webertj 
parents: 
16417 
diff
changeset
 | 
298  | 
end  |