src/HOL/HOLCF/Representable.thy
changeset 41285 efd23c1d9886
parent 41034 ce5d9e73fb98
child 41286 3d7685a4a5ff
equal deleted inserted replaced
41284:6d66975b711f 41285:efd23c1d9886
       
     1 (*  Title:      HOLCF/Representable.thy
       
     2     Author:     Brian Huffman
       
     3 *)
       
     4 
       
     5 header {* Representable domains *}
       
     6 
       
     7 theory Representable
       
     8 imports Algebraic Map_Functions Countable
       
     9 begin
       
    10 
       
    11 subsection {* Class of representable domains *}
       
    12 
       
    13 text {*
       
    14   We define a ``domain'' as a pcpo that is isomorphic to some
       
    15   algebraic deflation over the universal domain; this is equivalent
       
    16   to being omega-bifinite.
       
    17 
       
    18   A predomain is a cpo that, when lifted, becomes a domain.
       
    19 *}
       
    20 
       
    21 class predomain = cpo +
       
    22   fixes liftdefl :: "('a::cpo) itself \<Rightarrow> defl"
       
    23   fixes liftemb :: "'a\<^sub>\<bottom> \<rightarrow> udom"
       
    24   fixes liftprj :: "udom \<rightarrow> 'a\<^sub>\<bottom>"
       
    25   assumes predomain_ep: "ep_pair liftemb liftprj"
       
    26   assumes cast_liftdefl: "cast\<cdot>(liftdefl TYPE('a::cpo)) = liftemb oo liftprj"
       
    27 
       
    28 syntax "_LIFTDEFL" :: "type \<Rightarrow> logic"  ("(1LIFTDEFL/(1'(_')))")
       
    29 translations "LIFTDEFL('t)" \<rightleftharpoons> "CONST liftdefl TYPE('t)"
       
    30 
       
    31 class "domain" = predomain + pcpo +
       
    32   fixes emb :: "'a::cpo \<rightarrow> udom"
       
    33   fixes prj :: "udom \<rightarrow> 'a::cpo"
       
    34   fixes defl :: "'a itself \<Rightarrow> defl"
       
    35   assumes ep_pair_emb_prj: "ep_pair emb prj"
       
    36   assumes cast_DEFL: "cast\<cdot>(defl TYPE('a)) = emb oo prj"
       
    37 
       
    38 syntax "_DEFL" :: "type \<Rightarrow> defl"  ("(1DEFL/(1'(_')))")
       
    39 translations "DEFL('t)" \<rightleftharpoons> "CONST defl TYPE('t)"
       
    40 
       
    41 interpretation "domain": pcpo_ep_pair emb prj
       
    42   unfolding pcpo_ep_pair_def
       
    43   by (rule ep_pair_emb_prj)
       
    44 
       
    45 lemmas emb_inverse = domain.e_inverse
       
    46 lemmas emb_prj_below = domain.e_p_below
       
    47 lemmas emb_eq_iff = domain.e_eq_iff
       
    48 lemmas emb_strict = domain.e_strict
       
    49 lemmas prj_strict = domain.p_strict
       
    50 
       
    51 subsection {* Domains have a countable compact basis *}
       
    52 
       
    53 text {*
       
    54   Eventually it should be possible to generalize this to an unpointed
       
    55   variant of the domain class.
       
    56 *}
       
    57 
       
    58 interpretation compact_basis:
       
    59   ideal_completion below Rep_compact_basis "approximants::'a::domain \<Rightarrow> _"
       
    60 proof -
       
    61   obtain Y where Y: "\<forall>i. Y i \<sqsubseteq> Y (Suc i)"
       
    62   and DEFL: "DEFL('a) = (\<Squnion>i. defl_principal (Y i))"
       
    63     by (rule defl.obtain_principal_chain)
       
    64   def approx \<equiv> "\<lambda>i. (prj oo cast\<cdot>(defl_principal (Y i)) oo emb) :: 'a \<rightarrow> 'a"
       
    65   interpret defl_approx: approx_chain approx
       
    66   proof (rule approx_chain.intro)
       
    67     show "chain (\<lambda>i. approx i)"
       
    68       unfolding approx_def by (simp add: Y)
       
    69     show "(\<Squnion>i. approx i) = ID"
       
    70       unfolding approx_def
       
    71       by (simp add: lub_distribs Y DEFL [symmetric] cast_DEFL cfun_eq_iff)
       
    72     show "\<And>i. finite_deflation (approx i)"
       
    73       unfolding approx_def
       
    74       apply (rule domain.finite_deflation_p_d_e)
       
    75       apply (rule finite_deflation_cast)
       
    76       apply (rule defl.compact_principal)
       
    77       apply (rule below_trans [OF monofun_cfun_fun])
       
    78       apply (rule is_ub_thelub, simp add: Y)
       
    79       apply (simp add: lub_distribs Y DEFL [symmetric] cast_DEFL)
       
    80       done
       
    81   qed
       
    82   (* FIXME: why does show ?thesis fail here? *)
       
    83   show "ideal_completion below Rep_compact_basis (approximants::'a \<Rightarrow> _)" ..
       
    84 qed
       
    85 
       
    86 subsection {* Chains of approx functions *}
       
    87 
       
    88 definition u_approx :: "nat \<Rightarrow> udom\<^sub>\<bottom> \<rightarrow> udom\<^sub>\<bottom>"
       
    89   where "u_approx = (\<lambda>i. u_map\<cdot>(udom_approx i))"
       
    90 
       
    91 definition sfun_approx :: "nat \<Rightarrow> (udom \<rightarrow>! udom) \<rightarrow> (udom \<rightarrow>! udom)"
       
    92   where "sfun_approx = (\<lambda>i. sfun_map\<cdot>(udom_approx i)\<cdot>(udom_approx i))"
       
    93 
       
    94 definition prod_approx :: "nat \<Rightarrow> udom \<times> udom \<rightarrow> udom \<times> udom"
       
    95   where "prod_approx = (\<lambda>i. cprod_map\<cdot>(udom_approx i)\<cdot>(udom_approx i))"
       
    96 
       
    97 definition sprod_approx :: "nat \<Rightarrow> udom \<otimes> udom \<rightarrow> udom \<otimes> udom"
       
    98   where "sprod_approx = (\<lambda>i. sprod_map\<cdot>(udom_approx i)\<cdot>(udom_approx i))"
       
    99 
       
   100 definition ssum_approx :: "nat \<Rightarrow> udom \<oplus> udom \<rightarrow> udom \<oplus> udom"
       
   101   where "ssum_approx = (\<lambda>i. ssum_map\<cdot>(udom_approx i)\<cdot>(udom_approx i))"
       
   102 
       
   103 lemma approx_chain_lemma1:
       
   104   assumes "m\<cdot>ID = ID"
       
   105   assumes "\<And>d. finite_deflation d \<Longrightarrow> finite_deflation (m\<cdot>d)"
       
   106   shows "approx_chain (\<lambda>i. m\<cdot>(udom_approx i))"
       
   107 by (rule approx_chain.intro)
       
   108    (simp_all add: lub_distribs finite_deflation_udom_approx assms)
       
   109 
       
   110 lemma approx_chain_lemma2:
       
   111   assumes "m\<cdot>ID\<cdot>ID = ID"
       
   112   assumes "\<And>a b. \<lbrakk>finite_deflation a; finite_deflation b\<rbrakk>
       
   113     \<Longrightarrow> finite_deflation (m\<cdot>a\<cdot>b)"
       
   114   shows "approx_chain (\<lambda>i. m\<cdot>(udom_approx i)\<cdot>(udom_approx i))"
       
   115 by (rule approx_chain.intro)
       
   116    (simp_all add: lub_distribs finite_deflation_udom_approx assms)
       
   117 
       
   118 lemma u_approx: "approx_chain u_approx"
       
   119 using u_map_ID finite_deflation_u_map
       
   120 unfolding u_approx_def by (rule approx_chain_lemma1)
       
   121 
       
   122 lemma sfun_approx: "approx_chain sfun_approx"
       
   123 using sfun_map_ID finite_deflation_sfun_map
       
   124 unfolding sfun_approx_def by (rule approx_chain_lemma2)
       
   125 
       
   126 lemma prod_approx: "approx_chain prod_approx"
       
   127 using cprod_map_ID finite_deflation_cprod_map
       
   128 unfolding prod_approx_def by (rule approx_chain_lemma2)
       
   129 
       
   130 lemma sprod_approx: "approx_chain sprod_approx"
       
   131 using sprod_map_ID finite_deflation_sprod_map
       
   132 unfolding sprod_approx_def by (rule approx_chain_lemma2)
       
   133 
       
   134 lemma ssum_approx: "approx_chain ssum_approx"
       
   135 using ssum_map_ID finite_deflation_ssum_map
       
   136 unfolding ssum_approx_def by (rule approx_chain_lemma2)
       
   137 
       
   138 subsection {* Type combinators *}
       
   139 
       
   140 definition
       
   141   defl_fun1 ::
       
   142     "(nat \<Rightarrow> 'a \<rightarrow> 'a) \<Rightarrow> ((udom \<rightarrow> udom) \<rightarrow> ('a \<rightarrow> 'a)) \<Rightarrow> (defl \<rightarrow> defl)"
       
   143 where
       
   144   "defl_fun1 approx f =
       
   145     defl.basis_fun (\<lambda>a.
       
   146       defl_principal (Abs_fin_defl
       
   147         (udom_emb approx oo f\<cdot>(Rep_fin_defl a) oo udom_prj approx)))"
       
   148 
       
   149 definition
       
   150   defl_fun2 ::
       
   151     "(nat \<Rightarrow> 'a \<rightarrow> 'a) \<Rightarrow> ((udom \<rightarrow> udom) \<rightarrow> (udom \<rightarrow> udom) \<rightarrow> ('a \<rightarrow> 'a))
       
   152       \<Rightarrow> (defl \<rightarrow> defl \<rightarrow> defl)"
       
   153 where
       
   154   "defl_fun2 approx f =
       
   155     defl.basis_fun (\<lambda>a.
       
   156       defl.basis_fun (\<lambda>b.
       
   157         defl_principal (Abs_fin_defl
       
   158           (udom_emb approx oo
       
   159             f\<cdot>(Rep_fin_defl a)\<cdot>(Rep_fin_defl b) oo udom_prj approx))))"
       
   160 
       
   161 lemma cast_defl_fun1:
       
   162   assumes approx: "approx_chain approx"
       
   163   assumes f: "\<And>a. finite_deflation a \<Longrightarrow> finite_deflation (f\<cdot>a)"
       
   164   shows "cast\<cdot>(defl_fun1 approx f\<cdot>A) = udom_emb approx oo f\<cdot>(cast\<cdot>A) oo udom_prj approx"
       
   165 proof -
       
   166   have 1: "\<And>a. finite_deflation
       
   167         (udom_emb approx oo f\<cdot>(Rep_fin_defl a) oo udom_prj approx)"
       
   168     apply (rule ep_pair.finite_deflation_e_d_p)
       
   169     apply (rule approx_chain.ep_pair_udom [OF approx])
       
   170     apply (rule f, rule finite_deflation_Rep_fin_defl)
       
   171     done
       
   172   show ?thesis
       
   173     by (induct A rule: defl.principal_induct, simp)
       
   174        (simp only: defl_fun1_def
       
   175                    defl.basis_fun_principal
       
   176                    defl.basis_fun_mono
       
   177                    defl.principal_mono
       
   178                    Abs_fin_defl_mono [OF 1 1]
       
   179                    monofun_cfun below_refl
       
   180                    Rep_fin_defl_mono
       
   181                    cast_defl_principal
       
   182                    Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1])
       
   183 qed
       
   184 
       
   185 lemma cast_defl_fun2:
       
   186   assumes approx: "approx_chain approx"
       
   187   assumes f: "\<And>a b. finite_deflation a \<Longrightarrow> finite_deflation b \<Longrightarrow>
       
   188                 finite_deflation (f\<cdot>a\<cdot>b)"
       
   189   shows "cast\<cdot>(defl_fun2 approx f\<cdot>A\<cdot>B) =
       
   190     udom_emb approx oo f\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj approx"
       
   191 proof -
       
   192   have 1: "\<And>a b. finite_deflation (udom_emb approx oo
       
   193       f\<cdot>(Rep_fin_defl a)\<cdot>(Rep_fin_defl b) oo udom_prj approx)"
       
   194     apply (rule ep_pair.finite_deflation_e_d_p)
       
   195     apply (rule ep_pair_udom [OF approx])
       
   196     apply (rule f, (rule finite_deflation_Rep_fin_defl)+)
       
   197     done
       
   198   show ?thesis
       
   199     by (induct A B rule: defl.principal_induct2, simp, simp)
       
   200        (simp only: defl_fun2_def
       
   201                    defl.basis_fun_principal
       
   202                    defl.basis_fun_mono
       
   203                    defl.principal_mono
       
   204                    Abs_fin_defl_mono [OF 1 1]
       
   205                    monofun_cfun below_refl
       
   206                    Rep_fin_defl_mono
       
   207                    cast_defl_principal
       
   208                    Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1])
       
   209 qed
       
   210 
       
   211 definition u_defl :: "defl \<rightarrow> defl"
       
   212   where "u_defl = defl_fun1 u_approx u_map"
       
   213 
       
   214 definition sfun_defl :: "defl \<rightarrow> defl \<rightarrow> defl"
       
   215   where "sfun_defl = defl_fun2 sfun_approx sfun_map"
       
   216 
       
   217 definition prod_defl :: "defl \<rightarrow> defl \<rightarrow> defl"
       
   218   where "prod_defl = defl_fun2 prod_approx cprod_map"
       
   219 
       
   220 definition sprod_defl :: "defl \<rightarrow> defl \<rightarrow> defl"
       
   221   where "sprod_defl = defl_fun2 sprod_approx sprod_map"
       
   222 
       
   223 definition ssum_defl :: "defl \<rightarrow> defl \<rightarrow> defl"
       
   224 where "ssum_defl = defl_fun2 ssum_approx ssum_map"
       
   225 
       
   226 lemma cast_u_defl:
       
   227   "cast\<cdot>(u_defl\<cdot>A) =
       
   228     udom_emb u_approx oo u_map\<cdot>(cast\<cdot>A) oo udom_prj u_approx"
       
   229 using u_approx finite_deflation_u_map
       
   230 unfolding u_defl_def by (rule cast_defl_fun1)
       
   231 
       
   232 lemma cast_sfun_defl:
       
   233   "cast\<cdot>(sfun_defl\<cdot>A\<cdot>B) =
       
   234     udom_emb sfun_approx oo sfun_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj sfun_approx"
       
   235 using sfun_approx finite_deflation_sfun_map
       
   236 unfolding sfun_defl_def by (rule cast_defl_fun2)
       
   237 
       
   238 lemma cast_prod_defl:
       
   239   "cast\<cdot>(prod_defl\<cdot>A\<cdot>B) = udom_emb prod_approx oo
       
   240     cprod_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj prod_approx"
       
   241 using prod_approx finite_deflation_cprod_map
       
   242 unfolding prod_defl_def by (rule cast_defl_fun2)
       
   243 
       
   244 lemma cast_sprod_defl:
       
   245   "cast\<cdot>(sprod_defl\<cdot>A\<cdot>B) =
       
   246     udom_emb sprod_approx oo
       
   247       sprod_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo
       
   248         udom_prj sprod_approx"
       
   249 using sprod_approx finite_deflation_sprod_map
       
   250 unfolding sprod_defl_def by (rule cast_defl_fun2)
       
   251 
       
   252 lemma cast_ssum_defl:
       
   253   "cast\<cdot>(ssum_defl\<cdot>A\<cdot>B) =
       
   254     udom_emb ssum_approx oo ssum_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj ssum_approx"
       
   255 using ssum_approx finite_deflation_ssum_map
       
   256 unfolding ssum_defl_def by (rule cast_defl_fun2)
       
   257 
       
   258 subsection {* Lemma for proving domain instances *}
       
   259 
       
   260 text {*
       
   261   A class of domains where @{const liftemb}, @{const liftprj},
       
   262   and @{const liftdefl} are all defined in the standard way.
       
   263 *}
       
   264 
       
   265 class liftdomain = "domain" +
       
   266   assumes liftemb_eq: "liftemb = udom_emb u_approx oo u_map\<cdot>emb"
       
   267   assumes liftprj_eq: "liftprj = u_map\<cdot>prj oo udom_prj u_approx"
       
   268   assumes liftdefl_eq: "liftdefl TYPE('a::cpo) = u_defl\<cdot>DEFL('a)"
       
   269 
       
   270 text {* Temporarily relax type constraints. *}
       
   271 
       
   272 setup {*
       
   273   fold Sign.add_const_constraint
       
   274   [ (@{const_name defl}, SOME @{typ "'a::pcpo itself \<Rightarrow> defl"})
       
   275   , (@{const_name emb}, SOME @{typ "'a::pcpo \<rightarrow> udom"})
       
   276   , (@{const_name prj}, SOME @{typ "udom \<rightarrow> 'a::pcpo"})
       
   277   , (@{const_name liftdefl}, SOME @{typ "'a::pcpo itself \<Rightarrow> defl"})
       
   278   , (@{const_name liftemb}, SOME @{typ "'a::pcpo u \<rightarrow> udom"})
       
   279   , (@{const_name liftprj}, SOME @{typ "udom \<rightarrow> 'a::pcpo u"}) ]
       
   280 *}
       
   281 
       
   282 lemma liftdomain_class_intro:
       
   283   assumes liftemb: "(liftemb :: 'a u \<rightarrow> udom) = udom_emb u_approx oo u_map\<cdot>emb"
       
   284   assumes liftprj: "(liftprj :: udom \<rightarrow> 'a u) = u_map\<cdot>prj oo udom_prj u_approx"
       
   285   assumes liftdefl: "liftdefl TYPE('a) = u_defl\<cdot>DEFL('a)"
       
   286   assumes ep_pair: "ep_pair emb (prj :: udom \<rightarrow> 'a)"
       
   287   assumes cast_defl: "cast\<cdot>DEFL('a) = emb oo (prj :: udom \<rightarrow> 'a)"
       
   288   shows "OFCLASS('a, liftdomain_class)"
       
   289 proof
       
   290   show "ep_pair liftemb (liftprj :: udom \<rightarrow> 'a u)"
       
   291     unfolding liftemb liftprj
       
   292     by (intro ep_pair_comp ep_pair_u_map ep_pair ep_pair_udom u_approx)
       
   293   show "cast\<cdot>LIFTDEFL('a) = liftemb oo (liftprj :: udom \<rightarrow> 'a u)"
       
   294     unfolding liftemb liftprj liftdefl
       
   295     by (simp add: cfcomp1 cast_u_defl cast_defl u_map_map)
       
   296 next
       
   297 qed fact+
       
   298 
       
   299 text {* Restore original type constraints. *}
       
   300 
       
   301 setup {*
       
   302   fold Sign.add_const_constraint
       
   303   [ (@{const_name defl}, SOME @{typ "'a::domain itself \<Rightarrow> defl"})
       
   304   , (@{const_name emb}, SOME @{typ "'a::domain \<rightarrow> udom"})
       
   305   , (@{const_name prj}, SOME @{typ "udom \<rightarrow> 'a::domain"})
       
   306   , (@{const_name liftdefl}, SOME @{typ "'a::predomain itself \<Rightarrow> defl"})
       
   307   , (@{const_name liftemb}, SOME @{typ "'a::predomain u \<rightarrow> udom"})
       
   308   , (@{const_name liftprj}, SOME @{typ "udom \<rightarrow> 'a::predomain u"}) ]
       
   309 *}
       
   310 
       
   311 subsection {* Class instance proofs *}
       
   312 
       
   313 subsubsection {* Universal domain *}
       
   314 
       
   315 instantiation udom :: liftdomain
       
   316 begin
       
   317 
       
   318 definition [simp]:
       
   319   "emb = (ID :: udom \<rightarrow> udom)"
       
   320 
       
   321 definition [simp]:
       
   322   "prj = (ID :: udom \<rightarrow> udom)"
       
   323 
       
   324 definition
       
   325   "defl (t::udom itself) = (\<Squnion>i. defl_principal (Abs_fin_defl (udom_approx i)))"
       
   326 
       
   327 definition
       
   328   "(liftemb :: udom u \<rightarrow> udom) = udom_emb u_approx oo u_map\<cdot>emb"
       
   329 
       
   330 definition
       
   331   "(liftprj :: udom \<rightarrow> udom u) = u_map\<cdot>prj oo udom_prj u_approx"
       
   332 
       
   333 definition
       
   334   "liftdefl (t::udom itself) = u_defl\<cdot>DEFL(udom)"
       
   335 
       
   336 instance
       
   337 using liftemb_udom_def liftprj_udom_def liftdefl_udom_def
       
   338 proof (rule liftdomain_class_intro)
       
   339   show "ep_pair emb (prj :: udom \<rightarrow> udom)"
       
   340     by (simp add: ep_pair.intro)
       
   341   show "cast\<cdot>DEFL(udom) = emb oo (prj :: udom \<rightarrow> udom)"
       
   342     unfolding defl_udom_def
       
   343     apply (subst contlub_cfun_arg)
       
   344     apply (rule chainI)
       
   345     apply (rule defl.principal_mono)
       
   346     apply (simp add: below_fin_defl_def)
       
   347     apply (simp add: Abs_fin_defl_inverse finite_deflation_udom_approx)
       
   348     apply (rule chainE)
       
   349     apply (rule chain_udom_approx)
       
   350     apply (subst cast_defl_principal)
       
   351     apply (simp add: Abs_fin_defl_inverse finite_deflation_udom_approx)
       
   352     done
       
   353 qed
       
   354 
       
   355 end
       
   356 
       
   357 subsubsection {* Lifted cpo *}
       
   358 
       
   359 instantiation u :: (predomain) liftdomain
       
   360 begin
       
   361 
       
   362 definition
       
   363   "emb = liftemb"
       
   364 
       
   365 definition
       
   366   "prj = liftprj"
       
   367 
       
   368 definition
       
   369   "defl (t::'a u itself) = LIFTDEFL('a)"
       
   370 
       
   371 definition
       
   372   "(liftemb :: 'a u u \<rightarrow> udom) = udom_emb u_approx oo u_map\<cdot>emb"
       
   373 
       
   374 definition
       
   375   "(liftprj :: udom \<rightarrow> 'a u u) = u_map\<cdot>prj oo udom_prj u_approx"
       
   376 
       
   377 definition
       
   378   "liftdefl (t::'a u itself) = u_defl\<cdot>DEFL('a u)"
       
   379 
       
   380 instance
       
   381 using liftemb_u_def liftprj_u_def liftdefl_u_def
       
   382 proof (rule liftdomain_class_intro)
       
   383   show "ep_pair emb (prj :: udom \<rightarrow> 'a u)"
       
   384     unfolding emb_u_def prj_u_def
       
   385     by (rule predomain_ep)
       
   386   show "cast\<cdot>DEFL('a u) = emb oo (prj :: udom \<rightarrow> 'a u)"
       
   387     unfolding emb_u_def prj_u_def defl_u_def
       
   388     by (rule cast_liftdefl)
       
   389 qed
       
   390 
       
   391 end
       
   392 
       
   393 lemma DEFL_u: "DEFL('a::predomain u) = LIFTDEFL('a)"
       
   394 by (rule defl_u_def)
       
   395 
       
   396 subsubsection {* Strict function space *}
       
   397 
       
   398 instantiation sfun :: ("domain", "domain") liftdomain
       
   399 begin
       
   400 
       
   401 definition
       
   402   "emb = udom_emb sfun_approx oo sfun_map\<cdot>prj\<cdot>emb"
       
   403 
       
   404 definition
       
   405   "prj = sfun_map\<cdot>emb\<cdot>prj oo udom_prj sfun_approx"
       
   406 
       
   407 definition
       
   408   "defl (t::('a \<rightarrow>! 'b) itself) = sfun_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
       
   409 
       
   410 definition
       
   411   "(liftemb :: ('a \<rightarrow>! 'b) u \<rightarrow> udom) = udom_emb u_approx oo u_map\<cdot>emb"
       
   412 
       
   413 definition
       
   414   "(liftprj :: udom \<rightarrow> ('a \<rightarrow>! 'b) u) = u_map\<cdot>prj oo udom_prj u_approx"
       
   415 
       
   416 definition
       
   417   "liftdefl (t::('a \<rightarrow>! 'b) itself) = u_defl\<cdot>DEFL('a \<rightarrow>! 'b)"
       
   418 
       
   419 instance
       
   420 using liftemb_sfun_def liftprj_sfun_def liftdefl_sfun_def
       
   421 proof (rule liftdomain_class_intro)
       
   422   show "ep_pair emb (prj :: udom \<rightarrow> 'a \<rightarrow>! 'b)"
       
   423     unfolding emb_sfun_def prj_sfun_def
       
   424     using ep_pair_udom [OF sfun_approx]
       
   425     by (intro ep_pair_comp ep_pair_sfun_map ep_pair_emb_prj)
       
   426   show "cast\<cdot>DEFL('a \<rightarrow>! 'b) = emb oo (prj :: udom \<rightarrow> 'a \<rightarrow>! 'b)"
       
   427     unfolding emb_sfun_def prj_sfun_def defl_sfun_def cast_sfun_defl
       
   428     by (simp add: cast_DEFL oo_def sfun_eq_iff sfun_map_map)
       
   429 qed
       
   430 
       
   431 end
       
   432 
       
   433 lemma DEFL_sfun:
       
   434   "DEFL('a::domain \<rightarrow>! 'b::domain) = sfun_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
       
   435 by (rule defl_sfun_def)
       
   436 
       
   437 subsubsection {* Continuous function space *}
       
   438 
       
   439 text {*
       
   440   Types @{typ "'a \<rightarrow> 'b"} and @{typ "'a u \<rightarrow>! 'b"} are isomorphic.
       
   441 *}
       
   442 
       
   443 definition
       
   444   "encode_cfun = (\<Lambda> f. sfun_abs\<cdot>(fup\<cdot>f))"
       
   445 
       
   446 definition
       
   447   "decode_cfun = (\<Lambda> g x. sfun_rep\<cdot>g\<cdot>(up\<cdot>x))"
       
   448 
       
   449 lemma decode_encode_cfun [simp]: "decode_cfun\<cdot>(encode_cfun\<cdot>x) = x"
       
   450 unfolding encode_cfun_def decode_cfun_def
       
   451 by (simp add: eta_cfun)
       
   452 
       
   453 lemma encode_decode_cfun [simp]: "encode_cfun\<cdot>(decode_cfun\<cdot>y) = y"
       
   454 unfolding encode_cfun_def decode_cfun_def
       
   455 apply (simp add: sfun_eq_iff strictify_cancel)
       
   456 apply (rule cfun_eqI, case_tac x, simp_all)
       
   457 done
       
   458 
       
   459 instantiation cfun :: (predomain, "domain") liftdomain
       
   460 begin
       
   461 
       
   462 definition
       
   463   "emb = emb oo encode_cfun"
       
   464 
       
   465 definition
       
   466   "prj = decode_cfun oo prj"
       
   467 
       
   468 definition
       
   469   "defl (t::('a \<rightarrow> 'b) itself) = DEFL('a u \<rightarrow>! 'b)"
       
   470 
       
   471 definition
       
   472   "(liftemb :: ('a \<rightarrow> 'b) u \<rightarrow> udom) = udom_emb u_approx oo u_map\<cdot>emb"
       
   473 
       
   474 definition
       
   475   "(liftprj :: udom \<rightarrow> ('a \<rightarrow> 'b) u) = u_map\<cdot>prj oo udom_prj u_approx"
       
   476 
       
   477 definition
       
   478   "liftdefl (t::('a \<rightarrow> 'b) itself) = u_defl\<cdot>DEFL('a \<rightarrow> 'b)"
       
   479 
       
   480 instance
       
   481 using liftemb_cfun_def liftprj_cfun_def liftdefl_cfun_def
       
   482 proof (rule liftdomain_class_intro)
       
   483   have "ep_pair encode_cfun decode_cfun"
       
   484     by (rule ep_pair.intro, simp_all)
       
   485   thus "ep_pair emb (prj :: udom \<rightarrow> 'a \<rightarrow> 'b)"
       
   486     unfolding emb_cfun_def prj_cfun_def
       
   487     using ep_pair_emb_prj by (rule ep_pair_comp)
       
   488   show "cast\<cdot>DEFL('a \<rightarrow> 'b) = emb oo (prj :: udom \<rightarrow> 'a \<rightarrow> 'b)"
       
   489     unfolding emb_cfun_def prj_cfun_def defl_cfun_def
       
   490     by (simp add: cast_DEFL cfcomp1)
       
   491 qed
       
   492 
       
   493 end
       
   494 
       
   495 lemma DEFL_cfun:
       
   496   "DEFL('a::predomain \<rightarrow> 'b::domain) = DEFL('a u \<rightarrow>! 'b)"
       
   497 by (rule defl_cfun_def)
       
   498 
       
   499 subsubsection {* Strict product *}
       
   500 
       
   501 instantiation sprod :: ("domain", "domain") liftdomain
       
   502 begin
       
   503 
       
   504 definition
       
   505   "emb = udom_emb sprod_approx oo sprod_map\<cdot>emb\<cdot>emb"
       
   506 
       
   507 definition
       
   508   "prj = sprod_map\<cdot>prj\<cdot>prj oo udom_prj sprod_approx"
       
   509 
       
   510 definition
       
   511   "defl (t::('a \<otimes> 'b) itself) = sprod_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
       
   512 
       
   513 definition
       
   514   "(liftemb :: ('a \<otimes> 'b) u \<rightarrow> udom) = udom_emb u_approx oo u_map\<cdot>emb"
       
   515 
       
   516 definition
       
   517   "(liftprj :: udom \<rightarrow> ('a \<otimes> 'b) u) = u_map\<cdot>prj oo udom_prj u_approx"
       
   518 
       
   519 definition
       
   520   "liftdefl (t::('a \<otimes> 'b) itself) = u_defl\<cdot>DEFL('a \<otimes> 'b)"
       
   521 
       
   522 instance
       
   523 using liftemb_sprod_def liftprj_sprod_def liftdefl_sprod_def
       
   524 proof (rule liftdomain_class_intro)
       
   525   show "ep_pair emb (prj :: udom \<rightarrow> 'a \<otimes> 'b)"
       
   526     unfolding emb_sprod_def prj_sprod_def
       
   527     using ep_pair_udom [OF sprod_approx]
       
   528     by (intro ep_pair_comp ep_pair_sprod_map ep_pair_emb_prj)
       
   529 next
       
   530   show "cast\<cdot>DEFL('a \<otimes> 'b) = emb oo (prj :: udom \<rightarrow> 'a \<otimes> 'b)"
       
   531     unfolding emb_sprod_def prj_sprod_def defl_sprod_def cast_sprod_defl
       
   532     by (simp add: cast_DEFL oo_def cfun_eq_iff sprod_map_map)
       
   533 qed
       
   534 
       
   535 end
       
   536 
       
   537 lemma DEFL_sprod:
       
   538   "DEFL('a::domain \<otimes> 'b::domain) = sprod_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
       
   539 by (rule defl_sprod_def)
       
   540 
       
   541 subsubsection {* Cartesian product *}
       
   542 
       
   543 text {*
       
   544   Types @{typ "('a * 'b) u"} and @{typ "'a u \<otimes> 'b u"} are isomorphic.
       
   545 *}
       
   546 
       
   547 definition
       
   548   "encode_prod_u = (\<Lambda>(up\<cdot>(x, y)). (:up\<cdot>x, up\<cdot>y:))"
       
   549 
       
   550 definition
       
   551   "decode_prod_u = (\<Lambda>(:up\<cdot>x, up\<cdot>y:). up\<cdot>(x, y))"
       
   552 
       
   553 lemma decode_encode_prod_u [simp]: "decode_prod_u\<cdot>(encode_prod_u\<cdot>x) = x"
       
   554 unfolding encode_prod_u_def decode_prod_u_def
       
   555 by (case_tac x, simp, rename_tac y, case_tac y, simp)
       
   556 
       
   557 lemma encode_decode_prod_u [simp]: "encode_prod_u\<cdot>(decode_prod_u\<cdot>y) = y"
       
   558 unfolding encode_prod_u_def decode_prod_u_def
       
   559 apply (case_tac y, simp, rename_tac a b)
       
   560 apply (case_tac a, simp, case_tac b, simp, simp)
       
   561 done
       
   562 
       
   563 instantiation prod :: (predomain, predomain) predomain
       
   564 begin
       
   565 
       
   566 definition
       
   567   "liftemb = emb oo encode_prod_u"
       
   568 
       
   569 definition
       
   570   "liftprj = decode_prod_u oo prj"
       
   571 
       
   572 definition
       
   573   "liftdefl (t::('a \<times> 'b) itself) = DEFL('a\<^sub>\<bottom> \<otimes> 'b\<^sub>\<bottom>)"
       
   574 
       
   575 instance proof
       
   576   have "ep_pair encode_prod_u decode_prod_u"
       
   577     by (rule ep_pair.intro, simp_all)
       
   578   thus "ep_pair liftemb (liftprj :: udom \<rightarrow> ('a \<times> 'b) u)"
       
   579     unfolding liftemb_prod_def liftprj_prod_def
       
   580     using ep_pair_emb_prj by (rule ep_pair_comp)
       
   581   show "cast\<cdot>LIFTDEFL('a \<times> 'b) = liftemb oo (liftprj :: udom \<rightarrow> ('a \<times> 'b) u)"
       
   582     unfolding liftemb_prod_def liftprj_prod_def liftdefl_prod_def
       
   583     by (simp add: cast_DEFL cfcomp1)
       
   584 qed
       
   585 
       
   586 end
       
   587 
       
   588 instantiation prod :: ("domain", "domain") "domain"
       
   589 begin
       
   590 
       
   591 definition
       
   592   "emb = udom_emb prod_approx oo cprod_map\<cdot>emb\<cdot>emb"
       
   593 
       
   594 definition
       
   595   "prj = cprod_map\<cdot>prj\<cdot>prj oo udom_prj prod_approx"
       
   596 
       
   597 definition
       
   598   "defl (t::('a \<times> 'b) itself) = prod_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
       
   599 
       
   600 instance proof
       
   601   show "ep_pair emb (prj :: udom \<rightarrow> 'a \<times> 'b)"
       
   602     unfolding emb_prod_def prj_prod_def
       
   603     using ep_pair_udom [OF prod_approx]
       
   604     by (intro ep_pair_comp ep_pair_cprod_map ep_pair_emb_prj)
       
   605 next
       
   606   show "cast\<cdot>DEFL('a \<times> 'b) = emb oo (prj :: udom \<rightarrow> 'a \<times> 'b)"
       
   607     unfolding emb_prod_def prj_prod_def defl_prod_def cast_prod_defl
       
   608     by (simp add: cast_DEFL oo_def cfun_eq_iff cprod_map_map)
       
   609 qed
       
   610 
       
   611 end
       
   612 
       
   613 lemma DEFL_prod:
       
   614   "DEFL('a::domain \<times> 'b::domain) = prod_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
       
   615 by (rule defl_prod_def)
       
   616 
       
   617 lemma LIFTDEFL_prod:
       
   618   "LIFTDEFL('a::predomain \<times> 'b::predomain) = DEFL('a u \<otimes> 'b u)"
       
   619 by (rule liftdefl_prod_def)
       
   620 
       
   621 subsubsection {* Unit type *}
       
   622 
       
   623 instantiation unit :: liftdomain
       
   624 begin
       
   625 
       
   626 definition
       
   627   "emb = (\<bottom> :: unit \<rightarrow> udom)"
       
   628 
       
   629 definition
       
   630   "prj = (\<bottom> :: udom \<rightarrow> unit)"
       
   631 
       
   632 definition
       
   633   "defl (t::unit itself) = \<bottom>"
       
   634 
       
   635 definition
       
   636   "(liftemb :: unit u \<rightarrow> udom) = udom_emb u_approx oo u_map\<cdot>emb"
       
   637 
       
   638 definition
       
   639   "(liftprj :: udom \<rightarrow> unit u) = u_map\<cdot>prj oo udom_prj u_approx"
       
   640 
       
   641 definition
       
   642   "liftdefl (t::unit itself) = u_defl\<cdot>DEFL(unit)"
       
   643 
       
   644 instance
       
   645 using liftemb_unit_def liftprj_unit_def liftdefl_unit_def
       
   646 proof (rule liftdomain_class_intro)
       
   647   show "ep_pair emb (prj :: udom \<rightarrow> unit)"
       
   648     unfolding emb_unit_def prj_unit_def
       
   649     by (simp add: ep_pair.intro)
       
   650 next
       
   651   show "cast\<cdot>DEFL(unit) = emb oo (prj :: udom \<rightarrow> unit)"
       
   652     unfolding emb_unit_def prj_unit_def defl_unit_def by simp
       
   653 qed
       
   654 
       
   655 end
       
   656 
       
   657 subsubsection {* Discrete cpo *}
       
   658 
       
   659 definition discr_approx :: "nat \<Rightarrow> 'a::countable discr u \<rightarrow> 'a discr u"
       
   660   where "discr_approx = (\<lambda>i. \<Lambda>(up\<cdot>x). if to_nat (undiscr x) < i then up\<cdot>x else \<bottom>)"
       
   661 
       
   662 lemma chain_discr_approx [simp]: "chain discr_approx"
       
   663 unfolding discr_approx_def
       
   664 by (rule chainI, simp add: monofun_cfun monofun_LAM)
       
   665 
       
   666 lemma lub_discr_approx [simp]: "(\<Squnion>i. discr_approx i) = ID"
       
   667 apply (rule cfun_eqI)
       
   668 apply (simp add: contlub_cfun_fun)
       
   669 apply (simp add: discr_approx_def)
       
   670 apply (case_tac x, simp)
       
   671 apply (rule lub_eqI)
       
   672 apply (rule is_lubI)
       
   673 apply (rule ub_rangeI, simp)
       
   674 apply (drule ub_rangeD)
       
   675 apply (erule rev_below_trans)
       
   676 apply simp
       
   677 apply (rule lessI)
       
   678 done
       
   679 
       
   680 lemma inj_on_undiscr [simp]: "inj_on undiscr A"
       
   681 using Discr_undiscr by (rule inj_on_inverseI)
       
   682 
       
   683 lemma finite_deflation_discr_approx: "finite_deflation (discr_approx i)"
       
   684 proof
       
   685   fix x :: "'a discr u"
       
   686   show "discr_approx i\<cdot>x \<sqsubseteq> x"
       
   687     unfolding discr_approx_def
       
   688     by (cases x, simp, simp)
       
   689   show "discr_approx i\<cdot>(discr_approx i\<cdot>x) = discr_approx i\<cdot>x"
       
   690     unfolding discr_approx_def
       
   691     by (cases x, simp, simp)
       
   692   show "finite {x::'a discr u. discr_approx i\<cdot>x = x}"
       
   693   proof (rule finite_subset)
       
   694     let ?S = "insert (\<bottom>::'a discr u) ((\<lambda>x. up\<cdot>x) ` undiscr -` to_nat -` {..<i})"
       
   695     show "{x::'a discr u. discr_approx i\<cdot>x = x} \<subseteq> ?S"
       
   696       unfolding discr_approx_def
       
   697       by (rule subsetI, case_tac x, simp, simp split: split_if_asm)
       
   698     show "finite ?S"
       
   699       by (simp add: finite_vimageI)
       
   700   qed
       
   701 qed
       
   702 
       
   703 lemma discr_approx: "approx_chain discr_approx"
       
   704 using chain_discr_approx lub_discr_approx finite_deflation_discr_approx
       
   705 by (rule approx_chain.intro)
       
   706 
       
   707 instantiation discr :: (countable) predomain
       
   708 begin
       
   709 
       
   710 definition
       
   711   "liftemb = udom_emb discr_approx"
       
   712 
       
   713 definition
       
   714   "liftprj = udom_prj discr_approx"
       
   715 
       
   716 definition
       
   717   "liftdefl (t::'a discr itself) =
       
   718     (\<Squnion>i. defl_principal (Abs_fin_defl (liftemb oo discr_approx i oo liftprj)))"
       
   719 
       
   720 instance proof
       
   721   show "ep_pair liftemb (liftprj :: udom \<rightarrow> 'a discr u)"
       
   722     unfolding liftemb_discr_def liftprj_discr_def
       
   723     by (rule ep_pair_udom [OF discr_approx])
       
   724   show "cast\<cdot>LIFTDEFL('a discr) = liftemb oo (liftprj :: udom \<rightarrow> 'a discr u)"
       
   725     unfolding liftemb_discr_def liftprj_discr_def liftdefl_discr_def
       
   726     apply (subst contlub_cfun_arg)
       
   727     apply (rule chainI)
       
   728     apply (rule defl.principal_mono)
       
   729     apply (simp add: below_fin_defl_def)
       
   730     apply (simp add: Abs_fin_defl_inverse
       
   731         ep_pair.finite_deflation_e_d_p [OF ep_pair_udom [OF discr_approx]]
       
   732         approx_chain.finite_deflation_approx [OF discr_approx])
       
   733     apply (intro monofun_cfun below_refl)
       
   734     apply (rule chainE)
       
   735     apply (rule chain_discr_approx)
       
   736     apply (subst cast_defl_principal)
       
   737     apply (simp add: Abs_fin_defl_inverse
       
   738         ep_pair.finite_deflation_e_d_p [OF ep_pair_udom [OF discr_approx]]
       
   739         approx_chain.finite_deflation_approx [OF discr_approx])
       
   740     apply (simp add: lub_distribs)
       
   741     done
       
   742 qed
       
   743 
       
   744 end
       
   745 
       
   746 subsubsection {* Strict sum *}
       
   747 
       
   748 instantiation ssum :: ("domain", "domain") liftdomain
       
   749 begin
       
   750 
       
   751 definition
       
   752   "emb = udom_emb ssum_approx oo ssum_map\<cdot>emb\<cdot>emb"
       
   753 
       
   754 definition
       
   755   "prj = ssum_map\<cdot>prj\<cdot>prj oo udom_prj ssum_approx"
       
   756 
       
   757 definition
       
   758   "defl (t::('a \<oplus> 'b) itself) = ssum_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
       
   759 
       
   760 definition
       
   761   "(liftemb :: ('a \<oplus> 'b) u \<rightarrow> udom) = udom_emb u_approx oo u_map\<cdot>emb"
       
   762 
       
   763 definition
       
   764   "(liftprj :: udom \<rightarrow> ('a \<oplus> 'b) u) = u_map\<cdot>prj oo udom_prj u_approx"
       
   765 
       
   766 definition
       
   767   "liftdefl (t::('a \<oplus> 'b) itself) = u_defl\<cdot>DEFL('a \<oplus> 'b)"
       
   768 
       
   769 instance
       
   770 using liftemb_ssum_def liftprj_ssum_def liftdefl_ssum_def
       
   771 proof (rule liftdomain_class_intro)
       
   772   show "ep_pair emb (prj :: udom \<rightarrow> 'a \<oplus> 'b)"
       
   773     unfolding emb_ssum_def prj_ssum_def
       
   774     using ep_pair_udom [OF ssum_approx]
       
   775     by (intro ep_pair_comp ep_pair_ssum_map ep_pair_emb_prj)
       
   776   show "cast\<cdot>DEFL('a \<oplus> 'b) = emb oo (prj :: udom \<rightarrow> 'a \<oplus> 'b)"
       
   777     unfolding emb_ssum_def prj_ssum_def defl_ssum_def cast_ssum_defl
       
   778     by (simp add: cast_DEFL oo_def cfun_eq_iff ssum_map_map)
       
   779 qed
       
   780 
       
   781 end
       
   782 
       
   783 lemma DEFL_ssum:
       
   784   "DEFL('a::domain \<oplus> 'b::domain) = ssum_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
       
   785 by (rule defl_ssum_def)
       
   786 
       
   787 subsubsection {* Lifted HOL type *}
       
   788 
       
   789 instantiation lift :: (countable) liftdomain
       
   790 begin
       
   791 
       
   792 definition
       
   793   "emb = emb oo (\<Lambda> x. Rep_lift x)"
       
   794 
       
   795 definition
       
   796   "prj = (\<Lambda> y. Abs_lift y) oo prj"
       
   797 
       
   798 definition
       
   799   "defl (t::'a lift itself) = DEFL('a discr u)"
       
   800 
       
   801 definition
       
   802   "(liftemb :: 'a lift u \<rightarrow> udom) = udom_emb u_approx oo u_map\<cdot>emb"
       
   803 
       
   804 definition
       
   805   "(liftprj :: udom \<rightarrow> 'a lift u) = u_map\<cdot>prj oo udom_prj u_approx"
       
   806 
       
   807 definition
       
   808   "liftdefl (t::'a lift itself) = u_defl\<cdot>DEFL('a lift)"
       
   809 
       
   810 instance
       
   811 using liftemb_lift_def liftprj_lift_def liftdefl_lift_def
       
   812 proof (rule liftdomain_class_intro)
       
   813   note [simp] = cont_Rep_lift cont_Abs_lift Rep_lift_inverse Abs_lift_inverse
       
   814   have "ep_pair (\<Lambda>(x::'a lift). Rep_lift x) (\<Lambda> y. Abs_lift y)"
       
   815     by (simp add: ep_pair_def)
       
   816   thus "ep_pair emb (prj :: udom \<rightarrow> 'a lift)"
       
   817     unfolding emb_lift_def prj_lift_def
       
   818     using ep_pair_emb_prj by (rule ep_pair_comp)
       
   819   show "cast\<cdot>DEFL('a lift) = emb oo (prj :: udom \<rightarrow> 'a lift)"
       
   820     unfolding emb_lift_def prj_lift_def defl_lift_def cast_DEFL
       
   821     by (simp add: cfcomp1)
       
   822 qed
       
   823 
       
   824 end
       
   825 
       
   826 end