src/HOLCF/Cprod.thy
author huffman
Thu, 07 Oct 2010 13:54:43 -0700
changeset 39985 310f98585107
parent 39974 b525988432e9
child 39986 38677db30cad
permissions -rw-r--r--
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15600
a59f07556a8d fixed filename in header
huffman
parents: 15593
diff changeset
     1
(*  Title:      HOLCF/Cprod.thy
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     2
    Author:     Franz Regensburger
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     3
*)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     4
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     5
header {* The cpo of cartesian products *}
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     6
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     7
theory Cprod
39985
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
     8
imports Bifinite
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     9
begin
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    10
36452
d37c6eed8117 renamed command 'defaultsort' to 'default_sort';
wenzelm
parents: 35926
diff changeset
    11
default_sort cpo
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    12
35900
aa5dfb03eb1e remove LaTeX hyperref warnings by avoiding antiquotations within section headings
huffman
parents: 33399
diff changeset
    13
subsection {* Continuous case function for unit type *}
16008
861a255cf1e7 pcpo instance for type unit
huffman
parents: 15609
diff changeset
    14
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 18289
diff changeset
    15
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 18289
diff changeset
    16
  unit_when :: "'a \<rightarrow> unit \<rightarrow> 'a" where
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 18289
diff changeset
    17
  "unit_when = (\<Lambda> a _. a)"
16008
861a255cf1e7 pcpo instance for type unit
huffman
parents: 15609
diff changeset
    18
18289
56ddf617d6e8 add constant unit_when
huffman
parents: 18078
diff changeset
    19
translations
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 18289
diff changeset
    20
  "\<Lambda>(). t" == "CONST unit_when\<cdot>t"
18289
56ddf617d6e8 add constant unit_when
huffman
parents: 18078
diff changeset
    21
56ddf617d6e8 add constant unit_when
huffman
parents: 18078
diff changeset
    22
lemma unit_when [simp]: "unit_when\<cdot>a\<cdot>u = a"
56ddf617d6e8 add constant unit_when
huffman
parents: 18078
diff changeset
    23
by (simp add: unit_when_def)
56ddf617d6e8 add constant unit_when
huffman
parents: 18078
diff changeset
    24
35926
e6aec5d665f0 completely remove constants cpair, cfst, csnd
huffman
parents: 35922
diff changeset
    25
subsection {* Continuous version of split function *}
17834
28414668b43d added xsymbols syntax for pairs; cleaned up
huffman
parents: 17816
diff changeset
    26
25131
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 18289
diff changeset
    27
definition
2c8caac48ade modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents: 18289
diff changeset
    28
  csplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a * 'b) \<rightarrow> 'c" where
35922
bfa52a972745 define csplit using fst, snd
huffman
parents: 35900
diff changeset
    29
  "csplit = (\<Lambda> f p. f\<cdot>(fst p)\<cdot>(snd p))"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    30
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    31
translations
35926
e6aec5d665f0 completely remove constants cpair, cfst, csnd
huffman
parents: 35922
diff changeset
    32
  "\<Lambda>(CONST Pair x y). t" == "CONST csplit\<cdot>(\<Lambda> x y. t)"
17816
9942c5ed866a new syntax translations for continuous lambda abstraction
huffman
parents: 16916
diff changeset
    33
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    34
15593
24d770bbc44a reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    35
subsection {* Convert all lemmas to the continuous versions *}
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    36
18077
f1f4f951ec8d make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents: 17837
diff changeset
    37
lemma csplit1 [simp]: "csplit\<cdot>f\<cdot>\<bottom> = f\<cdot>\<bottom>\<cdot>\<bottom>"
f1f4f951ec8d make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents: 17837
diff changeset
    38
by (simp add: csplit_def)
f1f4f951ec8d make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents: 17837
diff changeset
    39
33399
768b2bb9e66a add (LAM (x, y). t) syntax and lemma csplit_Pair
huffman
parents: 31113
diff changeset
    40
lemma csplit_Pair [simp]: "csplit\<cdot>f\<cdot>(x, y) = f\<cdot>x\<cdot>y"
35922
bfa52a972745 define csplit using fst, snd
huffman
parents: 35900
diff changeset
    41
by (simp add: csplit_def)
33399
768b2bb9e66a add (LAM (x, y). t) syntax and lemma csplit_Pair
huffman
parents: 31113
diff changeset
    42
39985
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    43
subsection {* Map operator for product type *}
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    44
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    45
definition
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    46
  cprod_map :: "('a \<rightarrow> 'b) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> 'a \<times> 'c \<rightarrow> 'b \<times> 'd"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    47
where
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    48
  "cprod_map = (\<Lambda> f g p. (f\<cdot>(fst p), g\<cdot>(snd p)))"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    49
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    50
lemma cprod_map_Pair [simp]: "cprod_map\<cdot>f\<cdot>g\<cdot>(x, y) = (f\<cdot>x, g\<cdot>y)"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    51
unfolding cprod_map_def by simp
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    52
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    53
lemma cprod_map_ID: "cprod_map\<cdot>ID\<cdot>ID = ID"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    54
unfolding expand_cfun_eq by auto
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    55
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    56
lemma cprod_map_map:
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    57
  "cprod_map\<cdot>f1\<cdot>g1\<cdot>(cprod_map\<cdot>f2\<cdot>g2\<cdot>p) =
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    58
    cprod_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    59
by (induct p) simp
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    60
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    61
lemma ep_pair_cprod_map:
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    62
  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    63
  shows "ep_pair (cprod_map\<cdot>e1\<cdot>e2) (cprod_map\<cdot>p1\<cdot>p2)"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    64
proof
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    65
  interpret e1p1: ep_pair e1 p1 by fact
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    66
  interpret e2p2: ep_pair e2 p2 by fact
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    67
  fix x show "cprod_map\<cdot>p1\<cdot>p2\<cdot>(cprod_map\<cdot>e1\<cdot>e2\<cdot>x) = x"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    68
    by (induct x) simp
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    69
  fix y show "cprod_map\<cdot>e1\<cdot>e2\<cdot>(cprod_map\<cdot>p1\<cdot>p2\<cdot>y) \<sqsubseteq> y"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    70
    by (induct y) (simp add: e1p1.e_p_below e2p2.e_p_below)
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    71
qed
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    72
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    73
lemma deflation_cprod_map:
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    74
  assumes "deflation d1" and "deflation d2"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    75
  shows "deflation (cprod_map\<cdot>d1\<cdot>d2)"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    76
proof
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    77
  interpret d1: deflation d1 by fact
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    78
  interpret d2: deflation d2 by fact
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    79
  fix x
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    80
  show "cprod_map\<cdot>d1\<cdot>d2\<cdot>(cprod_map\<cdot>d1\<cdot>d2\<cdot>x) = cprod_map\<cdot>d1\<cdot>d2\<cdot>x"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    81
    by (induct x) (simp add: d1.idem d2.idem)
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    82
  show "cprod_map\<cdot>d1\<cdot>d2\<cdot>x \<sqsubseteq> x"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    83
    by (induct x) (simp add: d1.below d2.below)
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    84
qed
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    85
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    86
lemma finite_deflation_cprod_map:
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    87
  assumes "finite_deflation d1" and "finite_deflation d2"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    88
  shows "finite_deflation (cprod_map\<cdot>d1\<cdot>d2)"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    89
proof (rule finite_deflation_intro)
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    90
  interpret d1: finite_deflation d1 by fact
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    91
  interpret d2: finite_deflation d2 by fact
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    92
  have "deflation d1" and "deflation d2" by fact+
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    93
  thus "deflation (cprod_map\<cdot>d1\<cdot>d2)" by (rule deflation_cprod_map)
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    94
  have "{p. cprod_map\<cdot>d1\<cdot>d2\<cdot>p = p} \<subseteq> {x. d1\<cdot>x = x} \<times> {y. d2\<cdot>y = y}"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    95
    by clarsimp
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    96
  thus "finite {p. cprod_map\<cdot>d1\<cdot>d2\<cdot>p = p}"
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    97
    by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    98
qed
310f98585107 move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents: 39974
diff changeset
    99
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   100
subsection {* Cartesian product is an SFP domain *}
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   101
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   102
definition
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   103
  prod_approx :: "nat \<Rightarrow> udom \<times> udom \<rightarrow> udom \<times> udom"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   104
where
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   105
  "prod_approx = (\<lambda>i. cprod_map\<cdot>(udom_approx i)\<cdot>(udom_approx i))"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   106
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   107
lemma prod_approx: "approx_chain prod_approx"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   108
proof (rule approx_chain.intro)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   109
  show "chain (\<lambda>i. prod_approx i)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   110
    unfolding prod_approx_def by simp
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   111
  show "(\<Squnion>i. prod_approx i) = ID"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   112
    unfolding prod_approx_def
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   113
    by (simp add: lub_distribs cprod_map_ID)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   114
  show "\<And>i. finite_deflation (prod_approx i)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   115
    unfolding prod_approx_def
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   116
    by (intro finite_deflation_cprod_map finite_deflation_udom_approx)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   117
qed
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   118
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   119
definition prod_sfp :: "sfp \<rightarrow> sfp \<rightarrow> sfp"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   120
where "prod_sfp = sfp_fun2 prod_approx cprod_map"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   121
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   122
lemma cast_prod_sfp:
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   123
  "cast\<cdot>(prod_sfp\<cdot>A\<cdot>B) = udom_emb prod_approx oo
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   124
    cprod_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj prod_approx"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   125
unfolding prod_sfp_def
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   126
apply (rule cast_sfp_fun2 [OF prod_approx])
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   127
apply (erule (1) finite_deflation_cprod_map)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   128
done
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   129
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   130
instantiation prod :: (sfp, sfp) sfp
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   131
begin
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   132
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   133
definition
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   134
  "emb = udom_emb prod_approx oo cprod_map\<cdot>emb\<cdot>emb"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   135
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   136
definition
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   137
  "prj = cprod_map\<cdot>prj\<cdot>prj oo udom_prj prod_approx"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   138
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   139
definition
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   140
  "sfp (t::('a \<times> 'b) itself) = prod_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   141
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   142
instance proof
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   143
  show "ep_pair emb (prj :: udom \<rightarrow> 'a \<times> 'b)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   144
    unfolding emb_prod_def prj_prod_def
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   145
    using ep_pair_udom [OF prod_approx]
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   146
    by (intro ep_pair_comp ep_pair_cprod_map ep_pair_emb_prj)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   147
next
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   148
  show "cast\<cdot>SFP('a \<times> 'b) = emb oo (prj :: udom \<rightarrow> 'a \<times> 'b)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   149
    unfolding emb_prod_def prj_prod_def sfp_prod_def cast_prod_sfp
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   150
    by (simp add: cast_SFP oo_def expand_cfun_eq cprod_map_map)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   151
qed
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   152
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   153
end
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   154
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   155
lemma SFP_prod: "SFP('a::sfp \<times> 'b::sfp) = prod_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   156
by (rule sfp_prod_def)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   157
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   158
end