author | nipkow |
Fri, 10 Aug 2012 17:17:05 +0200 | |
changeset 48759 | ff570720ba1c |
parent 47818 | 151d137f1095 |
child 49095 | 7df19036392e |
permissions | -rw-r--r-- |
47613 | 1 |
(* Author: Tobias Nipkow *) |
2 |
||
3 |
theory Abs_Int0 |
|
4 |
imports Abs_Int_init |
|
5 |
begin |
|
6 |
||
7 |
subsection "Orderings" |
|
8 |
||
9 |
class preord = |
|
10 |
fixes le :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubseteq>" 50) |
|
11 |
assumes le_refl[simp]: "x \<sqsubseteq> x" |
|
12 |
and le_trans: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z" |
|
13 |
begin |
|
14 |
||
15 |
definition mono where "mono f = (\<forall>x y. x \<sqsubseteq> y \<longrightarrow> f x \<sqsubseteq> f y)" |
|
16 |
||
17 |
declare le_trans[trans] |
|
18 |
||
19 |
end |
|
20 |
||
21 |
text{* Note: no antisymmetry. Allows implementations where some abstract |
|
22 |
element is implemented by two different values @{prop "x \<noteq> y"} |
|
23 |
such that @{prop"x \<sqsubseteq> y"} and @{prop"y \<sqsubseteq> x"}. Antisymmetry is not |
|
24 |
needed because we never compare elements for equality but only for @{text"\<sqsubseteq>"}. |
|
25 |
*} |
|
26 |
||
27 |
class join = preord + |
|
28 |
fixes join :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<squnion>" 65) |
|
29 |
||
30 |
class SL_top = join + |
|
31 |
fixes Top :: "'a" ("\<top>") |
|
32 |
assumes join_ge1 [simp]: "x \<sqsubseteq> x \<squnion> y" |
|
33 |
and join_ge2 [simp]: "y \<sqsubseteq> x \<squnion> y" |
|
34 |
and join_least: "x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<squnion> y \<sqsubseteq> z" |
|
35 |
and top[simp]: "x \<sqsubseteq> \<top>" |
|
36 |
begin |
|
37 |
||
38 |
lemma join_le_iff[simp]: "x \<squnion> y \<sqsubseteq> z \<longleftrightarrow> x \<sqsubseteq> z \<and> y \<sqsubseteq> z" |
|
39 |
by (metis join_ge1 join_ge2 join_least le_trans) |
|
40 |
||
41 |
lemma le_join_disj: "x \<sqsubseteq> y \<or> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<squnion> z" |
|
42 |
by (metis join_ge1 join_ge2 le_trans) |
|
43 |
||
44 |
end |
|
45 |
||
46 |
instantiation "fun" :: (type, preord) preord |
|
47 |
begin |
|
48 |
||
49 |
definition "f \<sqsubseteq> g = (\<forall>x. f x \<sqsubseteq> g x)" |
|
50 |
||
51 |
instance |
|
52 |
proof |
|
53 |
case goal2 thus ?case by (metis le_fun_def preord_class.le_trans) |
|
54 |
qed (simp_all add: le_fun_def) |
|
55 |
||
56 |
end |
|
57 |
||
58 |
||
59 |
instantiation "fun" :: (type, SL_top) SL_top |
|
60 |
begin |
|
61 |
||
62 |
definition "f \<squnion> g = (\<lambda>x. f x \<squnion> g x)" |
|
63 |
definition "\<top> = (\<lambda>x. \<top>)" |
|
64 |
||
65 |
lemma join_apply[simp]: "(f \<squnion> g) x = f x \<squnion> g x" |
|
66 |
by (simp add: join_fun_def) |
|
67 |
||
68 |
instance |
|
69 |
proof |
|
70 |
qed (simp_all add: le_fun_def Top_fun_def) |
|
71 |
||
72 |
end |
|
73 |
||
74 |
||
75 |
instantiation acom :: (preord) preord |
|
76 |
begin |
|
77 |
||
78 |
fun le_acom :: "('a::preord)acom \<Rightarrow> 'a acom \<Rightarrow> bool" where |
|
79 |
"le_acom (SKIP {S}) (SKIP {S'}) = (S \<sqsubseteq> S')" | |
|
80 |
"le_acom (x ::= e {S}) (x' ::= e' {S'}) = (x=x' \<and> e=e' \<and> S \<sqsubseteq> S')" | |
|
81 |
"le_acom (C1;C2) (D1;D2) = (C1 \<sqsubseteq> D1 \<and> C2 \<sqsubseteq> D2)" | |
|
82 |
"le_acom (IF b THEN C1 ELSE C2 {S}) (IF b' THEN D1 ELSE D2 {S'}) = |
|
83 |
(b=b' \<and> C1 \<sqsubseteq> D1 \<and> C2 \<sqsubseteq> D2 \<and> S \<sqsubseteq> S')" | |
|
84 |
"le_acom ({I} WHILE b DO C {P}) ({I'} WHILE b' DO C' {P'}) = |
|
85 |
(b=b' \<and> C \<sqsubseteq> C' \<and> I \<sqsubseteq> I' \<and> P \<sqsubseteq> P')" | |
|
86 |
"le_acom _ _ = False" |
|
87 |
||
88 |
lemma [simp]: "SKIP {S} \<sqsubseteq> C \<longleftrightarrow> (\<exists>S'. C = SKIP {S'} \<and> S \<sqsubseteq> S')" |
|
89 |
by (cases C) auto |
|
90 |
||
91 |
lemma [simp]: "x ::= e {S} \<sqsubseteq> C \<longleftrightarrow> (\<exists>S'. C = x ::= e {S'} \<and> S \<sqsubseteq> S')" |
|
92 |
by (cases C) auto |
|
93 |
||
94 |
lemma [simp]: "C1;C2 \<sqsubseteq> C \<longleftrightarrow> (\<exists>D1 D2. C = D1;D2 \<and> C1 \<sqsubseteq> D1 \<and> C2 \<sqsubseteq> D2)" |
|
95 |
by (cases C) auto |
|
96 |
||
97 |
lemma [simp]: "IF b THEN C1 ELSE C2 {S} \<sqsubseteq> C \<longleftrightarrow> |
|
98 |
(\<exists>D1 D2 S'. C = IF b THEN D1 ELSE D2 {S'} \<and> C1 \<sqsubseteq> D1 \<and> C2 \<sqsubseteq> D2 \<and> S \<sqsubseteq> S')" |
|
99 |
by (cases C) auto |
|
100 |
||
101 |
lemma [simp]: "{I} WHILE b DO C {P} \<sqsubseteq> W \<longleftrightarrow> |
|
102 |
(\<exists>I' C' P'. W = {I'} WHILE b DO C' {P'} \<and> C \<sqsubseteq> C' \<and> I \<sqsubseteq> I' \<and> P \<sqsubseteq> P')" |
|
103 |
by (cases W) auto |
|
104 |
||
105 |
instance |
|
106 |
proof |
|
107 |
case goal1 thus ?case by (induct x) auto |
|
108 |
next |
|
109 |
case goal2 thus ?case |
|
110 |
apply(induct x y arbitrary: z rule: le_acom.induct) |
|
111 |
apply (auto intro: le_trans) |
|
112 |
done |
|
113 |
qed |
|
114 |
||
115 |
end |
|
116 |
||
117 |
||
118 |
instantiation option :: (preord)preord |
|
119 |
begin |
|
120 |
||
121 |
fun le_option where |
|
122 |
"Some x \<sqsubseteq> Some y = (x \<sqsubseteq> y)" | |
|
123 |
"None \<sqsubseteq> y = True" | |
|
124 |
"Some _ \<sqsubseteq> None = False" |
|
125 |
||
126 |
lemma [simp]: "(x \<sqsubseteq> None) = (x = None)" |
|
127 |
by (cases x) simp_all |
|
128 |
||
129 |
lemma [simp]: "(Some x \<sqsubseteq> u) = (\<exists>y. u = Some y \<and> x \<sqsubseteq> y)" |
|
130 |
by (cases u) auto |
|
131 |
||
132 |
instance proof |
|
133 |
case goal1 show ?case by(cases x, simp_all) |
|
134 |
next |
|
135 |
case goal2 thus ?case |
|
136 |
by(cases z, simp, cases y, simp, cases x, auto intro: le_trans) |
|
137 |
qed |
|
138 |
||
139 |
end |
|
140 |
||
141 |
instantiation option :: (join)join |
|
142 |
begin |
|
143 |
||
144 |
fun join_option where |
|
145 |
"Some x \<squnion> Some y = Some(x \<squnion> y)" | |
|
146 |
"None \<squnion> y = y" | |
|
147 |
"x \<squnion> None = x" |
|
148 |
||
149 |
lemma join_None2[simp]: "x \<squnion> None = x" |
|
150 |
by (cases x) simp_all |
|
151 |
||
152 |
instance .. |
|
153 |
||
154 |
end |
|
155 |
||
156 |
instantiation option :: (SL_top)SL_top |
|
157 |
begin |
|
158 |
||
159 |
definition "\<top> = Some \<top>" |
|
160 |
||
161 |
instance proof |
|
162 |
case goal1 thus ?case by(cases x, simp, cases y, simp_all) |
|
163 |
next |
|
164 |
case goal2 thus ?case by(cases y, simp, cases x, simp_all) |
|
165 |
next |
|
166 |
case goal3 thus ?case by(cases z, simp, cases y, simp, cases x, simp_all) |
|
167 |
next |
|
168 |
case goal4 thus ?case by(cases x, simp_all add: Top_option_def) |
|
169 |
qed |
|
170 |
||
171 |
end |
|
172 |
||
173 |
class Bot = preord + |
|
174 |
fixes Bot :: "'a" ("\<bottom>") |
|
175 |
assumes Bot[simp]: "\<bottom> \<sqsubseteq> x" |
|
176 |
||
177 |
instantiation option :: (preord)Bot |
|
178 |
begin |
|
179 |
||
180 |
definition Bot_option :: "'a option" where |
|
181 |
"\<bottom> = None" |
|
182 |
||
183 |
instance |
|
184 |
proof |
|
185 |
case goal1 thus ?case by(auto simp: Bot_option_def) |
|
186 |
qed |
|
187 |
||
188 |
end |
|
189 |
||
190 |
||
191 |
definition bot :: "com \<Rightarrow> 'a option acom" where |
|
192 |
"bot c = anno None c" |
|
193 |
||
194 |
lemma bot_least: "strip C = c \<Longrightarrow> bot c \<sqsubseteq> C" |
|
195 |
by(induct C arbitrary: c)(auto simp: bot_def) |
|
196 |
||
197 |
lemma strip_bot[simp]: "strip(bot c) = c" |
|
198 |
by(simp add: bot_def) |
|
199 |
||
200 |
||
201 |
subsubsection "Post-fixed point iteration" |
|
202 |
||
203 |
definition |
|
204 |
pfp :: "(('a::preord) \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a option" where |
|
205 |
"pfp f = while_option (\<lambda>x. \<not> f x \<sqsubseteq> x) f" |
|
206 |
||
207 |
lemma pfp_pfp: assumes "pfp f x0 = Some x" shows "f x \<sqsubseteq> x" |
|
208 |
using while_option_stop[OF assms[simplified pfp_def]] by simp |
|
209 |
||
210 |
lemma pfp_least: |
|
211 |
assumes mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y" |
|
212 |
and "f p \<sqsubseteq> p" and "x0 \<sqsubseteq> p" and "pfp f x0 = Some x" shows "x \<sqsubseteq> p" |
|
213 |
proof- |
|
214 |
{ fix x assume "x \<sqsubseteq> p" |
|
215 |
hence "f x \<sqsubseteq> f p" by(rule mono) |
|
216 |
from this `f p \<sqsubseteq> p` have "f x \<sqsubseteq> p" by(rule le_trans) |
|
217 |
} |
|
218 |
thus "x \<sqsubseteq> p" using assms(2-) while_option_rule[where P = "%x. x \<sqsubseteq> p"] |
|
219 |
unfolding pfp_def by blast |
|
220 |
qed |
|
221 |
||
222 |
definition |
|
223 |
lpfp :: "('a::preord option acom \<Rightarrow> 'a option acom) \<Rightarrow> com \<Rightarrow> 'a option acom option" |
|
224 |
where "lpfp f c = pfp f (bot c)" |
|
225 |
||
226 |
lemma lpfpc_pfp: "lpfp f c = Some p \<Longrightarrow> f p \<sqsubseteq> p" |
|
227 |
by(simp add: pfp_pfp lpfp_def) |
|
228 |
||
229 |
lemma strip_pfp: |
|
230 |
assumes "\<And>x. g(f x) = g x" and "pfp f x0 = Some x" shows "g x = g x0" |
|
231 |
using assms while_option_rule[where P = "%x. g x = g x0" and c = f] |
|
232 |
unfolding pfp_def by metis |
|
233 |
||
234 |
lemma strip_lpfp: assumes "\<And>C. strip(f C) = strip C" and "lpfp f c = Some C" |
|
235 |
shows "strip C = c" |
|
236 |
using assms(1) strip_pfp[OF _ assms(2)[simplified lpfp_def]] |
|
237 |
by(metis strip_bot) |
|
238 |
||
239 |
||
240 |
subsection "Abstract Interpretation" |
|
241 |
||
242 |
definition \<gamma>_fun :: "('a \<Rightarrow> 'b set) \<Rightarrow> ('c \<Rightarrow> 'a) \<Rightarrow> ('c \<Rightarrow> 'b)set" where |
|
243 |
"\<gamma>_fun \<gamma> F = {f. \<forall>x. f x \<in> \<gamma>(F x)}" |
|
244 |
||
245 |
fun \<gamma>_option :: "('a \<Rightarrow> 'b set) \<Rightarrow> 'a option \<Rightarrow> 'b set" where |
|
246 |
"\<gamma>_option \<gamma> None = {}" | |
|
247 |
"\<gamma>_option \<gamma> (Some a) = \<gamma> a" |
|
248 |
||
249 |
text{* The interface for abstract values: *} |
|
250 |
||
251 |
locale Val_abs = |
|
252 |
fixes \<gamma> :: "'av::SL_top \<Rightarrow> val set" |
|
253 |
assumes mono_gamma: "a \<sqsubseteq> b \<Longrightarrow> \<gamma> a \<subseteq> \<gamma> b" |
|
254 |
and gamma_Top[simp]: "\<gamma> \<top> = UNIV" |
|
255 |
fixes num' :: "val \<Rightarrow> 'av" |
|
256 |
and plus' :: "'av \<Rightarrow> 'av \<Rightarrow> 'av" |
|
257 |
assumes gamma_num': "n : \<gamma>(num' n)" |
|
258 |
and gamma_plus': |
|
259 |
"n1 : \<gamma> a1 \<Longrightarrow> n2 : \<gamma> a2 \<Longrightarrow> n1+n2 : \<gamma>(plus' a1 a2)" |
|
260 |
||
261 |
type_synonym 'av st = "(vname \<Rightarrow> 'av)" |
|
262 |
||
263 |
locale Abs_Int_Fun = Val_abs \<gamma> for \<gamma> :: "'av::SL_top \<Rightarrow> val set" |
|
264 |
begin |
|
265 |
||
266 |
fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where |
|
267 |
"aval' (N n) S = num' n" | |
|
268 |
"aval' (V x) S = S x" | |
|
269 |
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" |
|
270 |
||
271 |
fun step' :: "'av st option \<Rightarrow> 'av st option acom \<Rightarrow> 'av st option acom" |
|
272 |
where |
|
273 |
"step' S (SKIP {P}) = (SKIP {S})" | |
|
274 |
"step' S (x ::= e {P}) = |
|
275 |
x ::= e {case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(S(x := aval' e S))}" | |
|
276 |
"step' S (C1; C2) = step' S C1; step' (post C1) C2" | |
|
277 |
"step' S (IF b THEN C1 ELSE C2 {P}) = |
|
278 |
IF b THEN step' S C1 ELSE step' S C2 {post C1 \<squnion> post C2}" | |
|
279 |
"step' S ({Inv} WHILE b DO C {P}) = |
|
280 |
{S \<squnion> post C} WHILE b DO (step' Inv C) {Inv}" |
|
281 |
||
282 |
definition AI :: "com \<Rightarrow> 'av st option acom option" where |
|
283 |
"AI = lpfp (step' \<top>)" |
|
284 |
||
285 |
||
286 |
lemma strip_step'[simp]: "strip(step' S C) = strip C" |
|
287 |
by(induct C arbitrary: S) (simp_all add: Let_def) |
|
288 |
||
289 |
||
290 |
abbreviation \<gamma>\<^isub>f :: "'av st \<Rightarrow> state set" |
|
291 |
where "\<gamma>\<^isub>f == \<gamma>_fun \<gamma>" |
|
292 |
||
293 |
abbreviation \<gamma>\<^isub>o :: "'av st option \<Rightarrow> state set" |
|
294 |
where "\<gamma>\<^isub>o == \<gamma>_option \<gamma>\<^isub>f" |
|
295 |
||
296 |
abbreviation \<gamma>\<^isub>c :: "'av st option acom \<Rightarrow> state set acom" |
|
297 |
where "\<gamma>\<^isub>c == map_acom \<gamma>\<^isub>o" |
|
298 |
||
299 |
lemma gamma_f_Top[simp]: "\<gamma>\<^isub>f Top = UNIV" |
|
300 |
by(simp add: Top_fun_def \<gamma>_fun_def) |
|
301 |
||
302 |
lemma gamma_o_Top[simp]: "\<gamma>\<^isub>o Top = UNIV" |
|
303 |
by (simp add: Top_option_def) |
|
304 |
||
305 |
(* FIXME (maybe also le \<rightarrow> sqle?) *) |
|
306 |
||
307 |
lemma mono_gamma_f: "f1 \<sqsubseteq> f2 \<Longrightarrow> \<gamma>\<^isub>f f1 \<subseteq> \<gamma>\<^isub>f f2" |
|
308 |
by(auto simp: le_fun_def \<gamma>_fun_def dest: mono_gamma) |
|
309 |
||
310 |
lemma mono_gamma_o: |
|
311 |
"S1 \<sqsubseteq> S2 \<Longrightarrow> \<gamma>\<^isub>o S1 \<subseteq> \<gamma>\<^isub>o S2" |
|
312 |
by(induction S1 S2 rule: le_option.induct)(simp_all add: mono_gamma_f) |
|
313 |
||
314 |
lemma mono_gamma_c: "C1 \<sqsubseteq> C2 \<Longrightarrow> \<gamma>\<^isub>c C1 \<le> \<gamma>\<^isub>c C2" |
|
315 |
by (induction C1 C2 rule: le_acom.induct) (simp_all add:mono_gamma_o) |
|
316 |
||
317 |
text{* Soundness: *} |
|
318 |
||
319 |
lemma aval'_sound: "s : \<gamma>\<^isub>f S \<Longrightarrow> aval a s : \<gamma>(aval' a S)" |
|
320 |
by (induct a) (auto simp: gamma_num' gamma_plus' \<gamma>_fun_def) |
|
321 |
||
322 |
lemma in_gamma_update: |
|
323 |
"\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(S(x := a))" |
|
324 |
by(simp add: \<gamma>_fun_def) |
|
325 |
||
326 |
lemma step_preserves_le: |
|
327 |
"\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; C \<le> \<gamma>\<^isub>c C' \<rbrakk> \<Longrightarrow> step S C \<le> \<gamma>\<^isub>c (step' S' C')" |
|
328 |
proof(induction C arbitrary: C' S S') |
|
329 |
case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP) |
|
330 |
next |
|
331 |
case Assign thus ?case |
|
332 |
by (fastforce simp: Assign_le map_acom_Assign intro: aval'_sound in_gamma_update |
|
333 |
split: option.splits del:subsetD) |
|
334 |
next |
|
47818 | 335 |
case Seq thus ?case apply (auto simp: Seq_le map_acom_Seq) |
47613 | 336 |
by (metis le_post post_map_acom) |
337 |
next |
|
338 |
case (If b C1 C2 P) |
|
339 |
then obtain D1 D2 P' where |
|
340 |
"C' = IF b THEN D1 ELSE D2 {P'}" |
|
341 |
"P \<subseteq> \<gamma>\<^isub>o P'" "C1 \<le> \<gamma>\<^isub>c D1" "C2 \<le> \<gamma>\<^isub>c D2" |
|
342 |
by (fastforce simp: If_le map_acom_If) |
|
343 |
moreover have "post C1 \<subseteq> \<gamma>\<^isub>o(post D1 \<squnion> post D2)" |
|
344 |
by (metis (no_types) `C1 \<le> \<gamma>\<^isub>c D1` join_ge1 le_post mono_gamma_o order_trans post_map_acom) |
|
345 |
moreover have "post C2 \<subseteq> \<gamma>\<^isub>o(post D1 \<squnion> post D2)" |
|
346 |
by (metis (no_types) `C2 \<le> \<gamma>\<^isub>c D2` join_ge2 le_post mono_gamma_o order_trans post_map_acom) |
|
347 |
ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'` by (simp add: If.IH subset_iff) |
|
348 |
next |
|
349 |
case (While I b C1 P) |
|
350 |
then obtain D1 I' P' where |
|
351 |
"C' = {I'} WHILE b DO D1 {P'}" |
|
352 |
"I \<subseteq> \<gamma>\<^isub>o I'" "P \<subseteq> \<gamma>\<^isub>o P'" "C1 \<le> \<gamma>\<^isub>c D1" |
|
353 |
by (fastforce simp: map_acom_While While_le) |
|
354 |
moreover have "S \<union> post C1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post D1)" |
|
355 |
using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `C1 \<le> \<gamma>\<^isub>c D1`, simplified] |
|
356 |
by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans) |
|
357 |
ultimately show ?case by (simp add: While.IH subset_iff) |
|
358 |
qed |
|
359 |
||
360 |
lemma AI_sound: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c C" |
|
361 |
proof(simp add: CS_def AI_def) |
|
362 |
assume 1: "lpfp (step' \<top>) c = Some C" |
|
363 |
have 2: "step' \<top> C \<sqsubseteq> C" by(rule lpfpc_pfp[OF 1]) |
|
364 |
have 3: "strip (\<gamma>\<^isub>c (step' \<top> C)) = c" |
|
365 |
by(simp add: strip_lpfp[OF _ 1]) |
|
48759
ff570720ba1c
Improved complete lattice formalisation - no more index set.
nipkow
parents:
47818
diff
changeset
|
366 |
have "lfp c (step UNIV) \<le> \<gamma>\<^isub>c (step' \<top> C)" |
47613 | 367 |
proof(rule lfp_lowerbound[simplified,OF 3]) |
368 |
show "step UNIV (\<gamma>\<^isub>c (step' \<top> C)) \<le> \<gamma>\<^isub>c (step' \<top> C)" |
|
369 |
proof(rule step_preserves_le[OF _ _]) |
|
370 |
show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp |
|
371 |
show "\<gamma>\<^isub>c (step' \<top> C) \<le> \<gamma>\<^isub>c C" by(rule mono_gamma_c[OF 2]) |
|
372 |
qed |
|
373 |
qed |
|
48759
ff570720ba1c
Improved complete lattice formalisation - no more index set.
nipkow
parents:
47818
diff
changeset
|
374 |
with 2 show "lfp c (step UNIV) \<le> \<gamma>\<^isub>c C" |
47613 | 375 |
by (blast intro: mono_gamma_c order_trans) |
376 |
qed |
|
377 |
||
378 |
end |
|
379 |
||
380 |
||
381 |
subsubsection "Monotonicity" |
|
382 |
||
383 |
lemma mono_post: "C1 \<sqsubseteq> C2 \<Longrightarrow> post C1 \<sqsubseteq> post C2" |
|
384 |
by(induction C1 C2 rule: le_acom.induct) (auto) |
|
385 |
||
386 |
locale Abs_Int_Fun_mono = Abs_Int_Fun + |
|
387 |
assumes mono_plus': "a1 \<sqsubseteq> b1 \<Longrightarrow> a2 \<sqsubseteq> b2 \<Longrightarrow> plus' a1 a2 \<sqsubseteq> plus' b1 b2" |
|
388 |
begin |
|
389 |
||
390 |
lemma mono_aval': "S \<sqsubseteq> S' \<Longrightarrow> aval' e S \<sqsubseteq> aval' e S'" |
|
391 |
by(induction e)(auto simp: le_fun_def mono_plus') |
|
392 |
||
393 |
lemma mono_update: "a \<sqsubseteq> a' \<Longrightarrow> S \<sqsubseteq> S' \<Longrightarrow> S(x := a) \<sqsubseteq> S'(x := a')" |
|
394 |
by(simp add: le_fun_def) |
|
395 |
||
396 |
lemma mono_step': "S1 \<sqsubseteq> S2 \<Longrightarrow> C1 \<sqsubseteq> C2 \<Longrightarrow> step' S1 C1 \<sqsubseteq> step' S2 C2" |
|
397 |
apply(induction C1 C2 arbitrary: S1 S2 rule: le_acom.induct) |
|
398 |
apply (auto simp: Let_def mono_update mono_aval' mono_post le_join_disj |
|
399 |
split: option.split) |
|
400 |
done |
|
401 |
||
402 |
end |
|
403 |
||
404 |
text{* Problem: not executable because of the comparison of abstract states, |
|
405 |
i.e. functions, in the post-fixedpoint computation. *} |
|
406 |
||
407 |
end |