author | haftmann |
Thu, 19 Jun 2025 17:15:40 +0200 | |
changeset 82734 | 89347c0cc6a3 |
parent 81583 | b6df83045178 |
permissions | -rw-r--r-- |
42151 | 1 |
(* Title: HOL/HOLCF/Domain.thy |
15741 | 2 |
Author: Brian Huffman |
3 |
*) |
|
4 |
||
62175 | 5 |
section \<open>Domain package\<close> |
15741 | 6 |
|
7 |
theory Domain |
|
81573 | 8 |
imports Representable Map_Functions Fixrec |
46950
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents:
46947
diff
changeset
|
9 |
keywords |
63432 | 10 |
"lazy" "unsafe" and |
69913 | 11 |
"domaindef" "domain" :: thy_defn and |
12 |
"domain_isomorphism" :: thy_decl |
|
15741 | 13 |
begin |
14 |
||
81573 | 15 |
subsection \<open>Continuous isomorphisms\<close> |
16 |
||
17 |
text \<open>A locale for continuous isomorphisms\<close> |
|
18 |
||
19 |
locale iso = |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
20 |
fixes abs :: "'a::pcpo \<rightarrow> 'b::pcpo" |
81573 | 21 |
fixes rep :: "'b \<rightarrow> 'a" |
22 |
assumes abs_iso [simp]: "rep\<cdot>(abs\<cdot>x) = x" |
|
23 |
assumes rep_iso [simp]: "abs\<cdot>(rep\<cdot>y) = y" |
|
24 |
begin |
|
25 |
||
26 |
lemma swap: "iso rep abs" |
|
27 |
by (rule iso.intro [OF rep_iso abs_iso]) |
|
28 |
||
29 |
lemma abs_below: "(abs\<cdot>x \<sqsubseteq> abs\<cdot>y) = (x \<sqsubseteq> y)" |
|
30 |
proof |
|
31 |
assume "abs\<cdot>x \<sqsubseteq> abs\<cdot>y" |
|
32 |
then have "rep\<cdot>(abs\<cdot>x) \<sqsubseteq> rep\<cdot>(abs\<cdot>y)" by (rule monofun_cfun_arg) |
|
33 |
then show "x \<sqsubseteq> y" by simp |
|
34 |
next |
|
35 |
assume "x \<sqsubseteq> y" |
|
36 |
then show "abs\<cdot>x \<sqsubseteq> abs\<cdot>y" by (rule monofun_cfun_arg) |
|
37 |
qed |
|
38 |
||
39 |
lemma rep_below: "(rep\<cdot>x \<sqsubseteq> rep\<cdot>y) = (x \<sqsubseteq> y)" |
|
40 |
by (rule iso.abs_below [OF swap]) |
|
41 |
||
42 |
lemma abs_eq: "(abs\<cdot>x = abs\<cdot>y) = (x = y)" |
|
43 |
by (simp add: po_eq_conv abs_below) |
|
44 |
||
45 |
lemma rep_eq: "(rep\<cdot>x = rep\<cdot>y) = (x = y)" |
|
46 |
by (rule iso.abs_eq [OF swap]) |
|
47 |
||
48 |
lemma abs_strict: "abs\<cdot>\<bottom> = \<bottom>" |
|
49 |
proof - |
|
50 |
have "\<bottom> \<sqsubseteq> rep\<cdot>\<bottom>" .. |
|
51 |
then have "abs\<cdot>\<bottom> \<sqsubseteq> abs\<cdot>(rep\<cdot>\<bottom>)" by (rule monofun_cfun_arg) |
|
52 |
then have "abs\<cdot>\<bottom> \<sqsubseteq> \<bottom>" by simp |
|
53 |
then show ?thesis by (rule bottomI) |
|
54 |
qed |
|
55 |
||
56 |
lemma rep_strict: "rep\<cdot>\<bottom> = \<bottom>" |
|
57 |
by (rule iso.abs_strict [OF swap]) |
|
58 |
||
59 |
lemma abs_defin': "abs\<cdot>x = \<bottom> \<Longrightarrow> x = \<bottom>" |
|
60 |
proof - |
|
61 |
have "x = rep\<cdot>(abs\<cdot>x)" by simp |
|
62 |
also assume "abs\<cdot>x = \<bottom>" |
|
63 |
also note rep_strict |
|
64 |
finally show "x = \<bottom>" . |
|
65 |
qed |
|
66 |
||
67 |
lemma rep_defin': "rep\<cdot>z = \<bottom> \<Longrightarrow> z = \<bottom>" |
|
68 |
by (rule iso.abs_defin' [OF swap]) |
|
69 |
||
70 |
lemma abs_defined: "z \<noteq> \<bottom> \<Longrightarrow> abs\<cdot>z \<noteq> \<bottom>" |
|
71 |
by (erule contrapos_nn, erule abs_defin') |
|
72 |
||
73 |
lemma rep_defined: "z \<noteq> \<bottom> \<Longrightarrow> rep\<cdot>z \<noteq> \<bottom>" |
|
74 |
by (rule iso.abs_defined [OF iso.swap]) (rule iso_axioms) |
|
75 |
||
76 |
lemma abs_bottom_iff: "(abs\<cdot>x = \<bottom>) = (x = \<bottom>)" |
|
77 |
by (auto elim: abs_defin' intro: abs_strict) |
|
78 |
||
79 |
lemma rep_bottom_iff: "(rep\<cdot>x = \<bottom>) = (x = \<bottom>)" |
|
80 |
by (rule iso.abs_bottom_iff [OF iso.swap]) (rule iso_axioms) |
|
81 |
||
82 |
lemma casedist_rule: "rep\<cdot>x = \<bottom> \<or> P \<Longrightarrow> x = \<bottom> \<or> P" |
|
83 |
by (simp add: rep_bottom_iff) |
|
84 |
||
85 |
lemma compact_abs_rev: "compact (abs\<cdot>x) \<Longrightarrow> compact x" |
|
86 |
proof (unfold compact_def) |
|
87 |
assume "adm (\<lambda>y. abs\<cdot>x \<notsqsubseteq> y)" |
|
88 |
with cont_Rep_cfun2 |
|
89 |
have "adm (\<lambda>y. abs\<cdot>x \<notsqsubseteq> abs\<cdot>y)" by (rule adm_subst) |
|
90 |
then show "adm (\<lambda>y. x \<notsqsubseteq> y)" using abs_below by simp |
|
91 |
qed |
|
92 |
||
93 |
lemma compact_rep_rev: "compact (rep\<cdot>x) \<Longrightarrow> compact x" |
|
94 |
by (rule iso.compact_abs_rev [OF iso.swap]) (rule iso_axioms) |
|
95 |
||
96 |
lemma compact_abs: "compact x \<Longrightarrow> compact (abs\<cdot>x)" |
|
97 |
by (rule compact_rep_rev) simp |
|
98 |
||
99 |
lemma compact_rep: "compact x \<Longrightarrow> compact (rep\<cdot>x)" |
|
100 |
by (rule iso.compact_abs [OF iso.swap]) (rule iso_axioms) |
|
101 |
||
102 |
lemma iso_swap: "(x = abs\<cdot>y) = (rep\<cdot>x = y)" |
|
103 |
proof |
|
104 |
assume "x = abs\<cdot>y" |
|
105 |
then have "rep\<cdot>x = rep\<cdot>(abs\<cdot>y)" by simp |
|
106 |
then show "rep\<cdot>x = y" by simp |
|
107 |
next |
|
108 |
assume "rep\<cdot>x = y" |
|
109 |
then have "abs\<cdot>(rep\<cdot>x) = abs\<cdot>y" by simp |
|
110 |
then show "x = abs\<cdot>y" by simp |
|
111 |
qed |
|
112 |
||
113 |
end |
|
114 |
||
81577 | 115 |
|
81573 | 116 |
subsection \<open>Proofs about take functions\<close> |
117 |
||
118 |
text \<open> |
|
119 |
This section contains lemmas that are used in a module that supports |
|
120 |
the domain isomorphism package; the module contains proofs related |
|
121 |
to take functions and the finiteness predicate. |
|
122 |
\<close> |
|
123 |
||
124 |
lemma deflation_abs_rep: |
|
125 |
fixes abs and rep and d |
|
126 |
assumes abs_iso: "\<And>x. rep\<cdot>(abs\<cdot>x) = x" |
|
127 |
assumes rep_iso: "\<And>y. abs\<cdot>(rep\<cdot>y) = y" |
|
128 |
shows "deflation d \<Longrightarrow> deflation (abs oo d oo rep)" |
|
129 |
by (rule ep_pair.deflation_e_d_p) (simp add: ep_pair.intro assms) |
|
130 |
||
131 |
lemma deflation_chain_min: |
|
132 |
assumes chain: "chain d" |
|
133 |
assumes defl: "\<And>n. deflation (d n)" |
|
134 |
shows "d m\<cdot>(d n\<cdot>x) = d (min m n)\<cdot>x" |
|
135 |
proof (rule linorder_le_cases) |
|
136 |
assume "m \<le> n" |
|
137 |
with chain have "d m \<sqsubseteq> d n" by (rule chain_mono) |
|
138 |
then have "d m\<cdot>(d n\<cdot>x) = d m\<cdot>x" |
|
139 |
by (rule deflation_below_comp1 [OF defl defl]) |
|
140 |
moreover from \<open>m \<le> n\<close> have "min m n = m" by simp |
|
141 |
ultimately show ?thesis by simp |
|
142 |
next |
|
143 |
assume "n \<le> m" |
|
144 |
with chain have "d n \<sqsubseteq> d m" by (rule chain_mono) |
|
145 |
then have "d m\<cdot>(d n\<cdot>x) = d n\<cdot>x" |
|
146 |
by (rule deflation_below_comp2 [OF defl defl]) |
|
147 |
moreover from \<open>n \<le> m\<close> have "min m n = n" by simp |
|
148 |
ultimately show ?thesis by simp |
|
149 |
qed |
|
150 |
||
151 |
lemma lub_ID_take_lemma: |
|
152 |
assumes "chain t" and "(\<Squnion>n. t n) = ID" |
|
153 |
assumes "\<And>n. t n\<cdot>x = t n\<cdot>y" shows "x = y" |
|
154 |
proof - |
|
155 |
have "(\<Squnion>n. t n\<cdot>x) = (\<Squnion>n. t n\<cdot>y)" |
|
156 |
using assms(3) by simp |
|
157 |
then have "(\<Squnion>n. t n)\<cdot>x = (\<Squnion>n. t n)\<cdot>y" |
|
158 |
using assms(1) by (simp add: lub_distribs) |
|
159 |
then show "x = y" |
|
160 |
using assms(2) by simp |
|
161 |
qed |
|
162 |
||
163 |
lemma lub_ID_reach: |
|
164 |
assumes "chain t" and "(\<Squnion>n. t n) = ID" |
|
165 |
shows "(\<Squnion>n. t n\<cdot>x) = x" |
|
166 |
using assms by (simp add: lub_distribs) |
|
167 |
||
168 |
lemma lub_ID_take_induct: |
|
169 |
assumes "chain t" and "(\<Squnion>n. t n) = ID" |
|
170 |
assumes "adm P" and "\<And>n. P (t n\<cdot>x)" shows "P x" |
|
171 |
proof - |
|
172 |
from \<open>chain t\<close> have "chain (\<lambda>n. t n\<cdot>x)" by simp |
|
173 |
from \<open>adm P\<close> this \<open>\<And>n. P (t n\<cdot>x)\<close> have "P (\<Squnion>n. t n\<cdot>x)" by (rule admD) |
|
174 |
with \<open>chain t\<close> \<open>(\<Squnion>n. t n) = ID\<close> show "P x" by (simp add: lub_distribs) |
|
175 |
qed |
|
176 |
||
81577 | 177 |
|
81573 | 178 |
subsection \<open>Finiteness\<close> |
179 |
||
180 |
text \<open> |
|
181 |
Let a ``decisive'' function be a deflation that maps every input to |
|
182 |
either itself or bottom. Then if a domain's take functions are all |
|
183 |
decisive, then all values in the domain are finite. |
|
184 |
\<close> |
|
185 |
||
186 |
definition |
|
187 |
decisive :: "('a::pcpo \<rightarrow> 'a) \<Rightarrow> bool" |
|
188 |
where |
|
189 |
"decisive d \<longleftrightarrow> (\<forall>x. d\<cdot>x = x \<or> d\<cdot>x = \<bottom>)" |
|
190 |
||
191 |
lemma decisiveI: "(\<And>x. d\<cdot>x = x \<or> d\<cdot>x = \<bottom>) \<Longrightarrow> decisive d" |
|
192 |
unfolding decisive_def by simp |
|
193 |
||
194 |
lemma decisive_cases: |
|
195 |
assumes "decisive d" obtains "d\<cdot>x = x" | "d\<cdot>x = \<bottom>" |
|
196 |
using assms unfolding decisive_def by auto |
|
197 |
||
198 |
lemma decisive_bottom: "decisive \<bottom>" |
|
199 |
unfolding decisive_def by simp |
|
200 |
||
201 |
lemma decisive_ID: "decisive ID" |
|
202 |
unfolding decisive_def by simp |
|
203 |
||
204 |
lemma decisive_ssum_map: |
|
205 |
assumes f: "decisive f" |
|
206 |
assumes g: "decisive g" |
|
207 |
shows "decisive (ssum_map\<cdot>f\<cdot>g)" |
|
208 |
apply (rule decisiveI) |
|
209 |
subgoal for s |
|
210 |
apply (cases s, simp_all) |
|
211 |
apply (rule_tac x=x in decisive_cases [OF f], simp_all) |
|
212 |
apply (rule_tac x=y in decisive_cases [OF g], simp_all) |
|
213 |
done |
|
214 |
done |
|
215 |
||
216 |
lemma decisive_sprod_map: |
|
217 |
assumes f: "decisive f" |
|
218 |
assumes g: "decisive g" |
|
219 |
shows "decisive (sprod_map\<cdot>f\<cdot>g)" |
|
220 |
apply (rule decisiveI) |
|
221 |
subgoal for s |
|
222 |
apply (cases s, simp) |
|
223 |
subgoal for x y |
|
224 |
apply (rule decisive_cases [OF f, where x = x], simp_all) |
|
225 |
apply (rule decisive_cases [OF g, where x = y], simp_all) |
|
226 |
done |
|
227 |
done |
|
228 |
done |
|
229 |
||
230 |
lemma decisive_abs_rep: |
|
231 |
fixes abs rep |
|
232 |
assumes iso: "iso abs rep" |
|
233 |
assumes d: "decisive d" |
|
234 |
shows "decisive (abs oo d oo rep)" |
|
235 |
apply (rule decisiveI) |
|
236 |
subgoal for s |
|
237 |
apply (rule decisive_cases [OF d, where x="rep\<cdot>s"]) |
|
238 |
apply (simp add: iso.rep_iso [OF iso]) |
|
239 |
apply (simp add: iso.abs_strict [OF iso]) |
|
240 |
done |
|
241 |
done |
|
242 |
||
243 |
lemma lub_ID_finite: |
|
244 |
assumes chain: "chain d" |
|
245 |
assumes lub: "(\<Squnion>n. d n) = ID" |
|
246 |
assumes decisive: "\<And>n. decisive (d n)" |
|
247 |
shows "\<exists>n. d n\<cdot>x = x" |
|
248 |
proof - |
|
249 |
have 1: "chain (\<lambda>n. d n\<cdot>x)" using chain by simp |
|
250 |
have 2: "(\<Squnion>n. d n\<cdot>x) = x" using chain lub by (rule lub_ID_reach) |
|
251 |
have "\<forall>n. d n\<cdot>x = x \<or> d n\<cdot>x = \<bottom>" |
|
252 |
using decisive unfolding decisive_def by simp |
|
253 |
hence "range (\<lambda>n. d n\<cdot>x) \<subseteq> {x, \<bottom>}" |
|
254 |
by auto |
|
255 |
hence "finite (range (\<lambda>n. d n\<cdot>x))" |
|
256 |
by (rule finite_subset, simp) |
|
257 |
with 1 have "finite_chain (\<lambda>n. d n\<cdot>x)" |
|
258 |
by (rule finite_range_imp_finch) |
|
259 |
then have "\<exists>n. (\<Squnion>n. d n\<cdot>x) = d n\<cdot>x" |
|
260 |
unfolding finite_chain_def by (auto simp add: maxinch_is_thelub) |
|
261 |
with 2 show "\<exists>n. d n\<cdot>x = x" by (auto elim: sym) |
|
262 |
qed |
|
263 |
||
264 |
lemma lub_ID_finite_take_induct: |
|
265 |
assumes "chain d" and "(\<Squnion>n. d n) = ID" and "\<And>n. decisive (d n)" |
|
266 |
shows "(\<And>n. P (d n\<cdot>x)) \<Longrightarrow> P x" |
|
267 |
using lub_ID_finite [OF assms] by metis |
|
268 |
||
81577 | 269 |
|
81573 | 270 |
subsection \<open>Proofs about constructor functions\<close> |
271 |
||
272 |
text \<open>Lemmas for proving nchotomy rule:\<close> |
|
273 |
||
274 |
lemma ex_one_bottom_iff: |
|
275 |
"(\<exists>x. P x \<and> x \<noteq> \<bottom>) = P ONE" |
|
276 |
by simp |
|
277 |
||
278 |
lemma ex_up_bottom_iff: |
|
279 |
"(\<exists>x. P x \<and> x \<noteq> \<bottom>) = (\<exists>x. P (up\<cdot>x))" |
|
280 |
by (safe, case_tac x, auto) |
|
281 |
||
282 |
lemma ex_sprod_bottom_iff: |
|
283 |
"(\<exists>y. P y \<and> y \<noteq> \<bottom>) = |
|
284 |
(\<exists>x y. (P (:x, y:) \<and> x \<noteq> \<bottom>) \<and> y \<noteq> \<bottom>)" |
|
285 |
by (safe, case_tac y, auto) |
|
286 |
||
287 |
lemma ex_sprod_up_bottom_iff: |
|
288 |
"(\<exists>y. P y \<and> y \<noteq> \<bottom>) = |
|
289 |
(\<exists>x y. P (:up\<cdot>x, y:) \<and> y \<noteq> \<bottom>)" |
|
290 |
by (safe, case_tac y, simp, case_tac x, auto) |
|
291 |
||
292 |
lemma ex_ssum_bottom_iff: |
|
293 |
"(\<exists>x. P x \<and> x \<noteq> \<bottom>) = |
|
294 |
((\<exists>x. P (sinl\<cdot>x) \<and> x \<noteq> \<bottom>) \<or> |
|
295 |
(\<exists>x. P (sinr\<cdot>x) \<and> x \<noteq> \<bottom>))" |
|
296 |
by (safe, case_tac x, auto) |
|
297 |
||
298 |
lemma exh_start: "p = \<bottom> \<or> (\<exists>x. p = x \<and> x \<noteq> \<bottom>)" |
|
299 |
by auto |
|
300 |
||
301 |
lemmas ex_bottom_iffs = |
|
302 |
ex_ssum_bottom_iff |
|
303 |
ex_sprod_up_bottom_iff |
|
304 |
ex_sprod_bottom_iff |
|
305 |
ex_up_bottom_iff |
|
306 |
ex_one_bottom_iff |
|
307 |
||
308 |
text \<open>Rules for turning nchotomy into exhaust:\<close> |
|
309 |
||
310 |
lemma exh_casedist0: "\<lbrakk>R; R \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" (* like make_elim *) |
|
311 |
by auto |
|
312 |
||
313 |
lemma exh_casedist1: "((P \<or> Q \<Longrightarrow> R) \<Longrightarrow> S) \<equiv> (\<lbrakk>P \<Longrightarrow> R; Q \<Longrightarrow> R\<rbrakk> \<Longrightarrow> S)" |
|
314 |
by rule auto |
|
315 |
||
316 |
lemma exh_casedist2: "(\<exists>x. P x \<Longrightarrow> Q) \<equiv> (\<And>x. P x \<Longrightarrow> Q)" |
|
317 |
by rule auto |
|
318 |
||
319 |
lemma exh_casedist3: "(P \<and> Q \<Longrightarrow> R) \<equiv> (P \<Longrightarrow> Q \<Longrightarrow> R)" |
|
320 |
by rule auto |
|
321 |
||
322 |
lemmas exh_casedists = exh_casedist1 exh_casedist2 exh_casedist3 |
|
323 |
||
324 |
text \<open>Rules for proving constructor properties\<close> |
|
325 |
||
326 |
lemmas con_strict_rules = |
|
327 |
sinl_strict sinr_strict spair_strict1 spair_strict2 |
|
328 |
||
329 |
lemmas con_bottom_iff_rules = |
|
330 |
sinl_bottom_iff sinr_bottom_iff spair_bottom_iff up_defined ONE_defined |
|
331 |
||
332 |
lemmas con_below_iff_rules = |
|
333 |
sinl_below sinr_below sinl_below_sinr sinr_below_sinl con_bottom_iff_rules |
|
334 |
||
335 |
lemmas con_eq_iff_rules = |
|
336 |
sinl_eq sinr_eq sinl_eq_sinr sinr_eq_sinl con_bottom_iff_rules |
|
337 |
||
338 |
lemmas sel_strict_rules = |
|
339 |
cfcomp2 sscase1 sfst_strict ssnd_strict fup1 |
|
340 |
||
341 |
lemma sel_app_extra_rules: |
|
342 |
"sscase\<cdot>ID\<cdot>\<bottom>\<cdot>(sinr\<cdot>x) = \<bottom>" |
|
343 |
"sscase\<cdot>ID\<cdot>\<bottom>\<cdot>(sinl\<cdot>x) = x" |
|
344 |
"sscase\<cdot>\<bottom>\<cdot>ID\<cdot>(sinl\<cdot>x) = \<bottom>" |
|
345 |
"sscase\<cdot>\<bottom>\<cdot>ID\<cdot>(sinr\<cdot>x) = x" |
|
346 |
"fup\<cdot>ID\<cdot>(up\<cdot>x) = x" |
|
347 |
by (cases "x = \<bottom>", simp, simp)+ |
|
348 |
||
349 |
lemmas sel_app_rules = |
|
350 |
sel_strict_rules sel_app_extra_rules |
|
351 |
ssnd_spair sfst_spair up_defined spair_defined |
|
352 |
||
353 |
lemmas sel_bottom_iff_rules = |
|
354 |
cfcomp2 sfst_bottom_iff ssnd_bottom_iff |
|
355 |
||
356 |
lemmas take_con_rules = |
|
357 |
ssum_map_sinl' ssum_map_sinr' sprod_map_spair' u_map_up |
|
358 |
deflation_strict deflation_ID ID1 cfcomp2 |
|
359 |
||
81577 | 360 |
|
81573 | 361 |
subsection \<open>ML setup\<close> |
362 |
||
363 |
named_theorems domain_deflation "theorems like deflation a ==> deflation (foo_map$a)" |
|
364 |
and domain_map_ID "theorems like foo_map$ID = ID" |
|
365 |
||
366 |
ML_file \<open>Tools/Domain/domain_take_proofs.ML\<close> |
|
367 |
ML_file \<open>Tools/cont_consts.ML\<close> |
|
368 |
ML_file \<open>Tools/cont_proc.ML\<close> |
|
369 |
simproc_setup cont ("cont f") = \<open>K ContProc.cont_proc\<close> |
|
370 |
||
371 |
ML_file \<open>Tools/Domain/domain_constructors.ML\<close> |
|
372 |
ML_file \<open>Tools/Domain/domain_induction.ML\<close> |
|
373 |
||
40504 | 374 |
|
62175 | 375 |
subsection \<open>Representations of types\<close> |
40504 | 376 |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
377 |
lemma emb_prj: "emb\<cdot>((prj\<cdot>x)::'a::domain) = cast\<cdot>DEFL('a)\<cdot>x" |
40504 | 378 |
by (simp add: cast_DEFL) |
379 |
||
380 |
lemma emb_prj_emb: |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
381 |
fixes x :: "'a::domain" |
40504 | 382 |
assumes "DEFL('a) \<sqsubseteq> DEFL('b)" |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
383 |
shows "emb\<cdot>(prj\<cdot>(emb\<cdot>x) :: 'b::domain) = emb\<cdot>x" |
40504 | 384 |
unfolding emb_prj |
385 |
apply (rule cast.belowD) |
|
386 |
apply (rule monofun_cfun_arg [OF assms]) |
|
387 |
apply (simp add: cast_DEFL) |
|
388 |
done |
|
389 |
||
390 |
lemma prj_emb_prj: |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
391 |
assumes "DEFL('a::domain) \<sqsubseteq> DEFL('b::domain)" |
40504 | 392 |
shows "prj\<cdot>(emb\<cdot>(prj\<cdot>x :: 'b)) = (prj\<cdot>x :: 'a)" |
393 |
apply (rule emb_eq_iff [THEN iffD1]) |
|
394 |
apply (simp only: emb_prj) |
|
395 |
apply (rule deflation_below_comp1) |
|
396 |
apply (rule deflation_cast) |
|
397 |
apply (rule deflation_cast) |
|
398 |
apply (rule monofun_cfun_arg [OF assms]) |
|
399 |
done |
|
400 |
||
62175 | 401 |
text \<open>Isomorphism lemmas used internally by the domain package:\<close> |
40504 | 402 |
|
403 |
lemma domain_abs_iso: |
|
404 |
fixes abs and rep |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
405 |
assumes DEFL: "DEFL('b::domain) = DEFL('a::domain)" |
40504 | 406 |
assumes abs_def: "(abs :: 'a \<rightarrow> 'b) \<equiv> prj oo emb" |
407 |
assumes rep_def: "(rep :: 'b \<rightarrow> 'a) \<equiv> prj oo emb" |
|
408 |
shows "rep\<cdot>(abs\<cdot>x) = x" |
|
409 |
unfolding abs_def rep_def |
|
410 |
by (simp add: emb_prj_emb DEFL) |
|
411 |
||
412 |
lemma domain_rep_iso: |
|
413 |
fixes abs and rep |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
414 |
assumes DEFL: "DEFL('b::domain) = DEFL('a::domain)" |
40504 | 415 |
assumes abs_def: "(abs :: 'a \<rightarrow> 'b) \<equiv> prj oo emb" |
416 |
assumes rep_def: "(rep :: 'b \<rightarrow> 'a) \<equiv> prj oo emb" |
|
417 |
shows "abs\<cdot>(rep\<cdot>x) = x" |
|
418 |
unfolding abs_def rep_def |
|
419 |
by (simp add: emb_prj_emb DEFL) |
|
420 |
||
81577 | 421 |
|
62175 | 422 |
subsection \<open>Deflations as sets\<close> |
40504 | 423 |
|
41287
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
huffman
parents:
41285
diff
changeset
|
424 |
definition defl_set :: "'a::bifinite defl \<Rightarrow> 'a set" |
40504 | 425 |
where "defl_set A = {x. cast\<cdot>A\<cdot>x = x}" |
426 |
||
427 |
lemma adm_defl_set: "adm (\<lambda>x. x \<in> defl_set A)" |
|
428 |
unfolding defl_set_def by simp |
|
429 |
||
430 |
lemma defl_set_bottom: "\<bottom> \<in> defl_set A" |
|
431 |
unfolding defl_set_def by simp |
|
432 |
||
433 |
lemma defl_set_cast [simp]: "cast\<cdot>A\<cdot>x \<in> defl_set A" |
|
434 |
unfolding defl_set_def by simp |
|
435 |
||
436 |
lemma defl_set_subset_iff: "defl_set A \<subseteq> defl_set B \<longleftrightarrow> A \<sqsubseteq> B" |
|
437 |
apply (simp add: defl_set_def subset_eq cast_below_cast [symmetric]) |
|
438 |
apply (auto simp add: cast.belowI cast.belowD) |
|
439 |
done |
|
440 |
||
81577 | 441 |
|
62175 | 442 |
subsection \<open>Proving a subtype is representable\<close> |
40504 | 443 |
|
62175 | 444 |
text \<open>Temporarily relax type constraints.\<close> |
40504 | 445 |
|
62175 | 446 |
setup \<open> |
40504 | 447 |
fold Sign.add_const_constraint |
69597 | 448 |
[ (\<^const_name>\<open>defl\<close>, SOME \<^typ>\<open>'a::pcpo itself \<Rightarrow> udom defl\<close>) |
449 |
, (\<^const_name>\<open>emb\<close>, SOME \<^typ>\<open>'a::pcpo \<rightarrow> udom\<close>) |
|
450 |
, (\<^const_name>\<open>prj\<close>, SOME \<^typ>\<open>udom \<rightarrow> 'a::pcpo\<close>) |
|
451 |
, (\<^const_name>\<open>liftdefl\<close>, SOME \<^typ>\<open>'a::pcpo itself \<Rightarrow> udom u defl\<close>) |
|
452 |
, (\<^const_name>\<open>liftemb\<close>, SOME \<^typ>\<open>'a::pcpo u \<rightarrow> udom u\<close>) |
|
453 |
, (\<^const_name>\<open>liftprj\<close>, SOME \<^typ>\<open>udom u \<rightarrow> 'a::pcpo u\<close>) ] |
|
62175 | 454 |
\<close> |
40504 | 455 |
|
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
456 |
lemma typedef_domain_class: |
40504 | 457 |
fixes Rep :: "'a::pcpo \<Rightarrow> udom" |
458 |
fixes Abs :: "udom \<Rightarrow> 'a::pcpo" |
|
41287
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
huffman
parents:
41285
diff
changeset
|
459 |
fixes t :: "udom defl" |
40504 | 460 |
assumes type: "type_definition Rep Abs (defl_set t)" |
67399 | 461 |
assumes below: "(\<sqsubseteq>) \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y" |
40504 | 462 |
assumes emb: "emb \<equiv> (\<Lambda> x. Rep x)" |
463 |
assumes prj: "prj \<equiv> (\<Lambda> x. Abs (cast\<cdot>t\<cdot>x))" |
|
464 |
assumes defl: "defl \<equiv> (\<lambda> a::'a itself. t)" |
|
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
465 |
assumes liftemb: "(liftemb :: 'a u \<rightarrow> udom u) \<equiv> u_map\<cdot>emb" |
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
466 |
assumes liftprj: "(liftprj :: udom u \<rightarrow> 'a u) \<equiv> u_map\<cdot>prj" |
41436 | 467 |
assumes liftdefl: "(liftdefl :: 'a itself \<Rightarrow> _) \<equiv> (\<lambda>t. liftdefl_of\<cdot>DEFL('a))" |
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
468 |
shows "OFCLASS('a, domain_class)" |
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
469 |
proof |
40504 | 470 |
have emb_beta: "\<And>x. emb\<cdot>x = Rep x" |
471 |
unfolding emb |
|
472 |
apply (rule beta_cfun) |
|
40834
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
huffman
parents:
40830
diff
changeset
|
473 |
apply (rule typedef_cont_Rep [OF type below adm_defl_set cont_id]) |
40504 | 474 |
done |
475 |
have prj_beta: "\<And>y. prj\<cdot>y = Abs (cast\<cdot>t\<cdot>y)" |
|
476 |
unfolding prj |
|
477 |
apply (rule beta_cfun) |
|
478 |
apply (rule typedef_cont_Abs [OF type below adm_defl_set]) |
|
479 |
apply simp_all |
|
480 |
done |
|
481 |
have prj_emb: "\<And>x::'a. prj\<cdot>(emb\<cdot>x) = x" |
|
482 |
using type_definition.Rep [OF type] |
|
483 |
unfolding prj_beta emb_beta defl_set_def |
|
484 |
by (simp add: type_definition.Rep_inverse [OF type]) |
|
485 |
have emb_prj: "\<And>y. emb\<cdot>(prj\<cdot>y :: 'a) = cast\<cdot>t\<cdot>y" |
|
486 |
unfolding prj_beta emb_beta |
|
487 |
by (simp add: type_definition.Abs_inverse [OF type]) |
|
488 |
show "ep_pair (emb :: 'a \<rightarrow> udom) prj" |
|
61169 | 489 |
apply standard |
40504 | 490 |
apply (simp add: prj_emb) |
491 |
apply (simp add: emb_prj cast.below) |
|
492 |
done |
|
493 |
show "cast\<cdot>DEFL('a) = emb oo (prj :: udom \<rightarrow> 'a)" |
|
494 |
by (rule cfun_eqI, simp add: defl emb_prj) |
|
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
495 |
qed (simp_all only: liftemb liftprj liftdefl) |
40504 | 496 |
|
497 |
lemma typedef_DEFL: |
|
498 |
assumes "defl \<equiv> (\<lambda>a::'a::pcpo itself. t)" |
|
499 |
shows "DEFL('a::pcpo) = t" |
|
500 |
unfolding assms .. |
|
501 |
||
62175 | 502 |
text \<open>Restore original typing constraints.\<close> |
40504 | 503 |
|
62175 | 504 |
setup \<open> |
40504 | 505 |
fold Sign.add_const_constraint |
69597 | 506 |
[(\<^const_name>\<open>defl\<close>, SOME \<^typ>\<open>'a::domain itself \<Rightarrow> udom defl\<close>), |
507 |
(\<^const_name>\<open>emb\<close>, SOME \<^typ>\<open>'a::domain \<rightarrow> udom\<close>), |
|
508 |
(\<^const_name>\<open>prj\<close>, SOME \<^typ>\<open>udom \<rightarrow> 'a::domain\<close>), |
|
509 |
(\<^const_name>\<open>liftdefl\<close>, SOME \<^typ>\<open>'a::predomain itself \<Rightarrow> udom u defl\<close>), |
|
510 |
(\<^const_name>\<open>liftemb\<close>, SOME \<^typ>\<open>'a::predomain u \<rightarrow> udom u\<close>), |
|
511 |
(\<^const_name>\<open>liftprj\<close>, SOME \<^typ>\<open>udom u \<rightarrow> 'a::predomain u\<close>)] |
|
62175 | 512 |
\<close> |
40504 | 513 |
|
69605 | 514 |
ML_file \<open>Tools/domaindef.ML\<close> |
40504 | 515 |
|
81577 | 516 |
|
62175 | 517 |
subsection \<open>Isomorphic deflations\<close> |
40504 | 518 |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
519 |
definition isodefl :: "('a::domain \<rightarrow> 'a) \<Rightarrow> udom defl \<Rightarrow> bool" |
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
520 |
where "isodefl d t \<longleftrightarrow> cast\<cdot>t = emb oo d oo prj" |
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
521 |
|
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
522 |
definition isodefl' :: "('a::predomain \<rightarrow> 'a) \<Rightarrow> udom u defl \<Rightarrow> bool" |
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
523 |
where "isodefl' d t \<longleftrightarrow> cast\<cdot>t = liftemb oo u_map\<cdot>d oo liftprj" |
40504 | 524 |
|
525 |
lemma isodeflI: "(\<And>x. cast\<cdot>t\<cdot>x = emb\<cdot>(d\<cdot>(prj\<cdot>x))) \<Longrightarrow> isodefl d t" |
|
526 |
unfolding isodefl_def by (simp add: cfun_eqI) |
|
527 |
||
528 |
lemma cast_isodefl: "isodefl d t \<Longrightarrow> cast\<cdot>t = (\<Lambda> x. emb\<cdot>(d\<cdot>(prj\<cdot>x)))" |
|
529 |
unfolding isodefl_def by (simp add: cfun_eqI) |
|
530 |
||
531 |
lemma isodefl_strict: "isodefl d t \<Longrightarrow> d\<cdot>\<bottom> = \<bottom>" |
|
532 |
unfolding isodefl_def |
|
533 |
by (drule cfun_fun_cong [where x="\<bottom>"], simp) |
|
534 |
||
535 |
lemma isodefl_imp_deflation: |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
536 |
fixes d :: "'a::domain \<rightarrow> 'a" |
40504 | 537 |
assumes "isodefl d t" shows "deflation d" |
538 |
proof |
|
539 |
note assms [unfolded isodefl_def, simp] |
|
540 |
fix x :: 'a |
|
541 |
show "d\<cdot>(d\<cdot>x) = d\<cdot>x" |
|
542 |
using cast.idem [of t "emb\<cdot>x"] by simp |
|
543 |
show "d\<cdot>x \<sqsubseteq> x" |
|
544 |
using cast.below [of t "emb\<cdot>x"] by simp |
|
545 |
qed |
|
546 |
||
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
547 |
lemma isodefl_ID_DEFL: "isodefl (ID :: 'a \<rightarrow> 'a) DEFL('a::domain)" |
40504 | 548 |
unfolding isodefl_def by (simp add: cast_DEFL) |
549 |
||
550 |
lemma isodefl_LIFTDEFL: |
|
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
551 |
"isodefl' (ID :: 'a \<rightarrow> 'a) LIFTDEFL('a::predomain)" |
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
552 |
unfolding isodefl'_def by (simp add: cast_liftdefl u_map_ID) |
40504 | 553 |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
554 |
lemma isodefl_DEFL_imp_ID: "isodefl (d :: 'a \<rightarrow> 'a) DEFL('a::domain) \<Longrightarrow> d = ID" |
40504 | 555 |
unfolding isodefl_def |
556 |
apply (simp add: cast_DEFL) |
|
557 |
apply (simp add: cfun_eq_iff) |
|
558 |
apply (rule allI) |
|
559 |
apply (drule_tac x="emb\<cdot>x" in spec) |
|
560 |
apply simp |
|
561 |
done |
|
562 |
||
563 |
lemma isodefl_bottom: "isodefl \<bottom> \<bottom>" |
|
564 |
unfolding isodefl_def by (simp add: cfun_eq_iff) |
|
565 |
||
566 |
lemma adm_isodefl: |
|
567 |
"cont f \<Longrightarrow> cont g \<Longrightarrow> adm (\<lambda>x. isodefl (f x) (g x))" |
|
568 |
unfolding isodefl_def by simp |
|
569 |
||
570 |
lemma isodefl_lub: |
|
571 |
assumes "chain d" and "chain t" |
|
572 |
assumes "\<And>i. isodefl (d i) (t i)" |
|
573 |
shows "isodefl (\<Squnion>i. d i) (\<Squnion>i. t i)" |
|
41529 | 574 |
using assms unfolding isodefl_def |
40504 | 575 |
by (simp add: contlub_cfun_arg contlub_cfun_fun) |
576 |
||
577 |
lemma isodefl_fix: |
|
578 |
assumes "\<And>d t. isodefl d t \<Longrightarrow> isodefl (f\<cdot>d) (g\<cdot>t)" |
|
579 |
shows "isodefl (fix\<cdot>f) (fix\<cdot>g)" |
|
580 |
unfolding fix_def2 |
|
581 |
apply (rule isodefl_lub, simp, simp) |
|
582 |
apply (induct_tac i) |
|
583 |
apply (simp add: isodefl_bottom) |
|
584 |
apply (simp add: assms) |
|
585 |
done |
|
586 |
||
587 |
lemma isodefl_abs_rep: |
|
588 |
fixes abs and rep and d |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
589 |
assumes DEFL: "DEFL('b::domain) = DEFL('a::domain)" |
40504 | 590 |
assumes abs_def: "(abs :: 'a \<rightarrow> 'b) \<equiv> prj oo emb" |
591 |
assumes rep_def: "(rep :: 'b \<rightarrow> 'a) \<equiv> prj oo emb" |
|
592 |
shows "isodefl d t \<Longrightarrow> isodefl (abs oo d oo rep) t" |
|
593 |
unfolding isodefl_def |
|
594 |
by (simp add: cfun_eq_iff assms prj_emb_prj emb_prj_emb) |
|
595 |
||
41436 | 596 |
lemma isodefl'_liftdefl_of: "isodefl d t \<Longrightarrow> isodefl' d (liftdefl_of\<cdot>t)" |
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
597 |
unfolding isodefl_def isodefl'_def |
41436 | 598 |
by (simp add: cast_liftdefl_of u_map_oo liftemb_eq liftprj_eq) |
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
599 |
|
40592
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
600 |
lemma isodefl_sfun: |
40504 | 601 |
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow> |
40592
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
602 |
isodefl (sfun_map\<cdot>d1\<cdot>d2) (sfun_defl\<cdot>t1\<cdot>t2)" |
40504 | 603 |
apply (rule isodeflI) |
40592
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
604 |
apply (simp add: cast_sfun_defl cast_isodefl) |
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
605 |
apply (simp add: emb_sfun_def prj_sfun_def) |
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
606 |
apply (simp add: sfun_map_map isodefl_strict) |
40504 | 607 |
done |
608 |
||
609 |
lemma isodefl_ssum: |
|
610 |
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow> |
|
611 |
isodefl (ssum_map\<cdot>d1\<cdot>d2) (ssum_defl\<cdot>t1\<cdot>t2)" |
|
612 |
apply (rule isodeflI) |
|
613 |
apply (simp add: cast_ssum_defl cast_isodefl) |
|
614 |
apply (simp add: emb_ssum_def prj_ssum_def) |
|
615 |
apply (simp add: ssum_map_map isodefl_strict) |
|
616 |
done |
|
617 |
||
618 |
lemma isodefl_sprod: |
|
619 |
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow> |
|
620 |
isodefl (sprod_map\<cdot>d1\<cdot>d2) (sprod_defl\<cdot>t1\<cdot>t2)" |
|
621 |
apply (rule isodeflI) |
|
622 |
apply (simp add: cast_sprod_defl cast_isodefl) |
|
623 |
apply (simp add: emb_sprod_def prj_sprod_def) |
|
624 |
apply (simp add: sprod_map_map isodefl_strict) |
|
625 |
done |
|
626 |
||
41297 | 627 |
lemma isodefl_prod: |
40504 | 628 |
"isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow> |
41297 | 629 |
isodefl (prod_map\<cdot>d1\<cdot>d2) (prod_defl\<cdot>t1\<cdot>t2)" |
40504 | 630 |
apply (rule isodeflI) |
631 |
apply (simp add: cast_prod_defl cast_isodefl) |
|
632 |
apply (simp add: emb_prod_def prj_prod_def) |
|
41297 | 633 |
apply (simp add: prod_map_map cfcomp1) |
40504 | 634 |
done |
635 |
||
636 |
lemma isodefl_u: |
|
41437 | 637 |
"isodefl d t \<Longrightarrow> isodefl (u_map\<cdot>d) (u_defl\<cdot>t)" |
40504 | 638 |
apply (rule isodeflI) |
41437 | 639 |
apply (simp add: cast_u_defl cast_isodefl) |
640 |
apply (simp add: emb_u_def prj_u_def liftemb_eq liftprj_eq u_map_map) |
|
641 |
done |
|
642 |
||
643 |
lemma isodefl_u_liftdefl: |
|
644 |
"isodefl' d t \<Longrightarrow> isodefl (u_map\<cdot>d) (u_liftdefl\<cdot>t)" |
|
645 |
apply (rule isodeflI) |
|
646 |
apply (simp add: cast_u_liftdefl isodefl'_def) |
|
40504 | 647 |
apply (simp add: emb_u_def prj_u_def liftemb_eq liftprj_eq) |
648 |
done |
|
649 |
||
650 |
lemma encode_prod_u_map: |
|
41297 | 651 |
"encode_prod_u\<cdot>(u_map\<cdot>(prod_map\<cdot>f\<cdot>g)\<cdot>(decode_prod_u\<cdot>x)) |
40504 | 652 |
= sprod_map\<cdot>(u_map\<cdot>f)\<cdot>(u_map\<cdot>g)\<cdot>x" |
653 |
unfolding encode_prod_u_def decode_prod_u_def |
|
654 |
apply (case_tac x, simp, rename_tac a b) |
|
655 |
apply (case_tac a, simp, case_tac b, simp, simp) |
|
656 |
done |
|
657 |
||
41297 | 658 |
lemma isodefl_prod_u: |
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
659 |
assumes "isodefl' d1 t1" and "isodefl' d2 t2" |
41297 | 660 |
shows "isodefl' (prod_map\<cdot>d1\<cdot>d2) (prod_liftdefl\<cdot>t1\<cdot>t2)" |
41292
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
661 |
using assms unfolding isodefl'_def |
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
662 |
unfolding liftemb_prod_def liftprj_prod_def |
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
huffman
parents:
41290
diff
changeset
|
663 |
by (simp add: cast_prod_liftdefl cfcomp1 encode_prod_u_map sprod_map_map) |
40504 | 664 |
|
40592
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
665 |
lemma encode_cfun_map: |
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
666 |
"encode_cfun\<cdot>(cfun_map\<cdot>f\<cdot>g\<cdot>(decode_cfun\<cdot>x)) |
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
667 |
= sfun_map\<cdot>(u_map\<cdot>f)\<cdot>g\<cdot>x" |
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
668 |
unfolding encode_cfun_def decode_cfun_def |
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
669 |
apply (simp add: sfun_eq_iff cfun_map_def sfun_map_def) |
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
670 |
apply (rule cfun_eqI, rename_tac y, case_tac y, simp_all) |
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
671 |
done |
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
672 |
|
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
673 |
lemma isodefl_cfun: |
40830 | 674 |
assumes "isodefl (u_map\<cdot>d1) t1" and "isodefl d2 t2" |
675 |
shows "isodefl (cfun_map\<cdot>d1\<cdot>d2) (sfun_defl\<cdot>t1\<cdot>t2)" |
|
676 |
using isodefl_sfun [OF assms] unfolding isodefl_def |
|
677 |
by (simp add: emb_cfun_def prj_cfun_def cfcomp1 encode_cfun_map) |
|
40592
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
678 |
|
81577 | 679 |
|
62175 | 680 |
subsection \<open>Setting up the domain package\<close> |
40504 | 681 |
|
57945
cacb00a569e0
prefer 'named_theorems' over Named_Thms, with subtle change of semantics due to visual order vs. internal reverse order;
wenzelm
parents:
48891
diff
changeset
|
682 |
named_theorems domain_defl_simps "theorems like DEFL('a t) = t_defl$DEFL('a)" |
59028 | 683 |
and domain_isodefl "theorems like isodefl d t ==> isodefl (foo_map$d) (foo_defl$t)" |
57945
cacb00a569e0
prefer 'named_theorems' over Named_Thms, with subtle change of semantics due to visual order vs. internal reverse order;
wenzelm
parents:
48891
diff
changeset
|
684 |
|
69605 | 685 |
ML_file \<open>Tools/Domain/domain_isomorphism.ML\<close> |
686 |
ML_file \<open>Tools/Domain/domain_axioms.ML\<close> |
|
687 |
ML_file \<open>Tools/Domain/domain.ML\<close> |
|
40504 | 688 |
|
689 |
lemmas [domain_defl_simps] = |
|
40592
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
690 |
DEFL_cfun DEFL_sfun DEFL_ssum DEFL_sprod DEFL_prod DEFL_u |
41437 | 691 |
liftdefl_eq LIFTDEFL_prod u_liftdefl_liftdefl_of |
40504 | 692 |
|
693 |
lemmas [domain_map_ID] = |
|
41297 | 694 |
cfun_map_ID sfun_map_ID ssum_map_ID sprod_map_ID prod_map_ID u_map_ID |
40504 | 695 |
|
696 |
lemmas [domain_isodefl] = |
|
40592
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
697 |
isodefl_u isodefl_sfun isodefl_ssum isodefl_sprod |
41436 | 698 |
isodefl_cfun isodefl_prod isodefl_prod_u isodefl'_liftdefl_of |
41437 | 699 |
isodefl_u_liftdefl |
40504 | 700 |
|
701 |
lemmas [domain_deflation] = |
|
40592
f432973ce0f6
move strict function type into main HOLCF; instance cfun :: (predomain, domain) domain
huffman
parents:
40575
diff
changeset
|
702 |
deflation_cfun_map deflation_sfun_map deflation_ssum_map |
41297 | 703 |
deflation_sprod_map deflation_prod_map deflation_u_map |
40504 | 704 |
|
62175 | 705 |
setup \<open> |
40737
2037021f034f
remove map function names from domain package theory data
huffman
parents:
40592
diff
changeset
|
706 |
fold Domain_Take_Proofs.add_rec_type |
68357 | 707 |
[(\<^type_name>\<open>cfun\<close>, [true, true]), |
708 |
(\<^type_name>\<open>sfun\<close>, [true, true]), |
|
709 |
(\<^type_name>\<open>ssum\<close>, [true, true]), |
|
710 |
(\<^type_name>\<open>sprod\<close>, [true, true]), |
|
711 |
(\<^type_name>\<open>prod\<close>, [true, true]), |
|
712 |
(\<^type_name>\<open>u\<close>, [true])] |
|
62175 | 713 |
\<close> |
40504 | 714 |
|
15741 | 715 |
end |