# HG changeset patch # User huffman # Date 1292760362 28800 # Node ID efd23c1d988687cff1f1e3d92004863ef258d199 # Parent 6d66975b711fe0611f04b5ba4b06ab52988c0f0d renamed Bifinite.thy to Representable.thy diff -r 6d66975b711f -r efd23c1d9886 src/HOL/HOLCF/Bifinite.thy --- a/src/HOL/HOLCF/Bifinite.thy Fri Dec 17 16:43:45 2010 -0800 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,826 +0,0 @@ -(* Title: HOLCF/Bifinite.thy - Author: Brian Huffman -*) - -header {* Bifinite domains *} - -theory Bifinite -imports Algebraic Map_Functions Countable -begin - -subsection {* Class of bifinite domains *} - -text {* - We define a ``domain'' as a pcpo that is isomorphic to some - algebraic deflation over the universal domain; this is equivalent - to being omega-bifinite. - - A predomain is a cpo that, when lifted, becomes a domain. -*} - -class predomain = cpo + - fixes liftdefl :: "('a::cpo) itself \ defl" - fixes liftemb :: "'a\<^sub>\ \ udom" - fixes liftprj :: "udom \ 'a\<^sub>\" - assumes predomain_ep: "ep_pair liftemb liftprj" - assumes cast_liftdefl: "cast\(liftdefl TYPE('a::cpo)) = liftemb oo liftprj" - -syntax "_LIFTDEFL" :: "type \ logic" ("(1LIFTDEFL/(1'(_')))") -translations "LIFTDEFL('t)" \ "CONST liftdefl TYPE('t)" - -class "domain" = predomain + pcpo + - fixes emb :: "'a::cpo \ udom" - fixes prj :: "udom \ 'a::cpo" - fixes defl :: "'a itself \ defl" - assumes ep_pair_emb_prj: "ep_pair emb prj" - assumes cast_DEFL: "cast\(defl TYPE('a)) = emb oo prj" - -syntax "_DEFL" :: "type \ defl" ("(1DEFL/(1'(_')))") -translations "DEFL('t)" \ "CONST defl TYPE('t)" - -interpretation "domain": pcpo_ep_pair emb prj - unfolding pcpo_ep_pair_def - by (rule ep_pair_emb_prj) - -lemmas emb_inverse = domain.e_inverse -lemmas emb_prj_below = domain.e_p_below -lemmas emb_eq_iff = domain.e_eq_iff -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. -*} - -interpretation compact_basis: - ideal_completion below Rep_compact_basis "approximants::'a::domain \ _" -proof - - obtain Y where Y: "\i. Y i \ Y (Suc i)" - and DEFL: "DEFL('a) = (\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 - 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) - show "\i. finite_deflation (approx i)" - unfolding approx_def - apply (rule domain.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) - done - qed - (* FIXME: why does show ?thesis fail here? *) - show "ideal_completion below Rep_compact_basis (approximants::'a \ _)" .. -qed - -subsection {* Chains of approx functions *} - -definition u_approx :: "nat \ udom\<^sub>\ \ udom\<^sub>\" - where "u_approx = (\i. u_map\(udom_approx i))" - -definition sfun_approx :: "nat \ (udom \! udom) \ (udom \! udom)" - where "sfun_approx = (\i. sfun_map\(udom_approx i)\(udom_approx i))" - -definition prod_approx :: "nat \ udom \ udom \ udom \ udom" - where "prod_approx = (\i. cprod_map\(udom_approx i)\(udom_approx i))" - -definition sprod_approx :: "nat \ udom \ udom \ udom \ udom" - where "sprod_approx = (\i. sprod_map\(udom_approx i)\(udom_approx i))" - -definition ssum_approx :: "nat \ udom \ udom \ udom \ udom" - where "ssum_approx = (\i. ssum_map\(udom_approx i)\(udom_approx i))" - -lemma approx_chain_lemma1: - assumes "m\ID = ID" - assumes "\d. finite_deflation d \ finite_deflation (m\d)" - shows "approx_chain (\i. m\(udom_approx i))" -by (rule approx_chain.intro) - (simp_all add: lub_distribs finite_deflation_udom_approx assms) - -lemma approx_chain_lemma2: - assumes "m\ID\ID = ID" - assumes "\a b. \finite_deflation a; finite_deflation b\ - \ finite_deflation (m\a\b)" - shows "approx_chain (\i. m\(udom_approx i)\(udom_approx i))" -by (rule approx_chain.intro) - (simp_all add: lub_distribs finite_deflation_udom_approx assms) - -lemma u_approx: "approx_chain u_approx" -using u_map_ID finite_deflation_u_map -unfolding u_approx_def by (rule approx_chain_lemma1) - -lemma sfun_approx: "approx_chain sfun_approx" -using sfun_map_ID finite_deflation_sfun_map -unfolding sfun_approx_def by (rule approx_chain_lemma2) - -lemma prod_approx: "approx_chain prod_approx" -using cprod_map_ID finite_deflation_cprod_map -unfolding prod_approx_def by (rule approx_chain_lemma2) - -lemma sprod_approx: "approx_chain sprod_approx" -using sprod_map_ID finite_deflation_sprod_map -unfolding sprod_approx_def by (rule approx_chain_lemma2) - -lemma ssum_approx: "approx_chain ssum_approx" -using ssum_map_ID finite_deflation_ssum_map -unfolding ssum_approx_def by (rule approx_chain_lemma2) - -subsection {* Type combinators *} - -definition - defl_fun1 :: - "(nat \ 'a \ 'a) \ ((udom \ udom) \ ('a \ 'a)) \ (defl \ defl)" -where - "defl_fun1 approx f = - defl.basis_fun (\a. - defl_principal (Abs_fin_defl - (udom_emb approx oo f\(Rep_fin_defl a) oo udom_prj approx)))" - -definition - defl_fun2 :: - "(nat \ 'a \ 'a) \ ((udom \ udom) \ (udom \ udom) \ ('a \ 'a)) - \ (defl \ defl \ defl)" -where - "defl_fun2 approx f = - defl.basis_fun (\a. - defl.basis_fun (\b. - defl_principal (Abs_fin_defl - (udom_emb approx oo - f\(Rep_fin_defl a)\(Rep_fin_defl b) oo udom_prj approx))))" - -lemma cast_defl_fun1: - assumes approx: "approx_chain approx" - assumes f: "\a. finite_deflation a \ finite_deflation (f\a)" - shows "cast\(defl_fun1 approx f\A) = udom_emb approx oo f\(cast\A) oo udom_prj approx" -proof - - 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 f, rule finite_deflation_Rep_fin_defl) - done - show ?thesis - by (induct A rule: defl.principal_induct, simp) - (simp only: defl_fun1_def - defl.basis_fun_principal - defl.basis_fun_mono - defl.principal_mono - Abs_fin_defl_mono [OF 1 1] - monofun_cfun below_refl - Rep_fin_defl_mono - cast_defl_principal - Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1]) -qed - -lemma cast_defl_fun2: - assumes approx: "approx_chain approx" - assumes f: "\a b. finite_deflation a \ finite_deflation b \ - finite_deflation (f\a\b)" - shows "cast\(defl_fun2 approx f\A\B) = - udom_emb approx oo f\(cast\A)\(cast\B) oo udom_prj approx" -proof - - have 1: "\a b. finite_deflation (udom_emb approx oo - f\(Rep_fin_defl a)\(Rep_fin_defl b) oo udom_prj approx)" - apply (rule ep_pair.finite_deflation_e_d_p) - apply (rule ep_pair_udom [OF approx]) - apply (rule f, (rule finite_deflation_Rep_fin_defl)+) - done - show ?thesis - by (induct A B rule: defl.principal_induct2, simp, simp) - (simp only: defl_fun2_def - defl.basis_fun_principal - defl.basis_fun_mono - defl.principal_mono - Abs_fin_defl_mono [OF 1 1] - monofun_cfun below_refl - Rep_fin_defl_mono - cast_defl_principal - Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1]) -qed - -definition u_defl :: "defl \ defl" - where "u_defl = defl_fun1 u_approx u_map" - -definition sfun_defl :: "defl \ defl \ defl" - where "sfun_defl = defl_fun2 sfun_approx sfun_map" - -definition prod_defl :: "defl \ defl \ defl" - where "prod_defl = defl_fun2 prod_approx cprod_map" - -definition sprod_defl :: "defl \ defl \ defl" - where "sprod_defl = defl_fun2 sprod_approx sprod_map" - -definition ssum_defl :: "defl \ defl \ defl" -where "ssum_defl = defl_fun2 ssum_approx ssum_map" - -lemma cast_u_defl: - "cast\(u_defl\A) = - udom_emb u_approx oo u_map\(cast\A) oo udom_prj u_approx" -using u_approx finite_deflation_u_map -unfolding u_defl_def by (rule cast_defl_fun1) - -lemma cast_sfun_defl: - "cast\(sfun_defl\A\B) = - udom_emb sfun_approx oo sfun_map\(cast\A)\(cast\B) oo udom_prj sfun_approx" -using sfun_approx finite_deflation_sfun_map -unfolding sfun_defl_def by (rule cast_defl_fun2) - -lemma cast_prod_defl: - "cast\(prod_defl\A\B) = udom_emb prod_approx oo - cprod_map\(cast\A)\(cast\B) oo udom_prj prod_approx" -using prod_approx finite_deflation_cprod_map -unfolding prod_defl_def by (rule cast_defl_fun2) - -lemma cast_sprod_defl: - "cast\(sprod_defl\A\B) = - udom_emb sprod_approx oo - sprod_map\(cast\A)\(cast\B) oo - udom_prj sprod_approx" -using sprod_approx finite_deflation_sprod_map -unfolding sprod_defl_def by (rule cast_defl_fun2) - -lemma cast_ssum_defl: - "cast\(ssum_defl\A\B) = - udom_emb ssum_approx oo ssum_map\(cast\A)\(cast\B) oo udom_prj ssum_approx" -using ssum_approx finite_deflation_ssum_map -unfolding ssum_defl_def by (rule cast_defl_fun2) - -subsection {* Lemma for proving domain instances *} - -text {* - A class of domains where @{const liftemb}, @{const liftprj}, - and @{const liftdefl} are all defined in the standard way. -*} - -class liftdomain = "domain" + - assumes liftemb_eq: "liftemb = udom_emb u_approx oo u_map\emb" - assumes liftprj_eq: "liftprj = u_map\prj oo udom_prj u_approx" - assumes liftdefl_eq: "liftdefl TYPE('a::cpo) = u_defl\DEFL('a)" - -text {* Temporarily relax type constraints. *} - -setup {* - fold Sign.add_const_constraint - [ (@{const_name defl}, SOME @{typ "'a::pcpo itself \ defl"}) - , (@{const_name emb}, SOME @{typ "'a::pcpo \ udom"}) - , (@{const_name prj}, SOME @{typ "udom \ 'a::pcpo"}) - , (@{const_name liftdefl}, SOME @{typ "'a::pcpo itself \ defl"}) - , (@{const_name liftemb}, SOME @{typ "'a::pcpo u \ udom"}) - , (@{const_name liftprj}, SOME @{typ "udom \ 'a::pcpo u"}) ] -*} - -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" - assumes liftdefl: "liftdefl TYPE('a) = u_defl\DEFL('a)" - assumes ep_pair: "ep_pair emb (prj :: udom \ 'a)" - assumes cast_defl: "cast\DEFL('a) = emb oo (prj :: udom \ 'a)" - shows "OFCLASS('a, liftdomain_class)" -proof - show "ep_pair liftemb (liftprj :: udom \ 'a u)" - unfolding liftemb liftprj - by (intro ep_pair_comp ep_pair_u_map ep_pair ep_pair_udom u_approx) - show "cast\LIFTDEFL('a) = liftemb oo (liftprj :: udom \ 'a u)" - unfolding liftemb liftprj liftdefl - by (simp add: cfcomp1 cast_u_defl cast_defl u_map_map) -next -qed fact+ - -text {* Restore original type constraints. *} - -setup {* - fold Sign.add_const_constraint - [ (@{const_name defl}, SOME @{typ "'a::domain itself \ defl"}) - , (@{const_name emb}, SOME @{typ "'a::domain \ udom"}) - , (@{const_name prj}, SOME @{typ "udom \ 'a::domain"}) - , (@{const_name liftdefl}, SOME @{typ "'a::predomain itself \ defl"}) - , (@{const_name liftemb}, SOME @{typ "'a::predomain u \ udom"}) - , (@{const_name liftprj}, SOME @{typ "udom \ 'a::predomain u"}) ] -*} - -subsection {* Class instance proofs *} - -subsubsection {* Universal domain *} - -instantiation udom :: liftdomain -begin - -definition [simp]: - "emb = (ID :: udom \ udom)" - -definition [simp]: - "prj = (ID :: udom \ udom)" - -definition - "defl (t::udom itself) = (\i. defl_principal (Abs_fin_defl (udom_approx i)))" - -definition - "(liftemb :: udom u \ udom) = udom_emb u_approx oo u_map\emb" - -definition - "(liftprj :: udom \ udom u) = u_map\prj oo udom_prj u_approx" - -definition - "liftdefl (t::udom itself) = u_defl\DEFL(udom)" - -instance -using liftemb_udom_def liftprj_udom_def liftdefl_udom_def -proof (rule liftdomain_class_intro) - show "ep_pair emb (prj :: udom \ udom)" - by (simp add: ep_pair.intro) - show "cast\DEFL(udom) = emb oo (prj :: udom \ udom)" - unfolding defl_udom_def - apply (subst contlub_cfun_arg) - apply (rule chainI) - apply (rule defl.principal_mono) - apply (simp add: below_fin_defl_def) - apply (simp add: Abs_fin_defl_inverse finite_deflation_udom_approx) - apply (rule chainE) - apply (rule chain_udom_approx) - apply (subst cast_defl_principal) - apply (simp add: Abs_fin_defl_inverse finite_deflation_udom_approx) - done -qed - -end - -subsubsection {* Lifted cpo *} - -instantiation u :: (predomain) liftdomain -begin - -definition - "emb = liftemb" - -definition - "prj = liftprj" - -definition - "defl (t::'a u itself) = LIFTDEFL('a)" - -definition - "(liftemb :: 'a u u \ udom) = udom_emb u_approx oo u_map\emb" - -definition - "(liftprj :: udom \ 'a u u) = u_map\prj oo udom_prj u_approx" - -definition - "liftdefl (t::'a u itself) = u_defl\DEFL('a u)" - -instance -using liftemb_u_def liftprj_u_def liftdefl_u_def -proof (rule liftdomain_class_intro) - show "ep_pair emb (prj :: udom \ 'a u)" - unfolding emb_u_def prj_u_def - by (rule predomain_ep) - show "cast\DEFL('a u) = emb oo (prj :: udom \ 'a u)" - unfolding emb_u_def prj_u_def defl_u_def - by (rule cast_liftdefl) -qed - -end - -lemma DEFL_u: "DEFL('a::predomain u) = LIFTDEFL('a)" -by (rule defl_u_def) - -subsubsection {* Strict function space *} - -instantiation sfun :: ("domain", "domain") liftdomain -begin - -definition - "emb = udom_emb sfun_approx oo sfun_map\prj\emb" - -definition - "prj = sfun_map\emb\prj oo udom_prj sfun_approx" - -definition - "defl (t::('a \! 'b) itself) = sfun_defl\DEFL('a)\DEFL('b)" - -definition - "(liftemb :: ('a \! 'b) u \ udom) = udom_emb u_approx oo u_map\emb" - -definition - "(liftprj :: udom \ ('a \! 'b) u) = u_map\prj oo udom_prj u_approx" - -definition - "liftdefl (t::('a \! 'b) itself) = u_defl\DEFL('a \! 'b)" - -instance -using liftemb_sfun_def liftprj_sfun_def liftdefl_sfun_def -proof (rule liftdomain_class_intro) - show "ep_pair emb (prj :: udom \ 'a \! 'b)" - unfolding emb_sfun_def prj_sfun_def - using ep_pair_udom [OF sfun_approx] - by (intro ep_pair_comp ep_pair_sfun_map ep_pair_emb_prj) - show "cast\DEFL('a \! 'b) = emb oo (prj :: udom \ 'a \! 'b)" - unfolding emb_sfun_def prj_sfun_def defl_sfun_def cast_sfun_defl - by (simp add: cast_DEFL oo_def sfun_eq_iff sfun_map_map) -qed - -end - -lemma DEFL_sfun: - "DEFL('a::domain \! 'b::domain) = sfun_defl\DEFL('a)\DEFL('b)" -by (rule defl_sfun_def) - -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 - -definition - "emb = emb oo encode_cfun" - -definition - "prj = decode_cfun oo prj" - -definition - "defl (t::('a \ 'b) itself) = DEFL('a u \! 'b)" - -definition - "(liftemb :: ('a \ 'b) u \ udom) = udom_emb u_approx oo u_map\emb" - -definition - "(liftprj :: udom \ ('a \ 'b) u) = u_map\prj oo udom_prj u_approx" - -definition - "liftdefl (t::('a \ 'b) itself) = u_defl\DEFL('a \ 'b)" - -instance -using liftemb_cfun_def liftprj_cfun_def liftdefl_cfun_def -proof (rule liftdomain_class_intro) - have "ep_pair encode_cfun decode_cfun" - by (rule ep_pair.intro, simp_all) - thus "ep_pair emb (prj :: udom \ 'a \ 'b)" - unfolding emb_cfun_def prj_cfun_def - using ep_pair_emb_prj by (rule ep_pair_comp) - show "cast\DEFL('a \ 'b) = emb oo (prj :: udom \ 'a \ 'b)" - unfolding emb_cfun_def prj_cfun_def defl_cfun_def - by (simp add: cast_DEFL cfcomp1) -qed - -end - -lemma DEFL_cfun: - "DEFL('a::predomain \ 'b::domain) = DEFL('a u \! 'b)" -by (rule defl_cfun_def) - -subsubsection {* Strict product *} - -instantiation sprod :: ("domain", "domain") liftdomain -begin - -definition - "emb = udom_emb sprod_approx oo sprod_map\emb\emb" - -definition - "prj = sprod_map\prj\prj oo udom_prj sprod_approx" - -definition - "defl (t::('a \ 'b) itself) = sprod_defl\DEFL('a)\DEFL('b)" - -definition - "(liftemb :: ('a \ 'b) u \ udom) = udom_emb u_approx oo u_map\emb" - -definition - "(liftprj :: udom \ ('a \ 'b) u) = u_map\prj oo udom_prj u_approx" - -definition - "liftdefl (t::('a \ 'b) itself) = u_defl\DEFL('a \ 'b)" - -instance -using liftemb_sprod_def liftprj_sprod_def liftdefl_sprod_def -proof (rule liftdomain_class_intro) - show "ep_pair emb (prj :: udom \ 'a \ 'b)" - unfolding emb_sprod_def prj_sprod_def - using ep_pair_udom [OF sprod_approx] - by (intro ep_pair_comp ep_pair_sprod_map ep_pair_emb_prj) -next - show "cast\DEFL('a \ 'b) = emb oo (prj :: udom \ 'a \ 'b)" - unfolding emb_sprod_def prj_sprod_def defl_sprod_def cast_sprod_defl - by (simp add: cast_DEFL oo_def cfun_eq_iff sprod_map_map) -qed - -end - -lemma DEFL_sprod: - "DEFL('a::domain \ 'b::domain) = sprod_defl\DEFL('a)\DEFL('b)" -by (rule defl_sprod_def) - -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 - -definition - "liftemb = emb oo encode_prod_u" - -definition - "liftprj = decode_prod_u oo prj" - -definition - "liftdefl (t::('a \ 'b) itself) = DEFL('a\<^sub>\ \ 'b\<^sub>\)" - -instance proof - have "ep_pair encode_prod_u decode_prod_u" - by (rule ep_pair.intro, simp_all) - thus "ep_pair liftemb (liftprj :: udom \ ('a \ 'b) u)" - unfolding liftemb_prod_def liftprj_prod_def - using ep_pair_emb_prj by (rule ep_pair_comp) - show "cast\LIFTDEFL('a \ 'b) = liftemb oo (liftprj :: udom \ ('a \ 'b) u)" - unfolding liftemb_prod_def liftprj_prod_def liftdefl_prod_def - by (simp add: cast_DEFL cfcomp1) -qed - -end - -instantiation prod :: ("domain", "domain") "domain" -begin - -definition - "emb = udom_emb prod_approx oo cprod_map\emb\emb" - -definition - "prj = cprod_map\prj\prj oo udom_prj prod_approx" - -definition - "defl (t::('a \ 'b) itself) = prod_defl\DEFL('a)\DEFL('b)" - -instance proof - show "ep_pair emb (prj :: udom \ 'a \ 'b)" - unfolding emb_prod_def prj_prod_def - using ep_pair_udom [OF prod_approx] - by (intro ep_pair_comp ep_pair_cprod_map ep_pair_emb_prj) -next - show "cast\DEFL('a \ 'b) = emb oo (prj :: udom \ 'a \ 'b)" - unfolding emb_prod_def prj_prod_def defl_prod_def cast_prod_defl - by (simp add: cast_DEFL oo_def cfun_eq_iff cprod_map_map) -qed - -end - -lemma DEFL_prod: - "DEFL('a::domain \ 'b::domain) = prod_defl\DEFL('a)\DEFL('b)" -by (rule defl_prod_def) - -lemma LIFTDEFL_prod: - "LIFTDEFL('a::predomain \ 'b::predomain) = DEFL('a u \ 'b u)" -by (rule liftdefl_prod_def) - -subsubsection {* Unit type *} - -instantiation unit :: liftdomain -begin - -definition - "emb = (\ :: unit \ udom)" - -definition - "prj = (\ :: udom \ unit)" - -definition - "defl (t::unit itself) = \" - -definition - "(liftemb :: unit u \ udom) = udom_emb u_approx oo u_map\emb" - -definition - "(liftprj :: udom \ unit u) = u_map\prj oo udom_prj u_approx" - -definition - "liftdefl (t::unit itself) = u_defl\DEFL(unit)" - -instance -using liftemb_unit_def liftprj_unit_def liftdefl_unit_def -proof (rule liftdomain_class_intro) - show "ep_pair emb (prj :: udom \ unit)" - unfolding emb_unit_def prj_unit_def - by (simp add: ep_pair.intro) -next - show "cast\DEFL(unit) = emb oo (prj :: udom \ unit)" - unfolding emb_unit_def prj_unit_def defl_unit_def by simp -qed - -end - -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" - -definition - "liftprj = udom_prj discr_approx" - -definition - "liftdefl (t::'a discr itself) = - (\i. defl_principal (Abs_fin_defl (liftemb oo discr_approx i oo liftprj)))" - -instance proof - show "ep_pair liftemb (liftprj :: udom \ 'a discr u)" - unfolding liftemb_discr_def liftprj_discr_def - by (rule ep_pair_udom [OF discr_approx]) - show "cast\LIFTDEFL('a discr) = liftemb oo (liftprj :: udom \ 'a discr u)" - unfolding liftemb_discr_def liftprj_discr_def liftdefl_discr_def - apply (subst contlub_cfun_arg) - apply (rule chainI) - apply (rule defl.principal_mono) - apply (simp add: below_fin_defl_def) - apply (simp add: Abs_fin_defl_inverse - ep_pair.finite_deflation_e_d_p [OF ep_pair_udom [OF discr_approx]] - approx_chain.finite_deflation_approx [OF discr_approx]) - apply (intro monofun_cfun below_refl) - apply (rule chainE) - apply (rule chain_discr_approx) - apply (subst cast_defl_principal) - apply (simp add: Abs_fin_defl_inverse - ep_pair.finite_deflation_e_d_p [OF ep_pair_udom [OF discr_approx]] - approx_chain.finite_deflation_approx [OF discr_approx]) - apply (simp add: lub_distribs) - done -qed - -end - -subsubsection {* Strict sum *} - -instantiation ssum :: ("domain", "domain") liftdomain -begin - -definition - "emb = udom_emb ssum_approx oo ssum_map\emb\emb" - -definition - "prj = ssum_map\prj\prj oo udom_prj ssum_approx" - -definition - "defl (t::('a \ 'b) itself) = ssum_defl\DEFL('a)\DEFL('b)" - -definition - "(liftemb :: ('a \ 'b) u \ udom) = udom_emb u_approx oo u_map\emb" - -definition - "(liftprj :: udom \ ('a \ 'b) u) = u_map\prj oo udom_prj u_approx" - -definition - "liftdefl (t::('a \ 'b) itself) = u_defl\DEFL('a \ 'b)" - -instance -using liftemb_ssum_def liftprj_ssum_def liftdefl_ssum_def -proof (rule liftdomain_class_intro) - show "ep_pair emb (prj :: udom \ 'a \ 'b)" - unfolding emb_ssum_def prj_ssum_def - using ep_pair_udom [OF ssum_approx] - by (intro ep_pair_comp ep_pair_ssum_map ep_pair_emb_prj) - show "cast\DEFL('a \ 'b) = emb oo (prj :: udom \ 'a \ 'b)" - unfolding emb_ssum_def prj_ssum_def defl_ssum_def cast_ssum_defl - by (simp add: cast_DEFL oo_def cfun_eq_iff ssum_map_map) -qed - -end - -lemma DEFL_ssum: - "DEFL('a::domain \ 'b::domain) = ssum_defl\DEFL('a)\DEFL('b)" -by (rule defl_ssum_def) - -subsubsection {* Lifted HOL type *} - -instantiation lift :: (countable) liftdomain -begin - -definition - "emb = emb oo (\ x. Rep_lift x)" - -definition - "prj = (\ y. Abs_lift y) oo prj" - -definition - "defl (t::'a lift itself) = DEFL('a discr u)" - -definition - "(liftemb :: 'a lift u \ udom) = udom_emb u_approx oo u_map\emb" - -definition - "(liftprj :: udom \ 'a lift u) = u_map\prj oo udom_prj u_approx" - -definition - "liftdefl (t::'a lift itself) = u_defl\DEFL('a lift)" - -instance -using liftemb_lift_def liftprj_lift_def liftdefl_lift_def -proof (rule liftdomain_class_intro) - note [simp] = cont_Rep_lift cont_Abs_lift Rep_lift_inverse Abs_lift_inverse - have "ep_pair (\(x::'a lift). Rep_lift x) (\ y. Abs_lift y)" - by (simp add: ep_pair_def) - thus "ep_pair emb (prj :: udom \ 'a lift)" - unfolding emb_lift_def prj_lift_def - using ep_pair_emb_prj by (rule ep_pair_comp) - show "cast\DEFL('a lift) = emb oo (prj :: udom \ 'a lift)" - unfolding emb_lift_def prj_lift_def defl_lift_def cast_DEFL - by (simp add: cfcomp1) -qed - -end - -end diff -r 6d66975b711f -r efd23c1d9886 src/HOL/HOLCF/Compact_Basis.thy --- a/src/HOL/HOLCF/Compact_Basis.thy Fri Dec 17 16:43:45 2010 -0800 +++ b/src/HOL/HOLCF/Compact_Basis.thy Sun Dec 19 04:06:02 2010 -0800 @@ -5,7 +5,7 @@ header {* A compact basis for powerdomains *} theory Compact_Basis -imports Bifinite +imports Representable begin default_sort "domain" diff -r 6d66975b711f -r efd23c1d9886 src/HOL/HOLCF/Domain.thy --- a/src/HOL/HOLCF/Domain.thy Fri Dec 17 16:43:45 2010 -0800 +++ b/src/HOL/HOLCF/Domain.thy Sun Dec 19 04:06:02 2010 -0800 @@ -5,7 +5,7 @@ header {* Domain package *} theory Domain -imports Bifinite Domain_Aux +imports Representable Domain_Aux uses ("Tools/domaindef.ML") ("Tools/Domain/domain_isomorphism.ML") diff -r 6d66975b711f -r efd23c1d9886 src/HOL/HOLCF/IsaMakefile --- a/src/HOL/HOLCF/IsaMakefile Fri Dec 17 16:43:45 2010 -0800 +++ b/src/HOL/HOLCF/IsaMakefile Sun Dec 19 04:06:02 2010 -0800 @@ -37,7 +37,6 @@ ROOT.ML \ Adm.thy \ Algebraic.thy \ - Bifinite.thy \ Cfun.thy \ Compact_Basis.thy \ Completion.thy \ @@ -62,6 +61,7 @@ Porder.thy \ Powerdomains.thy \ Product_Cpo.thy \ + Representable.thy \ Sfun.thy \ Sprod.thy \ Ssum.thy \ diff -r 6d66975b711f -r efd23c1d9886 src/HOL/HOLCF/Representable.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HOLCF/Representable.thy Sun Dec 19 04:06:02 2010 -0800 @@ -0,0 +1,826 @@ +(* Title: HOLCF/Representable.thy + Author: Brian Huffman +*) + +header {* Representable domains *} + +theory Representable +imports Algebraic Map_Functions Countable +begin + +subsection {* Class of representable domains *} + +text {* + We define a ``domain'' as a pcpo that is isomorphic to some + algebraic deflation over the universal domain; this is equivalent + to being omega-bifinite. + + A predomain is a cpo that, when lifted, becomes a domain. +*} + +class predomain = cpo + + fixes liftdefl :: "('a::cpo) itself \ defl" + fixes liftemb :: "'a\<^sub>\ \ udom" + fixes liftprj :: "udom \ 'a\<^sub>\" + assumes predomain_ep: "ep_pair liftemb liftprj" + assumes cast_liftdefl: "cast\(liftdefl TYPE('a::cpo)) = liftemb oo liftprj" + +syntax "_LIFTDEFL" :: "type \ logic" ("(1LIFTDEFL/(1'(_')))") +translations "LIFTDEFL('t)" \ "CONST liftdefl TYPE('t)" + +class "domain" = predomain + pcpo + + fixes emb :: "'a::cpo \ udom" + fixes prj :: "udom \ 'a::cpo" + fixes defl :: "'a itself \ defl" + assumes ep_pair_emb_prj: "ep_pair emb prj" + assumes cast_DEFL: "cast\(defl TYPE('a)) = emb oo prj" + +syntax "_DEFL" :: "type \ defl" ("(1DEFL/(1'(_')))") +translations "DEFL('t)" \ "CONST defl TYPE('t)" + +interpretation "domain": pcpo_ep_pair emb prj + unfolding pcpo_ep_pair_def + by (rule ep_pair_emb_prj) + +lemmas emb_inverse = domain.e_inverse +lemmas emb_prj_below = domain.e_p_below +lemmas emb_eq_iff = domain.e_eq_iff +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. +*} + +interpretation compact_basis: + ideal_completion below Rep_compact_basis "approximants::'a::domain \ _" +proof - + obtain Y where Y: "\i. Y i \ Y (Suc i)" + and DEFL: "DEFL('a) = (\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 + 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) + show "\i. finite_deflation (approx i)" + unfolding approx_def + apply (rule domain.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) + done + qed + (* FIXME: why does show ?thesis fail here? *) + show "ideal_completion below Rep_compact_basis (approximants::'a \ _)" .. +qed + +subsection {* Chains of approx functions *} + +definition u_approx :: "nat \ udom\<^sub>\ \ udom\<^sub>\" + where "u_approx = (\i. u_map\(udom_approx i))" + +definition sfun_approx :: "nat \ (udom \! udom) \ (udom \! udom)" + where "sfun_approx = (\i. sfun_map\(udom_approx i)\(udom_approx i))" + +definition prod_approx :: "nat \ udom \ udom \ udom \ udom" + where "prod_approx = (\i. cprod_map\(udom_approx i)\(udom_approx i))" + +definition sprod_approx :: "nat \ udom \ udom \ udom \ udom" + where "sprod_approx = (\i. sprod_map\(udom_approx i)\(udom_approx i))" + +definition ssum_approx :: "nat \ udom \ udom \ udom \ udom" + where "ssum_approx = (\i. ssum_map\(udom_approx i)\(udom_approx i))" + +lemma approx_chain_lemma1: + assumes "m\ID = ID" + assumes "\d. finite_deflation d \ finite_deflation (m\d)" + shows "approx_chain (\i. m\(udom_approx i))" +by (rule approx_chain.intro) + (simp_all add: lub_distribs finite_deflation_udom_approx assms) + +lemma approx_chain_lemma2: + assumes "m\ID\ID = ID" + assumes "\a b. \finite_deflation a; finite_deflation b\ + \ finite_deflation (m\a\b)" + shows "approx_chain (\i. m\(udom_approx i)\(udom_approx i))" +by (rule approx_chain.intro) + (simp_all add: lub_distribs finite_deflation_udom_approx assms) + +lemma u_approx: "approx_chain u_approx" +using u_map_ID finite_deflation_u_map +unfolding u_approx_def by (rule approx_chain_lemma1) + +lemma sfun_approx: "approx_chain sfun_approx" +using sfun_map_ID finite_deflation_sfun_map +unfolding sfun_approx_def by (rule approx_chain_lemma2) + +lemma prod_approx: "approx_chain prod_approx" +using cprod_map_ID finite_deflation_cprod_map +unfolding prod_approx_def by (rule approx_chain_lemma2) + +lemma sprod_approx: "approx_chain sprod_approx" +using sprod_map_ID finite_deflation_sprod_map +unfolding sprod_approx_def by (rule approx_chain_lemma2) + +lemma ssum_approx: "approx_chain ssum_approx" +using ssum_map_ID finite_deflation_ssum_map +unfolding ssum_approx_def by (rule approx_chain_lemma2) + +subsection {* Type combinators *} + +definition + defl_fun1 :: + "(nat \ 'a \ 'a) \ ((udom \ udom) \ ('a \ 'a)) \ (defl \ defl)" +where + "defl_fun1 approx f = + defl.basis_fun (\a. + defl_principal (Abs_fin_defl + (udom_emb approx oo f\(Rep_fin_defl a) oo udom_prj approx)))" + +definition + defl_fun2 :: + "(nat \ 'a \ 'a) \ ((udom \ udom) \ (udom \ udom) \ ('a \ 'a)) + \ (defl \ defl \ defl)" +where + "defl_fun2 approx f = + defl.basis_fun (\a. + defl.basis_fun (\b. + defl_principal (Abs_fin_defl + (udom_emb approx oo + f\(Rep_fin_defl a)\(Rep_fin_defl b) oo udom_prj approx))))" + +lemma cast_defl_fun1: + assumes approx: "approx_chain approx" + assumes f: "\a. finite_deflation a \ finite_deflation (f\a)" + shows "cast\(defl_fun1 approx f\A) = udom_emb approx oo f\(cast\A) oo udom_prj approx" +proof - + 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 f, rule finite_deflation_Rep_fin_defl) + done + show ?thesis + by (induct A rule: defl.principal_induct, simp) + (simp only: defl_fun1_def + defl.basis_fun_principal + defl.basis_fun_mono + defl.principal_mono + Abs_fin_defl_mono [OF 1 1] + monofun_cfun below_refl + Rep_fin_defl_mono + cast_defl_principal + Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1]) +qed + +lemma cast_defl_fun2: + assumes approx: "approx_chain approx" + assumes f: "\a b. finite_deflation a \ finite_deflation b \ + finite_deflation (f\a\b)" + shows "cast\(defl_fun2 approx f\A\B) = + udom_emb approx oo f\(cast\A)\(cast\B) oo udom_prj approx" +proof - + have 1: "\a b. finite_deflation (udom_emb approx oo + f\(Rep_fin_defl a)\(Rep_fin_defl b) oo udom_prj approx)" + apply (rule ep_pair.finite_deflation_e_d_p) + apply (rule ep_pair_udom [OF approx]) + apply (rule f, (rule finite_deflation_Rep_fin_defl)+) + done + show ?thesis + by (induct A B rule: defl.principal_induct2, simp, simp) + (simp only: defl_fun2_def + defl.basis_fun_principal + defl.basis_fun_mono + defl.principal_mono + Abs_fin_defl_mono [OF 1 1] + monofun_cfun below_refl + Rep_fin_defl_mono + cast_defl_principal + Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1]) +qed + +definition u_defl :: "defl \ defl" + where "u_defl = defl_fun1 u_approx u_map" + +definition sfun_defl :: "defl \ defl \ defl" + where "sfun_defl = defl_fun2 sfun_approx sfun_map" + +definition prod_defl :: "defl \ defl \ defl" + where "prod_defl = defl_fun2 prod_approx cprod_map" + +definition sprod_defl :: "defl \ defl \ defl" + where "sprod_defl = defl_fun2 sprod_approx sprod_map" + +definition ssum_defl :: "defl \ defl \ defl" +where "ssum_defl = defl_fun2 ssum_approx ssum_map" + +lemma cast_u_defl: + "cast\(u_defl\A) = + udom_emb u_approx oo u_map\(cast\A) oo udom_prj u_approx" +using u_approx finite_deflation_u_map +unfolding u_defl_def by (rule cast_defl_fun1) + +lemma cast_sfun_defl: + "cast\(sfun_defl\A\B) = + udom_emb sfun_approx oo sfun_map\(cast\A)\(cast\B) oo udom_prj sfun_approx" +using sfun_approx finite_deflation_sfun_map +unfolding sfun_defl_def by (rule cast_defl_fun2) + +lemma cast_prod_defl: + "cast\(prod_defl\A\B) = udom_emb prod_approx oo + cprod_map\(cast\A)\(cast\B) oo udom_prj prod_approx" +using prod_approx finite_deflation_cprod_map +unfolding prod_defl_def by (rule cast_defl_fun2) + +lemma cast_sprod_defl: + "cast\(sprod_defl\A\B) = + udom_emb sprod_approx oo + sprod_map\(cast\A)\(cast\B) oo + udom_prj sprod_approx" +using sprod_approx finite_deflation_sprod_map +unfolding sprod_defl_def by (rule cast_defl_fun2) + +lemma cast_ssum_defl: + "cast\(ssum_defl\A\B) = + udom_emb ssum_approx oo ssum_map\(cast\A)\(cast\B) oo udom_prj ssum_approx" +using ssum_approx finite_deflation_ssum_map +unfolding ssum_defl_def by (rule cast_defl_fun2) + +subsection {* Lemma for proving domain instances *} + +text {* + A class of domains where @{const liftemb}, @{const liftprj}, + and @{const liftdefl} are all defined in the standard way. +*} + +class liftdomain = "domain" + + assumes liftemb_eq: "liftemb = udom_emb u_approx oo u_map\emb" + assumes liftprj_eq: "liftprj = u_map\prj oo udom_prj u_approx" + assumes liftdefl_eq: "liftdefl TYPE('a::cpo) = u_defl\DEFL('a)" + +text {* Temporarily relax type constraints. *} + +setup {* + fold Sign.add_const_constraint + [ (@{const_name defl}, SOME @{typ "'a::pcpo itself \ defl"}) + , (@{const_name emb}, SOME @{typ "'a::pcpo \ udom"}) + , (@{const_name prj}, SOME @{typ "udom \ 'a::pcpo"}) + , (@{const_name liftdefl}, SOME @{typ "'a::pcpo itself \ defl"}) + , (@{const_name liftemb}, SOME @{typ "'a::pcpo u \ udom"}) + , (@{const_name liftprj}, SOME @{typ "udom \ 'a::pcpo u"}) ] +*} + +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" + assumes liftdefl: "liftdefl TYPE('a) = u_defl\DEFL('a)" + assumes ep_pair: "ep_pair emb (prj :: udom \ 'a)" + assumes cast_defl: "cast\DEFL('a) = emb oo (prj :: udom \ 'a)" + shows "OFCLASS('a, liftdomain_class)" +proof + show "ep_pair liftemb (liftprj :: udom \ 'a u)" + unfolding liftemb liftprj + by (intro ep_pair_comp ep_pair_u_map ep_pair ep_pair_udom u_approx) + show "cast\LIFTDEFL('a) = liftemb oo (liftprj :: udom \ 'a u)" + unfolding liftemb liftprj liftdefl + by (simp add: cfcomp1 cast_u_defl cast_defl u_map_map) +next +qed fact+ + +text {* Restore original type constraints. *} + +setup {* + fold Sign.add_const_constraint + [ (@{const_name defl}, SOME @{typ "'a::domain itself \ defl"}) + , (@{const_name emb}, SOME @{typ "'a::domain \ udom"}) + , (@{const_name prj}, SOME @{typ "udom \ 'a::domain"}) + , (@{const_name liftdefl}, SOME @{typ "'a::predomain itself \ defl"}) + , (@{const_name liftemb}, SOME @{typ "'a::predomain u \ udom"}) + , (@{const_name liftprj}, SOME @{typ "udom \ 'a::predomain u"}) ] +*} + +subsection {* Class instance proofs *} + +subsubsection {* Universal domain *} + +instantiation udom :: liftdomain +begin + +definition [simp]: + "emb = (ID :: udom \ udom)" + +definition [simp]: + "prj = (ID :: udom \ udom)" + +definition + "defl (t::udom itself) = (\i. defl_principal (Abs_fin_defl (udom_approx i)))" + +definition + "(liftemb :: udom u \ udom) = udom_emb u_approx oo u_map\emb" + +definition + "(liftprj :: udom \ udom u) = u_map\prj oo udom_prj u_approx" + +definition + "liftdefl (t::udom itself) = u_defl\DEFL(udom)" + +instance +using liftemb_udom_def liftprj_udom_def liftdefl_udom_def +proof (rule liftdomain_class_intro) + show "ep_pair emb (prj :: udom \ udom)" + by (simp add: ep_pair.intro) + show "cast\DEFL(udom) = emb oo (prj :: udom \ udom)" + unfolding defl_udom_def + apply (subst contlub_cfun_arg) + apply (rule chainI) + apply (rule defl.principal_mono) + apply (simp add: below_fin_defl_def) + apply (simp add: Abs_fin_defl_inverse finite_deflation_udom_approx) + apply (rule chainE) + apply (rule chain_udom_approx) + apply (subst cast_defl_principal) + apply (simp add: Abs_fin_defl_inverse finite_deflation_udom_approx) + done +qed + +end + +subsubsection {* Lifted cpo *} + +instantiation u :: (predomain) liftdomain +begin + +definition + "emb = liftemb" + +definition + "prj = liftprj" + +definition + "defl (t::'a u itself) = LIFTDEFL('a)" + +definition + "(liftemb :: 'a u u \ udom) = udom_emb u_approx oo u_map\emb" + +definition + "(liftprj :: udom \ 'a u u) = u_map\prj oo udom_prj u_approx" + +definition + "liftdefl (t::'a u itself) = u_defl\DEFL('a u)" + +instance +using liftemb_u_def liftprj_u_def liftdefl_u_def +proof (rule liftdomain_class_intro) + show "ep_pair emb (prj :: udom \ 'a u)" + unfolding emb_u_def prj_u_def + by (rule predomain_ep) + show "cast\DEFL('a u) = emb oo (prj :: udom \ 'a u)" + unfolding emb_u_def prj_u_def defl_u_def + by (rule cast_liftdefl) +qed + +end + +lemma DEFL_u: "DEFL('a::predomain u) = LIFTDEFL('a)" +by (rule defl_u_def) + +subsubsection {* Strict function space *} + +instantiation sfun :: ("domain", "domain") liftdomain +begin + +definition + "emb = udom_emb sfun_approx oo sfun_map\prj\emb" + +definition + "prj = sfun_map\emb\prj oo udom_prj sfun_approx" + +definition + "defl (t::('a \! 'b) itself) = sfun_defl\DEFL('a)\DEFL('b)" + +definition + "(liftemb :: ('a \! 'b) u \ udom) = udom_emb u_approx oo u_map\emb" + +definition + "(liftprj :: udom \ ('a \! 'b) u) = u_map\prj oo udom_prj u_approx" + +definition + "liftdefl (t::('a \! 'b) itself) = u_defl\DEFL('a \! 'b)" + +instance +using liftemb_sfun_def liftprj_sfun_def liftdefl_sfun_def +proof (rule liftdomain_class_intro) + show "ep_pair emb (prj :: udom \ 'a \! 'b)" + unfolding emb_sfun_def prj_sfun_def + using ep_pair_udom [OF sfun_approx] + by (intro ep_pair_comp ep_pair_sfun_map ep_pair_emb_prj) + show "cast\DEFL('a \! 'b) = emb oo (prj :: udom \ 'a \! 'b)" + unfolding emb_sfun_def prj_sfun_def defl_sfun_def cast_sfun_defl + by (simp add: cast_DEFL oo_def sfun_eq_iff sfun_map_map) +qed + +end + +lemma DEFL_sfun: + "DEFL('a::domain \! 'b::domain) = sfun_defl\DEFL('a)\DEFL('b)" +by (rule defl_sfun_def) + +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 + +definition + "emb = emb oo encode_cfun" + +definition + "prj = decode_cfun oo prj" + +definition + "defl (t::('a \ 'b) itself) = DEFL('a u \! 'b)" + +definition + "(liftemb :: ('a \ 'b) u \ udom) = udom_emb u_approx oo u_map\emb" + +definition + "(liftprj :: udom \ ('a \ 'b) u) = u_map\prj oo udom_prj u_approx" + +definition + "liftdefl (t::('a \ 'b) itself) = u_defl\DEFL('a \ 'b)" + +instance +using liftemb_cfun_def liftprj_cfun_def liftdefl_cfun_def +proof (rule liftdomain_class_intro) + have "ep_pair encode_cfun decode_cfun" + by (rule ep_pair.intro, simp_all) + thus "ep_pair emb (prj :: udom \ 'a \ 'b)" + unfolding emb_cfun_def prj_cfun_def + using ep_pair_emb_prj by (rule ep_pair_comp) + show "cast\DEFL('a \ 'b) = emb oo (prj :: udom \ 'a \ 'b)" + unfolding emb_cfun_def prj_cfun_def defl_cfun_def + by (simp add: cast_DEFL cfcomp1) +qed + +end + +lemma DEFL_cfun: + "DEFL('a::predomain \ 'b::domain) = DEFL('a u \! 'b)" +by (rule defl_cfun_def) + +subsubsection {* Strict product *} + +instantiation sprod :: ("domain", "domain") liftdomain +begin + +definition + "emb = udom_emb sprod_approx oo sprod_map\emb\emb" + +definition + "prj = sprod_map\prj\prj oo udom_prj sprod_approx" + +definition + "defl (t::('a \ 'b) itself) = sprod_defl\DEFL('a)\DEFL('b)" + +definition + "(liftemb :: ('a \ 'b) u \ udom) = udom_emb u_approx oo u_map\emb" + +definition + "(liftprj :: udom \ ('a \ 'b) u) = u_map\prj oo udom_prj u_approx" + +definition + "liftdefl (t::('a \ 'b) itself) = u_defl\DEFL('a \ 'b)" + +instance +using liftemb_sprod_def liftprj_sprod_def liftdefl_sprod_def +proof (rule liftdomain_class_intro) + show "ep_pair emb (prj :: udom \ 'a \ 'b)" + unfolding emb_sprod_def prj_sprod_def + using ep_pair_udom [OF sprod_approx] + by (intro ep_pair_comp ep_pair_sprod_map ep_pair_emb_prj) +next + show "cast\DEFL('a \ 'b) = emb oo (prj :: udom \ 'a \ 'b)" + unfolding emb_sprod_def prj_sprod_def defl_sprod_def cast_sprod_defl + by (simp add: cast_DEFL oo_def cfun_eq_iff sprod_map_map) +qed + +end + +lemma DEFL_sprod: + "DEFL('a::domain \ 'b::domain) = sprod_defl\DEFL('a)\DEFL('b)" +by (rule defl_sprod_def) + +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 + +definition + "liftemb = emb oo encode_prod_u" + +definition + "liftprj = decode_prod_u oo prj" + +definition + "liftdefl (t::('a \ 'b) itself) = DEFL('a\<^sub>\ \ 'b\<^sub>\)" + +instance proof + have "ep_pair encode_prod_u decode_prod_u" + by (rule ep_pair.intro, simp_all) + thus "ep_pair liftemb (liftprj :: udom \ ('a \ 'b) u)" + unfolding liftemb_prod_def liftprj_prod_def + using ep_pair_emb_prj by (rule ep_pair_comp) + show "cast\LIFTDEFL('a \ 'b) = liftemb oo (liftprj :: udom \ ('a \ 'b) u)" + unfolding liftemb_prod_def liftprj_prod_def liftdefl_prod_def + by (simp add: cast_DEFL cfcomp1) +qed + +end + +instantiation prod :: ("domain", "domain") "domain" +begin + +definition + "emb = udom_emb prod_approx oo cprod_map\emb\emb" + +definition + "prj = cprod_map\prj\prj oo udom_prj prod_approx" + +definition + "defl (t::('a \ 'b) itself) = prod_defl\DEFL('a)\DEFL('b)" + +instance proof + show "ep_pair emb (prj :: udom \ 'a \ 'b)" + unfolding emb_prod_def prj_prod_def + using ep_pair_udom [OF prod_approx] + by (intro ep_pair_comp ep_pair_cprod_map ep_pair_emb_prj) +next + show "cast\DEFL('a \ 'b) = emb oo (prj :: udom \ 'a \ 'b)" + unfolding emb_prod_def prj_prod_def defl_prod_def cast_prod_defl + by (simp add: cast_DEFL oo_def cfun_eq_iff cprod_map_map) +qed + +end + +lemma DEFL_prod: + "DEFL('a::domain \ 'b::domain) = prod_defl\DEFL('a)\DEFL('b)" +by (rule defl_prod_def) + +lemma LIFTDEFL_prod: + "LIFTDEFL('a::predomain \ 'b::predomain) = DEFL('a u \ 'b u)" +by (rule liftdefl_prod_def) + +subsubsection {* Unit type *} + +instantiation unit :: liftdomain +begin + +definition + "emb = (\ :: unit \ udom)" + +definition + "prj = (\ :: udom \ unit)" + +definition + "defl (t::unit itself) = \" + +definition + "(liftemb :: unit u \ udom) = udom_emb u_approx oo u_map\emb" + +definition + "(liftprj :: udom \ unit u) = u_map\prj oo udom_prj u_approx" + +definition + "liftdefl (t::unit itself) = u_defl\DEFL(unit)" + +instance +using liftemb_unit_def liftprj_unit_def liftdefl_unit_def +proof (rule liftdomain_class_intro) + show "ep_pair emb (prj :: udom \ unit)" + unfolding emb_unit_def prj_unit_def + by (simp add: ep_pair.intro) +next + show "cast\DEFL(unit) = emb oo (prj :: udom \ unit)" + unfolding emb_unit_def prj_unit_def defl_unit_def by simp +qed + +end + +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" + +definition + "liftprj = udom_prj discr_approx" + +definition + "liftdefl (t::'a discr itself) = + (\i. defl_principal (Abs_fin_defl (liftemb oo discr_approx i oo liftprj)))" + +instance proof + show "ep_pair liftemb (liftprj :: udom \ 'a discr u)" + unfolding liftemb_discr_def liftprj_discr_def + by (rule ep_pair_udom [OF discr_approx]) + show "cast\LIFTDEFL('a discr) = liftemb oo (liftprj :: udom \ 'a discr u)" + unfolding liftemb_discr_def liftprj_discr_def liftdefl_discr_def + apply (subst contlub_cfun_arg) + apply (rule chainI) + apply (rule defl.principal_mono) + apply (simp add: below_fin_defl_def) + apply (simp add: Abs_fin_defl_inverse + ep_pair.finite_deflation_e_d_p [OF ep_pair_udom [OF discr_approx]] + approx_chain.finite_deflation_approx [OF discr_approx]) + apply (intro monofun_cfun below_refl) + apply (rule chainE) + apply (rule chain_discr_approx) + apply (subst cast_defl_principal) + apply (simp add: Abs_fin_defl_inverse + ep_pair.finite_deflation_e_d_p [OF ep_pair_udom [OF discr_approx]] + approx_chain.finite_deflation_approx [OF discr_approx]) + apply (simp add: lub_distribs) + done +qed + +end + +subsubsection {* Strict sum *} + +instantiation ssum :: ("domain", "domain") liftdomain +begin + +definition + "emb = udom_emb ssum_approx oo ssum_map\emb\emb" + +definition + "prj = ssum_map\prj\prj oo udom_prj ssum_approx" + +definition + "defl (t::('a \ 'b) itself) = ssum_defl\DEFL('a)\DEFL('b)" + +definition + "(liftemb :: ('a \ 'b) u \ udom) = udom_emb u_approx oo u_map\emb" + +definition + "(liftprj :: udom \ ('a \ 'b) u) = u_map\prj oo udom_prj u_approx" + +definition + "liftdefl (t::('a \ 'b) itself) = u_defl\DEFL('a \ 'b)" + +instance +using liftemb_ssum_def liftprj_ssum_def liftdefl_ssum_def +proof (rule liftdomain_class_intro) + show "ep_pair emb (prj :: udom \ 'a \ 'b)" + unfolding emb_ssum_def prj_ssum_def + using ep_pair_udom [OF ssum_approx] + by (intro ep_pair_comp ep_pair_ssum_map ep_pair_emb_prj) + show "cast\DEFL('a \ 'b) = emb oo (prj :: udom \ 'a \ 'b)" + unfolding emb_ssum_def prj_ssum_def defl_ssum_def cast_ssum_defl + by (simp add: cast_DEFL oo_def cfun_eq_iff ssum_map_map) +qed + +end + +lemma DEFL_ssum: + "DEFL('a::domain \ 'b::domain) = ssum_defl\DEFL('a)\DEFL('b)" +by (rule defl_ssum_def) + +subsubsection {* Lifted HOL type *} + +instantiation lift :: (countable) liftdomain +begin + +definition + "emb = emb oo (\ x. Rep_lift x)" + +definition + "prj = (\ y. Abs_lift y) oo prj" + +definition + "defl (t::'a lift itself) = DEFL('a discr u)" + +definition + "(liftemb :: 'a lift u \ udom) = udom_emb u_approx oo u_map\emb" + +definition + "(liftprj :: udom \ 'a lift u) = u_map\prj oo udom_prj u_approx" + +definition + "liftdefl (t::'a lift itself) = u_defl\DEFL('a lift)" + +instance +using liftemb_lift_def liftprj_lift_def liftdefl_lift_def +proof (rule liftdomain_class_intro) + note [simp] = cont_Rep_lift cont_Abs_lift Rep_lift_inverse Abs_lift_inverse + have "ep_pair (\(x::'a lift). Rep_lift x) (\ y. Abs_lift y)" + by (simp add: ep_pair_def) + thus "ep_pair emb (prj :: udom \ 'a lift)" + unfolding emb_lift_def prj_lift_def + using ep_pair_emb_prj by (rule ep_pair_comp) + show "cast\DEFL('a lift) = emb oo (prj :: udom \ 'a lift)" + unfolding emb_lift_def prj_lift_def defl_lift_def cast_DEFL + by (simp add: cfcomp1) +qed + +end + +end diff -r 6d66975b711f -r efd23c1d9886 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Fri Dec 17 16:43:45 2010 -0800 +++ b/src/HOL/IsaMakefile Sun Dec 19 04:06:02 2010 -0800 @@ -1404,7 +1404,6 @@ HOLCF/ROOT.ML \ HOLCF/Adm.thy \ HOLCF/Algebraic.thy \ - HOLCF/Bifinite.thy \ HOLCF/Cfun.thy \ HOLCF/Compact_Basis.thy \ HOLCF/Completion.thy \ @@ -1429,6 +1428,7 @@ HOLCF/Porder.thy \ HOLCF/Powerdomains.thy \ HOLCF/Product_Cpo.thy \ + HOLCF/Representable.thy \ HOLCF/Sfun.thy \ HOLCF/Sprod.thy \ HOLCF/Ssum.thy \