author | wenzelm |
Mon, 07 Jun 2010 18:09:18 +0200 | |
changeset 37363 | ca260a17e013 |
parent 37111 | 3f84f1f4de64 |
child 37678 | 0040bafffdef |
permissions | -rw-r--r-- |
29534 | 1 |
(* Title: HOLCF/Sum_Cpo.thy |
29130 | 2 |
Author: Brian Huffman |
3 |
*) |
|
4 |
||
5 |
header {* The cpo of disjoint sums *} |
|
6 |
||
29534 | 7 |
theory Sum_Cpo |
29130 | 8 |
imports Bifinite |
9 |
begin |
|
10 |
||
35900
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
huffman
parents:
31076
diff
changeset
|
11 |
subsection {* Ordering on sum type *} |
29130 | 12 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
13 |
instantiation "+" :: (below, below) below |
29130 | 14 |
begin |
15 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
16 |
definition below_sum_def: |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
17 |
"x \<sqsubseteq> y \<equiv> case x of |
29130 | 18 |
Inl a \<Rightarrow> (case y of Inl b \<Rightarrow> a \<sqsubseteq> b | Inr b \<Rightarrow> False) | |
19 |
Inr a \<Rightarrow> (case y of Inl b \<Rightarrow> False | Inr b \<Rightarrow> a \<sqsubseteq> b)" |
|
20 |
||
21 |
instance .. |
|
22 |
end |
|
23 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
24 |
lemma Inl_below_Inl [simp]: "Inl x \<sqsubseteq> Inl y = x \<sqsubseteq> y" |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
25 |
unfolding below_sum_def by simp |
29130 | 26 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
27 |
lemma Inr_below_Inr [simp]: "Inr x \<sqsubseteq> Inr y = x \<sqsubseteq> y" |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
28 |
unfolding below_sum_def by simp |
29130 | 29 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
30 |
lemma Inl_below_Inr [simp]: "\<not> Inl x \<sqsubseteq> Inr y" |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
31 |
unfolding below_sum_def by simp |
29130 | 32 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
33 |
lemma Inr_below_Inl [simp]: "\<not> Inr x \<sqsubseteq> Inl y" |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
34 |
unfolding below_sum_def by simp |
29130 | 35 |
|
36 |
lemma Inl_mono: "x \<sqsubseteq> y \<Longrightarrow> Inl x \<sqsubseteq> Inl y" |
|
37 |
by simp |
|
38 |
||
39 |
lemma Inr_mono: "x \<sqsubseteq> y \<Longrightarrow> Inr x \<sqsubseteq> Inr y" |
|
40 |
by simp |
|
41 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
42 |
lemma Inl_belowE: "\<lbrakk>Inl a \<sqsubseteq> x; \<And>b. \<lbrakk>x = Inl b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
29130 | 43 |
by (cases x, simp_all) |
44 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
45 |
lemma Inr_belowE: "\<lbrakk>Inr a \<sqsubseteq> x; \<And>b. \<lbrakk>x = Inr b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
29130 | 46 |
by (cases x, simp_all) |
47 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
48 |
lemmas sum_below_elims = Inl_belowE Inr_belowE |
29130 | 49 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
50 |
lemma sum_below_cases: |
29130 | 51 |
"\<lbrakk>x \<sqsubseteq> y; |
52 |
\<And>a b. \<lbrakk>x = Inl a; y = Inl b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R; |
|
53 |
\<And>a b. \<lbrakk>x = Inr a; y = Inr b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk> |
|
54 |
\<Longrightarrow> R" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
55 |
by (cases x, safe elim!: sum_below_elims, auto) |
29130 | 56 |
|
57 |
subsection {* Sum type is a complete partial order *} |
|
58 |
||
59 |
instance "+" :: (po, po) po |
|
60 |
proof |
|
61 |
fix x :: "'a + 'b" |
|
62 |
show "x \<sqsubseteq> x" |
|
63 |
by (induct x, simp_all) |
|
64 |
next |
|
65 |
fix x y :: "'a + 'b" |
|
66 |
assume "x \<sqsubseteq> y" and "y \<sqsubseteq> x" thus "x = y" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
67 |
by (induct x, auto elim!: sum_below_elims intro: below_antisym) |
29130 | 68 |
next |
69 |
fix x y z :: "'a + 'b" |
|
70 |
assume "x \<sqsubseteq> y" and "y \<sqsubseteq> z" thus "x \<sqsubseteq> z" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
71 |
by (induct x, auto elim!: sum_below_elims intro: below_trans) |
29130 | 72 |
qed |
73 |
||
74 |
lemma monofun_inv_Inl: "monofun (\<lambda>p. THE a. p = Inl a)" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
75 |
by (rule monofunI, erule sum_below_cases, simp_all) |
29130 | 76 |
|
77 |
lemma monofun_inv_Inr: "monofun (\<lambda>p. THE b. p = Inr b)" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
78 |
by (rule monofunI, erule sum_below_cases, simp_all) |
29130 | 79 |
|
80 |
lemma sum_chain_cases: |
|
81 |
assumes Y: "chain Y" |
|
82 |
assumes A: "\<And>A. \<lbrakk>chain A; Y = (\<lambda>i. Inl (A i))\<rbrakk> \<Longrightarrow> R" |
|
83 |
assumes B: "\<And>B. \<lbrakk>chain B; Y = (\<lambda>i. Inr (B i))\<rbrakk> \<Longrightarrow> R" |
|
84 |
shows "R" |
|
85 |
apply (cases "Y 0") |
|
86 |
apply (rule A) |
|
87 |
apply (rule ch2ch_monofun [OF monofun_inv_Inl Y]) |
|
88 |
apply (rule ext) |
|
89 |
apply (cut_tac j=i in chain_mono [OF Y le0], simp) |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
90 |
apply (erule Inl_belowE, simp) |
29130 | 91 |
apply (rule B) |
92 |
apply (rule ch2ch_monofun [OF monofun_inv_Inr Y]) |
|
93 |
apply (rule ext) |
|
94 |
apply (cut_tac j=i in chain_mono [OF Y le0], simp) |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
95 |
apply (erule Inr_belowE, simp) |
29130 | 96 |
done |
97 |
||
98 |
lemma is_lub_Inl: "range S <<| x \<Longrightarrow> range (\<lambda>i. Inl (S i)) <<| Inl x" |
|
99 |
apply (rule is_lubI) |
|
100 |
apply (rule ub_rangeI) |
|
101 |
apply (simp add: is_ub_lub) |
|
102 |
apply (frule ub_rangeD [where i=arbitrary]) |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
103 |
apply (erule Inl_belowE, simp) |
29130 | 104 |
apply (erule is_lub_lub) |
105 |
apply (rule ub_rangeI) |
|
106 |
apply (drule ub_rangeD, simp) |
|
107 |
done |
|
108 |
||
109 |
lemma is_lub_Inr: "range S <<| x \<Longrightarrow> range (\<lambda>i. Inr (S i)) <<| Inr x" |
|
110 |
apply (rule is_lubI) |
|
111 |
apply (rule ub_rangeI) |
|
112 |
apply (simp add: is_ub_lub) |
|
113 |
apply (frule ub_rangeD [where i=arbitrary]) |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
114 |
apply (erule Inr_belowE, simp) |
29130 | 115 |
apply (erule is_lub_lub) |
116 |
apply (rule ub_rangeI) |
|
117 |
apply (drule ub_rangeD, simp) |
|
118 |
done |
|
119 |
||
120 |
instance "+" :: (cpo, cpo) cpo |
|
121 |
apply intro_classes |
|
122 |
apply (erule sum_chain_cases, safe) |
|
123 |
apply (rule exI) |
|
124 |
apply (rule is_lub_Inl) |
|
125 |
apply (erule cpo_lubI) |
|
126 |
apply (rule exI) |
|
127 |
apply (rule is_lub_Inr) |
|
128 |
apply (erule cpo_lubI) |
|
129 |
done |
|
130 |
||
35900
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
huffman
parents:
31076
diff
changeset
|
131 |
subsection {* Continuity of \emph{Inl}, \emph{Inr}, and case function *} |
29130 | 132 |
|
133 |
lemma cont_Inl: "cont Inl" |
|
31041
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
134 |
by (intro contI is_lub_Inl cpo_lubI) |
29130 | 135 |
|
136 |
lemma cont_Inr: "cont Inr" |
|
31041
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
137 |
by (intro contI is_lub_Inr cpo_lubI) |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
138 |
|
37079
0cd15d8c90a0
remove cont2cont simproc; instead declare cont2cont rules as simp rules
huffman
parents:
35900
diff
changeset
|
139 |
lemmas cont2cont_Inl [simp, cont2cont] = cont_compose [OF cont_Inl] |
0cd15d8c90a0
remove cont2cont simproc; instead declare cont2cont rules as simp rules
huffman
parents:
35900
diff
changeset
|
140 |
lemmas cont2cont_Inr [simp, cont2cont] = cont_compose [OF cont_Inr] |
29130 | 141 |
|
142 |
lemmas ch2ch_Inl [simp] = ch2ch_cont [OF cont_Inl] |
|
143 |
lemmas ch2ch_Inr [simp] = ch2ch_cont [OF cont_Inr] |
|
144 |
||
145 |
lemmas lub_Inl = cont2contlubE [OF cont_Inl, symmetric] |
|
146 |
lemmas lub_Inr = cont2contlubE [OF cont_Inr, symmetric] |
|
147 |
||
148 |
lemma cont_sum_case1: |
|
149 |
assumes f: "\<And>a. cont (\<lambda>x. f x a)" |
|
150 |
assumes g: "\<And>b. cont (\<lambda>x. g x b)" |
|
151 |
shows "cont (\<lambda>x. case y of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)" |
|
152 |
by (induct y, simp add: f, simp add: g) |
|
153 |
||
154 |
lemma cont_sum_case2: "\<lbrakk>cont f; cont g\<rbrakk> \<Longrightarrow> cont (sum_case f g)" |
|
155 |
apply (rule contI) |
|
156 |
apply (erule sum_chain_cases) |
|
157 |
apply (simp add: cont2contlubE [OF cont_Inl, symmetric] contE) |
|
158 |
apply (simp add: cont2contlubE [OF cont_Inr, symmetric] contE) |
|
159 |
done |
|
160 |
||
31041
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
161 |
lemma cont2cont_sum_case: |
29130 | 162 |
assumes f1: "\<And>a. cont (\<lambda>x. f x a)" and f2: "\<And>x. cont (\<lambda>a. f x a)" |
163 |
assumes g1: "\<And>b. cont (\<lambda>x. g x b)" and g2: "\<And>x. cont (\<lambda>b. g x b)" |
|
164 |
assumes h: "cont (\<lambda>x. h x)" |
|
165 |
shows "cont (\<lambda>x. case h x of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)" |
|
31041
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
166 |
apply (rule cont_apply [OF h]) |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
167 |
apply (rule cont_sum_case2 [OF f2 g2]) |
29130 | 168 |
apply (rule cont_sum_case1 [OF f1 g1]) |
169 |
done |
|
170 |
||
37079
0cd15d8c90a0
remove cont2cont simproc; instead declare cont2cont rules as simp rules
huffman
parents:
35900
diff
changeset
|
171 |
lemma cont2cont_sum_case' [simp, cont2cont]: |
31041
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
172 |
assumes f: "cont (\<lambda>p. f (fst p) (snd p))" |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
173 |
assumes g: "cont (\<lambda>p. g (fst p) (snd p))" |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
174 |
assumes h: "cont (\<lambda>x. h x)" |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
175 |
shows "cont (\<lambda>x. case h x of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)" |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
176 |
proof - |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
177 |
note f1 = f [THEN cont_fst_snd_D1] |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
178 |
note f2 = f [THEN cont_fst_snd_D2] |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
179 |
note g1 = g [THEN cont_fst_snd_D1] |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
180 |
note g2 = g [THEN cont_fst_snd_D2] |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
181 |
show ?thesis |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
182 |
apply (rule cont_apply [OF h]) |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
183 |
apply (rule cont_sum_case2 [OF f2 g2]) |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
184 |
apply (rule cont_sum_case1 [OF f1 g1]) |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
185 |
done |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
186 |
qed |
85b4843d9939
replace cont2cont_apply with cont_apply; add new cont2cont lemmas
huffman
parents:
29534
diff
changeset
|
187 |
|
29130 | 188 |
subsection {* Compactness and chain-finiteness *} |
189 |
||
190 |
lemma compact_Inl: "compact a \<Longrightarrow> compact (Inl a)" |
|
191 |
apply (rule compactI2) |
|
192 |
apply (erule sum_chain_cases, safe) |
|
193 |
apply (simp add: lub_Inl) |
|
194 |
apply (erule (2) compactD2) |
|
195 |
apply (simp add: lub_Inr) |
|
196 |
done |
|
197 |
||
198 |
lemma compact_Inr: "compact a \<Longrightarrow> compact (Inr a)" |
|
199 |
apply (rule compactI2) |
|
200 |
apply (erule sum_chain_cases, safe) |
|
201 |
apply (simp add: lub_Inl) |
|
202 |
apply (simp add: lub_Inr) |
|
203 |
apply (erule (2) compactD2) |
|
204 |
done |
|
205 |
||
206 |
lemma compact_Inl_rev: "compact (Inl a) \<Longrightarrow> compact a" |
|
207 |
unfolding compact_def |
|
208 |
by (drule adm_subst [OF cont_Inl], simp) |
|
209 |
||
210 |
lemma compact_Inr_rev: "compact (Inr a) \<Longrightarrow> compact a" |
|
211 |
unfolding compact_def |
|
212 |
by (drule adm_subst [OF cont_Inr], simp) |
|
213 |
||
214 |
lemma compact_Inl_iff [simp]: "compact (Inl a) = compact a" |
|
215 |
by (safe elim!: compact_Inl compact_Inl_rev) |
|
216 |
||
217 |
lemma compact_Inr_iff [simp]: "compact (Inr a) = compact a" |
|
218 |
by (safe elim!: compact_Inr compact_Inr_rev) |
|
219 |
||
220 |
instance "+" :: (chfin, chfin) chfin |
|
221 |
apply intro_classes |
|
222 |
apply (erule compact_imp_max_in_chain) |
|
223 |
apply (case_tac "\<Squnion>i. Y i", simp_all) |
|
224 |
done |
|
225 |
||
226 |
instance "+" :: (finite_po, finite_po) finite_po .. |
|
227 |
||
228 |
instance "+" :: (discrete_cpo, discrete_cpo) discrete_cpo |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31041
diff
changeset
|
229 |
by intro_classes (simp add: below_sum_def split: sum.split) |
29130 | 230 |
|
231 |
subsection {* Sum type is a bifinite domain *} |
|
232 |
||
233 |
instantiation "+" :: (profinite, profinite) profinite |
|
234 |
begin |
|
235 |
||
236 |
definition |
|
237 |
approx_sum_def: "approx = |
|
238 |
(\<lambda>n. \<Lambda> x. case x of Inl a \<Rightarrow> Inl (approx n\<cdot>a) | Inr b \<Rightarrow> Inr (approx n\<cdot>b))" |
|
239 |
||
240 |
lemma approx_Inl [simp]: "approx n\<cdot>(Inl x) = Inl (approx n\<cdot>x)" |
|
241 |
unfolding approx_sum_def by simp |
|
242 |
||
243 |
lemma approx_Inr [simp]: "approx n\<cdot>(Inr x) = Inr (approx n\<cdot>x)" |
|
244 |
unfolding approx_sum_def by simp |
|
245 |
||
246 |
instance proof |
|
247 |
fix i :: nat and x :: "'a + 'b" |
|
248 |
show "chain (approx :: nat \<Rightarrow> 'a + 'b \<rightarrow> 'a + 'b)" |
|
249 |
unfolding approx_sum_def |
|
250 |
by (rule ch2ch_LAM, case_tac x, simp_all) |
|
251 |
show "(\<Squnion>i. approx i\<cdot>x) = x" |
|
252 |
by (induct x, simp_all add: lub_Inl lub_Inr) |
|
253 |
show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x" |
|
254 |
by (induct x, simp_all) |
|
255 |
have "{x::'a + 'b. approx i\<cdot>x = x} \<subseteq> |
|
256 |
{x::'a. approx i\<cdot>x = x} <+> {x::'b. approx i\<cdot>x = x}" |
|
257 |
by (rule subsetI, case_tac x, simp_all add: InlI InrI) |
|
258 |
thus "finite {x::'a + 'b. approx i\<cdot>x = x}" |
|
259 |
by (rule finite_subset, |
|
260 |
intro finite_Plus finite_fixes_approx) |
|
261 |
qed |
|
262 |
||
263 |
end |
|
264 |
||
265 |
end |