author | wenzelm |
Mon, 11 Feb 2013 14:39:04 +0100 | |
changeset 51085 | d90218288d51 |
parent 48909 | b2dea007e55e |
child 51506 | cd573f1a82b2 |
permissions | -rw-r--r-- |
44070 | 1 |
header "Constant Folding" |
2 |
||
3 |
theory Fold imports Sem_Equiv begin |
|
4 |
||
44850 | 5 |
subsection "Simple folding of arithmetic expressions" |
44070 | 6 |
|
45134 | 7 |
type_synonym |
45212 | 8 |
tab = "vname \<Rightarrow> val option" |
44070 | 9 |
|
10 |
(* maybe better as the composition of substitution and the existing simp_const(0) *) |
|
11 |
fun simp_const :: "aexp \<Rightarrow> tab \<Rightarrow> aexp" where |
|
12 |
"simp_const (N n) _ = N n" | |
|
13 |
"simp_const (V x) t = (case t x of None \<Rightarrow> V x | Some k \<Rightarrow> N k)" | |
|
14 |
"simp_const (Plus e1 e2) t = (case (simp_const e1 t, simp_const e2 t) of |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
15 |
(N n1, N n2) \<Rightarrow> N(n1+n2) | (e1',e2') \<Rightarrow> Plus e1' e2')" |
44070 | 16 |
|
17 |
definition "approx t s \<longleftrightarrow> (ALL x k. t x = Some k \<longrightarrow> s x = k)" |
|
18 |
||
19 |
theorem aval_simp_const[simp]: |
|
20 |
assumes "approx t s" |
|
21 |
shows "aval (simp_const a t) s = aval a s" |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
22 |
using assms |
44070 | 23 |
by (induct a) (auto simp: approx_def split: aexp.split option.split) |
24 |
||
25 |
theorem aval_simp_const_N: |
|
26 |
assumes "approx t s" |
|
27 |
shows "simp_const a t = N n \<Longrightarrow> aval a s = n" |
|
28 |
using assms |
|
29 |
by (induct a arbitrary: n) |
|
30 |
(auto simp: approx_def split: aexp.splits option.splits) |
|
31 |
||
32 |
definition |
|
33 |
"merge t1 t2 = (\<lambda>m. if t1 m = t2 m then t1 m else None)" |
|
34 |
||
45212 | 35 |
primrec lnames :: "com \<Rightarrow> vname set" where |
44070 | 36 |
"lnames SKIP = {}" | |
37 |
"lnames (x ::= a) = {x}" | |
|
38 |
"lnames (c1; c2) = lnames c1 \<union> lnames c2" | |
|
39 |
"lnames (IF b THEN c1 ELSE c2) = lnames c1 \<union> lnames c2" | |
|
40 |
"lnames (WHILE b DO c) = lnames c" |
|
41 |
||
42 |
primrec "defs" :: "com \<Rightarrow> tab \<Rightarrow> tab" where |
|
43 |
"defs SKIP t = t" | |
|
44 |
"defs (x ::= a) t = |
|
45 |
(case simp_const a t of N k \<Rightarrow> t(x \<mapsto> k) | _ \<Rightarrow> t(x:=None))" | |
|
46 |
"defs (c1;c2) t = (defs c2 o defs c1) t" | |
|
47 |
"defs (IF b THEN c1 ELSE c2) t = merge (defs c1 t) (defs c2 t)" | |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
48 |
"defs (WHILE b DO c) t = t |` (-lnames c)" |
44070 | 49 |
|
50 |
primrec fold where |
|
51 |
"fold SKIP _ = SKIP" | |
|
52 |
"fold (x ::= a) t = (x ::= (simp_const a t))" | |
|
53 |
"fold (c1;c2) t = (fold c1 t; fold c2 (defs c1 t))" | |
|
54 |
"fold (IF b THEN c1 ELSE c2) t = IF b THEN fold c1 t ELSE fold c2 t" | |
|
55 |
"fold (WHILE b DO c) t = WHILE b DO fold c (t |` (-lnames c))" |
|
56 |
||
57 |
lemma approx_merge: |
|
58 |
"approx t1 s \<or> approx t2 s \<Longrightarrow> approx (merge t1 t2) s" |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44850
diff
changeset
|
59 |
by (fastforce simp: merge_def approx_def) |
44070 | 60 |
|
61 |
lemma approx_map_le: |
|
62 |
"approx t2 s \<Longrightarrow> t1 \<subseteq>\<^sub>m t2 \<Longrightarrow> approx t1 s" |
|
63 |
by (clarsimp simp: approx_def map_le_def dom_def) |
|
64 |
||
65 |
lemma restrict_map_le [intro!, simp]: "t |` S \<subseteq>\<^sub>m t" |
|
66 |
by (clarsimp simp: restrict_map_def map_le_def) |
|
67 |
||
68 |
lemma merge_restrict: |
|
69 |
assumes "t1 |` S = t |` S" |
|
70 |
assumes "t2 |` S = t |` S" |
|
71 |
shows "merge t1 t2 |` S = t |` S" |
|
72 |
proof - |
|
73 |
from assms |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
74 |
have "\<forall>x. (t1 |` S) x = (t |` S) x" |
44070 | 75 |
and "\<forall>x. (t2 |` S) x = (t |` S) x" by auto |
76 |
thus ?thesis |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
77 |
by (auto simp: merge_def restrict_map_def |
44070 | 78 |
split: if_splits intro: ext) |
79 |
qed |
|
80 |
||
81 |
||
82 |
lemma defs_restrict: |
|
83 |
"defs c t |` (- lnames c) = t |` (- lnames c)" |
|
45015 | 84 |
proof (induction c arbitrary: t) |
47818 | 85 |
case (Seq c1 c2) |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
86 |
hence "defs c1 t |` (- lnames c1) = t |` (- lnames c1)" |
44070 | 87 |
by simp |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
88 |
hence "defs c1 t |` (- lnames c1) |` (-lnames c2) = |
44070 | 89 |
t |` (- lnames c1) |` (-lnames c2)" by simp |
90 |
moreover |
|
47818 | 91 |
from Seq |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
92 |
have "defs c2 (defs c1 t) |` (- lnames c2) = |
44070 | 93 |
defs c1 t |` (- lnames c2)" |
94 |
by simp |
|
95 |
hence "defs c2 (defs c1 t) |` (- lnames c2) |` (- lnames c1) = |
|
96 |
defs c1 t |` (- lnames c2) |` (- lnames c1)" |
|
97 |
by simp |
|
98 |
ultimately |
|
99 |
show ?case by (clarsimp simp: Int_commute) |
|
100 |
next |
|
101 |
case (If b c1 c2) |
|
102 |
hence "defs c1 t |` (- lnames c1) = t |` (- lnames c1)" by simp |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
103 |
hence "defs c1 t |` (- lnames c1) |` (-lnames c2) = |
44070 | 104 |
t |` (- lnames c1) |` (-lnames c2)" by simp |
105 |
moreover |
|
106 |
from If |
|
107 |
have "defs c2 t |` (- lnames c2) = t |` (- lnames c2)" by simp |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
108 |
hence "defs c2 t |` (- lnames c2) |` (-lnames c1) = |
44070 | 109 |
t |` (- lnames c2) |` (-lnames c1)" by simp |
110 |
ultimately |
|
111 |
show ?case by (auto simp: Int_commute intro: merge_restrict) |
|
112 |
qed (auto split: aexp.split) |
|
113 |
||
114 |
||
115 |
lemma big_step_pres_approx: |
|
116 |
"(c,s) \<Rightarrow> s' \<Longrightarrow> approx t s \<Longrightarrow> approx (defs c t) s'" |
|
45015 | 117 |
proof (induction arbitrary: t rule: big_step_induct) |
44070 | 118 |
case Skip thus ?case by simp |
119 |
next |
|
120 |
case Assign |
|
121 |
thus ?case |
|
122 |
by (clarsimp simp: aval_simp_const_N approx_def split: aexp.split) |
|
123 |
next |
|
47818 | 124 |
case (Seq c1 s1 s2 c2 s3) |
125 |
have "approx (defs c1 t) s2" by (rule Seq.IH(1)[OF Seq.prems]) |
|
126 |
hence "approx (defs c2 (defs c1 t)) s3" by (rule Seq.IH(2)) |
|
44070 | 127 |
thus ?case by simp |
128 |
next |
|
129 |
case (IfTrue b s c1 s') |
|
130 |
hence "approx (defs c1 t) s'" by simp |
|
131 |
thus ?case by (simp add: approx_merge) |
|
132 |
next |
|
133 |
case (IfFalse b s c2 s') |
|
134 |
hence "approx (defs c2 t) s'" by simp |
|
135 |
thus ?case by (simp add: approx_merge) |
|
136 |
next |
|
137 |
case WhileFalse |
|
138 |
thus ?case by (simp add: approx_def restrict_map_def) |
|
139 |
next |
|
140 |
case (WhileTrue b s1 c s2 s3) |
|
141 |
hence "approx (defs c t) s2" by simp |
|
142 |
with WhileTrue |
|
143 |
have "approx (defs c t |` (-lnames c)) s3" by simp |
|
144 |
thus ?case by (simp add: defs_restrict) |
|
145 |
qed |
|
146 |
||
147 |
||
148 |
lemma big_step_pres_approx_restrict: |
|
149 |
"(c,s) \<Rightarrow> s' \<Longrightarrow> approx (t |` (-lnames c)) s \<Longrightarrow> approx (t |` (-lnames c)) s'" |
|
45015 | 150 |
proof (induction arbitrary: t rule: big_step_induct) |
44070 | 151 |
case Assign |
152 |
thus ?case by (clarsimp simp: approx_def) |
|
153 |
next |
|
47818 | 154 |
case (Seq c1 s1 s2 c2 s3) |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
155 |
hence "approx (t |` (-lnames c2) |` (-lnames c1)) s1" |
44070 | 156 |
by (simp add: Int_commute) |
157 |
hence "approx (t |` (-lnames c2) |` (-lnames c1)) s2" |
|
47818 | 158 |
by (rule Seq) |
44070 | 159 |
hence "approx (t |` (-lnames c1) |` (-lnames c2)) s2" |
160 |
by (simp add: Int_commute) |
|
161 |
hence "approx (t |` (-lnames c1) |` (-lnames c2)) s3" |
|
47818 | 162 |
by (rule Seq) |
44070 | 163 |
thus ?case by simp |
164 |
next |
|
165 |
case (IfTrue b s c1 s' c2) |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
166 |
hence "approx (t |` (-lnames c2) |` (-lnames c1)) s" |
44070 | 167 |
by (simp add: Int_commute) |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
168 |
hence "approx (t |` (-lnames c2) |` (-lnames c1)) s'" |
44070 | 169 |
by (rule IfTrue) |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
170 |
thus ?case by (simp add: Int_commute) |
44070 | 171 |
next |
172 |
case (IfFalse b s c2 s' c1) |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
173 |
hence "approx (t |` (-lnames c1) |` (-lnames c2)) s" |
44070 | 174 |
by simp |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
175 |
hence "approx (t |` (-lnames c1) |` (-lnames c2)) s'" |
44070 | 176 |
by (rule IfFalse) |
177 |
thus ?case by simp |
|
178 |
qed auto |
|
179 |
||
180 |
||
181 |
declare assign_simp [simp] |
|
182 |
||
183 |
lemma approx_eq: |
|
184 |
"approx t \<Turnstile> c \<sim> fold c t" |
|
45015 | 185 |
proof (induction c arbitrary: t) |
44070 | 186 |
case SKIP show ?case by simp |
187 |
next |
|
188 |
case Assign |
|
189 |
show ?case by (simp add: equiv_up_to_def) |
|
190 |
next |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
191 |
case Seq |
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
192 |
thus ?case by (auto intro!: equiv_up_to_seq big_step_pres_approx) |
44070 | 193 |
next |
194 |
case If |
|
195 |
thus ?case by (auto intro!: equiv_up_to_if_weak) |
|
196 |
next |
|
197 |
case (While b c) |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
198 |
hence "approx (t |` (- lnames c)) \<Turnstile> |
44070 | 199 |
WHILE b DO c \<sim> WHILE b DO fold c (t |` (- lnames c))" |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
200 |
by (auto intro: equiv_up_to_while_weak big_step_pres_approx_restrict) |
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
201 |
thus ?case |
44070 | 202 |
by (auto intro: equiv_up_to_weaken approx_map_le) |
203 |
qed |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
204 |
|
44070 | 205 |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
206 |
lemma approx_empty [simp]: |
44070 | 207 |
"approx empty = (\<lambda>_. True)" |
208 |
by (auto intro!: ext simp: approx_def) |
|
209 |
||
210 |
lemma equiv_sym: |
|
211 |
"c \<sim> c' \<longleftrightarrow> c' \<sim> c" |
|
212 |
by (auto simp add: equiv_def) |
|
213 |
||
214 |
theorem constant_folding_equiv: |
|
215 |
"fold c empty \<sim> c" |
|
216 |
using approx_eq [of empty c] |
|
217 |
by (simp add: equiv_up_to_True equiv_sym) |
|
218 |
||
219 |
||
220 |
||
44850 | 221 |
subsection {* More ambitious folding including boolean expressions *} |
44070 | 222 |
|
223 |
||
224 |
fun bsimp_const :: "bexp \<Rightarrow> tab \<Rightarrow> bexp" where |
|
225 |
"bsimp_const (Less a1 a2) t = less (simp_const a1 t) (simp_const a2 t)" | |
|
226 |
"bsimp_const (And b1 b2) t = and (bsimp_const b1 t) (bsimp_const b2 t)" | |
|
227 |
"bsimp_const (Not b) t = not(bsimp_const b t)" | |
|
45200 | 228 |
"bsimp_const (Bc bc) _ = Bc bc" |
44070 | 229 |
|
230 |
theorem bvalsimp_const [simp]: |
|
231 |
assumes "approx t s" |
|
232 |
shows "bval (bsimp_const b t) s = bval b s" |
|
233 |
using assms by (induct b) auto |
|
234 |
||
45200 | 235 |
lemma not_Bc [simp]: "not (Bc v) = Bc (\<not>v)" |
44070 | 236 |
by (cases v) auto |
237 |
||
45200 | 238 |
lemma not_Bc_eq [simp]: "(not b = Bc v) = (b = Bc (\<not>v))" |
44070 | 239 |
by (cases b) auto |
240 |
||
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
241 |
lemma and_Bc_eq: |
45200 | 242 |
"(and a b = Bc v) = |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
243 |
(a = Bc False \<and> \<not>v \<or> b = Bc False \<and> \<not>v \<or> |
45200 | 244 |
(\<exists>v1 v2. a = Bc v1 \<and> b = Bc v2 \<and> v = v1 \<and> v2))" |
44070 | 245 |
by (rule and.induct) auto |
246 |
||
45200 | 247 |
lemma less_Bc_eq [simp]: |
248 |
"(less a b = Bc v) = (\<exists>n1 n2. a = N n1 \<and> b = N n2 \<and> v = (n1 < n2))" |
|
44070 | 249 |
by (rule less.induct) auto |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
250 |
|
45200 | 251 |
theorem bvalsimp_const_Bc: |
44070 | 252 |
assumes "approx t s" |
45200 | 253 |
shows "bsimp_const b t = Bc v \<Longrightarrow> bval b s = v" |
44070 | 254 |
using assms |
255 |
by (induct b arbitrary: v) |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
256 |
(auto simp: approx_def and_Bc_eq aval_simp_const_N |
44070 | 257 |
split: bexp.splits bool.splits) |
258 |
||
259 |
||
260 |
primrec "bdefs" :: "com \<Rightarrow> tab \<Rightarrow> tab" where |
|
261 |
"bdefs SKIP t = t" | |
|
262 |
"bdefs (x ::= a) t = |
|
263 |
(case simp_const a t of N k \<Rightarrow> t(x \<mapsto> k) | _ \<Rightarrow> t(x:=None))" | |
|
264 |
"bdefs (c1;c2) t = (bdefs c2 o bdefs c1) t" | |
|
265 |
"bdefs (IF b THEN c1 ELSE c2) t = (case bsimp_const b t of |
|
45200 | 266 |
Bc True \<Rightarrow> bdefs c1 t |
267 |
| Bc False \<Rightarrow> bdefs c2 t |
|
44070 | 268 |
| _ \<Rightarrow> merge (bdefs c1 t) (bdefs c2 t))" | |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
269 |
"bdefs (WHILE b DO c) t = t |` (-lnames c)" |
44070 | 270 |
|
271 |
||
272 |
primrec bfold where |
|
273 |
"bfold SKIP _ = SKIP" | |
|
274 |
"bfold (x ::= a) t = (x ::= (simp_const a t))" | |
|
275 |
"bfold (c1;c2) t = (bfold c1 t; bfold c2 (bdefs c1 t))" | |
|
276 |
"bfold (IF b THEN c1 ELSE c2) t = (case bsimp_const b t of |
|
45200 | 277 |
Bc True \<Rightarrow> bfold c1 t |
278 |
| Bc False \<Rightarrow> bfold c2 t |
|
44070 | 279 |
| _ \<Rightarrow> IF bsimp_const b t THEN bfold c1 t ELSE bfold c2 t)" | |
280 |
"bfold (WHILE b DO c) t = (case bsimp_const b t of |
|
45200 | 281 |
Bc False \<Rightarrow> SKIP |
44070 | 282 |
| _ \<Rightarrow> WHILE bsimp_const b (t |` (-lnames c)) DO bfold c (t |` (-lnames c)))" |
283 |
||
284 |
||
285 |
lemma bdefs_restrict: |
|
286 |
"bdefs c t |` (- lnames c) = t |` (- lnames c)" |
|
45015 | 287 |
proof (induction c arbitrary: t) |
47818 | 288 |
case (Seq c1 c2) |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
289 |
hence "bdefs c1 t |` (- lnames c1) = t |` (- lnames c1)" |
44070 | 290 |
by simp |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
291 |
hence "bdefs c1 t |` (- lnames c1) |` (-lnames c2) = |
44070 | 292 |
t |` (- lnames c1) |` (-lnames c2)" by simp |
293 |
moreover |
|
47818 | 294 |
from Seq |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
295 |
have "bdefs c2 (bdefs c1 t) |` (- lnames c2) = |
44070 | 296 |
bdefs c1 t |` (- lnames c2)" |
297 |
by simp |
|
298 |
hence "bdefs c2 (bdefs c1 t) |` (- lnames c2) |` (- lnames c1) = |
|
299 |
bdefs c1 t |` (- lnames c2) |` (- lnames c1)" |
|
300 |
by simp |
|
301 |
ultimately |
|
302 |
show ?case by (clarsimp simp: Int_commute) |
|
303 |
next |
|
304 |
case (If b c1 c2) |
|
305 |
hence "bdefs c1 t |` (- lnames c1) = t |` (- lnames c1)" by simp |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
306 |
hence "bdefs c1 t |` (- lnames c1) |` (-lnames c2) = |
44070 | 307 |
t |` (- lnames c1) |` (-lnames c2)" by simp |
308 |
moreover |
|
309 |
from If |
|
310 |
have "bdefs c2 t |` (- lnames c2) = t |` (- lnames c2)" by simp |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
311 |
hence "bdefs c2 t |` (- lnames c2) |` (-lnames c1) = |
44070 | 312 |
t |` (- lnames c2) |` (-lnames c1)" by simp |
313 |
ultimately |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
314 |
show ?case |
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
315 |
by (auto simp: Int_commute intro: merge_restrict |
44070 | 316 |
split: bexp.splits bool.splits) |
317 |
qed (auto split: aexp.split bexp.split bool.split) |
|
318 |
||
319 |
||
320 |
lemma big_step_pres_approx_b: |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
321 |
"(c,s) \<Rightarrow> s' \<Longrightarrow> approx t s \<Longrightarrow> approx (bdefs c t) s'" |
45015 | 322 |
proof (induction arbitrary: t rule: big_step_induct) |
44070 | 323 |
case Skip thus ?case by simp |
324 |
next |
|
325 |
case Assign |
|
326 |
thus ?case |
|
327 |
by (clarsimp simp: aval_simp_const_N approx_def split: aexp.split) |
|
328 |
next |
|
47818 | 329 |
case (Seq c1 s1 s2 c2 s3) |
330 |
have "approx (bdefs c1 t) s2" by (rule Seq.IH(1)[OF Seq.prems]) |
|
331 |
hence "approx (bdefs c2 (bdefs c1 t)) s3" by (rule Seq.IH(2)) |
|
44070 | 332 |
thus ?case by simp |
333 |
next |
|
334 |
case (IfTrue b s c1 s') |
|
335 |
hence "approx (bdefs c1 t) s'" by simp |
|
336 |
thus ?case using `bval b s` `approx t s` |
|
45200 | 337 |
by (clarsimp simp: approx_merge bvalsimp_const_Bc |
44070 | 338 |
split: bexp.splits bool.splits) |
339 |
next |
|
340 |
case (IfFalse b s c2 s') |
|
341 |
hence "approx (bdefs c2 t) s'" by simp |
|
342 |
thus ?case using `\<not>bval b s` `approx t s` |
|
45200 | 343 |
by (clarsimp simp: approx_merge bvalsimp_const_Bc |
44070 | 344 |
split: bexp.splits bool.splits) |
345 |
next |
|
346 |
case WhileFalse |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
347 |
thus ?case |
45200 | 348 |
by (clarsimp simp: bvalsimp_const_Bc approx_def restrict_map_def |
44070 | 349 |
split: bexp.splits bool.splits) |
350 |
next |
|
351 |
case (WhileTrue b s1 c s2 s3) |
|
352 |
hence "approx (bdefs c t) s2" by simp |
|
353 |
with WhileTrue |
|
354 |
have "approx (bdefs c t |` (- lnames c)) s3" |
|
355 |
by simp |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
356 |
thus ?case |
44070 | 357 |
by (simp add: bdefs_restrict) |
358 |
qed |
|
359 |
||
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
360 |
lemma bfold_equiv: |
44070 | 361 |
"approx t \<Turnstile> c \<sim> bfold c t" |
45015 | 362 |
proof (induction c arbitrary: t) |
44070 | 363 |
case SKIP show ?case by simp |
364 |
next |
|
365 |
case Assign |
|
366 |
thus ?case by (simp add: equiv_up_to_def) |
|
367 |
next |
|
47818 | 368 |
case Seq |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
369 |
thus ?case by (auto intro!: equiv_up_to_seq big_step_pres_approx_b) |
44070 | 370 |
next |
371 |
case (If b c1 c2) |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
372 |
hence "approx t \<Turnstile> IF b THEN c1 ELSE c2 \<sim> |
44070 | 373 |
IF Fold.bsimp_const b t THEN bfold c1 t ELSE bfold c2 t" |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
374 |
by (auto intro: equiv_up_to_if_weak simp: bequiv_up_to_def) |
44070 | 375 |
thus ?case using If |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
376 |
by (fastforce simp: bvalsimp_const_Bc split: bexp.splits bool.splits |
44070 | 377 |
intro: equiv_up_to_trans) |
378 |
next |
|
379 |
case (While b c) |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
380 |
hence "approx (t |` (- lnames c)) \<Turnstile> |
44070 | 381 |
WHILE b DO c \<sim> |
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
382 |
WHILE bsimp_const b (t |` (- lnames c)) |
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
383 |
DO bfold c (t |` (- lnames c))" (is "_ \<Turnstile> ?W \<sim> ?W'") |
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
384 |
by (auto intro: equiv_up_to_while_weak big_step_pres_approx_restrict |
44070 | 385 |
simp: bequiv_up_to_def) |
386 |
hence "approx t \<Turnstile> ?W \<sim> ?W'" |
|
387 |
by (auto intro: equiv_up_to_weaken approx_map_le) |
|
388 |
thus ?case |
|
48909
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
389 |
by (auto split: bexp.splits bool.splits |
b2dea007e55e
remove Hoare dependency from Fold.thy
Gerwin Klein <gerwin.klein@nicta.com.au>
parents:
47818
diff
changeset
|
390 |
intro: equiv_up_to_while_False |
45200 | 391 |
simp: bvalsimp_const_Bc) |
44070 | 392 |
qed |
393 |
||
394 |
||
395 |
theorem constant_bfolding_equiv: |
|
396 |
"bfold c empty \<sim> c" |
|
397 |
using bfold_equiv [of empty c] |
|
398 |
by (simp add: equiv_up_to_True equiv_sym) |
|
399 |
||
400 |
||
401 |
end |