10343
|
1 |
(* Title: HOL/IMP/Compiler.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Tobias Nipkow, TUM
|
|
4 |
Copyright 1996 TUM
|
|
5 |
|
|
6 |
A simple compiler for a simplistic machine.
|
|
7 |
*)
|
|
8 |
|
10342
|
9 |
theory Compiler = Natural:
|
|
10 |
|
|
11 |
datatype instr = ASIN loc aexp | JMPF bexp nat | JMPB nat
|
|
12 |
|
|
13 |
consts stepa1 :: "instr list => ((state*nat) * (state*nat))set"
|
|
14 |
|
|
15 |
syntax
|
|
16 |
"@stepa1" :: "[instr list,state,nat,state,nat] => bool"
|
|
17 |
("_ |- <_,_>/ -1-> <_,_>" [50,0,0,0,0] 50)
|
|
18 |
"@stepa" :: "[instr list,state,nat,state,nat] => bool"
|
|
19 |
("_ |-/ <_,_>/ -*-> <_,_>" [50,0,0,0,0] 50)
|
|
20 |
|
|
21 |
translations "P |- <s,m> -1-> <t,n>" == "((s,m),t,n) : stepa1 P"
|
|
22 |
"P |- <s,m> -*-> <t,n>" == "((s,m),t,n) : ((stepa1 P)^*)"
|
|
23 |
|
|
24 |
|
|
25 |
inductive "stepa1 P"
|
|
26 |
intros
|
|
27 |
ASIN: "P!n = ASIN x a ==> P |- <s,n> -1-> <s[x::= a s],Suc n>"
|
|
28 |
JMPFT: "[| P!n = JMPF b i; b s |] ==> P |- <s,n> -1-> <s,Suc n>"
|
|
29 |
JMPFF: "[| P!n = JMPF b i; ~b s; m=n+i |] ==> P |- <s,n> -1-> <s,m>"
|
|
30 |
JMPB: "[| P!n = JMB i |] ==> P |- <s,n> -1-> <s,n-i>"
|
|
31 |
|
|
32 |
consts compile :: "com => instr list"
|
|
33 |
primrec
|
|
34 |
"compile SKIP = []"
|
|
35 |
"compile (x:==a) = [ASIN x a]"
|
|
36 |
"compile (c1;c2) = compile c1 @ compile c2"
|
|
37 |
"compile (IF b THEN c1 ELSE c2) =
|
|
38 |
[JMPF b (length(compile c1)+2)] @ compile c1 @
|
|
39 |
[JMPF (%x. False) (length(compile c2)+1)] @ compile c2"
|
|
40 |
"compile (WHILE b DO c) = [JMPF b (length(compile c)+2)] @ compile c @
|
|
41 |
[JMPB (length(compile c)+1)]"
|
|
42 |
|
|
43 |
declare nth_append[simp];
|
|
44 |
|
|
45 |
lemma nth_tl[simp]: "tl(xs @ y # ys) ! (length xs + z) = ys!z";
|
|
46 |
apply(induct_tac xs);
|
|
47 |
by(auto);
|
|
48 |
|
|
49 |
theorem "<c,s> -c-> t ==>
|
|
50 |
!a z. a@compile c@z |- <s,length a> -*-> <t,length a + length(compile c)>";
|
|
51 |
apply(erule evalc.induct);
|
|
52 |
apply simp;
|
|
53 |
apply(force intro!: ASIN);
|
|
54 |
apply(intro strip);
|
|
55 |
apply(erule_tac x = a in allE);
|
|
56 |
apply(erule_tac x = "a@compile c0" in allE);
|
|
57 |
apply(erule_tac x = "compile c1@z" in allE);
|
|
58 |
apply(erule_tac x = z in allE);
|
|
59 |
apply(simp add:add_assoc[THEN sym]);
|
|
60 |
apply(blast intro:rtrancl_trans);
|
|
61 |
(* IF b THEN c0 ELSE c1; case b is true *)
|
|
62 |
apply(intro strip);
|
|
63 |
(* instantiate assumption sufficiently for later: *)
|
|
64 |
apply(erule_tac x = "a@[?I]" in allE);
|
|
65 |
apply(simp);
|
|
66 |
(* execute JMPF: *)
|
|
67 |
apply(rule rtrancl_into_rtrancl2);
|
|
68 |
apply(rule JMPFT);
|
|
69 |
apply(simp);
|
|
70 |
apply(blast);
|
|
71 |
apply assumption;
|
|
72 |
(* execute compile c0: *)
|
|
73 |
apply(rule rtrancl_trans);
|
|
74 |
apply(erule allE);
|
|
75 |
apply assumption;
|
|
76 |
(* execute JMPF: *)
|
|
77 |
apply(rule r_into_rtrancl);
|
|
78 |
apply(rule JMPFF);
|
|
79 |
apply(simp);
|
|
80 |
apply(blast);
|
|
81 |
apply(blast);
|
|
82 |
apply(simp);
|
|
83 |
(* end of case b is true *)
|
|
84 |
apply(intro strip);
|
|
85 |
apply(erule_tac x = "a@[?I]@compile c0@[?J]" in allE);
|
|
86 |
apply(simp add:add_assoc);
|
|
87 |
apply(rule rtrancl_into_rtrancl2);
|
|
88 |
apply(rule JMPFF);
|
|
89 |
apply(simp);
|
|
90 |
apply(blast);
|
|
91 |
apply assumption;
|
|
92 |
apply(simp);
|
|
93 |
apply(blast);
|
|
94 |
apply(force intro: JMPFF);
|
|
95 |
apply(intro strip);
|
|
96 |
apply(erule_tac x = "a@[?I]" in allE);
|
|
97 |
apply(erule_tac x = a in allE);
|
|
98 |
apply(simp);
|
|
99 |
apply(rule rtrancl_into_rtrancl2);
|
|
100 |
apply(rule JMPFT);
|
|
101 |
apply(simp);
|
|
102 |
apply(blast);
|
|
103 |
apply assumption;
|
|
104 |
apply(rule rtrancl_trans);
|
|
105 |
apply(erule allE);
|
|
106 |
apply assumption;
|
|
107 |
apply(rule rtrancl_into_rtrancl2);
|
|
108 |
apply(rule JMPB);
|
|
109 |
apply(simp);
|
|
110 |
apply(simp);
|
|
111 |
done
|
|
112 |
|
|
113 |
(* Missing: the other direction! *)
|
|
114 |
|
|
115 |
end
|