1 (* Title: HOLCF/Dsum.thy |
|
2 Author: Brian Huffman |
|
3 *) |
|
4 |
|
5 header {* The cpo of disjoint sums *} |
|
6 |
|
7 theory Dsum |
|
8 imports Bifinite |
|
9 begin |
|
10 |
|
11 subsection {* Ordering on type @{typ "'a + 'b"} *} |
|
12 |
|
13 instantiation "+" :: (sq_ord, sq_ord) sq_ord |
|
14 begin |
|
15 |
|
16 definition |
|
17 less_sum_def: "x \<sqsubseteq> y \<equiv> case x of |
|
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 |
|
24 lemma Inl_less_iff [simp]: "Inl x \<sqsubseteq> Inl y = x \<sqsubseteq> y" |
|
25 unfolding less_sum_def by simp |
|
26 |
|
27 lemma Inr_less_iff [simp]: "Inr x \<sqsubseteq> Inr y = x \<sqsubseteq> y" |
|
28 unfolding less_sum_def by simp |
|
29 |
|
30 lemma Inl_less_Inr [simp]: "\<not> Inl x \<sqsubseteq> Inr y" |
|
31 unfolding less_sum_def by simp |
|
32 |
|
33 lemma Inr_less_Inl [simp]: "\<not> Inr x \<sqsubseteq> Inl y" |
|
34 unfolding less_sum_def by simp |
|
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 |
|
42 lemma Inl_lessE: "\<lbrakk>Inl a \<sqsubseteq> x; \<And>b. \<lbrakk>x = Inl b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
|
43 by (cases x, simp_all) |
|
44 |
|
45 lemma Inr_lessE: "\<lbrakk>Inr a \<sqsubseteq> x; \<And>b. \<lbrakk>x = Inr b; a \<sqsubseteq> b\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" |
|
46 by (cases x, simp_all) |
|
47 |
|
48 lemmas sum_less_elims = Inl_lessE Inr_lessE |
|
49 |
|
50 lemma sum_less_cases: |
|
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" |
|
55 by (cases x, safe elim!: sum_less_elims, auto) |
|
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" |
|
67 by (induct x, auto elim!: sum_less_elims intro: antisym_less) |
|
68 next |
|
69 fix x y z :: "'a + 'b" |
|
70 assume "x \<sqsubseteq> y" and "y \<sqsubseteq> z" thus "x \<sqsubseteq> z" |
|
71 by (induct x, auto elim!: sum_less_elims intro: trans_less) |
|
72 qed |
|
73 |
|
74 lemma monofun_inv_Inl: "monofun (\<lambda>p. THE a. p = Inl a)" |
|
75 by (rule monofunI, erule sum_less_cases, simp_all) |
|
76 |
|
77 lemma monofun_inv_Inr: "monofun (\<lambda>p. THE b. p = Inr b)" |
|
78 by (rule monofunI, erule sum_less_cases, simp_all) |
|
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) |
|
90 apply (erule Inl_lessE, simp) |
|
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) |
|
95 apply (erule Inr_lessE, simp) |
|
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]) |
|
103 apply (erule Inl_lessE, simp) |
|
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]) |
|
114 apply (erule Inr_lessE, simp) |
|
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 |
|
131 subsection {* Continuity of @{term Inl}, @{term Inr}, @{term sum_case} *} |
|
132 |
|
133 lemma cont2cont_Inl [simp]: "cont f \<Longrightarrow> cont (\<lambda>x. Inl (f x))" |
|
134 by (fast intro: contI is_lub_Inl elim: contE) |
|
135 |
|
136 lemma cont2cont_Inr [simp]: "cont f \<Longrightarrow> cont (\<lambda>x. Inr (f x))" |
|
137 by (fast intro: contI is_lub_Inr elim: contE) |
|
138 |
|
139 lemma cont_Inl: "cont Inl" |
|
140 by (rule cont2cont_Inl [OF cont_id]) |
|
141 |
|
142 lemma cont_Inr: "cont Inr" |
|
143 by (rule cont2cont_Inr [OF cont_id]) |
|
144 |
|
145 lemmas ch2ch_Inl [simp] = ch2ch_cont [OF cont_Inl] |
|
146 lemmas ch2ch_Inr [simp] = ch2ch_cont [OF cont_Inr] |
|
147 |
|
148 lemmas lub_Inl = cont2contlubE [OF cont_Inl, symmetric] |
|
149 lemmas lub_Inr = cont2contlubE [OF cont_Inr, symmetric] |
|
150 |
|
151 lemma cont_sum_case1: |
|
152 assumes f: "\<And>a. cont (\<lambda>x. f x a)" |
|
153 assumes g: "\<And>b. cont (\<lambda>x. g x b)" |
|
154 shows "cont (\<lambda>x. case y of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)" |
|
155 by (induct y, simp add: f, simp add: g) |
|
156 |
|
157 lemma cont_sum_case2: "\<lbrakk>cont f; cont g\<rbrakk> \<Longrightarrow> cont (sum_case f g)" |
|
158 apply (rule contI) |
|
159 apply (erule sum_chain_cases) |
|
160 apply (simp add: cont2contlubE [OF cont_Inl, symmetric] contE) |
|
161 apply (simp add: cont2contlubE [OF cont_Inr, symmetric] contE) |
|
162 done |
|
163 |
|
164 lemma cont2cont_sum_case [simp]: |
|
165 assumes f1: "\<And>a. cont (\<lambda>x. f x a)" and f2: "\<And>x. cont (\<lambda>a. f x a)" |
|
166 assumes g1: "\<And>b. cont (\<lambda>x. g x b)" and g2: "\<And>x. cont (\<lambda>b. g x b)" |
|
167 assumes h: "cont (\<lambda>x. h x)" |
|
168 shows "cont (\<lambda>x. case h x of Inl a \<Rightarrow> f x a | Inr b \<Rightarrow> g x b)" |
|
169 apply (rule cont2cont_app2 [OF cont2cont_lambda _ h]) |
|
170 apply (rule cont_sum_case1 [OF f1 g1]) |
|
171 apply (rule cont_sum_case2 [OF f2 g2]) |
|
172 done |
|
173 |
|
174 subsection {* Compactness and chain-finiteness *} |
|
175 |
|
176 lemma compact_Inl: "compact a \<Longrightarrow> compact (Inl a)" |
|
177 apply (rule compactI2) |
|
178 apply (erule sum_chain_cases, safe) |
|
179 apply (simp add: lub_Inl) |
|
180 apply (erule (2) compactD2) |
|
181 apply (simp add: lub_Inr) |
|
182 done |
|
183 |
|
184 lemma compact_Inr: "compact a \<Longrightarrow> compact (Inr a)" |
|
185 apply (rule compactI2) |
|
186 apply (erule sum_chain_cases, safe) |
|
187 apply (simp add: lub_Inl) |
|
188 apply (simp add: lub_Inr) |
|
189 apply (erule (2) compactD2) |
|
190 done |
|
191 |
|
192 lemma compact_Inl_rev: "compact (Inl a) \<Longrightarrow> compact a" |
|
193 unfolding compact_def |
|
194 by (drule adm_subst [OF cont_Inl], simp) |
|
195 |
|
196 lemma compact_Inr_rev: "compact (Inr a) \<Longrightarrow> compact a" |
|
197 unfolding compact_def |
|
198 by (drule adm_subst [OF cont_Inr], simp) |
|
199 |
|
200 lemma compact_Inl_iff [simp]: "compact (Inl a) = compact a" |
|
201 by (safe elim!: compact_Inl compact_Inl_rev) |
|
202 |
|
203 lemma compact_Inr_iff [simp]: "compact (Inr a) = compact a" |
|
204 by (safe elim!: compact_Inr compact_Inr_rev) |
|
205 |
|
206 instance "+" :: (chfin, chfin) chfin |
|
207 apply intro_classes |
|
208 apply (erule compact_imp_max_in_chain) |
|
209 apply (case_tac "\<Squnion>i. Y i", simp_all) |
|
210 done |
|
211 |
|
212 instance "+" :: (finite_po, finite_po) finite_po .. |
|
213 |
|
214 instance "+" :: (discrete_cpo, discrete_cpo) discrete_cpo |
|
215 by intro_classes (simp add: less_sum_def split: sum.split) |
|
216 |
|
217 subsection {* Sum type is a bifinite domain *} |
|
218 |
|
219 instantiation "+" :: (profinite, profinite) profinite |
|
220 begin |
|
221 |
|
222 definition |
|
223 approx_sum_def: "approx = |
|
224 (\<lambda>n. \<Lambda> x. case x of Inl a \<Rightarrow> Inl (approx n\<cdot>a) | Inr b \<Rightarrow> Inr (approx n\<cdot>b))" |
|
225 |
|
226 lemma approx_Inl [simp]: "approx n\<cdot>(Inl x) = Inl (approx n\<cdot>x)" |
|
227 unfolding approx_sum_def by simp |
|
228 |
|
229 lemma approx_Inr [simp]: "approx n\<cdot>(Inr x) = Inr (approx n\<cdot>x)" |
|
230 unfolding approx_sum_def by simp |
|
231 |
|
232 instance proof |
|
233 fix i :: nat and x :: "'a + 'b" |
|
234 show "chain (approx :: nat \<Rightarrow> 'a + 'b \<rightarrow> 'a + 'b)" |
|
235 unfolding approx_sum_def |
|
236 by (rule ch2ch_LAM, case_tac x, simp_all) |
|
237 show "(\<Squnion>i. approx i\<cdot>x) = x" |
|
238 by (induct x, simp_all add: lub_Inl lub_Inr) |
|
239 show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x" |
|
240 by (induct x, simp_all) |
|
241 have "{x::'a + 'b. approx i\<cdot>x = x} \<subseteq> |
|
242 {x::'a. approx i\<cdot>x = x} <+> {x::'b. approx i\<cdot>x = x}" |
|
243 by (rule subsetI, case_tac x, simp_all add: InlI InrI) |
|
244 thus "finite {x::'a + 'b. approx i\<cdot>x = x}" |
|
245 by (rule finite_subset, |
|
246 intro finite_Plus finite_fixes_approx) |
|
247 qed |
|
248 |
|
249 end |
|
250 |
|
251 end |
|