move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
authorhuffman
Wed, 10 Nov 2010 17:56:08 -0800
changeset 40502 8e92772bc0e8
parent 40501 2ed71459136e
child 40503 4094d788b904
move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
src/HOLCF/Algebraic.thy
src/HOLCF/Bifinite.thy
src/HOLCF/Cfun.thy
src/HOLCF/Completion.thy
src/HOLCF/Cprod.thy
src/HOLCF/Deflation.thy
src/HOLCF/Domain_Aux.thy
src/HOLCF/Fixrec.thy
src/HOLCF/IsaMakefile
src/HOLCF/Map_Functions.thy
src/HOLCF/Plain_HOLCF.thy
src/HOLCF/Sprod.thy
src/HOLCF/Ssum.thy
src/HOLCF/Up.thy
--- a/src/HOLCF/Algebraic.thy	Wed Nov 10 14:59:52 2010 -0800
+++ b/src/HOLCF/Algebraic.thy	Wed Nov 10 17:56:08 2010 -0800
@@ -5,7 +5,7 @@
 header {* Algebraic deflations *}
 
 theory Algebraic
-imports Universal
+imports Universal Map_Functions
 begin
 
 subsection {* Type constructor for finite deflations *}
--- a/src/HOLCF/Bifinite.thy	Wed Nov 10 14:59:52 2010 -0800
+++ b/src/HOLCF/Bifinite.thy	Wed Nov 10 17:56:08 2010 -0800
@@ -5,7 +5,7 @@
 header {* Bifinite domains *}
 
 theory Bifinite
-imports Algebraic Cprod Sprod Ssum Up Lift One Tr Countable
+imports Algebraic Map_Functions Countable
 begin
 
 subsection {* Class of bifinite domains *}
--- a/src/HOLCF/Cfun.thy	Wed Nov 10 14:59:52 2010 -0800
+++ b/src/HOLCF/Cfun.thy	Wed Nov 10 17:56:08 2010 -0800
@@ -479,24 +479,6 @@
 lemma assoc_oo: "f oo (g oo h) = (f oo g) oo h"
 by (rule cfun_eqI, simp)
 
-subsection {* Map operator for continuous function space *}
-
-definition
-  cfun_map :: "('b \<rightarrow> 'a) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> ('a \<rightarrow> 'c) \<rightarrow> ('b \<rightarrow> 'd)"
-where
-  "cfun_map = (\<Lambda> a b f x. b\<cdot>(f\<cdot>(a\<cdot>x)))"
-
-lemma cfun_map_beta [simp]: "cfun_map\<cdot>a\<cdot>b\<cdot>f\<cdot>x = b\<cdot>(f\<cdot>(a\<cdot>x))"
-unfolding cfun_map_def by simp
-
-lemma cfun_map_ID: "cfun_map\<cdot>ID\<cdot>ID = ID"
-unfolding cfun_eq_iff by simp
-
-lemma cfun_map_map:
-  "cfun_map\<cdot>f1\<cdot>g1\<cdot>(cfun_map\<cdot>f2\<cdot>g2\<cdot>p) =
-    cfun_map\<cdot>(\<Lambda> x. f2\<cdot>(f1\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
-by (rule cfun_eqI) simp
-
 subsection {* Strictified functions *}
 
 default_sort pcpo
--- a/src/HOLCF/Completion.thy	Wed Nov 10 14:59:52 2010 -0800
+++ b/src/HOLCF/Completion.thy	Wed Nov 10 17:56:08 2010 -0800
@@ -5,7 +5,7 @@
 header {* Defining algebraic domains by ideal completion *}
 
 theory Completion
-imports Cfun
+imports Plain_HOLCF
 begin
 
 subsection {* Ideals over a preorder *}
--- a/src/HOLCF/Cprod.thy	Wed Nov 10 14:59:52 2010 -0800
+++ b/src/HOLCF/Cprod.thy	Wed Nov 10 17:56:08 2010 -0800
@@ -5,7 +5,7 @@
 header {* The cpo of cartesian products *}
 
 theory Cprod
-imports Deflation
+imports Cfun
 begin
 
 default_sort cpo
@@ -40,61 +40,4 @@
 lemma csplit_Pair [simp]: "csplit\<cdot>f\<cdot>(x, y) = f\<cdot>x\<cdot>y"
 by (simp add: csplit_def)
 
-subsection {* Map operator for product type *}
-
-definition
-  cprod_map :: "('a \<rightarrow> 'b) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> 'a \<times> 'c \<rightarrow> 'b \<times> 'd"
-where
-  "cprod_map = (\<Lambda> f g p. (f\<cdot>(fst p), g\<cdot>(snd p)))"
-
-lemma cprod_map_Pair [simp]: "cprod_map\<cdot>f\<cdot>g\<cdot>(x, y) = (f\<cdot>x, g\<cdot>y)"
-unfolding cprod_map_def by simp
-
-lemma cprod_map_ID: "cprod_map\<cdot>ID\<cdot>ID = ID"
-unfolding cfun_eq_iff by auto
-
-lemma cprod_map_map:
-  "cprod_map\<cdot>f1\<cdot>g1\<cdot>(cprod_map\<cdot>f2\<cdot>g2\<cdot>p) =
-    cprod_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
-by (induct p) simp
-
-lemma ep_pair_cprod_map:
-  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
-  shows "ep_pair (cprod_map\<cdot>e1\<cdot>e2) (cprod_map\<cdot>p1\<cdot>p2)"
-proof
-  interpret e1p1: ep_pair e1 p1 by fact
-  interpret e2p2: ep_pair e2 p2 by fact
-  fix x show "cprod_map\<cdot>p1\<cdot>p2\<cdot>(cprod_map\<cdot>e1\<cdot>e2\<cdot>x) = x"
-    by (induct x) simp
-  fix y show "cprod_map\<cdot>e1\<cdot>e2\<cdot>(cprod_map\<cdot>p1\<cdot>p2\<cdot>y) \<sqsubseteq> y"
-    by (induct y) (simp add: e1p1.e_p_below e2p2.e_p_below)
-qed
-
-lemma deflation_cprod_map:
-  assumes "deflation d1" and "deflation d2"
-  shows "deflation (cprod_map\<cdot>d1\<cdot>d2)"
-proof
-  interpret d1: deflation d1 by fact
-  interpret d2: deflation d2 by fact
-  fix x
-  show "cprod_map\<cdot>d1\<cdot>d2\<cdot>(cprod_map\<cdot>d1\<cdot>d2\<cdot>x) = cprod_map\<cdot>d1\<cdot>d2\<cdot>x"
-    by (induct x) (simp add: d1.idem d2.idem)
-  show "cprod_map\<cdot>d1\<cdot>d2\<cdot>x \<sqsubseteq> x"
-    by (induct x) (simp add: d1.below d2.below)
-qed
-
-lemma finite_deflation_cprod_map:
-  assumes "finite_deflation d1" and "finite_deflation d2"
-  shows "finite_deflation (cprod_map\<cdot>d1\<cdot>d2)"
-proof (rule finite_deflation_intro)
-  interpret d1: finite_deflation d1 by fact
-  interpret d2: finite_deflation d2 by fact
-  have "deflation d1" and "deflation d2" by fact+
-  thus "deflation (cprod_map\<cdot>d1\<cdot>d2)" by (rule deflation_cprod_map)
-  have "{p. cprod_map\<cdot>d1\<cdot>d2\<cdot>p = p} \<subseteq> {x. d1\<cdot>x = x} \<times> {y. d2\<cdot>y = y}"
-    by clarsimp
-  thus "finite {p. cprod_map\<cdot>d1\<cdot>d2\<cdot>p = p}"
-    by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
-qed
-
 end
--- a/src/HOLCF/Deflation.thy	Wed Nov 10 14:59:52 2010 -0800
+++ b/src/HOLCF/Deflation.thy	Wed Nov 10 17:56:08 2010 -0800
@@ -5,7 +5,7 @@
 header {* Continuous deflations and ep-pairs *}
 
 theory Deflation
-imports Cfun
+imports Plain_HOLCF
 begin
 
 default_sort cpo
@@ -405,93 +405,4 @@
 
 end
 
-subsection {* Map operator for continuous functions *}
-
-lemma ep_pair_cfun_map:
-  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
-  shows "ep_pair (cfun_map\<cdot>p1\<cdot>e2) (cfun_map\<cdot>e1\<cdot>p2)"
-proof
-  interpret e1p1: ep_pair e1 p1 by fact
-  interpret e2p2: ep_pair e2 p2 by fact
-  fix f show "cfun_map\<cdot>e1\<cdot>p2\<cdot>(cfun_map\<cdot>p1\<cdot>e2\<cdot>f) = f"
-    by (simp add: cfun_eq_iff)
-  fix g show "cfun_map\<cdot>p1\<cdot>e2\<cdot>(cfun_map\<cdot>e1\<cdot>p2\<cdot>g) \<sqsubseteq> g"
-    apply (rule cfun_belowI, simp)
-    apply (rule below_trans [OF e2p2.e_p_below])
-    apply (rule monofun_cfun_arg)
-    apply (rule e1p1.e_p_below)
-    done
-qed
-
-lemma deflation_cfun_map:
-  assumes "deflation d1" and "deflation d2"
-  shows "deflation (cfun_map\<cdot>d1\<cdot>d2)"
-proof
-  interpret d1: deflation d1 by fact
-  interpret d2: deflation d2 by fact
-  fix f
-  show "cfun_map\<cdot>d1\<cdot>d2\<cdot>(cfun_map\<cdot>d1\<cdot>d2\<cdot>f) = cfun_map\<cdot>d1\<cdot>d2\<cdot>f"
-    by (simp add: cfun_eq_iff d1.idem d2.idem)
-  show "cfun_map\<cdot>d1\<cdot>d2\<cdot>f \<sqsubseteq> f"
-    apply (rule cfun_belowI, simp)
-    apply (rule below_trans [OF d2.below])
-    apply (rule monofun_cfun_arg)
-    apply (rule d1.below)
-    done
-qed
-
-lemma finite_range_cfun_map:
-  assumes a: "finite (range (\<lambda>x. a\<cdot>x))"
-  assumes b: "finite (range (\<lambda>y. b\<cdot>y))"
-  shows "finite (range (\<lambda>f. cfun_map\<cdot>a\<cdot>b\<cdot>f))"  (is "finite (range ?h)")
-proof (rule finite_imageD)
-  let ?f = "\<lambda>g. range (\<lambda>x. (a\<cdot>x, g\<cdot>x))"
-  show "finite (?f ` range ?h)"
-  proof (rule finite_subset)
-    let ?B = "Pow (range (\<lambda>x. a\<cdot>x) \<times> range (\<lambda>y. b\<cdot>y))"
-    show "?f ` range ?h \<subseteq> ?B"
-      by clarsimp
-    show "finite ?B"
-      by (simp add: a b)
-  qed
-  show "inj_on ?f (range ?h)"
-  proof (rule inj_onI, rule cfun_eqI, clarsimp)
-    fix x f g
-    assume "range (\<lambda>x. (a\<cdot>x, b\<cdot>(f\<cdot>(a\<cdot>x)))) = range (\<lambda>x. (a\<cdot>x, b\<cdot>(g\<cdot>(a\<cdot>x))))"
-    hence "range (\<lambda>x. (a\<cdot>x, b\<cdot>(f\<cdot>(a\<cdot>x)))) \<subseteq> range (\<lambda>x. (a\<cdot>x, b\<cdot>(g\<cdot>(a\<cdot>x))))"
-      by (rule equalityD1)
-    hence "(a\<cdot>x, b\<cdot>(f\<cdot>(a\<cdot>x))) \<in> range (\<lambda>x. (a\<cdot>x, b\<cdot>(g\<cdot>(a\<cdot>x))))"
-      by (simp add: subset_eq)
-    then obtain y where "(a\<cdot>x, b\<cdot>(f\<cdot>(a\<cdot>x))) = (a\<cdot>y, b\<cdot>(g\<cdot>(a\<cdot>y)))"
-      by (rule rangeE)
-    thus "b\<cdot>(f\<cdot>(a\<cdot>x)) = b\<cdot>(g\<cdot>(a\<cdot>x))"
-      by clarsimp
-  qed
-qed
-
-lemma finite_deflation_cfun_map:
-  assumes "finite_deflation d1" and "finite_deflation d2"
-  shows "finite_deflation (cfun_map\<cdot>d1\<cdot>d2)"
-proof (rule finite_deflation_intro)
-  interpret d1: finite_deflation d1 by fact
-  interpret d2: finite_deflation d2 by fact
-  have "deflation d1" and "deflation d2" by fact+
-  thus "deflation (cfun_map\<cdot>d1\<cdot>d2)" by (rule deflation_cfun_map)
-  have "finite (range (\<lambda>f. cfun_map\<cdot>d1\<cdot>d2\<cdot>f))"
-    using d1.finite_range d2.finite_range
-    by (rule finite_range_cfun_map)
-  thus "finite {f. cfun_map\<cdot>d1\<cdot>d2\<cdot>f = f}"
-    by (rule finite_range_imp_finite_fixes)
-qed
-
-text {* Finite deflations are compact elements of the function space *}
-
-lemma finite_deflation_imp_compact: "finite_deflation d \<Longrightarrow> compact d"
-apply (frule finite_deflation_imp_deflation)
-apply (subgoal_tac "compact (cfun_map\<cdot>d\<cdot>d\<cdot>d)")
-apply (simp add: cfun_map_def deflation.idem eta_cfun)
-apply (rule finite_deflation.compact)
-apply (simp only: finite_deflation_cfun_map)
-done
-
 end
--- a/src/HOLCF/Domain_Aux.thy	Wed Nov 10 14:59:52 2010 -0800
+++ b/src/HOLCF/Domain_Aux.thy	Wed Nov 10 17:56:08 2010 -0800
@@ -5,7 +5,7 @@
 header {* Domain package support *}
 
 theory Domain_Aux
-imports Ssum Sprod Fixrec
+imports Map_Functions Fixrec
 uses
   ("Tools/Domain/domain_take_proofs.ML")
 begin
--- a/src/HOLCF/Fixrec.thy	Wed Nov 10 14:59:52 2010 -0800
+++ b/src/HOLCF/Fixrec.thy	Wed Nov 10 17:56:08 2010 -0800
@@ -5,7 +5,7 @@
 header "Package for defining recursive functions in HOLCF"
 
 theory Fixrec
-imports Cprod Sprod Ssum Up One Tr Fix
+imports Plain_HOLCF
 uses
   ("Tools/holcf_library.ML")
   ("Tools/fixrec.ML")
--- a/src/HOLCF/IsaMakefile	Wed Nov 10 14:59:52 2010 -0800
+++ b/src/HOLCF/IsaMakefile	Wed Nov 10 17:56:08 2010 -0800
@@ -54,9 +54,11 @@
   HOLCF.thy \
   Lift.thy \
   LowerPD.thy \
+  Map_Functions.thy \
   One.thy \
   Pcpodef.thy \
   Pcpo.thy \
+  Plain_HOLCF.thy \
   Porder.thy \
   Powerdomains.thy \
   Product_Cpo.thy \
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Map_Functions.thy	Wed Nov 10 17:56:08 2010 -0800
@@ -0,0 +1,383 @@
+(*  Title:      HOLCF/Map_Functions.thy
+    Author:     Brian Huffman
+*)
+
+header {* Map functions for various types *}
+
+theory Map_Functions
+imports Deflation
+begin
+
+subsection {* Map operator for continuous function space *}
+
+default_sort cpo
+
+definition
+  cfun_map :: "('b \<rightarrow> 'a) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> ('a \<rightarrow> 'c) \<rightarrow> ('b \<rightarrow> 'd)"
+where
+  "cfun_map = (\<Lambda> a b f x. b\<cdot>(f\<cdot>(a\<cdot>x)))"
+
+lemma cfun_map_beta [simp]: "cfun_map\<cdot>a\<cdot>b\<cdot>f\<cdot>x = b\<cdot>(f\<cdot>(a\<cdot>x))"
+unfolding cfun_map_def by simp
+
+lemma cfun_map_ID: "cfun_map\<cdot>ID\<cdot>ID = ID"
+unfolding cfun_eq_iff by simp
+
+lemma cfun_map_map:
+  "cfun_map\<cdot>f1\<cdot>g1\<cdot>(cfun_map\<cdot>f2\<cdot>g2\<cdot>p) =
+    cfun_map\<cdot>(\<Lambda> x. f2\<cdot>(f1\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
+by (rule cfun_eqI) simp
+
+lemma ep_pair_cfun_map:
+  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
+  shows "ep_pair (cfun_map\<cdot>p1\<cdot>e2) (cfun_map\<cdot>e1\<cdot>p2)"
+proof
+  interpret e1p1: ep_pair e1 p1 by fact
+  interpret e2p2: ep_pair e2 p2 by fact
+  fix f show "cfun_map\<cdot>e1\<cdot>p2\<cdot>(cfun_map\<cdot>p1\<cdot>e2\<cdot>f) = f"
+    by (simp add: cfun_eq_iff)
+  fix g show "cfun_map\<cdot>p1\<cdot>e2\<cdot>(cfun_map\<cdot>e1\<cdot>p2\<cdot>g) \<sqsubseteq> g"
+    apply (rule cfun_belowI, simp)
+    apply (rule below_trans [OF e2p2.e_p_below])
+    apply (rule monofun_cfun_arg)
+    apply (rule e1p1.e_p_below)
+    done
+qed
+
+lemma deflation_cfun_map:
+  assumes "deflation d1" and "deflation d2"
+  shows "deflation (cfun_map\<cdot>d1\<cdot>d2)"
+proof
+  interpret d1: deflation d1 by fact
+  interpret d2: deflation d2 by fact
+  fix f
+  show "cfun_map\<cdot>d1\<cdot>d2\<cdot>(cfun_map\<cdot>d1\<cdot>d2\<cdot>f) = cfun_map\<cdot>d1\<cdot>d2\<cdot>f"
+    by (simp add: cfun_eq_iff d1.idem d2.idem)
+  show "cfun_map\<cdot>d1\<cdot>d2\<cdot>f \<sqsubseteq> f"
+    apply (rule cfun_belowI, simp)
+    apply (rule below_trans [OF d2.below])
+    apply (rule monofun_cfun_arg)
+    apply (rule d1.below)
+    done
+qed
+
+lemma finite_range_cfun_map:
+  assumes a: "finite (range (\<lambda>x. a\<cdot>x))"
+  assumes b: "finite (range (\<lambda>y. b\<cdot>y))"
+  shows "finite (range (\<lambda>f. cfun_map\<cdot>a\<cdot>b\<cdot>f))"  (is "finite (range ?h)")
+proof (rule finite_imageD)
+  let ?f = "\<lambda>g. range (\<lambda>x. (a\<cdot>x, g\<cdot>x))"
+  show "finite (?f ` range ?h)"
+  proof (rule finite_subset)
+    let ?B = "Pow (range (\<lambda>x. a\<cdot>x) \<times> range (\<lambda>y. b\<cdot>y))"
+    show "?f ` range ?h \<subseteq> ?B"
+      by clarsimp
+    show "finite ?B"
+      by (simp add: a b)
+  qed
+  show "inj_on ?f (range ?h)"
+  proof (rule inj_onI, rule cfun_eqI, clarsimp)
+    fix x f g
+    assume "range (\<lambda>x. (a\<cdot>x, b\<cdot>(f\<cdot>(a\<cdot>x)))) = range (\<lambda>x. (a\<cdot>x, b\<cdot>(g\<cdot>(a\<cdot>x))))"
+    hence "range (\<lambda>x. (a\<cdot>x, b\<cdot>(f\<cdot>(a\<cdot>x)))) \<subseteq> range (\<lambda>x. (a\<cdot>x, b\<cdot>(g\<cdot>(a\<cdot>x))))"
+      by (rule equalityD1)
+    hence "(a\<cdot>x, b\<cdot>(f\<cdot>(a\<cdot>x))) \<in> range (\<lambda>x. (a\<cdot>x, b\<cdot>(g\<cdot>(a\<cdot>x))))"
+      by (simp add: subset_eq)
+    then obtain y where "(a\<cdot>x, b\<cdot>(f\<cdot>(a\<cdot>x))) = (a\<cdot>y, b\<cdot>(g\<cdot>(a\<cdot>y)))"
+      by (rule rangeE)
+    thus "b\<cdot>(f\<cdot>(a\<cdot>x)) = b\<cdot>(g\<cdot>(a\<cdot>x))"
+      by clarsimp
+  qed
+qed
+
+lemma finite_deflation_cfun_map:
+  assumes "finite_deflation d1" and "finite_deflation d2"
+  shows "finite_deflation (cfun_map\<cdot>d1\<cdot>d2)"
+proof (rule finite_deflation_intro)
+  interpret d1: finite_deflation d1 by fact
+  interpret d2: finite_deflation d2 by fact
+  have "deflation d1" and "deflation d2" by fact+
+  thus "deflation (cfun_map\<cdot>d1\<cdot>d2)" by (rule deflation_cfun_map)
+  have "finite (range (\<lambda>f. cfun_map\<cdot>d1\<cdot>d2\<cdot>f))"
+    using d1.finite_range d2.finite_range
+    by (rule finite_range_cfun_map)
+  thus "finite {f. cfun_map\<cdot>d1\<cdot>d2\<cdot>f = f}"
+    by (rule finite_range_imp_finite_fixes)
+qed
+
+text {* Finite deflations are compact elements of the function space *}
+
+lemma finite_deflation_imp_compact: "finite_deflation d \<Longrightarrow> compact d"
+apply (frule finite_deflation_imp_deflation)
+apply (subgoal_tac "compact (cfun_map\<cdot>d\<cdot>d\<cdot>d)")
+apply (simp add: cfun_map_def deflation.idem eta_cfun)
+apply (rule finite_deflation.compact)
+apply (simp only: finite_deflation_cfun_map)
+done
+
+subsection {* Map operator for product type *}
+
+definition
+  cprod_map :: "('a \<rightarrow> 'b) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> 'a \<times> 'c \<rightarrow> 'b \<times> 'd"
+where
+  "cprod_map = (\<Lambda> f g p. (f\<cdot>(fst p), g\<cdot>(snd p)))"
+
+lemma cprod_map_Pair [simp]: "cprod_map\<cdot>f\<cdot>g\<cdot>(x, y) = (f\<cdot>x, g\<cdot>y)"
+unfolding cprod_map_def by simp
+
+lemma cprod_map_ID: "cprod_map\<cdot>ID\<cdot>ID = ID"
+unfolding cfun_eq_iff by auto
+
+lemma cprod_map_map:
+  "cprod_map\<cdot>f1\<cdot>g1\<cdot>(cprod_map\<cdot>f2\<cdot>g2\<cdot>p) =
+    cprod_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
+by (induct p) simp
+
+lemma ep_pair_cprod_map:
+  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
+  shows "ep_pair (cprod_map\<cdot>e1\<cdot>e2) (cprod_map\<cdot>p1\<cdot>p2)"
+proof
+  interpret e1p1: ep_pair e1 p1 by fact
+  interpret e2p2: ep_pair e2 p2 by fact
+  fix x show "cprod_map\<cdot>p1\<cdot>p2\<cdot>(cprod_map\<cdot>e1\<cdot>e2\<cdot>x) = x"
+    by (induct x) simp
+  fix y show "cprod_map\<cdot>e1\<cdot>e2\<cdot>(cprod_map\<cdot>p1\<cdot>p2\<cdot>y) \<sqsubseteq> y"
+    by (induct y) (simp add: e1p1.e_p_below e2p2.e_p_below)
+qed
+
+lemma deflation_cprod_map:
+  assumes "deflation d1" and "deflation d2"
+  shows "deflation (cprod_map\<cdot>d1\<cdot>d2)"
+proof
+  interpret d1: deflation d1 by fact
+  interpret d2: deflation d2 by fact
+  fix x
+  show "cprod_map\<cdot>d1\<cdot>d2\<cdot>(cprod_map\<cdot>d1\<cdot>d2\<cdot>x) = cprod_map\<cdot>d1\<cdot>d2\<cdot>x"
+    by (induct x) (simp add: d1.idem d2.idem)
+  show "cprod_map\<cdot>d1\<cdot>d2\<cdot>x \<sqsubseteq> x"
+    by (induct x) (simp add: d1.below d2.below)
+qed
+
+lemma finite_deflation_cprod_map:
+  assumes "finite_deflation d1" and "finite_deflation d2"
+  shows "finite_deflation (cprod_map\<cdot>d1\<cdot>d2)"
+proof (rule finite_deflation_intro)
+  interpret d1: finite_deflation d1 by fact
+  interpret d2: finite_deflation d2 by fact
+  have "deflation d1" and "deflation d2" by fact+
+  thus "deflation (cprod_map\<cdot>d1\<cdot>d2)" by (rule deflation_cprod_map)
+  have "{p. cprod_map\<cdot>d1\<cdot>d2\<cdot>p = p} \<subseteq> {x. d1\<cdot>x = x} \<times> {y. d2\<cdot>y = y}"
+    by clarsimp
+  thus "finite {p. cprod_map\<cdot>d1\<cdot>d2\<cdot>p = p}"
+    by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
+qed
+
+subsection {* Map function for lifted cpo *}
+
+definition
+  u_map :: "('a \<rightarrow> 'b) \<rightarrow> 'a u \<rightarrow> 'b u"
+where
+  "u_map = (\<Lambda> f. fup\<cdot>(up oo f))"
+
+lemma u_map_strict [simp]: "u_map\<cdot>f\<cdot>\<bottom> = \<bottom>"
+unfolding u_map_def by simp
+
+lemma u_map_up [simp]: "u_map\<cdot>f\<cdot>(up\<cdot>x) = up\<cdot>(f\<cdot>x)"
+unfolding u_map_def by simp
+
+lemma u_map_ID: "u_map\<cdot>ID = ID"
+unfolding u_map_def by (simp add: cfun_eq_iff eta_cfun)
+
+lemma u_map_map: "u_map\<cdot>f\<cdot>(u_map\<cdot>g\<cdot>p) = u_map\<cdot>(\<Lambda> x. f\<cdot>(g\<cdot>x))\<cdot>p"
+by (induct p) simp_all
+
+lemma ep_pair_u_map: "ep_pair e p \<Longrightarrow> ep_pair (u_map\<cdot>e) (u_map\<cdot>p)"
+apply default
+apply (case_tac x, simp, simp add: ep_pair.e_inverse)
+apply (case_tac y, simp, simp add: ep_pair.e_p_below)
+done
+
+lemma deflation_u_map: "deflation d \<Longrightarrow> deflation (u_map\<cdot>d)"
+apply default
+apply (case_tac x, simp, simp add: deflation.idem)
+apply (case_tac x, simp, simp add: deflation.below)
+done
+
+lemma finite_deflation_u_map:
+  assumes "finite_deflation d" shows "finite_deflation (u_map\<cdot>d)"
+proof (rule finite_deflation_intro)
+  interpret d: finite_deflation d by fact
+  have "deflation d" by fact
+  thus "deflation (u_map\<cdot>d)" by (rule deflation_u_map)
+  have "{x. u_map\<cdot>d\<cdot>x = x} \<subseteq> insert \<bottom> ((\<lambda>x. up\<cdot>x) ` {x. d\<cdot>x = x})"
+    by (rule subsetI, case_tac x, simp_all)
+  thus "finite {x. u_map\<cdot>d\<cdot>x = x}"
+    by (rule finite_subset, simp add: d.finite_fixes)
+qed
+
+subsection {* Map function for strict products *}
+
+default_sort pcpo
+
+definition
+  sprod_map :: "('a \<rightarrow> 'b) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> 'a \<otimes> 'c \<rightarrow> 'b \<otimes> 'd"
+where
+  "sprod_map = (\<Lambda> f g. ssplit\<cdot>(\<Lambda> x y. (:f\<cdot>x, g\<cdot>y:)))"
+
+lemma sprod_map_strict [simp]: "sprod_map\<cdot>a\<cdot>b\<cdot>\<bottom> = \<bottom>"
+unfolding sprod_map_def by simp
+
+lemma sprod_map_spair [simp]:
+  "x \<noteq> \<bottom> \<Longrightarrow> y \<noteq> \<bottom> \<Longrightarrow> sprod_map\<cdot>f\<cdot>g\<cdot>(:x, y:) = (:f\<cdot>x, g\<cdot>y:)"
+by (simp add: sprod_map_def)
+
+lemma sprod_map_spair':
+  "f\<cdot>\<bottom> = \<bottom> \<Longrightarrow> g\<cdot>\<bottom> = \<bottom> \<Longrightarrow> sprod_map\<cdot>f\<cdot>g\<cdot>(:x, y:) = (:f\<cdot>x, g\<cdot>y:)"
+by (cases "x = \<bottom> \<or> y = \<bottom>") auto
+
+lemma sprod_map_ID: "sprod_map\<cdot>ID\<cdot>ID = ID"
+unfolding sprod_map_def by (simp add: cfun_eq_iff eta_cfun)
+
+lemma sprod_map_map:
+  "\<lbrakk>f1\<cdot>\<bottom> = \<bottom>; g1\<cdot>\<bottom> = \<bottom>\<rbrakk> \<Longrightarrow>
+    sprod_map\<cdot>f1\<cdot>g1\<cdot>(sprod_map\<cdot>f2\<cdot>g2\<cdot>p) =
+     sprod_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
+apply (induct p, simp)
+apply (case_tac "f2\<cdot>x = \<bottom>", simp)
+apply (case_tac "g2\<cdot>y = \<bottom>", simp)
+apply simp
+done
+
+lemma ep_pair_sprod_map:
+  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
+  shows "ep_pair (sprod_map\<cdot>e1\<cdot>e2) (sprod_map\<cdot>p1\<cdot>p2)"
+proof
+  interpret e1p1: pcpo_ep_pair e1 p1 unfolding pcpo_ep_pair_def by fact
+  interpret e2p2: pcpo_ep_pair e2 p2 unfolding pcpo_ep_pair_def by fact
+  fix x show "sprod_map\<cdot>p1\<cdot>p2\<cdot>(sprod_map\<cdot>e1\<cdot>e2\<cdot>x) = x"
+    by (induct x) simp_all
+  fix y show "sprod_map\<cdot>e1\<cdot>e2\<cdot>(sprod_map\<cdot>p1\<cdot>p2\<cdot>y) \<sqsubseteq> y"
+    apply (induct y, simp)
+    apply (case_tac "p1\<cdot>x = \<bottom>", simp, case_tac "p2\<cdot>y = \<bottom>", simp)
+    apply (simp add: monofun_cfun e1p1.e_p_below e2p2.e_p_below)
+    done
+qed
+
+lemma deflation_sprod_map:
+  assumes "deflation d1" and "deflation d2"
+  shows "deflation (sprod_map\<cdot>d1\<cdot>d2)"
+proof
+  interpret d1: deflation d1 by fact
+  interpret d2: deflation d2 by fact
+  fix x
+  show "sprod_map\<cdot>d1\<cdot>d2\<cdot>(sprod_map\<cdot>d1\<cdot>d2\<cdot>x) = sprod_map\<cdot>d1\<cdot>d2\<cdot>x"
+    apply (induct x, simp)
+    apply (case_tac "d1\<cdot>x = \<bottom>", simp, case_tac "d2\<cdot>y = \<bottom>", simp)
+    apply (simp add: d1.idem d2.idem)
+    done
+  show "sprod_map\<cdot>d1\<cdot>d2\<cdot>x \<sqsubseteq> x"
+    apply (induct x, simp)
+    apply (simp add: monofun_cfun d1.below d2.below)
+    done
+qed
+
+lemma finite_deflation_sprod_map:
+  assumes "finite_deflation d1" and "finite_deflation d2"
+  shows "finite_deflation (sprod_map\<cdot>d1\<cdot>d2)"
+proof (rule finite_deflation_intro)
+  interpret d1: finite_deflation d1 by fact
+  interpret d2: finite_deflation d2 by fact
+  have "deflation d1" and "deflation d2" by fact+
+  thus "deflation (sprod_map\<cdot>d1\<cdot>d2)" by (rule deflation_sprod_map)
+  have "{x. sprod_map\<cdot>d1\<cdot>d2\<cdot>x = x} \<subseteq> insert \<bottom>
+        ((\<lambda>(x, y). (:x, y:)) ` ({x. d1\<cdot>x = x} \<times> {y. d2\<cdot>y = y}))"
+    by (rule subsetI, case_tac x, auto simp add: spair_eq_iff)
+  thus "finite {x. sprod_map\<cdot>d1\<cdot>d2\<cdot>x = x}"
+    by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
+qed
+
+subsection {* Map function for strict sums *}
+
+definition
+  ssum_map :: "('a \<rightarrow> 'b) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> 'a \<oplus> 'c \<rightarrow> 'b \<oplus> 'd"
+where
+  "ssum_map = (\<Lambda> f g. sscase\<cdot>(sinl oo f)\<cdot>(sinr oo g))"
+
+lemma ssum_map_strict [simp]: "ssum_map\<cdot>f\<cdot>g\<cdot>\<bottom> = \<bottom>"
+unfolding ssum_map_def by simp
+
+lemma ssum_map_sinl [simp]: "x \<noteq> \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinl\<cdot>x) = sinl\<cdot>(f\<cdot>x)"
+unfolding ssum_map_def by simp
+
+lemma ssum_map_sinr [simp]: "x \<noteq> \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinr\<cdot>x) = sinr\<cdot>(g\<cdot>x)"
+unfolding ssum_map_def by simp
+
+lemma ssum_map_sinl': "f\<cdot>\<bottom> = \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinl\<cdot>x) = sinl\<cdot>(f\<cdot>x)"
+by (cases "x = \<bottom>") simp_all
+
+lemma ssum_map_sinr': "g\<cdot>\<bottom> = \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinr\<cdot>x) = sinr\<cdot>(g\<cdot>x)"
+by (cases "x = \<bottom>") simp_all
+
+lemma ssum_map_ID: "ssum_map\<cdot>ID\<cdot>ID = ID"
+unfolding ssum_map_def by (simp add: cfun_eq_iff eta_cfun)
+
+lemma ssum_map_map:
+  "\<lbrakk>f1\<cdot>\<bottom> = \<bottom>; g1\<cdot>\<bottom> = \<bottom>\<rbrakk> \<Longrightarrow>
+    ssum_map\<cdot>f1\<cdot>g1\<cdot>(ssum_map\<cdot>f2\<cdot>g2\<cdot>p) =
+     ssum_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
+apply (induct p, simp)
+apply (case_tac "f2\<cdot>x = \<bottom>", simp, simp)
+apply (case_tac "g2\<cdot>y = \<bottom>", simp, simp)
+done
+
+lemma ep_pair_ssum_map:
+  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
+  shows "ep_pair (ssum_map\<cdot>e1\<cdot>e2) (ssum_map\<cdot>p1\<cdot>p2)"
+proof
+  interpret e1p1: pcpo_ep_pair e1 p1 unfolding pcpo_ep_pair_def by fact
+  interpret e2p2: pcpo_ep_pair e2 p2 unfolding pcpo_ep_pair_def by fact
+  fix x show "ssum_map\<cdot>p1\<cdot>p2\<cdot>(ssum_map\<cdot>e1\<cdot>e2\<cdot>x) = x"
+    by (induct x) simp_all
+  fix y show "ssum_map\<cdot>e1\<cdot>e2\<cdot>(ssum_map\<cdot>p1\<cdot>p2\<cdot>y) \<sqsubseteq> y"
+    apply (induct y, simp)
+    apply (case_tac "p1\<cdot>x = \<bottom>", simp, simp add: e1p1.e_p_below)
+    apply (case_tac "p2\<cdot>y = \<bottom>", simp, simp add: e2p2.e_p_below)
+    done
+qed
+
+lemma deflation_ssum_map:
+  assumes "deflation d1" and "deflation d2"
+  shows "deflation (ssum_map\<cdot>d1\<cdot>d2)"
+proof
+  interpret d1: deflation d1 by fact
+  interpret d2: deflation d2 by fact
+  fix x
+  show "ssum_map\<cdot>d1\<cdot>d2\<cdot>(ssum_map\<cdot>d1\<cdot>d2\<cdot>x) = ssum_map\<cdot>d1\<cdot>d2\<cdot>x"
+    apply (induct x, simp)
+    apply (case_tac "d1\<cdot>x = \<bottom>", simp, simp add: d1.idem)
+    apply (case_tac "d2\<cdot>y = \<bottom>", simp, simp add: d2.idem)
+    done
+  show "ssum_map\<cdot>d1\<cdot>d2\<cdot>x \<sqsubseteq> x"
+    apply (induct x, simp)
+    apply (case_tac "d1\<cdot>x = \<bottom>", simp, simp add: d1.below)
+    apply (case_tac "d2\<cdot>y = \<bottom>", simp, simp add: d2.below)
+    done
+qed
+
+lemma finite_deflation_ssum_map:
+  assumes "finite_deflation d1" and "finite_deflation d2"
+  shows "finite_deflation (ssum_map\<cdot>d1\<cdot>d2)"
+proof (rule finite_deflation_intro)
+  interpret d1: finite_deflation d1 by fact
+  interpret d2: finite_deflation d2 by fact
+  have "deflation d1" and "deflation d2" by fact+
+  thus "deflation (ssum_map\<cdot>d1\<cdot>d2)" by (rule deflation_ssum_map)
+  have "{x. ssum_map\<cdot>d1\<cdot>d2\<cdot>x = x} \<subseteq>
+        (\<lambda>x. sinl\<cdot>x) ` {x. d1\<cdot>x = x} \<union>
+        (\<lambda>x. sinr\<cdot>x) ` {x. d2\<cdot>x = x} \<union> {\<bottom>}"
+    by (rule subsetI, case_tac x, simp_all)
+  thus "finite {x. ssum_map\<cdot>d1\<cdot>d2\<cdot>x = x}"
+    by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOLCF/Plain_HOLCF.thy	Wed Nov 10 17:56:08 2010 -0800
@@ -0,0 +1,15 @@
+(*  Title:      HOLCF/Plain_HOLCF.thy
+    Author:     Brian Huffman
+*)
+
+header {* Plain HOLCF *}
+
+theory Plain_HOLCF
+imports Cfun Cprod Sprod Ssum Up Discrete Lift One Tr Fix
+begin
+
+text {*
+  Basic HOLCF concepts and types; does not include definition packages.
+*}
+
+end
--- a/src/HOLCF/Sprod.thy	Wed Nov 10 14:59:52 2010 -0800
+++ b/src/HOLCF/Sprod.thy	Wed Nov 10 17:56:08 2010 -0800
@@ -1,11 +1,12 @@
 (*  Title:      HOLCF/Sprod.thy
-    Author:     Franz Regensburger and Brian Huffman
+    Author:     Franz Regensburger
+    Author:     Brian Huffman
 *)
 
 header {* The type of strict products *}
 
 theory Sprod
-imports Deflation
+imports Cfun
 begin
 
 default_sort pcpo
@@ -210,83 +211,4 @@
     done
 qed
 
-subsection {* Map function for strict products *}
-
-definition
-  sprod_map :: "('a \<rightarrow> 'b) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> 'a \<otimes> 'c \<rightarrow> 'b \<otimes> 'd"
-where
-  "sprod_map = (\<Lambda> f g. ssplit\<cdot>(\<Lambda> x y. (:f\<cdot>x, g\<cdot>y:)))"
-
-lemma sprod_map_strict [simp]: "sprod_map\<cdot>a\<cdot>b\<cdot>\<bottom> = \<bottom>"
-unfolding sprod_map_def by simp
-
-lemma sprod_map_spair [simp]:
-  "x \<noteq> \<bottom> \<Longrightarrow> y \<noteq> \<bottom> \<Longrightarrow> sprod_map\<cdot>f\<cdot>g\<cdot>(:x, y:) = (:f\<cdot>x, g\<cdot>y:)"
-by (simp add: sprod_map_def)
-
-lemma sprod_map_spair':
-  "f\<cdot>\<bottom> = \<bottom> \<Longrightarrow> g\<cdot>\<bottom> = \<bottom> \<Longrightarrow> sprod_map\<cdot>f\<cdot>g\<cdot>(:x, y:) = (:f\<cdot>x, g\<cdot>y:)"
-by (cases "x = \<bottom> \<or> y = \<bottom>") auto
-
-lemma sprod_map_ID: "sprod_map\<cdot>ID\<cdot>ID = ID"
-unfolding sprod_map_def by (simp add: cfun_eq_iff eta_cfun)
-
-lemma sprod_map_map:
-  "\<lbrakk>f1\<cdot>\<bottom> = \<bottom>; g1\<cdot>\<bottom> = \<bottom>\<rbrakk> \<Longrightarrow>
-    sprod_map\<cdot>f1\<cdot>g1\<cdot>(sprod_map\<cdot>f2\<cdot>g2\<cdot>p) =
-     sprod_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
-apply (induct p, simp)
-apply (case_tac "f2\<cdot>x = \<bottom>", simp)
-apply (case_tac "g2\<cdot>y = \<bottom>", simp)
-apply simp
-done
-
-lemma ep_pair_sprod_map:
-  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
-  shows "ep_pair (sprod_map\<cdot>e1\<cdot>e2) (sprod_map\<cdot>p1\<cdot>p2)"
-proof
-  interpret e1p1: pcpo_ep_pair e1 p1 unfolding pcpo_ep_pair_def by fact
-  interpret e2p2: pcpo_ep_pair e2 p2 unfolding pcpo_ep_pair_def by fact
-  fix x show "sprod_map\<cdot>p1\<cdot>p2\<cdot>(sprod_map\<cdot>e1\<cdot>e2\<cdot>x) = x"
-    by (induct x) simp_all
-  fix y show "sprod_map\<cdot>e1\<cdot>e2\<cdot>(sprod_map\<cdot>p1\<cdot>p2\<cdot>y) \<sqsubseteq> y"
-    apply (induct y, simp)
-    apply (case_tac "p1\<cdot>x = \<bottom>", simp, case_tac "p2\<cdot>y = \<bottom>", simp)
-    apply (simp add: monofun_cfun e1p1.e_p_below e2p2.e_p_below)
-    done
-qed
-
-lemma deflation_sprod_map:
-  assumes "deflation d1" and "deflation d2"
-  shows "deflation (sprod_map\<cdot>d1\<cdot>d2)"
-proof
-  interpret d1: deflation d1 by fact
-  interpret d2: deflation d2 by fact
-  fix x
-  show "sprod_map\<cdot>d1\<cdot>d2\<cdot>(sprod_map\<cdot>d1\<cdot>d2\<cdot>x) = sprod_map\<cdot>d1\<cdot>d2\<cdot>x"
-    apply (induct x, simp)
-    apply (case_tac "d1\<cdot>x = \<bottom>", simp, case_tac "d2\<cdot>y = \<bottom>", simp)
-    apply (simp add: d1.idem d2.idem)
-    done
-  show "sprod_map\<cdot>d1\<cdot>d2\<cdot>x \<sqsubseteq> x"
-    apply (induct x, simp)
-    apply (simp add: monofun_cfun d1.below d2.below)
-    done
-qed
-
-lemma finite_deflation_sprod_map:
-  assumes "finite_deflation d1" and "finite_deflation d2"
-  shows "finite_deflation (sprod_map\<cdot>d1\<cdot>d2)"
-proof (rule finite_deflation_intro)
-  interpret d1: finite_deflation d1 by fact
-  interpret d2: finite_deflation d2 by fact
-  have "deflation d1" and "deflation d2" by fact+
-  thus "deflation (sprod_map\<cdot>d1\<cdot>d2)" by (rule deflation_sprod_map)
-  have "{x. sprod_map\<cdot>d1\<cdot>d2\<cdot>x = x} \<subseteq> insert \<bottom>
-        ((\<lambda>(x, y). (:x, y:)) ` ({x. d1\<cdot>x = x} \<times> {y. d2\<cdot>y = y}))"
-    by (rule subsetI, case_tac x, auto simp add: spair_eq_iff)
-  thus "finite {x. sprod_map\<cdot>d1\<cdot>d2\<cdot>x = x}"
-    by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
-qed
-
 end
--- a/src/HOLCF/Ssum.thy	Wed Nov 10 14:59:52 2010 -0800
+++ b/src/HOLCF/Ssum.thy	Wed Nov 10 17:56:08 2010 -0800
@@ -1,5 +1,6 @@
 (*  Title:      HOLCF/Ssum.thy
-    Author:     Franz Regensburger and Brian Huffman
+    Author:     Franz Regensburger
+    Author:     Brian Huffman
 *)
 
 header {* The type of strict sums *}
@@ -194,88 +195,4 @@
 apply (case_tac y, simp_all add: flat_below_iff)
 done
 
-subsection {* Map function for strict sums *}
-
-definition
-  ssum_map :: "('a \<rightarrow> 'b) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> 'a \<oplus> 'c \<rightarrow> 'b \<oplus> 'd"
-where
-  "ssum_map = (\<Lambda> f g. sscase\<cdot>(sinl oo f)\<cdot>(sinr oo g))"
-
-lemma ssum_map_strict [simp]: "ssum_map\<cdot>f\<cdot>g\<cdot>\<bottom> = \<bottom>"
-unfolding ssum_map_def by simp
-
-lemma ssum_map_sinl [simp]: "x \<noteq> \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinl\<cdot>x) = sinl\<cdot>(f\<cdot>x)"
-unfolding ssum_map_def by simp
-
-lemma ssum_map_sinr [simp]: "x \<noteq> \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinr\<cdot>x) = sinr\<cdot>(g\<cdot>x)"
-unfolding ssum_map_def by simp
-
-lemma ssum_map_sinl': "f\<cdot>\<bottom> = \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinl\<cdot>x) = sinl\<cdot>(f\<cdot>x)"
-by (cases "x = \<bottom>") simp_all
-
-lemma ssum_map_sinr': "g\<cdot>\<bottom> = \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinr\<cdot>x) = sinr\<cdot>(g\<cdot>x)"
-by (cases "x = \<bottom>") simp_all
-
-lemma ssum_map_ID: "ssum_map\<cdot>ID\<cdot>ID = ID"
-unfolding ssum_map_def by (simp add: cfun_eq_iff eta_cfun)
-
-lemma ssum_map_map:
-  "\<lbrakk>f1\<cdot>\<bottom> = \<bottom>; g1\<cdot>\<bottom> = \<bottom>\<rbrakk> \<Longrightarrow>
-    ssum_map\<cdot>f1\<cdot>g1\<cdot>(ssum_map\<cdot>f2\<cdot>g2\<cdot>p) =
-     ssum_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
-apply (induct p, simp)
-apply (case_tac "f2\<cdot>x = \<bottom>", simp, simp)
-apply (case_tac "g2\<cdot>y = \<bottom>", simp, simp)
-done
-
-lemma ep_pair_ssum_map:
-  assumes "ep_pair e1 p1" and "ep_pair e2 p2"
-  shows "ep_pair (ssum_map\<cdot>e1\<cdot>e2) (ssum_map\<cdot>p1\<cdot>p2)"
-proof
-  interpret e1p1: pcpo_ep_pair e1 p1 unfolding pcpo_ep_pair_def by fact
-  interpret e2p2: pcpo_ep_pair e2 p2 unfolding pcpo_ep_pair_def by fact
-  fix x show "ssum_map\<cdot>p1\<cdot>p2\<cdot>(ssum_map\<cdot>e1\<cdot>e2\<cdot>x) = x"
-    by (induct x) simp_all
-  fix y show "ssum_map\<cdot>e1\<cdot>e2\<cdot>(ssum_map\<cdot>p1\<cdot>p2\<cdot>y) \<sqsubseteq> y"
-    apply (induct y, simp)
-    apply (case_tac "p1\<cdot>x = \<bottom>", simp, simp add: e1p1.e_p_below)
-    apply (case_tac "p2\<cdot>y = \<bottom>", simp, simp add: e2p2.e_p_below)
-    done
-qed
-
-lemma deflation_ssum_map:
-  assumes "deflation d1" and "deflation d2"
-  shows "deflation (ssum_map\<cdot>d1\<cdot>d2)"
-proof
-  interpret d1: deflation d1 by fact
-  interpret d2: deflation d2 by fact
-  fix x
-  show "ssum_map\<cdot>d1\<cdot>d2\<cdot>(ssum_map\<cdot>d1\<cdot>d2\<cdot>x) = ssum_map\<cdot>d1\<cdot>d2\<cdot>x"
-    apply (induct x, simp)
-    apply (case_tac "d1\<cdot>x = \<bottom>", simp, simp add: d1.idem)
-    apply (case_tac "d2\<cdot>y = \<bottom>", simp, simp add: d2.idem)
-    done
-  show "ssum_map\<cdot>d1\<cdot>d2\<cdot>x \<sqsubseteq> x"
-    apply (induct x, simp)
-    apply (case_tac "d1\<cdot>x = \<bottom>", simp, simp add: d1.below)
-    apply (case_tac "d2\<cdot>y = \<bottom>", simp, simp add: d2.below)
-    done
-qed
-
-lemma finite_deflation_ssum_map:
-  assumes "finite_deflation d1" and "finite_deflation d2"
-  shows "finite_deflation (ssum_map\<cdot>d1\<cdot>d2)"
-proof (rule finite_deflation_intro)
-  interpret d1: finite_deflation d1 by fact
-  interpret d2: finite_deflation d2 by fact
-  have "deflation d1" and "deflation d2" by fact+
-  thus "deflation (ssum_map\<cdot>d1\<cdot>d2)" by (rule deflation_ssum_map)
-  have "{x. ssum_map\<cdot>d1\<cdot>d2\<cdot>x = x} \<subseteq>
-        (\<lambda>x. sinl\<cdot>x) ` {x. d1\<cdot>x = x} \<union>
-        (\<lambda>x. sinr\<cdot>x) ` {x. d2\<cdot>x = x} \<union> {\<bottom>}"
-    by (rule subsetI, case_tac x, simp_all)
-  thus "finite {x. ssum_map\<cdot>d1\<cdot>d2\<cdot>x = x}"
-    by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
-qed
-
 end
--- a/src/HOLCF/Up.thy	Wed Nov 10 14:59:52 2010 -0800
+++ b/src/HOLCF/Up.thy	Wed Nov 10 17:56:08 2010 -0800
@@ -1,11 +1,12 @@
 (*  Title:      HOLCF/Up.thy
-    Author:     Franz Regensburger and Brian Huffman
+    Author:     Franz Regensburger
+    Author:     Brian Huffman
 *)
 
 header {* The type of lifted values *}
 
 theory Up
-imports Deflation
+imports Cfun
 begin
 
 default_sort cpo
@@ -259,47 +260,4 @@
 lemma fup3 [simp]: "fup\<cdot>up\<cdot>x = x"
 by (cases x, simp_all)
 
-subsection {* Map function for lifted cpo *}
-
-definition
-  u_map :: "('a \<rightarrow> 'b) \<rightarrow> 'a u \<rightarrow> 'b u"
-where
-  "u_map = (\<Lambda> f. fup\<cdot>(up oo f))"
-
-lemma u_map_strict [simp]: "u_map\<cdot>f\<cdot>\<bottom> = \<bottom>"
-unfolding u_map_def by simp
-
-lemma u_map_up [simp]: "u_map\<cdot>f\<cdot>(up\<cdot>x) = up\<cdot>(f\<cdot>x)"
-unfolding u_map_def by simp
-
-lemma u_map_ID: "u_map\<cdot>ID = ID"
-unfolding u_map_def by (simp add: cfun_eq_iff eta_cfun)
-
-lemma u_map_map: "u_map\<cdot>f\<cdot>(u_map\<cdot>g\<cdot>p) = u_map\<cdot>(\<Lambda> x. f\<cdot>(g\<cdot>x))\<cdot>p"
-by (induct p) simp_all
-
-lemma ep_pair_u_map: "ep_pair e p \<Longrightarrow> ep_pair (u_map\<cdot>e) (u_map\<cdot>p)"
-apply default
-apply (case_tac x, simp, simp add: ep_pair.e_inverse)
-apply (case_tac y, simp, simp add: ep_pair.e_p_below)
-done
-
-lemma deflation_u_map: "deflation d \<Longrightarrow> deflation (u_map\<cdot>d)"
-apply default
-apply (case_tac x, simp, simp add: deflation.idem)
-apply (case_tac x, simp, simp add: deflation.below)
-done
-
-lemma finite_deflation_u_map:
-  assumes "finite_deflation d" shows "finite_deflation (u_map\<cdot>d)"
-proof (rule finite_deflation_intro)
-  interpret d: finite_deflation d by fact
-  have "deflation d" by fact
-  thus "deflation (u_map\<cdot>d)" by (rule deflation_u_map)
-  have "{x. u_map\<cdot>d\<cdot>x = x} \<subseteq> insert \<bottom> ((\<lambda>x. up\<cdot>x) ` {x. d\<cdot>x = x})"
-    by (rule subsetI, case_tac x, simp_all)
-  thus "finite {x. u_map\<cdot>d\<cdot>x = x}"
-    by (rule finite_subset, simp add: d.finite_fixes)
-qed
-
 end