author | nipkow |
Thu, 11 Apr 2013 15:10:22 +0200 | |
changeset 51694 | 6ae88642962f |
parent 51628 | 0a6d576da295 |
child 51710 | f692657e0f71 |
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 |
||
51625 | 9 |
text{* The basic type classes @{class order}, @{class semilattice_sup} and @{class top} are |
10 |
defined in @{theory Main}, more precisely in theories @{theory Orderings} and @{theory Lattices}. |
|
11 |
If you view this theory with jedit, just click on the names to get there. *} |
|
47613 | 12 |
|
51389 | 13 |
class semilattice = semilattice_sup + top |
47613 | 14 |
|
15 |
||
51389 | 16 |
instance "fun" :: (type, semilattice) semilattice .. |
47613 | 17 |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
18 |
instantiation option :: (order)order |
47613 | 19 |
begin |
20 |
||
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
21 |
fun less_eq_option where |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
22 |
"Some x \<le> Some y = (x \<le> y)" | |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
23 |
"None \<le> y = True" | |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
24 |
"Some _ \<le> None = False" |
47613 | 25 |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
26 |
definition less_option where "x < (y::'a option) = (x \<le> y \<and> \<not> y \<le> x)" |
47613 | 27 |
|
51628 | 28 |
lemma le_None[simp]: "(x \<le> None) = (x = None)" |
47613 | 29 |
by (cases x) simp_all |
30 |
||
51628 | 31 |
lemma Some_le[simp]: "(Some x \<le> u) = (\<exists>y. u = Some y \<and> x \<le> y)" |
47613 | 32 |
by (cases u) auto |
33 |
||
34 |
instance proof |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
35 |
case goal1 show ?case by(rule less_option_def) |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
36 |
next |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
37 |
case goal2 show ?case by(cases x, simp_all) |
47613 | 38 |
next |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
39 |
case goal3 thus ?case by(cases z, simp, cases y, simp, cases x, auto) |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
40 |
next |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
41 |
case goal4 thus ?case by(cases y, simp, cases x, auto) |
47613 | 42 |
qed |
43 |
||
44 |
end |
|
45 |
||
51389 | 46 |
instantiation option :: (sup)sup |
47613 | 47 |
begin |
48 |
||
51389 | 49 |
fun sup_option where |
47613 | 50 |
"Some x \<squnion> Some y = Some(x \<squnion> y)" | |
51 |
"None \<squnion> y = y" | |
|
52 |
"x \<squnion> None = x" |
|
53 |
||
51389 | 54 |
lemma sup_None2[simp]: "x \<squnion> None = x" |
47613 | 55 |
by (cases x) simp_all |
56 |
||
57 |
instance .. |
|
58 |
||
59 |
end |
|
60 |
||
49396 | 61 |
instantiation option :: (semilattice)semilattice |
47613 | 62 |
begin |
63 |
||
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
64 |
definition top_option where "\<top> = Some \<top>" |
47613 | 65 |
|
66 |
instance proof |
|
51389 | 67 |
case goal4 show ?case by(cases a, simp_all add: top_option_def) |
47613 | 68 |
next |
51389 | 69 |
case goal1 thus ?case by(cases x, simp, cases y, simp_all) |
47613 | 70 |
next |
51389 | 71 |
case goal2 thus ?case by(cases y, simp, cases x, simp_all) |
47613 | 72 |
next |
51389 | 73 |
case goal3 thus ?case by(cases z, simp, cases y, simp, cases x, simp_all) |
47613 | 74 |
qed |
75 |
||
76 |
end |
|
77 |
||
51390 | 78 |
lemma [simp]: "(Some x < Some y) = (x < y)" |
79 |
by(auto simp: less_le) |
|
80 |
||
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
81 |
instantiation option :: (order)bot |
47613 | 82 |
begin |
83 |
||
49396 | 84 |
definition bot_option :: "'a option" where |
47613 | 85 |
"\<bottom> = None" |
86 |
||
87 |
instance |
|
88 |
proof |
|
49396 | 89 |
case goal1 thus ?case by(auto simp: bot_option_def) |
47613 | 90 |
qed |
91 |
||
92 |
end |
|
93 |
||
94 |
||
95 |
definition bot :: "com \<Rightarrow> 'a option acom" where |
|
96 |
"bot c = anno None c" |
|
97 |
||
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
98 |
lemma bot_least: "strip C = c \<Longrightarrow> bot c \<le> C" |
47613 | 99 |
by(induct C arbitrary: c)(auto simp: bot_def) |
100 |
||
101 |
lemma strip_bot[simp]: "strip(bot c) = c" |
|
102 |
by(simp add: bot_def) |
|
103 |
||
104 |
||
105 |
subsubsection "Post-fixed point iteration" |
|
106 |
||
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
107 |
definition pfp :: "(('a::order) \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a option" where |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
108 |
"pfp f = while_option (\<lambda>x. \<not> f x \<le> x) f" |
47613 | 109 |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
110 |
lemma pfp_pfp: assumes "pfp f x0 = Some x" shows "f x \<le> x" |
47613 | 111 |
using while_option_stop[OF assms[simplified pfp_def]] by simp |
112 |
||
49464 | 113 |
lemma while_least: |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
114 |
fixes q :: "'a::order" |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
115 |
assumes "\<forall>x\<in>L.\<forall>y\<in>L. x \<le> y \<longrightarrow> f x \<le> f y" and "\<forall>x. x \<in> L \<longrightarrow> f x \<in> L" |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
116 |
and "\<forall>x \<in> L. b \<le> x" and "b \<in> L" and "f q \<le> q" and "q \<in> L" |
49464 | 117 |
and "while_option P f b = Some p" |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
118 |
shows "p \<le> q" |
49464 | 119 |
using while_option_rule[OF _ assms(7)[unfolded pfp_def], |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
120 |
where P = "%x. x \<in> L \<and> x \<le> q"] |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
121 |
by (metis assms(1-6) order_trans) |
47613 | 122 |
|
49464 | 123 |
lemma pfp_inv: |
124 |
"pfp f x = Some y \<Longrightarrow> (\<And>x. P x \<Longrightarrow> P(f x)) \<Longrightarrow> P x \<Longrightarrow> P y" |
|
125 |
unfolding pfp_def by (metis (lifting) while_option_rule) |
|
47613 | 126 |
|
127 |
lemma strip_pfp: |
|
128 |
assumes "\<And>x. g(f x) = g x" and "pfp f x0 = Some x" shows "g x = g x0" |
|
49464 | 129 |
using pfp_inv[OF assms(2), where P = "%x. g x = g x0"] assms(1) by simp |
47613 | 130 |
|
131 |
||
132 |
subsection "Abstract Interpretation" |
|
133 |
||
134 |
definition \<gamma>_fun :: "('a \<Rightarrow> 'b set) \<Rightarrow> ('c \<Rightarrow> 'a) \<Rightarrow> ('c \<Rightarrow> 'b)set" where |
|
135 |
"\<gamma>_fun \<gamma> F = {f. \<forall>x. f x \<in> \<gamma>(F x)}" |
|
136 |
||
137 |
fun \<gamma>_option :: "('a \<Rightarrow> 'b set) \<Rightarrow> 'a option \<Rightarrow> 'b set" where |
|
138 |
"\<gamma>_option \<gamma> None = {}" | |
|
139 |
"\<gamma>_option \<gamma> (Some a) = \<gamma> a" |
|
140 |
||
141 |
text{* The interface for abstract values: *} |
|
142 |
||
143 |
locale Val_abs = |
|
49396 | 144 |
fixes \<gamma> :: "'av::semilattice \<Rightarrow> val set" |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
145 |
assumes mono_gamma: "a \<le> b \<Longrightarrow> \<gamma> a \<le> \<gamma> b" |
47613 | 146 |
and gamma_Top[simp]: "\<gamma> \<top> = UNIV" |
147 |
fixes num' :: "val \<Rightarrow> 'av" |
|
148 |
and plus' :: "'av \<Rightarrow> 'av \<Rightarrow> 'av" |
|
51036 | 149 |
assumes gamma_num': "i \<in> \<gamma>(num' i)" |
150 |
and gamma_plus': "i1 \<in> \<gamma> a1 \<Longrightarrow> i2 \<in> \<gamma> a2 \<Longrightarrow> i1+i2 \<in> \<gamma>(plus' a1 a2)" |
|
47613 | 151 |
|
152 |
type_synonym 'av st = "(vname \<Rightarrow> 'av)" |
|
153 |
||
49396 | 154 |
locale Abs_Int_Fun = Val_abs \<gamma> for \<gamma> :: "'av::semilattice \<Rightarrow> val set" |
47613 | 155 |
begin |
156 |
||
157 |
fun aval' :: "aexp \<Rightarrow> 'av st \<Rightarrow> 'av" where |
|
50896 | 158 |
"aval' (N i) S = num' i" | |
47613 | 159 |
"aval' (V x) S = S x" | |
160 |
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)" |
|
161 |
||
51694 | 162 |
definition "fa x e S = (case S of None \<Rightarrow> None | Some S \<Rightarrow> Some(S(x := aval' e S)))" |
163 |
||
164 |
definition "step' = Step fa (\<lambda>b S. S)" |
|
51389 | 165 |
|
47613 | 166 |
definition AI :: "com \<Rightarrow> 'av st option acom option" where |
49464 | 167 |
"AI c = pfp (step' \<top>) (bot c)" |
47613 | 168 |
|
169 |
||
49497 | 170 |
abbreviation \<gamma>\<^isub>s :: "'av st \<Rightarrow> state set" |
171 |
where "\<gamma>\<^isub>s == \<gamma>_fun \<gamma>" |
|
47613 | 172 |
|
173 |
abbreviation \<gamma>\<^isub>o :: "'av st option \<Rightarrow> state set" |
|
49497 | 174 |
where "\<gamma>\<^isub>o == \<gamma>_option \<gamma>\<^isub>s" |
47613 | 175 |
|
176 |
abbreviation \<gamma>\<^isub>c :: "'av st option acom \<Rightarrow> state set acom" |
|
177 |
where "\<gamma>\<^isub>c == map_acom \<gamma>\<^isub>o" |
|
178 |
||
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
179 |
lemma gamma_s_Top[simp]: "\<gamma>\<^isub>s \<top> = UNIV" |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
180 |
by(simp add: top_fun_def \<gamma>_fun_def) |
47613 | 181 |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
182 |
lemma gamma_o_Top[simp]: "\<gamma>\<^isub>o \<top> = UNIV" |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
183 |
by (simp add: top_option_def) |
47613 | 184 |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
185 |
lemma mono_gamma_s: "f1 \<le> f2 \<Longrightarrow> \<gamma>\<^isub>s f1 \<subseteq> \<gamma>\<^isub>s f2" |
47613 | 186 |
by(auto simp: le_fun_def \<gamma>_fun_def dest: mono_gamma) |
187 |
||
188 |
lemma mono_gamma_o: |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
189 |
"S1 \<le> S2 \<Longrightarrow> \<gamma>\<^isub>o S1 \<subseteq> \<gamma>\<^isub>o S2" |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
190 |
by(induction S1 S2 rule: less_eq_option.induct)(simp_all add: mono_gamma_s) |
47613 | 191 |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
192 |
lemma mono_gamma_c: "C1 \<le> C2 \<Longrightarrow> \<gamma>\<^isub>c C1 \<le> \<gamma>\<^isub>c C2" |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
193 |
by (induction C1 C2 rule: less_eq_acom.induct) (simp_all add:mono_gamma_o) |
47613 | 194 |
|
195 |
text{* Soundness: *} |
|
196 |
||
49497 | 197 |
lemma aval'_sound: "s : \<gamma>\<^isub>s S \<Longrightarrow> aval a s : \<gamma>(aval' a S)" |
47613 | 198 |
by (induct a) (auto simp: gamma_num' gamma_plus' \<gamma>_fun_def) |
199 |
||
51390 | 200 |
lemma in_gamma_update: "\<lbrakk> s : \<gamma>\<^isub>s S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>s(S(x := a))" |
47613 | 201 |
by(simp add: \<gamma>_fun_def) |
202 |
||
51390 | 203 |
lemma gamma_Step_subcomm: |
204 |
assumes "!!x e S. f1 x e (\<gamma>\<^isub>o S) \<subseteq> \<gamma>\<^isub>o (f2 x e S)" "!!b S. g1 b (\<gamma>\<^isub>o S) \<subseteq> \<gamma>\<^isub>o (g2 b S)" |
|
205 |
shows "Step f1 g1 (\<gamma>\<^isub>o S) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (Step f2 g2 S C)" |
|
206 |
proof(induction C arbitrary: S) |
|
207 |
qed (auto simp: mono_gamma_o assms) |
|
208 |
||
50986 | 209 |
lemma step_step': "step (\<gamma>\<^isub>o S) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (step' S C)" |
51390 | 210 |
unfolding step_def step'_def |
51694 | 211 |
by(rule gamma_Step_subcomm) |
212 |
(auto simp: aval'_sound in_gamma_update fa_def split: option.splits) |
|
47613 | 213 |
|
214 |
lemma AI_sound: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c C" |
|
215 |
proof(simp add: CS_def AI_def) |
|
49464 | 216 |
assume 1: "pfp (step' \<top>) (bot c) = Some C" |
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
217 |
have pfp': "step' \<top> C \<le> C" by(rule pfp_pfp[OF 1]) |
50986 | 218 |
have 2: "step (\<gamma>\<^isub>o \<top>) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c C" --"transfer the pfp'" |
219 |
proof(rule order_trans) |
|
220 |
show "step (\<gamma>\<^isub>o \<top>) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (step' \<top> C)" by(rule step_step') |
|
221 |
show "... \<le> \<gamma>\<^isub>c C" by (metis mono_gamma_c[OF pfp']) |
|
47613 | 222 |
qed |
51390 | 223 |
have 3: "strip (\<gamma>\<^isub>c C) = c" by(simp add: strip_pfp[OF _ 1] step'_def) |
50986 | 224 |
have "lfp c (step (\<gamma>\<^isub>o \<top>)) \<le> \<gamma>\<^isub>c C" |
225 |
by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^isub>o \<top>)", OF 3 2]) |
|
226 |
thus "lfp c (step UNIV) \<le> \<gamma>\<^isub>c C" by simp |
|
47613 | 227 |
qed |
228 |
||
229 |
end |
|
230 |
||
231 |
||
232 |
subsubsection "Monotonicity" |
|
233 |
||
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
234 |
lemma mono_post: "C1 \<le> C2 \<Longrightarrow> post C1 \<le> post C2" |
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
235 |
by(induction C1 C2 rule: less_eq_acom.induct) (auto) |
47613 | 236 |
|
237 |
locale Abs_Int_Fun_mono = Abs_Int_Fun + |
|
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
238 |
assumes mono_plus': "a1 \<le> b1 \<Longrightarrow> a2 \<le> b2 \<Longrightarrow> plus' a1 a2 \<le> plus' b1 b2" |
47613 | 239 |
begin |
240 |
||
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
241 |
lemma mono_aval': "S \<le> S' \<Longrightarrow> aval' e S \<le> aval' e S'" |
47613 | 242 |
by(induction e)(auto simp: le_fun_def mono_plus') |
243 |
||
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
244 |
lemma mono_update: "a \<le> a' \<Longrightarrow> S \<le> S' \<Longrightarrow> S(x := a) \<le> S'(x := a')" |
47613 | 245 |
by(simp add: le_fun_def) |
246 |
||
51359
00b45c7e831f
major redesign: order instead of preorder, new definition of intervals as quotients
nipkow
parents:
51036
diff
changeset
|
247 |
lemma mono_step': "S1 \<le> S2 \<Longrightarrow> C1 \<le> C2 \<Longrightarrow> step' S1 C1 \<le> step' S2 C2" |
51390 | 248 |
unfolding step'_def |
51694 | 249 |
by(rule mono2_Step) |
250 |
(auto simp: mono_update mono_aval' fa_def split: option.split) |
|
47613 | 251 |
|
252 |
end |
|
253 |
||
254 |
text{* Problem: not executable because of the comparison of abstract states, |
|
255 |
i.e. functions, in the post-fixedpoint computation. *} |
|
256 |
||
257 |
end |