src/HOL/HOLCF/Powerdomains.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 81577 a712bf5ccab0
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 41436
diff changeset
     1
(*  Title:      HOL/HOLCF/Powerdomains.thy
35473
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
     3
*)
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
     4
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
     5
section \<open>Powerdomains\<close>
35473
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
     6
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
     7
theory Powerdomains
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 39973
diff changeset
     8
imports ConvexPD Domain
35473
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
     9
begin
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
    10
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    11
subsection \<open>Universal domain embeddings\<close>
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    12
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    13
definition "upper_emb = udom_emb (\<lambda>i. upper_map\<cdot>(udom_approx i))"
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    14
definition "upper_prj = udom_prj (\<lambda>i. upper_map\<cdot>(udom_approx i))"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    15
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    16
definition "lower_emb = udom_emb (\<lambda>i. lower_map\<cdot>(udom_approx i))"
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    17
definition "lower_prj = udom_prj (\<lambda>i. lower_map\<cdot>(udom_approx i))"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    18
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    19
definition "convex_emb = udom_emb (\<lambda>i. convex_map\<cdot>(udom_approx i))"
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    20
definition "convex_prj = udom_prj (\<lambda>i. convex_map\<cdot>(udom_approx i))"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    21
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    22
lemma ep_pair_upper: "ep_pair upper_emb upper_prj"
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    23
  unfolding upper_emb_def upper_prj_def
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    24
  by (simp add: ep_pair_udom approx_chain_upper_map)
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    25
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    26
lemma ep_pair_lower: "ep_pair lower_emb lower_prj"
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    27
  unfolding lower_emb_def lower_prj_def
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    28
  by (simp add: ep_pair_udom approx_chain_lower_map)
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    29
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    30
lemma ep_pair_convex: "ep_pair convex_emb convex_prj"
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    31
  unfolding convex_emb_def convex_prj_def
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    32
  by (simp add: ep_pair_udom approx_chain_convex_map)
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    33
81577
a712bf5ccab0 tuned whitespace;
wenzelm
parents: 69597
diff changeset
    34
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    35
subsection \<open>Deflation combinators\<close>
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    36
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    37
definition upper_defl :: "udom defl \<rightarrow> udom defl"
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    38
  where "upper_defl = defl_fun1 upper_emb upper_prj upper_map"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    39
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    40
definition lower_defl :: "udom defl \<rightarrow> udom defl"
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    41
  where "lower_defl = defl_fun1 lower_emb lower_prj lower_map"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    42
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    43
definition convex_defl :: "udom defl \<rightarrow> udom defl"
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    44
  where "convex_defl = defl_fun1 convex_emb convex_prj convex_map"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    45
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    46
lemma cast_upper_defl:
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    47
  "cast\<cdot>(upper_defl\<cdot>A) = upper_emb oo upper_map\<cdot>(cast\<cdot>A) oo upper_prj"
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    48
using ep_pair_upper finite_deflation_upper_map
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    49
unfolding upper_defl_def by (rule cast_defl_fun1)
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    50
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    51
lemma cast_lower_defl:
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    52
  "cast\<cdot>(lower_defl\<cdot>A) = lower_emb oo lower_map\<cdot>(cast\<cdot>A) oo lower_prj"
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    53
using ep_pair_lower finite_deflation_lower_map
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    54
unfolding lower_defl_def by (rule cast_defl_fun1)
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    55
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    56
lemma cast_convex_defl:
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    57
  "cast\<cdot>(convex_defl\<cdot>A) = convex_emb oo convex_map\<cdot>(cast\<cdot>A) oo convex_prj"
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    58
using ep_pair_convex finite_deflation_convex_map
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    59
unfolding convex_defl_def by (rule cast_defl_fun1)
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    60
81577
a712bf5ccab0 tuned whitespace;
wenzelm
parents: 69597
diff changeset
    61
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    62
subsection \<open>Domain class instances\<close>
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    63
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
    64
instantiation upper_pd :: ("domain") "domain"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    65
begin
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    66
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    67
definition
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    68
  "emb = upper_emb oo upper_map\<cdot>emb"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    69
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    70
definition
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    71
  "prj = upper_map\<cdot>prj oo upper_prj"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    72
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    73
definition
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    74
  "defl (t::'a upper_pd itself) = upper_defl\<cdot>DEFL('a)"
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    75
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    76
definition
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
    77
  "(liftemb :: 'a upper_pd u \<rightarrow> udom u) = u_map\<cdot>emb"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    78
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    79
definition
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
    80
  "(liftprj :: udom u \<rightarrow> 'a upper_pd u) = u_map\<cdot>prj"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    81
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
    82
definition
41436
480978f80eae rename constant pdefl to liftdefl_of
huffman
parents: 41292
diff changeset
    83
  "liftdefl (t::'a upper_pd itself) = liftdefl_of\<cdot>DEFL('a upper_pd)"
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
    84
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
    85
instance proof
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    86
  show "ep_pair emb (prj :: udom \<rightarrow> 'a upper_pd)"
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    87
    unfolding emb_upper_pd_def prj_upper_pd_def
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
    88
    by (intro ep_pair_comp ep_pair_upper ep_pair_upper_map ep_pair_emb_prj)
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    89
next
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    90
  show "cast\<cdot>DEFL('a upper_pd) = emb oo (prj :: udom \<rightarrow> 'a upper_pd)"
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    91
    unfolding emb_upper_pd_def prj_upper_pd_def defl_upper_pd_def cast_upper_defl
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    92
    by (simp add: cast_DEFL oo_def cfun_eq_iff upper_map_map)
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
    93
qed (fact liftemb_upper_pd_def liftprj_upper_pd_def liftdefl_upper_pd_def)+
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    94
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    95
end
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    96
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
    97
instantiation lower_pd :: ("domain") "domain"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    98
begin
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
    99
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   100
definition
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
   101
  "emb = lower_emb oo lower_map\<cdot>emb"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   102
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   103
definition
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
   104
  "prj = lower_map\<cdot>prj oo lower_prj"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   105
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   106
definition
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   107
  "defl (t::'a lower_pd itself) = lower_defl\<cdot>DEFL('a)"
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   108
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   109
definition
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
   110
  "(liftemb :: 'a lower_pd u \<rightarrow> udom u) = u_map\<cdot>emb"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   111
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   112
definition
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
   113
  "(liftprj :: udom u \<rightarrow> 'a lower_pd u) = u_map\<cdot>prj"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   114
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
   115
definition
41436
480978f80eae rename constant pdefl to liftdefl_of
huffman
parents: 41292
diff changeset
   116
  "liftdefl (t::'a lower_pd itself) = liftdefl_of\<cdot>DEFL('a lower_pd)"
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
   117
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
   118
instance proof
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   119
  show "ep_pair emb (prj :: udom \<rightarrow> 'a lower_pd)"
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   120
    unfolding emb_lower_pd_def prj_lower_pd_def
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
   121
    by (intro ep_pair_comp ep_pair_lower ep_pair_lower_map ep_pair_emb_prj)
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   122
next
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   123
  show "cast\<cdot>DEFL('a lower_pd) = emb oo (prj :: udom \<rightarrow> 'a lower_pd)"
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   124
    unfolding emb_lower_pd_def prj_lower_pd_def defl_lower_pd_def cast_lower_defl
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   125
    by (simp add: cast_DEFL oo_def cfun_eq_iff lower_map_map)
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
   126
qed (fact liftemb_lower_pd_def liftprj_lower_pd_def liftdefl_lower_pd_def)+
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   127
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   128
end
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   129
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
   130
instantiation convex_pd :: ("domain") "domain"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   131
begin
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   132
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   133
definition
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
   134
  "emb = convex_emb oo convex_map\<cdot>emb"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   135
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   136
definition
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
   137
  "prj = convex_map\<cdot>prj oo convex_prj"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   138
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   139
definition
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   140
  "defl (t::'a convex_pd itself) = convex_defl\<cdot>DEFL('a)"
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   141
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   142
definition
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
   143
  "(liftemb :: 'a convex_pd u \<rightarrow> udom u) = u_map\<cdot>emb"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   144
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   145
definition
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
   146
  "(liftprj :: udom u \<rightarrow> 'a convex_pd u) = u_map\<cdot>prj"
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   147
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
   148
definition
41436
480978f80eae rename constant pdefl to liftdefl_of
huffman
parents: 41292
diff changeset
   149
  "liftdefl (t::'a convex_pd itself) = liftdefl_of\<cdot>DEFL('a convex_pd)"
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
   150
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
   151
instance proof
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   152
  show "ep_pair emb (prj :: udom \<rightarrow> 'a convex_pd)"
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   153
    unfolding emb_convex_pd_def prj_convex_pd_def
41290
e9c9577d88b5 replace foo_approx functions with foo_emb, foo_prj functions for universal domain embeddings
huffman
parents: 41289
diff changeset
   154
    by (intro ep_pair_comp ep_pair_convex ep_pair_convex_map ep_pair_emb_prj)
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   155
next
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   156
  show "cast\<cdot>DEFL('a convex_pd) = emb oo (prj :: udom \<rightarrow> 'a convex_pd)"
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   157
    unfolding emb_convex_pd_def prj_convex_pd_def defl_convex_pd_def cast_convex_defl
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   158
    by (simp add: cast_DEFL oo_def cfun_eq_iff convex_map_map)
41292
2b7bc8d9fd6e use deflations over type 'udom u' to represent predomains;
huffman
parents: 41290
diff changeset
   159
qed (fact liftemb_convex_pd_def liftprj_convex_pd_def liftdefl_convex_pd_def)+
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   160
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   161
end
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   162
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   163
lemma DEFL_upper: "DEFL('a::domain upper_pd) = upper_defl\<cdot>DEFL('a)"
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   164
by (rule defl_upper_pd_def)
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   165
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   166
lemma DEFL_lower: "DEFL('a::domain lower_pd) = lower_defl\<cdot>DEFL('a)"
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   167
by (rule defl_lower_pd_def)
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   168
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   169
lemma DEFL_convex: "DEFL('a::domain convex_pd) = convex_defl\<cdot>DEFL('a)"
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   170
by (rule defl_convex_pd_def)
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   171
81577
a712bf5ccab0 tuned whitespace;
wenzelm
parents: 69597
diff changeset
   172
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   173
subsection \<open>Isomorphic deflations\<close>
41289
f655912ac235 minimize imports; move domain class instances for powerdomain types into Powerdomains.thy
huffman
parents: 40774
diff changeset
   174
35473
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   175
lemma isodefl_upper:
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   176
  "isodefl d t \<Longrightarrow> isodefl (upper_map\<cdot>d) (upper_defl\<cdot>t)"
35473
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   177
apply (rule isodeflI)
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   178
apply (simp add: cast_upper_defl cast_isodefl)
35473
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   179
apply (simp add: emb_upper_pd_def prj_upper_pd_def)
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   180
apply (simp add: upper_map_map)
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   181
done
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   182
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   183
lemma isodefl_lower:
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   184
  "isodefl d t \<Longrightarrow> isodefl (lower_map\<cdot>d) (lower_defl\<cdot>t)"
35473
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   185
apply (rule isodeflI)
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   186
apply (simp add: cast_lower_defl cast_isodefl)
35473
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   187
apply (simp add: emb_lower_pd_def prj_lower_pd_def)
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   188
apply (simp add: lower_map_map)
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   189
done
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   190
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   191
lemma isodefl_convex:
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   192
  "isodefl d t \<Longrightarrow> isodefl (convex_map\<cdot>d) (convex_defl\<cdot>t)"
35473
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   193
apply (rule isodeflI)
39989
ad60d7311f43 renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents: 39974
diff changeset
   194
apply (simp add: cast_convex_defl cast_isodefl)
35473
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   195
apply (simp add: emb_convex_pd_def prj_convex_pd_def)
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   196
apply (simp add: convex_map_map)
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   197
done
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   198
81577
a712bf5ccab0 tuned whitespace;
wenzelm
parents: 69597
diff changeset
   199
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   200
subsection \<open>Domain package setup for powerdomains\<close>
35473
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   201
40216
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 39989
diff changeset
   202
lemmas [domain_defl_simps] = DEFL_upper DEFL_lower DEFL_convex
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 39989
diff changeset
   203
lemmas [domain_map_ID] = upper_map_ID lower_map_ID convex_map_ID
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 39989
diff changeset
   204
lemmas [domain_isodefl] = isodefl_upper isodefl_lower isodefl_convex
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 39989
diff changeset
   205
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 39989
diff changeset
   206
lemmas [domain_deflation] =
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 39989
diff changeset
   207
  deflation_upper_map deflation_lower_map deflation_convex_map
366309dfaf60 use Named_Thms instead of Theory_Data for some domain package theorems
huffman
parents: 39989
diff changeset
   208
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   209
setup \<open>
40737
2037021f034f remove map function names from domain package theory data
huffman
parents: 40487
diff changeset
   210
  fold Domain_Take_Proofs.add_rec_type
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62175
diff changeset
   211
    [(\<^type_name>\<open>upper_pd\<close>, [true]),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62175
diff changeset
   212
     (\<^type_name>\<open>lower_pd\<close>, [true]),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 62175
diff changeset
   213
     (\<^type_name>\<open>convex_pd\<close>, [true])]
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   214
\<close>
35473
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   215
c4d3d65856dd move some powerdomain stuff into a new file
huffman
parents:
diff changeset
   216
end