|
1 (* Title: HOL/HOLCF/Cpo.thy |
|
2 Author: Franz Regensburger |
|
3 Author: Tobias Nipkow |
|
4 Author: Brian Huffman |
|
5 |
|
6 Foundations of HOLCF: complete partial orders etc. |
|
7 *) |
|
8 |
|
9 theory Cpo |
|
10 imports Main |
|
11 begin |
|
12 |
|
13 section \<open>Partial orders\<close> |
|
14 |
|
15 declare [[typedef_overloaded]] |
|
16 |
|
17 |
|
18 subsection \<open>Type class for partial orders\<close> |
|
19 |
|
20 class below = |
|
21 fixes below :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
|
22 begin |
|
23 |
|
24 notation (ASCII) |
|
25 below (infix \<open><<\<close> 50) |
|
26 |
|
27 notation |
|
28 below (infix \<open>\<sqsubseteq>\<close> 50) |
|
29 |
|
30 abbreviation not_below :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix \<open>\<notsqsubseteq>\<close> 50) |
|
31 where "not_below x y \<equiv> \<not> below x y" |
|
32 |
|
33 notation (ASCII) |
|
34 not_below (infix \<open>~<<\<close> 50) |
|
35 |
|
36 lemma below_eq_trans: "a \<sqsubseteq> b \<Longrightarrow> b = c \<Longrightarrow> a \<sqsubseteq> c" |
|
37 by (rule subst) |
|
38 |
|
39 lemma eq_below_trans: "a = b \<Longrightarrow> b \<sqsubseteq> c \<Longrightarrow> a \<sqsubseteq> c" |
|
40 by (rule ssubst) |
|
41 |
|
42 end |
|
43 |
|
44 class po = below + |
|
45 assumes below_refl [iff]: "x \<sqsubseteq> x" |
|
46 assumes below_trans: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z" |
|
47 assumes below_antisym: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> x \<Longrightarrow> x = y" |
|
48 begin |
|
49 |
|
50 lemma eq_imp_below: "x = y \<Longrightarrow> x \<sqsubseteq> y" |
|
51 by simp |
|
52 |
|
53 lemma box_below: "a \<sqsubseteq> b \<Longrightarrow> c \<sqsubseteq> a \<Longrightarrow> b \<sqsubseteq> d \<Longrightarrow> c \<sqsubseteq> d" |
|
54 by (rule below_trans [OF below_trans]) |
|
55 |
|
56 lemma po_eq_conv: "x = y \<longleftrightarrow> x \<sqsubseteq> y \<and> y \<sqsubseteq> x" |
|
57 by (fast intro!: below_antisym) |
|
58 |
|
59 lemma rev_below_trans: "y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z" |
|
60 by (rule below_trans) |
|
61 |
|
62 lemma not_below2not_eq: "x \<notsqsubseteq> y \<Longrightarrow> x \<noteq> y" |
|
63 by auto |
|
64 |
|
65 end |
|
66 |
|
67 lemmas HOLCF_trans_rules [trans] = |
|
68 below_trans |
|
69 below_antisym |
|
70 below_eq_trans |
|
71 eq_below_trans |
|
72 |
|
73 context po |
|
74 begin |
|
75 |
|
76 subsection \<open>Upper bounds\<close> |
|
77 |
|
78 definition is_ub :: "'a set \<Rightarrow> 'a \<Rightarrow> bool" (infix \<open><|\<close> 55) |
|
79 where "S <| x \<longleftrightarrow> (\<forall>y\<in>S. y \<sqsubseteq> x)" |
|
80 |
|
81 lemma is_ubI: "(\<And>x. x \<in> S \<Longrightarrow> x \<sqsubseteq> u) \<Longrightarrow> S <| u" |
|
82 by (simp add: is_ub_def) |
|
83 |
|
84 lemma is_ubD: "\<lbrakk>S <| u; x \<in> S\<rbrakk> \<Longrightarrow> x \<sqsubseteq> u" |
|
85 by (simp add: is_ub_def) |
|
86 |
|
87 lemma ub_imageI: "(\<And>x. x \<in> S \<Longrightarrow> f x \<sqsubseteq> u) \<Longrightarrow> (\<lambda>x. f x) ` S <| u" |
|
88 unfolding is_ub_def by fast |
|
89 |
|
90 lemma ub_imageD: "\<lbrakk>f ` S <| u; x \<in> S\<rbrakk> \<Longrightarrow> f x \<sqsubseteq> u" |
|
91 unfolding is_ub_def by fast |
|
92 |
|
93 lemma ub_rangeI: "(\<And>i. S i \<sqsubseteq> x) \<Longrightarrow> range S <| x" |
|
94 unfolding is_ub_def by fast |
|
95 |
|
96 lemma ub_rangeD: "range S <| x \<Longrightarrow> S i \<sqsubseteq> x" |
|
97 unfolding is_ub_def by fast |
|
98 |
|
99 lemma is_ub_empty [simp]: "{} <| u" |
|
100 unfolding is_ub_def by fast |
|
101 |
|
102 lemma is_ub_insert [simp]: "(insert x A) <| y = (x \<sqsubseteq> y \<and> A <| y)" |
|
103 unfolding is_ub_def by fast |
|
104 |
|
105 lemma is_ub_upward: "\<lbrakk>S <| x; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> S <| y" |
|
106 unfolding is_ub_def by (fast intro: below_trans) |
|
107 |
|
108 |
|
109 subsection \<open>Least upper bounds\<close> |
|
110 |
|
111 definition is_lub :: "'a set \<Rightarrow> 'a \<Rightarrow> bool" (infix \<open><<|\<close> 55) |
|
112 where "S <<| x \<longleftrightarrow> S <| x \<and> (\<forall>u. S <| u \<longrightarrow> x \<sqsubseteq> u)" |
|
113 |
|
114 definition lub :: "'a set \<Rightarrow> 'a" |
|
115 where "lub S = (THE x. S <<| x)" |
|
116 |
|
117 end |
|
118 |
|
119 syntax (ASCII) |
|
120 "_BLub" :: "[pttrn, 'a set, 'b] \<Rightarrow> 'b" (\<open>(\<open>indent=3 notation=\<open>binder LUB\<close>\<close>LUB _:_./ _)\<close> [0,0, 10] 10) |
|
121 |
|
122 syntax |
|
123 "_BLub" :: "[pttrn, 'a set, 'b] \<Rightarrow> 'b" (\<open>(\<open>indent=3 notation=\<open>binder \<Squnion>\<close>\<close>\<Squnion>_\<in>_./ _)\<close> [0,0, 10] 10) |
|
124 |
|
125 syntax_consts |
|
126 "_BLub" \<rightleftharpoons> lub |
|
127 |
|
128 translations |
|
129 "LUB x:A. t" \<rightleftharpoons> "CONST lub ((\<lambda>x. t) ` A)" |
|
130 |
|
131 context po |
|
132 begin |
|
133 |
|
134 abbreviation Lub (binder \<open>\<Squnion>\<close> 10) |
|
135 where "\<Squnion>n. t n \<equiv> lub (range t)" |
|
136 |
|
137 notation (ASCII) |
|
138 Lub (binder \<open>LUB \<close> 10) |
|
139 |
|
140 text \<open>access to some definition as inference rule\<close> |
|
141 |
|
142 lemma is_lubD1: "S <<| x \<Longrightarrow> S <| x" |
|
143 unfolding is_lub_def by fast |
|
144 |
|
145 lemma is_lubD2: "\<lbrakk>S <<| x; S <| u\<rbrakk> \<Longrightarrow> x \<sqsubseteq> u" |
|
146 unfolding is_lub_def by fast |
|
147 |
|
148 lemma is_lubI: "\<lbrakk>S <| x; \<And>u. S <| u \<Longrightarrow> x \<sqsubseteq> u\<rbrakk> \<Longrightarrow> S <<| x" |
|
149 unfolding is_lub_def by fast |
|
150 |
|
151 lemma is_lub_below_iff: "S <<| x \<Longrightarrow> x \<sqsubseteq> u \<longleftrightarrow> S <| u" |
|
152 unfolding is_lub_def is_ub_def by (metis below_trans) |
|
153 |
|
154 text \<open>lubs are unique\<close> |
|
155 |
|
156 lemma is_lub_unique: "S <<| x \<Longrightarrow> S <<| y \<Longrightarrow> x = y" |
|
157 unfolding is_lub_def is_ub_def by (blast intro: below_antisym) |
|
158 |
|
159 text \<open>technical lemmas about \<^term>\<open>lub\<close> and \<^term>\<open>is_lub\<close>\<close> |
|
160 |
|
161 lemma is_lub_lub: "M <<| x \<Longrightarrow> M <<| lub M" |
|
162 unfolding lub_def by (rule theI [OF _ is_lub_unique]) |
|
163 |
|
164 lemma lub_eqI: "M <<| l \<Longrightarrow> lub M = l" |
|
165 by (rule is_lub_unique [OF is_lub_lub]) |
|
166 |
|
167 lemma is_lub_singleton [simp]: "{x} <<| x" |
|
168 by (simp add: is_lub_def) |
|
169 |
|
170 lemma lub_singleton [simp]: "lub {x} = x" |
|
171 by (rule is_lub_singleton [THEN lub_eqI]) |
|
172 |
|
173 lemma is_lub_bin: "x \<sqsubseteq> y \<Longrightarrow> {x, y} <<| y" |
|
174 by (simp add: is_lub_def) |
|
175 |
|
176 lemma lub_bin: "x \<sqsubseteq> y \<Longrightarrow> lub {x, y} = y" |
|
177 by (rule is_lub_bin [THEN lub_eqI]) |
|
178 |
|
179 lemma is_lub_maximal: "S <| x \<Longrightarrow> x \<in> S \<Longrightarrow> S <<| x" |
|
180 by (erule is_lubI, erule (1) is_ubD) |
|
181 |
|
182 lemma lub_maximal: "S <| x \<Longrightarrow> x \<in> S \<Longrightarrow> lub S = x" |
|
183 by (rule is_lub_maximal [THEN lub_eqI]) |
|
184 |
|
185 |
|
186 subsection \<open>Countable chains\<close> |
|
187 |
|
188 definition chain :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool" |
|
189 where \<comment> \<open>Here we use countable chains and I prefer to code them as functions!\<close> |
|
190 "chain Y = (\<forall>i. Y i \<sqsubseteq> Y (Suc i))" |
|
191 |
|
192 lemma chainI: "(\<And>i. Y i \<sqsubseteq> Y (Suc i)) \<Longrightarrow> chain Y" |
|
193 unfolding chain_def by fast |
|
194 |
|
195 lemma chainE: "chain Y \<Longrightarrow> Y i \<sqsubseteq> Y (Suc i)" |
|
196 unfolding chain_def by fast |
|
197 |
|
198 text \<open>chains are monotone functions\<close> |
|
199 |
|
200 lemma chain_mono_less: "chain Y \<Longrightarrow> i < j \<Longrightarrow> Y i \<sqsubseteq> Y j" |
|
201 by (erule less_Suc_induct, erule chainE, erule below_trans) |
|
202 |
|
203 lemma chain_mono: "chain Y \<Longrightarrow> i \<le> j \<Longrightarrow> Y i \<sqsubseteq> Y j" |
|
204 by (cases "i = j") (simp_all add: chain_mono_less) |
|
205 |
|
206 lemma chain_shift: "chain Y \<Longrightarrow> chain (\<lambda>i. Y (i + j))" |
|
207 by (rule chainI, simp, erule chainE) |
|
208 |
|
209 text \<open>technical lemmas about (least) upper bounds of chains\<close> |
|
210 |
|
211 lemma is_lub_rangeD1: "range S <<| x \<Longrightarrow> S i \<sqsubseteq> x" |
|
212 by (rule is_lubD1 [THEN ub_rangeD]) |
|
213 |
|
214 lemma is_ub_range_shift: "chain S \<Longrightarrow> range (\<lambda>i. S (i + j)) <| x = range S <| x" |
|
215 apply (rule iffI) |
|
216 apply (rule ub_rangeI) |
|
217 apply (rule_tac y="S (i + j)" in below_trans) |
|
218 apply (erule chain_mono) |
|
219 apply (rule le_add1) |
|
220 apply (erule ub_rangeD) |
|
221 apply (rule ub_rangeI) |
|
222 apply (erule ub_rangeD) |
|
223 done |
|
224 |
|
225 lemma is_lub_range_shift: "chain S \<Longrightarrow> range (\<lambda>i. S (i + j)) <<| x = range S <<| x" |
|
226 by (simp add: is_lub_def is_ub_range_shift) |
|
227 |
|
228 text \<open>the lub of a constant chain is the constant\<close> |
|
229 |
|
230 lemma chain_const [simp]: "chain (\<lambda>i. c)" |
|
231 by (simp add: chainI) |
|
232 |
|
233 lemma is_lub_const: "range (\<lambda>x. c) <<| c" |
|
234 by (blast dest: ub_rangeD intro: is_lubI ub_rangeI) |
|
235 |
|
236 lemma lub_const [simp]: "(\<Squnion>i. c) = c" |
|
237 by (rule is_lub_const [THEN lub_eqI]) |
|
238 |
|
239 |
|
240 subsection \<open>Finite chains\<close> |
|
241 |
|
242 definition max_in_chain :: "nat \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> bool" |
|
243 where \<comment> \<open>finite chains, needed for monotony of continuous functions\<close> |
|
244 "max_in_chain i C \<longleftrightarrow> (\<forall>j. i \<le> j \<longrightarrow> C i = C j)" |
|
245 |
|
246 definition finite_chain :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool" |
|
247 where "finite_chain C = (chain C \<and> (\<exists>i. max_in_chain i C))" |
|
248 |
|
249 text \<open>results about finite chains\<close> |
|
250 |
|
251 lemma max_in_chainI: "(\<And>j. i \<le> j \<Longrightarrow> Y i = Y j) \<Longrightarrow> max_in_chain i Y" |
|
252 unfolding max_in_chain_def by fast |
|
253 |
|
254 lemma max_in_chainD: "max_in_chain i Y \<Longrightarrow> i \<le> j \<Longrightarrow> Y i = Y j" |
|
255 unfolding max_in_chain_def by fast |
|
256 |
|
257 lemma finite_chainI: "chain C \<Longrightarrow> max_in_chain i C \<Longrightarrow> finite_chain C" |
|
258 unfolding finite_chain_def by fast |
|
259 |
|
260 lemma finite_chainE: "\<lbrakk>finite_chain C; \<And>i. \<lbrakk>chain C; max_in_chain i C\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
|
261 unfolding finite_chain_def by fast |
|
262 |
|
263 lemma lub_finch1: "chain C \<Longrightarrow> max_in_chain i C \<Longrightarrow> range C <<| C i" |
|
264 apply (rule is_lubI) |
|
265 apply (rule ub_rangeI, rename_tac j) |
|
266 apply (rule_tac x=i and y=j in linorder_le_cases) |
|
267 apply (drule (1) max_in_chainD, simp) |
|
268 apply (erule (1) chain_mono) |
|
269 apply (erule ub_rangeD) |
|
270 done |
|
271 |
|
272 lemma lub_finch2: "finite_chain C \<Longrightarrow> range C <<| C (LEAST i. max_in_chain i C)" |
|
273 apply (erule finite_chainE) |
|
274 apply (erule LeastI2 [where Q="\<lambda>i. range C <<| C i"]) |
|
275 apply (erule (1) lub_finch1) |
|
276 done |
|
277 |
|
278 lemma finch_imp_finite_range: "finite_chain Y \<Longrightarrow> finite (range Y)" |
|
279 apply (erule finite_chainE) |
|
280 apply (rule_tac B="Y ` {..i}" in finite_subset) |
|
281 apply (rule subsetI) |
|
282 apply (erule rangeE, rename_tac j) |
|
283 apply (rule_tac x=i and y=j in linorder_le_cases) |
|
284 apply (subgoal_tac "Y j = Y i", simp) |
|
285 apply (simp add: max_in_chain_def) |
|
286 apply simp |
|
287 apply simp |
|
288 done |
|
289 |
|
290 lemma finite_range_has_max: |
|
291 fixes f :: "nat \<Rightarrow> 'a" |
|
292 and r :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
|
293 assumes mono: "\<And>i j. i \<le> j \<Longrightarrow> r (f i) (f j)" |
|
294 assumes finite_range: "finite (range f)" |
|
295 shows "\<exists>k. \<forall>i. r (f i) (f k)" |
|
296 proof (intro exI allI) |
|
297 fix i :: nat |
|
298 let ?j = "LEAST k. f k = f i" |
|
299 let ?k = "Max ((\<lambda>x. LEAST k. f k = x) ` range f)" |
|
300 have "?j \<le> ?k" |
|
301 proof (rule Max_ge) |
|
302 show "finite ((\<lambda>x. LEAST k. f k = x) ` range f)" |
|
303 using finite_range by (rule finite_imageI) |
|
304 show "?j \<in> (\<lambda>x. LEAST k. f k = x) ` range f" |
|
305 by (intro imageI rangeI) |
|
306 qed |
|
307 hence "r (f ?j) (f ?k)" |
|
308 by (rule mono) |
|
309 also have "f ?j = f i" |
|
310 by (rule LeastI, rule refl) |
|
311 finally show "r (f i) (f ?k)" . |
|
312 qed |
|
313 |
|
314 lemma finite_range_imp_finch: "chain Y \<Longrightarrow> finite (range Y) \<Longrightarrow> finite_chain Y" |
|
315 apply (subgoal_tac "\<exists>k. \<forall>i. Y i \<sqsubseteq> Y k") |
|
316 apply (erule exE) |
|
317 apply (rule finite_chainI, assumption) |
|
318 apply (rule max_in_chainI) |
|
319 apply (rule below_antisym) |
|
320 apply (erule (1) chain_mono) |
|
321 apply (erule spec) |
|
322 apply (rule finite_range_has_max) |
|
323 apply (erule (1) chain_mono) |
|
324 apply assumption |
|
325 done |
|
326 |
|
327 lemma bin_chain: "x \<sqsubseteq> y \<Longrightarrow> chain (\<lambda>i. if i=0 then x else y)" |
|
328 by (rule chainI) simp |
|
329 |
|
330 lemma bin_chainmax: "x \<sqsubseteq> y \<Longrightarrow> max_in_chain (Suc 0) (\<lambda>i. if i=0 then x else y)" |
|
331 by (simp add: max_in_chain_def) |
|
332 |
|
333 lemma is_lub_bin_chain: "x \<sqsubseteq> y \<Longrightarrow> range (\<lambda>i::nat. if i=0 then x else y) <<| y" |
|
334 apply (frule bin_chain) |
|
335 apply (drule bin_chainmax) |
|
336 apply (drule (1) lub_finch1) |
|
337 apply simp |
|
338 done |
|
339 |
|
340 text \<open>the maximal element in a chain is its lub\<close> |
|
341 |
|
342 lemma lub_chain_maxelem: "Y i = c \<Longrightarrow> \<forall>i. Y i \<sqsubseteq> c \<Longrightarrow> lub (range Y) = c" |
|
343 by (blast dest: ub_rangeD intro: lub_eqI is_lubI ub_rangeI) |
|
344 |
|
345 end |
|
346 |
|
347 |
|
348 section \<open>Classes cpo and pcpo\<close> |
|
349 |
|
350 subsection \<open>Complete partial orders\<close> |
|
351 |
|
352 text \<open>The class cpo of chain complete partial orders\<close> |
|
353 |
|
354 class cpo = po + |
|
355 assumes cpo: "chain S \<Longrightarrow> \<exists>x. range S <<| x" |
|
356 begin |
|
357 |
|
358 text \<open>in cpo's everthing equal to THE lub has lub properties for every chain\<close> |
|
359 |
|
360 lemma cpo_lubI: "chain S \<Longrightarrow> range S <<| (\<Squnion>i. S i)" |
|
361 by (fast dest: cpo elim: is_lub_lub) |
|
362 |
|
363 lemma thelubE: "\<lbrakk>chain S; (\<Squnion>i. S i) = l\<rbrakk> \<Longrightarrow> range S <<| l" |
|
364 by (blast dest: cpo intro: is_lub_lub) |
|
365 |
|
366 text \<open>Properties of the lub\<close> |
|
367 |
|
368 lemma is_ub_thelub: "chain S \<Longrightarrow> S x \<sqsubseteq> (\<Squnion>i. S i)" |
|
369 by (blast dest: cpo intro: is_lub_lub [THEN is_lub_rangeD1]) |
|
370 |
|
371 lemma is_lub_thelub: "\<lbrakk>chain S; range S <| x\<rbrakk> \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x" |
|
372 by (blast dest: cpo intro: is_lub_lub [THEN is_lubD2]) |
|
373 |
|
374 lemma lub_below_iff: "chain S \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x \<longleftrightarrow> (\<forall>i. S i \<sqsubseteq> x)" |
|
375 by (simp add: is_lub_below_iff [OF cpo_lubI] is_ub_def) |
|
376 |
|
377 lemma lub_below: "\<lbrakk>chain S; \<And>i. S i \<sqsubseteq> x\<rbrakk> \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x" |
|
378 by (simp add: lub_below_iff) |
|
379 |
|
380 lemma below_lub: "\<lbrakk>chain S; x \<sqsubseteq> S i\<rbrakk> \<Longrightarrow> x \<sqsubseteq> (\<Squnion>i. S i)" |
|
381 by (erule below_trans, erule is_ub_thelub) |
|
382 |
|
383 lemma lub_range_mono: "\<lbrakk>range X \<subseteq> range Y; chain Y; chain X\<rbrakk> \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" |
|
384 apply (erule lub_below) |
|
385 apply (subgoal_tac "\<exists>j. X i = Y j") |
|
386 apply clarsimp |
|
387 apply (erule is_ub_thelub) |
|
388 apply auto |
|
389 done |
|
390 |
|
391 lemma lub_range_shift: "chain Y \<Longrightarrow> (\<Squnion>i. Y (i + j)) = (\<Squnion>i. Y i)" |
|
392 apply (rule below_antisym) |
|
393 apply (rule lub_range_mono) |
|
394 apply fast |
|
395 apply assumption |
|
396 apply (erule chain_shift) |
|
397 apply (rule lub_below) |
|
398 apply assumption |
|
399 apply (rule_tac i="i" in below_lub) |
|
400 apply (erule chain_shift) |
|
401 apply (erule chain_mono) |
|
402 apply (rule le_add1) |
|
403 done |
|
404 |
|
405 lemma maxinch_is_thelub: "chain Y \<Longrightarrow> max_in_chain i Y = ((\<Squnion>i. Y i) = Y i)" |
|
406 apply (rule iffI) |
|
407 apply (fast intro!: lub_eqI lub_finch1) |
|
408 apply (unfold max_in_chain_def) |
|
409 apply (safe intro!: below_antisym) |
|
410 apply (fast elim!: chain_mono) |
|
411 apply (drule sym) |
|
412 apply (force elim!: is_ub_thelub) |
|
413 done |
|
414 |
|
415 text \<open>the \<open>\<sqsubseteq>\<close> relation between two chains is preserved by their lubs\<close> |
|
416 |
|
417 lemma lub_mono: "\<lbrakk>chain X; chain Y; \<And>i. X i \<sqsubseteq> Y i\<rbrakk> \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" |
|
418 by (fast elim: lub_below below_lub) |
|
419 |
|
420 text \<open>the = relation between two chains is preserved by their lubs\<close> |
|
421 |
|
422 lemma lub_eq: "(\<And>i. X i = Y i) \<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)" |
|
423 by simp |
|
424 |
|
425 lemma ch2ch_lub: |
|
426 assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" |
|
427 assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" |
|
428 shows "chain (\<lambda>i. \<Squnion>j. Y i j)" |
|
429 apply (rule chainI) |
|
430 apply (rule lub_mono [OF 2 2]) |
|
431 apply (rule chainE [OF 1]) |
|
432 done |
|
433 |
|
434 lemma diag_lub: |
|
435 assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" |
|
436 assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" |
|
437 shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>i. Y i i)" |
|
438 proof (rule below_antisym) |
|
439 have 3: "chain (\<lambda>i. Y i i)" |
|
440 apply (rule chainI) |
|
441 apply (rule below_trans) |
|
442 apply (rule chainE [OF 1]) |
|
443 apply (rule chainE [OF 2]) |
|
444 done |
|
445 have 4: "chain (\<lambda>i. \<Squnion>j. Y i j)" |
|
446 by (rule ch2ch_lub [OF 1 2]) |
|
447 show "(\<Squnion>i. \<Squnion>j. Y i j) \<sqsubseteq> (\<Squnion>i. Y i i)" |
|
448 apply (rule lub_below [OF 4]) |
|
449 apply (rule lub_below [OF 2]) |
|
450 apply (rule below_lub [OF 3]) |
|
451 apply (rule below_trans) |
|
452 apply (rule chain_mono [OF 1 max.cobounded1]) |
|
453 apply (rule chain_mono [OF 2 max.cobounded2]) |
|
454 done |
|
455 show "(\<Squnion>i. Y i i) \<sqsubseteq> (\<Squnion>i. \<Squnion>j. Y i j)" |
|
456 apply (rule lub_mono [OF 3 4]) |
|
457 apply (rule is_ub_thelub [OF 2]) |
|
458 done |
|
459 qed |
|
460 |
|
461 lemma ex_lub: |
|
462 assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" |
|
463 assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" |
|
464 shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>j. \<Squnion>i. Y i j)" |
|
465 by (simp add: diag_lub 1 2) |
|
466 |
|
467 end |
|
468 |
|
469 |
|
470 subsection \<open>Pointed cpos\<close> |
|
471 |
|
472 text \<open>The class pcpo of pointed cpos\<close> |
|
473 |
|
474 class pcpo = cpo + |
|
475 assumes least: "\<exists>x. \<forall>y. x \<sqsubseteq> y" |
|
476 begin |
|
477 |
|
478 definition bottom :: "'a" (\<open>\<bottom>\<close>) |
|
479 where "bottom = (THE x. \<forall>y. x \<sqsubseteq> y)" |
|
480 |
|
481 lemma minimal [iff]: "\<bottom> \<sqsubseteq> x" |
|
482 unfolding bottom_def |
|
483 apply (rule the1I2) |
|
484 apply (rule ex_ex1I) |
|
485 apply (rule least) |
|
486 apply (blast intro: below_antisym) |
|
487 apply simp |
|
488 done |
|
489 |
|
490 end |
|
491 |
|
492 text \<open>Old "UU" syntax:\<close> |
|
493 abbreviation (input) "UU \<equiv> bottom" |
|
494 |
|
495 text \<open>Simproc to rewrite \<^term>\<open>\<bottom> = x\<close> to \<^term>\<open>x = \<bottom>\<close>.\<close> |
|
496 setup \<open>Reorient_Proc.add (fn \<^Const_>\<open>bottom _\<close> => true | _ => false)\<close> |
|
497 simproc_setup reorient_bottom ("\<bottom> = x") = \<open>K Reorient_Proc.proc\<close> |
|
498 |
|
499 text \<open>useful lemmas about \<^term>\<open>\<bottom>\<close>\<close> |
|
500 |
|
501 lemma below_bottom_iff [simp]: "x \<sqsubseteq> \<bottom> \<longleftrightarrow> x = \<bottom>" |
|
502 by (simp add: po_eq_conv) |
|
503 |
|
504 lemma eq_bottom_iff: "x = \<bottom> \<longleftrightarrow> x \<sqsubseteq> \<bottom>" |
|
505 by simp |
|
506 |
|
507 lemma bottomI: "x \<sqsubseteq> \<bottom> \<Longrightarrow> x = \<bottom>" |
|
508 by (subst eq_bottom_iff) |
|
509 |
|
510 lemma lub_eq_bottom_iff: "chain Y \<Longrightarrow> (\<Squnion>i. Y i) = \<bottom> \<longleftrightarrow> (\<forall>i. Y i = \<bottom>)" |
|
511 by (simp only: eq_bottom_iff lub_below_iff) |
|
512 |
|
513 |
|
514 subsection \<open>Chain-finite and flat cpos\<close> |
|
515 |
|
516 text \<open>further useful classes for HOLCF domains\<close> |
|
517 |
|
518 class chfin = po + |
|
519 assumes chfin: "chain Y \<Longrightarrow> \<exists>n. max_in_chain n Y" |
|
520 begin |
|
521 |
|
522 subclass cpo |
|
523 apply standard |
|
524 apply (frule chfin) |
|
525 apply (blast intro: lub_finch1) |
|
526 done |
|
527 |
|
528 lemma chfin2finch: "chain Y \<Longrightarrow> finite_chain Y" |
|
529 by (simp add: chfin finite_chain_def) |
|
530 |
|
531 end |
|
532 |
|
533 class flat = pcpo + |
|
534 assumes ax_flat: "x \<sqsubseteq> y \<Longrightarrow> x = \<bottom> \<or> x = y" |
|
535 begin |
|
536 |
|
537 subclass chfin |
|
538 proof |
|
539 fix Y |
|
540 assume *: "chain Y" |
|
541 show "\<exists>n. max_in_chain n Y" |
|
542 apply (unfold max_in_chain_def) |
|
543 apply (cases "\<forall>i. Y i = \<bottom>") |
|
544 apply simp |
|
545 apply simp |
|
546 apply (erule exE) |
|
547 apply (rule_tac x="i" in exI) |
|
548 apply clarify |
|
549 using * apply (blast dest: chain_mono ax_flat) |
|
550 done |
|
551 qed |
|
552 |
|
553 lemma flat_below_iff: "x \<sqsubseteq> y \<longleftrightarrow> x = \<bottom> \<or> x = y" |
|
554 by (safe dest!: ax_flat) |
|
555 |
|
556 lemma flat_eq: "a \<noteq> \<bottom> \<Longrightarrow> a \<sqsubseteq> b = (a = b)" |
|
557 by (safe dest!: ax_flat) |
|
558 |
|
559 end |
|
560 |
|
561 subsection \<open>Discrete cpos\<close> |
|
562 |
|
563 class discrete_cpo = below + |
|
564 assumes discrete_cpo [simp]: "x \<sqsubseteq> y \<longleftrightarrow> x = y" |
|
565 begin |
|
566 |
|
567 subclass po |
|
568 by standard simp_all |
|
569 |
|
570 text \<open>In a discrete cpo, every chain is constant\<close> |
|
571 |
|
572 lemma discrete_chain_const: |
|
573 assumes S: "chain S" |
|
574 shows "\<exists>x. S = (\<lambda>i. x)" |
|
575 proof (intro exI ext) |
|
576 fix i :: nat |
|
577 from S le0 have "S 0 \<sqsubseteq> S i" by (rule chain_mono) |
|
578 then have "S 0 = S i" by simp |
|
579 then show "S i = S 0" by (rule sym) |
|
580 qed |
|
581 |
|
582 subclass chfin |
|
583 proof |
|
584 fix S :: "nat \<Rightarrow> 'a" |
|
585 assume S: "chain S" |
|
586 then have "\<exists>x. S = (\<lambda>i. x)" |
|
587 by (rule discrete_chain_const) |
|
588 then have "max_in_chain 0 S" |
|
589 by (auto simp: max_in_chain_def) |
|
590 then show "\<exists>i. max_in_chain i S" .. |
|
591 qed |
|
592 |
|
593 end |
|
594 |
|
595 |
|
596 section \<open>Continuity and monotonicity\<close> |
|
597 |
|
598 text \<open> |
|
599 Now we change the default class! Form now on all untyped type variables are |
|
600 of default class po |
|
601 \<close> |
|
602 |
|
603 default_sort po |
|
604 |
|
605 subsection \<open>Definitions\<close> |
|
606 |
|
607 definition monofun :: "('a \<Rightarrow> 'b) \<Rightarrow> bool" \<comment> \<open>monotonicity\<close> |
|
608 where "monofun f \<longleftrightarrow> (\<forall>x y. x \<sqsubseteq> y \<longrightarrow> f x \<sqsubseteq> f y)" |
|
609 |
|
610 definition cont :: "('a::cpo \<Rightarrow> 'b::cpo) \<Rightarrow> bool" |
|
611 where "cont f = (\<forall>Y. chain Y \<longrightarrow> range (\<lambda>i. f (Y i)) <<| f (\<Squnion>i. Y i))" |
|
612 |
|
613 lemma contI: "(\<And>Y. chain Y \<Longrightarrow> range (\<lambda>i. f (Y i)) <<| f (\<Squnion>i. Y i)) \<Longrightarrow> cont f" |
|
614 by (simp add: cont_def) |
|
615 |
|
616 lemma contE: "cont f \<Longrightarrow> chain Y \<Longrightarrow> range (\<lambda>i. f (Y i)) <<| f (\<Squnion>i. Y i)" |
|
617 by (simp add: cont_def) |
|
618 |
|
619 lemma monofunI: "(\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y) \<Longrightarrow> monofun f" |
|
620 by (simp add: monofun_def) |
|
621 |
|
622 lemma monofunE: "monofun f \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y" |
|
623 by (simp add: monofun_def) |
|
624 |
|
625 |
|
626 subsection \<open>Equivalence of alternate definition\<close> |
|
627 |
|
628 text \<open>monotone functions map chains to chains\<close> |
|
629 |
|
630 lemma ch2ch_monofun: "monofun f \<Longrightarrow> chain Y \<Longrightarrow> chain (\<lambda>i. f (Y i))" |
|
631 apply (rule chainI) |
|
632 apply (erule monofunE) |
|
633 apply (erule chainE) |
|
634 done |
|
635 |
|
636 text \<open>monotone functions map upper bound to upper bounds\<close> |
|
637 |
|
638 lemma ub2ub_monofun: "monofun f \<Longrightarrow> range Y <| u \<Longrightarrow> range (\<lambda>i. f (Y i)) <| f u" |
|
639 apply (rule ub_rangeI) |
|
640 apply (erule monofunE) |
|
641 apply (erule ub_rangeD) |
|
642 done |
|
643 |
|
644 text \<open>a lemma about binary chains\<close> |
|
645 |
|
646 lemma binchain_cont: "cont f \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> range (\<lambda>i::nat. f (if i = 0 then x else y)) <<| f y" |
|
647 apply (subgoal_tac "f (\<Squnion>i::nat. if i = 0 then x else y) = f y") |
|
648 apply (erule subst) |
|
649 apply (erule contE) |
|
650 apply (erule bin_chain) |
|
651 apply (rule_tac f=f in arg_cong) |
|
652 apply (erule is_lub_bin_chain [THEN lub_eqI]) |
|
653 done |
|
654 |
|
655 text \<open>continuity implies monotonicity\<close> |
|
656 |
|
657 lemma cont2mono: "cont f \<Longrightarrow> monofun f" |
|
658 apply (rule monofunI) |
|
659 apply (drule (1) binchain_cont) |
|
660 apply (drule_tac i=0 in is_lub_rangeD1) |
|
661 apply simp |
|
662 done |
|
663 |
|
664 lemmas cont2monofunE = cont2mono [THEN monofunE] |
|
665 |
|
666 lemmas ch2ch_cont = cont2mono [THEN ch2ch_monofun] |
|
667 |
|
668 text \<open>continuity implies preservation of lubs\<close> |
|
669 |
|
670 lemma cont2contlubE: "cont f \<Longrightarrow> chain Y \<Longrightarrow> f (\<Squnion>i. Y i) = (\<Squnion>i. f (Y i))" |
|
671 apply (rule lub_eqI [symmetric]) |
|
672 apply (erule (1) contE) |
|
673 done |
|
674 |
|
675 lemma contI2: |
|
676 fixes f :: "'a::cpo \<Rightarrow> 'b::cpo" |
|
677 assumes mono: "monofun f" |
|
678 assumes below: "\<And>Y. \<lbrakk>chain Y; chain (\<lambda>i. f (Y i))\<rbrakk> \<Longrightarrow> f (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. f (Y i))" |
|
679 shows "cont f" |
|
680 proof (rule contI) |
|
681 fix Y :: "nat \<Rightarrow> 'a" |
|
682 assume Y: "chain Y" |
|
683 with mono have fY: "chain (\<lambda>i. f (Y i))" |
|
684 by (rule ch2ch_monofun) |
|
685 have "(\<Squnion>i. f (Y i)) = f (\<Squnion>i. Y i)" |
|
686 apply (rule below_antisym) |
|
687 apply (rule lub_below [OF fY]) |
|
688 apply (rule monofunE [OF mono]) |
|
689 apply (rule is_ub_thelub [OF Y]) |
|
690 apply (rule below [OF Y fY]) |
|
691 done |
|
692 with fY show "range (\<lambda>i. f (Y i)) <<| f (\<Squnion>i. Y i)" |
|
693 by (rule thelubE) |
|
694 qed |
|
695 |
|
696 |
|
697 subsection \<open>Collection of continuity rules\<close> |
|
698 |
|
699 named_theorems cont2cont "continuity intro rule" |
|
700 |
|
701 |
|
702 subsection \<open>Continuity of basic functions\<close> |
|
703 |
|
704 text \<open>The identity function is continuous\<close> |
|
705 |
|
706 lemma cont_id [simp, cont2cont]: "cont (\<lambda>x. x)" |
|
707 apply (rule contI) |
|
708 apply (erule cpo_lubI) |
|
709 done |
|
710 |
|
711 text \<open>constant functions are continuous\<close> |
|
712 |
|
713 lemma cont_const [simp, cont2cont]: "cont (\<lambda>x. c)" |
|
714 using is_lub_const by (rule contI) |
|
715 |
|
716 text \<open>application of functions is continuous\<close> |
|
717 |
|
718 lemma cont_apply: |
|
719 fixes f :: "'a::cpo \<Rightarrow> 'b::cpo \<Rightarrow> 'c::cpo" and t :: "'a \<Rightarrow> 'b" |
|
720 assumes 1: "cont (\<lambda>x. t x)" |
|
721 assumes 2: "\<And>x. cont (\<lambda>y. f x y)" |
|
722 assumes 3: "\<And>y. cont (\<lambda>x. f x y)" |
|
723 shows "cont (\<lambda>x. (f x) (t x))" |
|
724 proof (rule contI2 [OF monofunI]) |
|
725 fix x y :: "'a" |
|
726 assume "x \<sqsubseteq> y" |
|
727 then show "f x (t x) \<sqsubseteq> f y (t y)" |
|
728 by (auto intro: cont2monofunE [OF 1] |
|
729 cont2monofunE [OF 2] |
|
730 cont2monofunE [OF 3] |
|
731 below_trans) |
|
732 next |
|
733 fix Y :: "nat \<Rightarrow> 'a" |
|
734 assume "chain Y" |
|
735 then show "f (\<Squnion>i. Y i) (t (\<Squnion>i. Y i)) \<sqsubseteq> (\<Squnion>i. f (Y i) (t (Y i)))" |
|
736 by (simp only: cont2contlubE [OF 1] ch2ch_cont [OF 1] |
|
737 cont2contlubE [OF 2] ch2ch_cont [OF 2] |
|
738 cont2contlubE [OF 3] ch2ch_cont [OF 3] |
|
739 diag_lub below_refl) |
|
740 qed |
|
741 |
|
742 lemma cont_compose: "cont c \<Longrightarrow> cont (\<lambda>x. f x) \<Longrightarrow> cont (\<lambda>x. c (f x))" |
|
743 by (rule cont_apply [OF _ _ cont_const]) |
|
744 |
|
745 text \<open>Least upper bounds preserve continuity\<close> |
|
746 |
|
747 lemma cont2cont_lub [simp]: |
|
748 assumes chain: "\<And>x. chain (\<lambda>i. F i x)" |
|
749 and cont: "\<And>i. cont (\<lambda>x. F i x)" |
|
750 shows "cont (\<lambda>x. \<Squnion>i. F i x)" |
|
751 apply (rule contI2) |
|
752 apply (simp add: monofunI cont2monofunE [OF cont] lub_mono chain) |
|
753 apply (simp add: cont2contlubE [OF cont]) |
|
754 apply (simp add: diag_lub ch2ch_cont [OF cont] chain) |
|
755 done |
|
756 |
|
757 text \<open>if-then-else is continuous\<close> |
|
758 |
|
759 lemma cont_if [simp, cont2cont]: "cont f \<Longrightarrow> cont g \<Longrightarrow> cont (\<lambda>x. if b then f x else g x)" |
|
760 by (induct b) simp_all |
|
761 |
|
762 |
|
763 subsection \<open>Finite chains and flat pcpos\<close> |
|
764 |
|
765 text \<open>Monotone functions map finite chains to finite chains.\<close> |
|
766 |
|
767 lemma monofun_finch2finch: "monofun f \<Longrightarrow> finite_chain Y \<Longrightarrow> finite_chain (\<lambda>n. f (Y n))" |
|
768 by (force simp add: finite_chain_def ch2ch_monofun max_in_chain_def) |
|
769 |
|
770 text \<open>The same holds for continuous functions.\<close> |
|
771 |
|
772 lemma cont_finch2finch: "cont f \<Longrightarrow> finite_chain Y \<Longrightarrow> finite_chain (\<lambda>n. f (Y n))" |
|
773 by (rule cont2mono [THEN monofun_finch2finch]) |
|
774 |
|
775 text \<open>All monotone functions with chain-finite domain are continuous.\<close> |
|
776 |
|
777 lemma chfindom_monofun2cont: "monofun f \<Longrightarrow> cont f" |
|
778 for f :: "'a::chfin \<Rightarrow> 'b::cpo" |
|
779 apply (erule contI2) |
|
780 apply (frule chfin2finch) |
|
781 apply (clarsimp simp add: finite_chain_def) |
|
782 apply (subgoal_tac "max_in_chain i (\<lambda>i. f (Y i))") |
|
783 apply (simp add: maxinch_is_thelub ch2ch_monofun) |
|
784 apply (force simp add: max_in_chain_def) |
|
785 done |
|
786 |
|
787 text \<open>All strict functions with flat domain are continuous.\<close> |
|
788 |
|
789 lemma flatdom_strict2mono: "f \<bottom> = \<bottom> \<Longrightarrow> monofun f" |
|
790 for f :: "'a::flat \<Rightarrow> 'b::pcpo" |
|
791 apply (rule monofunI) |
|
792 apply (drule ax_flat) |
|
793 apply auto |
|
794 done |
|
795 |
|
796 lemma flatdom_strict2cont: "f \<bottom> = \<bottom> \<Longrightarrow> cont f" |
|
797 for f :: "'a::flat \<Rightarrow> 'b::pcpo" |
|
798 by (rule flatdom_strict2mono [THEN chfindom_monofun2cont]) |
|
799 |
|
800 text \<open>All functions with discrete domain are continuous.\<close> |
|
801 |
|
802 lemma cont_discrete_cpo [simp, cont2cont]: "cont f" |
|
803 for f :: "'a::discrete_cpo \<Rightarrow> 'b::cpo" |
|
804 apply (rule contI) |
|
805 apply (drule discrete_chain_const, clarify) |
|
806 apply simp |
|
807 done |
|
808 |
|
809 section \<open>Admissibility and compactness\<close> |
|
810 |
|
811 default_sort cpo |
|
812 |
|
813 subsection \<open>Definitions\<close> |
|
814 |
|
815 definition adm :: "('a::cpo \<Rightarrow> bool) \<Rightarrow> bool" |
|
816 where "adm P \<longleftrightarrow> (\<forall>Y. chain Y \<longrightarrow> (\<forall>i. P (Y i)) \<longrightarrow> P (\<Squnion>i. Y i))" |
|
817 |
|
818 lemma admI: "(\<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i)\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)) \<Longrightarrow> adm P" |
|
819 unfolding adm_def by fast |
|
820 |
|
821 lemma admD: "adm P \<Longrightarrow> chain Y \<Longrightarrow> (\<And>i. P (Y i)) \<Longrightarrow> P (\<Squnion>i. Y i)" |
|
822 unfolding adm_def by fast |
|
823 |
|
824 lemma admD2: "adm (\<lambda>x. \<not> P x) \<Longrightarrow> chain Y \<Longrightarrow> P (\<Squnion>i. Y i) \<Longrightarrow> \<exists>i. P (Y i)" |
|
825 unfolding adm_def by fast |
|
826 |
|
827 lemma triv_admI: "\<forall>x. P x \<Longrightarrow> adm P" |
|
828 by (rule admI) (erule spec) |
|
829 |
|
830 |
|
831 subsection \<open>Admissibility on chain-finite types\<close> |
|
832 |
|
833 text \<open>For chain-finite (easy) types every formula is admissible.\<close> |
|
834 |
|
835 lemma adm_chfin [simp]: "adm P" |
|
836 for P :: "'a::chfin \<Rightarrow> bool" |
|
837 by (rule admI, frule chfin, auto simp add: maxinch_is_thelub) |
|
838 |
|
839 |
|
840 subsection \<open>Admissibility of special formulae and propagation\<close> |
|
841 |
|
842 lemma adm_const [simp]: "adm (\<lambda>x. t)" |
|
843 by (rule admI, simp) |
|
844 |
|
845 lemma adm_conj [simp]: "adm (\<lambda>x. P x) \<Longrightarrow> adm (\<lambda>x. Q x) \<Longrightarrow> adm (\<lambda>x. P x \<and> Q x)" |
|
846 by (fast intro: admI elim: admD) |
|
847 |
|
848 lemma adm_all [simp]: "(\<And>y. adm (\<lambda>x. P x y)) \<Longrightarrow> adm (\<lambda>x. \<forall>y. P x y)" |
|
849 by (fast intro: admI elim: admD) |
|
850 |
|
851 lemma adm_ball [simp]: "(\<And>y. y \<in> A \<Longrightarrow> adm (\<lambda>x. P x y)) \<Longrightarrow> adm (\<lambda>x. \<forall>y\<in>A. P x y)" |
|
852 by (fast intro: admI elim: admD) |
|
853 |
|
854 text \<open>Admissibility for disjunction is hard to prove. It requires 2 lemmas.\<close> |
|
855 |
|
856 lemma adm_disj_lemma1: |
|
857 assumes adm: "adm P" |
|
858 assumes chain: "chain Y" |
|
859 assumes P: "\<forall>i. \<exists>j\<ge>i. P (Y j)" |
|
860 shows "P (\<Squnion>i. Y i)" |
|
861 proof - |
|
862 define f where "f i = (LEAST j. i \<le> j \<and> P (Y j))" for i |
|
863 have chain': "chain (\<lambda>i. Y (f i))" |
|
864 unfolding f_def |
|
865 apply (rule chainI) |
|
866 apply (rule chain_mono [OF chain]) |
|
867 apply (rule Least_le) |
|
868 apply (rule LeastI2_ex) |
|
869 apply (simp_all add: P) |
|
870 done |
|
871 have f1: "\<And>i. i \<le> f i" and f2: "\<And>i. P (Y (f i))" |
|
872 using LeastI_ex [OF P [rule_format]] by (simp_all add: f_def) |
|
873 have lub_eq: "(\<Squnion>i. Y i) = (\<Squnion>i. Y (f i))" |
|
874 apply (rule below_antisym) |
|
875 apply (rule lub_mono [OF chain chain']) |
|
876 apply (rule chain_mono [OF chain f1]) |
|
877 apply (rule lub_range_mono [OF _ chain chain']) |
|
878 apply clarsimp |
|
879 done |
|
880 show "P (\<Squnion>i. Y i)" |
|
881 unfolding lub_eq using adm chain' f2 by (rule admD) |
|
882 qed |
|
883 |
|
884 lemma adm_disj_lemma2: "\<forall>n::nat. P n \<or> Q n \<Longrightarrow> (\<forall>i. \<exists>j\<ge>i. P j) \<or> (\<forall>i. \<exists>j\<ge>i. Q j)" |
|
885 apply (erule contrapos_pp) |
|
886 apply (clarsimp, rename_tac a b) |
|
887 apply (rule_tac x="max a b" in exI) |
|
888 apply simp |
|
889 done |
|
890 |
|
891 lemma adm_disj [simp]: "adm (\<lambda>x. P x) \<Longrightarrow> adm (\<lambda>x. Q x) \<Longrightarrow> adm (\<lambda>x. P x \<or> Q x)" |
|
892 apply (rule admI) |
|
893 apply (erule adm_disj_lemma2 [THEN disjE]) |
|
894 apply (erule (2) adm_disj_lemma1 [THEN disjI1]) |
|
895 apply (erule (2) adm_disj_lemma1 [THEN disjI2]) |
|
896 done |
|
897 |
|
898 lemma adm_imp [simp]: "adm (\<lambda>x. \<not> P x) \<Longrightarrow> adm (\<lambda>x. Q x) \<Longrightarrow> adm (\<lambda>x. P x \<longrightarrow> Q x)" |
|
899 by (subst imp_conv_disj) (rule adm_disj) |
|
900 |
|
901 lemma adm_iff [simp]: "adm (\<lambda>x. P x \<longrightarrow> Q x) \<Longrightarrow> adm (\<lambda>x. Q x \<longrightarrow> P x) \<Longrightarrow> adm (\<lambda>x. P x \<longleftrightarrow> Q x)" |
|
902 by (subst iff_conv_conj_imp) (rule adm_conj) |
|
903 |
|
904 text \<open>admissibility and continuity\<close> |
|
905 |
|
906 lemma adm_below [simp]: "cont (\<lambda>x. u x) \<Longrightarrow> cont (\<lambda>x. v x) \<Longrightarrow> adm (\<lambda>x. u x \<sqsubseteq> v x)" |
|
907 by (simp add: adm_def cont2contlubE lub_mono ch2ch_cont) |
|
908 |
|
909 lemma adm_eq [simp]: "cont (\<lambda>x. u x) \<Longrightarrow> cont (\<lambda>x. v x) \<Longrightarrow> adm (\<lambda>x. u x = v x)" |
|
910 by (simp add: po_eq_conv) |
|
911 |
|
912 lemma adm_subst: "cont (\<lambda>x. t x) \<Longrightarrow> adm P \<Longrightarrow> adm (\<lambda>x. P (t x))" |
|
913 by (simp add: adm_def cont2contlubE ch2ch_cont) |
|
914 |
|
915 lemma adm_not_below [simp]: "cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. t x \<notsqsubseteq> u)" |
|
916 by (rule admI) (simp add: cont2contlubE ch2ch_cont lub_below_iff) |
|
917 |
|
918 |
|
919 subsection \<open>Compactness\<close> |
|
920 |
|
921 definition compact :: "'a::cpo \<Rightarrow> bool" |
|
922 where "compact k = adm (\<lambda>x. k \<notsqsubseteq> x)" |
|
923 |
|
924 lemma compactI: "adm (\<lambda>x. k \<notsqsubseteq> x) \<Longrightarrow> compact k" |
|
925 unfolding compact_def . |
|
926 |
|
927 lemma compactD: "compact k \<Longrightarrow> adm (\<lambda>x. k \<notsqsubseteq> x)" |
|
928 unfolding compact_def . |
|
929 |
|
930 lemma compactI2: "(\<And>Y. \<lbrakk>chain Y; x \<sqsubseteq> (\<Squnion>i. Y i)\<rbrakk> \<Longrightarrow> \<exists>i. x \<sqsubseteq> Y i) \<Longrightarrow> compact x" |
|
931 unfolding compact_def adm_def by fast |
|
932 |
|
933 lemma compactD2: "compact x \<Longrightarrow> chain Y \<Longrightarrow> x \<sqsubseteq> (\<Squnion>i. Y i) \<Longrightarrow> \<exists>i. x \<sqsubseteq> Y i" |
|
934 unfolding compact_def adm_def by fast |
|
935 |
|
936 lemma compact_below_lub_iff: "compact x \<Longrightarrow> chain Y \<Longrightarrow> x \<sqsubseteq> (\<Squnion>i. Y i) \<longleftrightarrow> (\<exists>i. x \<sqsubseteq> Y i)" |
|
937 by (fast intro: compactD2 elim: below_lub) |
|
938 |
|
939 lemma compact_chfin [simp]: "compact x" |
|
940 for x :: "'a::chfin" |
|
941 by (rule compactI [OF adm_chfin]) |
|
942 |
|
943 lemma compact_imp_max_in_chain: "chain Y \<Longrightarrow> compact (\<Squnion>i. Y i) \<Longrightarrow> \<exists>i. max_in_chain i Y" |
|
944 apply (drule (1) compactD2, simp) |
|
945 apply (erule exE, rule_tac x=i in exI) |
|
946 apply (rule max_in_chainI) |
|
947 apply (rule below_antisym) |
|
948 apply (erule (1) chain_mono) |
|
949 apply (erule (1) below_trans [OF is_ub_thelub]) |
|
950 done |
|
951 |
|
952 text \<open>admissibility and compactness\<close> |
|
953 |
|
954 lemma adm_compact_not_below [simp]: |
|
955 "compact k \<Longrightarrow> cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. k \<notsqsubseteq> t x)" |
|
956 unfolding compact_def by (rule adm_subst) |
|
957 |
|
958 lemma adm_neq_compact [simp]: "compact k \<Longrightarrow> cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. t x \<noteq> k)" |
|
959 by (simp add: po_eq_conv) |
|
960 |
|
961 lemma adm_compact_neq [simp]: "compact k \<Longrightarrow> cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. k \<noteq> t x)" |
|
962 by (simp add: po_eq_conv) |
|
963 |
|
964 lemma compact_bottom [simp, intro]: "compact \<bottom>" |
|
965 by (rule compactI) simp |
|
966 |
|
967 text \<open>Any upward-closed predicate is admissible.\<close> |
|
968 |
|
969 lemma adm_upward: |
|
970 assumes P: "\<And>x y. \<lbrakk>P x; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> P y" |
|
971 shows "adm P" |
|
972 by (rule admI, drule spec, erule P, erule is_ub_thelub) |
|
973 |
|
974 lemmas adm_lemmas = |
|
975 adm_const adm_conj adm_all adm_ball adm_disj adm_imp adm_iff |
|
976 adm_below adm_eq adm_not_below |
|
977 adm_compact_not_below adm_compact_neq adm_neq_compact |
|
978 |
|
979 section \<open>Class instances for the full function space\<close> |
|
980 |
|
981 subsection \<open>Full function space is a partial order\<close> |
|
982 |
|
983 instantiation "fun" :: (type, below) below |
|
984 begin |
|
985 |
|
986 definition below_fun_def: "(\<sqsubseteq>) \<equiv> (\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x)" |
|
987 |
|
988 instance .. |
|
989 end |
|
990 |
|
991 instance "fun" :: (type, po) po |
|
992 proof |
|
993 fix f :: "'a \<Rightarrow> 'b" |
|
994 show "f \<sqsubseteq> f" |
|
995 by (simp add: below_fun_def) |
|
996 next |
|
997 fix f g :: "'a \<Rightarrow> 'b" |
|
998 assume "f \<sqsubseteq> g" and "g \<sqsubseteq> f" then show "f = g" |
|
999 by (simp add: below_fun_def fun_eq_iff below_antisym) |
|
1000 next |
|
1001 fix f g h :: "'a \<Rightarrow> 'b" |
|
1002 assume "f \<sqsubseteq> g" and "g \<sqsubseteq> h" then show "f \<sqsubseteq> h" |
|
1003 unfolding below_fun_def by (fast elim: below_trans) |
|
1004 qed |
|
1005 |
|
1006 lemma fun_below_iff: "f \<sqsubseteq> g \<longleftrightarrow> (\<forall>x. f x \<sqsubseteq> g x)" |
|
1007 by (simp add: below_fun_def) |
|
1008 |
|
1009 lemma fun_belowI: "(\<And>x. f x \<sqsubseteq> g x) \<Longrightarrow> f \<sqsubseteq> g" |
|
1010 by (simp add: below_fun_def) |
|
1011 |
|
1012 lemma fun_belowD: "f \<sqsubseteq> g \<Longrightarrow> f x \<sqsubseteq> g x" |
|
1013 by (simp add: below_fun_def) |
|
1014 |
|
1015 |
|
1016 subsection \<open>Full function space is chain complete\<close> |
|
1017 |
|
1018 text \<open>Properties of chains of functions.\<close> |
|
1019 |
|
1020 lemma fun_chain_iff: "chain S \<longleftrightarrow> (\<forall>x. chain (\<lambda>i. S i x))" |
|
1021 by (auto simp: chain_def fun_below_iff) |
|
1022 |
|
1023 lemma ch2ch_fun: "chain S \<Longrightarrow> chain (\<lambda>i. S i x)" |
|
1024 by (simp add: chain_def below_fun_def) |
|
1025 |
|
1026 lemma ch2ch_lambda: "(\<And>x. chain (\<lambda>i. S i x)) \<Longrightarrow> chain S" |
|
1027 by (simp add: chain_def below_fun_def) |
|
1028 |
|
1029 text \<open>Type \<^typ>\<open>'a::type \<Rightarrow> 'b::cpo\<close> is chain complete\<close> |
|
1030 |
|
1031 lemma is_lub_lambda: "(\<And>x. range (\<lambda>i. Y i x) <<| f x) \<Longrightarrow> range Y <<| f" |
|
1032 by (simp add: is_lub_def is_ub_def below_fun_def) |
|
1033 |
|
1034 lemma is_lub_fun: "chain S \<Longrightarrow> range S <<| (\<lambda>x. \<Squnion>i. S i x)" |
|
1035 for S :: "nat \<Rightarrow> 'a::type \<Rightarrow> 'b::cpo" |
|
1036 apply (rule is_lub_lambda) |
|
1037 apply (rule cpo_lubI) |
|
1038 apply (erule ch2ch_fun) |
|
1039 done |
|
1040 |
|
1041 lemma lub_fun: "chain S \<Longrightarrow> (\<Squnion>i. S i) = (\<lambda>x. \<Squnion>i. S i x)" |
|
1042 for S :: "nat \<Rightarrow> 'a::type \<Rightarrow> 'b::cpo" |
|
1043 by (rule is_lub_fun [THEN lub_eqI]) |
|
1044 |
|
1045 instance "fun" :: (type, cpo) cpo |
|
1046 by intro_classes (rule exI, erule is_lub_fun) |
|
1047 |
|
1048 instance "fun" :: (type, discrete_cpo) discrete_cpo |
|
1049 proof |
|
1050 fix f g :: "'a \<Rightarrow> 'b" |
|
1051 show "f \<sqsubseteq> g \<longleftrightarrow> f = g" |
|
1052 by (simp add: fun_below_iff fun_eq_iff) |
|
1053 qed |
|
1054 |
|
1055 |
|
1056 subsection \<open>Full function space is pointed\<close> |
|
1057 |
|
1058 lemma minimal_fun: "(\<lambda>x. \<bottom>) \<sqsubseteq> f" |
|
1059 by (simp add: below_fun_def) |
|
1060 |
|
1061 instance "fun" :: (type, pcpo) pcpo |
|
1062 by standard (fast intro: minimal_fun) |
|
1063 |
|
1064 lemma inst_fun_pcpo: "\<bottom> = (\<lambda>x. \<bottom>)" |
|
1065 by (rule minimal_fun [THEN bottomI, symmetric]) |
|
1066 |
|
1067 lemma app_strict [simp]: "\<bottom> x = \<bottom>" |
|
1068 by (simp add: inst_fun_pcpo) |
|
1069 |
|
1070 lemma lambda_strict: "(\<lambda>x. \<bottom>) = \<bottom>" |
|
1071 by (rule bottomI, rule minimal_fun) |
|
1072 |
|
1073 |
|
1074 subsection \<open>Propagation of monotonicity and continuity\<close> |
|
1075 |
|
1076 text \<open>The lub of a chain of monotone functions is monotone.\<close> |
|
1077 |
|
1078 lemma adm_monofun: "adm monofun" |
|
1079 by (rule admI) (simp add: lub_fun fun_chain_iff monofun_def lub_mono) |
|
1080 |
|
1081 text \<open>The lub of a chain of continuous functions is continuous.\<close> |
|
1082 |
|
1083 lemma adm_cont: "adm cont" |
|
1084 by (rule admI) (simp add: lub_fun fun_chain_iff) |
|
1085 |
|
1086 text \<open>Function application preserves monotonicity and continuity.\<close> |
|
1087 |
|
1088 lemma mono2mono_fun: "monofun f \<Longrightarrow> monofun (\<lambda>x. f x y)" |
|
1089 by (simp add: monofun_def fun_below_iff) |
|
1090 |
|
1091 lemma cont2cont_fun: "cont f \<Longrightarrow> cont (\<lambda>x. f x y)" |
|
1092 apply (rule contI2) |
|
1093 apply (erule cont2mono [THEN mono2mono_fun]) |
|
1094 apply (simp add: cont2contlubE lub_fun ch2ch_cont) |
|
1095 done |
|
1096 |
|
1097 lemma cont_fun: "cont (\<lambda>f. f x)" |
|
1098 using cont_id by (rule cont2cont_fun) |
|
1099 |
|
1100 text \<open> |
|
1101 Lambda abstraction preserves monotonicity and continuity. |
|
1102 (Note \<open>(\<lambda>x. \<lambda>y. f x y) = f\<close>.) |
|
1103 \<close> |
|
1104 |
|
1105 lemma mono2mono_lambda: "(\<And>y. monofun (\<lambda>x. f x y)) \<Longrightarrow> monofun f" |
|
1106 by (simp add: monofun_def fun_below_iff) |
|
1107 |
|
1108 lemma cont2cont_lambda [simp]: |
|
1109 assumes f: "\<And>y. cont (\<lambda>x. f x y)" |
|
1110 shows "cont f" |
|
1111 by (rule contI, rule is_lub_lambda, rule contE [OF f]) |
|
1112 |
|
1113 text \<open>What D.A.Schmidt calls continuity of abstraction; never used here\<close> |
|
1114 |
|
1115 lemma contlub_lambda: "(\<And>x. chain (\<lambda>i. S i x)) \<Longrightarrow> (\<lambda>x. \<Squnion>i. S i x) = (\<Squnion>i. (\<lambda>x. S i x))" |
|
1116 for S :: "nat \<Rightarrow> 'a::type \<Rightarrow> 'b::cpo" |
|
1117 by (simp add: lub_fun ch2ch_lambda) |
|
1118 |
|
1119 section \<open>The cpo of cartesian products\<close> |
|
1120 |
|
1121 default_sort cpo |
|
1122 |
|
1123 |
|
1124 subsection \<open>Unit type is a pcpo\<close> |
|
1125 |
|
1126 instantiation unit :: discrete_cpo |
|
1127 begin |
|
1128 |
|
1129 definition below_unit_def [simp]: "x \<sqsubseteq> (y::unit) \<longleftrightarrow> True" |
|
1130 |
|
1131 instance |
|
1132 by standard simp |
|
1133 |
|
1134 end |
|
1135 |
|
1136 instance unit :: pcpo |
|
1137 by standard simp |
|
1138 |
|
1139 |
|
1140 subsection \<open>Product type is a partial order\<close> |
|
1141 |
|
1142 instantiation prod :: (below, below) below |
|
1143 begin |
|
1144 |
|
1145 definition below_prod_def: "(\<sqsubseteq>) \<equiv> \<lambda>p1 p2. (fst p1 \<sqsubseteq> fst p2 \<and> snd p1 \<sqsubseteq> snd p2)" |
|
1146 |
|
1147 instance .. |
|
1148 |
|
1149 end |
|
1150 |
|
1151 instance prod :: (po, po) po |
|
1152 proof |
|
1153 fix x :: "'a \<times> 'b" |
|
1154 show "x \<sqsubseteq> x" |
|
1155 by (simp add: below_prod_def) |
|
1156 next |
|
1157 fix x y :: "'a \<times> 'b" |
|
1158 assume "x \<sqsubseteq> y" "y \<sqsubseteq> x" |
|
1159 then show "x = y" |
|
1160 unfolding below_prod_def prod_eq_iff |
|
1161 by (fast intro: below_antisym) |
|
1162 next |
|
1163 fix x y z :: "'a \<times> 'b" |
|
1164 assume "x \<sqsubseteq> y" "y \<sqsubseteq> z" |
|
1165 then show "x \<sqsubseteq> z" |
|
1166 unfolding below_prod_def |
|
1167 by (fast intro: below_trans) |
|
1168 qed |
|
1169 |
|
1170 |
|
1171 subsection \<open>Monotonicity of \emph{Pair}, \emph{fst}, \emph{snd}\<close> |
|
1172 |
|
1173 lemma prod_belowI: "fst p \<sqsubseteq> fst q \<Longrightarrow> snd p \<sqsubseteq> snd q \<Longrightarrow> p \<sqsubseteq> q" |
|
1174 by (simp add: below_prod_def) |
|
1175 |
|
1176 lemma Pair_below_iff [simp]: "(a, b) \<sqsubseteq> (c, d) \<longleftrightarrow> a \<sqsubseteq> c \<and> b \<sqsubseteq> d" |
|
1177 by (simp add: below_prod_def) |
|
1178 |
|
1179 text \<open>Pair \<open>(_,_)\<close> is monotone in both arguments\<close> |
|
1180 |
|
1181 lemma monofun_pair1: "monofun (\<lambda>x. (x, y))" |
|
1182 by (simp add: monofun_def) |
|
1183 |
|
1184 lemma monofun_pair2: "monofun (\<lambda>y. (x, y))" |
|
1185 by (simp add: monofun_def) |
|
1186 |
|
1187 lemma monofun_pair: "x1 \<sqsubseteq> x2 \<Longrightarrow> y1 \<sqsubseteq> y2 \<Longrightarrow> (x1, y1) \<sqsubseteq> (x2, y2)" |
|
1188 by simp |
|
1189 |
|
1190 lemma ch2ch_Pair [simp]: "chain X \<Longrightarrow> chain Y \<Longrightarrow> chain (\<lambda>i. (X i, Y i))" |
|
1191 by (rule chainI, simp add: chainE) |
|
1192 |
|
1193 text \<open>\<^term>\<open>fst\<close> and \<^term>\<open>snd\<close> are monotone\<close> |
|
1194 |
|
1195 lemma fst_monofun: "x \<sqsubseteq> y \<Longrightarrow> fst x \<sqsubseteq> fst y" |
|
1196 by (simp add: below_prod_def) |
|
1197 |
|
1198 lemma snd_monofun: "x \<sqsubseteq> y \<Longrightarrow> snd x \<sqsubseteq> snd y" |
|
1199 by (simp add: below_prod_def) |
|
1200 |
|
1201 lemma monofun_fst: "monofun fst" |
|
1202 by (simp add: monofun_def below_prod_def) |
|
1203 |
|
1204 lemma monofun_snd: "monofun snd" |
|
1205 by (simp add: monofun_def below_prod_def) |
|
1206 |
|
1207 lemmas ch2ch_fst [simp] = ch2ch_monofun [OF monofun_fst] |
|
1208 |
|
1209 lemmas ch2ch_snd [simp] = ch2ch_monofun [OF monofun_snd] |
|
1210 |
|
1211 lemma prod_chain_cases: |
|
1212 assumes chain: "chain Y" |
|
1213 obtains A B |
|
1214 where "chain A" and "chain B" and "Y = (\<lambda>i. (A i, B i))" |
|
1215 proof |
|
1216 from chain show "chain (\<lambda>i. fst (Y i))" |
|
1217 by (rule ch2ch_fst) |
|
1218 from chain show "chain (\<lambda>i. snd (Y i))" |
|
1219 by (rule ch2ch_snd) |
|
1220 show "Y = (\<lambda>i. (fst (Y i), snd (Y i)))" |
|
1221 by simp |
|
1222 qed |
|
1223 |
|
1224 |
|
1225 subsection \<open>Product type is a cpo\<close> |
|
1226 |
|
1227 lemma is_lub_Pair: "range A <<| x \<Longrightarrow> range B <<| y \<Longrightarrow> range (\<lambda>i. (A i, B i)) <<| (x, y)" |
|
1228 by (simp add: is_lub_def is_ub_def below_prod_def) |
|
1229 |
|
1230 lemma lub_Pair: "chain A \<Longrightarrow> chain B \<Longrightarrow> (\<Squnion>i. (A i, B i)) = (\<Squnion>i. A i, \<Squnion>i. B i)" |
|
1231 for A :: "nat \<Rightarrow> 'a::cpo" and B :: "nat \<Rightarrow> 'b::cpo" |
|
1232 by (fast intro: lub_eqI is_lub_Pair elim: thelubE) |
|
1233 |
|
1234 lemma is_lub_prod: |
|
1235 fixes S :: "nat \<Rightarrow> ('a::cpo \<times> 'b::cpo)" |
|
1236 assumes "chain S" |
|
1237 shows "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" |
|
1238 using assms by (auto elim: prod_chain_cases simp: is_lub_Pair cpo_lubI) |
|
1239 |
|
1240 lemma lub_prod: "chain S \<Longrightarrow> (\<Squnion>i. S i) = (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" |
|
1241 for S :: "nat \<Rightarrow> 'a::cpo \<times> 'b::cpo" |
|
1242 by (rule is_lub_prod [THEN lub_eqI]) |
|
1243 |
|
1244 instance prod :: (cpo, cpo) cpo |
|
1245 proof |
|
1246 fix S :: "nat \<Rightarrow> ('a \<times> 'b)" |
|
1247 assume "chain S" |
|
1248 then have "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" |
|
1249 by (rule is_lub_prod) |
|
1250 then show "\<exists>x. range S <<| x" .. |
|
1251 qed |
|
1252 |
|
1253 instance prod :: (discrete_cpo, discrete_cpo) discrete_cpo |
|
1254 proof |
|
1255 fix x y :: "'a \<times> 'b" |
|
1256 show "x \<sqsubseteq> y \<longleftrightarrow> x = y" |
|
1257 by (simp add: below_prod_def prod_eq_iff) |
|
1258 qed |
|
1259 |
|
1260 |
|
1261 subsection \<open>Product type is pointed\<close> |
|
1262 |
|
1263 lemma minimal_prod: "(\<bottom>, \<bottom>) \<sqsubseteq> p" |
|
1264 by (simp add: below_prod_def) |
|
1265 |
|
1266 instance prod :: (pcpo, pcpo) pcpo |
|
1267 by intro_classes (fast intro: minimal_prod) |
|
1268 |
|
1269 lemma inst_prod_pcpo: "\<bottom> = (\<bottom>, \<bottom>)" |
|
1270 by (rule minimal_prod [THEN bottomI, symmetric]) |
|
1271 |
|
1272 lemma Pair_bottom_iff [simp]: "(x, y) = \<bottom> \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>" |
|
1273 by (simp add: inst_prod_pcpo) |
|
1274 |
|
1275 lemma fst_strict [simp]: "fst \<bottom> = \<bottom>" |
|
1276 unfolding inst_prod_pcpo by (rule fst_conv) |
|
1277 |
|
1278 lemma snd_strict [simp]: "snd \<bottom> = \<bottom>" |
|
1279 unfolding inst_prod_pcpo by (rule snd_conv) |
|
1280 |
|
1281 lemma Pair_strict [simp]: "(\<bottom>, \<bottom>) = \<bottom>" |
|
1282 by simp |
|
1283 |
|
1284 lemma split_strict [simp]: "case_prod f \<bottom> = f \<bottom> \<bottom>" |
|
1285 by (simp add: split_def) |
|
1286 |
|
1287 |
|
1288 subsection \<open>Continuity of \emph{Pair}, \emph{fst}, \emph{snd}\<close> |
|
1289 |
|
1290 lemma cont_pair1: "cont (\<lambda>x. (x, y))" |
|
1291 apply (rule contI) |
|
1292 apply (rule is_lub_Pair) |
|
1293 apply (erule cpo_lubI) |
|
1294 apply (rule is_lub_const) |
|
1295 done |
|
1296 |
|
1297 lemma cont_pair2: "cont (\<lambda>y. (x, y))" |
|
1298 apply (rule contI) |
|
1299 apply (rule is_lub_Pair) |
|
1300 apply (rule is_lub_const) |
|
1301 apply (erule cpo_lubI) |
|
1302 done |
|
1303 |
|
1304 lemma cont_fst: "cont fst" |
|
1305 apply (rule contI) |
|
1306 apply (simp add: lub_prod) |
|
1307 apply (erule cpo_lubI [OF ch2ch_fst]) |
|
1308 done |
|
1309 |
|
1310 lemma cont_snd: "cont snd" |
|
1311 apply (rule contI) |
|
1312 apply (simp add: lub_prod) |
|
1313 apply (erule cpo_lubI [OF ch2ch_snd]) |
|
1314 done |
|
1315 |
|
1316 lemma cont2cont_Pair [simp, cont2cont]: |
|
1317 assumes f: "cont (\<lambda>x. f x)" |
|
1318 assumes g: "cont (\<lambda>x. g x)" |
|
1319 shows "cont (\<lambda>x. (f x, g x))" |
|
1320 apply (rule cont_apply [OF f cont_pair1]) |
|
1321 apply (rule cont_apply [OF g cont_pair2]) |
|
1322 apply (rule cont_const) |
|
1323 done |
|
1324 |
|
1325 lemmas cont2cont_fst [simp, cont2cont] = cont_compose [OF cont_fst] |
|
1326 |
|
1327 lemmas cont2cont_snd [simp, cont2cont] = cont_compose [OF cont_snd] |
|
1328 |
|
1329 lemma cont2cont_case_prod: |
|
1330 assumes f1: "\<And>a b. cont (\<lambda>x. f x a b)" |
|
1331 assumes f2: "\<And>x b. cont (\<lambda>a. f x a b)" |
|
1332 assumes f3: "\<And>x a. cont (\<lambda>b. f x a b)" |
|
1333 assumes g: "cont (\<lambda>x. g x)" |
|
1334 shows "cont (\<lambda>x. case g x of (a, b) \<Rightarrow> f x a b)" |
|
1335 unfolding split_def |
|
1336 apply (rule cont_apply [OF g]) |
|
1337 apply (rule cont_apply [OF cont_fst f2]) |
|
1338 apply (rule cont_apply [OF cont_snd f3]) |
|
1339 apply (rule cont_const) |
|
1340 apply (rule f1) |
|
1341 done |
|
1342 |
|
1343 lemma prod_contI: |
|
1344 assumes f1: "\<And>y. cont (\<lambda>x. f (x, y))" |
|
1345 assumes f2: "\<And>x. cont (\<lambda>y. f (x, y))" |
|
1346 shows "cont f" |
|
1347 proof - |
|
1348 have "cont (\<lambda>(x, y). f (x, y))" |
|
1349 by (intro cont2cont_case_prod f1 f2 cont2cont) |
|
1350 then show "cont f" |
|
1351 by (simp only: case_prod_eta) |
|
1352 qed |
|
1353 |
|
1354 lemma prod_cont_iff: "cont f \<longleftrightarrow> (\<forall>y. cont (\<lambda>x. f (x, y))) \<and> (\<forall>x. cont (\<lambda>y. f (x, y)))" |
|
1355 apply safe |
|
1356 apply (erule cont_compose [OF _ cont_pair1]) |
|
1357 apply (erule cont_compose [OF _ cont_pair2]) |
|
1358 apply (simp only: prod_contI) |
|
1359 done |
|
1360 |
|
1361 lemma cont2cont_case_prod' [simp, cont2cont]: |
|
1362 assumes f: "cont (\<lambda>p. f (fst p) (fst (snd p)) (snd (snd p)))" |
|
1363 assumes g: "cont (\<lambda>x. g x)" |
|
1364 shows "cont (\<lambda>x. case_prod (f x) (g x))" |
|
1365 using assms by (simp add: cont2cont_case_prod prod_cont_iff) |
|
1366 |
|
1367 text \<open>The simple version (due to Joachim Breitner) is needed if |
|
1368 either element type of the pair is not a cpo.\<close> |
|
1369 |
|
1370 lemma cont2cont_split_simple [simp, cont2cont]: |
|
1371 assumes "\<And>a b. cont (\<lambda>x. f x a b)" |
|
1372 shows "cont (\<lambda>x. case p of (a, b) \<Rightarrow> f x a b)" |
|
1373 using assms by (cases p) auto |
|
1374 |
|
1375 text \<open>Admissibility of predicates on product types.\<close> |
|
1376 |
|
1377 lemma adm_case_prod [simp]: |
|
1378 assumes "adm (\<lambda>x. P x (fst (f x)) (snd (f x)))" |
|
1379 shows "adm (\<lambda>x. case f x of (a, b) \<Rightarrow> P x a b)" |
|
1380 unfolding case_prod_beta using assms . |
|
1381 |
|
1382 |
|
1383 subsection \<open>Compactness and chain-finiteness\<close> |
|
1384 |
|
1385 lemma fst_below_iff: "fst x \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> (y, snd x)" |
|
1386 for x :: "'a \<times> 'b" |
|
1387 by (simp add: below_prod_def) |
|
1388 |
|
1389 lemma snd_below_iff: "snd x \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> (fst x, y)" |
|
1390 for x :: "'a \<times> 'b" |
|
1391 by (simp add: below_prod_def) |
|
1392 |
|
1393 lemma compact_fst: "compact x \<Longrightarrow> compact (fst x)" |
|
1394 by (rule compactI) (simp add: fst_below_iff) |
|
1395 |
|
1396 lemma compact_snd: "compact x \<Longrightarrow> compact (snd x)" |
|
1397 by (rule compactI) (simp add: snd_below_iff) |
|
1398 |
|
1399 lemma compact_Pair: "compact x \<Longrightarrow> compact y \<Longrightarrow> compact (x, y)" |
|
1400 by (rule compactI) (simp add: below_prod_def) |
|
1401 |
|
1402 lemma compact_Pair_iff [simp]: "compact (x, y) \<longleftrightarrow> compact x \<and> compact y" |
|
1403 apply (safe intro!: compact_Pair) |
|
1404 apply (drule compact_fst, simp) |
|
1405 apply (drule compact_snd, simp) |
|
1406 done |
|
1407 |
|
1408 instance prod :: (chfin, chfin) chfin |
|
1409 apply intro_classes |
|
1410 apply (erule compact_imp_max_in_chain) |
|
1411 apply (case_tac "\<Squnion>i. Y i", simp) |
|
1412 done |
|
1413 |
|
1414 section \<open>Discrete cpo types\<close> |
|
1415 |
|
1416 datatype 'a discr = Discr "'a :: type" |
|
1417 |
|
1418 subsection \<open>Discrete cpo class instance\<close> |
|
1419 |
|
1420 instantiation discr :: (type) discrete_cpo |
|
1421 begin |
|
1422 |
|
1423 definition "((\<sqsubseteq>) :: 'a discr \<Rightarrow> 'a discr \<Rightarrow> bool) = (=)" |
|
1424 |
|
1425 instance |
|
1426 by standard (simp add: below_discr_def) |
|
1427 |
|
1428 end |
|
1429 |
|
1430 |
|
1431 subsection \<open>\emph{undiscr}\<close> |
|
1432 |
|
1433 definition undiscr :: "('a::type)discr \<Rightarrow> 'a" |
|
1434 where "undiscr x = (case x of Discr y \<Rightarrow> y)" |
|
1435 |
|
1436 lemma undiscr_Discr [simp]: "undiscr (Discr x) = x" |
|
1437 by (simp add: undiscr_def) |
|
1438 |
|
1439 lemma Discr_undiscr [simp]: "Discr (undiscr y) = y" |
|
1440 by (induct y) simp |
|
1441 |
|
1442 end |