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