author | bulwahn |
Thu, 12 Nov 2009 20:38:57 +0100 | |
changeset 33649 | 854173fcd21c |
parent 30235 | 58d147683393 |
permissions | -rw-r--r-- |
12516 | 1 |
(* Title: HOL/MicroJava/BV/Opt.thy |
10496 | 2 |
ID: $Id$ |
3 |
Author: Tobias Nipkow |
|
4 |
Copyright 2000 TUM |
|
5 |
||
6 |
More about options |
|
7 |
*) |
|
8 |
||
12911 | 9 |
header {* \isaheader{More about Options} *} |
10496 | 10 |
|
27680 | 11 |
theory Opt |
12 |
imports Err |
|
13 |
begin |
|
10496 | 14 |
|
15 |
constdefs |
|
13006 | 16 |
le :: "'a ord \<Rightarrow> 'a option ord" |
17 |
"le r o1 o2 == case o2 of None \<Rightarrow> o1=None | |
|
18 |
Some y \<Rightarrow> (case o1 of None \<Rightarrow> True |
|
19 |
| Some x \<Rightarrow> x <=_r y)" |
|
10496 | 20 |
|
13006 | 21 |
opt :: "'a set \<Rightarrow> 'a option set" |
10496 | 22 |
"opt A == insert None {x . ? y:A. x = Some y}" |
23 |
||
13006 | 24 |
sup :: "'a ebinop \<Rightarrow> 'a option ebinop" |
10496 | 25 |
"sup f o1 o2 == |
13006 | 26 |
case o1 of None \<Rightarrow> OK o2 | Some x \<Rightarrow> (case o2 of None \<Rightarrow> OK o1 |
27 |
| Some y \<Rightarrow> (case f x y of Err \<Rightarrow> Err | OK z \<Rightarrow> OK (Some z)))" |
|
10496 | 28 |
|
13006 | 29 |
esl :: "'a esl \<Rightarrow> 'a option esl" |
10496 | 30 |
"esl == %(A,r,f). (opt A, le r, sup f)" |
31 |
||
32 |
lemma unfold_le_opt: |
|
33 |
"o1 <=_(le r) o2 = |
|
13006 | 34 |
(case o2 of None \<Rightarrow> o1=None | |
35 |
Some y \<Rightarrow> (case o1 of None \<Rightarrow> True | Some x \<Rightarrow> x <=_r y))" |
|
10496 | 36 |
apply (unfold lesub_def le_def) |
37 |
apply (rule refl) |
|
38 |
done |
|
39 |
||
40 |
lemma le_opt_refl: |
|
13006 | 41 |
"order r \<Longrightarrow> o1 <=_(le r) o1" |
10496 | 42 |
by (simp add: unfold_le_opt split: option.split) |
43 |
||
44 |
lemma le_opt_trans [rule_format]: |
|
13006 | 45 |
"order r \<Longrightarrow> |
46 |
o1 <=_(le r) o2 \<longrightarrow> o2 <=_(le r) o3 \<longrightarrow> o1 <=_(le r) o3" |
|
10496 | 47 |
apply (simp add: unfold_le_opt split: option.split) |
48 |
apply (blast intro: order_trans) |
|
49 |
done |
|
50 |
||
51 |
lemma le_opt_antisym [rule_format]: |
|
13006 | 52 |
"order r \<Longrightarrow> o1 <=_(le r) o2 \<longrightarrow> o2 <=_(le r) o1 \<longrightarrow> o1=o2" |
10496 | 53 |
apply (simp add: unfold_le_opt split: option.split) |
54 |
apply (blast intro: order_antisym) |
|
55 |
done |
|
56 |
||
57 |
lemma order_le_opt [intro!,simp]: |
|
13006 | 58 |
"order r \<Longrightarrow> order(le r)" |
22068 | 59 |
apply (subst Semilat.order_def) |
10496 | 60 |
apply (blast intro: le_opt_refl le_opt_trans le_opt_antisym) |
61 |
done |
|
62 |
||
63 |
lemma None_bot [iff]: |
|
64 |
"None <=_(le r) ox" |
|
65 |
apply (unfold lesub_def le_def) |
|
66 |
apply (simp split: option.split) |
|
67 |
done |
|
68 |
||
69 |
lemma Some_le [iff]: |
|
70 |
"(Some x <=_(le r) ox) = (? y. ox = Some y & x <=_r y)" |
|
71 |
apply (unfold lesub_def le_def) |
|
72 |
apply (simp split: option.split) |
|
73 |
done |
|
74 |
||
75 |
lemma le_None [iff]: |
|
76 |
"(ox <=_(le r) None) = (ox = None)"; |
|
77 |
apply (unfold lesub_def le_def) |
|
78 |
apply (simp split: option.split) |
|
79 |
done |
|
80 |
||
81 |
||
82 |
lemma OK_None_bot [iff]: |
|
83 |
"OK None <=_(Err.le (le r)) x" |
|
84 |
by (simp add: lesub_def Err.le_def le_def split: option.split err.split) |
|
85 |
||
86 |
lemma sup_None1 [iff]: |
|
87 |
"x +_(sup f) None = OK x" |
|
88 |
by (simp add: plussub_def sup_def split: option.split) |
|
89 |
||
90 |
lemma sup_None2 [iff]: |
|
91 |
"None +_(sup f) x = OK x" |
|
92 |
by (simp add: plussub_def sup_def split: option.split) |
|
93 |
||
94 |
||
95 |
lemma None_in_opt [iff]: |
|
96 |
"None : opt A" |
|
97 |
by (simp add: opt_def) |
|
98 |
||
99 |
lemma Some_in_opt [iff]: |
|
100 |
"(Some x : opt A) = (x:A)" |
|
101 |
apply (unfold opt_def) |
|
102 |
apply auto |
|
103 |
done |
|
104 |
||
105 |
||
13062 | 106 |
lemma semilat_opt [intro, simp]: |
13006 | 107 |
"\<And>L. err_semilat L \<Longrightarrow> err_semilat (Opt.esl L)" |
10496 | 108 |
proof (unfold Opt.esl_def Err.sl_def, simp add: split_tupled_all) |
109 |
||
110 |
fix A r f |
|
111 |
assume s: "semilat (err A, Err.le r, lift2 f)" |
|
112 |
||
113 |
let ?A0 = "err A" |
|
114 |
let ?r0 = "Err.le r" |
|
115 |
let ?f0 = "lift2 f" |
|
116 |
||
117 |
from s |
|
118 |
obtain |
|
119 |
ord: "order ?r0" and |
|
120 |
clo: "closed ?A0 ?f0" and |
|
121 |
ub1: "\<forall>x\<in>?A0. \<forall>y\<in>?A0. x <=_?r0 x +_?f0 y" and |
|
122 |
ub2: "\<forall>x\<in>?A0. \<forall>y\<in>?A0. y <=_?r0 x +_?f0 y" and |
|
123 |
lub: "\<forall>x\<in>?A0. \<forall>y\<in>?A0. \<forall>z\<in>?A0. x <=_?r0 z \<and> y <=_?r0 z \<longrightarrow> x +_?f0 y <=_?r0 z" |
|
124 |
by (unfold semilat_def) simp |
|
125 |
||
126 |
let ?A = "err (opt A)" |
|
127 |
let ?r = "Err.le (Opt.le r)" |
|
128 |
let ?f = "lift2 (Opt.sup f)" |
|
129 |
||
130 |
from ord |
|
131 |
have "order ?r" |
|
132 |
by simp |
|
133 |
||
134 |
moreover |
|
135 |
||
136 |
have "closed ?A ?f" |
|
137 |
proof (unfold closed_def, intro strip) |
|
138 |
fix x y |
|
139 |
assume x: "x : ?A" |
|
140 |
assume y: "y : ?A" |
|
141 |
||
11085 | 142 |
{ fix a b |
10496 | 143 |
assume ab: "x = OK a" "y = OK b" |
144 |
||
145 |
with x |
|
13006 | 146 |
have a: "\<And>c. a = Some c \<Longrightarrow> c : A" |
10496 | 147 |
by (clarsimp simp add: opt_def) |
148 |
||
149 |
from ab y |
|
13006 | 150 |
have b: "\<And>d. b = Some d \<Longrightarrow> d : A" |
10496 | 151 |
by (clarsimp simp add: opt_def) |
152 |
||
153 |
{ fix c d assume "a = Some c" "b = Some d" |
|
154 |
with ab x y |
|
155 |
have "c:A & d:A" |
|
156 |
by (simp add: err_def opt_def Bex_def) |
|
157 |
with clo |
|
158 |
have "f c d : err A" |
|
159 |
by (simp add: closed_def plussub_def err_def lift2_def) |
|
160 |
moreover |
|
161 |
fix z assume "f c d = OK z" |
|
162 |
ultimately |
|
163 |
have "z : A" by simp |
|
164 |
} note f_closed = this |
|
165 |
||
166 |
have "sup f a b : ?A" |
|
167 |
proof (cases a) |
|
168 |
case None |
|
169 |
thus ?thesis |
|
170 |
by (simp add: sup_def opt_def) (cases b, simp, simp add: b Bex_def) |
|
171 |
next |
|
172 |
case Some |
|
173 |
thus ?thesis |
|
174 |
by (auto simp add: sup_def opt_def Bex_def a b f_closed split: err.split option.split) |
|
175 |
qed |
|
176 |
} |
|
177 |
||
178 |
thus "x +_?f y : ?A" |
|
179 |
by (simp add: plussub_def lift2_def split: err.split) |
|
180 |
qed |
|
181 |
||
182 |
moreover |
|
183 |
||
11085 | 184 |
{ fix a b c |
185 |
assume "a \<in> opt A" "b \<in> opt A" "a +_(sup f) b = OK c" |
|
186 |
moreover |
|
187 |
from ord have "order r" by simp |
|
188 |
moreover |
|
189 |
{ fix x y z |
|
190 |
assume "x \<in> A" "y \<in> A" |
|
191 |
hence "OK x \<in> err A \<and> OK y \<in> err A" by simp |
|
192 |
with ub1 ub2 |
|
193 |
have "(OK x) <=_(Err.le r) (OK x) +_(lift2 f) (OK y) \<and> |
|
194 |
(OK y) <=_(Err.le r) (OK x) +_(lift2 f) (OK y)" |
|
195 |
by blast |
|
196 |
moreover |
|
197 |
assume "x +_f y = OK z" |
|
198 |
ultimately |
|
199 |
have "x <=_r z \<and> y <=_r z" |
|
200 |
by (auto simp add: plussub_def lift2_def Err.le_def lesub_def) |
|
201 |
} |
|
202 |
ultimately |
|
203 |
have "a <=_(le r) c \<and> b <=_(le r) c" |
|
204 |
by (auto simp add: sup_def le_def lesub_def plussub_def |
|
205 |
dest: order_refl split: option.splits err.splits) |
|
206 |
} |
|
207 |
||
208 |
hence "(\<forall>x\<in>?A. \<forall>y\<in>?A. x <=_?r x +_?f y) \<and> (\<forall>x\<in>?A. \<forall>y\<in>?A. y <=_?r x +_?f y)" |
|
209 |
by (auto simp add: lesub_def plussub_def Err.le_def lift2_def split: err.split) |
|
10496 | 210 |
|
211 |
moreover |
|
212 |
||
213 |
have "\<forall>x\<in>?A. \<forall>y\<in>?A. \<forall>z\<in>?A. x <=_?r z \<and> y <=_?r z \<longrightarrow> x +_?f y <=_?r z" |
|
214 |
proof (intro strip, elim conjE) |
|
215 |
fix x y z |
|
216 |
assume xyz: "x : ?A" "y : ?A" "z : ?A" |
|
217 |
assume xz: "x <=_?r z" |
|
218 |
assume yz: "y <=_?r z" |
|
219 |
||
220 |
{ fix a b c |
|
221 |
assume ok: "x = OK a" "y = OK b" "z = OK c" |
|
222 |
||
223 |
{ fix d e g |
|
224 |
assume some: "a = Some d" "b = Some e" "c = Some g" |
|
225 |
||
226 |
with ok xyz |
|
227 |
obtain "OK d:err A" "OK e:err A" "OK g:err A" |
|
228 |
by simp |
|
229 |
with lub |
|
13006 | 230 |
have "\<lbrakk> (OK d) <=_(Err.le r) (OK g); (OK e) <=_(Err.le r) (OK g) \<rbrakk> |
231 |
\<Longrightarrow> (OK d) +_(lift2 f) (OK e) <=_(Err.le r) (OK g)" |
|
10496 | 232 |
by blast |
13006 | 233 |
hence "\<lbrakk> d <=_r g; e <=_r g \<rbrakk> \<Longrightarrow> \<exists>y. d +_f e = OK y \<and> y <=_r g" |
10496 | 234 |
by simp |
235 |
||
236 |
with ok some xyz xz yz |
|
237 |
have "x +_?f y <=_?r z" |
|
238 |
by (auto simp add: sup_def le_def lesub_def lift2_def plussub_def Err.le_def) |
|
239 |
} note this [intro!] |
|
240 |
||
241 |
from ok xyz xz yz |
|
242 |
have "x +_?f y <=_?r z" |
|
243 |
by - (cases a, simp, cases b, simp, cases c, simp, blast) |
|
244 |
} |
|
245 |
||
246 |
with xyz xz yz |
|
247 |
show "x +_?f y <=_?r z" |
|
248 |
by - (cases x, simp, cases y, simp, cases z, simp+) |
|
249 |
qed |
|
250 |
||
251 |
ultimately |
|
252 |
||
253 |
show "semilat (?A,?r,?f)" |
|
254 |
by (unfold semilat_def) simp |
|
255 |
qed |
|
256 |
||
257 |
lemma top_le_opt_Some [iff]: |
|
258 |
"top (le r) (Some T) = top r T" |
|
259 |
apply (unfold top_def) |
|
260 |
apply (rule iffI) |
|
261 |
apply blast |
|
262 |
apply (rule allI) |
|
263 |
apply (case_tac "x") |
|
264 |
apply simp+ |
|
265 |
done |
|
266 |
||
267 |
lemma Top_le_conv: |
|
13006 | 268 |
"\<lbrakk> order r; top r T \<rbrakk> \<Longrightarrow> (T <=_r x) = (x = T)" |
10496 | 269 |
apply (unfold top_def) |
270 |
apply (blast intro: order_antisym) |
|
271 |
done |
|
272 |
||
273 |
||
274 |
lemma acc_le_optI [intro!]: |
|
13006 | 275 |
"acc r \<Longrightarrow> acc(le r)" |
10496 | 276 |
apply (unfold acc_def lesub_def le_def lesssub_def) |
22271 | 277 |
apply (simp add: wfP_eq_minimal split: option.split) |
10496 | 278 |
apply clarify |
279 |
apply (case_tac "? a. Some a : Q") |
|
280 |
apply (erule_tac x = "{a . Some a : Q}" in allE) |
|
281 |
apply blast |
|
282 |
apply (case_tac "x") |
|
283 |
apply blast |
|
284 |
apply blast |
|
285 |
done |
|
286 |
||
287 |
lemma option_map_in_optionI: |
|
13006 | 288 |
"\<lbrakk> ox : opt S; !x:S. ox = Some x \<longrightarrow> f x : S \<rbrakk> |
30235
58d147683393
Made Option a separate theory and renamed option_map to Option.map
nipkow
parents:
27680
diff
changeset
|
289 |
\<Longrightarrow> Option.map f ox : opt S"; |
58d147683393
Made Option a separate theory and renamed option_map to Option.map
nipkow
parents:
27680
diff
changeset
|
290 |
apply (unfold Option.map_def) |
10496 | 291 |
apply (simp split: option.split) |
292 |
apply blast |
|
293 |
done |
|
294 |
||
295 |
end |