# HG changeset patch # User huffman # Date 1292764531 28800 # Node ID 3d7685a4a5ff191aa24d6298ad49c47fbe626bc0 # Parent efd23c1d988687cff1f1e3d92004863ef258d199 reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9) diff -r efd23c1d9886 -r 3d7685a4a5ff NEWS --- a/NEWS Sun Dec 19 04:06:02 2010 -0800 +++ b/NEWS Sun Dec 19 05:15:31 2010 -0800 @@ -487,8 +487,8 @@ Accordingly, users of the definitional package must remove any 'default_sort rep' declarations. INCOMPATIBILITY. -* The old type classes 'profinite' and 'bifinite', along with the -overloaded constant 'approx' have been removed. INCOMPATIBILITY. +* The 'bifinite' class no longer fixes a constant 'approx'; the class +now just asserts that such a function exists. INCOMPATIBILITY. * The type 'udom alg_defl' has been replaced by the non-parameterized type 'defl'. HOLCF no longer defines an embedding of type defl into diff -r efd23c1d9886 -r 3d7685a4a5ff src/HOL/HOLCF/Algebraic.thy --- a/src/HOL/HOLCF/Algebraic.thy Sun Dec 19 04:06:02 2010 -0800 +++ b/src/HOL/HOLCF/Algebraic.thy Sun Dec 19 05:15:31 2010 -0800 @@ -97,9 +97,10 @@ "defl_principal t = Abs_defl {u. u \ t}" lemma fin_defl_countable: "\f::fin_defl \ nat. inj f" -proof - have *: "\d. finite (approx_chain.place udom_approx ` - Rep_compact_basis -` {x. Rep_fin_defl d\x = x})" +proof - + obtain f :: "udom compact_basis \ nat" where inj_f: "inj f" + using compact_basis.countable .. + have *: "\d. finite (f ` Rep_compact_basis -` {x. Rep_fin_defl d\x = x})" apply (rule finite_imageI) apply (rule finite_vimageI) apply (rule Rep_fin_defl.finite_fixes) @@ -107,11 +108,11 @@ done have range_eq: "range Rep_compact_basis = {x. compact x}" using type_definition_compact_basis by (rule type_definition.Rep_range) - show "inj (\d. set_encode - (approx_chain.place udom_approx ` Rep_compact_basis -` {x. Rep_fin_defl d\x = x}))" + have "inj (\d. set_encode + (f ` Rep_compact_basis -` {x. Rep_fin_defl d\x = x}))" apply (rule inj_onI) apply (simp only: set_encode_eq *) - apply (simp only: inj_image_eq_iff approx_chain.inj_place [OF udom_approx]) + apply (simp only: inj_image_eq_iff inj_f) apply (drule_tac f="image Rep_compact_basis" in arg_cong) apply (simp del: vimage_Collect_eq add: range_eq set_eq_iff) apply (rule Rep_fin_defl_inject [THEN iffD1]) @@ -121,6 +122,7 @@ apply (rule Rep_fin_defl.compact_belowI, rename_tac z) apply (drule_tac x=z in spec, simp) done + thus ?thesis by - (rule exI) qed interpretation defl: ideal_completion below defl_principal Rep_defl diff -r efd23c1d9886 -r 3d7685a4a5ff src/HOL/HOLCF/Bifinite.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HOLCF/Bifinite.thy Sun Dec 19 05:15:31 2010 -0800 @@ -0,0 +1,275 @@ +(* Title: HOLCF/Bifinite.thy + Author: Brian Huffman +*) + +header {* Profinite and bifinite cpos *} + +theory Bifinite +imports Map_Functions Countable +begin + +default_sort cpo + +subsection {* Chains of finite deflations *} + +locale approx_chain = + fixes approx :: "nat \ 'a \ 'a" + assumes chain_approx [simp]: "chain (\i. approx i)" + assumes lub_approx [simp]: "(\i. approx i) = ID" + assumes finite_deflation_approx [simp]: "\i. finite_deflation (approx i)" +begin + +lemma deflation_approx: "deflation (approx i)" +using finite_deflation_approx by (rule finite_deflation_imp_deflation) + +lemma approx_idem: "approx i\(approx i\x) = approx i\x" +using deflation_approx by (rule deflation.idem) + +lemma approx_below: "approx i\x \ x" +using deflation_approx by (rule deflation.below) + +lemma finite_range_approx: "finite (range (\x. approx i\x))" +apply (rule finite_deflation.finite_range) +apply (rule finite_deflation_approx) +done + +lemma compact_approx: "compact (approx n\x)" +apply (rule finite_deflation.compact) +apply (rule finite_deflation_approx) +done + +lemma compact_eq_approx: "compact x \ \i. approx i\x = x" +by (rule admD2, simp_all) + +end + +subsection {* Omega-profinite and bifinite domains *} + +class bifinite = pcpo + + assumes bifinite: "\(a::nat \ 'a \ 'a). approx_chain a" + +class profinite = cpo + + assumes profinite: "\(a::nat \ 'a\<^sub>\ \ 'a\<^sub>\). approx_chain a" + +subsection {* Building approx chains *} + +lemma approx_chain_iso: + assumes a: "approx_chain a" + assumes [simp]: "\x. f\(g\x) = x" + assumes [simp]: "\y. g\(f\y) = y" + shows "approx_chain (\i. f oo a i oo g)" +proof - + have 1: "f oo g = ID" by (simp add: cfun_eqI) + have 2: "ep_pair f g" by (simp add: ep_pair_def) + from 1 2 show ?thesis + using a unfolding approx_chain_def + by (simp add: lub_APP ep_pair.finite_deflation_e_d_p) +qed + +lemma approx_chain_u_map: + assumes "approx_chain a" + shows "approx_chain (\i. u_map\(a i))" + using assms unfolding approx_chain_def + by (simp add: lub_APP u_map_ID finite_deflation_u_map) + +lemma approx_chain_sfun_map: + assumes "approx_chain a" and "approx_chain b" + shows "approx_chain (\i. sfun_map\(a i)\(b i))" + using assms unfolding approx_chain_def + by (simp add: lub_APP sfun_map_ID finite_deflation_sfun_map) + +lemma approx_chain_sprod_map: + assumes "approx_chain a" and "approx_chain b" + shows "approx_chain (\i. sprod_map\(a i)\(b i))" + using assms unfolding approx_chain_def + by (simp add: lub_APP sprod_map_ID finite_deflation_sprod_map) + +lemma approx_chain_ssum_map: + assumes "approx_chain a" and "approx_chain b" + shows "approx_chain (\i. ssum_map\(a i)\(b i))" + using assms unfolding approx_chain_def + by (simp add: lub_APP ssum_map_ID finite_deflation_ssum_map) + +lemma approx_chain_cfun_map: + assumes "approx_chain a" and "approx_chain b" + shows "approx_chain (\i. cfun_map\(a i)\(b i))" + using assms unfolding approx_chain_def + by (simp add: lub_APP cfun_map_ID finite_deflation_cfun_map) + +lemma approx_chain_cprod_map: + assumes "approx_chain a" and "approx_chain b" + shows "approx_chain (\i. cprod_map\(a i)\(b i))" + using assms unfolding approx_chain_def + by (simp add: lub_APP cprod_map_ID finite_deflation_cprod_map) + +text {* Approx chains for countable discrete types. *} + +definition discr_approx :: "nat \ 'a::countable discr u \ 'a discr u" + where "discr_approx = (\i. \(up\x). if to_nat (undiscr x) < i then up\x else \)" + +lemma chain_discr_approx [simp]: "chain discr_approx" +unfolding discr_approx_def +by (rule chainI, simp add: monofun_cfun monofun_LAM) + +lemma lub_discr_approx [simp]: "(\i. discr_approx i) = ID" +apply (rule cfun_eqI) +apply (simp add: contlub_cfun_fun) +apply (simp add: discr_approx_def) +apply (case_tac x, simp) +apply (rule lub_eqI) +apply (rule is_lubI) +apply (rule ub_rangeI, simp) +apply (drule ub_rangeD) +apply (erule rev_below_trans) +apply simp +apply (rule lessI) +done + +lemma inj_on_undiscr [simp]: "inj_on undiscr A" +using Discr_undiscr by (rule inj_on_inverseI) + +lemma finite_deflation_discr_approx: "finite_deflation (discr_approx i)" +proof + fix x :: "'a discr u" + show "discr_approx i\x \ x" + unfolding discr_approx_def + by (cases x, simp, simp) + show "discr_approx i\(discr_approx i\x) = discr_approx i\x" + unfolding discr_approx_def + by (cases x, simp, simp) + show "finite {x::'a discr u. discr_approx i\x = x}" + proof (rule finite_subset) + let ?S = "insert (\::'a discr u) ((\x. up\x) ` undiscr -` to_nat -` {..x = x} \ ?S" + unfolding discr_approx_def + by (rule subsetI, case_tac x, simp, simp split: split_if_asm) + show "finite ?S" + by (simp add: finite_vimageI) + qed +qed + +lemma discr_approx: "approx_chain discr_approx" +using chain_discr_approx lub_discr_approx finite_deflation_discr_approx +by (rule approx_chain.intro) + +subsection {* Class instance proofs *} + +instance bifinite \ profinite +proof + show "\(a::nat \ 'a\<^sub>\ \ 'a\<^sub>\). approx_chain a" + using bifinite [where 'a='a] + by (fast intro!: approx_chain_u_map) +qed + +instance u :: (profinite) bifinite + by default (rule profinite) + +text {* + Types @{typ "'a \ 'b"} and @{typ "'a u \! 'b"} are isomorphic. +*} + +definition "encode_cfun = (\ f. sfun_abs\(fup\f))" + +definition "decode_cfun = (\ g x. sfun_rep\g\(up\x))" + +lemma decode_encode_cfun [simp]: "decode_cfun\(encode_cfun\x) = x" +unfolding encode_cfun_def decode_cfun_def +by (simp add: eta_cfun) + +lemma encode_decode_cfun [simp]: "encode_cfun\(decode_cfun\y) = y" +unfolding encode_cfun_def decode_cfun_def +apply (simp add: sfun_eq_iff strictify_cancel) +apply (rule cfun_eqI, case_tac x, simp_all) +done + +instance cfun :: (profinite, bifinite) bifinite +proof + obtain a :: "nat \ 'a\<^sub>\ \ 'a\<^sub>\" where a: "approx_chain a" + using profinite .. + obtain b :: "nat \ 'b \ 'b" where b: "approx_chain b" + using bifinite .. + have "approx_chain (\i. decode_cfun oo sfun_map\(a i)\(b i) oo encode_cfun)" + using a b by (simp add: approx_chain_iso approx_chain_sfun_map) + thus "\(a::nat \ ('a \ 'b) \ ('a \ 'b)). approx_chain a" + by - (rule exI) +qed + +text {* + Types @{typ "('a * 'b) u"} and @{typ "'a u \ 'b u"} are isomorphic. +*} + +definition "encode_prod_u = (\(up\(x, y)). (:up\x, up\y:))" + +definition "decode_prod_u = (\(:up\x, up\y:). up\(x, y))" + +lemma decode_encode_prod_u [simp]: "decode_prod_u\(encode_prod_u\x) = x" +unfolding encode_prod_u_def decode_prod_u_def +by (case_tac x, simp, rename_tac y, case_tac y, simp) + +lemma encode_decode_prod_u [simp]: "encode_prod_u\(decode_prod_u\y) = y" +unfolding encode_prod_u_def decode_prod_u_def +apply (case_tac y, simp, rename_tac a b) +apply (case_tac a, simp, case_tac b, simp, simp) +done + +instance prod :: (profinite, profinite) profinite +proof + obtain a :: "nat \ 'a\<^sub>\ \ 'a\<^sub>\" where a: "approx_chain a" + using profinite .. + obtain b :: "nat \ 'b\<^sub>\ \ 'b\<^sub>\" where b: "approx_chain b" + using profinite .. + have "approx_chain (\i. decode_prod_u oo sprod_map\(a i)\(b i) oo encode_prod_u)" + using a b by (simp add: approx_chain_iso approx_chain_sprod_map) + thus "\(a::nat \ ('a \ 'b)\<^sub>\ \ ('a \ 'b)\<^sub>\). approx_chain a" + by - (rule exI) +qed + +instance prod :: (bifinite, bifinite) bifinite +proof + show "\(a::nat \ ('a \ 'b) \ ('a \ 'b)). approx_chain a" + using bifinite [where 'a='a] and bifinite [where 'a='b] + by (fast intro!: approx_chain_cprod_map) +qed + +instance sfun :: (bifinite, bifinite) bifinite +proof + show "\(a::nat \ ('a \! 'b) \ ('a \! 'b)). approx_chain a" + using bifinite [where 'a='a] and bifinite [where 'a='b] + by (fast intro!: approx_chain_sfun_map) +qed + +instance sprod :: (bifinite, bifinite) bifinite +proof + show "\(a::nat \ ('a \ 'b) \ ('a \ 'b)). approx_chain a" + using bifinite [where 'a='a] and bifinite [where 'a='b] + by (fast intro!: approx_chain_sprod_map) +qed + +instance ssum :: (bifinite, bifinite) bifinite +proof + show "\(a::nat \ ('a \ 'b) \ ('a \ 'b)). approx_chain a" + using bifinite [where 'a='a] and bifinite [where 'a='b] + by (fast intro!: approx_chain_ssum_map) +qed + +lemma approx_chain_unit: "approx_chain (\ :: nat \ unit \ unit)" +by (simp add: approx_chain_def cfun_eq_iff finite_deflation_UU) + +instance unit :: bifinite + by default (fast intro!: approx_chain_unit) + +instance discr :: (countable) profinite + by default (fast intro!: discr_approx) + +instance lift :: (countable) bifinite +proof + note [simp] = cont_Abs_lift cont_Rep_lift Rep_lift_inverse Abs_lift_inverse + obtain a :: "nat \ ('a discr)\<^sub>\ \ ('a discr)\<^sub>\" where a: "approx_chain a" + using profinite .. + hence "approx_chain (\i. (\ y. Abs_lift y) oo a i oo (\ x. Rep_lift x))" + by (rule approx_chain_iso) simp_all + thus "\(a::nat \ 'a lift \ 'a lift). approx_chain a" + by - (rule exI) +qed + +end diff -r efd23c1d9886 -r 3d7685a4a5ff src/HOL/HOLCF/ConvexPD.thy --- a/src/HOL/HOLCF/ConvexPD.thy Sun Dec 19 04:06:02 2010 -0800 +++ b/src/HOL/HOLCF/ConvexPD.thy Sun Dec 19 05:15:31 2010 -0800 @@ -468,6 +468,19 @@ subsection {* Convex powerdomain is a domain *} +lemma approx_chain_convex_map: + assumes "approx_chain a" + shows "approx_chain (\i. convex_map\(a i))" + using assms unfolding approx_chain_def + by (simp add: lub_APP convex_map_ID finite_deflation_convex_map) + +instance convex_pd :: ("domain") bifinite +proof + show "\(a::nat \ 'a convex_pd \ 'a convex_pd). approx_chain a" + using bifinite [where 'a='a] + by (fast intro!: approx_chain_convex_map) +qed + definition convex_approx :: "nat \ udom convex_pd \ udom convex_pd" where diff -r efd23c1d9886 -r 3d7685a4a5ff src/HOL/HOLCF/IsaMakefile --- a/src/HOL/HOLCF/IsaMakefile Sun Dec 19 04:06:02 2010 -0800 +++ b/src/HOL/HOLCF/IsaMakefile Sun Dec 19 05:15:31 2010 -0800 @@ -37,6 +37,7 @@ ROOT.ML \ Adm.thy \ Algebraic.thy \ + Bifinite.thy \ Cfun.thy \ Compact_Basis.thy \ Completion.thy \ diff -r efd23c1d9886 -r 3d7685a4a5ff src/HOL/HOLCF/Library/Defl_Bifinite.thy --- a/src/HOL/HOLCF/Library/Defl_Bifinite.thy Sun Dec 19 04:06:02 2010 -0800 +++ b/src/HOL/HOLCF/Library/Defl_Bifinite.thy Sun Dec 19 05:15:31 2010 -0800 @@ -609,6 +609,9 @@ subsection {* Algebraic deflations are a bifinite domain *} +instance defl :: bifinite + by default (fast intro!: defl_approx) + instantiation defl :: liftdomain begin diff -r efd23c1d9886 -r 3d7685a4a5ff src/HOL/HOLCF/LowerPD.thy --- a/src/HOL/HOLCF/LowerPD.thy Sun Dec 19 04:06:02 2010 -0800 +++ b/src/HOL/HOLCF/LowerPD.thy Sun Dec 19 05:15:31 2010 -0800 @@ -460,6 +460,19 @@ subsection {* Lower powerdomain is a domain *} +lemma approx_chain_lower_map: + assumes "approx_chain a" + shows "approx_chain (\i. lower_map\(a i))" + using assms unfolding approx_chain_def + by (simp add: lub_APP lower_map_ID finite_deflation_lower_map) + +instance lower_pd :: ("domain") bifinite +proof + show "\(a::nat \ 'a lower_pd \ 'a lower_pd). approx_chain a" + using bifinite [where 'a='a] + by (fast intro!: approx_chain_lower_map) +qed + definition lower_approx :: "nat \ udom lower_pd \ udom lower_pd" where diff -r efd23c1d9886 -r 3d7685a4a5ff src/HOL/HOLCF/Representable.thy --- a/src/HOL/HOLCF/Representable.thy Sun Dec 19 04:06:02 2010 -0800 +++ b/src/HOL/HOLCF/Representable.thy Sun Dec 19 05:15:31 2010 -0800 @@ -48,41 +48,44 @@ lemmas emb_strict = domain.e_strict lemmas prj_strict = domain.p_strict -subsection {* Domains have a countable compact basis *} - -text {* - Eventually it should be possible to generalize this to an unpointed - variant of the domain class. -*} +subsection {* Domains are bifinite *} -interpretation compact_basis: - ideal_completion below Rep_compact_basis "approximants::'a::domain \ _" +lemma approx_chain_ep_cast: + assumes ep: "ep_pair (e::'a \ udom) (p::udom \ 'a)" + assumes cast_t: "cast\t = e oo p" + shows "\(a::nat \ 'a \ 'a). approx_chain a" proof - + interpret ep_pair e p by fact obtain Y where Y: "\i. Y i \ Y (Suc i)" - and DEFL: "DEFL('a) = (\i. defl_principal (Y i))" + and t: "t = (\i. defl_principal (Y i))" by (rule defl.obtain_principal_chain) - def approx \ "\i. (prj oo cast\(defl_principal (Y i)) oo emb) :: 'a \ 'a" - interpret defl_approx: approx_chain approx + def approx \ "\i. (p oo cast\(defl_principal (Y i)) oo e) :: 'a \ 'a" + have "approx_chain approx" proof (rule approx_chain.intro) show "chain (\i. approx i)" unfolding approx_def by (simp add: Y) show "(\i. approx i) = ID" unfolding approx_def - by (simp add: lub_distribs Y DEFL [symmetric] cast_DEFL cfun_eq_iff) + by (simp add: lub_distribs Y t [symmetric] cast_t cfun_eq_iff) show "\i. finite_deflation (approx i)" unfolding approx_def - apply (rule domain.finite_deflation_p_d_e) + apply (rule finite_deflation_p_d_e) apply (rule finite_deflation_cast) apply (rule defl.compact_principal) apply (rule below_trans [OF monofun_cfun_fun]) apply (rule is_ub_thelub, simp add: Y) - apply (simp add: lub_distribs Y DEFL [symmetric] cast_DEFL) + apply (simp add: lub_distribs Y t [symmetric] cast_t) done qed - (* FIXME: why does show ?thesis fail here? *) - show "ideal_completion below Rep_compact_basis (approximants::'a \ _)" .. + thus "\(a::nat \ 'a \ 'a). approx_chain a" by - (rule exI) qed +instance "domain" \ bifinite +by default (rule approx_chain_ep_cast [OF ep_pair_emb_prj cast_DEFL]) + +instance predomain \ profinite +by default (rule approx_chain_ep_cast [OF predomain_ep cast_liftdefl]) + subsection {* Chains of approx functions *} definition u_approx :: "nat \ udom\<^sub>\ \ udom\<^sub>\" @@ -137,6 +140,8 @@ subsection {* Type combinators *} +default_sort bifinite + definition defl_fun1 :: "(nat \ 'a \ 'a) \ ((udom \ udom) \ ('a \ 'a)) \ (defl \ defl)" @@ -166,7 +171,7 @@ have 1: "\a. finite_deflation (udom_emb approx oo f\(Rep_fin_defl a) oo udom_prj approx)" apply (rule ep_pair.finite_deflation_e_d_p) - apply (rule approx_chain.ep_pair_udom [OF approx]) + apply (rule ep_pair_udom [OF approx]) apply (rule f, rule finite_deflation_Rep_fin_defl) done show ?thesis @@ -279,6 +284,8 @@ , (@{const_name liftprj}, SOME @{typ "udom \ 'a::pcpo u"}) ] *} +default_sort pcpo + lemma liftdomain_class_intro: assumes liftemb: "(liftemb :: 'a u \ udom) = udom_emb u_approx oo u_map\emb" assumes liftprj: "(liftprj :: udom \ 'a u) = u_map\prj oo udom_prj u_approx" @@ -436,26 +443,6 @@ subsubsection {* Continuous function space *} -text {* - Types @{typ "'a \ 'b"} and @{typ "'a u \! 'b"} are isomorphic. -*} - -definition - "encode_cfun = (\ f. sfun_abs\(fup\f))" - -definition - "decode_cfun = (\ g x. sfun_rep\g\(up\x))" - -lemma decode_encode_cfun [simp]: "decode_cfun\(encode_cfun\x) = x" -unfolding encode_cfun_def decode_cfun_def -by (simp add: eta_cfun) - -lemma encode_decode_cfun [simp]: "encode_cfun\(decode_cfun\y) = y" -unfolding encode_cfun_def decode_cfun_def -apply (simp add: sfun_eq_iff strictify_cancel) -apply (rule cfun_eqI, case_tac x, simp_all) -done - instantiation cfun :: (predomain, "domain") liftdomain begin @@ -540,26 +527,6 @@ subsubsection {* Cartesian product *} -text {* - Types @{typ "('a * 'b) u"} and @{typ "'a u \ 'b u"} are isomorphic. -*} - -definition - "encode_prod_u = (\(up\(x, y)). (:up\x, up\y:))" - -definition - "decode_prod_u = (\(:up\x, up\y:). up\(x, y))" - -lemma decode_encode_prod_u [simp]: "decode_prod_u\(encode_prod_u\x) = x" -unfolding encode_prod_u_def decode_prod_u_def -by (case_tac x, simp, rename_tac y, case_tac y, simp) - -lemma encode_decode_prod_u [simp]: "encode_prod_u\(decode_prod_u\y) = y" -unfolding encode_prod_u_def decode_prod_u_def -apply (case_tac y, simp, rename_tac a b) -apply (case_tac a, simp, case_tac b, simp, simp) -done - instantiation prod :: (predomain, predomain) predomain begin @@ -656,66 +623,18 @@ subsubsection {* Discrete cpo *} -definition discr_approx :: "nat \ 'a::countable discr u \ 'a discr u" - where "discr_approx = (\i. \(up\x). if to_nat (undiscr x) < i then up\x else \)" - -lemma chain_discr_approx [simp]: "chain discr_approx" -unfolding discr_approx_def -by (rule chainI, simp add: monofun_cfun monofun_LAM) - -lemma lub_discr_approx [simp]: "(\i. discr_approx i) = ID" -apply (rule cfun_eqI) -apply (simp add: contlub_cfun_fun) -apply (simp add: discr_approx_def) -apply (case_tac x, simp) -apply (rule lub_eqI) -apply (rule is_lubI) -apply (rule ub_rangeI, simp) -apply (drule ub_rangeD) -apply (erule rev_below_trans) -apply simp -apply (rule lessI) -done - -lemma inj_on_undiscr [simp]: "inj_on undiscr A" -using Discr_undiscr by (rule inj_on_inverseI) - -lemma finite_deflation_discr_approx: "finite_deflation (discr_approx i)" -proof - fix x :: "'a discr u" - show "discr_approx i\x \ x" - unfolding discr_approx_def - by (cases x, simp, simp) - show "discr_approx i\(discr_approx i\x) = discr_approx i\x" - unfolding discr_approx_def - by (cases x, simp, simp) - show "finite {x::'a discr u. discr_approx i\x = x}" - proof (rule finite_subset) - let ?S = "insert (\::'a discr u) ((\x. up\x) ` undiscr -` to_nat -` {..x = x} \ ?S" - unfolding discr_approx_def - by (rule subsetI, case_tac x, simp, simp split: split_if_asm) - show "finite ?S" - by (simp add: finite_vimageI) - qed -qed - -lemma discr_approx: "approx_chain discr_approx" -using chain_discr_approx lub_discr_approx finite_deflation_discr_approx -by (rule approx_chain.intro) - instantiation discr :: (countable) predomain begin definition - "liftemb = udom_emb discr_approx" + "(liftemb :: 'a discr u \ udom) = udom_emb discr_approx" definition - "liftprj = udom_prj discr_approx" + "(liftprj :: udom \ 'a discr u) = udom_prj discr_approx" definition "liftdefl (t::'a discr itself) = - (\i. defl_principal (Abs_fin_defl (liftemb oo discr_approx i oo liftprj)))" + (\i. defl_principal (Abs_fin_defl (liftemb oo discr_approx i oo (liftprj::udom \ 'a discr u))))" instance proof show "ep_pair liftemb (liftprj :: udom \ 'a discr u)" diff -r efd23c1d9886 -r 3d7685a4a5ff src/HOL/HOLCF/Universal.thy --- a/src/HOL/HOLCF/Universal.thy Sun Dec 19 04:06:02 2010 -0800 +++ b/src/HOL/HOLCF/Universal.thy Sun Dec 19 05:15:31 2010 -0800 @@ -5,7 +5,7 @@ header {* A universal bifinite domain *} theory Universal -imports Completion Deflation Nat_Bijection +imports Bifinite Completion Nat_Bijection begin subsection {* Basis for universal domain *} @@ -287,11 +287,8 @@ text {* We use a locale to parameterize the construction over a chain of approx functions on the type to be embedded. *} -locale approx_chain = - fixes approx :: "nat \ 'a::pcpo \ 'a" - assumes chain_approx [simp]: "chain (\i. approx i)" - assumes lub_approx [simp]: "(\i. approx i) = ID" - assumes finite_deflation_approx: "\i. finite_deflation (approx i)" +locale bifinite_approx_chain = approx_chain + + constrains approx :: "nat \ 'a::bifinite \ 'a" begin subsubsection {* Choosing a maximal element from a finite set *} @@ -408,30 +405,6 @@ apply (simp add: choose_pos.simps) done -subsubsection {* Properties of approx function *} - -lemma deflation_approx: "deflation (approx i)" -using finite_deflation_approx by (rule finite_deflation_imp_deflation) - -lemma approx_idem: "approx i\(approx i\x) = approx i\x" -using deflation_approx by (rule deflation.idem) - -lemma approx_below: "approx i\x \ x" -using deflation_approx by (rule deflation.below) - -lemma finite_range_approx: "finite (range (\x. approx i\x))" -apply (rule finite_deflation.finite_range) -apply (rule finite_deflation_approx) -done - -lemma compact_approx: "compact (approx n\x)" -apply (rule finite_deflation.compact) -apply (rule finite_deflation_approx) -done - -lemma compact_eq_approx: "compact x \ \i. approx i\x = x" -by (rule admD2, simp_all) - subsubsection {* Compact basis take function *} primrec @@ -804,11 +777,8 @@ apply (rule ubasis_until_less) done -end - -sublocale approx_chain \ compact_basis!: - ideal_completion below Rep_compact_basis - "approximants :: 'a \ 'a compact_basis set" +lemma ideal_completion: + "ideal_completion below Rep_compact_basis (approximants :: 'a \ _)" proof fix w :: "'a" show "below.ideal (approximants w)" @@ -873,9 +843,23 @@ by (rule exI, rule inj_place) qed +end + +interpretation compact_basis!: + ideal_completion below Rep_compact_basis + "approximants :: 'a::bifinite \ 'a compact_basis set" +proof - + obtain a :: "nat \ 'a \ 'a" where "approx_chain a" + using bifinite .. + hence "bifinite_approx_chain a" + unfolding bifinite_approx_chain_def . + thus "ideal_completion below Rep_compact_basis (approximants :: 'a \ _)" + by (rule bifinite_approx_chain.ideal_completion) +qed + subsubsection {* EP-pair from any bifinite domain into \emph{udom} *} -context approx_chain begin +context bifinite_approx_chain begin definition udom_emb :: "'a \ udom" @@ -915,10 +899,11 @@ end -abbreviation "udom_emb \ approx_chain.udom_emb" -abbreviation "udom_prj \ approx_chain.udom_prj" +abbreviation "udom_emb \ bifinite_approx_chain.udom_emb" +abbreviation "udom_prj \ bifinite_approx_chain.udom_prj" -lemmas ep_pair_udom = approx_chain.ep_pair_udom +lemmas ep_pair_udom = + bifinite_approx_chain.ep_pair_udom [unfolded bifinite_approx_chain_def] subsection {* Chain of approx functions for type \emph{udom} *} @@ -1001,7 +986,7 @@ apply (simp add: ubasis_until_same ubasis_le_refl) done -lemma udom_approx: "approx_chain udom_approx" +lemma udom_approx [simp]: "approx_chain udom_approx" proof show "chain (\i. udom_approx i)" by (rule chain_udom_approx) @@ -1009,6 +994,9 @@ by (rule lub_udom_approx) qed +instance udom :: bifinite + by default (fast intro: udom_approx) + hide_const (open) node end diff -r efd23c1d9886 -r 3d7685a4a5ff src/HOL/HOLCF/UpperPD.thy --- a/src/HOL/HOLCF/UpperPD.thy Sun Dec 19 04:06:02 2010 -0800 +++ b/src/HOL/HOLCF/UpperPD.thy Sun Dec 19 05:15:31 2010 -0800 @@ -455,6 +455,19 @@ subsection {* Upper powerdomain is a domain *} +lemma approx_chain_upper_map: + assumes "approx_chain a" + shows "approx_chain (\i. upper_map\(a i))" + using assms unfolding approx_chain_def + by (simp add: lub_APP upper_map_ID finite_deflation_upper_map) + +instance upper_pd :: ("domain") bifinite +proof + show "\(a::nat \ 'a upper_pd \ 'a upper_pd). approx_chain a" + using bifinite [where 'a='a] + by (fast intro!: approx_chain_upper_map) +qed + definition upper_approx :: "nat \ udom upper_pd \ udom upper_pd" where diff -r efd23c1d9886 -r 3d7685a4a5ff src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Sun Dec 19 04:06:02 2010 -0800 +++ b/src/HOL/IsaMakefile Sun Dec 19 05:15:31 2010 -0800 @@ -1404,6 +1404,7 @@ HOLCF/ROOT.ML \ HOLCF/Adm.thy \ HOLCF/Algebraic.thy \ + HOLCF/Bifinite.thy \ HOLCF/Cfun.thy \ HOLCF/Compact_Basis.thy \ HOLCF/Completion.thy \