major reorganization/simplification of HOLCF type classes:
authorhuffman
Wed, 06 Oct 2010 10:49:27 -0700
changeset 39974 b525988432e9
parent 39973 c62b4ff97bfc
child 39975 7c50d5ca5c04
child 39981 fdff0444fa7d
major reorganization/simplification of HOLCF type classes: removed profinite/bifinite classes and approx function; universal domain uses approx_chain locale instead of bifinite class; ideal_completion locale does not use 'take' functions, requires countable basis instead; replaced type 'udom alg_defl' with type 'sfp'; replaced class 'rep' with class 'sfp'; renamed REP('a) to SFP('a);
src/HOLCF/Algebraic.thy
src/HOLCF/Bifinite.thy
src/HOLCF/CompactBasis.thy
src/HOLCF/Completion.thy
src/HOLCF/ConvexPD.thy
src/HOLCF/Cprod.thy
src/HOLCF/Eventual.thy
src/HOLCF/HOLCF.thy
src/HOLCF/IsaMakefile
src/HOLCF/Library/Strict_Fun.thy
src/HOLCF/Library/Sum_Cpo.thy
src/HOLCF/Lift.thy
src/HOLCF/LowerPD.thy
src/HOLCF/Powerdomains.thy
src/HOLCF/Representable.thy
src/HOLCF/Sprod.thy
src/HOLCF/Ssum.thy
src/HOLCF/Tools/Domain/domain_extender.ML
src/HOLCF/Tools/Domain/domain_isomorphism.ML
src/HOLCF/Tools/repdef.ML
src/HOLCF/Tutorial/New_Domain.thy
src/HOLCF/Universal.thy
src/HOLCF/Up.thy
src/HOLCF/UpperPD.thy
src/HOLCF/ex/Domain_Proofs.thy
src/HOLCF/ex/Powerdomain_ex.thy
--- a/src/HOLCF/Algebraic.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Algebraic.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -2,303 +2,18 @@
     Author:     Brian Huffman
 *)
 
-header {* Algebraic deflations *}
+header {* Algebraic deflations and SFP domains *}
 
 theory Algebraic
-imports Completion Fix Eventual Bifinite
-begin
-
-subsection {* Constructing finite deflations by iteration *}
-
-lemma le_Suc_induct:
-  assumes le: "i \<le> j"
-  assumes step: "\<And>i. P i (Suc i)"
-  assumes refl: "\<And>i. P i i"
-  assumes trans: "\<And>i j k. \<lbrakk>P i j; P j k\<rbrakk> \<Longrightarrow> P i k"
-  shows "P i j"
-proof (cases "i = j")
-  assume "i = j"
-  thus "P i j" by (simp add: refl)
-next
-  assume "i \<noteq> j"
-  with le have "i < j" by simp
-  thus "P i j" using step trans by (rule less_Suc_induct)
-qed
-
-definition
-  eventual_iterate :: "('a \<rightarrow> 'a::cpo) \<Rightarrow> ('a \<rightarrow> 'a)"
-where
-  "eventual_iterate f = eventual (\<lambda>n. iterate n\<cdot>f)"
-
-text {* A pre-deflation is like a deflation, but not idempotent. *}
-
-locale pre_deflation =
-  fixes f :: "'a \<rightarrow> 'a::cpo"
-  assumes below: "\<And>x. f\<cdot>x \<sqsubseteq> x"
-  assumes finite_range: "finite (range (\<lambda>x. f\<cdot>x))"
+imports Universal Bifinite
 begin
 
-lemma iterate_below: "iterate i\<cdot>f\<cdot>x \<sqsubseteq> x"
-by (induct i, simp_all add: below_trans [OF below])
-
-lemma iterate_fixed: "f\<cdot>x = x \<Longrightarrow> iterate i\<cdot>f\<cdot>x = x"
-by (induct i, simp_all)
-
-lemma antichain_iterate_app: "i \<le> j \<Longrightarrow> iterate j\<cdot>f\<cdot>x \<sqsubseteq> iterate i\<cdot>f\<cdot>x"
-apply (erule le_Suc_induct)
-apply (simp add: below)
-apply (rule below_refl)
-apply (erule (1) below_trans)
-done
-
-lemma finite_range_iterate_app: "finite (range (\<lambda>i. iterate i\<cdot>f\<cdot>x))"
-proof (rule finite_subset)
-  show "range (\<lambda>i. iterate i\<cdot>f\<cdot>x) \<subseteq> insert x (range (\<lambda>x. f\<cdot>x))"
-    by (clarify, case_tac i, simp_all)
-  show "finite (insert x (range (\<lambda>x. f\<cdot>x)))"
-    by (simp add: finite_range)
-qed
-
-lemma eventually_constant_iterate_app:
-  "eventually_constant (\<lambda>i. iterate i\<cdot>f\<cdot>x)"
-unfolding eventually_constant_def MOST_nat_le
-proof -
-  let ?Y = "\<lambda>i. iterate i\<cdot>f\<cdot>x"
-  have "\<exists>j. \<forall>k. ?Y j \<sqsubseteq> ?Y k"
-    apply (rule finite_range_has_max)
-    apply (erule antichain_iterate_app)
-    apply (rule finite_range_iterate_app)
-    done
-  then obtain j where j: "\<And>k. ?Y j \<sqsubseteq> ?Y k" by fast
-  show "\<exists>z m. \<forall>n\<ge>m. ?Y n = z"
-  proof (intro exI allI impI)
-    fix k
-    assume "j \<le> k"
-    hence "?Y k \<sqsubseteq> ?Y j" by (rule antichain_iterate_app)
-    also have "?Y j \<sqsubseteq> ?Y k" by (rule j)
-    finally show "?Y k = ?Y j" .
-  qed
-qed
-
-lemma eventually_constant_iterate:
-  "eventually_constant (\<lambda>n. iterate n\<cdot>f)"
-proof -
-  have "\<forall>y\<in>range (\<lambda>x. f\<cdot>x). eventually_constant (\<lambda>i. iterate i\<cdot>f\<cdot>y)"
-    by (simp add: eventually_constant_iterate_app)
-  hence "\<forall>y\<in>range (\<lambda>x. f\<cdot>x). MOST i. MOST j. iterate j\<cdot>f\<cdot>y = iterate i\<cdot>f\<cdot>y"
-    unfolding eventually_constant_MOST_MOST .
-  hence "MOST i. MOST j. \<forall>y\<in>range (\<lambda>x. f\<cdot>x). iterate j\<cdot>f\<cdot>y = iterate i\<cdot>f\<cdot>y"
-    by (simp only: MOST_finite_Ball_distrib [OF finite_range])
-  hence "MOST i. MOST j. \<forall>x. iterate j\<cdot>f\<cdot>(f\<cdot>x) = iterate i\<cdot>f\<cdot>(f\<cdot>x)"
-    by simp
-  hence "MOST i. MOST j. \<forall>x. iterate (Suc j)\<cdot>f\<cdot>x = iterate (Suc i)\<cdot>f\<cdot>x"
-    by (simp only: iterate_Suc2)
-  hence "MOST i. MOST j. iterate (Suc j)\<cdot>f = iterate (Suc i)\<cdot>f"
-    by (simp only: expand_cfun_eq)
-  hence "eventually_constant (\<lambda>i. iterate (Suc i)\<cdot>f)"
-    unfolding eventually_constant_MOST_MOST .
-  thus "eventually_constant (\<lambda>i. iterate i\<cdot>f)"
-    by (rule eventually_constant_SucD)
-qed
-
-abbreviation
-  d :: "'a \<rightarrow> 'a"
-where
-  "d \<equiv> eventual_iterate f"
-
-lemma MOST_d: "MOST n. P (iterate n\<cdot>f) \<Longrightarrow> P d"
-unfolding eventual_iterate_def
-using eventually_constant_iterate by (rule MOST_eventual)
-
-lemma f_d: "f\<cdot>(d\<cdot>x) = d\<cdot>x"
-apply (rule MOST_d)
-apply (subst iterate_Suc [symmetric])
-apply (rule eventually_constant_MOST_Suc_eq)
-apply (rule eventually_constant_iterate_app)
-done
-
-lemma d_fixed_iff: "d\<cdot>x = x \<longleftrightarrow> f\<cdot>x = x"
-proof
-  assume "d\<cdot>x = x"
-  with f_d [where x=x]
-  show "f\<cdot>x = x" by simp
-next
-  assume f: "f\<cdot>x = x"
-  have "\<forall>n. iterate n\<cdot>f\<cdot>x = x"
-    by (rule allI, rule nat.induct, simp, simp add: f)
-  hence "MOST n. iterate n\<cdot>f\<cdot>x = x"
-    by (rule ALL_MOST)
-  thus "d\<cdot>x = x"
-    by (rule MOST_d)
-qed
-
-lemma finite_deflation_d: "finite_deflation d"
-proof
-  fix x :: 'a
-  have "d \<in> range (\<lambda>n. iterate n\<cdot>f)"
-    unfolding eventual_iterate_def
-    using eventually_constant_iterate
-    by (rule eventual_mem_range)
-  then obtain n where n: "d = iterate n\<cdot>f" ..
-  have "iterate n\<cdot>f\<cdot>(d\<cdot>x) = d\<cdot>x"
-    using f_d by (rule iterate_fixed)
-  thus "d\<cdot>(d\<cdot>x) = d\<cdot>x"
-    by (simp add: n)
-next
-  fix x :: 'a
-  show "d\<cdot>x \<sqsubseteq> x"
-    by (rule MOST_d, simp add: iterate_below)
-next
-  from finite_range
-  have "finite {x. f\<cdot>x = x}"
-    by (rule finite_range_imp_finite_fixes)
-  thus "finite {x. d\<cdot>x = x}"
-    by (simp add: d_fixed_iff)
-qed
-
-lemma deflation_d: "deflation d"
-using finite_deflation_d
-by (rule finite_deflation_imp_deflation)
-
-end
-
-lemma finite_deflation_eventual_iterate:
-  "pre_deflation d \<Longrightarrow> finite_deflation (eventual_iterate d)"
-by (rule pre_deflation.finite_deflation_d)
-
-lemma pre_deflation_oo:
-  assumes "finite_deflation d"
-  assumes f: "\<And>x. f\<cdot>x \<sqsubseteq> x"
-  shows "pre_deflation (d oo f)"
-proof
-  interpret d: finite_deflation d by fact
-  fix x
-  show "\<And>x. (d oo f)\<cdot>x \<sqsubseteq> x"
-    by (simp, rule below_trans [OF d.below f])
-  show "finite (range (\<lambda>x. (d oo f)\<cdot>x))"
-    by (rule finite_subset [OF _ d.finite_range], auto)
-qed
-
-lemma eventual_iterate_oo_fixed_iff:
-  assumes "finite_deflation d"
-  assumes f: "\<And>x. f\<cdot>x \<sqsubseteq> x"
-  shows "eventual_iterate (d oo f)\<cdot>x = x \<longleftrightarrow> d\<cdot>x = x \<and> f\<cdot>x = x"
-proof -
-  interpret d: finite_deflation d by fact
-  let ?e = "d oo f"
-  interpret e: pre_deflation "d oo f"
-    using `finite_deflation d` f
-    by (rule pre_deflation_oo)
-  let ?g = "eventual (\<lambda>n. iterate n\<cdot>?e)"
-  show ?thesis
-    apply (subst e.d_fixed_iff)
-    apply simp
-    apply safe
-    apply (erule subst)
-    apply (rule d.idem)
-    apply (rule below_antisym)
-    apply (rule f)
-    apply (erule subst, rule d.below)
-    apply simp
-    done
-qed
-
-lemma eventual_mono:
-  assumes A: "eventually_constant A"
-  assumes B: "eventually_constant B"
-  assumes below: "\<And>n. A n \<sqsubseteq> B n"
-  shows "eventual A \<sqsubseteq> eventual B"
-proof -
-  from A have "MOST n. A n = eventual A"
-    by (rule MOST_eq_eventual)
-  then have "MOST n. eventual A \<sqsubseteq> B n"
-    by (rule MOST_mono) (erule subst, rule below)
-  with B show "eventual A \<sqsubseteq> eventual B"
-    by (rule MOST_eventual)
-qed
-
-lemma eventual_iterate_mono:
-  assumes f: "pre_deflation f" and g: "pre_deflation g" and "f \<sqsubseteq> g"
-  shows "eventual_iterate f \<sqsubseteq> eventual_iterate g"
-unfolding eventual_iterate_def
-apply (rule eventual_mono)
-apply (rule pre_deflation.eventually_constant_iterate [OF f])
-apply (rule pre_deflation.eventually_constant_iterate [OF g])
-apply (rule monofun_cfun_arg [OF `f \<sqsubseteq> g`])
-done
-
-lemma cont2cont_eventual_iterate_oo:
-  assumes d: "finite_deflation d"
-  assumes cont: "cont f" and below: "\<And>x y. f x\<cdot>y \<sqsubseteq> y"
-  shows "cont (\<lambda>x. eventual_iterate (d oo f x))"
-    (is "cont ?e")
-proof (rule contI2)
-  show "monofun ?e"
-    apply (rule monofunI)
-    apply (rule eventual_iterate_mono)
-    apply (rule pre_deflation_oo [OF d below])
-    apply (rule pre_deflation_oo [OF d below])
-    apply (rule monofun_cfun_arg)
-    apply (erule cont2monofunE [OF cont])
-    done
-next
-  fix Y :: "nat \<Rightarrow> 'b"
-  assume Y: "chain Y"
-  with cont have fY: "chain (\<lambda>i. f (Y i))"
-    by (rule ch2ch_cont)
-  assume eY: "chain (\<lambda>i. ?e (Y i))"
-  have lub_below: "\<And>x. f (\<Squnion>i. Y i)\<cdot>x \<sqsubseteq> x"
-    by (rule admD [OF _ Y], simp add: cont, rule below)
-  have "deflation (?e (\<Squnion>i. Y i))"
-    apply (rule pre_deflation.deflation_d)
-    apply (rule pre_deflation_oo [OF d lub_below])
-    done
-  then show "?e (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. ?e (Y i))"
-  proof (rule deflation.belowI)
-    fix x :: 'a
-    assume "?e (\<Squnion>i. Y i)\<cdot>x = x"
-    hence "d\<cdot>x = x" and "f (\<Squnion>i. Y i)\<cdot>x = x"
-      by (simp_all add: eventual_iterate_oo_fixed_iff [OF d lub_below])
-    hence "(\<Squnion>i. f (Y i)\<cdot>x) = x"
-      apply (simp only: cont2contlubE [OF cont Y])
-      apply (simp only: contlub_cfun_fun [OF fY])
-      done
-    have "compact (d\<cdot>x)"
-      using d by (rule finite_deflation.compact)
-    then have "compact x"
-      using `d\<cdot>x = x` by simp
-    then have "compact (\<Squnion>i. f (Y i)\<cdot>x)"
-      using `(\<Squnion>i. f (Y i)\<cdot>x) = x` by simp
-    then have "\<exists>n. max_in_chain n (\<lambda>i. f (Y i)\<cdot>x)"
-      by - (rule compact_imp_max_in_chain, simp add: fY, assumption)
-    then obtain n where n: "max_in_chain n (\<lambda>i. f (Y i)\<cdot>x)" ..
-    then have "f (Y n)\<cdot>x = x"
-      using `(\<Squnion>i. f (Y i)\<cdot>x) = x` fY by (simp add: maxinch_is_thelub)
-    with `d\<cdot>x = x` have "?e (Y n)\<cdot>x = x"
-      by (simp add: eventual_iterate_oo_fixed_iff [OF d below])
-    moreover have "?e (Y n)\<cdot>x \<sqsubseteq> (\<Squnion>i. ?e (Y i)\<cdot>x)"
-      by (rule is_ub_thelub, simp add: eY)
-    ultimately have "x \<sqsubseteq> (\<Squnion>i. ?e (Y i))\<cdot>x"
-      by (simp add: contlub_cfun_fun eY)
-    also have "(\<Squnion>i. ?e (Y i))\<cdot>x \<sqsubseteq> x"
-      apply (rule deflation.below)
-      apply (rule admD [OF adm_deflation eY])
-      apply (rule pre_deflation.deflation_d)
-      apply (rule pre_deflation_oo [OF d below])
-      done
-    finally show "(\<Squnion>i. ?e (Y i))\<cdot>x = x" ..
-  qed
-qed
-
-
 subsection {* Type constructor for finite deflations *}
 
-default_sort profinite
+typedef (open) fin_defl = "{d::udom \<rightarrow> udom. finite_deflation d}"
+by (fast intro: finite_deflation_UU)
 
-typedef (open) 'a fin_defl = "{d::'a \<rightarrow> 'a. finite_deflation d}"
-by (fast intro: finite_deflation_approx)
-
-instantiation fin_defl :: (profinite) below
+instantiation fin_defl :: below
 begin
 
 definition below_fin_defl_def:
@@ -307,8 +22,9 @@
 instance ..
 end
 
-instance fin_defl :: (profinite) po
-by (rule typedef_po [OF type_definition_fin_defl below_fin_defl_def])
+instance fin_defl :: po
+using type_definition_fin_defl below_fin_defl_def
+by (rule typedef_po)
 
 lemma finite_deflation_Rep_fin_defl: "finite_deflation (Rep_fin_defl d)"
 using Rep_fin_defl by simp
@@ -337,269 +53,146 @@
 apply (rule fin_defl_belowI, simp)
 done
 
+lemma Rep_fin_defl_mono: "a \<sqsubseteq> b \<Longrightarrow> Rep_fin_defl a \<sqsubseteq> Rep_fin_defl b"
+unfolding below_fin_defl_def .
+
 lemma Abs_fin_defl_mono:
   "\<lbrakk>finite_deflation a; finite_deflation b; a \<sqsubseteq> b\<rbrakk>
     \<Longrightarrow> Abs_fin_defl a \<sqsubseteq> Abs_fin_defl b"
 unfolding below_fin_defl_def
 by (simp add: Abs_fin_defl_inverse)
 
-
-subsection {* Take function for finite deflations *}
-
-definition
-  defl_approx :: "nat \<Rightarrow> ('a \<rightarrow> 'a) \<Rightarrow> ('a \<rightarrow> 'a)"
-where
-  "defl_approx i d = eventual_iterate (approx i oo d)"
-
-lemma finite_deflation_defl_approx:
-  "deflation d \<Longrightarrow> finite_deflation (defl_approx i d)"
-unfolding defl_approx_def
-apply (rule pre_deflation.finite_deflation_d)
-apply (rule pre_deflation_oo)
-apply (rule finite_deflation_approx)
-apply (erule deflation.below)
-done
-
-lemma deflation_defl_approx:
-  "deflation d \<Longrightarrow> deflation (defl_approx i d)"
-apply (rule finite_deflation_imp_deflation)
-apply (erule finite_deflation_defl_approx)
-done
-
-lemma defl_approx_fixed_iff:
-  "deflation d \<Longrightarrow> defl_approx i d\<cdot>x = x \<longleftrightarrow> approx i\<cdot>x = x \<and> d\<cdot>x = x"
-unfolding defl_approx_def
-apply (rule eventual_iterate_oo_fixed_iff)
-apply (rule finite_deflation_approx)
-apply (erule deflation.below)
-done
-
-lemma defl_approx_below:
-  "\<lbrakk>a \<sqsubseteq> b; deflation a; deflation b\<rbrakk> \<Longrightarrow> defl_approx i a \<sqsubseteq> defl_approx i b"
-apply (rule deflation.belowI)
-apply (erule deflation_defl_approx)
-apply (simp add: defl_approx_fixed_iff)
-apply (erule (1) deflation.belowD)
-apply (erule conjunct2)
-done
-
-lemma cont2cont_defl_approx:
-  assumes cont: "cont f" and below: "\<And>x y. f x\<cdot>y \<sqsubseteq> y"
-  shows "cont (\<lambda>x. defl_approx i (f x))"
-unfolding defl_approx_def
-using finite_deflation_approx assms
-by (rule cont2cont_eventual_iterate_oo)
-
-definition
-  fd_take :: "nat \<Rightarrow> 'a fin_defl \<Rightarrow> 'a fin_defl"
-where
-  "fd_take i d = Abs_fin_defl (defl_approx i (Rep_fin_defl d))"
-
-lemma Rep_fin_defl_fd_take:
-  "Rep_fin_defl (fd_take i d) = defl_approx i (Rep_fin_defl d)"
-unfolding fd_take_def
-apply (rule Abs_fin_defl_inverse [unfolded mem_Collect_eq])
-apply (rule finite_deflation_defl_approx)
-apply (rule deflation_Rep_fin_defl)
-done
+lemma (in finite_deflation) compact_belowI:
+  assumes "\<And>x. compact x \<Longrightarrow> d\<cdot>x = x \<Longrightarrow> f\<cdot>x = x" shows "d \<sqsubseteq> f"
+by (rule belowI, rule assms, erule subst, rule compact)
 
-lemma fd_take_fixed_iff:
-  "Rep_fin_defl (fd_take i d)\<cdot>x = x \<longleftrightarrow>
-    approx i\<cdot>x = x \<and> Rep_fin_defl d\<cdot>x = x"
-unfolding Rep_fin_defl_fd_take
-apply (rule defl_approx_fixed_iff)
-apply (rule deflation_Rep_fin_defl)
-done
-
-lemma fd_take_below: "fd_take n d \<sqsubseteq> d"
-apply (rule fin_defl_belowI)
-apply (simp add: fd_take_fixed_iff)
-done
-
-lemma fd_take_idem: "fd_take n (fd_take n d) = fd_take n d"
-apply (rule fin_defl_eqI)
-apply (simp add: fd_take_fixed_iff)
-done
-
-lemma fd_take_mono: "a \<sqsubseteq> b \<Longrightarrow> fd_take n a \<sqsubseteq> fd_take n b"
-apply (rule fin_defl_belowI)
-apply (simp add: fd_take_fixed_iff)
-apply (simp add: fin_defl_belowD)
-done
-
-lemma approx_fixed_le_lemma: "\<lbrakk>i \<le> j; approx i\<cdot>x = x\<rbrakk> \<Longrightarrow> approx j\<cdot>x = x"
-by (erule subst, simp add: min_def)
-
-lemma fd_take_chain: "m \<le> n \<Longrightarrow> fd_take m a \<sqsubseteq> fd_take n a"
-apply (rule fin_defl_belowI)
-apply (simp add: fd_take_fixed_iff)
-apply (simp add: approx_fixed_le_lemma)
-done
-
-lemma finite_range_fd_take: "finite (range (fd_take n))"
-apply (rule finite_imageD [where f="\<lambda>a. {x. Rep_fin_defl a\<cdot>x = x}"])
-apply (rule finite_subset [where B="Pow {x. approx n\<cdot>x = x}"])
-apply (clarify, simp add: fd_take_fixed_iff)
-apply (simp add: finite_fixes_approx)
-apply (rule inj_onI, clarify)
-apply (simp add: set_eq_iff fin_defl_eqI)
-done
-
-lemma fd_take_covers: "\<exists>n. fd_take n a = a"
-apply (rule_tac x=
-  "Max ((\<lambda>x. LEAST n. approx n\<cdot>x = x) ` {x. Rep_fin_defl a\<cdot>x = x})" in exI)
-apply (rule below_antisym)
-apply (rule fd_take_below)
-apply (rule fin_defl_belowI)
-apply (simp add: fd_take_fixed_iff)
-apply (rule approx_fixed_le_lemma)
-apply (rule Max_ge)
-apply (rule finite_imageI)
-apply (rule Rep_fin_defl.finite_fixes)
-apply (rule imageI)
-apply (erule CollectI)
-apply (rule LeastI_ex)
-apply (rule profinite_compact_eq_approx)
-apply (erule subst)
-apply (rule Rep_fin_defl.compact)
-done
-
-interpretation fin_defl: basis_take below fd_take
-apply default
-apply (rule fd_take_below)
-apply (rule fd_take_idem)
-apply (erule fd_take_mono)
-apply (rule fd_take_chain, simp)
-apply (rule finite_range_fd_take)
-apply (rule fd_take_covers)
-done
-
+lemma compact_Rep_fin_defl [simp]: "compact (Rep_fin_defl a)"
+using finite_deflation_Rep_fin_defl
+by (rule finite_deflation_imp_compact)
 
 subsection {* Defining algebraic deflations by ideal completion *}
 
-typedef (open) 'a alg_defl =
-  "{S::'a fin_defl set. below.ideal S}"
+text {*
+  An SFP domain is one that can be represented as the limit of a
+  sequence of finite posets.  Here we use omega-algebraic deflations
+  (i.e. countable ideals of finite deflations) to model sequences of
+  finite posets.
+*}
+
+typedef (open) sfp = "{S::fin_defl set. below.ideal S}"
 by (fast intro: below.ideal_principal)
 
-instantiation alg_defl :: (profinite) below
+instantiation sfp :: below
 begin
 
 definition
-  "x \<sqsubseteq> y \<longleftrightarrow> Rep_alg_defl x \<subseteq> Rep_alg_defl y"
+  "x \<sqsubseteq> y \<longleftrightarrow> Rep_sfp x \<subseteq> Rep_sfp y"
 
 instance ..
 end
 
-instance alg_defl :: (profinite) po
-by (rule below.typedef_ideal_po
-    [OF type_definition_alg_defl below_alg_defl_def])
+instance sfp :: po
+using type_definition_sfp below_sfp_def
+by (rule below.typedef_ideal_po)
 
-instance alg_defl :: (profinite) cpo
-by (rule below.typedef_ideal_cpo
-    [OF type_definition_alg_defl below_alg_defl_def])
+instance sfp :: cpo
+using type_definition_sfp below_sfp_def
+by (rule below.typedef_ideal_cpo)
 
-lemma Rep_alg_defl_lub:
-  "chain Y \<Longrightarrow> Rep_alg_defl (\<Squnion>i. Y i) = (\<Union>i. Rep_alg_defl (Y i))"
-by (rule below.typedef_ideal_rep_contlub
-    [OF type_definition_alg_defl below_alg_defl_def])
+lemma Rep_sfp_lub:
+  "chain Y \<Longrightarrow> Rep_sfp (\<Squnion>i. Y i) = (\<Union>i. Rep_sfp (Y i))"
+using type_definition_sfp below_sfp_def
+by (rule below.typedef_ideal_rep_contlub)
 
-lemma ideal_Rep_alg_defl: "below.ideal (Rep_alg_defl xs)"
-by (rule Rep_alg_defl [unfolded mem_Collect_eq])
+lemma ideal_Rep_sfp: "below.ideal (Rep_sfp xs)"
+by (rule Rep_sfp [unfolded mem_Collect_eq])
 
 definition
-  alg_defl_principal :: "'a fin_defl \<Rightarrow> 'a alg_defl" where
-  "alg_defl_principal t = Abs_alg_defl {u. u \<sqsubseteq> t}"
+  sfp_principal :: "fin_defl \<Rightarrow> sfp" where
+  "sfp_principal t = Abs_sfp {u. u \<sqsubseteq> t}"
 
-lemma Rep_alg_defl_principal:
-  "Rep_alg_defl (alg_defl_principal t) = {u. u \<sqsubseteq> t}"
-unfolding alg_defl_principal_def
-by (simp add: Abs_alg_defl_inverse below.ideal_principal)
+lemma Rep_sfp_principal:
+  "Rep_sfp (sfp_principal t) = {u. u \<sqsubseteq> t}"
+unfolding sfp_principal_def
+by (simp add: Abs_sfp_inverse below.ideal_principal)
 
-interpretation alg_defl:
-  ideal_completion below fd_take alg_defl_principal Rep_alg_defl
+lemma fin_defl_countable: "\<exists>f::fin_defl \<Rightarrow> nat. inj f"
+proof
+  have *: "\<And>d. finite (approx_chain.place udom_approx `
+               Rep_compact_basis -` {x. Rep_fin_defl d\<cdot>x = x})"
+    apply (rule finite_imageI)
+    apply (rule finite_vimageI)
+    apply (rule Rep_fin_defl.finite_fixes)
+    apply (simp add: inj_on_def Rep_compact_basis_inject)
+    done
+  have range_eq: "range Rep_compact_basis = {x. compact x}"
+    using type_definition_compact_basis by (rule type_definition.Rep_range)
+  show "inj (\<lambda>d. set_encode
+    (approx_chain.place udom_approx ` Rep_compact_basis -` {x. Rep_fin_defl d\<cdot>x = x}))"
+    apply (rule inj_onI)
+    apply (simp only: set_encode_eq *)
+    apply (simp only: inj_image_eq_iff approx_chain.inj_place [OF udom_approx])
+    apply (drule_tac f="image Rep_compact_basis" in arg_cong)
+    apply (simp del: vimage_Collect_eq add: range_eq set_eq_iff)
+    apply (rule Rep_fin_defl_inject [THEN iffD1])
+    apply (rule below_antisym)
+    apply (rule Rep_fin_defl.compact_belowI, rename_tac z)
+    apply (drule_tac x=z in spec, simp)
+    apply (rule Rep_fin_defl.compact_belowI, rename_tac z)
+    apply (drule_tac x=z in spec, simp)
+    done
+qed
+
+interpretation sfp: ideal_completion below sfp_principal Rep_sfp
 apply default
-apply (rule ideal_Rep_alg_defl)
-apply (erule Rep_alg_defl_lub)
-apply (rule Rep_alg_defl_principal)
-apply (simp only: below_alg_defl_def)
+apply (rule ideal_Rep_sfp)
+apply (erule Rep_sfp_lub)
+apply (rule Rep_sfp_principal)
+apply (simp only: below_sfp_def)
+apply (rule fin_defl_countable)
 done
 
 text {* Algebraic deflations are pointed *}
 
-lemma alg_defl_minimal:
-  "alg_defl_principal (Abs_fin_defl \<bottom>) \<sqsubseteq> x"
-apply (induct x rule: alg_defl.principal_induct, simp)
-apply (rule alg_defl.principal_mono)
-apply (induct_tac a)
-apply (rule Abs_fin_defl_mono)
-apply (rule finite_deflation_UU)
-apply simp
-apply (rule minimal)
+lemma sfp_minimal: "sfp_principal (Abs_fin_defl \<bottom>) \<sqsubseteq> x"
+apply (induct x rule: sfp.principal_induct, simp)
+apply (rule sfp.principal_mono)
+apply (simp add: below_fin_defl_def)
+apply (simp add: Abs_fin_defl_inverse finite_deflation_UU)
 done
 
-instance alg_defl :: (bifinite) pcpo
-by intro_classes (fast intro: alg_defl_minimal)
-
-lemma inst_alg_defl_pcpo: "\<bottom> = alg_defl_principal (Abs_fin_defl \<bottom>)"
-by (rule alg_defl_minimal [THEN UU_I, symmetric])
-
-text {* Algebraic deflations are profinite *}
-
-instantiation alg_defl :: (profinite) profinite
-begin
-
-definition
-  approx_alg_defl_def: "approx = alg_defl.completion_approx"
+instance sfp :: pcpo
+by intro_classes (fast intro: sfp_minimal)
 
-instance
-apply (intro_classes, unfold approx_alg_defl_def)
-apply (rule alg_defl.chain_completion_approx)
-apply (rule alg_defl.lub_completion_approx)
-apply (rule alg_defl.completion_approx_idem)
-apply (rule alg_defl.finite_fixes_completion_approx)
-done
-
-end
-
-instance alg_defl :: (bifinite) bifinite ..
-
-lemma approx_alg_defl_principal [simp]:
-  "approx n\<cdot>(alg_defl_principal t) = alg_defl_principal (fd_take n t)"
-unfolding approx_alg_defl_def
-by (rule alg_defl.completion_approx_principal)
-
-lemma approx_eq_alg_defl_principal:
-  "\<exists>t\<in>Rep_alg_defl xs. approx n\<cdot>xs = alg_defl_principal (fd_take n t)"
-unfolding approx_alg_defl_def
-by (rule alg_defl.completion_approx_eq_principal)
-
+lemma inst_sfp_pcpo: "\<bottom> = sfp_principal (Abs_fin_defl \<bottom>)"
+by (rule sfp_minimal [THEN UU_I, symmetric])
 
 subsection {* Applying algebraic deflations *}
 
 definition
-  cast :: "'a alg_defl \<rightarrow> 'a \<rightarrow> 'a"
+  cast :: "sfp \<rightarrow> udom \<rightarrow> udom"
 where
-  "cast = alg_defl.basis_fun Rep_fin_defl"
+  "cast = sfp.basis_fun Rep_fin_defl"
 
-lemma cast_alg_defl_principal:
-  "cast\<cdot>(alg_defl_principal a) = Rep_fin_defl a"
+lemma cast_sfp_principal:
+  "cast\<cdot>(sfp_principal a) = Rep_fin_defl a"
 unfolding cast_def
-apply (rule alg_defl.basis_fun_principal)
+apply (rule sfp.basis_fun_principal)
 apply (simp only: below_fin_defl_def)
 done
 
 lemma deflation_cast: "deflation (cast\<cdot>d)"
-apply (induct d rule: alg_defl.principal_induct)
+apply (induct d rule: sfp.principal_induct)
 apply (rule adm_subst [OF _ adm_deflation], simp)
-apply (simp add: cast_alg_defl_principal)
+apply (simp add: cast_sfp_principal)
 apply (rule finite_deflation_imp_deflation)
 apply (rule finite_deflation_Rep_fin_defl)
 done
 
 lemma finite_deflation_cast:
   "compact d \<Longrightarrow> finite_deflation (cast\<cdot>d)"
-apply (drule alg_defl.compact_imp_principal, clarify)
-apply (simp add: cast_alg_defl_principal)
+apply (drule sfp.compact_imp_principal, clarify)
+apply (simp add: cast_sfp_principal)
 apply (rule finite_deflation_Rep_fin_defl)
 done
 
@@ -608,43 +201,33 @@
 
 declare cast.idem [simp]
 
-lemma cast_approx: "cast\<cdot>(approx n\<cdot>A) = defl_approx n (cast\<cdot>A)"
-apply (rule alg_defl.principal_induct)
-apply (rule adm_eq)
-apply simp
-apply (simp add: cont2cont_defl_approx cast.below)
-apply (simp only: approx_alg_defl_principal)
-apply (simp only: cast_alg_defl_principal)
-apply (simp only: Rep_fin_defl_fd_take)
+lemma compact_cast [simp]: "compact d \<Longrightarrow> compact (cast\<cdot>d)"
+apply (rule finite_deflation_imp_compact)
+apply (erule finite_deflation_cast)
 done
 
-lemma cast_approx_fixed_iff:
-  "cast\<cdot>(approx i\<cdot>A)\<cdot>x = x \<longleftrightarrow> approx i\<cdot>x = x \<and> cast\<cdot>A\<cdot>x = x"
-apply (simp only: cast_approx)
-apply (rule defl_approx_fixed_iff)
-apply (rule deflation_cast)
+lemma cast_below_cast: "cast\<cdot>A \<sqsubseteq> cast\<cdot>B \<longleftrightarrow> A \<sqsubseteq> B"
+apply (induct A rule: sfp.principal_induct, simp)
+apply (induct B rule: sfp.principal_induct, simp)
+apply (simp add: cast_sfp_principal below_fin_defl_def)
 done
 
-lemma defl_approx_cast: "defl_approx i (cast\<cdot>A) = cast\<cdot>(approx i\<cdot>A)"
-by (rule cast_approx [symmetric])
-
-lemma cast_below_cast_iff: "cast\<cdot>A \<sqsubseteq> cast\<cdot>B \<longleftrightarrow> A \<sqsubseteq> B"
-apply (induct A rule: alg_defl.principal_induct, simp)
-apply (induct B rule: alg_defl.principal_induct)
-apply (simp add: cast_alg_defl_principal)
-apply (simp add: finite_deflation_imp_compact finite_deflation_Rep_fin_defl)
-apply (simp add: cast_alg_defl_principal below_fin_defl_def)
+lemma compact_cast_iff: "compact (cast\<cdot>d) \<longleftrightarrow> compact d"
+apply (rule iffI)
+apply (simp only: compact_def cast_below_cast [symmetric])
+apply (erule adm_subst [OF cont_Rep_CFun2])
+apply (erule compact_cast)
 done
 
 lemma cast_below_imp_below: "cast\<cdot>A \<sqsubseteq> cast\<cdot>B \<Longrightarrow> A \<sqsubseteq> B"
-by (simp only: cast_below_cast_iff)
+by (simp only: cast_below_cast)
 
 lemma cast_eq_imp_eq: "cast\<cdot>A = cast\<cdot>B \<Longrightarrow> A = B"
 by (simp add: below_antisym cast_below_imp_below)
 
 lemma cast_strict1 [simp]: "cast\<cdot>\<bottom> = \<bottom>"
-apply (subst inst_alg_defl_pcpo)
-apply (subst cast_alg_defl_principal)
+apply (subst inst_sfp_pcpo)
+apply (subst cast_sfp_principal)
 apply (rule Abs_fin_defl_inverse)
 apply (simp add: finite_deflation_UU)
 done
@@ -652,119 +235,271 @@
 lemma cast_strict2 [simp]: "cast\<cdot>A\<cdot>\<bottom> = \<bottom>"
 by (rule cast.below [THEN UU_I])
 
-
 subsection {* Deflation membership relation *}
 
 definition
-  in_deflation :: "'a::profinite \<Rightarrow> 'a alg_defl \<Rightarrow> bool" (infixl ":::" 50)
+  in_sfp :: "udom \<Rightarrow> sfp \<Rightarrow> bool" (infixl ":::" 50)
 where
   "x ::: A \<longleftrightarrow> cast\<cdot>A\<cdot>x = x"
 
-lemma adm_in_deflation: "adm (\<lambda>x. x ::: A)"
-unfolding in_deflation_def by simp
+lemma adm_in_sfp: "adm (\<lambda>x. x ::: A)"
+unfolding in_sfp_def by simp
 
-lemma in_deflationI: "cast\<cdot>A\<cdot>x = x \<Longrightarrow> x ::: A"
-unfolding in_deflation_def .
+lemma in_sfpI: "cast\<cdot>A\<cdot>x = x \<Longrightarrow> x ::: A"
+unfolding in_sfp_def .
 
 lemma cast_fixed: "x ::: A \<Longrightarrow> cast\<cdot>A\<cdot>x = x"
-unfolding in_deflation_def .
+unfolding in_sfp_def .
 
-lemma cast_in_deflation [simp]: "cast\<cdot>A\<cdot>x ::: A"
-unfolding in_deflation_def by (rule cast.idem)
+lemma cast_in_sfp [simp]: "cast\<cdot>A\<cdot>x ::: A"
+unfolding in_sfp_def by (rule cast.idem)
 
-lemma bottom_in_deflation [simp]: "\<bottom> ::: A"
-unfolding in_deflation_def by (rule cast_strict2)
+lemma bottom_in_sfp [simp]: "\<bottom> ::: A"
+unfolding in_sfp_def by (rule cast_strict2)
 
-lemma subdeflationD: "A \<sqsubseteq> B \<Longrightarrow> x ::: A \<Longrightarrow> x ::: B"
-unfolding in_deflation_def
+lemma sfp_belowD: "A \<sqsubseteq> B \<Longrightarrow> x ::: A \<Longrightarrow> x ::: B"
+unfolding in_sfp_def
  apply (rule deflation.belowD)
    apply (rule deflation_cast)
   apply (erule monofun_cfun_arg)
  apply assumption
 done
 
-lemma rev_subdeflationD: "x ::: A \<Longrightarrow> A \<sqsubseteq> B \<Longrightarrow> x ::: B"
-by (rule subdeflationD)
+lemma rev_sfp_belowD: "x ::: A \<Longrightarrow> A \<sqsubseteq> B \<Longrightarrow> x ::: B"
+by (rule sfp_belowD)
 
-lemma subdeflationI: "(\<And>x. x ::: A \<Longrightarrow> x ::: B) \<Longrightarrow> A \<sqsubseteq> B"
+lemma sfp_belowI: "(\<And>x. x ::: A \<Longrightarrow> x ::: B) \<Longrightarrow> A \<sqsubseteq> B"
 apply (rule cast_below_imp_below)
 apply (rule cast.belowI)
-apply (simp add: in_deflation_def)
-done
-
-text "Identity deflation:"
-
-lemma "cast\<cdot>(\<Squnion>i. alg_defl_principal (Abs_fin_defl (approx i)))\<cdot>x = x"
-apply (subst contlub_cfun_arg)
-apply (rule chainI)
-apply (rule alg_defl.principal_mono)
-apply (rule Abs_fin_defl_mono)
-apply (rule finite_deflation_approx)
-apply (rule finite_deflation_approx)
-apply (rule chainE)
-apply (rule chain_approx)
-apply (simp add: cast_alg_defl_principal
-  Abs_fin_defl_inverse finite_deflation_approx)
+apply (simp add: in_sfp_def)
 done
 
-subsection {* Bifinite domains and algebraic deflations *}
+subsection {* Class of SFP domains *}
+
+text {*
+  We define a SFP domain as a pcpo that is isomorphic to some
+  algebraic deflation over the universal domain.
+*}
+
+class sfp = pcpo +
+  fixes emb :: "'a::pcpo \<rightarrow> udom"
+  fixes prj :: "udom \<rightarrow> 'a::pcpo"
+  fixes sfp :: "'a itself \<Rightarrow> sfp"
+  assumes ep_pair_emb_prj: "ep_pair emb prj"
+  assumes cast_SFP: "cast\<cdot>(sfp TYPE('a)) = emb oo prj"
+
+syntax "_SFP" :: "type \<Rightarrow> sfp"  ("(1SFP/(1'(_')))")
+translations "SFP('t)" \<rightleftharpoons> "CONST sfp TYPE('t)"
+
+interpretation sfp:
+  pcpo_ep_pair "emb :: 'a::sfp \<rightarrow> udom" "prj :: udom \<rightarrow> 'a::sfp"
+  unfolding pcpo_ep_pair_def
+  by (rule ep_pair_emb_prj)
+
+lemmas emb_inverse = sfp.e_inverse
+lemmas emb_prj_below = sfp.e_p_below
+lemmas emb_eq_iff = sfp.e_eq_iff
+lemmas emb_strict = sfp.e_strict
+lemmas prj_strict = sfp.p_strict
+
+subsection {* SFP domains have a countable compact basis *}
+
+text {*
+  Eventually it should be possible to generalize this to an unpointed
+  variant of the sfp class.
+*}
 
-text {* This lemma says that if we have an ep-pair from
-a bifinite domain into a universal domain, then e oo p
-is an algebraic deflation. *}
+interpretation compact_basis:
+  ideal_completion below Rep_compact_basis "approximants::'a::sfp \<Rightarrow> _"
+proof -
+  obtain Y where Y: "\<forall>i. Y i \<sqsubseteq> Y (Suc i)"
+  and SFP: "SFP('a) = (\<Squnion>i. sfp_principal (Y i))"
+    by (rule sfp.obtain_principal_chain)
+  def approx \<equiv> "\<lambda>i. (prj oo cast\<cdot>(sfp_principal (Y i)) oo emb) :: 'a \<rightarrow> 'a"
+  interpret sfp_approx: approx_chain approx
+  proof (rule approx_chain.intro)
+    show "chain (\<lambda>i. approx i)"
+      unfolding approx_def by (simp add: Y)
+    show "(\<Squnion>i. approx i) = ID"
+      unfolding approx_def
+      by (simp add: lub_distribs Y SFP [symmetric] cast_SFP expand_cfun_eq)
+    show "\<And>i. finite_deflation (approx i)"
+      unfolding approx_def
+      apply (rule sfp.finite_deflation_p_d_e)
+      apply (rule finite_deflation_cast)
+      apply (rule sfp.compact_principal)
+      apply (rule below_trans [OF monofun_cfun_fun])
+      apply (rule is_ub_thelub, simp add: Y)
+      apply (simp add: lub_distribs Y SFP [symmetric] cast_SFP)
+      done
+  qed
+  (* FIXME: why does show ?thesis fail here? *)
+  show "ideal_completion below Rep_compact_basis (approximants::'a \<Rightarrow> _)" ..
+qed
+
+subsection {* Type combinators *}
+
+definition
+  sfp_fun1 ::
+    "(nat \<Rightarrow> 'a \<rightarrow> 'a) \<Rightarrow> ((udom \<rightarrow> udom) \<rightarrow> ('a \<rightarrow> 'a)) \<Rightarrow> (sfp \<rightarrow> sfp)"
+where
+  "sfp_fun1 approx f =
+    sfp.basis_fun (\<lambda>a.
+      sfp_principal (Abs_fin_defl
+        (udom_emb approx oo f\<cdot>(Rep_fin_defl a) oo udom_prj approx)))"
 
-lemma
-  assumes "ep_pair e p"
-  constrains e :: "'a::profinite \<rightarrow> 'b::profinite"
-  shows "\<exists>d. cast\<cdot>d = e oo p"
-proof
-  interpret ep_pair e p by fact
-  let ?a = "\<lambda>i. e oo approx i oo p"
-  have a: "\<And>i. finite_deflation (?a i)"
-    apply (rule finite_deflation_e_d_p)
-    apply (rule finite_deflation_approx)
+definition
+  sfp_fun2 ::
+    "(nat \<Rightarrow> 'a \<rightarrow> 'a) \<Rightarrow> ((udom \<rightarrow> udom) \<rightarrow> (udom \<rightarrow> udom) \<rightarrow> ('a \<rightarrow> 'a))
+      \<Rightarrow> (sfp \<rightarrow> sfp \<rightarrow> sfp)"
+where
+  "sfp_fun2 approx f =
+    sfp.basis_fun (\<lambda>a.
+      sfp.basis_fun (\<lambda>b.
+        sfp_principal (Abs_fin_defl
+          (udom_emb approx oo
+            f\<cdot>(Rep_fin_defl a)\<cdot>(Rep_fin_defl b) oo udom_prj approx))))"
+
+lemma cast_sfp_fun1:
+  assumes approx: "approx_chain approx"
+  assumes f: "\<And>a. finite_deflation a \<Longrightarrow> finite_deflation (f\<cdot>a)"
+  shows "cast\<cdot>(sfp_fun1 approx f\<cdot>A) = udom_emb approx oo f\<cdot>(cast\<cdot>A) oo udom_prj approx"
+proof -
+  have 1: "\<And>a. finite_deflation
+        (udom_emb approx oo f\<cdot>(Rep_fin_defl a) oo udom_prj approx)"
+    apply (rule ep_pair.finite_deflation_e_d_p)
+    apply (rule approx_chain.ep_pair_udom [OF approx])
+    apply (rule f, rule finite_deflation_Rep_fin_defl)
     done
-  let ?d = "\<Squnion>i. alg_defl_principal (Abs_fin_defl (?a i))"
-  show "cast\<cdot>?d = e oo p"
+  show ?thesis
+    by (induct A rule: sfp.principal_induct, simp)
+       (simp only: sfp_fun1_def
+                   sfp.basis_fun_principal
+                   sfp.basis_fun_mono
+                   sfp.principal_mono
+                   Abs_fin_defl_mono [OF 1 1]
+                   monofun_cfun below_refl
+                   Rep_fin_defl_mono
+                   cast_sfp_principal
+                   Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1])
+qed
+
+lemma cast_sfp_fun2:
+  assumes approx: "approx_chain approx"
+  assumes f: "\<And>a b. finite_deflation a \<Longrightarrow> finite_deflation b \<Longrightarrow>
+                finite_deflation (f\<cdot>a\<cdot>b)"
+  shows "cast\<cdot>(sfp_fun2 approx f\<cdot>A\<cdot>B) =
+    udom_emb approx oo f\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj approx"
+proof -
+  have 1: "\<And>a b. finite_deflation (udom_emb approx oo
+      f\<cdot>(Rep_fin_defl a)\<cdot>(Rep_fin_defl b) oo udom_prj approx)"
+    apply (rule ep_pair.finite_deflation_e_d_p)
+    apply (rule ep_pair_udom [OF approx])
+    apply (rule f, (rule finite_deflation_Rep_fin_defl)+)
+    done
+  show ?thesis
+    by (induct A B rule: sfp.principal_induct2, simp, simp)
+       (simp only: sfp_fun2_def
+                   sfp.basis_fun_principal
+                   sfp.basis_fun_mono
+                   sfp.principal_mono
+                   Abs_fin_defl_mono [OF 1 1]
+                   monofun_cfun below_refl
+                   Rep_fin_defl_mono
+                   cast_sfp_principal
+                   Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1])
+qed
+
+subsection {* Instance for universal domain type *}
+
+instantiation udom :: sfp
+begin
+
+definition [simp]:
+  "emb = (ID :: udom \<rightarrow> udom)"
+
+definition [simp]:
+  "prj = (ID :: udom \<rightarrow> udom)"
+
+definition
+  "sfp (t::udom itself) = (\<Squnion>i. sfp_principal (Abs_fin_defl (udom_approx i)))"
+
+instance proof
+  show "ep_pair emb (prj :: udom \<rightarrow> udom)"
+    by (simp add: ep_pair.intro)
+next
+  show "cast\<cdot>SFP(udom) = emb oo (prj :: udom \<rightarrow> udom)"
+    unfolding sfp_udom_def
     apply (subst contlub_cfun_arg)
     apply (rule chainI)
-    apply (rule alg_defl.principal_mono)
-    apply (rule Abs_fin_defl_mono [OF a a])
-    apply (rule chainE, simp)
-    apply (subst cast_alg_defl_principal)
-    apply (simp add: Abs_fin_defl_inverse a)
-    apply (simp add: expand_cfun_eq lub_distribs)
-    done
-qed
-
-text {* This lemma says that if we have an ep-pair
-from a cpo into a bifinite domain, and e oo p is
-an algebraic deflation, then the cpo is bifinite. *}
-
-lemma
-  assumes "ep_pair e p"
-  constrains e :: "'a::cpo \<rightarrow> 'b::profinite"
-  assumes d: "\<And>x. cast\<cdot>d\<cdot>x = e\<cdot>(p\<cdot>x)"
-  obtains a :: "nat \<Rightarrow> 'a \<rightarrow> 'a" where
-    "\<And>i. finite_deflation (a i)"
-    "(\<Squnion>i. a i) = ID"
-proof
-  interpret ep_pair e p by fact
-  let ?a = "\<lambda>i. p oo cast\<cdot>(approx i\<cdot>d) oo e"
-  show "\<And>i. finite_deflation (?a i)"
-    apply (rule finite_deflation_p_d_e)
-    apply (rule finite_deflation_cast)
-    apply (rule compact_approx)
-    apply (rule below_eq_trans [OF _ d])
-    apply (rule monofun_cfun_fun)
-    apply (rule monofun_cfun_arg)
-    apply (rule approx_below)
-    done
-  show "(\<Squnion>i. ?a i) = ID"
-    apply (rule ext_cfun, simp)
-    apply (simp add: lub_distribs)
-    apply (simp add: d)
+    apply (rule sfp.principal_mono)
+    apply (simp add: below_fin_defl_def)
+    apply (simp add: Abs_fin_defl_inverse finite_deflation_udom_approx)
+    apply (rule chainE)
+    apply (rule chain_udom_approx)
+    apply (subst cast_sfp_principal)
+    apply (simp add: Abs_fin_defl_inverse finite_deflation_udom_approx)
     done
 qed
 
 end
+
+subsection {* Instance for continuous function space *}
+
+definition
+  cfun_approx :: "nat \<Rightarrow> (udom \<rightarrow> udom) \<rightarrow> (udom \<rightarrow> udom)"
+where
+  "cfun_approx = (\<lambda>i. cfun_map\<cdot>(udom_approx i)\<cdot>(udom_approx i))"
+
+lemma cfun_approx: "approx_chain cfun_approx"
+proof (rule approx_chain.intro)
+  show "chain (\<lambda>i. cfun_approx i)"
+    unfolding cfun_approx_def by simp
+  show "(\<Squnion>i. cfun_approx i) = ID"
+    unfolding cfun_approx_def
+    by (simp add: lub_distribs cfun_map_ID)
+  show "\<And>i. finite_deflation (cfun_approx i)"
+    unfolding cfun_approx_def
+    by (intro finite_deflation_cfun_map finite_deflation_udom_approx)
+qed
+
+definition cfun_sfp :: "sfp \<rightarrow> sfp \<rightarrow> sfp"
+where "cfun_sfp = sfp_fun2 cfun_approx cfun_map"
+
+lemma cast_cfun_sfp:
+  "cast\<cdot>(cfun_sfp\<cdot>A\<cdot>B) =
+    udom_emb cfun_approx oo cfun_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj cfun_approx"
+unfolding cfun_sfp_def
+apply (rule cast_sfp_fun2 [OF cfun_approx])
+apply (erule (1) finite_deflation_cfun_map)
+done
+
+instantiation cfun :: (sfp, sfp) sfp
+begin
+
+definition
+  "emb = udom_emb cfun_approx oo cfun_map\<cdot>prj\<cdot>emb"
+
+definition
+  "prj = cfun_map\<cdot>emb\<cdot>prj oo udom_prj cfun_approx"
+
+definition
+  "sfp (t::('a \<rightarrow> 'b) itself) = cfun_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
+
+instance proof
+  show "ep_pair emb (prj :: udom \<rightarrow> 'a \<rightarrow> 'b)"
+    unfolding emb_cfun_def prj_cfun_def
+    using ep_pair_udom [OF cfun_approx]
+    by (intro ep_pair_comp ep_pair_cfun_map ep_pair_emb_prj)
+next
+  show "cast\<cdot>SFP('a \<rightarrow> 'b) = emb oo (prj :: udom \<rightarrow> 'a \<rightarrow> 'b)"
+    unfolding emb_cfun_def prj_cfun_def sfp_cfun_def cast_cfun_sfp
+    by (simp add: cast_SFP oo_def expand_cfun_eq cfun_map_map)
+qed
+
+end
+
+lemma SFP_cfun: "SFP('a::sfp \<rightarrow> 'b::sfp) = cfun_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
+by (rule sfp_cfun_def)
+
+end
--- a/src/HOLCF/Bifinite.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Bifinite.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -8,103 +8,7 @@
 imports Deflation
 begin
 
-subsection {* Omega-profinite and bifinite domains *}
-
-class profinite =
-  fixes approx :: "nat \<Rightarrow> 'a \<rightarrow> 'a"
-  assumes chain_approx [simp]: "chain approx"
-  assumes lub_approx_app [simp]: "(\<Squnion>i. approx i\<cdot>x) = x"
-  assumes approx_idem: "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
-  assumes finite_fixes_approx: "finite {x. approx i\<cdot>x = x}"
-
-class bifinite = profinite + pcpo
-
-lemma approx_below: "approx i\<cdot>x \<sqsubseteq> x"
-proof -
-  have "chain (\<lambda>i. approx i\<cdot>x)" by simp
-  hence "approx i\<cdot>x \<sqsubseteq> (\<Squnion>i. approx i\<cdot>x)" by (rule is_ub_thelub)
-  thus "approx i\<cdot>x \<sqsubseteq> x" by simp
-qed
-
-lemma finite_deflation_approx: "finite_deflation (approx i)"
-proof
-  fix x :: 'a
-  show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
-    by (rule approx_idem)
-  show "approx i\<cdot>x \<sqsubseteq> x"
-    by (rule approx_below)
-  show "finite {x. approx i\<cdot>x = x}"
-    by (rule finite_fixes_approx)
-qed
-
-interpretation approx: finite_deflation "approx i"
-by (rule finite_deflation_approx)
-
-lemma (in deflation) deflation: "deflation d" ..
-
-lemma deflation_approx: "deflation (approx i)"
-by (rule approx.deflation)
-
-lemma lub_approx [simp]: "(\<Squnion>i. approx i) = (\<Lambda> x. x)"
-by (rule ext_cfun, simp add: contlub_cfun_fun)
-
-lemma approx_strict [simp]: "approx i\<cdot>\<bottom> = \<bottom>"
-by (rule UU_I, rule approx_below)
-
-lemma approx_approx1:
-  "i \<le> j \<Longrightarrow> approx i\<cdot>(approx j\<cdot>x) = approx i\<cdot>x"
-apply (rule deflation_below_comp1 [OF deflation_approx deflation_approx])
-apply (erule chain_mono [OF chain_approx])
-done
-
-lemma approx_approx2:
-  "j \<le> i \<Longrightarrow> approx i\<cdot>(approx j\<cdot>x) = approx j\<cdot>x"
-apply (rule deflation_below_comp2 [OF deflation_approx deflation_approx])
-apply (erule chain_mono [OF chain_approx])
-done
-
-lemma approx_approx [simp]:
-  "approx i\<cdot>(approx j\<cdot>x) = approx (min i j)\<cdot>x"
-apply (rule_tac x=i and y=j in linorder_le_cases)
-apply (simp add: approx_approx1 min_def)
-apply (simp add: approx_approx2 min_def)
-done
-
-lemma finite_image_approx: "finite ((\<lambda>x. approx n\<cdot>x) ` A)"
-by (rule approx.finite_image)
-
-lemma finite_range_approx: "finite (range (\<lambda>x. approx i\<cdot>x))"
-by (rule approx.finite_range)
-
-lemma compact_approx [simp]: "compact (approx n\<cdot>x)"
-by (rule approx.compact)
-
-lemma profinite_compact_eq_approx: "compact x \<Longrightarrow> \<exists>i. approx i\<cdot>x = x"
-by (rule admD2, simp_all)
-
-lemma profinite_compact_iff: "compact x \<longleftrightarrow> (\<exists>n. approx n\<cdot>x = x)"
- apply (rule iffI)
-  apply (erule profinite_compact_eq_approx)
- apply (erule exE)
- apply (erule subst)
- apply (rule compact_approx)
-done
-
-lemma approx_induct:
-  assumes adm: "adm P" and P: "\<And>n x. P (approx n\<cdot>x)"
-  shows "P x"
-proof -
-  have "P (\<Squnion>n. approx n\<cdot>x)"
-    by (rule admD [OF adm], simp, simp add: P)
-  thus "P x" by simp
-qed
-
-lemma profinite_below_ext: "(\<And>i. approx i\<cdot>x \<sqsubseteq> approx i\<cdot>y) \<Longrightarrow> x \<sqsubseteq> y"
-apply (subgoal_tac "(\<Squnion>i. approx i\<cdot>x) \<sqsubseteq> (\<Squnion>i. approx i\<cdot>y)", simp)
-apply (rule lub_mono, simp, simp, simp)
-done
-
-subsection {* Instance for product type *}
+subsection {* Map operator for product type *}
 
 definition
   cprod_map :: "('a \<rightarrow> 'b) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> 'a \<times> 'c \<rightarrow> 'b \<times> 'd"
@@ -161,46 +65,7 @@
     by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
 qed
 
-instantiation prod :: (profinite, profinite) profinite
-begin
-
-definition
-  approx_prod_def:
-    "approx = (\<lambda>n. cprod_map\<cdot>(approx n)\<cdot>(approx n))"
-
-instance proof
-  fix i :: nat and x :: "'a \<times> 'b"
-  show "chain (approx :: nat \<Rightarrow> 'a \<times> 'b \<rightarrow> 'a \<times> 'b)"
-    unfolding approx_prod_def by simp
-  show "(\<Squnion>i. approx i\<cdot>x) = x"
-    unfolding approx_prod_def cprod_map_def
-    by (simp add: lub_distribs thelub_Pair)
-  show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
-    unfolding approx_prod_def cprod_map_def by simp
-  have "{x::'a \<times> 'b. approx i\<cdot>x = x} \<subseteq>
-        {x::'a. approx i\<cdot>x = x} \<times> {x::'b. approx i\<cdot>x = x}"
-    unfolding approx_prod_def by clarsimp
-  thus "finite {x::'a \<times> 'b. approx i\<cdot>x = x}"
-    by (rule finite_subset,
-        intro finite_cartesian_product finite_fixes_approx)
-qed
-
-end
-
-instance prod :: (bifinite, bifinite) bifinite ..
-
-lemma approx_Pair [simp]:
-  "approx i\<cdot>(x, y) = (approx i\<cdot>x, approx i\<cdot>y)"
-unfolding approx_prod_def by simp
-
-lemma fst_approx: "fst (approx i\<cdot>p) = approx i\<cdot>(fst p)"
-by (induct p, simp)
-
-lemma snd_approx: "snd (approx i\<cdot>p) = approx i\<cdot>(snd p)"
-by (induct p, simp)
-
-
-subsection {* Instance for continuous function space *}
+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)"
@@ -305,39 +170,4 @@
 apply (simp only: finite_deflation_cfun_map)
 done
 
-instantiation cfun :: (profinite, profinite) profinite
-begin
-
-definition
-  approx_cfun_def:
-    "approx = (\<lambda>n. cfun_map\<cdot>(approx n)\<cdot>(approx n))"
-
-instance proof
-  show "chain (approx :: nat \<Rightarrow> ('a \<rightarrow> 'b) \<rightarrow> ('a \<rightarrow> 'b))"
-    unfolding approx_cfun_def by simp
-next
-  fix x :: "'a \<rightarrow> 'b"
-  show "(\<Squnion>i. approx i\<cdot>x) = x"
-    unfolding approx_cfun_def cfun_map_def
-    by (simp add: lub_distribs eta_cfun)
-next
-  fix i :: nat and x :: "'a \<rightarrow> 'b"
-  show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
-    unfolding approx_cfun_def cfun_map_def by simp
-next
-  fix i :: nat
-  show "finite {x::'a \<rightarrow> 'b. approx i\<cdot>x = x}"
-    unfolding approx_cfun_def
-    by (intro finite_deflation.finite_fixes
-              finite_deflation_cfun_map
-              finite_deflation_approx)
-qed
-
 end
-
-instance cfun :: (profinite, bifinite) bifinite ..
-
-lemma approx_cfun: "approx n\<cdot>f\<cdot>x = approx n\<cdot>(f\<cdot>(approx n\<cdot>x))"
-by (simp add: approx_cfun_def)
-
-end
--- a/src/HOLCF/CompactBasis.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/CompactBasis.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -2,174 +2,18 @@
     Author:     Brian Huffman
 *)
 
-header {* Compact bases of domains *}
+header {* A compact basis for powerdomains *}
 
 theory CompactBasis
-imports Completion Bifinite
-begin
-
-subsection {* Compact bases of bifinite domains *}
-
-default_sort profinite
-
-typedef (open) 'a compact_basis = "{x::'a::profinite. compact x}"
-by (fast intro: compact_approx)
-
-lemma compact_Rep_compact_basis: "compact (Rep_compact_basis a)"
-by (rule Rep_compact_basis [unfolded mem_Collect_eq])
-
-instantiation compact_basis :: (profinite) below
+imports Algebraic
 begin
 
-definition
-  compact_le_def:
-    "(op \<sqsubseteq>) \<equiv> (\<lambda>x y. Rep_compact_basis x \<sqsubseteq> Rep_compact_basis y)"
-
-instance ..
-
-end
-
-instance compact_basis :: (profinite) po
-by (rule typedef_po
-    [OF type_definition_compact_basis compact_le_def])
-
-text {* Take function for compact basis *}
-
-definition
-  compact_take :: "nat \<Rightarrow> 'a compact_basis \<Rightarrow> 'a compact_basis" where
-  "compact_take = (\<lambda>n a. Abs_compact_basis (approx n\<cdot>(Rep_compact_basis a)))"
-
-lemma Rep_compact_take:
-  "Rep_compact_basis (compact_take n a) = approx n\<cdot>(Rep_compact_basis a)"
-unfolding compact_take_def
-by (simp add: Abs_compact_basis_inverse)
-
-lemmas approx_Rep_compact_basis = Rep_compact_take [symmetric]
-
-interpretation compact_basis:
-  basis_take below compact_take
-proof
-  fix n :: nat and a :: "'a compact_basis"
-  show "compact_take n a \<sqsubseteq> a"
-    unfolding compact_le_def
-    by (simp add: Rep_compact_take approx_below)
-next
-  fix n :: nat and a :: "'a compact_basis"
-  show "compact_take n (compact_take n a) = compact_take n a"
-    by (simp add: Rep_compact_basis_inject [symmetric] Rep_compact_take)
-next
-  fix n :: nat and a b :: "'a compact_basis"
-  assume "a \<sqsubseteq> b" thus "compact_take n a \<sqsubseteq> compact_take n b"
-    unfolding compact_le_def Rep_compact_take
-    by (rule monofun_cfun_arg)
-next
-  fix n :: nat and a :: "'a compact_basis"
-  show "\<And>n a. compact_take n a \<sqsubseteq> compact_take (Suc n) a"
-    unfolding compact_le_def Rep_compact_take
-    by (rule chainE, simp)
-next
-  fix n :: nat
-  show "finite (range (compact_take n))"
-    apply (rule finite_imageD [where f="Rep_compact_basis"])
-    apply (rule finite_subset [where B="range (\<lambda>x. approx n\<cdot>x)"])
-    apply (clarsimp simp add: Rep_compact_take)
-    apply (rule finite_range_approx)
-    apply (rule inj_onI, simp add: Rep_compact_basis_inject)
-    done
-next
-  fix a :: "'a compact_basis"
-  show "\<exists>n. compact_take n a = a"
-    apply (simp add: Rep_compact_basis_inject [symmetric])
-    apply (simp add: Rep_compact_take)
-    apply (rule profinite_compact_eq_approx)
-    apply (rule compact_Rep_compact_basis)
-    done
-qed
-
-text {* Ideal completion *}
-
-definition
-  approximants :: "'a \<Rightarrow> 'a compact_basis set" where
-  "approximants = (\<lambda>x. {a. Rep_compact_basis a \<sqsubseteq> x})"
-
-interpretation compact_basis:
-  ideal_completion below compact_take Rep_compact_basis approximants
-proof
-  fix w :: 'a
-  show "preorder.ideal below (approximants w)"
-  proof (rule below.idealI)
-    show "\<exists>x. x \<in> approximants w"
-      unfolding approximants_def
-      apply (rule_tac x="Abs_compact_basis (approx 0\<cdot>w)" in exI)
-      apply (simp add: Abs_compact_basis_inverse approx_below)
-      done
-  next
-    fix x y :: "'a compact_basis"
-    assume "x \<in> approximants w" "y \<in> approximants w"
-    thus "\<exists>z \<in> approximants w. x \<sqsubseteq> z \<and> y \<sqsubseteq> z"
-      unfolding approximants_def
-      apply simp
-      apply (cut_tac a=x in compact_Rep_compact_basis)
-      apply (cut_tac a=y in compact_Rep_compact_basis)
-      apply (drule profinite_compact_eq_approx)
-      apply (drule profinite_compact_eq_approx)
-      apply (clarify, rename_tac i j)
-      apply (rule_tac x="Abs_compact_basis (approx (max i j)\<cdot>w)" in exI)
-      apply (simp add: compact_le_def)
-      apply (simp add: Abs_compact_basis_inverse approx_below)
-      apply (erule subst, erule subst)
-      apply (simp add: monofun_cfun chain_mono [OF chain_approx])
-      done
-  next
-    fix x y :: "'a compact_basis"
-    assume "x \<sqsubseteq> y" "y \<in> approximants w" thus "x \<in> approximants w"
-      unfolding approximants_def
-      apply simp
-      apply (simp add: compact_le_def)
-      apply (erule (1) below_trans)
-      done
-  qed
-next
-  fix Y :: "nat \<Rightarrow> 'a"
-  assume Y: "chain Y"
-  show "approximants (\<Squnion>i. Y i) = (\<Union>i. approximants (Y i))"
-    unfolding approximants_def
-    apply safe
-    apply (simp add: compactD2 [OF compact_Rep_compact_basis Y])
-    apply (erule below_trans, rule is_ub_thelub [OF Y])
-    done
-next
-  fix a :: "'a compact_basis"
-  show "approximants (Rep_compact_basis a) = {b. b \<sqsubseteq> a}"
-    unfolding approximants_def compact_le_def ..
-next
-  fix x y :: "'a"
-  assume "approximants x \<subseteq> approximants y" thus "x \<sqsubseteq> y"
-    apply (subgoal_tac "(\<Squnion>i. approx i\<cdot>x) \<sqsubseteq> y", simp)
-    apply (rule admD, simp, simp)
-    apply (drule_tac c="Abs_compact_basis (approx i\<cdot>x)" in subsetD)
-    apply (simp add: approximants_def Abs_compact_basis_inverse approx_below)
-    apply (simp add: approximants_def Abs_compact_basis_inverse)
-    done
-qed
-
-text {* minimal compact element *}
-
-definition
-  compact_bot :: "'a::bifinite compact_basis" where
-  "compact_bot = Abs_compact_basis \<bottom>"
-
-lemma Rep_compact_bot: "Rep_compact_basis compact_bot = \<bottom>"
-unfolding compact_bot_def by (simp add: Abs_compact_basis_inverse)
-
-lemma compact_bot_minimal [simp]: "compact_bot \<sqsubseteq> a"
-unfolding compact_le_def Rep_compact_bot by simp
-
+default_sort sfp
 
 subsection {* A compact basis for powerdomains *}
 
 typedef 'a pd_basis =
-  "{S::'a::profinite compact_basis set. finite S \<and> S \<noteq> {}}"
+  "{S::'a compact_basis set. finite S \<and> S \<noteq> {}}"
 by (rule_tac x="{arbitrary}" in exI, simp)
 
 lemma finite_Rep_pd_basis [simp]: "finite (Rep_pd_basis u)"
@@ -178,7 +22,21 @@
 lemma Rep_pd_basis_nonempty [simp]: "Rep_pd_basis u \<noteq> {}"
 by (insert Rep_pd_basis [of u, unfolded pd_basis_def]) simp
 
-text {* unit and plus *}
+text {* The powerdomain basis type is countable. *}
+
+lemma pd_basis_countable: "\<exists>f::'a::sfp pd_basis \<Rightarrow> nat. inj f"
+proof -
+  obtain g :: "'a compact_basis \<Rightarrow> nat" where "inj g"
+    using compact_basis.countable ..
+  hence image_g_eq: "\<And>A B. g ` A = g ` B \<longleftrightarrow> A = B"
+    by (rule inj_image_eq_iff)
+  have "inj (\<lambda>t. set_encode (g ` Rep_pd_basis t))"
+    by (simp add: inj_on_def set_encode_eq image_g_eq Rep_pd_basis_inject)
+  thus ?thesis by - (rule exI)
+  (* FIXME: why doesn't ".." or "by (rule exI)" work? *)
+qed
+
+subsection {* Unit and plus constructors *}
 
 definition
   PDUnit :: "'a compact_basis \<Rightarrow> 'a pd_basis" where
@@ -229,7 +87,7 @@
 apply (rule PDUnit, erule PDPlus [OF PDUnit])
 done
 
-text {* fold-pd *}
+subsection {* Fold operator *}
 
 definition
   fold_pd ::
@@ -250,52 +108,26 @@
     by (simp add: image_Un fold1_Un2)
 qed
 
-text {* Take function for powerdomain basis *}
-
-definition
-  pd_take :: "nat \<Rightarrow> 'a pd_basis \<Rightarrow> 'a pd_basis" where
-  "pd_take n = (\<lambda>t. Abs_pd_basis (compact_take n ` Rep_pd_basis t))"
+subsection {* Lemmas for proving if-and-only-if inequalities *}
 
-lemma Rep_pd_take:
-  "Rep_pd_basis (pd_take n t) = compact_take n ` Rep_pd_basis t"
-unfolding pd_take_def
-apply (rule Abs_pd_basis_inverse)
-apply (simp add: pd_basis_def)
-done
-
-lemma pd_take_simps [simp]:
-  "pd_take n (PDUnit a) = PDUnit (compact_take n a)"
-  "pd_take n (PDPlus t u) = PDPlus (pd_take n t) (pd_take n u)"
-apply (simp_all add: Rep_pd_basis_inject [symmetric])
-apply (simp_all add: Rep_pd_take Rep_PDUnit Rep_PDPlus image_Un)
+lemma chain_max_below_iff:
+  assumes Y: "chain Y" shows "Y (max i j) \<sqsubseteq> x \<longleftrightarrow> Y i \<sqsubseteq> x \<and> Y j \<sqsubseteq> x"
+apply auto
+apply (erule below_trans [OF chain_mono [OF Y le_maxI1]])
+apply (erule below_trans [OF chain_mono [OF Y le_maxI2]])
+apply (simp add: max_def)
 done
 
-lemma pd_take_idem: "pd_take n (pd_take n t) = pd_take n t"
-apply (induct t rule: pd_basis_induct)
-apply (simp add: compact_basis.take_take)
-apply simp
-done
-
-lemma finite_range_pd_take: "finite (range (pd_take n))"
-apply (rule finite_imageD [where f="Rep_pd_basis"])
-apply (rule finite_subset [where B="Pow (range (compact_take n))"])
-apply (clarsimp simp add: Rep_pd_take)
-apply (simp add: compact_basis.finite_range_take)
-apply (rule inj_onI, simp add: Rep_pd_basis_inject)
-done
+lemma all_ex_below_disj_iff:
+  assumes "chain X" and "chain Y"
+  shows "(\<forall>i. \<exists>j. X i \<sqsubseteq> Z j \<or> Y i \<sqsubseteq> Z j) \<longleftrightarrow>
+         (\<forall>i. \<exists>j. X i \<sqsubseteq> Z j) \<or> (\<forall>i. \<exists>j. Y i \<sqsubseteq> Z j)"
+by (metis chain_max_below_iff assms)
 
-lemma pd_take_covers: "\<exists>n. pd_take n t = t"
-apply (subgoal_tac "\<exists>n. \<forall>m\<ge>n. pd_take m t = t", fast)
-apply (induct t rule: pd_basis_induct)
-apply (cut_tac a=a in compact_basis.take_covers)
-apply (clarify, rule_tac x=n in exI)
-apply (clarify, simp)
-apply (rule below_antisym)
-apply (rule compact_basis.take_less)
-apply (drule_tac a=a in compact_basis.take_chain_le)
-apply simp
-apply (clarify, rename_tac i j)
-apply (rule_tac x="max i j" in exI, simp)
-done
+lemma all_ex_below_conj_iff:
+  assumes "chain X" and "chain Y" and "chain Z"
+  shows "(\<forall>i. \<exists>j. X i \<sqsubseteq> Z j \<and> Y i \<sqsubseteq> Z j) \<longleftrightarrow>
+         (\<forall>i. \<exists>j. X i \<sqsubseteq> Z j) \<and> (\<forall>i. \<exists>j. Y i \<sqsubseteq> Z j)"
+oops
 
 end
--- a/src/HOLCF/Completion.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Completion.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -2,7 +2,7 @@
     Author:     Brian Huffman
 *)
 
-header {* Defining bifinite domains by ideal completion *}
+header {* Defining algebraic domains by ideal completion *}
 
 theory Completion
 imports Cfun
@@ -164,64 +164,27 @@
 
 subsection {* Lemmas about least upper bounds *}
 
-lemma finite_directed_contains_lub:
-  "\<lbrakk>finite S; directed S\<rbrakk> \<Longrightarrow> \<exists>u\<in>S. S <<| u"
-apply (drule (1) directed_finiteD, rule subset_refl)
-apply (erule bexE)
-apply (rule rev_bexI, assumption)
-apply (erule (1) is_lub_maximal)
-done
-
-lemma lub_finite_directed_in_self:
-  "\<lbrakk>finite S; directed S\<rbrakk> \<Longrightarrow> lub S \<in> S"
-apply (drule (1) finite_directed_contains_lub, clarify)
-apply (drule thelubI, simp)
-done
-
-lemma finite_directed_has_lub: "\<lbrakk>finite S; directed S\<rbrakk> \<Longrightarrow> \<exists>u. S <<| u"
-by (drule (1) finite_directed_contains_lub, fast)
-
-lemma is_ub_thelub0: "\<lbrakk>\<exists>u. S <<| u; x \<in> S\<rbrakk> \<Longrightarrow> x \<sqsubseteq> lub S"
+lemma is_ub_thelub_ex: "\<lbrakk>\<exists>u. S <<| u; x \<in> S\<rbrakk> \<Longrightarrow> x \<sqsubseteq> lub S"
 apply (erule exE, drule lubI)
 apply (drule is_lubD1)
 apply (erule (1) is_ubD)
 done
 
-lemma is_lub_thelub0: "\<lbrakk>\<exists>u. S <<| u; S <| x\<rbrakk> \<Longrightarrow> lub S \<sqsubseteq> x"
+lemma is_lub_thelub_ex: "\<lbrakk>\<exists>u. S <<| u; S <| x\<rbrakk> \<Longrightarrow> lub S \<sqsubseteq> x"
 by (erule exE, drule lubI, erule is_lub_lub)
 
 subsection {* Locale for ideal completion *}
 
-locale basis_take = preorder +
-  fixes take :: "nat \<Rightarrow> 'a::type \<Rightarrow> 'a"
-  assumes take_less: "take n a \<preceq> a"
-  assumes take_take: "take n (take n a) = take n a"
-  assumes take_mono: "a \<preceq> b \<Longrightarrow> take n a \<preceq> take n b"
-  assumes take_chain: "take n a \<preceq> take (Suc n) a"
-  assumes finite_range_take: "finite (range (take n))"
-  assumes take_covers: "\<exists>n. take n a = a"
-begin
-
-lemma take_chain_less: "m < n \<Longrightarrow> take m a \<preceq> take n a"
-by (erule less_Suc_induct, rule take_chain, erule (1) r_trans)
-
-lemma take_chain_le: "m \<le> n \<Longrightarrow> take m a \<preceq> take n a"
-by (cases "m = n", simp add: r_refl, simp add: take_chain_less)
-
-end
-
-locale ideal_completion = basis_take +
+locale ideal_completion = preorder +
   fixes principal :: "'a::type \<Rightarrow> 'b::cpo"
   fixes rep :: "'b::cpo \<Rightarrow> 'a::type set"
-  assumes ideal_rep: "\<And>x. preorder.ideal r (rep x)"
+  assumes ideal_rep: "\<And>x. ideal (rep x)"
   assumes rep_contlub: "\<And>Y. chain Y \<Longrightarrow> rep (\<Squnion>i. Y i) = (\<Union>i. rep (Y i))"
   assumes rep_principal: "\<And>a. rep (principal a) = {b. b \<preceq> a}"
   assumes subset_repD: "\<And>x y. rep x \<subseteq> rep y \<Longrightarrow> x \<sqsubseteq> y"
+  assumes countable: "\<exists>f::'a \<Rightarrow> nat. inj f"
 begin
 
-lemma finite_take_rep: "finite (take n ` rep x)"
-by (rule finite_subset [OF image_mono [OF subset_UNIV] finite_range_take])
-
 lemma rep_mono: "x \<sqsubseteq> y \<Longrightarrow> rep x \<subseteq> rep y"
 apply (frule bin_chain)
 apply (drule rep_contlub)
@@ -251,12 +214,19 @@
 lemma principal_eq_iff: "principal a = principal b \<longleftrightarrow> a \<preceq> b \<and> b \<preceq> a"
 unfolding po_eq_conv [where 'a='b] principal_below_iff ..
 
+lemma eq_iff: "x = y \<longleftrightarrow> rep x = rep y"
+unfolding po_eq_conv below_def by auto
+
 lemma repD: "a \<in> rep x \<Longrightarrow> principal a \<sqsubseteq> x"
 by (simp add: rep_eq)
 
 lemma principal_mono: "a \<preceq> b \<Longrightarrow> principal a \<sqsubseteq> principal b"
 by (simp only: principal_below_iff)
 
+lemma ch2ch_principal [simp]:
+  "\<forall>i. Y i \<preceq> Y (Suc i) \<Longrightarrow> chain (\<lambda>i. principal (Y i))"
+by (simp add: chainI principal_mono)
+
 lemma belowI: "(\<And>a. principal a \<sqsubseteq> x \<Longrightarrow> principal a \<sqsubseteq> u) \<Longrightarrow> x \<sqsubseteq> u"
 unfolding principal_below_iff_mem_rep
 by (simp add: below_def subset_eq)
@@ -271,68 +241,155 @@
 apply (simp add: rep_eq)
 done
 
+subsubsection {* Principal ideals approximate all elements *}
+
+lemma compact_principal [simp]: "compact (principal a)"
+by (rule compactI2, simp add: principal_below_iff_mem_rep rep_contlub)
+
+text {* Construct a chain whose lub is the same as a given ideal *}
+
+lemma obtain_principal_chain:
+  obtains Y where "\<forall>i. Y i \<preceq> Y (Suc i)" and "x = (\<Squnion>i. principal (Y i))"
+proof -
+  obtain count :: "'a \<Rightarrow> nat" where inj: "inj count"
+    using countable ..
+  def enum \<equiv> "\<lambda>i. THE a. count a = i"
+  have enum_count [simp]: "\<And>x. enum (count x) = x"
+    unfolding enum_def by (simp add: inj_eq [OF inj])
+  def a \<equiv> "LEAST i. enum i \<in> rep x"
+  def b \<equiv> "\<lambda>i. LEAST j. enum j \<in> rep x \<and> \<not> enum j \<preceq> enum i"
+  def c \<equiv> "\<lambda>i j. LEAST k. enum k \<in> rep x \<and> enum i \<preceq> enum k \<and> enum j \<preceq> enum k"
+  def P \<equiv> "\<lambda>i. \<exists>j. enum j \<in> rep x \<and> \<not> enum j \<preceq> enum i"
+  def X \<equiv> "nat_rec a (\<lambda>n i. if P i then c i (b i) else i)"
+  have X_0: "X 0 = a" unfolding X_def by simp
+  have X_Suc: "\<And>n. X (Suc n) = (if P (X n) then c (X n) (b (X n)) else X n)"
+    unfolding X_def by simp
+  have a_mem: "enum a \<in> rep x"
+    unfolding a_def
+    apply (rule LeastI_ex)
+    apply (cut_tac ideal_rep [of x])
+    apply (drule idealD1)
+    apply (clarify, rename_tac a)
+    apply (rule_tac x="count a" in exI, simp)
+    done
+  have b: "\<And>i. P i \<Longrightarrow> enum i \<in> rep x
+    \<Longrightarrow> enum (b i) \<in> rep x \<and> \<not> enum (b i) \<preceq> enum i"
+    unfolding P_def b_def by (erule LeastI2_ex, simp)
+  have c: "\<And>i j. enum i \<in> rep x \<Longrightarrow> enum j \<in> rep x
+    \<Longrightarrow> enum (c i j) \<in> rep x \<and> enum i \<preceq> enum (c i j) \<and> enum j \<preceq> enum (c i j)"
+    unfolding c_def
+    apply (drule (1) idealD2 [OF ideal_rep], clarify)
+    apply (rule_tac a="count z" in LeastI2, simp, simp)
+    done
+  have X_mem: "\<And>n. enum (X n) \<in> rep x"
+    apply (induct_tac n)
+    apply (simp add: X_0 a_mem)
+    apply (clarsimp simp add: X_Suc, rename_tac n)
+    apply (simp add: b c)
+    done
+  have X_chain: "\<And>n. enum (X n) \<preceq> enum (X (Suc n))"
+    apply (clarsimp simp add: X_Suc r_refl)
+    apply (simp add: b c X_mem)
+    done
+  have less_b: "\<And>n i. n < b i \<Longrightarrow> enum n \<in> rep x \<Longrightarrow> enum n \<preceq> enum i"
+    unfolding b_def by (drule not_less_Least, simp)
+  have X_covers: "\<And>n. \<forall>k\<le>n. enum k \<in> rep x \<longrightarrow> enum k \<preceq> enum (X n)"
+    apply (induct_tac n)
+    apply (clarsimp simp add: X_0 a_def)
+    apply (drule_tac k=0 in Least_le, simp add: r_refl)
+    apply (clarsimp, rename_tac n k)
+    apply (erule le_SucE)
+    apply (rule r_trans [OF _ X_chain], simp)
+    apply (case_tac "P (X n)", simp add: X_Suc)
+    apply (rule_tac x="b (X n)" and y="Suc n" in linorder_cases)
+    apply (simp only: less_Suc_eq_le)
+    apply (drule spec, drule (1) mp, simp add: b X_mem)
+    apply (simp add: c X_mem)
+    apply (drule (1) less_b)
+    apply (erule r_trans)
+    apply (simp add: b c X_mem)
+    apply (simp add: X_Suc)
+    apply (simp add: P_def)
+    done
+  have 1: "\<forall>i. enum (X i) \<preceq> enum (X (Suc i))"
+    by (simp add: X_chain)
+  have 2: "x = (\<Squnion>n. principal (enum (X n)))"
+    apply (simp add: eq_iff rep_contlub 1 rep_principal)
+    apply (auto, rename_tac a)
+    apply (subgoal_tac "\<exists>i. a = enum i", erule exE)
+    apply (rule_tac x=i in exI, simp add: X_covers)
+    apply (rule_tac x="count a" in exI, simp)
+    apply (erule idealD3 [OF ideal_rep])
+    apply (rule X_mem)
+    done
+  from 1 2 show ?thesis ..
+qed
+
+lemma principal_induct:
+  assumes adm: "adm P"
+  assumes P: "\<And>a. P (principal a)"
+  shows "P x"
+apply (rule obtain_principal_chain [of x])
+apply (simp add: admD [OF adm] P)
+done
+
+lemma principal_induct2:
+  "\<lbrakk>\<And>y. adm (\<lambda>x. P x y); \<And>x. adm (\<lambda>y. P x y);
+    \<And>a b. P (principal a) (principal b)\<rbrakk> \<Longrightarrow> P x y"
+apply (rule_tac x=y in spec)
+apply (rule_tac x=x in principal_induct, simp)
+apply (rule allI, rename_tac y)
+apply (rule_tac x=y in principal_induct, simp)
+apply simp
+done
+
+lemma compact_imp_principal: "compact x \<Longrightarrow> \<exists>a. x = principal a"
+apply (rule obtain_principal_chain [of x])
+apply (drule adm_compact_neq [OF _ cont_id])
+apply (subgoal_tac "chain (\<lambda>i. principal (Y i))")
+apply (drule (2) admD2, fast, simp)
+done
+
+lemma obtain_compact_chain:
+  obtains Y :: "nat \<Rightarrow> 'b"
+  where "chain Y" and "\<forall>i. compact (Y i)" and "x = (\<Squnion>i. Y i)"
+apply (rule obtain_principal_chain [of x])
+apply (rule_tac Y="\<lambda>i. principal (Y i)" in that, simp_all)
+done
+
 subsection {* Defining functions in terms of basis elements *}
 
 definition
   basis_fun :: "('a::type \<Rightarrow> 'c::cpo) \<Rightarrow> 'b \<rightarrow> 'c" where
   "basis_fun = (\<lambda>f. (\<Lambda> x. lub (f ` rep x)))"
 
-lemma basis_fun_lemma0:
-  fixes f :: "'a::type \<Rightarrow> 'c::cpo"
-  assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"
-  shows "\<exists>u. f ` take i ` rep x <<| u"
-apply (rule finite_directed_has_lub)
-apply (rule finite_imageI)
-apply (rule finite_take_rep)
-apply (subst image_image)
-apply (rule directed_image_ideal)
-apply (rule ideal_rep)
-apply (rule f_mono)
-apply (erule take_mono)
-done
-
-lemma basis_fun_lemma1:
-  fixes f :: "'a::type \<Rightarrow> 'c::cpo"
-  assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"
-  shows "chain (\<lambda>i. lub (f ` take i ` rep x))"
- apply (rule chainI)
- apply (rule is_lub_thelub0)
-  apply (rule basis_fun_lemma0, erule f_mono)
- apply (rule is_ubI, clarsimp, rename_tac a)
- apply (rule below_trans [OF f_mono [OF take_chain]])
- apply (rule is_ub_thelub0)
-  apply (rule basis_fun_lemma0, erule f_mono)
- apply simp
-done
-
-lemma basis_fun_lemma2:
-  fixes f :: "'a::type \<Rightarrow> 'c::cpo"
-  assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"
-  shows "f ` rep x <<| (\<Squnion>i. lub (f ` take i ` rep x))"
- apply (rule is_lubI)
- apply (rule ub_imageI, rename_tac a)
-  apply (cut_tac a=a in take_covers, erule exE, rename_tac i)
-  apply (erule subst)
-  apply (rule rev_below_trans)
-   apply (rule_tac x=i in is_ub_thelub)
-   apply (rule basis_fun_lemma1, erule f_mono)
-  apply (rule is_ub_thelub0)
-   apply (rule basis_fun_lemma0, erule f_mono)
-  apply simp
- apply (rule is_lub_thelub [OF _ ub_rangeI])
-  apply (rule basis_fun_lemma1, erule f_mono)
- apply (rule is_lub_thelub0)
-  apply (rule basis_fun_lemma0, erule f_mono)
- apply (rule is_ubI, clarsimp, rename_tac a)
- apply (rule below_trans [OF f_mono [OF take_less]])
- apply (erule (1) ub_imageD)
-done
-
 lemma basis_fun_lemma:
   fixes f :: "'a::type \<Rightarrow> 'c::cpo"
   assumes f_mono: "\<And>a b. a \<preceq> b \<Longrightarrow> f a \<sqsubseteq> f b"
   shows "\<exists>u. f ` rep x <<| u"
-by (rule exI, rule basis_fun_lemma2, erule f_mono)
+proof -
+  obtain Y where Y: "\<forall>i. Y i \<preceq> Y (Suc i)"
+  and x: "x = (\<Squnion>i. principal (Y i))"
+    by (rule obtain_principal_chain [of x])
+  have chain: "chain (\<lambda>i. f (Y i))"
+    by (rule chainI, simp add: f_mono Y)
+  have rep_x: "rep x = (\<Union>n. {a. a \<preceq> Y n})"
+    by (simp add: x rep_contlub Y rep_principal)
+  have "f ` rep x <<| (\<Squnion>n. f (Y n))"
+    apply (rule is_lubI)
+    apply (rule ub_imageI, rename_tac a)
+    apply (clarsimp simp add: rep_x)
+    apply (drule f_mono)
+    apply (erule below_trans)
+    apply (rule is_ub_thelub [OF chain])
+    apply (rule is_lub_thelub [OF chain])
+    apply (rule ub_rangeI)
+    apply (drule_tac x="Y i" in ub_imageD)
+    apply (simp add: rep_x, fast intro: r_refl)
+    apply assumption
+    done
+  thus ?thesis ..
+qed
 
 lemma basis_fun_beta:
   fixes f :: "'a::type \<Rightarrow> 'c::cpo"
@@ -345,13 +402,13 @@
   show cont: "cont (\<lambda>x. lub (f ` rep x))"
     apply (rule contI2)
      apply (rule monofunI)
-     apply (rule is_lub_thelub0 [OF lub ub_imageI])
-     apply (rule is_ub_thelub0 [OF lub imageI])
+     apply (rule is_lub_thelub_ex [OF lub ub_imageI])
+     apply (rule is_ub_thelub_ex [OF lub imageI])
      apply (erule (1) subsetD [OF rep_mono])
-    apply (rule is_lub_thelub0 [OF lub ub_imageI])
+    apply (rule is_lub_thelub_ex [OF lub ub_imageI])
     apply (simp add: rep_contlub, clarify)
     apply (erule rev_below_trans [OF is_ub_thelub])
-    apply (erule is_ub_thelub0 [OF lub imageI])
+    apply (erule is_ub_thelub_ex [OF lub imageI])
     done
 qed
 
@@ -371,113 +428,15 @@
   shows "basis_fun f \<sqsubseteq> basis_fun g"
  apply (rule below_cfun_ext)
  apply (simp only: basis_fun_beta f_mono g_mono)
- apply (rule is_lub_thelub0)
+ apply (rule is_lub_thelub_ex)
   apply (rule basis_fun_lemma, erule f_mono)
  apply (rule ub_imageI, rename_tac a)
  apply (rule below_trans [OF below])
- apply (rule is_ub_thelub0)
+ apply (rule is_ub_thelub_ex)
   apply (rule basis_fun_lemma, erule g_mono)
  apply (erule imageI)
 done
 
-lemma compact_principal [simp]: "compact (principal a)"
-by (rule compactI2, simp add: principal_below_iff_mem_rep rep_contlub)
-
-subsection {* Bifiniteness of ideal completions *}
-
-definition
-  completion_approx :: "nat \<Rightarrow> 'b \<rightarrow> 'b" where
-  "completion_approx = (\<lambda>i. basis_fun (\<lambda>a. principal (take i a)))"
-
-lemma completion_approx_beta:
-  "completion_approx i\<cdot>x = (\<Squnion>a\<in>rep x. principal (take i a))"
-unfolding completion_approx_def
-by (simp add: basis_fun_beta principal_mono take_mono)
-
-lemma completion_approx_principal:
-  "completion_approx i\<cdot>(principal a) = principal (take i a)"
-unfolding completion_approx_def
-by (simp add: basis_fun_principal principal_mono take_mono)
-
-lemma chain_completion_approx: "chain completion_approx"
-unfolding completion_approx_def
-apply (rule chainI)
-apply (rule basis_fun_mono)
-apply (erule principal_mono [OF take_mono])
-apply (erule principal_mono [OF take_mono])
-apply (rule principal_mono [OF take_chain])
-done
-
-lemma lub_completion_approx: "(\<Squnion>i. completion_approx i\<cdot>x) = x"
-unfolding completion_approx_beta
- apply (subst image_image [where f=principal, symmetric])
- apply (rule unique_lub [OF _ lub_principal_rep])
- apply (rule basis_fun_lemma2, erule principal_mono)
-done
-
-lemma completion_approx_eq_principal:
-  "\<exists>a\<in>rep x. completion_approx i\<cdot>x = principal (take i a)"
-unfolding completion_approx_beta
- apply (subst image_image [where f=principal, symmetric])
- apply (subgoal_tac "finite (principal ` take i ` rep x)")
-  apply (subgoal_tac "directed (principal ` take i ` rep x)")
-   apply (drule (1) lub_finite_directed_in_self, fast)
-  apply (subst image_image)
-  apply (rule directed_image_ideal)
-   apply (rule ideal_rep)
-  apply (erule principal_mono [OF take_mono])
- apply (rule finite_imageI)
- apply (rule finite_take_rep)
-done
-
-lemma completion_approx_idem:
-  "completion_approx i\<cdot>(completion_approx i\<cdot>x) = completion_approx i\<cdot>x"
-using completion_approx_eq_principal [where i=i and x=x]
-by (auto simp add: completion_approx_principal take_take)
-
-lemma finite_fixes_completion_approx:
-  "finite {x. completion_approx i\<cdot>x = x}" (is "finite ?S")
-apply (subgoal_tac "?S \<subseteq> principal ` range (take i)")
-apply (erule finite_subset)
-apply (rule finite_imageI)
-apply (rule finite_range_take)
-apply (clarify, erule subst)
-apply (cut_tac x=x and i=i in completion_approx_eq_principal)
-apply fast
-done
-
-lemma principal_induct:
-  assumes adm: "adm P"
-  assumes P: "\<And>a. P (principal a)"
-  shows "P x"
- apply (subgoal_tac "P (\<Squnion>i. completion_approx i\<cdot>x)")
- apply (simp add: lub_completion_approx)
- apply (rule admD [OF adm])
-  apply (simp add: chain_completion_approx)
- apply (cut_tac x=x and i=i in completion_approx_eq_principal)
- apply (clarify, simp add: P)
-done
-
-lemma principal_induct2:
-  "\<lbrakk>\<And>y. adm (\<lambda>x. P x y); \<And>x. adm (\<lambda>y. P x y);
-    \<And>a b. P (principal a) (principal b)\<rbrakk> \<Longrightarrow> P x y"
-apply (rule_tac x=y in spec)
-apply (rule_tac x=x in principal_induct, simp)
-apply (rule allI, rename_tac y)
-apply (rule_tac x=y in principal_induct, simp)
-apply simp
-done
-
-lemma compact_imp_principal: "compact x \<Longrightarrow> \<exists>a. x = principal a"
-apply (drule adm_compact_neq [OF _ cont_id])
-apply (drule admD2 [where Y="\<lambda>n. completion_approx n\<cdot>x"])
-apply (simp add: chain_completion_approx)
-apply (simp add: lub_completion_approx)
-apply (erule exE, erule ssubst)
-apply (cut_tac i=i and x=x in completion_approx_eq_principal)
-apply (clarify, erule exI)
-done
-
 end
 
 end
--- a/src/HOLCF/ConvexPD.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/ConvexPD.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -116,27 +116,6 @@
 apply (simp add: 4)
 done
 
-lemma pd_take_convex_chain:
-  "pd_take n t \<le>\<natural> pd_take (Suc n) t"
-apply (induct t rule: pd_basis_induct)
-apply (simp add: compact_basis.take_chain)
-apply (simp add: PDPlus_convex_mono)
-done
-
-lemma pd_take_convex_le: "pd_take i t \<le>\<natural> t"
-apply (induct t rule: pd_basis_induct)
-apply (simp add: compact_basis.take_less)
-apply (simp add: PDPlus_convex_mono)
-done
-
-lemma pd_take_convex_mono:
-  "t \<le>\<natural> u \<Longrightarrow> pd_take n t \<le>\<natural> pd_take n u"
-apply (erule convex_le_induct)
-apply (erule (1) convex_le_trans)
-apply (simp add: compact_basis.take_mono)
-apply (simp add: PDPlus_convex_mono)
-done
-
 
 subsection {* Type definition *}
 
@@ -144,7 +123,7 @@
   "{S::'a pd_basis set. convex_le.ideal S}"
 by (fast intro: convex_le.ideal_principal)
 
-instantiation convex_pd :: (profinite) below
+instantiation convex_pd :: (sfp) below
 begin
 
 definition
@@ -153,18 +132,18 @@
 instance ..
 end
 
-instance convex_pd :: (profinite) po
-by (rule convex_le.typedef_ideal_po
-    [OF type_definition_convex_pd below_convex_pd_def])
+instance convex_pd :: (sfp) po
+using type_definition_convex_pd below_convex_pd_def
+by (rule convex_le.typedef_ideal_po)
 
-instance convex_pd :: (profinite) cpo
-by (rule convex_le.typedef_ideal_cpo
-    [OF type_definition_convex_pd below_convex_pd_def])
+instance convex_pd :: (sfp) cpo
+using type_definition_convex_pd below_convex_pd_def
+by (rule convex_le.typedef_ideal_cpo)
 
 lemma Rep_convex_pd_lub:
   "chain Y \<Longrightarrow> Rep_convex_pd (\<Squnion>i. Y i) = (\<Union>i. Rep_convex_pd (Y i))"
-by (rule convex_le.typedef_ideal_rep_contlub
-    [OF type_definition_convex_pd below_convex_pd_def])
+using type_definition_convex_pd below_convex_pd_def
+by (rule convex_le.typedef_ideal_rep_contlub)
 
 lemma ideal_Rep_convex_pd: "convex_le.ideal (Rep_convex_pd xs)"
 by (rule Rep_convex_pd [unfolded mem_Collect_eq])
@@ -179,18 +158,13 @@
 by (simp add: Abs_convex_pd_inverse convex_le.ideal_principal)
 
 interpretation convex_pd:
-  ideal_completion convex_le pd_take convex_principal Rep_convex_pd
+  ideal_completion convex_le convex_principal Rep_convex_pd
 apply unfold_locales
-apply (rule pd_take_convex_le)
-apply (rule pd_take_idem)
-apply (erule pd_take_convex_mono)
-apply (rule pd_take_convex_chain)
-apply (rule finite_range_pd_take)
-apply (rule pd_take_covers)
 apply (rule ideal_Rep_convex_pd)
 apply (erule Rep_convex_pd_lub)
 apply (rule Rep_convex_principal)
 apply (simp only: below_convex_pd_def)
+apply (rule pd_basis_countable)
 done
 
 text {* Convex powerdomain is pointed *}
@@ -198,42 +172,12 @@
 lemma convex_pd_minimal: "convex_principal (PDUnit compact_bot) \<sqsubseteq> ys"
 by (induct ys rule: convex_pd.principal_induct, simp, simp)
 
-instance convex_pd :: (bifinite) pcpo
+instance convex_pd :: (sfp) pcpo
 by intro_classes (fast intro: convex_pd_minimal)
 
 lemma inst_convex_pd_pcpo: "\<bottom> = convex_principal (PDUnit compact_bot)"
 by (rule convex_pd_minimal [THEN UU_I, symmetric])
 
-text {* Convex powerdomain is profinite *}
-
-instantiation convex_pd :: (profinite) profinite
-begin
-
-definition
-  approx_convex_pd_def: "approx = convex_pd.completion_approx"
-
-instance
-apply (intro_classes, unfold approx_convex_pd_def)
-apply (rule convex_pd.chain_completion_approx)
-apply (rule convex_pd.lub_completion_approx)
-apply (rule convex_pd.completion_approx_idem)
-apply (rule convex_pd.finite_fixes_completion_approx)
-done
-
-end
-
-instance convex_pd :: (bifinite) bifinite ..
-
-lemma approx_convex_principal [simp]:
-  "approx n\<cdot>(convex_principal t) = convex_principal (pd_take n t)"
-unfolding approx_convex_pd_def
-by (rule convex_pd.completion_approx_principal)
-
-lemma approx_eq_convex_principal:
-  "\<exists>t\<in>Rep_convex_pd xs. approx n\<cdot>xs = convex_principal (pd_take n t)"
-unfolding approx_convex_pd_def
-by (rule convex_pd.completion_approx_eq_principal)
-
 
 subsection {* Monadic unit and plus *}
 
@@ -269,16 +213,6 @@
 by (simp add: convex_pd.basis_fun_principal
     convex_pd.basis_fun_mono PDPlus_convex_mono)
 
-lemma approx_convex_unit [simp]:
-  "approx n\<cdot>{x}\<natural> = {approx n\<cdot>x}\<natural>"
-apply (induct x rule: compact_basis.principal_induct, simp)
-apply (simp add: approx_Rep_compact_basis)
-done
-
-lemma approx_convex_plus [simp]:
-  "approx n\<cdot>(xs +\<natural> ys) = approx n\<cdot>xs +\<natural> approx n\<cdot>ys"
-by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp)
-
 interpretation convex_add: semilattice convex_add proof
   fix xs ys zs :: "'a convex_pd"
   show "(xs +\<natural> ys) +\<natural> zs = xs +\<natural> (ys +\<natural> zs)"
@@ -336,14 +270,20 @@
 unfolding po_eq_conv by simp
 
 lemma convex_unit_strict [simp]: "{\<bottom>}\<natural> = \<bottom>"
-unfolding inst_convex_pd_pcpo Rep_compact_bot [symmetric] by simp
+using convex_unit_Rep_compact_basis [of compact_bot]
+by (simp add: inst_convex_pd_pcpo)
 
 lemma convex_unit_strict_iff [simp]: "{x}\<natural> = \<bottom> \<longleftrightarrow> x = \<bottom>"
 unfolding convex_unit_strict [symmetric] by (rule convex_unit_eq_iff)
 
-lemma compact_convex_unit_iff [simp]:
-  "compact {x}\<natural> \<longleftrightarrow> compact x"
-unfolding profinite_compact_iff by simp
+lemma compact_convex_unit: "compact x \<Longrightarrow> compact {x}\<natural>"
+by (auto dest!: compact_basis.compact_imp_principal)
+
+lemma compact_convex_unit_iff [simp]: "compact {x}\<natural> \<longleftrightarrow> compact x"
+apply (safe elim!: compact_convex_unit)
+apply (simp only: compact_def convex_unit_below_iff [symmetric])
+apply (erule adm_subst [OF cont_Rep_CFun2])
+done
 
 lemma compact_convex_plus [simp]:
   "\<lbrakk>compact xs; compact ys\<rbrakk> \<Longrightarrow> compact (xs +\<natural> ys)"
@@ -441,32 +381,20 @@
 unfolding convex_unit_strict [symmetric] by (rule convex_bind_unit)
 
 
-subsection {* Map and join *}
+subsection {* Map *}
 
 definition
   convex_map :: "('a \<rightarrow> 'b) \<rightarrow> 'a convex_pd \<rightarrow> 'b convex_pd" where
   "convex_map = (\<Lambda> f xs. convex_bind\<cdot>xs\<cdot>(\<Lambda> x. {f\<cdot>x}\<natural>))"
 
-definition
-  convex_join :: "'a convex_pd convex_pd \<rightarrow> 'a convex_pd" where
-  "convex_join = (\<Lambda> xss. convex_bind\<cdot>xss\<cdot>(\<Lambda> xs. xs))"
-
 lemma convex_map_unit [simp]:
-  "convex_map\<cdot>f\<cdot>(convex_unit\<cdot>x) = convex_unit\<cdot>(f\<cdot>x)"
+  "convex_map\<cdot>f\<cdot>{x}\<natural> = {f\<cdot>x}\<natural>"
 unfolding convex_map_def by simp
 
 lemma convex_map_plus [simp]:
   "convex_map\<cdot>f\<cdot>(xs +\<natural> ys) = convex_map\<cdot>f\<cdot>xs +\<natural> convex_map\<cdot>f\<cdot>ys"
 unfolding convex_map_def by simp
 
-lemma convex_join_unit [simp]:
-  "convex_join\<cdot>{xs}\<natural> = xs"
-unfolding convex_join_def by simp
-
-lemma convex_join_plus [simp]:
-  "convex_join\<cdot>(xss +\<natural> yss) = convex_join\<cdot>xss +\<natural> convex_join\<cdot>yss"
-unfolding convex_join_def by simp
-
 lemma convex_map_ident: "convex_map\<cdot>(\<Lambda> x. x)\<cdot>xs = xs"
 by (induct xs rule: convex_pd_induct, simp_all)
 
@@ -477,6 +405,137 @@
   "convex_map\<cdot>f\<cdot>(convex_map\<cdot>g\<cdot>xs) = convex_map\<cdot>(\<Lambda> x. f\<cdot>(g\<cdot>x))\<cdot>xs"
 by (induct xs rule: convex_pd_induct, simp_all)
 
+lemma ep_pair_convex_map: "ep_pair e p \<Longrightarrow> ep_pair (convex_map\<cdot>e) (convex_map\<cdot>p)"
+apply default
+apply (induct_tac x rule: convex_pd_induct, simp_all add: ep_pair.e_inverse)
+apply (induct_tac y rule: convex_pd_induct)
+apply (simp_all add: ep_pair.e_p_below monofun_cfun)
+done
+
+lemma deflation_convex_map: "deflation d \<Longrightarrow> deflation (convex_map\<cdot>d)"
+apply default
+apply (induct_tac x rule: convex_pd_induct, simp_all add: deflation.idem)
+apply (induct_tac x rule: convex_pd_induct)
+apply (simp_all add: deflation.below monofun_cfun)
+done
+
+(* FIXME: long proof! *)
+lemma finite_deflation_convex_map:
+  assumes "finite_deflation d" shows "finite_deflation (convex_map\<cdot>d)"
+proof (rule finite_deflation_intro)
+  interpret d: finite_deflation d by fact
+  have "deflation d" by fact
+  thus "deflation (convex_map\<cdot>d)" by (rule deflation_convex_map)
+  have "finite (range (\<lambda>x. d\<cdot>x))" by (rule d.finite_range)
+  hence "finite (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))"
+    by (rule finite_vimageI, simp add: inj_on_def Rep_compact_basis_inject)
+  hence "finite (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x)))" by simp
+  hence "finite (Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))"
+    by (rule finite_vimageI, simp add: inj_on_def Rep_pd_basis_inject)
+  hence *: "finite (convex_principal ` Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))" by simp
+  hence "finite (range (\<lambda>xs. convex_map\<cdot>d\<cdot>xs))"
+    apply (rule rev_finite_subset)
+    apply clarsimp
+    apply (induct_tac xs rule: convex_pd.principal_induct)
+    apply (simp add: adm_mem_finite *)
+    apply (rename_tac t, induct_tac t rule: pd_basis_induct)
+    apply (simp only: convex_unit_Rep_compact_basis [symmetric] convex_map_unit)
+    apply simp
+    apply (subgoal_tac "\<exists>b. d\<cdot>(Rep_compact_basis a) = Rep_compact_basis b")
+    apply clarsimp
+    apply (rule imageI)
+    apply (rule vimageI2)
+    apply (simp add: Rep_PDUnit)
+    apply (rule range_eqI)
+    apply (erule sym)
+    apply (rule exI)
+    apply (rule Abs_compact_basis_inverse [symmetric])
+    apply (simp add: d.compact)
+    apply (simp only: convex_plus_principal [symmetric] convex_map_plus)
+    apply clarsimp
+    apply (rule imageI)
+    apply (rule vimageI2)
+    apply (simp add: Rep_PDPlus)
+    done
+  thus "finite {xs. convex_map\<cdot>d\<cdot>xs = xs}"
+    by (rule finite_range_imp_finite_fixes)
+qed
+
+subsection {* Convex powerdomain is an SFP domain *}
+
+definition
+  convex_approx :: "nat \<Rightarrow> udom convex_pd \<rightarrow> udom convex_pd"
+where
+  "convex_approx = (\<lambda>i. convex_map\<cdot>(udom_approx i))"
+
+lemma convex_approx: "approx_chain convex_approx"
+proof (rule approx_chain.intro)
+  show "chain (\<lambda>i. convex_approx i)"
+    unfolding convex_approx_def by simp
+  show "(\<Squnion>i. convex_approx i) = ID"
+    unfolding convex_approx_def
+    by (simp add: lub_distribs convex_map_ID)
+  show "\<And>i. finite_deflation (convex_approx i)"
+    unfolding convex_approx_def
+    by (intro finite_deflation_convex_map finite_deflation_udom_approx)
+qed
+
+definition convex_sfp :: "sfp \<rightarrow> sfp"
+where "convex_sfp = sfp_fun1 convex_approx convex_map"
+
+lemma cast_convex_sfp:
+  "cast\<cdot>(convex_sfp\<cdot>A) =
+    udom_emb convex_approx oo convex_map\<cdot>(cast\<cdot>A) oo udom_prj convex_approx"
+unfolding convex_sfp_def
+apply (rule cast_sfp_fun1 [OF convex_approx])
+apply (erule finite_deflation_convex_map)
+done
+
+instantiation convex_pd :: (sfp) sfp
+begin
+
+definition
+  "emb = udom_emb convex_approx oo convex_map\<cdot>emb"
+
+definition
+  "prj = convex_map\<cdot>prj oo udom_prj convex_approx"
+
+definition
+  "sfp (t::'a convex_pd itself) = convex_sfp\<cdot>SFP('a)"
+
+instance proof
+  show "ep_pair emb (prj :: udom \<rightarrow> 'a convex_pd)"
+    unfolding emb_convex_pd_def prj_convex_pd_def
+    using ep_pair_udom [OF convex_approx]
+    by (intro ep_pair_comp ep_pair_convex_map ep_pair_emb_prj)
+next
+  show "cast\<cdot>SFP('a convex_pd) = emb oo (prj :: udom \<rightarrow> 'a convex_pd)"
+    unfolding emb_convex_pd_def prj_convex_pd_def sfp_convex_pd_def cast_convex_sfp
+    by (simp add: cast_SFP oo_def expand_cfun_eq convex_map_map)
+qed
+
+end
+
+text {* SFP of type constructor = type combinator *}
+
+lemma SFP_convex: "SFP('a convex_pd) = convex_sfp\<cdot>SFP('a)"
+by (rule sfp_convex_pd_def)
+
+
+subsection {* Join *}
+
+definition
+  convex_join :: "'a convex_pd convex_pd \<rightarrow> 'a convex_pd" where
+  "convex_join = (\<Lambda> xss. convex_bind\<cdot>xss\<cdot>(\<Lambda> xs. xs))"
+
+lemma convex_join_unit [simp]:
+  "convex_join\<cdot>{xs}\<natural> = xs"
+unfolding convex_join_def by simp
+
+lemma convex_join_plus [simp]:
+  "convex_join\<cdot>(xss +\<natural> yss) = convex_join\<cdot>xss +\<natural> convex_join\<cdot>yss"
+unfolding convex_join_def by simp
+
 lemma convex_join_map_unit:
   "convex_join\<cdot>(convex_map\<cdot>convex_unit\<cdot>xs) = xs"
 by (induct xs rule: convex_pd_induct, simp_all)
@@ -490,24 +549,6 @@
    convex_map\<cdot>f\<cdot>(convex_join\<cdot>xss)"
 by (induct xss rule: convex_pd_induct, simp_all)
 
-lemma convex_map_approx: "convex_map\<cdot>(approx n)\<cdot>xs = approx n\<cdot>xs"
-by (induct xs rule: convex_pd_induct, simp_all)
-
-lemma ep_pair_convex_map:
-  "ep_pair e p \<Longrightarrow> ep_pair (convex_map\<cdot>e) (convex_map\<cdot>p)"
-apply default
-apply (induct_tac x rule: convex_pd_induct, simp_all add: ep_pair.e_inverse)
-apply (induct_tac y rule: convex_pd_induct)
-apply (simp_all add: ep_pair.e_p_below monofun_cfun)
-done
-
-lemma deflation_convex_map: "deflation d \<Longrightarrow> deflation (convex_map\<cdot>d)"
-apply default
-apply (induct_tac x rule: convex_pd_induct, simp_all add: deflation.idem)
-apply (induct_tac x rule: convex_pd_induct)
-apply (simp_all add: deflation.below monofun_cfun)
-done
-
 
 subsection {* Conversions to other powerdomains *}
 
@@ -536,10 +577,6 @@
   "convex_to_upper\<cdot>(xs +\<natural> ys) = convex_to_upper\<cdot>xs +\<sharp> convex_to_upper\<cdot>ys"
 by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp)
 
-lemma approx_convex_to_upper:
-  "approx i\<cdot>(convex_to_upper\<cdot>xs) = convex_to_upper\<cdot>(approx i\<cdot>xs)"
-by (induct xs rule: convex_pd_induct, simp, simp, simp)
-
 lemma convex_to_upper_bind [simp]:
   "convex_to_upper\<cdot>(convex_bind\<cdot>xs\<cdot>f) =
     upper_bind\<cdot>(convex_to_upper\<cdot>xs)\<cdot>(convex_to_upper oo f)"
@@ -579,10 +616,6 @@
   "convex_to_lower\<cdot>(xs +\<natural> ys) = convex_to_lower\<cdot>xs +\<flat> convex_to_lower\<cdot>ys"
 by (induct xs ys rule: convex_pd.principal_induct2, simp, simp, simp)
 
-lemma approx_convex_to_lower:
-  "approx i\<cdot>(convex_to_lower\<cdot>xs) = convex_to_lower\<cdot>(approx i\<cdot>xs)"
-by (induct xs rule: convex_pd_induct, simp, simp, simp)
-
 lemma convex_to_lower_bind [simp]:
   "convex_to_lower\<cdot>(convex_bind\<cdot>xs\<cdot>f) =
     lower_bind\<cdot>(convex_to_lower\<cdot>xs)\<cdot>(convex_to_lower oo f)"
--- a/src/HOLCF/Cprod.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Cprod.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -5,7 +5,7 @@
 header {* The cpo of cartesian products *}
 
 theory Cprod
-imports Bifinite
+imports Algebraic
 begin
 
 default_sort cpo
@@ -40,4 +40,62 @@
 lemma csplit_Pair [simp]: "csplit\<cdot>f\<cdot>(x, y) = f\<cdot>x\<cdot>y"
 by (simp add: csplit_def)
 
+subsection {* Cartesian product is an SFP domain *}
+
+definition
+  prod_approx :: "nat \<Rightarrow> udom \<times> udom \<rightarrow> udom \<times> udom"
+where
+  "prod_approx = (\<lambda>i. cprod_map\<cdot>(udom_approx i)\<cdot>(udom_approx i))"
+
+lemma prod_approx: "approx_chain prod_approx"
+proof (rule approx_chain.intro)
+  show "chain (\<lambda>i. prod_approx i)"
+    unfolding prod_approx_def by simp
+  show "(\<Squnion>i. prod_approx i) = ID"
+    unfolding prod_approx_def
+    by (simp add: lub_distribs cprod_map_ID)
+  show "\<And>i. finite_deflation (prod_approx i)"
+    unfolding prod_approx_def
+    by (intro finite_deflation_cprod_map finite_deflation_udom_approx)
+qed
+
+definition prod_sfp :: "sfp \<rightarrow> sfp \<rightarrow> sfp"
+where "prod_sfp = sfp_fun2 prod_approx cprod_map"
+
+lemma cast_prod_sfp:
+  "cast\<cdot>(prod_sfp\<cdot>A\<cdot>B) = udom_emb prod_approx oo
+    cprod_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj prod_approx"
+unfolding prod_sfp_def
+apply (rule cast_sfp_fun2 [OF prod_approx])
+apply (erule (1) finite_deflation_cprod_map)
+done
+
+instantiation prod :: (sfp, sfp) sfp
+begin
+
+definition
+  "emb = udom_emb prod_approx oo cprod_map\<cdot>emb\<cdot>emb"
+
+definition
+  "prj = cprod_map\<cdot>prj\<cdot>prj oo udom_prj prod_approx"
+
+definition
+  "sfp (t::('a \<times> 'b) itself) = prod_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
+
+instance proof
+  show "ep_pair emb (prj :: udom \<rightarrow> 'a \<times> 'b)"
+    unfolding emb_prod_def prj_prod_def
+    using ep_pair_udom [OF prod_approx]
+    by (intro ep_pair_comp ep_pair_cprod_map ep_pair_emb_prj)
+next
+  show "cast\<cdot>SFP('a \<times> 'b) = emb oo (prj :: udom \<rightarrow> 'a \<times> 'b)"
+    unfolding emb_prod_def prj_prod_def sfp_prod_def cast_prod_sfp
+    by (simp add: cast_SFP oo_def expand_cfun_eq cprod_map_map)
+qed
+
 end
+
+lemma SFP_prod: "SFP('a::sfp \<times> 'b::sfp) = prod_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
+by (rule sfp_prod_def)
+
+end
--- a/src/HOLCF/Eventual.thy	Tue Oct 05 17:53:00 2010 -0700
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,153 +0,0 @@
-(*  Title:      HOLCF/Eventual.thy
-    Author:     Brian Huffman
-*)
-
-header {* Eventually-constant sequences *}
-
-theory Eventual
-imports Infinite_Set
-begin
-
-subsection {* Lemmas about MOST *}
-
-lemma MOST_INFM:
-  assumes inf: "infinite (UNIV::'a set)"
-  shows "MOST x::'a. P x \<Longrightarrow> INFM x::'a. P x"
-  unfolding Alm_all_def Inf_many_def
-  apply (auto simp add: Collect_neg_eq)
-  apply (drule (1) finite_UnI)
-  apply (simp add: Compl_partition2 inf)
-  done
-
-lemma MOST_SucI: "MOST n. P n \<Longrightarrow> MOST n. P (Suc n)"
-by (rule MOST_inj [OF _ inj_Suc])
-
-lemma MOST_SucD: "MOST n. P (Suc n) \<Longrightarrow> MOST n. P n"
-unfolding MOST_nat
-apply (clarify, rule_tac x="Suc m" in exI, clarify)
-apply (erule Suc_lessE, simp)
-done
-
-lemma MOST_Suc_iff: "(MOST n. P (Suc n)) \<longleftrightarrow> (MOST n. P n)"
-by (rule iffI [OF MOST_SucD MOST_SucI])
-
-lemma INFM_finite_Bex_distrib:
-  "finite A \<Longrightarrow> (INFM y. \<exists>x\<in>A. P x y) \<longleftrightarrow> (\<exists>x\<in>A. INFM y. P x y)"
-by (induct set: finite, simp, simp add: INFM_disj_distrib)
-
-lemma MOST_finite_Ball_distrib:
-  "finite A \<Longrightarrow> (MOST y. \<forall>x\<in>A. P x y) \<longleftrightarrow> (\<forall>x\<in>A. MOST y. P x y)"
-by (induct set: finite, simp, simp add: MOST_conj_distrib)
-
-lemma MOST_ge_nat: "MOST n::nat. m \<le> n"
-unfolding MOST_nat_le by fast
-
-subsection {* Eventually constant sequences *}
-
-definition
-  eventually_constant :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool"
-where
-  "eventually_constant S = (\<exists>x. MOST i. S i = x)"
-
-lemma eventually_constant_MOST_MOST:
-  "eventually_constant S \<longleftrightarrow> (MOST m. MOST n. S n = S m)"
-unfolding eventually_constant_def MOST_nat
-apply safe
-apply (rule_tac x=m in exI, clarify)
-apply (rule_tac x=m in exI, clarify)
-apply simp
-apply fast
-done
-
-lemma eventually_constantI: "MOST i. S i = x \<Longrightarrow> eventually_constant S"
-unfolding eventually_constant_def by fast
-
-lemma eventually_constant_comp:
-  "eventually_constant (\<lambda>i. S i) \<Longrightarrow> eventually_constant (\<lambda>i. f (S i))"
-unfolding eventually_constant_def
-apply (erule exE, rule_tac x="f x" in exI)
-apply (erule MOST_mono, simp)
-done
-
-lemma eventually_constant_Suc_iff:
-  "eventually_constant (\<lambda>i. S (Suc i)) \<longleftrightarrow> eventually_constant (\<lambda>i. S i)"
-unfolding eventually_constant_def
-by (subst MOST_Suc_iff, rule refl)
-
-lemma eventually_constant_SucD:
-  "eventually_constant (\<lambda>i. S (Suc i)) \<Longrightarrow> eventually_constant (\<lambda>i. S i)"
-by (rule eventually_constant_Suc_iff [THEN iffD1])
-
-subsection {* Limits of eventually constant sequences *}
-
-definition
-  eventual :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a" where
-  "eventual S = (THE x. MOST i. S i = x)"
-
-lemma eventual_eqI: "MOST i. S i = x \<Longrightarrow> eventual S = x"
-unfolding eventual_def
-apply (rule the_equality, assumption)
-apply (rename_tac y)
-apply (subgoal_tac "MOST i::nat. y = x", simp)
-apply (erule MOST_rev_mp)
-apply (erule MOST_rev_mp)
-apply simp
-done
-
-lemma MOST_eq_eventual:
-  "eventually_constant S \<Longrightarrow> MOST i. S i = eventual S"
-unfolding eventually_constant_def
-by (erule exE, simp add: eventual_eqI)
-
-lemma eventual_mem_range:
-  "eventually_constant S \<Longrightarrow> eventual S \<in> range S"
-apply (drule MOST_eq_eventual)
-apply (simp only: MOST_nat_le, clarify)
-apply (drule spec, drule mp, rule order_refl)
-apply (erule range_eqI [OF sym])
-done
-
-lemma eventually_constant_MOST_iff:
-  assumes S: "eventually_constant S"
-  shows "(MOST n. P (S n)) \<longleftrightarrow> P (eventual S)"
-apply (subgoal_tac "(MOST n. P (S n)) \<longleftrightarrow> (MOST n::nat. P (eventual S))")
-apply simp
-apply (rule iffI)
-apply (rule MOST_rev_mp [OF MOST_eq_eventual [OF S]])
-apply (erule MOST_mono, force)
-apply (rule MOST_rev_mp [OF MOST_eq_eventual [OF S]])
-apply (erule MOST_mono, simp)
-done
-
-lemma MOST_eventual:
-  "\<lbrakk>eventually_constant S; MOST n. P (S n)\<rbrakk> \<Longrightarrow> P (eventual S)"
-proof -
-  assume "eventually_constant S"
-  hence "MOST n. S n = eventual S"
-    by (rule MOST_eq_eventual)
-  moreover assume "MOST n. P (S n)"
-  ultimately have "MOST n. S n = eventual S \<and> P (S n)"
-    by (rule MOST_conj_distrib [THEN iffD2, OF conjI])
-  hence "MOST n::nat. P (eventual S)"
-    by (rule MOST_mono) auto
-  thus ?thesis by simp
-qed
-
-lemma eventually_constant_MOST_Suc_eq:
-  "eventually_constant S \<Longrightarrow> MOST n. S (Suc n) = S n"
-apply (drule MOST_eq_eventual)
-apply (frule MOST_Suc_iff [THEN iffD2])
-apply (erule MOST_rev_mp)
-apply (erule MOST_rev_mp)
-apply simp
-done
-
-lemma eventual_comp:
-  "eventually_constant S \<Longrightarrow> eventual (\<lambda>i. f (S i)) = f (eventual (\<lambda>i. S i))"
-apply (rule eventual_eqI)
-apply (rule MOST_mono)
-apply (erule MOST_eq_eventual)
-apply simp
-done
-
-end
--- a/src/HOLCF/HOLCF.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/HOLCF.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -45,8 +45,6 @@
 lemmas expand_cfun_less = expand_cfun_below
 lemmas less_cfun_ext = below_cfun_ext
 lemmas injection_less = injection_below
-lemmas approx_less = approx_below
-lemmas profinite_less_ext = profinite_below_ext
 lemmas less_up_def = below_up_def
 lemmas not_Iup_less = not_Iup_below
 lemmas Iup_less = Iup_below
--- a/src/HOLCF/IsaMakefile	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/IsaMakefile	Wed Oct 06 10:49:27 2010 -0700
@@ -48,7 +48,6 @@
   Deflation.thy \
   Domain.thy \
   Domain_Aux.thy \
-  Eventual.thy \
   Ffun.thy \
   Fixrec.thy \
   Fix.thy \
--- a/src/HOLCF/Library/Strict_Fun.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Library/Strict_Fun.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -1,4 +1,4 @@
-(*  Title:      HOLCF/ex/Strict_Fun.thy
+(*  Title:      HOLCF/Library/Strict_Fun.thy
     Author:     Brian Huffman
 *)
 
@@ -143,96 +143,78 @@
          deflation_strict `deflation d1` `deflation d2`)
 qed
 
-subsection {* Strict function space is bifinite *}
+subsection {* Strict function space is an SFP domain *}
+
+definition
+  sfun_approx :: "nat \<Rightarrow> (udom \<rightarrow>! udom) \<rightarrow> (udom \<rightarrow>! udom)"
+where
+  "sfun_approx = (\<lambda>i. sfun_map\<cdot>(udom_approx i)\<cdot>(udom_approx i))"
 
-instantiation sfun :: (bifinite, bifinite) bifinite
+lemma sfun_approx: "approx_chain sfun_approx"
+proof (rule approx_chain.intro)
+  show "chain (\<lambda>i. sfun_approx i)"
+    unfolding sfun_approx_def by simp
+  show "(\<Squnion>i. sfun_approx i) = ID"
+    unfolding sfun_approx_def
+    by (simp add: lub_distribs sfun_map_ID)
+  show "\<And>i. finite_deflation (sfun_approx i)"
+    unfolding sfun_approx_def
+    by (intro finite_deflation_sfun_map finite_deflation_udom_approx)
+qed
+
+definition sfun_sfp :: "sfp \<rightarrow> sfp \<rightarrow> sfp"
+where "sfun_sfp = sfp_fun2 sfun_approx sfun_map"
+
+lemma cast_sfun_sfp:
+  "cast\<cdot>(sfun_sfp\<cdot>A\<cdot>B) =
+    udom_emb sfun_approx oo sfun_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj sfun_approx"
+unfolding sfun_sfp_def
+apply (rule cast_sfp_fun2 [OF sfun_approx])
+apply (erule (1) finite_deflation_sfun_map)
+done
+
+instantiation sfun :: (sfp, sfp) sfp
 begin
 
 definition
-  "approx = (\<lambda>i. sfun_map\<cdot>(approx i)\<cdot>(approx i))"
+  "emb = udom_emb sfun_approx oo sfun_map\<cdot>prj\<cdot>emb"
+
+definition
+  "prj = sfun_map\<cdot>emb\<cdot>prj oo udom_prj sfun_approx"
+
+definition
+  "sfp (t::('a \<rightarrow>! 'b) itself) = sfun_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
 
 instance proof
-  show "chain (approx :: nat \<Rightarrow> ('a \<rightarrow>! 'b) \<rightarrow> ('a \<rightarrow>! 'b))"
-    unfolding approx_sfun_def by simp
-next
-  fix x :: "'a \<rightarrow>! 'b"
-  show "(\<Squnion>i. approx i\<cdot>x) = x"
-    unfolding approx_sfun_def
-    by (simp add: lub_distribs sfun_map_ID [unfolded ID_def])
+  show "ep_pair emb (prj :: udom \<rightarrow> 'a \<rightarrow>! 'b)"
+    unfolding emb_sfun_def prj_sfun_def
+    using ep_pair_udom [OF sfun_approx]
+    by (intro ep_pair_comp ep_pair_sfun_map ep_pair_emb_prj)
 next
-  fix i :: nat and x :: "'a \<rightarrow>! 'b"
-  show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
-    unfolding approx_sfun_def
-    by (intro deflation.idem deflation_sfun_map deflation_approx)
-next
-  fix i :: nat
-  show "finite {x::'a \<rightarrow>! 'b. approx i\<cdot>x = x}"
-    unfolding approx_sfun_def
-    by (intro finite_deflation.finite_fixes
-              finite_deflation_sfun_map
-              finite_deflation_approx)
+  show "cast\<cdot>SFP('a \<rightarrow>! 'b) = emb oo (prj :: udom \<rightarrow> 'a \<rightarrow>! 'b)"
+    unfolding emb_sfun_def prj_sfun_def sfp_sfun_def cast_sfun_sfp
+    by (simp add: cast_SFP oo_def expand_cfun_eq sfun_map_map)
 qed
 
 end
 
-subsection {* Strict function space is representable *}
-
-instantiation sfun :: (rep, rep) rep
-begin
-
-definition
-  "emb = udom_emb oo sfun_map\<cdot>prj\<cdot>emb"
-
-definition
-  "prj = sfun_map\<cdot>emb\<cdot>prj oo udom_prj"
-
-instance
-apply (default, unfold emb_sfun_def prj_sfun_def)
-apply (rule ep_pair_comp)
-apply (rule ep_pair_sfun_map)
-apply (rule ep_pair_emb_prj)
-apply (rule ep_pair_emb_prj)
-apply (rule ep_pair_udom)
-done
-
-end
+text {* SFP of type constructor = type combinator *}
 
-text {*
-  A deflation constructor lets us configure the domain package to work
-  with the strict function space type constructor.
-*}
-
-definition
-  sfun_defl :: "TypeRep \<rightarrow> TypeRep \<rightarrow> TypeRep"
-where
-  "sfun_defl = TypeRep_fun2 sfun_map"
-
-lemma cast_sfun_defl:
-  "cast\<cdot>(sfun_defl\<cdot>A\<cdot>B) = udom_emb oo sfun_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj"
-unfolding sfun_defl_def
-apply (rule cast_TypeRep_fun2)
-apply (erule (1) finite_deflation_sfun_map)
-done
-
-lemma REP_sfun: "REP('a::rep \<rightarrow>! 'b::rep) = sfun_defl\<cdot>REP('a)\<cdot>REP('b)"
-apply (rule cast_eq_imp_eq, rule ext_cfun)
-apply (simp add: cast_REP cast_sfun_defl)
-apply (simp only: prj_sfun_def emb_sfun_def)
-apply (simp add: sfun_map_def cfun_map_def strictify_cancel)
-done
+lemma SFP_sfun: "SFP('a::sfp \<rightarrow>! 'b::sfp) = sfun_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
+by (rule sfp_sfun_def)
 
 lemma isodefl_sfun:
   "isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>
-    isodefl (sfun_map\<cdot>d1\<cdot>d2) (sfun_defl\<cdot>t1\<cdot>t2)"
+    isodefl (sfun_map\<cdot>d1\<cdot>d2) (sfun_sfp\<cdot>t1\<cdot>t2)"
 apply (rule isodeflI)
-apply (simp add: cast_sfun_defl cast_isodefl)
+apply (simp add: cast_sfun_sfp cast_isodefl)
 apply (simp add: emb_sfun_def prj_sfun_def)
 apply (simp add: sfun_map_map deflation_strict [OF isodefl_imp_deflation])
 done
 
 setup {*
   Domain_Isomorphism.add_type_constructor
-    (@{type_name "sfun"}, @{term sfun_defl}, @{const_name sfun_map}, @{thm REP_sfun},
+    (@{type_name "sfun"}, @{term sfun_sfp}, @{const_name sfun_map}, @{thm SFP_sfun},
        @{thm isodefl_sfun}, @{thm sfun_map_ID}, @{thm deflation_sfun_map})
 *}
 
--- a/src/HOLCF/Library/Sum_Cpo.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Library/Sum_Cpo.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -5,7 +5,7 @@
 header {* The cpo of disjoint sums *}
 
 theory Sum_Cpo
-imports Bifinite
+imports HOLCF
 begin
 
 subsection {* Ordering on sum type *}
@@ -218,38 +218,4 @@
 instance sum :: (discrete_cpo, discrete_cpo) discrete_cpo
 by intro_classes (simp add: below_sum_def split: sum.split)
 
-subsection {* Sum type is a bifinite domain *}
-
-instantiation sum :: (profinite, profinite) profinite
-begin
-
-definition
-  approx_sum_def: "approx =
-    (\<lambda>n. \<Lambda> x. case x of Inl a \<Rightarrow> Inl (approx n\<cdot>a) | Inr b \<Rightarrow> Inr (approx n\<cdot>b))"
-
-lemma approx_Inl [simp]: "approx n\<cdot>(Inl x) = Inl (approx n\<cdot>x)"
-  unfolding approx_sum_def by simp
-
-lemma approx_Inr [simp]: "approx n\<cdot>(Inr x) = Inr (approx n\<cdot>x)"
-  unfolding approx_sum_def by simp
-
-instance proof
-  fix i :: nat and x :: "'a + 'b"
-  show "chain (approx :: nat \<Rightarrow> 'a + 'b \<rightarrow> 'a + 'b)"
-    unfolding approx_sum_def
-    by (rule ch2ch_LAM, case_tac x, simp_all)
-  show "(\<Squnion>i. approx i\<cdot>x) = x"
-    by (induct x, simp_all add: lub_Inl lub_Inr)
-  show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
-    by (induct x, simp_all)
-  have "{x::'a + 'b. approx i\<cdot>x = x} \<subseteq>
-        {x::'a. approx i\<cdot>x = x} <+> {x::'b. approx i\<cdot>x = x}"
-    by (rule subsetI, case_tac x, simp_all add: InlI InrI)
-  thus "finite {x::'a + 'b. approx i\<cdot>x = x}"
-    by (rule finite_subset,
-        intro finite_Plus finite_fixes_approx)
-qed
-
 end
-
-end
--- a/src/HOLCF/Lift.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Lift.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -171,51 +171,89 @@
 by (cases x, simp_all)
 
 
-subsection {* Lifted countable types are bifinite *}
-
-instantiation lift :: (countable) bifinite
-begin
+subsection {* Lifted countable types are SFP domains *}
 
 definition
-  approx_lift_def:
-    "approx = (\<lambda>n. FLIFT x. if to_nat x < n then Def x else \<bottom>)"
+  lift_approx :: "nat \<Rightarrow> 'a::countable lift \<rightarrow> 'a lift"
+where
+  "lift_approx = (\<lambda>i. FLIFT x. if to_nat x < i then Def x else \<bottom>)"
+
+lemma chain_lift_approx [simp]: "chain lift_approx"
+  unfolding lift_approx_def
+  by (rule chainI, simp add: FLIFT_mono)
 
-instance proof
-  fix x :: "'a lift"
-  show "chain (approx :: nat \<Rightarrow> 'a lift \<rightarrow> 'a lift)"
-    unfolding approx_lift_def
-    by (rule chainI, simp add: FLIFT_mono)
-next
-  fix x :: "'a lift"
-  show "(\<Squnion>i. approx i\<cdot>x) = x"
-    unfolding approx_lift_def
-    apply (cases x, simp)
-    apply (rule thelubI)
-    apply (rule is_lubI)
-     apply (rule ub_rangeI, simp)
-    apply (drule ub_rangeD)
-    apply (erule rev_below_trans)
-    apply simp
-    apply (rule lessI)
-    done
-next
-  fix i :: nat and x :: "'a lift"
-  show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
-    unfolding approx_lift_def
+lemma lub_lift_approx [simp]: "(\<Squnion>i. lift_approx i) = ID"
+apply (rule ext_cfun)
+apply (simp add: contlub_cfun_fun)
+apply (simp add: lift_approx_def)
+apply (case_tac x, simp)
+apply (rule thelubI)
+apply (rule is_lubI)
+apply (rule ub_rangeI, simp)
+apply (drule ub_rangeD)
+apply (erule rev_below_trans)
+apply simp
+apply (rule lessI)
+done
+
+lemma finite_deflation_lift_approx: "finite_deflation (lift_approx i)"
+proof
+  fix x
+  show "lift_approx i\<cdot>x \<sqsubseteq> x"
+    unfolding lift_approx_def
     by (cases x, simp, simp)
-next
-  fix i :: nat
-  show "finite {x::'a lift. approx i\<cdot>x = x}"
+  show "lift_approx i\<cdot>(lift_approx i\<cdot>x) = lift_approx i\<cdot>x"
+    unfolding lift_approx_def
+    by (cases x, simp, simp)
+  show "finite {x::'a lift. lift_approx i\<cdot>x = x}"
   proof (rule finite_subset)
     let ?S = "insert (\<bottom>::'a lift) (Def ` to_nat -` {..<i})"
-    show "{x::'a lift. approx i\<cdot>x = x} \<subseteq> ?S"
-      unfolding approx_lift_def
+    show "{x::'a lift. lift_approx i\<cdot>x = x} \<subseteq> ?S"
+      unfolding lift_approx_def
       by (rule subsetI, case_tac x, simp, simp split: split_if_asm)
     show "finite ?S"
       by (simp add: finite_vimageI)
   qed
 qed
 
+lemma lift_approx: "approx_chain lift_approx"
+using chain_lift_approx lub_lift_approx finite_deflation_lift_approx
+by (rule approx_chain.intro)
+
+instantiation lift :: (countable) sfp
+begin
+
+definition
+  "emb = udom_emb lift_approx"
+
+definition
+  "prj = udom_prj lift_approx"
+
+definition
+  "sfp (t::'a lift itself) =
+    (\<Squnion>i. sfp_principal (Abs_fin_defl (emb oo lift_approx i oo prj)))"
+
+instance proof
+  show ep: "ep_pair emb (prj :: udom \<rightarrow> 'a lift)"
+    unfolding emb_lift_def prj_lift_def
+    by (rule ep_pair_udom [OF lift_approx])
+  show "cast\<cdot>SFP('a lift) = emb oo (prj :: udom \<rightarrow> 'a lift)"
+    unfolding sfp_lift_def
+    apply (subst contlub_cfun_arg)
+    apply (rule chainI)
+    apply (rule sfp.principal_mono)
+    apply (simp add: below_fin_defl_def)
+    apply (simp add: Abs_fin_defl_inverse finite_deflation_lift_approx
+                     ep_pair.finite_deflation_e_d_p [OF ep])
+    apply (intro monofun_cfun below_refl)
+    apply (rule chainE)
+    apply (rule chain_lift_approx)
+    apply (subst cast_sfp_principal)
+    apply (simp add: Abs_fin_defl_inverse finite_deflation_lift_approx
+                     ep_pair.finite_deflation_e_d_p [OF ep] lub_distribs)
+    done
+qed
+
 end
 
 end
--- a/src/HOLCF/LowerPD.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/LowerPD.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -71,27 +71,6 @@
 apply (simp add: lower_le_PDPlus_iff 3)
 done
 
-lemma pd_take_lower_chain:
-  "pd_take n t \<le>\<flat> pd_take (Suc n) t"
-apply (induct t rule: pd_basis_induct)
-apply (simp add: compact_basis.take_chain)
-apply (simp add: PDPlus_lower_mono)
-done
-
-lemma pd_take_lower_le: "pd_take i t \<le>\<flat> t"
-apply (induct t rule: pd_basis_induct)
-apply (simp add: compact_basis.take_less)
-apply (simp add: PDPlus_lower_mono)
-done
-
-lemma pd_take_lower_mono:
-  "t \<le>\<flat> u \<Longrightarrow> pd_take n t \<le>\<flat> pd_take n u"
-apply (erule lower_le_induct)
-apply (simp add: compact_basis.take_mono)
-apply (simp add: lower_le_PDUnit_PDPlus_iff)
-apply (simp add: lower_le_PDPlus_iff)
-done
-
 
 subsection {* Type definition *}
 
@@ -99,7 +78,7 @@
   "{S::'a pd_basis set. lower_le.ideal S}"
 by (fast intro: lower_le.ideal_principal)
 
-instantiation lower_pd :: (profinite) below
+instantiation lower_pd :: (sfp) below
 begin
 
 definition
@@ -108,11 +87,11 @@
 instance ..
 end
 
-instance lower_pd :: (profinite) po
+instance lower_pd :: (sfp) po
 by (rule lower_le.typedef_ideal_po
     [OF type_definition_lower_pd below_lower_pd_def])
 
-instance lower_pd :: (profinite) cpo
+instance lower_pd :: (sfp) cpo
 by (rule lower_le.typedef_ideal_cpo
     [OF type_definition_lower_pd below_lower_pd_def])
 
@@ -134,18 +113,13 @@
 by (simp add: Abs_lower_pd_inverse lower_le.ideal_principal)
 
 interpretation lower_pd:
-  ideal_completion lower_le pd_take lower_principal Rep_lower_pd
+  ideal_completion lower_le lower_principal Rep_lower_pd
 apply unfold_locales
-apply (rule pd_take_lower_le)
-apply (rule pd_take_idem)
-apply (erule pd_take_lower_mono)
-apply (rule pd_take_lower_chain)
-apply (rule finite_range_pd_take)
-apply (rule pd_take_covers)
 apply (rule ideal_Rep_lower_pd)
 apply (erule Rep_lower_pd_lub)
 apply (rule Rep_lower_principal)
 apply (simp only: below_lower_pd_def)
+apply (rule pd_basis_countable)
 done
 
 text {* Lower powerdomain is pointed *}
@@ -153,42 +127,12 @@
 lemma lower_pd_minimal: "lower_principal (PDUnit compact_bot) \<sqsubseteq> ys"
 by (induct ys rule: lower_pd.principal_induct, simp, simp)
 
-instance lower_pd :: (bifinite) pcpo
+instance lower_pd :: (sfp) pcpo
 by intro_classes (fast intro: lower_pd_minimal)
 
 lemma inst_lower_pd_pcpo: "\<bottom> = lower_principal (PDUnit compact_bot)"
 by (rule lower_pd_minimal [THEN UU_I, symmetric])
 
-text {* Lower powerdomain is profinite *}
-
-instantiation lower_pd :: (profinite) profinite
-begin
-
-definition
-  approx_lower_pd_def: "approx = lower_pd.completion_approx"
-
-instance
-apply (intro_classes, unfold approx_lower_pd_def)
-apply (rule lower_pd.chain_completion_approx)
-apply (rule lower_pd.lub_completion_approx)
-apply (rule lower_pd.completion_approx_idem)
-apply (rule lower_pd.finite_fixes_completion_approx)
-done
-
-end
-
-instance lower_pd :: (bifinite) bifinite ..
-
-lemma approx_lower_principal [simp]:
-  "approx n\<cdot>(lower_principal t) = lower_principal (pd_take n t)"
-unfolding approx_lower_pd_def
-by (rule lower_pd.completion_approx_principal)
-
-lemma approx_eq_lower_principal:
-  "\<exists>t\<in>Rep_lower_pd xs. approx n\<cdot>xs = lower_principal (pd_take n t)"
-unfolding approx_lower_pd_def
-by (rule lower_pd.completion_approx_eq_principal)
-
 
 subsection {* Monadic unit and plus *}
 
@@ -224,16 +168,6 @@
 by (simp add: lower_pd.basis_fun_principal
     lower_pd.basis_fun_mono PDPlus_lower_mono)
 
-lemma approx_lower_unit [simp]:
-  "approx n\<cdot>{x}\<flat> = {approx n\<cdot>x}\<flat>"
-apply (induct x rule: compact_basis.principal_induct, simp)
-apply (simp add: approx_Rep_compact_basis)
-done
-
-lemma approx_lower_plus [simp]:
-  "approx n\<cdot>(xs +\<flat> ys) = (approx n\<cdot>xs) +\<flat> (approx n\<cdot>ys)"
-by (induct xs ys rule: lower_pd.principal_induct2, simp, simp, simp)
-
 interpretation lower_add: semilattice lower_add proof
   fix xs ys zs :: "'a lower_pd"
   show "(xs +\<flat> ys) +\<flat> zs = xs +\<flat> (ys +\<flat> zs)"
@@ -309,7 +243,8 @@
 by (simp add: po_eq_conv)
 
 lemma lower_unit_strict [simp]: "{\<bottom>}\<flat> = \<bottom>"
-unfolding inst_lower_pd_pcpo Rep_compact_bot [symmetric] by simp
+using lower_unit_Rep_compact_basis [of compact_bot]
+by (simp add: inst_lower_pd_pcpo)
 
 lemma lower_unit_strict_iff [simp]: "{x}\<flat> = \<bottom> \<longleftrightarrow> x = \<bottom>"
 unfolding lower_unit_strict [symmetric] by (rule lower_unit_eq_iff)
@@ -332,8 +267,14 @@
 apply (simp add: lower_plus_least)
 done
 
+lemma compact_lower_unit: "compact x \<Longrightarrow> compact {x}\<flat>"
+by (auto dest!: compact_basis.compact_imp_principal)
+
 lemma compact_lower_unit_iff [simp]: "compact {x}\<flat> \<longleftrightarrow> compact x"
-unfolding profinite_compact_iff by simp
+apply (safe elim!: compact_lower_unit)
+apply (simp only: compact_def lower_unit_below_iff [symmetric])
+apply (erule adm_subst [OF cont_Rep_CFun2])
+done
 
 lemma compact_lower_plus [simp]:
   "\<lbrakk>compact xs; compact ys\<rbrakk> \<Longrightarrow> compact (xs +\<flat> ys)"
@@ -429,16 +370,12 @@
 unfolding lower_unit_strict [symmetric] by (rule lower_bind_unit)
 
 
-subsection {* Map and join *}
+subsection {* Map *}
 
 definition
   lower_map :: "('a \<rightarrow> 'b) \<rightarrow> 'a lower_pd \<rightarrow> 'b lower_pd" where
   "lower_map = (\<Lambda> f xs. lower_bind\<cdot>xs\<cdot>(\<Lambda> x. {f\<cdot>x}\<flat>))"
 
-definition
-  lower_join :: "'a lower_pd lower_pd \<rightarrow> 'a lower_pd" where
-  "lower_join = (\<Lambda> xss. lower_bind\<cdot>xss\<cdot>(\<Lambda> xs. xs))"
-
 lemma lower_map_unit [simp]:
   "lower_map\<cdot>f\<cdot>{x}\<flat> = {f\<cdot>x}\<flat>"
 unfolding lower_map_def by simp
@@ -447,14 +384,6 @@
   "lower_map\<cdot>f\<cdot>(xs +\<flat> ys) = lower_map\<cdot>f\<cdot>xs +\<flat> lower_map\<cdot>f\<cdot>ys"
 unfolding lower_map_def by simp
 
-lemma lower_join_unit [simp]:
-  "lower_join\<cdot>{xs}\<flat> = xs"
-unfolding lower_join_def by simp
-
-lemma lower_join_plus [simp]:
-  "lower_join\<cdot>(xss +\<flat> yss) = lower_join\<cdot>xss +\<flat> lower_join\<cdot>yss"
-unfolding lower_join_def by simp
-
 lemma lower_map_ident: "lower_map\<cdot>(\<Lambda> x. x)\<cdot>xs = xs"
 by (induct xs rule: lower_pd_induct, simp_all)
 
@@ -465,22 +394,6 @@
   "lower_map\<cdot>f\<cdot>(lower_map\<cdot>g\<cdot>xs) = lower_map\<cdot>(\<Lambda> x. f\<cdot>(g\<cdot>x))\<cdot>xs"
 by (induct xs rule: lower_pd_induct, simp_all)
 
-lemma lower_join_map_unit:
-  "lower_join\<cdot>(lower_map\<cdot>lower_unit\<cdot>xs) = xs"
-by (induct xs rule: lower_pd_induct, simp_all)
-
-lemma lower_join_map_join:
-  "lower_join\<cdot>(lower_map\<cdot>lower_join\<cdot>xsss) = lower_join\<cdot>(lower_join\<cdot>xsss)"
-by (induct xsss rule: lower_pd_induct, simp_all)
-
-lemma lower_join_map_map:
-  "lower_join\<cdot>(lower_map\<cdot>(lower_map\<cdot>f)\<cdot>xss) =
-   lower_map\<cdot>f\<cdot>(lower_join\<cdot>xss)"
-by (induct xss rule: lower_pd_induct, simp_all)
-
-lemma lower_map_approx: "lower_map\<cdot>(approx n)\<cdot>xs = approx n\<cdot>xs"
-by (induct xs rule: lower_pd_induct, simp_all)
-
 lemma ep_pair_lower_map: "ep_pair e p \<Longrightarrow> ep_pair (lower_map\<cdot>e) (lower_map\<cdot>p)"
 apply default
 apply (induct_tac x rule: lower_pd_induct, simp_all add: ep_pair.e_inverse)
@@ -495,4 +408,134 @@
 apply (simp_all add: deflation.below monofun_cfun)
 done
 
+(* FIXME: long proof! *)
+lemma finite_deflation_lower_map:
+  assumes "finite_deflation d" shows "finite_deflation (lower_map\<cdot>d)"
+proof (rule finite_deflation_intro)
+  interpret d: finite_deflation d by fact
+  have "deflation d" by fact
+  thus "deflation (lower_map\<cdot>d)" by (rule deflation_lower_map)
+  have "finite (range (\<lambda>x. d\<cdot>x))" by (rule d.finite_range)
+  hence "finite (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))"
+    by (rule finite_vimageI, simp add: inj_on_def Rep_compact_basis_inject)
+  hence "finite (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x)))" by simp
+  hence "finite (Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))"
+    by (rule finite_vimageI, simp add: inj_on_def Rep_pd_basis_inject)
+  hence *: "finite (lower_principal ` Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))" by simp
+  hence "finite (range (\<lambda>xs. lower_map\<cdot>d\<cdot>xs))"
+    apply (rule rev_finite_subset)
+    apply clarsimp
+    apply (induct_tac xs rule: lower_pd.principal_induct)
+    apply (simp add: adm_mem_finite *)
+    apply (rename_tac t, induct_tac t rule: pd_basis_induct)
+    apply (simp only: lower_unit_Rep_compact_basis [symmetric] lower_map_unit)
+    apply simp
+    apply (subgoal_tac "\<exists>b. d\<cdot>(Rep_compact_basis a) = Rep_compact_basis b")
+    apply clarsimp
+    apply (rule imageI)
+    apply (rule vimageI2)
+    apply (simp add: Rep_PDUnit)
+    apply (rule range_eqI)
+    apply (erule sym)
+    apply (rule exI)
+    apply (rule Abs_compact_basis_inverse [symmetric])
+    apply (simp add: d.compact)
+    apply (simp only: lower_plus_principal [symmetric] lower_map_plus)
+    apply clarsimp
+    apply (rule imageI)
+    apply (rule vimageI2)
+    apply (simp add: Rep_PDPlus)
+    done
+  thus "finite {xs. lower_map\<cdot>d\<cdot>xs = xs}"
+    by (rule finite_range_imp_finite_fixes)
+qed
+
+subsection {* Lower powerdomain is an SFP domain *}
+
+definition
+  lower_approx :: "nat \<Rightarrow> udom lower_pd \<rightarrow> udom lower_pd"
+where
+  "lower_approx = (\<lambda>i. lower_map\<cdot>(udom_approx i))"
+
+lemma lower_approx: "approx_chain lower_approx"
+proof (rule approx_chain.intro)
+  show "chain (\<lambda>i. lower_approx i)"
+    unfolding lower_approx_def by simp
+  show "(\<Squnion>i. lower_approx i) = ID"
+    unfolding lower_approx_def
+    by (simp add: lub_distribs lower_map_ID)
+  show "\<And>i. finite_deflation (lower_approx i)"
+    unfolding lower_approx_def
+    by (intro finite_deflation_lower_map finite_deflation_udom_approx)
+qed
+
+definition lower_sfp :: "sfp \<rightarrow> sfp"
+where "lower_sfp = sfp_fun1 lower_approx lower_map"
+
+lemma cast_lower_sfp:
+  "cast\<cdot>(lower_sfp\<cdot>A) =
+    udom_emb lower_approx oo lower_map\<cdot>(cast\<cdot>A) oo udom_prj lower_approx"
+unfolding lower_sfp_def
+apply (rule cast_sfp_fun1 [OF lower_approx])
+apply (erule finite_deflation_lower_map)
+done
+
+instantiation lower_pd :: (sfp) sfp
+begin
+
+definition
+  "emb = udom_emb lower_approx oo lower_map\<cdot>emb"
+
+definition
+  "prj = lower_map\<cdot>prj oo udom_prj lower_approx"
+
+definition
+  "sfp (t::'a lower_pd itself) = lower_sfp\<cdot>SFP('a)"
+
+instance proof
+  show "ep_pair emb (prj :: udom \<rightarrow> 'a lower_pd)"
+    unfolding emb_lower_pd_def prj_lower_pd_def
+    using ep_pair_udom [OF lower_approx]
+    by (intro ep_pair_comp ep_pair_lower_map ep_pair_emb_prj)
+next
+  show "cast\<cdot>SFP('a lower_pd) = emb oo (prj :: udom \<rightarrow> 'a lower_pd)"
+    unfolding emb_lower_pd_def prj_lower_pd_def sfp_lower_pd_def cast_lower_sfp
+    by (simp add: cast_SFP oo_def expand_cfun_eq lower_map_map)
+qed
+
 end
+
+text {* SFP of type constructor = type combinator *}
+
+lemma SFP_lower: "SFP('a lower_pd) = lower_sfp\<cdot>SFP('a)"
+by (rule sfp_lower_pd_def)
+
+
+subsection {* Join *}
+
+definition
+  lower_join :: "'a lower_pd lower_pd \<rightarrow> 'a lower_pd" where
+  "lower_join = (\<Lambda> xss. lower_bind\<cdot>xss\<cdot>(\<Lambda> xs. xs))"
+
+lemma lower_join_unit [simp]:
+  "lower_join\<cdot>{xs}\<flat> = xs"
+unfolding lower_join_def by simp
+
+lemma lower_join_plus [simp]:
+  "lower_join\<cdot>(xss +\<flat> yss) = lower_join\<cdot>xss +\<flat> lower_join\<cdot>yss"
+unfolding lower_join_def by simp
+
+lemma lower_join_map_unit:
+  "lower_join\<cdot>(lower_map\<cdot>lower_unit\<cdot>xs) = xs"
+by (induct xs rule: lower_pd_induct, simp_all)
+
+lemma lower_join_map_join:
+  "lower_join\<cdot>(lower_map\<cdot>lower_join\<cdot>xsss) = lower_join\<cdot>(lower_join\<cdot>xsss)"
+by (induct xsss rule: lower_pd_induct, simp_all)
+
+lemma lower_join_map_map:
+  "lower_join\<cdot>(lower_map\<cdot>(lower_map\<cdot>f)\<cdot>xss) =
+   lower_map\<cdot>f\<cdot>(lower_join\<cdot>xss)"
+by (induct xss rule: lower_pd_induct, simp_all)
+
+end
--- a/src/HOLCF/Powerdomains.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Powerdomains.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -5,290 +5,29 @@
 header {* Powerdomains *}
 
 theory Powerdomains
-imports Representable ConvexPD
-begin
-
-subsection {* Powerdomains are representable *}
-
-text "Upper powerdomain of a representable type is representable."
-
-instantiation upper_pd :: (rep) rep
-begin
-
-definition emb_upper_pd_def: "emb = udom_emb oo upper_map\<cdot>emb"
-definition prj_upper_pd_def: "prj = upper_map\<cdot>prj oo udom_prj"
-
-instance
- apply (intro_classes, unfold emb_upper_pd_def prj_upper_pd_def)
- apply (intro ep_pair_comp ep_pair_upper_map ep_pair_emb_prj ep_pair_udom)
-done
-
-end
-
-text "Lower powerdomain of a representable type is representable."
-
-instantiation lower_pd :: (rep) rep
-begin
-
-definition emb_lower_pd_def: "emb = udom_emb oo lower_map\<cdot>emb"
-definition prj_lower_pd_def: "prj = lower_map\<cdot>prj oo udom_prj"
-
-instance
- apply (intro_classes, unfold emb_lower_pd_def prj_lower_pd_def)
- apply (intro ep_pair_comp ep_pair_lower_map ep_pair_emb_prj ep_pair_udom)
-done
-
-end
-
-text "Convex powerdomain of a representable type is representable."
-
-instantiation convex_pd :: (rep) rep
+imports ConvexPD Domain
 begin
 
-definition emb_convex_pd_def: "emb = udom_emb oo convex_map\<cdot>emb"
-definition prj_convex_pd_def: "prj = convex_map\<cdot>prj oo udom_prj"
-
-instance
- apply (intro_classes, unfold emb_convex_pd_def prj_convex_pd_def)
- apply (intro ep_pair_comp ep_pair_convex_map ep_pair_emb_prj ep_pair_udom)
-done
-
-end
-
-subsection {* Finite deflation lemmas *}
-
-text "TODO: move these lemmas somewhere else"
-
-lemma finite_compact_range_imp_finite_range:
-  fixes d :: "'a::profinite \<rightarrow> 'b::cpo"
-  assumes "finite ((\<lambda>x. d\<cdot>x) ` {x. compact x})"
-  shows "finite (range (\<lambda>x. d\<cdot>x))"
-proof (rule finite_subset [OF _ prems])
-  {
-    fix x :: 'a
-    have "range (\<lambda>i. d\<cdot>(approx i\<cdot>x)) \<subseteq> (\<lambda>x. d\<cdot>x) ` {x. compact x}"
-      by auto
-    hence "finite (range (\<lambda>i. d\<cdot>(approx i\<cdot>x)))"
-      using prems by (rule finite_subset)
-    hence "finite_chain (\<lambda>i. d\<cdot>(approx i\<cdot>x))"
-      by (simp add: finite_range_imp_finch)
-    hence "\<exists>i. (\<Squnion>i. d\<cdot>(approx i\<cdot>x)) = d\<cdot>(approx i\<cdot>x)"
-      by (simp add: finite_chain_def maxinch_is_thelub)
-    hence "\<exists>i. d\<cdot>x = d\<cdot>(approx i\<cdot>x)"
-      by (simp add: lub_distribs)
-    hence "d\<cdot>x \<in> (\<lambda>x. d\<cdot>x) ` {x. compact x}"
-      by auto
-  }
-  thus "range (\<lambda>x. d\<cdot>x) \<subseteq> (\<lambda>x. d\<cdot>x) ` {x. compact x}"
-    by clarsimp
-qed
-
-lemma finite_deflation_upper_map:
-  assumes "finite_deflation d" shows "finite_deflation (upper_map\<cdot>d)"
-proof (rule finite_deflation_intro)
-  interpret d: finite_deflation d by fact
-  have "deflation d" by fact
-  thus "deflation (upper_map\<cdot>d)" by (rule deflation_upper_map)
-  have "finite (range (\<lambda>x. d\<cdot>x))" by (rule d.finite_range)
-  hence "finite (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))"
-    by (rule finite_vimageI, simp add: inj_on_def Rep_compact_basis_inject)
-  hence "finite (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x)))" by simp
-  hence "finite (Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))"
-    by (rule finite_vimageI, simp add: inj_on_def Rep_pd_basis_inject)
-  hence "finite (upper_principal ` Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))" by simp
-  hence "finite ((\<lambda>xs. upper_map\<cdot>d\<cdot>xs) ` range upper_principal)"
-    apply (rule finite_subset [COMP swap_prems_rl])
-    apply (clarsimp, rename_tac t)
-    apply (induct_tac t rule: pd_basis_induct)
-    apply (simp only: upper_unit_Rep_compact_basis [symmetric] upper_map_unit)
-    apply (subgoal_tac "\<exists>b. d\<cdot>(Rep_compact_basis a) = Rep_compact_basis b")
-    apply clarsimp
-    apply (rule imageI)
-    apply (rule vimageI2)
-    apply (simp add: Rep_PDUnit)
-    apply (rule image_eqI)
-    apply (erule sym)
-    apply simp
-    apply (rule exI)
-    apply (rule Abs_compact_basis_inverse [symmetric])
-    apply (simp add: d.compact)
-    apply (simp only: upper_plus_principal [symmetric] upper_map_plus)
-    apply clarsimp
-    apply (rule imageI)
-    apply (rule vimageI2)
-    apply (simp add: Rep_PDPlus)
-    done
-  moreover have "{xs::'a upper_pd. compact xs} = range upper_principal"
-    by (auto dest: upper_pd.compact_imp_principal)
-  ultimately have "finite ((\<lambda>xs. upper_map\<cdot>d\<cdot>xs) ` {xs::'a upper_pd. compact xs})"
-    by simp
-  hence "finite (range (\<lambda>xs. upper_map\<cdot>d\<cdot>xs))"
-    by (rule finite_compact_range_imp_finite_range)
-  thus "finite {xs. upper_map\<cdot>d\<cdot>xs = xs}"
-    by (rule finite_range_imp_finite_fixes)
-qed
-
-lemma finite_deflation_lower_map:
-  assumes "finite_deflation d" shows "finite_deflation (lower_map\<cdot>d)"
-proof (rule finite_deflation_intro)
-  interpret d: finite_deflation d by fact
-  have "deflation d" by fact
-  thus "deflation (lower_map\<cdot>d)" by (rule deflation_lower_map)
-  have "finite (range (\<lambda>x. d\<cdot>x))" by (rule d.finite_range)
-  hence "finite (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))"
-    by (rule finite_vimageI, simp add: inj_on_def Rep_compact_basis_inject)
-  hence "finite (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x)))" by simp
-  hence "finite (Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))"
-    by (rule finite_vimageI, simp add: inj_on_def Rep_pd_basis_inject)
-  hence "finite (lower_principal ` Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))" by simp
-  hence "finite ((\<lambda>xs. lower_map\<cdot>d\<cdot>xs) ` range lower_principal)"
-    apply (rule finite_subset [COMP swap_prems_rl])
-    apply (clarsimp, rename_tac t)
-    apply (induct_tac t rule: pd_basis_induct)
-    apply (simp only: lower_unit_Rep_compact_basis [symmetric] lower_map_unit)
-    apply (subgoal_tac "\<exists>b. d\<cdot>(Rep_compact_basis a) = Rep_compact_basis b")
-    apply clarsimp
-    apply (rule imageI)
-    apply (rule vimageI2)
-    apply (simp add: Rep_PDUnit)
-    apply (rule image_eqI)
-    apply (erule sym)
-    apply simp
-    apply (rule exI)
-    apply (rule Abs_compact_basis_inverse [symmetric])
-    apply (simp add: d.compact)
-    apply (simp only: lower_plus_principal [symmetric] lower_map_plus)
-    apply clarsimp
-    apply (rule imageI)
-    apply (rule vimageI2)
-    apply (simp add: Rep_PDPlus)
-    done
-  moreover have "{xs::'a lower_pd. compact xs} = range lower_principal"
-    by (auto dest: lower_pd.compact_imp_principal)
-  ultimately have "finite ((\<lambda>xs. lower_map\<cdot>d\<cdot>xs) ` {xs::'a lower_pd. compact xs})"
-    by simp
-  hence "finite (range (\<lambda>xs. lower_map\<cdot>d\<cdot>xs))"
-    by (rule finite_compact_range_imp_finite_range)
-  thus "finite {xs. lower_map\<cdot>d\<cdot>xs = xs}"
-    by (rule finite_range_imp_finite_fixes)
-qed
-
-lemma finite_deflation_convex_map:
-  assumes "finite_deflation d" shows "finite_deflation (convex_map\<cdot>d)"
-proof (rule finite_deflation_intro)
-  interpret d: finite_deflation d by fact
-  have "deflation d" by fact
-  thus "deflation (convex_map\<cdot>d)" by (rule deflation_convex_map)
-  have "finite (range (\<lambda>x. d\<cdot>x))" by (rule d.finite_range)
-  hence "finite (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))"
-    by (rule finite_vimageI, simp add: inj_on_def Rep_compact_basis_inject)
-  hence "finite (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x)))" by simp
-  hence "finite (Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))"
-    by (rule finite_vimageI, simp add: inj_on_def Rep_pd_basis_inject)
-  hence "finite (convex_principal ` Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))" by simp
-  hence "finite ((\<lambda>xs. convex_map\<cdot>d\<cdot>xs) ` range convex_principal)"
-    apply (rule finite_subset [COMP swap_prems_rl])
-    apply (clarsimp, rename_tac t)
-    apply (induct_tac t rule: pd_basis_induct)
-    apply (simp only: convex_unit_Rep_compact_basis [symmetric] convex_map_unit)
-    apply (subgoal_tac "\<exists>b. d\<cdot>(Rep_compact_basis a) = Rep_compact_basis b")
-    apply clarsimp
-    apply (rule imageI)
-    apply (rule vimageI2)
-    apply (simp add: Rep_PDUnit)
-    apply (rule image_eqI)
-    apply (erule sym)
-    apply simp
-    apply (rule exI)
-    apply (rule Abs_compact_basis_inverse [symmetric])
-    apply (simp add: d.compact)
-    apply (simp only: convex_plus_principal [symmetric] convex_map_plus)
-    apply clarsimp
-    apply (rule imageI)
-    apply (rule vimageI2)
-    apply (simp add: Rep_PDPlus)
-    done
-  moreover have "{xs::'a convex_pd. compact xs} = range convex_principal"
-    by (auto dest: convex_pd.compact_imp_principal)
-  ultimately have "finite ((\<lambda>xs. convex_map\<cdot>d\<cdot>xs) ` {xs::'a convex_pd. compact xs})"
-    by simp
-  hence "finite (range (\<lambda>xs. convex_map\<cdot>d\<cdot>xs))"
-    by (rule finite_compact_range_imp_finite_range)
-  thus "finite {xs. convex_map\<cdot>d\<cdot>xs = xs}"
-    by (rule finite_range_imp_finite_fixes)
-qed
-
-subsection {* Deflation combinators *}
-
-definition "upper_defl = TypeRep_fun1 upper_map"
-definition "lower_defl = TypeRep_fun1 lower_map"
-definition "convex_defl = TypeRep_fun1 convex_map"
-
-lemma cast_upper_defl:
-  "cast\<cdot>(upper_defl\<cdot>A) = udom_emb oo upper_map\<cdot>(cast\<cdot>A) oo udom_prj"
-unfolding upper_defl_def
-apply (rule cast_TypeRep_fun1)
-apply (erule finite_deflation_upper_map)
-done
-
-lemma cast_lower_defl:
-  "cast\<cdot>(lower_defl\<cdot>A) = udom_emb oo lower_map\<cdot>(cast\<cdot>A) oo udom_prj"
-unfolding lower_defl_def
-apply (rule cast_TypeRep_fun1)
-apply (erule finite_deflation_lower_map)
-done
-
-lemma cast_convex_defl:
-  "cast\<cdot>(convex_defl\<cdot>A) = udom_emb oo convex_map\<cdot>(cast\<cdot>A) oo udom_prj"
-unfolding convex_defl_def
-apply (rule cast_TypeRep_fun1)
-apply (erule finite_deflation_convex_map)
-done
-
-lemma REP_upper: "REP('a upper_pd) = upper_defl\<cdot>REP('a)"
-apply (rule cast_eq_imp_eq, rule ext_cfun)
-apply (simp add: cast_REP cast_upper_defl)
-apply (simp add: prj_upper_pd_def)
-apply (simp add: emb_upper_pd_def)
-apply (simp add: upper_map_map cfcomp1)
-done
-
-lemma REP_lower: "REP('a lower_pd) = lower_defl\<cdot>REP('a)"
-apply (rule cast_eq_imp_eq, rule ext_cfun)
-apply (simp add: cast_REP cast_lower_defl)
-apply (simp add: prj_lower_pd_def)
-apply (simp add: emb_lower_pd_def)
-apply (simp add: lower_map_map cfcomp1)
-done
-
-lemma REP_convex: "REP('a convex_pd) = convex_defl\<cdot>REP('a)"
-apply (rule cast_eq_imp_eq, rule ext_cfun)
-apply (simp add: cast_REP cast_convex_defl)
-apply (simp add: prj_convex_pd_def)
-apply (simp add: emb_convex_pd_def)
-apply (simp add: convex_map_map cfcomp1)
-done
-
 lemma isodefl_upper:
-  "isodefl d t \<Longrightarrow> isodefl (upper_map\<cdot>d) (upper_defl\<cdot>t)"
+  "isodefl d t \<Longrightarrow> isodefl (upper_map\<cdot>d) (upper_sfp\<cdot>t)"
 apply (rule isodeflI)
-apply (simp add: cast_upper_defl cast_isodefl)
+apply (simp add: cast_upper_sfp cast_isodefl)
 apply (simp add: emb_upper_pd_def prj_upper_pd_def)
 apply (simp add: upper_map_map)
 done
 
 lemma isodefl_lower:
-  "isodefl d t \<Longrightarrow> isodefl (lower_map\<cdot>d) (lower_defl\<cdot>t)"
+  "isodefl d t \<Longrightarrow> isodefl (lower_map\<cdot>d) (lower_sfp\<cdot>t)"
 apply (rule isodeflI)
-apply (simp add: cast_lower_defl cast_isodefl)
+apply (simp add: cast_lower_sfp cast_isodefl)
 apply (simp add: emb_lower_pd_def prj_lower_pd_def)
 apply (simp add: lower_map_map)
 done
 
 lemma isodefl_convex:
-  "isodefl d t \<Longrightarrow> isodefl (convex_map\<cdot>d) (convex_defl\<cdot>t)"
+  "isodefl d t \<Longrightarrow> isodefl (convex_map\<cdot>d) (convex_sfp\<cdot>t)"
 apply (rule isodeflI)
-apply (simp add: cast_convex_defl cast_isodefl)
+apply (simp add: cast_convex_sfp cast_isodefl)
 apply (simp add: emb_convex_pd_def prj_convex_pd_def)
 apply (simp add: convex_map_map)
 done
@@ -297,16 +36,16 @@
 
 setup {*
   fold Domain_Isomorphism.add_type_constructor
-    [(@{type_name "upper_pd"}, @{term upper_defl}, @{const_name upper_map},
-        @{thm REP_upper}, @{thm isodefl_upper}, @{thm upper_map_ID},
+    [(@{type_name "upper_pd"}, @{term upper_sfp}, @{const_name upper_map},
+        @{thm SFP_upper}, @{thm isodefl_upper}, @{thm upper_map_ID},
           @{thm deflation_upper_map}),
 
-     (@{type_name "lower_pd"}, @{term lower_defl}, @{const_name lower_map},
-        @{thm REP_lower}, @{thm isodefl_lower}, @{thm lower_map_ID},
+     (@{type_name "lower_pd"}, @{term lower_sfp}, @{const_name lower_map},
+        @{thm SFP_lower}, @{thm isodefl_lower}, @{thm lower_map_ID},
           @{thm deflation_lower_map}),
 
-     (@{type_name "convex_pd"}, @{term convex_defl}, @{const_name convex_map},
-        @{thm REP_convex}, @{thm isodefl_convex}, @{thm convex_map_ID},
+     (@{type_name "convex_pd"}, @{term convex_sfp}, @{const_name convex_map},
+        @{thm SFP_convex}, @{thm isodefl_convex}, @{thm convex_map_ID},
           @{thm deflation_convex_map})]
 *}
 
--- a/src/HOLCF/Representable.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Representable.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -5,97 +5,30 @@
 header {* Representable Types *}
 
 theory Representable
-imports Algebraic Universal Ssum Sprod One Fixrec Domain_Aux
+imports Algebraic Universal Ssum One Fixrec Domain_Aux
 uses
   ("Tools/repdef.ML")
   ("Tools/Domain/domain_isomorphism.ML")
 begin
 
-subsection {* Class of representable types *}
-
-text "Overloaded embedding and projection functions between
-      a representable type and the universal domain."
-
-class rep = bifinite +
-  fixes emb :: "'a::pcpo \<rightarrow> udom"
-  fixes prj :: "udom \<rightarrow> 'a::pcpo"
-  assumes ep_pair_emb_prj: "ep_pair emb prj"
-
-interpretation rep:
-  pcpo_ep_pair
-    "emb :: 'a::rep \<rightarrow> udom"
-    "prj :: udom \<rightarrow> 'a::rep"
-  unfolding pcpo_ep_pair_def
-  by (rule ep_pair_emb_prj)
-
-lemmas emb_inverse = rep.e_inverse
-lemmas emb_prj_below = rep.e_p_below
-lemmas emb_eq_iff = rep.e_eq_iff
-lemmas emb_strict = rep.e_strict
-lemmas prj_strict = rep.p_strict
-
-
-subsection {* Making \emph{rep} the default class *}
-
-text {*
-  From now on, free type variables are assumed to be in class
-  @{term rep}, unless specified otherwise.
-*}
-
-default_sort rep
+default_sort sfp
 
 subsection {* Representations of types *}
 
-text "A TypeRep is an algebraic deflation over the universe of values."
-
-types TypeRep = "udom alg_defl"
-translations (type) "TypeRep" \<leftharpoondown> (type) "udom alg_defl"
-
-definition
-  Rep_of :: "'a::rep itself \<Rightarrow> TypeRep"
-where
-  "Rep_of TYPE('a::rep) =
-    (\<Squnion>i. alg_defl_principal (Abs_fin_defl
-      (emb oo (approx i :: 'a \<rightarrow> 'a) oo prj)))"
-
-syntax "_REP" :: "type \<Rightarrow> TypeRep"  ("(1REP/(1'(_')))")
-translations "REP('t)" \<rightleftharpoons> "CONST Rep_of TYPE('t)"
+lemma emb_prj: "emb\<cdot>((prj\<cdot>x)::'a::sfp) = cast\<cdot>SFP('a)\<cdot>x"
+by (simp add: cast_SFP)
 
-lemma cast_REP:
-  "cast\<cdot>REP('a::rep) = (emb::'a \<rightarrow> udom) oo (prj::udom \<rightarrow> 'a)"
-proof -
-  let ?a = "\<lambda>i. emb oo approx i oo (prj::udom \<rightarrow> 'a)"
-  have a: "\<And>i. finite_deflation (?a i)"
-    apply (rule rep.finite_deflation_e_d_p)
-    apply (rule finite_deflation_approx)
-    done
-  show ?thesis
-    unfolding Rep_of_def
-    apply (subst contlub_cfun_arg)
-    apply (rule chainI)
-    apply (rule alg_defl.principal_mono)
-    apply (rule Abs_fin_defl_mono [OF a a])
-    apply (rule chainE, simp)
-    apply (subst cast_alg_defl_principal)
-    apply (simp add: Abs_fin_defl_inverse a)
-    apply (simp add: expand_cfun_eq lub_distribs)
-    done
-qed
-
-lemma emb_prj: "emb\<cdot>((prj\<cdot>x)::'a::rep) = cast\<cdot>REP('a)\<cdot>x"
-by (simp add: cast_REP)
-
-lemma in_REP_iff:
-  "x ::: REP('a::rep) \<longleftrightarrow> emb\<cdot>((prj\<cdot>x)::'a) = x"
-by (simp add: in_deflation_def cast_REP)
+lemma in_SFP_iff:
+  "x ::: SFP('a::sfp) \<longleftrightarrow> emb\<cdot>((prj\<cdot>x)::'a) = x"
+by (simp add: in_sfp_def cast_SFP)
 
 lemma prj_inverse:
-  "x ::: REP('a::rep) \<Longrightarrow> emb\<cdot>((prj\<cdot>x)::'a) = x"
-by (simp only: in_REP_iff)
+  "x ::: SFP('a::sfp) \<Longrightarrow> emb\<cdot>((prj\<cdot>x)::'a) = x"
+by (simp only: in_SFP_iff)
 
-lemma emb_in_REP [simp]:
-  "emb\<cdot>(x::'a::rep) ::: REP('a)"
-by (simp add: in_REP_iff)
+lemma emb_in_SFP [simp]:
+  "emb\<cdot>(x::'a::sfp) ::: SFP('a)"
+by (simp add: in_SFP_iff)
 
 subsection {* Coerce operator *}
 
@@ -115,16 +48,16 @@
 by (rule ext_cfun, simp add: beta_coerce)
 
 lemma emb_coerce:
-  "REP('a) \<sqsubseteq> REP('b)
+  "SFP('a) \<sqsubseteq> SFP('b)
    \<Longrightarrow> emb\<cdot>((coerce::'a \<rightarrow> 'b)\<cdot>x) = emb\<cdot>x"
  apply (simp add: beta_coerce)
  apply (rule prj_inverse)
- apply (erule subdeflationD)
- apply (rule emb_in_REP)
+ apply (erule sfp_belowD)
+ apply (rule emb_in_SFP)
 done
 
 lemma coerce_prj:
-  "REP('a) \<sqsubseteq> REP('b)
+  "SFP('a) \<sqsubseteq> SFP('b)
    \<Longrightarrow> (coerce::'b \<rightarrow> 'a)\<cdot>(prj\<cdot>x) = prj\<cdot>x"
  apply (simp add: coerce_def)
  apply (rule emb_eq_iff [THEN iffD1])
@@ -136,21 +69,21 @@
 done
 
 lemma coerce_coerce [simp]:
-  "REP('a) \<sqsubseteq> REP('b)
+  "SFP('a) \<sqsubseteq> SFP('b)
    \<Longrightarrow> coerce\<cdot>((coerce::'a \<rightarrow> 'b)\<cdot>x) = coerce\<cdot>x"
-by (simp add: beta_coerce prj_inverse subdeflationD)
+by (simp add: beta_coerce prj_inverse sfp_belowD)
 
 lemma coerce_inverse:
-  "emb\<cdot>(x::'a) ::: REP('b) \<Longrightarrow> coerce\<cdot>(coerce\<cdot>x :: 'b) = x"
+  "emb\<cdot>(x::'a) ::: SFP('b) \<Longrightarrow> coerce\<cdot>(coerce\<cdot>x :: 'b) = x"
 by (simp only: beta_coerce prj_inverse emb_inverse)
 
 lemma coerce_type:
-  "REP('a) \<sqsubseteq> REP('b)
-   \<Longrightarrow> emb\<cdot>((coerce::'a \<rightarrow> 'b)\<cdot>x) ::: REP('a)"
-by (simp add: beta_coerce prj_inverse subdeflationD)
+  "SFP('a) \<sqsubseteq> SFP('b)
+   \<Longrightarrow> emb\<cdot>((coerce::'a \<rightarrow> 'b)\<cdot>x) ::: SFP('a)"
+by (simp add: beta_coerce prj_inverse sfp_belowD)
 
 lemma ep_pair_coerce:
-  "REP('a) \<sqsubseteq> REP('b)
+  "SFP('a) \<sqsubseteq> SFP('b)
    \<Longrightarrow> ep_pair (coerce::'a \<rightarrow> 'b) (coerce::'b \<rightarrow> 'a)"
  apply (rule ep_pair.intro)
   apply simp
@@ -165,30 +98,29 @@
 
 lemma domain_abs_iso:
   fixes abs and rep
-  assumes REP: "REP('b) = REP('a)"
+  assumes SFP: "SFP('b) = SFP('a)"
   assumes abs_def: "abs \<equiv> (coerce :: 'a \<rightarrow> 'b)"
   assumes rep_def: "rep \<equiv> (coerce :: 'b \<rightarrow> 'a)"
   shows "rep\<cdot>(abs\<cdot>x) = x"
-unfolding abs_def rep_def by (simp add: REP)
+unfolding abs_def rep_def by (simp add: SFP)
 
 lemma domain_rep_iso:
   fixes abs and rep
-  assumes REP: "REP('b) = REP('a)"
+  assumes SFP: "SFP('b) = SFP('a)"
   assumes abs_def: "abs \<equiv> (coerce :: 'a \<rightarrow> 'b)"
   assumes rep_def: "rep \<equiv> (coerce :: 'b \<rightarrow> 'a)"
   shows "abs\<cdot>(rep\<cdot>x) = x"
-unfolding abs_def rep_def by (simp add: REP [symmetric])
+unfolding abs_def rep_def by (simp add: SFP [symmetric])
 
 
 subsection {* Proving a subtype is representable *}
 
 text {*
-  Temporarily relax type constraints for @{term "approx"},
-  @{term emb}, and @{term prj}.
+  Temporarily relax type constraints for @{term emb}, and @{term prj}.
 *}
 
 setup {* Sign.add_const_constraint
-  (@{const_name "approx"}, SOME @{typ "nat \<Rightarrow> 'a::cpo \<rightarrow> 'a"}) *}
+  (@{const_name sfp}, SOME @{typ "'a::pcpo itself \<Rightarrow> sfp"}) *}
 
 setup {* Sign.add_const_constraint
   (@{const_name emb}, SOME @{typ "'a::pcpo \<rightarrow> udom"}) *}
@@ -196,118 +128,19 @@
 setup {* Sign.add_const_constraint
   (@{const_name prj}, SOME @{typ "udom \<rightarrow> 'a::pcpo"}) *}
 
-definition
-  repdef_approx ::
-    "('a::pcpo \<Rightarrow> udom) \<Rightarrow> (udom \<Rightarrow> 'a) \<Rightarrow> udom alg_defl \<Rightarrow> nat \<Rightarrow> 'a \<rightarrow> 'a"
-where
-  "repdef_approx Rep Abs t = (\<lambda>i. \<Lambda> x. Abs (cast\<cdot>(approx i\<cdot>t)\<cdot>(Rep x)))"
-
 lemma typedef_rep_class:
   fixes Rep :: "'a::pcpo \<Rightarrow> udom"
   fixes Abs :: "udom \<Rightarrow> 'a::pcpo"
-  fixes t :: TypeRep
+  fixes t :: sfp
   assumes type: "type_definition Rep Abs {x. x ::: t}"
   assumes below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
   assumes emb: "emb \<equiv> (\<Lambda> x. Rep x)"
   assumes prj: "prj \<equiv> (\<Lambda> x. Abs (cast\<cdot>t\<cdot>x))"
-  assumes approx: "(approx :: nat \<Rightarrow> 'a \<rightarrow> 'a) \<equiv> repdef_approx Rep Abs t"
-  shows "OFCLASS('a, rep_class)"
+  assumes sfp: "sfp \<equiv> (\<lambda> a::'a itself. t)"
+  shows "OFCLASS('a, sfp_class)"
 proof
   have adm: "adm (\<lambda>x. x \<in> {x. x ::: t})"
-    by (simp add: adm_in_deflation)
-  have emb_beta: "\<And>x. emb\<cdot>x = Rep x"
-    unfolding emb
-    apply (rule beta_cfun)
-    apply (rule typedef_cont_Rep [OF type below adm])
-    done
-  have prj_beta: "\<And>y. prj\<cdot>y = Abs (cast\<cdot>t\<cdot>y)"
-    unfolding prj
-    apply (rule beta_cfun)
-    apply (rule typedef_cont_Abs [OF type below adm])
-    apply simp_all
-    done
-  have cast_cast_approx:
-    "\<And>i x. cast\<cdot>t\<cdot>(cast\<cdot>(approx i\<cdot>t)\<cdot>x) = cast\<cdot>(approx i\<cdot>t)\<cdot>x"
-    apply (rule cast_fixed)
-    apply (rule subdeflationD)
-    apply (rule approx.below)
-    apply (rule cast_in_deflation)
-    done
-  have approx':
-    "approx = (\<lambda>i. \<Lambda>(x::'a). prj\<cdot>(cast\<cdot>(approx i\<cdot>t)\<cdot>(emb\<cdot>x)))"
-    unfolding approx repdef_approx_def
-    apply (subst cast_cast_approx [symmetric])
-    apply (simp add: prj_beta [symmetric] emb_beta [symmetric])
-    done
-  have emb_in_deflation: "\<And>x::'a. emb\<cdot>x ::: t"
-    using type_definition.Rep [OF type]
-    by (simp add: emb_beta)
-  have prj_emb: "\<And>x::'a. prj\<cdot>(emb\<cdot>x) = x"
-    unfolding prj_beta
-    apply (simp add: cast_fixed [OF emb_in_deflation])
-    apply (simp add: emb_beta type_definition.Rep_inverse [OF type])
-    done
-  have emb_prj: "\<And>y. emb\<cdot>(prj\<cdot>y :: 'a) = cast\<cdot>t\<cdot>y"
-    unfolding prj_beta emb_beta
-    by (simp add: type_definition.Abs_inverse [OF type])
-  show "ep_pair (emb :: 'a \<rightarrow> udom) prj"
-    apply default
-    apply (simp add: prj_emb)
-    apply (simp add: emb_prj cast.below)
-    done
-  show "chain (approx :: nat \<Rightarrow> 'a \<rightarrow> 'a)"
-    unfolding approx' by simp
-  show "\<And>x::'a. (\<Squnion>i. approx i\<cdot>x) = x"
-    unfolding approx'
-    apply (simp add: lub_distribs)
-    apply (subst cast_fixed [OF emb_in_deflation])
-    apply (rule prj_emb)
-    done
-  show "\<And>(i::nat) (x::'a). approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
-    unfolding approx'
-    apply simp
-    apply (simp add: emb_prj)
-    apply (simp add: cast_cast_approx)
-    done
-  show "\<And>i::nat. finite {x::'a. approx i\<cdot>x = x}"
-    apply (rule_tac B="(\<lambda>x. prj\<cdot>x) ` {x. cast\<cdot>(approx i\<cdot>t)\<cdot>x = x}"
-           in finite_subset)
-    apply (clarsimp simp add: approx')
-    apply (drule_tac f="\<lambda>x. emb\<cdot>x" in arg_cong)
-    apply (rule image_eqI)
-    apply (rule prj_emb [symmetric])
-    apply (simp add: emb_prj)
-    apply (simp add: cast_cast_approx)
-    apply (rule finite_imageI)
-    apply (simp add: cast_approx_fixed_iff)
-    apply (simp add: Collect_conj_eq)
-    apply (simp add: finite_fixes_approx)
-    done
-qed
-
-text {* Restore original typing constraints *}
-
-setup {* Sign.add_const_constraint
-  (@{const_name "approx"}, SOME @{typ "nat \<Rightarrow> 'a::profinite \<rightarrow> 'a"}) *}
-
-setup {* Sign.add_const_constraint
-  (@{const_name emb}, SOME @{typ "'a::rep \<rightarrow> udom"}) *}
-
-setup {* Sign.add_const_constraint
-  (@{const_name prj}, SOME @{typ "udom \<rightarrow> 'a::rep"}) *}
-
-lemma typedef_REP:
-  fixes Rep :: "'a::rep \<Rightarrow> udom"
-  fixes Abs :: "udom \<Rightarrow> 'a::rep"
-  fixes t :: TypeRep
-  assumes type: "type_definition Rep Abs {x. x ::: t}"
-  assumes below: "op \<sqsubseteq> \<equiv> \<lambda>x y. Rep x \<sqsubseteq> Rep y"
-  assumes emb: "emb \<equiv> (\<Lambda> x. Rep x)"
-  assumes prj: "prj \<equiv> (\<Lambda> x. Abs (cast\<cdot>t\<cdot>x))"
-  shows "REP('a) = t"
-proof -
-  have adm: "adm (\<lambda>x. x \<in> {x. x ::: t})"
-    by (simp add: adm_in_deflation)
+    by (simp add: adm_in_sfp)
   have emb_beta: "\<And>x. emb\<cdot>x = Rep x"
     unfolding emb
     apply (rule beta_cfun)
@@ -319,313 +152,51 @@
     apply (rule typedef_cont_Abs [OF type below adm])
     apply simp_all
     done
-  have emb_in_deflation: "\<And>x::'a. emb\<cdot>x ::: t"
+  have emb_in_sfp: "\<And>x::'a. emb\<cdot>x ::: t"
     using type_definition.Rep [OF type]
     by (simp add: emb_beta)
   have prj_emb: "\<And>x::'a. prj\<cdot>(emb\<cdot>x) = x"
     unfolding prj_beta
-    apply (simp add: cast_fixed [OF emb_in_deflation])
+    apply (simp add: cast_fixed [OF emb_in_sfp])
     apply (simp add: emb_beta type_definition.Rep_inverse [OF type])
     done
   have emb_prj: "\<And>y. emb\<cdot>(prj\<cdot>y :: 'a) = cast\<cdot>t\<cdot>y"
     unfolding prj_beta emb_beta
     by (simp add: type_definition.Abs_inverse [OF type])
-  show "REP('a) = t"
-    apply (rule cast_eq_imp_eq, rule ext_cfun)
-    apply (simp add: cast_REP emb_prj)
+  show "ep_pair (emb :: 'a \<rightarrow> udom) prj"
+    apply default
+    apply (simp add: prj_emb)
+    apply (simp add: emb_prj cast.below)
     done
+  show "cast\<cdot>SFP('a) = emb oo (prj :: udom \<rightarrow> 'a)"
+    by (rule ext_cfun, simp add: sfp emb_prj)
 qed
 
-lemma adm_mem_Collect_in_deflation: "adm (\<lambda>x. x \<in> {x. x ::: A})"
-unfolding mem_Collect_eq by (rule adm_in_deflation)
+lemma typedef_SFP:
+  assumes "sfp \<equiv> (\<lambda>a::'a::pcpo itself. t)"
+  shows "SFP('a::pcpo) = t"
+unfolding assms ..
+
+text {* Restore original typing constraints *}
+
+setup {* Sign.add_const_constraint
+  (@{const_name sfp}, SOME @{typ "'a::sfp itself \<Rightarrow> sfp"}) *}
+
+setup {* Sign.add_const_constraint
+  (@{const_name emb}, SOME @{typ "'a::sfp \<rightarrow> udom"}) *}
+
+setup {* Sign.add_const_constraint
+  (@{const_name prj}, SOME @{typ "udom \<rightarrow> 'a::sfp"}) *}
+
+lemma adm_mem_Collect_in_sfp: "adm (\<lambda>x. x \<in> {x. x ::: A})"
+unfolding mem_Collect_eq by (rule adm_in_sfp)
 
 use "Tools/repdef.ML"
 
-
-subsection {* Instances of class \emph{rep} *}
-
-subsubsection {* Universal Domain *}
-
-text "The Universal Domain itself is trivially representable."
-
-instantiation udom :: rep
-begin
-
-definition emb_udom_def [simp]: "emb = (ID :: udom \<rightarrow> udom)"
-definition prj_udom_def [simp]: "prj = (ID :: udom \<rightarrow> udom)"
-
-instance
- apply (intro_classes)
- apply (simp_all add: ep_pair.intro)
-done
-
-end
-
-subsubsection {* Lifted types *}
-
-instantiation lift :: (countable) rep
-begin
-
-definition emb_lift_def:
-  "emb = udom_emb oo (FLIFT x. Def (to_nat x))"
-
-definition prj_lift_def:
-  "prj = (FLIFT n. if (\<exists>x::'a::countable. n = to_nat x)
-                    then Def (THE x::'a. n = to_nat x) else \<bottom>) oo udom_prj"
-
-instance
- apply (intro_classes, unfold emb_lift_def prj_lift_def)
- apply (rule ep_pair_comp [OF _ ep_pair_udom])
- apply (rule ep_pair.intro)
-  apply (case_tac x, simp, simp)
- apply (case_tac y, simp, clarsimp)
-done
-
-end
-
-subsubsection {* Representable type constructors *}
-
-text "Functions between representable types are representable."
-
-instantiation cfun :: (rep, rep) rep
-begin
-
-definition emb_cfun_def: "emb = udom_emb oo cfun_map\<cdot>prj\<cdot>emb"
-definition prj_cfun_def: "prj = cfun_map\<cdot>emb\<cdot>prj oo udom_prj"
-
-instance
- apply (intro_classes, unfold emb_cfun_def prj_cfun_def)
- apply (intro ep_pair_comp ep_pair_cfun_map ep_pair_emb_prj ep_pair_udom)
-done
-
-end
-
-text "Strict products of representable types are representable."
-
-instantiation sprod :: (rep, rep) rep
-begin
-
-definition emb_sprod_def: "emb = udom_emb oo sprod_map\<cdot>emb\<cdot>emb"
-definition prj_sprod_def: "prj = sprod_map\<cdot>prj\<cdot>prj oo udom_prj"
-
-instance
- apply (intro_classes, unfold emb_sprod_def prj_sprod_def)
- apply (intro ep_pair_comp ep_pair_sprod_map ep_pair_emb_prj ep_pair_udom)
-done
-
-end
-
-text "Strict sums of representable types are representable."
-
-instantiation ssum :: (rep, rep) rep
-begin
-
-definition emb_ssum_def: "emb = udom_emb oo ssum_map\<cdot>emb\<cdot>emb"
-definition prj_ssum_def: "prj = ssum_map\<cdot>prj\<cdot>prj oo udom_prj"
-
-instance
- apply (intro_classes, unfold emb_ssum_def prj_ssum_def)
- apply (intro ep_pair_comp ep_pair_ssum_map ep_pair_emb_prj ep_pair_udom)
-done
-
-end
-
-text "Up of a representable type is representable."
-
-instantiation "u" :: (rep) rep
-begin
-
-definition emb_u_def: "emb = udom_emb oo u_map\<cdot>emb"
-definition prj_u_def: "prj = u_map\<cdot>prj oo udom_prj"
-
-instance
- apply (intro_classes, unfold emb_u_def prj_u_def)
- apply (intro ep_pair_comp ep_pair_u_map ep_pair_emb_prj ep_pair_udom)
-done
-
-end
-
-text "Cartesian products of representable types are representable."
-
-instantiation prod :: (rep, rep) rep
-begin
-
-definition emb_cprod_def: "emb = udom_emb oo cprod_map\<cdot>emb\<cdot>emb"
-definition prj_cprod_def: "prj = cprod_map\<cdot>prj\<cdot>prj oo udom_prj"
-
-instance
- apply (intro_classes, unfold emb_cprod_def prj_cprod_def)
- apply (intro ep_pair_comp ep_pair_cprod_map ep_pair_emb_prj ep_pair_udom)
-done
-
-end
-
-subsection {* Type combinators *}
-
-definition
-  TypeRep_fun1 ::
-    "((udom \<rightarrow> udom) \<rightarrow> ('a \<rightarrow> 'a))
-      \<Rightarrow> (TypeRep \<rightarrow> TypeRep)"
-where
-  "TypeRep_fun1 f =
-    alg_defl.basis_fun (\<lambda>a.
-      alg_defl_principal (
-        Abs_fin_defl (udom_emb oo f\<cdot>(Rep_fin_defl a) oo udom_prj)))"
-
-definition
-  TypeRep_fun2 ::
-    "((udom \<rightarrow> udom) \<rightarrow> (udom \<rightarrow> udom) \<rightarrow> ('a \<rightarrow> 'a))
-      \<Rightarrow> (TypeRep \<rightarrow> TypeRep \<rightarrow> TypeRep)"
-where
-  "TypeRep_fun2 f =
-    alg_defl.basis_fun (\<lambda>a.
-      alg_defl.basis_fun (\<lambda>b.
-        alg_defl_principal (
-          Abs_fin_defl (udom_emb oo
-            f\<cdot>(Rep_fin_defl a)\<cdot>(Rep_fin_defl b) oo udom_prj))))"
-
-definition "cfun_defl = TypeRep_fun2 cfun_map"
-definition "ssum_defl = TypeRep_fun2 ssum_map"
-definition "sprod_defl = TypeRep_fun2 sprod_map"
-definition "cprod_defl = TypeRep_fun2 cprod_map"
-definition "u_defl = TypeRep_fun1 u_map"
-
-lemma Rep_fin_defl_mono: "a \<sqsubseteq> b \<Longrightarrow> Rep_fin_defl a \<sqsubseteq> Rep_fin_defl b"
-unfolding below_fin_defl_def .
-
-lemma cast_TypeRep_fun1:
-  assumes f: "\<And>a. finite_deflation a \<Longrightarrow> finite_deflation (f\<cdot>a)"
-  shows "cast\<cdot>(TypeRep_fun1 f\<cdot>A) = udom_emb oo f\<cdot>(cast\<cdot>A) oo udom_prj"
-proof -
-  have 1: "\<And>a. finite_deflation (udom_emb oo f\<cdot>(Rep_fin_defl a) oo udom_prj)"
-    apply (rule ep_pair.finite_deflation_e_d_p [OF ep_pair_udom])
-    apply (rule f, rule finite_deflation_Rep_fin_defl)
-    done
-  show ?thesis
-    by (induct A rule: alg_defl.principal_induct, simp)
-       (simp only: TypeRep_fun1_def
-                   alg_defl.basis_fun_principal
-                   alg_defl.basis_fun_mono
-                   alg_defl.principal_mono
-                   Abs_fin_defl_mono [OF 1 1]
-                   monofun_cfun below_refl
-                   Rep_fin_defl_mono
-                   cast_alg_defl_principal
-                   Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1])
-qed
-
-lemma cast_TypeRep_fun2:
-  assumes f: "\<And>a b. finite_deflation a \<Longrightarrow> finite_deflation b \<Longrightarrow>
-                finite_deflation (f\<cdot>a\<cdot>b)"
-  shows "cast\<cdot>(TypeRep_fun2 f\<cdot>A\<cdot>B) =
-    udom_emb oo f\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj"
-proof -
-  have 1: "\<And>a b. finite_deflation
-           (udom_emb oo f\<cdot>(Rep_fin_defl a)\<cdot>(Rep_fin_defl b) oo udom_prj)"
-    apply (rule ep_pair.finite_deflation_e_d_p [OF ep_pair_udom])
-    apply (rule f, (rule finite_deflation_Rep_fin_defl)+)
-    done
-  show ?thesis
-    by (induct A B rule: alg_defl.principal_induct2, simp, simp)
-       (simp only: TypeRep_fun2_def
-                   alg_defl.basis_fun_principal
-                   alg_defl.basis_fun_mono
-                   alg_defl.principal_mono
-                   Abs_fin_defl_mono [OF 1 1]
-                   monofun_cfun below_refl
-                   Rep_fin_defl_mono
-                   cast_alg_defl_principal
-                   Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1])
-qed
-
-lemma cast_cfun_defl:
-  "cast\<cdot>(cfun_defl\<cdot>A\<cdot>B) = udom_emb oo cfun_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj"
-unfolding cfun_defl_def
-apply (rule cast_TypeRep_fun2)
-apply (erule (1) finite_deflation_cfun_map)
-done
-
-lemma cast_ssum_defl:
-  "cast\<cdot>(ssum_defl\<cdot>A\<cdot>B) = udom_emb oo ssum_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj"
-unfolding ssum_defl_def
-apply (rule cast_TypeRep_fun2)
-apply (erule (1) finite_deflation_ssum_map)
-done
-
-lemma cast_sprod_defl:
-  "cast\<cdot>(sprod_defl\<cdot>A\<cdot>B) = udom_emb oo sprod_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj"
-unfolding sprod_defl_def
-apply (rule cast_TypeRep_fun2)
-apply (erule (1) finite_deflation_sprod_map)
-done
-
-lemma cast_cprod_defl:
-  "cast\<cdot>(cprod_defl\<cdot>A\<cdot>B) = udom_emb oo cprod_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj"
-unfolding cprod_defl_def
-apply (rule cast_TypeRep_fun2)
-apply (erule (1) finite_deflation_cprod_map)
-done
-
-lemma cast_u_defl:
-  "cast\<cdot>(u_defl\<cdot>A) = udom_emb oo u_map\<cdot>(cast\<cdot>A) oo udom_prj"
-unfolding u_defl_def
-apply (rule cast_TypeRep_fun1)
-apply (erule finite_deflation_u_map)
-done
-
-text {* REP of type constructor = type combinator *}
-
-lemma REP_cfun: "REP('a \<rightarrow> 'b) = cfun_defl\<cdot>REP('a)\<cdot>REP('b)"
-apply (rule cast_eq_imp_eq, rule ext_cfun)
-apply (simp add: cast_REP cast_cfun_defl)
-apply (simp add: cfun_map_def)
-apply (simp only: prj_cfun_def emb_cfun_def)
-apply (simp add: expand_cfun_eq ep_pair.e_eq_iff [OF ep_pair_udom])
-done
-
-lemma REP_ssum: "REP('a \<oplus> 'b) = ssum_defl\<cdot>REP('a)\<cdot>REP('b)"
-apply (rule cast_eq_imp_eq, rule ext_cfun)
-apply (simp add: cast_REP cast_ssum_defl)
-apply (simp add: prj_ssum_def)
-apply (simp add: emb_ssum_def)
-apply (simp add: ssum_map_map cfcomp1)
-done
-
-lemma REP_sprod: "REP('a \<otimes> 'b) = sprod_defl\<cdot>REP('a)\<cdot>REP('b)"
-apply (rule cast_eq_imp_eq, rule ext_cfun)
-apply (simp add: cast_REP cast_sprod_defl)
-apply (simp add: prj_sprod_def)
-apply (simp add: emb_sprod_def)
-apply (simp add: sprod_map_map cfcomp1)
-done
-
-lemma REP_cprod: "REP('a \<times> 'b) = cprod_defl\<cdot>REP('a)\<cdot>REP('b)"
-apply (rule cast_eq_imp_eq, rule ext_cfun)
-apply (simp add: cast_REP cast_cprod_defl)
-apply (simp add: prj_cprod_def)
-apply (simp add: emb_cprod_def)
-apply (simp add: cprod_map_map cfcomp1)
-done
-
-lemma REP_up: "REP('a u) = u_defl\<cdot>REP('a)"
-apply (rule cast_eq_imp_eq, rule ext_cfun)
-apply (simp add: cast_REP cast_u_defl)
-apply (simp add: prj_u_def)
-apply (simp add: emb_u_def)
-apply (simp add: u_map_map cfcomp1)
-done
-
-lemmas REP_simps =
-  REP_cfun
-  REP_ssum
-  REP_sprod
-  REP_cprod
-  REP_up
-
 subsection {* Isomorphic deflations *}
 
 definition
-  isodefl :: "('a::rep \<rightarrow> 'a) \<Rightarrow> udom alg_defl \<Rightarrow> bool"
+  isodefl :: "('a::sfp \<rightarrow> 'a) \<Rightarrow> sfp \<Rightarrow> bool"
 where
   "isodefl d t \<longleftrightarrow> cast\<cdot>t = emb oo d oo prj"
 
@@ -640,10 +211,10 @@
 by (drule cfun_fun_cong [where x="\<bottom>"], simp)
 
 lemma isodefl_imp_deflation:
-  fixes d :: "'a::rep \<rightarrow> 'a"
+  fixes d :: "'a::sfp \<rightarrow> 'a"
   assumes "isodefl d t" shows "deflation d"
 proof
-  note prems [unfolded isodefl_def, simp]
+  note assms [unfolded isodefl_def, simp]
   fix x :: 'a
   show "d\<cdot>(d\<cdot>x) = d\<cdot>x"
     using cast.idem [of t "emb\<cdot>x"] by simp
@@ -651,12 +222,12 @@
     using cast.below [of t "emb\<cdot>x"] by simp
 qed
 
-lemma isodefl_ID_REP: "isodefl (ID :: 'a \<rightarrow> 'a) REP('a)"
-unfolding isodefl_def by (simp add: cast_REP)
+lemma isodefl_ID_SFP: "isodefl (ID :: 'a \<rightarrow> 'a) SFP('a)"
+unfolding isodefl_def by (simp add: cast_SFP)
 
-lemma isodefl_REP_imp_ID: "isodefl (d :: 'a \<rightarrow> 'a) REP('a) \<Longrightarrow> d = ID"
+lemma isodefl_SFP_imp_ID: "isodefl (d :: 'a \<rightarrow> 'a) SFP('a) \<Longrightarrow> d = ID"
 unfolding isodefl_def
-apply (simp add: cast_REP)
+apply (simp add: cast_SFP)
 apply (simp add: expand_cfun_eq)
 apply (rule allI)
 apply (drule_tac x="emb\<cdot>x" in spec)
@@ -684,66 +255,66 @@
 apply (rule isodefl_lub, simp, simp)
 apply (induct_tac i)
 apply (simp add: isodefl_bottom)
-apply (simp add: prems)
+apply (simp add: assms)
 done
 
 lemma isodefl_coerce:
   fixes d :: "'a \<rightarrow> 'a"
-  assumes REP: "REP('b) = REP('a)"
+  assumes SFP: "SFP('b) = SFP('a)"
   shows "isodefl d t \<Longrightarrow> isodefl (coerce oo d oo coerce :: 'b \<rightarrow> 'b) t"
 unfolding isodefl_def
 apply (simp add: expand_cfun_eq)
-apply (simp add: emb_coerce coerce_prj REP)
+apply (simp add: emb_coerce coerce_prj SFP)
 done
 
 lemma isodefl_abs_rep:
   fixes abs and rep and d
-  assumes REP: "REP('b) = REP('a)"
+  assumes SFP: "SFP('b) = SFP('a)"
   assumes abs_def: "abs \<equiv> (coerce :: 'a \<rightarrow> 'b)"
   assumes rep_def: "rep \<equiv> (coerce :: 'b \<rightarrow> 'a)"
   shows "isodefl d t \<Longrightarrow> isodefl (abs oo d oo rep) t"
-unfolding abs_def rep_def using REP by (rule isodefl_coerce)
+unfolding abs_def rep_def using SFP by (rule isodefl_coerce)
 
 lemma isodefl_cfun:
   "isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>
-    isodefl (cfun_map\<cdot>d1\<cdot>d2) (cfun_defl\<cdot>t1\<cdot>t2)"
+    isodefl (cfun_map\<cdot>d1\<cdot>d2) (cfun_sfp\<cdot>t1\<cdot>t2)"
 apply (rule isodeflI)
-apply (simp add: cast_cfun_defl cast_isodefl)
+apply (simp add: cast_cfun_sfp cast_isodefl)
 apply (simp add: emb_cfun_def prj_cfun_def)
 apply (simp add: cfun_map_map cfcomp1)
 done
 
 lemma isodefl_ssum:
   "isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>
-    isodefl (ssum_map\<cdot>d1\<cdot>d2) (ssum_defl\<cdot>t1\<cdot>t2)"
+    isodefl (ssum_map\<cdot>d1\<cdot>d2) (ssum_sfp\<cdot>t1\<cdot>t2)"
 apply (rule isodeflI)
-apply (simp add: cast_ssum_defl cast_isodefl)
+apply (simp add: cast_ssum_sfp cast_isodefl)
 apply (simp add: emb_ssum_def prj_ssum_def)
 apply (simp add: ssum_map_map isodefl_strict)
 done
 
 lemma isodefl_sprod:
   "isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>
-    isodefl (sprod_map\<cdot>d1\<cdot>d2) (sprod_defl\<cdot>t1\<cdot>t2)"
+    isodefl (sprod_map\<cdot>d1\<cdot>d2) (sprod_sfp\<cdot>t1\<cdot>t2)"
 apply (rule isodeflI)
-apply (simp add: cast_sprod_defl cast_isodefl)
+apply (simp add: cast_sprod_sfp cast_isodefl)
 apply (simp add: emb_sprod_def prj_sprod_def)
 apply (simp add: sprod_map_map isodefl_strict)
 done
 
 lemma isodefl_cprod:
   "isodefl d1 t1 \<Longrightarrow> isodefl d2 t2 \<Longrightarrow>
-    isodefl (cprod_map\<cdot>d1\<cdot>d2) (cprod_defl\<cdot>t1\<cdot>t2)"
+    isodefl (cprod_map\<cdot>d1\<cdot>d2) (prod_sfp\<cdot>t1\<cdot>t2)"
 apply (rule isodeflI)
-apply (simp add: cast_cprod_defl cast_isodefl)
-apply (simp add: emb_cprod_def prj_cprod_def)
+apply (simp add: cast_prod_sfp cast_isodefl)
+apply (simp add: emb_prod_def prj_prod_def)
 apply (simp add: cprod_map_map cfcomp1)
 done
 
 lemma isodefl_u:
-  "isodefl d t \<Longrightarrow> isodefl (u_map\<cdot>d) (u_defl\<cdot>t)"
+  "isodefl d t \<Longrightarrow> isodefl (u_map\<cdot>d) (u_sfp\<cdot>t)"
 apply (rule isodeflI)
-apply (simp add: cast_u_defl cast_isodefl)
+apply (simp add: cast_u_sfp cast_isodefl)
 apply (simp add: emb_u_def prj_u_def)
 apply (simp add: u_map_map)
 done
@@ -754,19 +325,19 @@
 
 setup {*
   fold Domain_Isomorphism.add_type_constructor
-    [(@{type_name cfun}, @{term cfun_defl}, @{const_name cfun_map}, @{thm REP_cfun},
+    [(@{type_name cfun}, @{term cfun_sfp}, @{const_name cfun_map}, @{thm SFP_cfun},
         @{thm isodefl_cfun}, @{thm cfun_map_ID}, @{thm deflation_cfun_map}),
 
-     (@{type_name ssum}, @{term ssum_defl}, @{const_name ssum_map}, @{thm REP_ssum},
+     (@{type_name ssum}, @{term ssum_sfp}, @{const_name ssum_map}, @{thm SFP_ssum},
         @{thm isodefl_ssum}, @{thm ssum_map_ID}, @{thm deflation_ssum_map}),
 
-     (@{type_name sprod}, @{term sprod_defl}, @{const_name sprod_map}, @{thm REP_sprod},
+     (@{type_name sprod}, @{term sprod_sfp}, @{const_name sprod_map}, @{thm SFP_sprod},
         @{thm isodefl_sprod}, @{thm sprod_map_ID}, @{thm deflation_sprod_map}),
 
-     (@{type_name prod}, @{term cprod_defl}, @{const_name cprod_map}, @{thm REP_cprod},
+     (@{type_name prod}, @{term cprod_sfp}, @{const_name cprod_map}, @{thm SFP_prod},
         @{thm isodefl_cprod}, @{thm cprod_map_ID}, @{thm deflation_cprod_map}),
 
-     (@{type_name "u"}, @{term u_defl}, @{const_name u_map}, @{thm REP_up},
+     (@{type_name "u"}, @{term u_sfp}, @{const_name u_map}, @{thm SFP_u},
         @{thm isodefl_u}, @{thm u_map_ID}, @{thm deflation_u_map})]
 *}
 
--- a/src/HOLCF/Sprod.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Sprod.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -5,7 +5,7 @@
 header {* The type of strict products *}
 
 theory Sprod
-imports Bifinite
+imports Algebraic
 begin
 
 default_sort pcpo
@@ -310,37 +310,66 @@
     by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
 qed
 
-subsection {* Strict product is a bifinite domain *}
+subsection {* Strict product is an SFP domain *}
+
+definition
+  sprod_approx :: "nat \<Rightarrow> udom \<otimes> udom \<rightarrow> udom \<otimes> udom"
+where
+  "sprod_approx = (\<lambda>i. sprod_map\<cdot>(udom_approx i)\<cdot>(udom_approx i))"
 
-instantiation sprod :: (bifinite, bifinite) bifinite
+lemma sprod_approx: "approx_chain sprod_approx"
+proof (rule approx_chain.intro)
+  show "chain (\<lambda>i. sprod_approx i)"
+    unfolding sprod_approx_def by simp
+  show "(\<Squnion>i. sprod_approx i) = ID"
+    unfolding sprod_approx_def
+    by (simp add: lub_distribs sprod_map_ID)
+  show "\<And>i. finite_deflation (sprod_approx i)"
+    unfolding sprod_approx_def
+    by (intro finite_deflation_sprod_map finite_deflation_udom_approx)
+qed
+
+definition sprod_sfp :: "sfp \<rightarrow> sfp \<rightarrow> sfp"
+where "sprod_sfp = sfp_fun2 sprod_approx sprod_map"
+
+lemma cast_sprod_sfp:
+  "cast\<cdot>(sprod_sfp\<cdot>A\<cdot>B) =
+    udom_emb sprod_approx oo
+      sprod_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo
+        udom_prj sprod_approx"
+unfolding sprod_sfp_def
+apply (rule cast_sfp_fun2 [OF sprod_approx])
+apply (erule (1) finite_deflation_sprod_map)
+done
+
+instantiation sprod :: (sfp, sfp) sfp
 begin
 
 definition
-  approx_sprod_def:
-    "approx = (\<lambda>n. sprod_map\<cdot>(approx n)\<cdot>(approx n))"
+  "emb = udom_emb sprod_approx oo sprod_map\<cdot>emb\<cdot>emb"
+
+definition
+  "prj = sprod_map\<cdot>prj\<cdot>prj oo udom_prj sprod_approx"
+
+definition
+  "sfp (t::('a \<otimes> 'b) itself) = sprod_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
 
 instance proof
-  fix i :: nat and x :: "'a \<otimes> 'b"
-  show "chain (approx :: nat \<Rightarrow> 'a \<otimes> 'b \<rightarrow> 'a \<otimes> 'b)"
-    unfolding approx_sprod_def by simp
-  show "(\<Squnion>i. approx i\<cdot>x) = x"
-    unfolding approx_sprod_def sprod_map_def
-    by (simp add: lub_distribs eta_cfun)
-  show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
-    unfolding approx_sprod_def sprod_map_def
-    by (simp add: ssplit_def strictify_conv_if)
-  show "finite {x::'a \<otimes> 'b. approx i\<cdot>x = x}"
-    unfolding approx_sprod_def
-    by (intro finite_deflation.finite_fixes
-              finite_deflation_sprod_map
-              finite_deflation_approx)
+  show "ep_pair emb (prj :: udom \<rightarrow> 'a \<otimes> 'b)"
+    unfolding emb_sprod_def prj_sprod_def
+    using ep_pair_udom [OF sprod_approx]
+    by (intro ep_pair_comp ep_pair_sprod_map ep_pair_emb_prj)
+next
+  show "cast\<cdot>SFP('a \<otimes> 'b) = emb oo (prj :: udom \<rightarrow> 'a \<otimes> 'b)"
+    unfolding emb_sprod_def prj_sprod_def sfp_sprod_def cast_sprod_sfp
+    by (simp add: cast_SFP oo_def expand_cfun_eq sprod_map_map)
 qed
 
 end
 
-lemma approx_spair [simp]:
-  "approx i\<cdot>(:x, y:) = (:approx i\<cdot>x, approx i\<cdot>y:)"
-unfolding approx_sprod_def sprod_map_def
-by (simp add: ssplit_def strictify_conv_if)
+text {* SFP of type constructor = type combinator *}
+
+lemma SFP_sprod: "SFP('a::sfp \<otimes> 'b::sfp) = sprod_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
+by (rule sfp_sprod_def)
 
 end
--- a/src/HOLCF/Ssum.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Ssum.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -295,37 +295,64 @@
     by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
 qed
 
-subsection {* Strict sum is a bifinite domain *}
+subsection {* Strict sum is an SFP domain *}
+
+definition
+  ssum_approx :: "nat \<Rightarrow> udom \<oplus> udom \<rightarrow> udom \<oplus> udom"
+where
+  "ssum_approx = (\<lambda>i. ssum_map\<cdot>(udom_approx i)\<cdot>(udom_approx i))"
 
-instantiation ssum :: (bifinite, bifinite) bifinite
+lemma ssum_approx: "approx_chain ssum_approx"
+proof (rule approx_chain.intro)
+  show "chain (\<lambda>i. ssum_approx i)"
+    unfolding ssum_approx_def by simp
+  show "(\<Squnion>i. ssum_approx i) = ID"
+    unfolding ssum_approx_def
+    by (simp add: lub_distribs ssum_map_ID)
+  show "\<And>i. finite_deflation (ssum_approx i)"
+    unfolding ssum_approx_def
+    by (intro finite_deflation_ssum_map finite_deflation_udom_approx)
+qed
+
+definition ssum_sfp :: "sfp \<rightarrow> sfp \<rightarrow> sfp"
+where "ssum_sfp = sfp_fun2 ssum_approx ssum_map"
+
+lemma cast_ssum_sfp:
+  "cast\<cdot>(ssum_sfp\<cdot>A\<cdot>B) =
+    udom_emb ssum_approx oo ssum_map\<cdot>(cast\<cdot>A)\<cdot>(cast\<cdot>B) oo udom_prj ssum_approx"
+unfolding ssum_sfp_def
+apply (rule cast_sfp_fun2 [OF ssum_approx])
+apply (erule (1) finite_deflation_ssum_map)
+done
+
+instantiation ssum :: (sfp, sfp) sfp
 begin
 
 definition
-  approx_ssum_def:
-    "approx = (\<lambda>n. ssum_map\<cdot>(approx n)\<cdot>(approx n))"
+  "emb = udom_emb ssum_approx oo ssum_map\<cdot>emb\<cdot>emb"
 
-lemma approx_sinl [simp]: "approx i\<cdot>(sinl\<cdot>x) = sinl\<cdot>(approx i\<cdot>x)"
-unfolding approx_ssum_def by (cases "x = \<bottom>") simp_all
+definition
+  "prj = ssum_map\<cdot>prj\<cdot>prj oo udom_prj ssum_approx"
 
-lemma approx_sinr [simp]: "approx i\<cdot>(sinr\<cdot>x) = sinr\<cdot>(approx i\<cdot>x)"
-unfolding approx_ssum_def by (cases "x = \<bottom>") simp_all
+definition
+  "sfp (t::('a \<oplus> 'b) itself) = ssum_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
 
 instance proof
-  fix i :: nat and x :: "'a \<oplus> 'b"
-  show "chain (approx :: nat \<Rightarrow> 'a \<oplus> 'b \<rightarrow> 'a \<oplus> 'b)"
-    unfolding approx_ssum_def by simp
-  show "(\<Squnion>i. approx i\<cdot>x) = x"
-    unfolding approx_ssum_def
-    by (cases x, simp_all add: lub_distribs)
-  show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
-    by (cases x, simp add: approx_ssum_def, simp, simp)
-  show "finite {x::'a \<oplus> 'b. approx i\<cdot>x = x}"
-    unfolding approx_ssum_def
-    by (intro finite_deflation.finite_fixes
-              finite_deflation_ssum_map
-              finite_deflation_approx)
+  show "ep_pair emb (prj :: udom \<rightarrow> 'a \<oplus> 'b)"
+    unfolding emb_ssum_def prj_ssum_def
+    using ep_pair_udom [OF ssum_approx]
+    by (intro ep_pair_comp ep_pair_ssum_map ep_pair_emb_prj)
+next
+  show "cast\<cdot>SFP('a \<oplus> 'b) = emb oo (prj :: udom \<rightarrow> 'a \<oplus> 'b)"
+    unfolding emb_ssum_def prj_ssum_def sfp_ssum_def cast_ssum_sfp
+    by (simp add: cast_SFP oo_def expand_cfun_eq ssum_map_map)
 qed
 
 end
 
+text {* SFP of type constructor = type combinator *}
+
+lemma SFP_ssum: "SFP('a::sfp \<oplus> 'b::sfp) = ssum_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
+by (rule sfp_ssum_def)
+
 end
--- a/src/HOLCF/Tools/Domain/domain_extender.ML	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Tools/Domain/domain_extender.ML	Wed Oct 06 10:49:27 2010 -0700
@@ -144,7 +144,7 @@
 
     fun thy_type  (dname,tvars,mx) = (dname, length tvars, mx);
     fun thy_arity (dname,tvars,mx) =
-      (Sign.full_name thy dname, map (snd o dest_TFree) tvars, @{sort rep});
+      (Sign.full_name thy dname, map (snd o dest_TFree) tvars, arg_sort false);
 
     (* this theory is used just for parsing and error checking *)
     val tmp_thy = thy
@@ -213,7 +213,7 @@
   end;
 
 fun pcpo_arg lazy = if lazy then @{sort cpo} else @{sort pcpo};
-fun rep_arg lazy = @{sort rep};
+fun rep_arg lazy = @{sort sfp};
 
 val add_domain =
     gen_add_domain Sign.certify_typ Domain_Axioms.add_axioms pcpo_arg;
--- a/src/HOLCF/Tools/Domain/domain_isomorphism.ML	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Tools/Domain/domain_isomorphism.ML	Wed Oct 06 10:49:27 2010 -0700
@@ -48,7 +48,7 @@
 
 structure DeflData = Theory_Data
 (
-  (* terms like "foo_defl" *)
+  (* terms like "foo_sfp" *)
   type T = term Symtab.table;
   val empty = Symtab.empty;
   val extend = I;
@@ -57,7 +57,7 @@
 
 structure RepData = Theory_Data
 (
-  (* theorems like "REP('a foo) = foo_defl$REP('a)" *)
+  (* theorems like "SFP('a foo) = foo_sfp$SFP('a)" *)
   type T = thm list;
   val empty = [];
   val extend = I;
@@ -83,11 +83,11 @@
 );
 
 fun add_type_constructor
-  (tname, defl_const, map_name, REP_thm,
+  (tname, defl_const, map_name, SFP_thm,
    isodefl_thm, map_ID_thm, defl_map_thm) =
     DeflData.map (Symtab.insert (K true) (tname, defl_const))
     #> Domain_Take_Proofs.add_map_function (tname, map_name, defl_map_thm)
-    #> RepData.map (Thm.add_thm REP_thm)
+    #> RepData.map (Thm.add_thm SFP_thm)
     #> IsodeflData.map (Thm.add_thm isodefl_thm)
     #> MapIdData.map (Thm.add_thm map_ID_thm);
 
@@ -104,19 +104,19 @@
 infixr 6 ->>;
 infix -->>;
 
-val deflT = @{typ "udom alg_defl"};
+val sfpT = @{typ "sfp"};
 
 fun mapT (T as Type (_, Ts)) =
     (map (fn T => T ->> T) Ts) -->> (T ->> T)
   | mapT T = T ->> T;
 
-fun mk_Rep_of T =
-  Const (@{const_name Rep_of}, Term.itselfT T --> deflT) $ Logic.mk_type T;
+fun mk_SFP T =
+  Const (@{const_name sfp}, Term.itselfT T --> sfpT) $ Logic.mk_type T;
 
 fun coerce_const T = Const (@{const_name coerce}, T);
 
 fun isodefl_const T =
-  Const (@{const_name isodefl}, (T ->> T) --> deflT --> HOLogic.boolT);
+  Const (@{const_name isodefl}, (T ->> T) --> sfpT --> HOLogic.boolT);
 
 fun mk_deflation t =
   Const (@{const_name deflation}, Term.fastype_of t --> boolT) $ t;
@@ -218,13 +218,13 @@
   let
     fun is_closed_typ (Type (_, Ts)) = forall is_closed_typ Ts
       | is_closed_typ _ = false;
-    fun defl_of (TFree (a, _)) = Free (Library.unprefix "'" a, deflT)
+    fun defl_of (TFree (a, _)) = Free (Library.unprefix "'" a, sfpT)
       | defl_of (TVar _) = error ("defl_of_typ: TVar")
       | defl_of (T as Type (c, Ts)) =
         case Symtab.lookup tab c of
           SOME t => list_ccomb (t, map defl_of Ts)
         | NONE => if is_closed_typ T
-                  then mk_Rep_of T
+                  then mk_SFP T
                   else error ("defl_of_typ: type variable under unsupported type constructor " ^ c);
   in defl_of T end;
 
@@ -443,7 +443,7 @@
     (* declare deflation combinator constants *)
     fun declare_defl_const (vs, tbind, mx, rhs, morphs) thy =
       let
-        val defl_type = map (K deflT) vs -->> deflT;
+        val defl_type = map (K sfpT) vs -->> sfpT;
         val defl_bind = Binding.suffix_name "_defl" tbind;
       in
         Sign.declare_const ((defl_bind, defl_type), NoSyn) thy
@@ -470,34 +470,34 @@
     fun make_repdef ((vs, tbind, mx, _, _), defl_const) thy =
       let
         fun tfree a = TFree (a, the (AList.lookup (op =) sorts a))
-        val reps = map (mk_Rep_of o tfree) vs;
+        val reps = map (mk_SFP o tfree) vs;
         val defl = list_ccomb (defl_const, reps);
-        val ((_, _, _, {REP, ...}), thy) =
+        val ((_, _, _, {SFP, ...}), thy) =
           Repdef.add_repdef false NONE (tbind, map (rpair dummyS) vs, mx) defl NONE thy;
       in
-        (REP, thy)
+        (SFP, thy)
       end;
-    val (REP_thms, thy) = fold_map make_repdef (doms ~~ defl_consts) thy;
-    val thy = RepData.map (fold Thm.add_thm REP_thms) thy;
+    val (SFP_thms, thy) = fold_map make_repdef (doms ~~ defl_consts) thy;
+    val thy = RepData.map (fold Thm.add_thm SFP_thms) thy;
 
-    (* prove REP equations *)
-    fun mk_REP_eq_thm (lhsT, rhsT) =
+    (* prove SFP equations *)
+    fun mk_SFP_eq_thm (lhsT, rhsT) =
       let
-        val goal = mk_eqs (mk_Rep_of lhsT, mk_Rep_of rhsT);
-        val REP_simps = RepData.get thy;
+        val goal = mk_eqs (mk_SFP lhsT, mk_SFP rhsT);
+        val SFP_simps = RepData.get thy;
         val tac =
-          rewrite_goals_tac (map mk_meta_eq REP_simps)
+          rewrite_goals_tac (map mk_meta_eq SFP_simps)
           THEN resolve_tac defl_unfold_thms 1;
       in
         Goal.prove_global thy [] [] goal (K tac)
       end;
-    val REP_eq_thms = map mk_REP_eq_thm dom_eqns;
+    val SFP_eq_thms = map mk_SFP_eq_thm dom_eqns;
 
-    (* register REP equations *)
-    val REP_eq_binds = map (Binding.prefix_name "REP_eq_") dbinds;
+    (* register SFP equations *)
+    val SFP_eq_binds = map (Binding.prefix_name "SFP_eq_") dbinds;
     val (_, thy) = thy |>
       (Global_Theory.add_thms o map Thm.no_attributes)
-        (REP_eq_binds ~~ REP_eq_thms);
+        (SFP_eq_binds ~~ SFP_eq_thms);
 
     (* define rep/abs functions *)
     fun mk_rep_abs ((tbind, morphs), (lhsT, rhsT)) thy =
@@ -516,10 +516,10 @@
       |>> ListPair.unzip;
 
     (* prove isomorphism and isodefl rules *)
-    fun mk_iso_thms ((tbind, REP_eq), (rep_def, abs_def)) thy =
+    fun mk_iso_thms ((tbind, SFP_eq), (rep_def, abs_def)) thy =
       let
         fun make thm =
-            Drule.zero_var_indexes (thm OF [REP_eq, abs_def, rep_def]);
+            Drule.zero_var_indexes (thm OF [SFP_eq, abs_def, rep_def]);
         val rep_iso_thm = make @{thm domain_rep_iso};
         val abs_iso_thm = make @{thm domain_abs_iso};
         val isodefl_thm = make @{thm isodefl_abs_rep};
@@ -532,7 +532,7 @@
       end;
     val ((iso_thms, isodefl_abs_rep_thms), thy) =
       thy
-      |> fold_map mk_iso_thms (dbinds ~~ REP_eq_thms ~~ rep_abs_defs)
+      |> fold_map mk_iso_thms (dbinds ~~ SFP_eq_thms ~~ rep_abs_defs)
       |>> ListPair.unzip;
 
     (* collect info about rep/abs *)
@@ -561,7 +561,7 @@
     val isodefl_thm =
       let
         fun unprime a = Library.unprefix "'" a;
-        fun mk_d T = Free ("d" ^ unprime (fst (dest_TFree T)), deflT);
+        fun mk_d T = Free ("d" ^ unprime (fst (dest_TFree T)), sfpT);
         fun mk_f T = Free ("f" ^ unprime (fst (dest_TFree T)), T ->> T);
         fun mk_assm T = mk_trp (isodefl_const T $ mk_f T $ mk_d T);
         fun mk_goal ((map_const, defl_const), (T, rhsT)) =
@@ -579,9 +579,9 @@
           @{thms adm_conj adm_isodefl cont2cont_fst cont2cont_snd cont_id};
         val bottom_rules =
           @{thms fst_strict snd_strict isodefl_bottom simp_thms};
-        val REP_simps = map (fn th => th RS sym) (RepData.get thy);
+        val SFP_simps = map (fn th => th RS sym) (RepData.get thy);
         val isodefl_rules =
-          @{thms conjI isodefl_ID_REP}
+          @{thms conjI isodefl_ID_SFP}
           @ isodefl_abs_rep_thms
           @ IsodeflData.get thy;
       in
@@ -595,7 +595,7 @@
            simp_tac (HOL_basic_ss addsimps bottom_rules) 1,
            simp_tac beta_ss 1,
            simp_tac (HOL_basic_ss addsimps @{thms fst_conv snd_conv}) 1,
-           simp_tac (HOL_basic_ss addsimps REP_simps) 1,
+           simp_tac (HOL_basic_ss addsimps SFP_simps) 1,
            REPEAT (etac @{thm conjE} 1),
            REPEAT (resolve_tac (isodefl_rules @ prems) 1 ORELSE atac 1)])
       end;
@@ -613,23 +613,23 @@
 
     (* prove map_ID theorems *)
     fun prove_map_ID_thm
-        (((map_const, (lhsT, _)), REP_thm), isodefl_thm) =
+        (((map_const, (lhsT, _)), SFP_thm), isodefl_thm) =
       let
         val Ts = snd (dest_Type lhsT);
         val lhs = list_ccomb (map_const, map mk_ID Ts);
         val goal = mk_eqs (lhs, mk_ID lhsT);
         val tac = EVERY
-          [rtac @{thm isodefl_REP_imp_ID} 1,
-           stac REP_thm 1,
+          [rtac @{thm isodefl_SFP_imp_ID} 1,
+           stac SFP_thm 1,
            rtac isodefl_thm 1,
-           REPEAT (rtac @{thm isodefl_ID_REP} 1)];
+           REPEAT (rtac @{thm isodefl_ID_SFP} 1)];
       in
         Goal.prove_global thy [] [] goal (K tac)
       end;
     val map_ID_binds = map (Binding.suffix_name "_map_ID") dbinds;
     val map_ID_thms =
       map prove_map_ID_thm
-        (map_consts ~~ dom_eqns ~~ REP_thms ~~ isodefl_thms);
+        (map_consts ~~ dom_eqns ~~ SFP_thms ~~ isodefl_thms);
     val (_, thy) = thy |>
       (Global_Theory.add_thms o map Thm.no_attributes)
         (map_ID_binds ~~ map_ID_thms);
--- a/src/HOLCF/Tools/repdef.ML	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Tools/repdef.ML	Wed Oct 06 10:49:27 2010 -0700
@@ -7,7 +7,7 @@
 signature REPDEF =
 sig
   type rep_info =
-    { emb_def: thm, prj_def: thm, approx_def: thm, REP: thm }
+    { emb_def: thm, prj_def: thm, sfp_def: thm, SFP: thm }
 
   val add_repdef: bool -> binding option -> binding * (string * sort) list * mixfix ->
     term -> (binding * binding) option -> theory ->
@@ -28,20 +28,19 @@
 (** type definitions **)
 
 type rep_info =
-  { emb_def: thm, prj_def: thm, approx_def: thm, REP: thm };
+  { emb_def: thm, prj_def: thm, sfp_def: thm, SFP: thm };
 
 (* building types and terms *)
 
 val udomT = @{typ udom};
-fun alg_deflT T = Type (@{type_name alg_defl}, [T]);
+val sfpT = @{typ sfp};
 fun emb_const T = Const (@{const_name emb}, T ->> udomT);
 fun prj_const T = Const (@{const_name prj}, udomT ->> T);
-fun approx_const T = Const (@{const_name approx}, natT --> (T ->> T));
+fun sfp_const T = Const (@{const_name sfp}, Term.itselfT T --> sfpT);
 
-fun cast_const T = Const (@{const_name cast}, alg_deflT T ->> T ->> T);
 fun mk_cast (t, x) =
   capply_const (udomT, udomT)
-  $ (capply_const (alg_deflT udomT, udomT ->> udomT) $ cast_const udomT $ t)
+  $ (capply_const (sfpT, udomT ->> udomT) $ @{const cast} $ t)
   $ x;
 
 (* manipulating theorems *)
@@ -71,8 +70,8 @@
     val tmp_ctxt = tmp_ctxt |> Variable.declare_constraints defl;
 
     val deflT = Term.fastype_of defl;
-    val _ = if deflT = @{typ "udom alg_defl"} then ()
-            else error ("Not type udom alg_defl: " ^ quote (Syntax.string_of_typ tmp_ctxt deflT));
+    val _ = if deflT = @{typ "sfp"} then ()
+            else error ("Not type sfp: " ^ quote (Syntax.string_of_typ tmp_ctxt deflT));
 
     (*lhs*)
     val lhs_tfrees = map (ProofContext.check_tfree tmp_ctxt) raw_args;
@@ -85,12 +84,12 @@
       |> the_default (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name);
 
     (*set*)
-    val in_defl = @{term "in_deflation :: udom => udom alg_defl => bool"};
+    val in_defl = @{term "in_sfp :: udom => sfp => bool"};
     val set = HOLogic.Collect_const udomT $ Abs ("x", udomT, in_defl $ Bound 0 $ defl);
 
     (*pcpodef*)
-    val tac1 = rtac @{thm CollectI} 1 THEN rtac @{thm bottom_in_deflation} 1;
-    val tac2 = rtac @{thm adm_mem_Collect_in_deflation} 1;
+    val tac1 = rtac @{thm CollectI} 1 THEN rtac @{thm bottom_in_sfp} 1;
+    val tac2 = rtac @{thm adm_mem_Collect_in_sfp} 1;
     val ((info, cpo_info, pcpo_info), thy) = thy
       |> Pcpodef.add_pcpodef def (SOME name) typ set (SOME morphs) (tac1, tac2);
 
@@ -100,52 +99,48 @@
     val emb_eqn = Logic.mk_equals (emb_const newT, cabs_const (newT, udomT) $ Rep_const);
     val prj_eqn = Logic.mk_equals (prj_const newT, cabs_const (udomT, newT) $
       Abs ("x", udomT, Abs_const $ mk_cast (defl, Bound 0)));
-    val repdef_approx_const =
-      Const (@{const_name repdef_approx}, (newT --> udomT) --> (udomT --> newT)
-        --> alg_deflT udomT --> natT --> (newT ->> newT));
-    val approx_eqn = Logic.mk_equals (approx_const newT,
-      repdef_approx_const $ Rep_const $ Abs_const $ defl);
+    val sfp_eqn = Logic.mk_equals (sfp_const newT,
+      Abs ("x", Term.itselfT newT, defl));
     val name_def = Binding.suffix_name "_def" name;
     val emb_bind = (Binding.prefix_name "emb_" name_def, []);
     val prj_bind = (Binding.prefix_name "prj_" name_def, []);
-    val approx_bind = (Binding.prefix_name "approx_" name_def, []);
+    val sfp_bind = (Binding.prefix_name "sfp_" name_def, []);
 
     (*instantiate class rep*)
     val lthy = thy
-      |> Class.instantiation ([full_tname], lhs_tfrees, @{sort rep});
+      |> Class.instantiation ([full_tname], lhs_tfrees, @{sort sfp});
     val ((_, (_, emb_ldef)), lthy) =
         Specification.definition (NONE, (emb_bind, emb_eqn)) lthy;
     val ((_, (_, prj_ldef)), lthy) =
         Specification.definition (NONE, (prj_bind, prj_eqn)) lthy;
-    val ((_, (_, approx_ldef)), lthy) =
-        Specification.definition (NONE, (approx_bind, approx_eqn)) lthy;
+    val ((_, (_, sfp_ldef)), lthy) =
+        Specification.definition (NONE, (sfp_bind, sfp_eqn)) lthy;
     val ctxt_thy = ProofContext.init_global (ProofContext.theory_of lthy);
     val emb_def = singleton (ProofContext.export lthy ctxt_thy) emb_ldef;
     val prj_def = singleton (ProofContext.export lthy ctxt_thy) prj_ldef;
-    val approx_def = singleton (ProofContext.export lthy ctxt_thy) approx_ldef;
+    val sfp_def = singleton (ProofContext.export lthy ctxt_thy) sfp_ldef;
     val type_definition_thm =
       MetaSimplifier.rewrite_rule
         (the_list (#set_def (#2 info)))
         (#type_definition (#2 info));
     val typedef_thms =
-      [type_definition_thm, #below_def cpo_info, emb_def, prj_def, approx_def];
+      [type_definition_thm, #below_def cpo_info, emb_def, prj_def, sfp_def];
     val thy = lthy
       |> Class.prove_instantiation_instance
           (K (Tactic.rtac (@{thm typedef_rep_class} OF typedef_thms) 1))
       |> Local_Theory.exit_global;
 
     (*other theorems*)
-    val typedef_thms' = map (Thm.transfer thy)
-      [type_definition_thm, #below_def cpo_info, emb_def, prj_def];
-    val (REP_thm, thy) = thy
+    val sfp_thm' = Thm.transfer thy sfp_def;
+    val (SFP_thm, thy) = thy
       |> Sign.add_path (Binding.name_of name)
       |> Global_Theory.add_thm
-         ((Binding.prefix_name "REP_" name,
-          Drule.zero_var_indexes (@{thm typedef_REP} OF typedef_thms')), [])
+         ((Binding.prefix_name "SFP_" name,
+          Drule.zero_var_indexes (@{thm typedef_SFP} OF [sfp_thm'])), [])
       ||> Sign.restore_naming thy;
 
     val rep_info =
-      { emb_def = emb_def, prj_def = prj_def, approx_def = approx_def, REP = REP_thm };
+      { emb_def = emb_def, prj_def = prj_def, sfp_def = sfp_def, SFP = SFP_thm };
   in
     ((info, cpo_info, pcpo_info, rep_info), thy)
   end
--- a/src/HOLCF/Tutorial/New_Domain.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Tutorial/New_Domain.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -9,14 +9,14 @@
 begin
 
 text {*
-  The definitional domain package only works with representable domains,
-  i.e. types in class @{text rep}.
+  The definitional domain package only works with SFP domains,
+  i.e. types in class @{text sfp}.
 *}
 
-default_sort rep
+default_sort sfp
 
 text {*
-  Provided that @{text rep} is the default sort, the @{text new_domain}
+  Provided that @{text sfp} is the default sort, the @{text new_domain}
   package should work with any type definition supported by the old
   domain package.
 *}
--- a/src/HOLCF/Universal.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Universal.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -5,10 +5,12 @@
 header {* A universal bifinite domain *}
 
 theory Universal
-imports CompactBasis Nat_Bijection
+imports Completion Deflation Nat_Bijection
 begin
 
-subsection {* Basis datatype *}
+subsection {* Basis for universal domain *}
+
+subsubsection {* Basis datatype *}
 
 types ubasis = nat
 
@@ -75,7 +77,7 @@
  apply (simp add: 2 node_gt1 node_gt2)
 done
 
-subsection {* Basis ordering *}
+subsubsection {* Basis ordering *}
 
 inductive
   ubasis_le :: "nat \<Rightarrow> nat \<Rightarrow> bool"
@@ -95,6 +97,12 @@
 apply (erule ubasis_le_lower)
 done
 
+interpretation udom: preorder ubasis_le
+apply default
+apply (rule ubasis_le_refl)
+apply (erule (1) ubasis_le_trans)
+done
+
 subsubsection {* Generic take function *}
 
 function
@@ -187,66 +195,6 @@
 apply simp
 done
 
-subsubsection {* Take function for \emph{ubasis} *}
-
-definition
-  ubasis_take :: "nat \<Rightarrow> ubasis \<Rightarrow> ubasis"
-where
-  "ubasis_take n = ubasis_until (\<lambda>x. x \<le> n)"
-
-lemma ubasis_take_le: "ubasis_take n x \<le> n"
-unfolding ubasis_take_def by (rule ubasis_until, rule le0)
-
-lemma ubasis_take_same: "x \<le> n \<Longrightarrow> ubasis_take n x = x"
-unfolding ubasis_take_def by (rule ubasis_until_same)
-
-lemma ubasis_take_idem: "ubasis_take n (ubasis_take n x) = ubasis_take n x"
-by (rule ubasis_take_same [OF ubasis_take_le])
-
-lemma ubasis_take_0 [simp]: "ubasis_take 0 x = 0"
-unfolding ubasis_take_def by (simp add: ubasis_until_0)
-
-lemma ubasis_take_less: "ubasis_le (ubasis_take n x) x"
-unfolding ubasis_take_def by (rule ubasis_until_less)
-
-lemma ubasis_take_chain: "ubasis_le (ubasis_take n x) (ubasis_take (Suc n) x)"
-unfolding ubasis_take_def by (rule ubasis_until_chain) simp
-
-lemma ubasis_take_mono:
-  assumes "ubasis_le x y"
-  shows "ubasis_le (ubasis_take n x) (ubasis_take n y)"
-unfolding ubasis_take_def
- apply (rule ubasis_until_mono [OF _ prems])
- apply (frule (2) order_less_le_trans [OF node_gt2])
- apply (erule order_less_imp_le)
-done
-
-lemma finite_range_ubasis_take: "finite (range (ubasis_take n))"
-apply (rule finite_subset [where B="{..n}"])
-apply (simp add: subset_eq ubasis_take_le)
-apply simp
-done
-
-lemma ubasis_take_covers: "\<exists>n. ubasis_take n x = x"
-apply (rule exI [where x=x])
-apply (simp add: ubasis_take_same)
-done
-
-interpretation udom: preorder ubasis_le
-apply default
-apply (rule ubasis_le_refl)
-apply (erule (1) ubasis_le_trans)
-done
-
-interpretation udom: basis_take ubasis_le ubasis_take
-apply default
-apply (rule ubasis_take_less)
-apply (rule ubasis_take_idem)
-apply (erule ubasis_take_mono)
-apply (rule ubasis_take_chain)
-apply (rule finite_range_ubasis_take)
-apply (rule ubasis_take_covers)
-done
 
 subsection {* Defining the universal domain by ideal completion *}
 
@@ -263,17 +211,17 @@
 end
 
 instance udom :: po
-by (rule udom.typedef_ideal_po
-    [OF type_definition_udom below_udom_def])
+using type_definition_udom below_udom_def
+by (rule udom.typedef_ideal_po)
 
 instance udom :: cpo
-by (rule udom.typedef_ideal_cpo
-    [OF type_definition_udom below_udom_def])
+using type_definition_udom below_udom_def
+by (rule udom.typedef_ideal_cpo)
 
 lemma Rep_udom_lub:
   "chain Y \<Longrightarrow> Rep_udom (\<Squnion>i. Y i) = (\<Union>i. Rep_udom (Y i))"
-by (rule udom.typedef_ideal_rep_contlub
-    [OF type_definition_udom below_udom_def])
+using type_definition_udom below_udom_def
+by (rule udom.typedef_ideal_rep_contlub)
 
 lemma ideal_Rep_udom: "udom.ideal (Rep_udom xs)"
 by (rule Rep_udom [unfolded mem_Collect_eq])
@@ -288,12 +236,13 @@
 by (simp add: Abs_udom_inverse udom.ideal_principal)
 
 interpretation udom:
-  ideal_completion ubasis_le ubasis_take udom_principal Rep_udom
+  ideal_completion ubasis_le udom_principal Rep_udom
 apply unfold_locales
 apply (rule ideal_Rep_udom)
 apply (erule Rep_udom_lub)
 apply (rule Rep_udom_principal)
 apply (simp only: below_udom_def)
+apply (rule exI, rule inj_on_id)
 done
 
 text {* Universal domain is pointed *}
@@ -309,43 +258,60 @@
 lemma inst_udom_pcpo: "\<bottom> = udom_principal 0"
 by (rule udom_minimal [THEN UU_I, symmetric])
 
-text {* Universal domain is bifinite *}
+
+subsection {* Compact bases of domains *}
 
-instantiation udom :: bifinite
+typedef (open) 'a compact_basis = "{x::'a::pcpo. compact x}"
+by auto
+
+lemma compact_Rep_compact_basis: "compact (Rep_compact_basis a)"
+by (rule Rep_compact_basis [unfolded mem_Collect_eq])
+
+instantiation compact_basis :: (pcpo) below
 begin
 
 definition
-  approx_udom_def: "approx = udom.completion_approx"
+  compact_le_def:
+    "(op \<sqsubseteq>) \<equiv> (\<lambda>x y. Rep_compact_basis x \<sqsubseteq> Rep_compact_basis y)"
 
-instance
-apply (intro_classes, unfold approx_udom_def)
-apply (rule udom.chain_completion_approx)
-apply (rule udom.lub_completion_approx)
-apply (rule udom.completion_approx_idem)
-apply (rule udom.finite_fixes_completion_approx)
-done
-
+instance ..
 end
 
-lemma approx_udom_principal [simp]:
-  "approx n\<cdot>(udom_principal x) = udom_principal (ubasis_take n x)"
-unfolding approx_udom_def
-by (rule udom.completion_approx_principal)
+instance compact_basis :: (pcpo) po
+using type_definition_compact_basis compact_le_def
+by (rule typedef_po)
+
+definition
+  approximants :: "'a \<Rightarrow> 'a compact_basis set" where
+  "approximants = (\<lambda>x. {a. Rep_compact_basis a \<sqsubseteq> x})"
 
-lemma approx_eq_udom_principal:
-  "\<exists>a\<in>Rep_udom x. approx n\<cdot>x = udom_principal (ubasis_take n a)"
-unfolding approx_udom_def
-by (rule udom.completion_approx_eq_principal)
+definition
+  compact_bot :: "'a::pcpo compact_basis" where
+  "compact_bot = Abs_compact_basis \<bottom>"
+
+lemma Rep_compact_bot [simp]: "Rep_compact_basis compact_bot = \<bottom>"
+unfolding compact_bot_def by (simp add: Abs_compact_basis_inverse)
+
+lemma compact_bot_minimal [simp]: "compact_bot \<sqsubseteq> a"
+unfolding compact_le_def Rep_compact_bot by simp
 
 
 subsection {* Universality of \emph{udom} *}
 
-default_sort bifinite
+text {* We use a locale to parameterize the construction over a chain
+of approx functions on the type to be embedded. *}
+
+locale approx_chain =
+  fixes approx :: "nat \<Rightarrow> 'a::pcpo \<rightarrow> 'a"
+  assumes chain_approx [simp]: "chain (\<lambda>i. approx i)"
+  assumes lub_approx [simp]: "(\<Squnion>i. approx i) = ID"
+  assumes finite_deflation_approx: "\<And>i. finite_deflation (approx i)"
+begin
 
 subsubsection {* Choosing a maximal element from a finite set *}
 
 lemma finite_has_maximal:
-  fixes A :: "'a::po set"
+  fixes A :: "'a compact_basis set"
   shows "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> \<exists>x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y"
 proof (induct rule: finite_ne_induct)
   case (singleton x)
@@ -456,43 +422,86 @@
  apply (simp add: choose_pos.simps)
 done
 
-subsubsection {* Rank of basis elements *}
+subsubsection {* Properties of approx function *}
+
+lemma deflation_approx: "deflation (approx i)"
+using finite_deflation_approx by (rule finite_deflation_imp_deflation)
+
+lemma approx_idem: "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
+using deflation_approx by (rule deflation.idem)
+
+lemma approx_below: "approx i\<cdot>x \<sqsubseteq> x"
+using deflation_approx by (rule deflation.below)
+
+lemma finite_range_approx: "finite (range (\<lambda>x. approx i\<cdot>x))"
+apply (rule finite_deflation.finite_range)
+apply (rule finite_deflation_approx)
+done
+
+lemma compact_approx: "compact (approx n\<cdot>x)"
+apply (rule finite_deflation.compact)
+apply (rule finite_deflation_approx)
+done
+
+lemma compact_eq_approx: "compact x \<Longrightarrow> \<exists>i. approx i\<cdot>x = x"
+by (rule admD2, simp_all)
+
+subsubsection {* Compact basis take function *}
 
 primrec
-  cb_take :: "nat \<Rightarrow> 'a compact_basis \<Rightarrow> 'a compact_basis"
-where
+  cb_take :: "nat \<Rightarrow> 'a compact_basis \<Rightarrow> 'a compact_basis" where
   "cb_take 0 = (\<lambda>x. compact_bot)"
-| "cb_take (Suc n) = compact_take n"
+| "cb_take (Suc n) = (\<lambda>a. Abs_compact_basis (approx n\<cdot>(Rep_compact_basis a)))"
+
+declare cb_take.simps [simp del]
+
+lemma cb_take_zero [simp]: "cb_take 0 a = compact_bot"
+by (simp only: cb_take.simps)
+
+lemma Rep_cb_take:
+  "Rep_compact_basis (cb_take (Suc n) a) = approx n\<cdot>(Rep_compact_basis a)"
+by (simp add: Abs_compact_basis_inverse cb_take.simps(2) compact_approx)
+
+lemmas approx_Rep_compact_basis = Rep_cb_take [symmetric]
 
 lemma cb_take_covers: "\<exists>n. cb_take n x = x"
-apply (rule exE [OF compact_basis.take_covers [where a=x]])
-apply (rename_tac n, rule_tac x="Suc n" in exI, simp)
+apply (subgoal_tac "\<exists>n. cb_take (Suc n) x = x", fast)
+apply (simp add: Rep_compact_basis_inject [symmetric])
+apply (simp add: Rep_cb_take)
+apply (rule compact_eq_approx)
+apply (rule compact_Rep_compact_basis)
 done
 
 lemma cb_take_less: "cb_take n x \<sqsubseteq> x"
-by (cases n, simp, simp add: compact_basis.take_less)
+unfolding compact_le_def
+by (cases n, simp, simp add: Rep_cb_take approx_below)
 
 lemma cb_take_idem: "cb_take n (cb_take n x) = cb_take n x"
-by (cases n, simp, simp add: compact_basis.take_take)
+unfolding Rep_compact_basis_inject [symmetric]
+by (cases n, simp, simp add: Rep_cb_take approx_idem)
 
 lemma cb_take_mono: "x \<sqsubseteq> y \<Longrightarrow> cb_take n x \<sqsubseteq> cb_take n y"
-by (cases n, simp, simp add: compact_basis.take_mono)
+unfolding compact_le_def
+by (cases n, simp, simp add: Rep_cb_take monofun_cfun_arg)
 
 lemma cb_take_chain_le: "m \<le> n \<Longrightarrow> cb_take m x \<sqsubseteq> cb_take n x"
-apply (cases m, simp)
-apply (cases n, simp)
-apply (simp add: compact_basis.take_chain_le)
+unfolding compact_le_def
+apply (cases m, simp, cases n, simp)
+apply (simp add: Rep_cb_take, rule chain_mono, simp, simp)
 done
 
-lemma range_const: "range (\<lambda>x. c) = {c}"
-by auto
-
 lemma finite_range_cb_take: "finite (range (cb_take n))"
 apply (cases n)
-apply (simp add: range_const)
-apply (simp add: compact_basis.finite_range_take)
+apply (subgoal_tac "range (cb_take 0) = {compact_bot}", simp, force)
+apply (rule finite_imageD [where f="Rep_compact_basis"])
+apply (rule finite_subset [where B="range (\<lambda>x. approx (n - 1)\<cdot>x)"])
+apply (clarsimp simp add: Rep_cb_take)
+apply (rule finite_range_approx)
+apply (rule inj_onI, simp add: Rep_compact_basis_inject)
 done
 
+subsubsection {* Rank of basis elements *}
+
 definition
   rank :: "'a compact_basis \<Rightarrow> nat"
 where
@@ -809,22 +818,86 @@
  apply (rule ubasis_until_less)
 done
 
-hide_const (open)
-  node
-  choose
-  choose_pos
-  place
-  sub
+end
+
+sublocale approx_chain \<subseteq> compact_basis!:
+  ideal_completion below Rep_compact_basis
+    "approximants :: 'a \<Rightarrow> 'a compact_basis set"
+proof
+  fix w :: "'a"
+  show "below.ideal (approximants w)"
+  proof (rule below.idealI)
+    show "\<exists>x. x \<in> approximants w"
+      unfolding approximants_def
+      apply (rule_tac x="Abs_compact_basis (approx 0\<cdot>w)" in exI)
+      apply (simp add: Abs_compact_basis_inverse approx_below compact_approx)
+      done
+  next
+    fix x y :: "'a compact_basis"
+    assume "x \<in> approximants w" "y \<in> approximants w"
+    thus "\<exists>z \<in> approximants w. x \<sqsubseteq> z \<and> y \<sqsubseteq> z"
+      unfolding approximants_def
+      apply simp
+      apply (cut_tac a=x in compact_Rep_compact_basis)
+      apply (cut_tac a=y in compact_Rep_compact_basis)
+      apply (drule compact_eq_approx)
+      apply (drule compact_eq_approx)
+      apply (clarify, rename_tac i j)
+      apply (rule_tac x="Abs_compact_basis (approx (max i j)\<cdot>w)" in exI)
+      apply (simp add: compact_le_def)
+      apply (simp add: Abs_compact_basis_inverse approx_below compact_approx)
+      apply (erule subst, erule subst)
+      apply (simp add: monofun_cfun chain_mono [OF chain_approx])
+      done
+  next
+    fix x y :: "'a compact_basis"
+    assume "x \<sqsubseteq> y" "y \<in> approximants w" thus "x \<in> approximants w"
+      unfolding approximants_def
+      apply simp
+      apply (simp add: compact_le_def)
+      apply (erule (1) below_trans)
+      done
+  qed
+next
+  fix Y :: "nat \<Rightarrow> 'a"
+  assume Y: "chain Y"
+  show "approximants (\<Squnion>i. Y i) = (\<Union>i. approximants (Y i))"
+    unfolding approximants_def
+    apply safe
+    apply (simp add: compactD2 [OF compact_Rep_compact_basis Y])
+    apply (erule below_trans, rule is_ub_thelub [OF Y])
+    done
+next
+  fix a :: "'a compact_basis"
+  show "approximants (Rep_compact_basis a) = {b. b \<sqsubseteq> a}"
+    unfolding approximants_def compact_le_def ..
+next
+  fix x y :: "'a"
+  assume "approximants x \<subseteq> approximants y" thus "x \<sqsubseteq> y"
+    apply (subgoal_tac "(\<Squnion>i. approx i\<cdot>x) \<sqsubseteq> y")
+    apply (simp add: lub_distribs)
+    apply (rule admD, simp, simp)
+    apply (drule_tac c="Abs_compact_basis (approx i\<cdot>x)" in subsetD)
+    apply (simp add: approximants_def Abs_compact_basis_inverse
+                     approx_below compact_approx)
+    apply (simp add: approximants_def Abs_compact_basis_inverse compact_approx)
+    done
+next
+  show "\<exists>f::'a compact_basis \<Rightarrow> nat. inj f"
+    by (rule exI, rule inj_place)
+qed
 
 subsubsection {* EP-pair from any bifinite domain into \emph{udom} *}
 
+context approx_chain begin
+
 definition
-  udom_emb :: "'a::bifinite \<rightarrow> udom"
+  udom_emb :: "'a \<rightarrow> udom"
 where
   "udom_emb = compact_basis.basis_fun (\<lambda>x. udom_principal (basis_emb x))"
 
 definition
-  udom_prj :: "udom \<rightarrow> 'a::bifinite"
+  udom_prj :: "udom \<rightarrow> 'a"
 where
   "udom_prj = udom.basis_fun (\<lambda>x. Rep_compact_basis (basis_prj x))"
 
@@ -855,3 +928,103 @@
 done
 
 end
+
+abbreviation "udom_emb \<equiv> approx_chain.udom_emb"
+abbreviation "udom_prj \<equiv> approx_chain.udom_prj"
+
+lemmas ep_pair_udom = approx_chain.ep_pair_udom
+
+subsection {* Chain of approx functions for type \emph{udom} *}
+
+definition
+  udom_approx :: "nat \<Rightarrow> udom \<rightarrow> udom"
+where
+  "udom_approx i =
+    udom.basis_fun (\<lambda>x. udom_principal (ubasis_until (\<lambda>y. y \<le> i) x))"
+
+lemma udom_approx_mono:
+  "ubasis_le a b \<Longrightarrow>
+    udom_principal (ubasis_until (\<lambda>y. y \<le> i) a) \<sqsubseteq>
+    udom_principal (ubasis_until (\<lambda>y. y \<le> i) b)"
+apply (rule udom.principal_mono)
+apply (rule ubasis_until_mono)
+apply (frule (2) order_less_le_trans [OF node_gt2])
+apply (erule order_less_imp_le)
+apply assumption
+done
+
+lemma adm_mem_finite: "\<lbrakk>cont f; finite S\<rbrakk> \<Longrightarrow> adm (\<lambda>x. f x \<in> S)"
+by (erule adm_subst, induct set: finite, simp_all)
+
+lemma udom_approx_principal:
+  "udom_approx i\<cdot>(udom_principal x) =
+    udom_principal (ubasis_until (\<lambda>y. y \<le> i) x)"
+unfolding udom_approx_def
+apply (rule udom.basis_fun_principal)
+apply (erule udom_approx_mono)
+done
+
+lemma finite_deflation_udom_approx: "finite_deflation (udom_approx i)"
+proof
+  fix x show "udom_approx i\<cdot>(udom_approx i\<cdot>x) = udom_approx i\<cdot>x"
+    by (induct x rule: udom.principal_induct, simp)
+       (simp add: udom_approx_principal ubasis_until_idem)
+next
+  fix x show "udom_approx i\<cdot>x \<sqsubseteq> x"
+    by (induct x rule: udom.principal_induct, simp)
+       (simp add: udom_approx_principal ubasis_until_less)
+next
+  have *: "finite (range (\<lambda>x. udom_principal (ubasis_until (\<lambda>y. y \<le> i) x)))"
+    apply (subst range_composition [where f=udom_principal])
+    apply (simp add: finite_range_ubasis_until)
+    done
+  show "finite {x. udom_approx i\<cdot>x = x}"
+    apply (rule finite_range_imp_finite_fixes)
+    apply (rule rev_finite_subset [OF *])
+    apply (clarsimp, rename_tac x)
+    apply (induct_tac x rule: udom.principal_induct)
+    apply (simp add: adm_mem_finite *)
+    apply (simp add: udom_approx_principal)
+    done
+qed
+
+interpretation udom_approx: finite_deflation "udom_approx i"
+by (rule finite_deflation_udom_approx)
+
+lemma chain_udom_approx [simp]: "chain (\<lambda>i. udom_approx i)"
+unfolding udom_approx_def
+apply (rule chainI)
+apply (rule udom.basis_fun_mono)
+apply (erule udom_approx_mono)
+apply (erule udom_approx_mono)
+apply (rule udom.principal_mono)
+apply (rule ubasis_until_chain, simp)
+done
+
+lemma lub_udom_approx [simp]: "(\<Squnion>i. udom_approx i) = ID"
+apply (rule ext_cfun, simp add: contlub_cfun_fun)
+apply (rule below_antisym)
+apply (rule is_lub_thelub)
+apply (simp)
+apply (rule ub_rangeI)
+apply (rule udom_approx.below)
+apply (rule_tac x=x in udom.principal_induct)
+apply (simp add: lub_distribs)
+apply (rule rev_below_trans)
+apply (rule_tac x=a in is_ub_thelub)
+apply simp
+apply (simp add: udom_approx_principal)
+apply (simp add: ubasis_until_same ubasis_le_refl)
+done
+ 
+lemma udom_approx: "approx_chain udom_approx"
+proof
+  show "chain (\<lambda>i. udom_approx i)"
+    by (rule chain_udom_approx)
+  show "(\<Squnion>i. udom_approx i) = ID"
+    by (rule lub_udom_approx)
+qed
+
+hide_const (open) node
+
+end
--- a/src/HOLCF/Up.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/Up.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -5,7 +5,7 @@
 header {* The type of lifted values *}
 
 theory Up
-imports Bifinite
+imports Algebraic
 begin
 
 default_sort cpo
@@ -332,35 +332,62 @@
     by (rule finite_subset, simp add: d.finite_fixes)
 qed
 
-subsection {* Lifted cpo is a bifinite domain *}
+subsection {* Lifted cpo is an SFP domain *}
+
+definition u_approx :: "nat \<Rightarrow> udom\<^sub>\<bottom> \<rightarrow> udom\<^sub>\<bottom>"
+where "u_approx = (\<lambda>i. u_map\<cdot>(udom_approx i))"
 
-instantiation u :: (profinite) bifinite
+lemma u_approx: "approx_chain u_approx"
+proof (rule approx_chain.intro)
+  show "chain (\<lambda>i. u_approx i)"
+    unfolding u_approx_def by simp
+  show "(\<Squnion>i. u_approx i) = ID"
+    unfolding u_approx_def
+    by (simp add: lub_distribs u_map_ID)
+  show "\<And>i. finite_deflation (u_approx i)"
+    unfolding u_approx_def
+    by (intro finite_deflation_u_map finite_deflation_udom_approx)
+qed
+
+definition u_sfp :: "sfp \<rightarrow> sfp"
+where "u_sfp = sfp_fun1 u_approx u_map"
+
+lemma cast_u_sfp:
+  "cast\<cdot>(u_sfp\<cdot>A) =
+    udom_emb u_approx oo u_map\<cdot>(cast\<cdot>A) oo udom_prj u_approx"
+unfolding u_sfp_def
+apply (rule cast_sfp_fun1 [OF u_approx])
+apply (erule finite_deflation_u_map)
+done
+
+instantiation u :: (sfp) sfp
 begin
 
 definition
-  approx_up_def:
-    "approx = (\<lambda>n. u_map\<cdot>(approx n))"
+  "emb = udom_emb u_approx oo u_map\<cdot>emb"
+
+definition
+  "prj = u_map\<cdot>prj oo udom_prj u_approx"
+
+definition
+  "sfp (t::'a u itself) = u_sfp\<cdot>SFP('a)"
 
 instance proof
-  fix i :: nat and x :: "'a u"
-  show "chain (approx :: nat \<Rightarrow> 'a u \<rightarrow> 'a u)"
-    unfolding approx_up_def by simp
-  show "(\<Squnion>i. approx i\<cdot>x) = x"
-    unfolding approx_up_def
-    by (induct x, simp, simp add: lub_distribs)
-  show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x"
-    unfolding approx_up_def
-    by (induct x) simp_all
-  show "finite {x::'a u. approx i\<cdot>x = x}"
-    unfolding approx_up_def
-    by (intro finite_deflation.finite_fixes
-              finite_deflation_u_map
-              finite_deflation_approx)
+  show "ep_pair emb (prj :: udom \<rightarrow> 'a u)"
+    unfolding emb_u_def prj_u_def
+    using ep_pair_udom [OF u_approx]
+    by (intro ep_pair_comp ep_pair_u_map ep_pair_emb_prj)
+next
+  show "cast\<cdot>SFP('a u) = emb oo (prj :: udom \<rightarrow> 'a u)"
+    unfolding emb_u_def prj_u_def sfp_u_def cast_u_sfp
+    by (simp add: cast_SFP oo_def expand_cfun_eq u_map_map)
 qed
 
 end
 
-lemma approx_up [simp]: "approx i\<cdot>(up\<cdot>x) = up\<cdot>(approx i\<cdot>x)"
-unfolding approx_up_def by simp
+text {* SFP of type constructor = type combinator *}
+
+lemma SFP_u: "SFP('a::sfp u) = u_sfp\<cdot>SFP('a)"
+by (rule sfp_u_def)
 
 end
--- a/src/HOLCF/UpperPD.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/UpperPD.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -69,27 +69,6 @@
 apply (simp add: upper_le_PDPlus_iff 3)
 done
 
-lemma pd_take_upper_chain:
-  "pd_take n t \<le>\<sharp> pd_take (Suc n) t"
-apply (induct t rule: pd_basis_induct)
-apply (simp add: compact_basis.take_chain)
-apply (simp add: PDPlus_upper_mono)
-done
-
-lemma pd_take_upper_le: "pd_take i t \<le>\<sharp> t"
-apply (induct t rule: pd_basis_induct)
-apply (simp add: compact_basis.take_less)
-apply (simp add: PDPlus_upper_mono)
-done
-
-lemma pd_take_upper_mono:
-  "t \<le>\<sharp> u \<Longrightarrow> pd_take n t \<le>\<sharp> pd_take n u"
-apply (erule upper_le_induct)
-apply (simp add: compact_basis.take_mono)
-apply (simp add: upper_le_PDPlus_PDUnit_iff)
-apply (simp add: upper_le_PDPlus_iff)
-done
-
 
 subsection {* Type definition *}
 
@@ -97,7 +76,7 @@
   "{S::'a pd_basis set. upper_le.ideal S}"
 by (fast intro: upper_le.ideal_principal)
 
-instantiation upper_pd :: (profinite) below
+instantiation upper_pd :: (sfp) below
 begin
 
 definition
@@ -106,18 +85,18 @@
 instance ..
 end
 
-instance upper_pd :: (profinite) po
-by (rule upper_le.typedef_ideal_po
-    [OF type_definition_upper_pd below_upper_pd_def])
+instance upper_pd :: (sfp) po
+using type_definition_upper_pd below_upper_pd_def
+by (rule upper_le.typedef_ideal_po)
 
-instance upper_pd :: (profinite) cpo
-by (rule upper_le.typedef_ideal_cpo
-    [OF type_definition_upper_pd below_upper_pd_def])
+instance upper_pd :: (sfp) cpo
+using type_definition_upper_pd below_upper_pd_def
+by (rule upper_le.typedef_ideal_cpo)
 
 lemma Rep_upper_pd_lub:
   "chain Y \<Longrightarrow> Rep_upper_pd (\<Squnion>i. Y i) = (\<Union>i. Rep_upper_pd (Y i))"
-by (rule upper_le.typedef_ideal_rep_contlub
-    [OF type_definition_upper_pd below_upper_pd_def])
+using type_definition_upper_pd below_upper_pd_def
+by (rule upper_le.typedef_ideal_rep_contlub)
 
 lemma ideal_Rep_upper_pd: "upper_le.ideal (Rep_upper_pd xs)"
 by (rule Rep_upper_pd [unfolded mem_Collect_eq])
@@ -132,18 +111,13 @@
 by (simp add: Abs_upper_pd_inverse upper_le.ideal_principal)
 
 interpretation upper_pd:
-  ideal_completion upper_le pd_take upper_principal Rep_upper_pd
+  ideal_completion upper_le upper_principal Rep_upper_pd
 apply unfold_locales
-apply (rule pd_take_upper_le)
-apply (rule pd_take_idem)
-apply (erule pd_take_upper_mono)
-apply (rule pd_take_upper_chain)
-apply (rule finite_range_pd_take)
-apply (rule pd_take_covers)
 apply (rule ideal_Rep_upper_pd)
 apply (erule Rep_upper_pd_lub)
 apply (rule Rep_upper_principal)
 apply (simp only: below_upper_pd_def)
+apply (rule pd_basis_countable)
 done
 
 text {* Upper powerdomain is pointed *}
@@ -151,42 +125,12 @@
 lemma upper_pd_minimal: "upper_principal (PDUnit compact_bot) \<sqsubseteq> ys"
 by (induct ys rule: upper_pd.principal_induct, simp, simp)
 
-instance upper_pd :: (bifinite) pcpo
+instance upper_pd :: (sfp) pcpo
 by intro_classes (fast intro: upper_pd_minimal)
 
 lemma inst_upper_pd_pcpo: "\<bottom> = upper_principal (PDUnit compact_bot)"
 by (rule upper_pd_minimal [THEN UU_I, symmetric])
 
-text {* Upper powerdomain is profinite *}
-
-instantiation upper_pd :: (profinite) profinite
-begin
-
-definition
-  approx_upper_pd_def: "approx = upper_pd.completion_approx"
-
-instance
-apply (intro_classes, unfold approx_upper_pd_def)
-apply (rule upper_pd.chain_completion_approx)
-apply (rule upper_pd.lub_completion_approx)
-apply (rule upper_pd.completion_approx_idem)
-apply (rule upper_pd.finite_fixes_completion_approx)
-done
-
-end
-
-instance upper_pd :: (bifinite) bifinite ..
-
-lemma approx_upper_principal [simp]:
-  "approx n\<cdot>(upper_principal t) = upper_principal (pd_take n t)"
-unfolding approx_upper_pd_def
-by (rule upper_pd.completion_approx_principal)
-
-lemma approx_eq_upper_principal:
-  "\<exists>t\<in>Rep_upper_pd xs. approx n\<cdot>xs = upper_principal (pd_take n t)"
-unfolding approx_upper_pd_def
-by (rule upper_pd.completion_approx_eq_principal)
-
 
 subsection {* Monadic unit and plus *}
 
@@ -222,16 +166,6 @@
 by (simp add: upper_pd.basis_fun_principal
     upper_pd.basis_fun_mono PDPlus_upper_mono)
 
-lemma approx_upper_unit [simp]:
-  "approx n\<cdot>{x}\<sharp> = {approx n\<cdot>x}\<sharp>"
-apply (induct x rule: compact_basis.principal_induct, simp)
-apply (simp add: approx_Rep_compact_basis)
-done
-
-lemma approx_upper_plus [simp]:
-  "approx n\<cdot>(xs +\<sharp> ys) = (approx n\<cdot>xs) +\<sharp> (approx n\<cdot>ys)"
-by (induct xs ys rule: upper_pd.principal_induct2, simp, simp, simp)
-
 interpretation upper_add: semilattice upper_add proof
   fix xs ys zs :: "'a upper_pd"
   show "(xs +\<sharp> ys) +\<sharp> zs = xs +\<sharp> (ys +\<sharp> zs)"
@@ -307,7 +241,8 @@
 unfolding po_eq_conv by simp
 
 lemma upper_unit_strict [simp]: "{\<bottom>}\<sharp> = \<bottom>"
-unfolding inst_upper_pd_pcpo Rep_compact_bot [symmetric] by simp
+using upper_unit_Rep_compact_basis [of compact_bot]
+by (simp add: inst_upper_pd_pcpo)
 
 lemma upper_plus_strict1 [simp]: "\<bottom> +\<sharp> ys = \<bottom>"
 by (rule UU_I, rule upper_plus_below1)
@@ -328,8 +263,14 @@
 apply auto
 done
 
+lemma compact_upper_unit: "compact x \<Longrightarrow> compact {x}\<sharp>"
+by (auto dest!: compact_basis.compact_imp_principal)
+
 lemma compact_upper_unit_iff [simp]: "compact {x}\<sharp> \<longleftrightarrow> compact x"
-unfolding profinite_compact_iff by simp
+apply (safe elim!: compact_upper_unit)
+apply (simp only: compact_def upper_unit_below_iff [symmetric])
+apply (erule adm_subst [OF cont_Rep_CFun2])
+done
 
 lemma compact_upper_plus [simp]:
   "\<lbrakk>compact xs; compact ys\<rbrakk> \<Longrightarrow> compact (xs +\<sharp> ys)"
@@ -424,16 +365,12 @@
 unfolding upper_unit_strict [symmetric] by (rule upper_bind_unit)
 
 
-subsection {* Map and join *}
+subsection {* Map *}
 
 definition
   upper_map :: "('a \<rightarrow> 'b) \<rightarrow> 'a upper_pd \<rightarrow> 'b upper_pd" where
   "upper_map = (\<Lambda> f xs. upper_bind\<cdot>xs\<cdot>(\<Lambda> x. {f\<cdot>x}\<sharp>))"
 
-definition
-  upper_join :: "'a upper_pd upper_pd \<rightarrow> 'a upper_pd" where
-  "upper_join = (\<Lambda> xss. upper_bind\<cdot>xss\<cdot>(\<Lambda> xs. xs))"
-
 lemma upper_map_unit [simp]:
   "upper_map\<cdot>f\<cdot>{x}\<sharp> = {f\<cdot>x}\<sharp>"
 unfolding upper_map_def by simp
@@ -442,14 +379,6 @@
   "upper_map\<cdot>f\<cdot>(xs +\<sharp> ys) = upper_map\<cdot>f\<cdot>xs +\<sharp> upper_map\<cdot>f\<cdot>ys"
 unfolding upper_map_def by simp
 
-lemma upper_join_unit [simp]:
-  "upper_join\<cdot>{xs}\<sharp> = xs"
-unfolding upper_join_def by simp
-
-lemma upper_join_plus [simp]:
-  "upper_join\<cdot>(xss +\<sharp> yss) = upper_join\<cdot>xss +\<sharp> upper_join\<cdot>yss"
-unfolding upper_join_def by simp
-
 lemma upper_map_ident: "upper_map\<cdot>(\<Lambda> x. x)\<cdot>xs = xs"
 by (induct xs rule: upper_pd_induct, simp_all)
 
@@ -460,22 +389,6 @@
   "upper_map\<cdot>f\<cdot>(upper_map\<cdot>g\<cdot>xs) = upper_map\<cdot>(\<Lambda> x. f\<cdot>(g\<cdot>x))\<cdot>xs"
 by (induct xs rule: upper_pd_induct, simp_all)
 
-lemma upper_join_map_unit:
-  "upper_join\<cdot>(upper_map\<cdot>upper_unit\<cdot>xs) = xs"
-by (induct xs rule: upper_pd_induct, simp_all)
-
-lemma upper_join_map_join:
-  "upper_join\<cdot>(upper_map\<cdot>upper_join\<cdot>xsss) = upper_join\<cdot>(upper_join\<cdot>xsss)"
-by (induct xsss rule: upper_pd_induct, simp_all)
-
-lemma upper_join_map_map:
-  "upper_join\<cdot>(upper_map\<cdot>(upper_map\<cdot>f)\<cdot>xss) =
-   upper_map\<cdot>f\<cdot>(upper_join\<cdot>xss)"
-by (induct xss rule: upper_pd_induct, simp_all)
-
-lemma upper_map_approx: "upper_map\<cdot>(approx n)\<cdot>xs = approx n\<cdot>xs"
-by (induct xs rule: upper_pd_induct, simp_all)
-
 lemma ep_pair_upper_map: "ep_pair e p \<Longrightarrow> ep_pair (upper_map\<cdot>e) (upper_map\<cdot>p)"
 apply default
 apply (induct_tac x rule: upper_pd_induct, simp_all add: ep_pair.e_inverse)
@@ -490,4 +403,134 @@
 apply (simp_all add: deflation.below monofun_cfun)
 done
 
+(* FIXME: long proof! *)
+lemma finite_deflation_upper_map:
+  assumes "finite_deflation d" shows "finite_deflation (upper_map\<cdot>d)"
+proof (rule finite_deflation_intro)
+  interpret d: finite_deflation d by fact
+  have "deflation d" by fact
+  thus "deflation (upper_map\<cdot>d)" by (rule deflation_upper_map)
+  have "finite (range (\<lambda>x. d\<cdot>x))" by (rule d.finite_range)
+  hence "finite (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))"
+    by (rule finite_vimageI, simp add: inj_on_def Rep_compact_basis_inject)
+  hence "finite (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x)))" by simp
+  hence "finite (Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))"
+    by (rule finite_vimageI, simp add: inj_on_def Rep_pd_basis_inject)
+  hence *: "finite (upper_principal ` Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))" by simp
+  hence "finite (range (\<lambda>xs. upper_map\<cdot>d\<cdot>xs))"
+    apply (rule rev_finite_subset)
+    apply clarsimp
+    apply (induct_tac xs rule: upper_pd.principal_induct)
+    apply (simp add: adm_mem_finite *)
+    apply (rename_tac t, induct_tac t rule: pd_basis_induct)
+    apply (simp only: upper_unit_Rep_compact_basis [symmetric] upper_map_unit)
+    apply simp
+    apply (subgoal_tac "\<exists>b. d\<cdot>(Rep_compact_basis a) = Rep_compact_basis b")
+    apply clarsimp
+    apply (rule imageI)
+    apply (rule vimageI2)
+    apply (simp add: Rep_PDUnit)
+    apply (rule range_eqI)
+    apply (erule sym)
+    apply (rule exI)
+    apply (rule Abs_compact_basis_inverse [symmetric])
+    apply (simp add: d.compact)
+    apply (simp only: upper_plus_principal [symmetric] upper_map_plus)
+    apply clarsimp
+    apply (rule imageI)
+    apply (rule vimageI2)
+    apply (simp add: Rep_PDPlus)
+    done
+  thus "finite {xs. upper_map\<cdot>d\<cdot>xs = xs}"
+    by (rule finite_range_imp_finite_fixes)
+qed
+
+subsection {* Upper powerdomain is an SFP domain *}
+
+definition
+  upper_approx :: "nat \<Rightarrow> udom upper_pd \<rightarrow> udom upper_pd"
+where
+  "upper_approx = (\<lambda>i. upper_map\<cdot>(udom_approx i))"
+
+lemma upper_approx: "approx_chain upper_approx"
+proof (rule approx_chain.intro)
+  show "chain (\<lambda>i. upper_approx i)"
+    unfolding upper_approx_def by simp
+  show "(\<Squnion>i. upper_approx i) = ID"
+    unfolding upper_approx_def
+    by (simp add: lub_distribs upper_map_ID)
+  show "\<And>i. finite_deflation (upper_approx i)"
+    unfolding upper_approx_def
+    by (intro finite_deflation_upper_map finite_deflation_udom_approx)
+qed
+
+definition upper_sfp :: "sfp \<rightarrow> sfp"
+where "upper_sfp = sfp_fun1 upper_approx upper_map"
+
+lemma cast_upper_sfp:
+  "cast\<cdot>(upper_sfp\<cdot>A) =
+    udom_emb upper_approx oo upper_map\<cdot>(cast\<cdot>A) oo udom_prj upper_approx"
+unfolding upper_sfp_def
+apply (rule cast_sfp_fun1 [OF upper_approx])
+apply (erule finite_deflation_upper_map)
+done
+
+instantiation upper_pd :: (sfp) sfp
+begin
+
+definition
+  "emb = udom_emb upper_approx oo upper_map\<cdot>emb"
+
+definition
+  "prj = upper_map\<cdot>prj oo udom_prj upper_approx"
+
+definition
+  "sfp (t::'a upper_pd itself) = upper_sfp\<cdot>SFP('a)"
+
+instance proof
+  show "ep_pair emb (prj :: udom \<rightarrow> 'a upper_pd)"
+    unfolding emb_upper_pd_def prj_upper_pd_def
+    using ep_pair_udom [OF upper_approx]
+    by (intro ep_pair_comp ep_pair_upper_map ep_pair_emb_prj)
+next
+  show "cast\<cdot>SFP('a upper_pd) = emb oo (prj :: udom \<rightarrow> 'a upper_pd)"
+    unfolding emb_upper_pd_def prj_upper_pd_def sfp_upper_pd_def cast_upper_sfp
+    by (simp add: cast_SFP oo_def expand_cfun_eq upper_map_map)
+qed
+
 end
+
+text {* SFP of type constructor = type combinator *}
+
+lemma SFP_upper: "SFP('a upper_pd) = upper_sfp\<cdot>SFP('a)"
+by (rule sfp_upper_pd_def)
+
+
+subsection {* Join *}
+
+definition
+  upper_join :: "'a upper_pd upper_pd \<rightarrow> 'a upper_pd" where
+  "upper_join = (\<Lambda> xss. upper_bind\<cdot>xss\<cdot>(\<Lambda> xs. xs))"
+
+lemma upper_join_unit [simp]:
+  "upper_join\<cdot>{xs}\<sharp> = xs"
+unfolding upper_join_def by simp
+
+lemma upper_join_plus [simp]:
+  "upper_join\<cdot>(xss +\<sharp> yss) = upper_join\<cdot>xss +\<sharp> upper_join\<cdot>yss"
+unfolding upper_join_def by simp
+
+lemma upper_join_map_unit:
+  "upper_join\<cdot>(upper_map\<cdot>upper_unit\<cdot>xs) = xs"
+by (induct xs rule: upper_pd_induct, simp_all)
+
+lemma upper_join_map_join:
+  "upper_join\<cdot>(upper_map\<cdot>upper_join\<cdot>xsss) = upper_join\<cdot>(upper_join\<cdot>xsss)"
+by (induct xsss rule: upper_pd_induct, simp_all)
+
+lemma upper_join_map_map:
+  "upper_join\<cdot>(upper_map\<cdot>(upper_map\<cdot>f)\<cdot>xss) =
+   upper_map\<cdot>f\<cdot>(upper_join\<cdot>xss)"
+by (induct xss rule: upper_pd_induct, simp_all)
+
+end
--- a/src/HOLCF/ex/Domain_Proofs.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/ex/Domain_Proofs.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -8,7 +8,7 @@
 imports HOLCF
 begin
 
-default_sort rep
+default_sort sfp
 
 (*
 
@@ -28,50 +28,50 @@
 text {* Start with the one-step non-recursive version *}
 
 definition
-  foo_bar_baz_deflF ::
-    "TypeRep \<rightarrow> TypeRep \<times> TypeRep \<times> TypeRep \<rightarrow> TypeRep \<times> TypeRep \<times> TypeRep"
+  foo_bar_baz_sfpF ::
+    "sfp \<rightarrow> sfp \<times> sfp \<times> sfp \<rightarrow> sfp \<times> sfp \<times> sfp"
 where
-  "foo_bar_baz_deflF = (\<Lambda> a. Abs_CFun (\<lambda>(t1, t2, t3). 
-    ( ssum_defl\<cdot>REP(one)\<cdot>(sprod_defl\<cdot>(u_defl\<cdot>a)\<cdot>(u_defl\<cdot>t2))
-    , u_defl\<cdot>(cfun_defl\<cdot>t3\<cdot>REP(tr))
-    , u_defl\<cdot>(cfun_defl\<cdot>(convex_defl\<cdot>t1)\<cdot>REP(tr)))))"
+  "foo_bar_baz_sfpF = (\<Lambda> a. Abs_CFun (\<lambda>(t1, t2, t3). 
+    ( ssum_sfp\<cdot>SFP(one)\<cdot>(sprod_sfp\<cdot>(u_sfp\<cdot>a)\<cdot>(u_sfp\<cdot>t2))
+    , u_sfp\<cdot>(cfun_sfp\<cdot>t3\<cdot>SFP(tr))
+    , u_sfp\<cdot>(cfun_sfp\<cdot>(convex_sfp\<cdot>t1)\<cdot>SFP(tr)))))"
 
-lemma foo_bar_baz_deflF_beta:
-  "foo_bar_baz_deflF\<cdot>a\<cdot>t =
-    ( ssum_defl\<cdot>REP(one)\<cdot>(sprod_defl\<cdot>(u_defl\<cdot>a)\<cdot>(u_defl\<cdot>(fst (snd t))))
-    , u_defl\<cdot>(cfun_defl\<cdot>(snd (snd t))\<cdot>REP(tr))
-    , u_defl\<cdot>(cfun_defl\<cdot>(convex_defl\<cdot>(fst t))\<cdot>REP(tr)))"
-unfolding foo_bar_baz_deflF_def
+lemma foo_bar_baz_sfpF_beta:
+  "foo_bar_baz_sfpF\<cdot>a\<cdot>t =
+    ( ssum_sfp\<cdot>SFP(one)\<cdot>(sprod_sfp\<cdot>(u_sfp\<cdot>a)\<cdot>(u_sfp\<cdot>(fst (snd t))))
+    , u_sfp\<cdot>(cfun_sfp\<cdot>(snd (snd t))\<cdot>SFP(tr))
+    , u_sfp\<cdot>(cfun_sfp\<cdot>(convex_sfp\<cdot>(fst t))\<cdot>SFP(tr)))"
+unfolding foo_bar_baz_sfpF_def
 by (simp add: split_def)
 
 text {* Individual type combinators are projected from the fixed point. *}
 
-definition foo_defl :: "TypeRep \<rightarrow> TypeRep"
-where "foo_defl = (\<Lambda> a. fst (fix\<cdot>(foo_bar_baz_deflF\<cdot>a)))"
+definition foo_sfp :: "sfp \<rightarrow> sfp"
+where "foo_sfp = (\<Lambda> a. fst (fix\<cdot>(foo_bar_baz_sfpF\<cdot>a)))"
 
-definition bar_defl :: "TypeRep \<rightarrow> TypeRep"
-where "bar_defl = (\<Lambda> a. fst (snd (fix\<cdot>(foo_bar_baz_deflF\<cdot>a))))"
+definition bar_sfp :: "sfp \<rightarrow> sfp"
+where "bar_sfp = (\<Lambda> a. fst (snd (fix\<cdot>(foo_bar_baz_sfpF\<cdot>a))))"
 
-definition baz_defl :: "TypeRep \<rightarrow> TypeRep"
-where "baz_defl = (\<Lambda> a. snd (snd (fix\<cdot>(foo_bar_baz_deflF\<cdot>a))))"
+definition baz_sfp :: "sfp \<rightarrow> sfp"
+where "baz_sfp = (\<Lambda> a. snd (snd (fix\<cdot>(foo_bar_baz_sfpF\<cdot>a))))"
 
 lemma defl_apply_thms:
-  "foo_defl\<cdot>a = fst (fix\<cdot>(foo_bar_baz_deflF\<cdot>a))"
-  "bar_defl\<cdot>a = fst (snd (fix\<cdot>(foo_bar_baz_deflF\<cdot>a)))"
-  "baz_defl\<cdot>a = snd (snd (fix\<cdot>(foo_bar_baz_deflF\<cdot>a)))"
-unfolding foo_defl_def bar_defl_def baz_defl_def by simp_all
+  "foo_sfp\<cdot>a = fst (fix\<cdot>(foo_bar_baz_sfpF\<cdot>a))"
+  "bar_sfp\<cdot>a = fst (snd (fix\<cdot>(foo_bar_baz_sfpF\<cdot>a)))"
+  "baz_sfp\<cdot>a = snd (snd (fix\<cdot>(foo_bar_baz_sfpF\<cdot>a)))"
+unfolding foo_sfp_def bar_sfp_def baz_sfp_def by simp_all
 
 text {* Unfold rules for each combinator. *}
 
-lemma foo_defl_unfold:
-  "foo_defl\<cdot>a = ssum_defl\<cdot>REP(one)\<cdot>(sprod_defl\<cdot>(u_defl\<cdot>a)\<cdot>(u_defl\<cdot>(bar_defl\<cdot>a)))"
-unfolding defl_apply_thms by (subst fix_eq, simp add: foo_bar_baz_deflF_beta)
+lemma foo_sfp_unfold:
+  "foo_sfp\<cdot>a = ssum_sfp\<cdot>SFP(one)\<cdot>(sprod_sfp\<cdot>(u_sfp\<cdot>a)\<cdot>(u_sfp\<cdot>(bar_sfp\<cdot>a)))"
+unfolding defl_apply_thms by (subst fix_eq, simp add: foo_bar_baz_sfpF_beta)
 
-lemma bar_defl_unfold: "bar_defl\<cdot>a = u_defl\<cdot>(cfun_defl\<cdot>(baz_defl\<cdot>a)\<cdot>REP(tr))"
-unfolding defl_apply_thms by (subst fix_eq, simp add: foo_bar_baz_deflF_beta)
+lemma bar_sfp_unfold: "bar_sfp\<cdot>a = u_sfp\<cdot>(cfun_sfp\<cdot>(baz_sfp\<cdot>a)\<cdot>SFP(tr))"
+unfolding defl_apply_thms by (subst fix_eq, simp add: foo_bar_baz_sfpF_beta)
 
-lemma baz_defl_unfold: "baz_defl\<cdot>a = u_defl\<cdot>(cfun_defl\<cdot>(convex_defl\<cdot>(foo_defl\<cdot>a))\<cdot>REP(tr))"
-unfolding defl_apply_thms by (subst fix_eq, simp add: foo_bar_baz_deflF_beta)
+lemma baz_sfp_unfold: "baz_sfp\<cdot>a = u_sfp\<cdot>(cfun_sfp\<cdot>(convex_sfp\<cdot>(foo_sfp\<cdot>a))\<cdot>SFP(tr))"
+unfolding defl_apply_thms by (subst fix_eq, simp add: foo_bar_baz_sfpF_beta)
 
 text "The automation for the previous steps will be quite similar to
 how the fixrec package works."
@@ -82,28 +82,28 @@
 
 text {* Use @{text pcpodef} with the appropriate type combinator. *}
 
-pcpodef (open) 'a foo = "{x. x ::: foo_defl\<cdot>REP('a)}"
-by (simp_all add: adm_in_deflation)
+pcpodef (open) 'a foo = "{x. x ::: foo_sfp\<cdot>SFP('a)}"
+by (simp_all add: adm_in_sfp)
 
-pcpodef (open) 'a bar = "{x. x ::: bar_defl\<cdot>REP('a)}"
-by (simp_all add: adm_in_deflation)
+pcpodef (open) 'a bar = "{x. x ::: bar_sfp\<cdot>SFP('a)}"
+by (simp_all add: adm_in_sfp)
 
-pcpodef (open) 'a baz = "{x. x ::: baz_defl\<cdot>REP('a)}"
-by (simp_all add: adm_in_deflation)
+pcpodef (open) 'a baz = "{x. x ::: baz_sfp\<cdot>SFP('a)}"
+by (simp_all add: adm_in_sfp)
 
 text {* Prove rep instance using lemma @{text typedef_rep_class}. *}
 
-instantiation foo :: (rep) rep
+instantiation foo :: (sfp) sfp
 begin
 
 definition emb_foo :: "'a foo \<rightarrow> udom"
 where "emb_foo \<equiv> (\<Lambda> x. Rep_foo x)"
 
 definition prj_foo :: "udom \<rightarrow> 'a foo"
-where "prj_foo \<equiv> (\<Lambda> y. Abs_foo (cast\<cdot>(foo_defl\<cdot>REP('a))\<cdot>y))"
+where "prj_foo \<equiv> (\<Lambda> y. Abs_foo (cast\<cdot>(foo_sfp\<cdot>SFP('a))\<cdot>y))"
 
-definition approx_foo :: "nat \<Rightarrow> 'a foo \<rightarrow> 'a foo"
-where "approx_foo \<equiv> repdef_approx Rep_foo Abs_foo (foo_defl\<cdot>REP('a))"
+definition sfp_foo :: "'a foo itself \<Rightarrow> sfp"
+where "sfp_foo \<equiv> \<lambda>a. foo_sfp\<cdot>SFP('a)"
 
 instance
 apply (rule typedef_rep_class)
@@ -111,22 +111,22 @@
 apply (rule below_foo_def)
 apply (rule emb_foo_def)
 apply (rule prj_foo_def)
-apply (rule approx_foo_def)
+apply (rule sfp_foo_def)
 done
 
 end
 
-instantiation bar :: (rep) rep
+instantiation bar :: (sfp) sfp
 begin
 
 definition emb_bar :: "'a bar \<rightarrow> udom"
 where "emb_bar \<equiv> (\<Lambda> x. Rep_bar x)"
 
 definition prj_bar :: "udom \<rightarrow> 'a bar"
-where "prj_bar \<equiv> (\<Lambda> y. Abs_bar (cast\<cdot>(bar_defl\<cdot>REP('a))\<cdot>y))"
+where "prj_bar \<equiv> (\<Lambda> y. Abs_bar (cast\<cdot>(bar_sfp\<cdot>SFP('a))\<cdot>y))"
 
-definition approx_bar :: "nat \<Rightarrow> 'a bar \<rightarrow> 'a bar"
-where "approx_bar \<equiv> repdef_approx Rep_bar Abs_bar (bar_defl\<cdot>REP('a))"
+definition sfp_bar :: "'a bar itself \<Rightarrow> sfp"
+where "sfp_bar \<equiv> \<lambda>a. bar_sfp\<cdot>SFP('a)"
 
 instance
 apply (rule typedef_rep_class)
@@ -134,22 +134,22 @@
 apply (rule below_bar_def)
 apply (rule emb_bar_def)
 apply (rule prj_bar_def)
-apply (rule approx_bar_def)
+apply (rule sfp_bar_def)
 done
 
 end
 
-instantiation baz :: (rep) rep
+instantiation baz :: (sfp) sfp
 begin
 
 definition emb_baz :: "'a baz \<rightarrow> udom"
 where "emb_baz \<equiv> (\<Lambda> x. Rep_baz x)"
 
 definition prj_baz :: "udom \<rightarrow> 'a baz"
-where "prj_baz \<equiv> (\<Lambda> y. Abs_baz (cast\<cdot>(baz_defl\<cdot>REP('a))\<cdot>y))"
+where "prj_baz \<equiv> (\<Lambda> y. Abs_baz (cast\<cdot>(baz_sfp\<cdot>SFP('a))\<cdot>y))"
 
-definition approx_baz :: "nat \<Rightarrow> 'a baz \<rightarrow> 'a baz"
-where "approx_baz \<equiv> repdef_approx Rep_baz Abs_baz (baz_defl\<cdot>REP('a))"
+definition sfp_baz :: "'a baz itself \<Rightarrow> sfp"
+where "sfp_baz \<equiv> \<lambda>a. baz_sfp\<cdot>SFP('a)"
 
 instance
 apply (rule typedef_rep_class)
@@ -157,50 +157,44 @@
 apply (rule below_baz_def)
 apply (rule emb_baz_def)
 apply (rule prj_baz_def)
-apply (rule approx_baz_def)
+apply (rule sfp_baz_def)
 done
 
 end
 
-text {* Prove REP rules using lemma @{text typedef_REP}. *}
+text {* Prove SFP rules using lemma @{text typedef_SFP}. *}
 
-lemma REP_foo: "REP('a foo) = foo_defl\<cdot>REP('a)"
-apply (rule typedef_REP)
-apply (rule type_definition_foo)
-apply (rule below_foo_def)
-apply (rule emb_foo_def)
-apply (rule prj_foo_def)
+lemma SFP_foo: "SFP('a foo) = foo_sfp\<cdot>SFP('a)"
+apply (rule typedef_SFP)
+apply (rule sfp_foo_def)
 done
 
-lemma REP_bar: "REP('a bar) = bar_defl\<cdot>REP('a)"
-apply (rule typedef_REP)
-apply (rule type_definition_bar)
-apply (rule below_bar_def)
-apply (rule emb_bar_def)
-apply (rule prj_bar_def)
+lemma SFP_bar: "SFP('a bar) = bar_sfp\<cdot>SFP('a)"
+apply (rule typedef_SFP)
+apply (rule sfp_bar_def)
+done
+
+lemma SFP_baz: "SFP('a baz) = baz_sfp\<cdot>SFP('a)"
+apply (rule typedef_SFP)
+apply (rule sfp_baz_def)
 done
 
-lemma REP_baz: "REP('a baz) = baz_defl\<cdot>REP('a)"
-apply (rule typedef_REP)
-apply (rule type_definition_baz)
-apply (rule below_baz_def)
-apply (rule emb_baz_def)
-apply (rule prj_baz_def)
-done
+text {* Prove SFP equations using type combinator unfold lemmas. *}
 
-text {* Prove REP equations using type combinator unfold lemmas. *}
+lemmas SFP_simps =
+  SFP_ssum SFP_sprod SFP_u SFP_cfun
 
-lemma REP_foo': "REP('a foo) = REP(one \<oplus> 'a\<^sub>\<bottom> \<otimes> ('a bar)\<^sub>\<bottom>)"
-unfolding REP_foo REP_bar REP_baz REP_simps
-by (rule foo_defl_unfold)
+lemma SFP_foo': "SFP('a foo) = SFP(one \<oplus> 'a\<^sub>\<bottom> \<otimes> ('a bar)\<^sub>\<bottom>)"
+unfolding SFP_foo SFP_bar SFP_baz SFP_simps
+by (rule foo_sfp_unfold)
 
-lemma REP_bar': "REP('a bar) = REP(('a baz \<rightarrow> tr)\<^sub>\<bottom>)"
-unfolding REP_foo REP_bar REP_baz REP_simps
-by (rule bar_defl_unfold)
+lemma SFP_bar': "SFP('a bar) = SFP(('a baz \<rightarrow> tr)\<^sub>\<bottom>)"
+unfolding SFP_foo SFP_bar SFP_baz SFP_simps
+by (rule bar_sfp_unfold)
 
-lemma REP_baz': "REP('a baz) = REP(('a foo convex_pd \<rightarrow> tr)\<^sub>\<bottom>)"
-unfolding REP_foo REP_bar REP_baz REP_simps REP_convex
-by (rule baz_defl_unfold)
+lemma SFP_baz': "SFP('a baz) = SFP(('a foo convex_pd \<rightarrow> tr)\<^sub>\<bottom>)"
+unfolding SFP_foo SFP_bar SFP_baz SFP_simps SFP_convex
+by (rule baz_sfp_unfold)
 
 (********************************************************************)
 
@@ -229,36 +223,36 @@
 text {* Prove isomorphism rules. *}
 
 lemma foo_abs_iso: "foo_rep\<cdot>(foo_abs\<cdot>x) = x"
-by (rule domain_abs_iso [OF REP_foo' foo_abs_def foo_rep_def])
+by (rule domain_abs_iso [OF SFP_foo' foo_abs_def foo_rep_def])
 
 lemma foo_rep_iso: "foo_abs\<cdot>(foo_rep\<cdot>x) = x"
-by (rule domain_rep_iso [OF REP_foo' foo_abs_def foo_rep_def])
+by (rule domain_rep_iso [OF SFP_foo' foo_abs_def foo_rep_def])
 
 lemma bar_abs_iso: "bar_rep\<cdot>(bar_abs\<cdot>x) = x"
-by (rule domain_abs_iso [OF REP_bar' bar_abs_def bar_rep_def])
+by (rule domain_abs_iso [OF SFP_bar' bar_abs_def bar_rep_def])
 
 lemma bar_rep_iso: "bar_abs\<cdot>(bar_rep\<cdot>x) = x"
-by (rule domain_rep_iso [OF REP_bar' bar_abs_def bar_rep_def])
+by (rule domain_rep_iso [OF SFP_bar' bar_abs_def bar_rep_def])
 
 lemma baz_abs_iso: "baz_rep\<cdot>(baz_abs\<cdot>x) = x"
-by (rule domain_abs_iso [OF REP_baz' baz_abs_def baz_rep_def])
+by (rule domain_abs_iso [OF SFP_baz' baz_abs_def baz_rep_def])
 
 lemma baz_rep_iso: "baz_abs\<cdot>(baz_rep\<cdot>x) = x"
-by (rule domain_rep_iso [OF REP_baz' baz_abs_def baz_rep_def])
+by (rule domain_rep_iso [OF SFP_baz' baz_abs_def baz_rep_def])
 
 text {* Prove isodefl rules using @{text isodefl_coerce}. *}
 
 lemma isodefl_foo_abs:
   "isodefl d t \<Longrightarrow> isodefl (foo_abs oo d oo foo_rep) t"
-by (rule isodefl_abs_rep [OF REP_foo' foo_abs_def foo_rep_def])
+by (rule isodefl_abs_rep [OF SFP_foo' foo_abs_def foo_rep_def])
 
 lemma isodefl_bar_abs:
   "isodefl d t \<Longrightarrow> isodefl (bar_abs oo d oo bar_rep) t"
-by (rule isodefl_abs_rep [OF REP_bar' bar_abs_def bar_rep_def])
+by (rule isodefl_abs_rep [OF SFP_bar' bar_abs_def bar_rep_def])
 
 lemma isodefl_baz_abs:
   "isodefl d t \<Longrightarrow> isodefl (baz_abs oo d oo baz_rep) t"
-by (rule isodefl_abs_rep [OF REP_baz' baz_abs_def baz_rep_def])
+by (rule isodefl_abs_rep [OF SFP_baz' baz_abs_def baz_rep_def])
 
 (********************************************************************)
 
@@ -322,15 +316,15 @@
 lemma isodefl_foo_bar_baz:
   assumes isodefl_d: "isodefl d t"
   shows
-  "isodefl (foo_map\<cdot>d) (foo_defl\<cdot>t) \<and>
-  isodefl (bar_map\<cdot>d) (bar_defl\<cdot>t) \<and>
-  isodefl (baz_map\<cdot>d) (baz_defl\<cdot>t)"
+  "isodefl (foo_map\<cdot>d) (foo_sfp\<cdot>t) \<and>
+  isodefl (bar_map\<cdot>d) (bar_sfp\<cdot>t) \<and>
+  isodefl (baz_map\<cdot>d) (baz_sfp\<cdot>t)"
 unfolding map_apply_thms defl_apply_thms
  apply (rule parallel_fix_ind)
    apply (intro adm_conj adm_isodefl cont2cont_fst cont2cont_snd cont_id)
   apply (simp only: fst_strict snd_strict isodefl_bottom simp_thms)
  apply (simp only: foo_bar_baz_mapF_beta
-                   foo_bar_baz_deflF_beta
+                   foo_bar_baz_sfpF_beta
                    fst_conv snd_conv)
  apply (elim conjE)
  apply (intro
@@ -338,7 +332,7 @@
   isodefl_foo_abs
   isodefl_bar_abs
   isodefl_baz_abs
-  isodefl_ssum isodefl_sprod isodefl_ID_REP
+  isodefl_ssum isodefl_sprod isodefl_ID_SFP
   isodefl_u isodefl_convex isodefl_cfun
   isodefl_d
  )
@@ -349,27 +343,27 @@
 lemmas isodefl_bar = isodefl_foo_bar_baz [THEN conjunct2, THEN conjunct1]
 lemmas isodefl_baz = isodefl_foo_bar_baz [THEN conjunct2, THEN conjunct2]
 
-text {* Prove map ID lemmas, using isodefl_REP_imp_ID *}
+text {* Prove map ID lemmas, using isodefl_SFP_imp_ID *}
 
 lemma foo_map_ID: "foo_map\<cdot>ID = ID"
-apply (rule isodefl_REP_imp_ID)
-apply (subst REP_foo)
+apply (rule isodefl_SFP_imp_ID)
+apply (subst SFP_foo)
 apply (rule isodefl_foo)
-apply (rule isodefl_ID_REP)
+apply (rule isodefl_ID_SFP)
 done
 
 lemma bar_map_ID: "bar_map\<cdot>ID = ID"
-apply (rule isodefl_REP_imp_ID)
-apply (subst REP_bar)
+apply (rule isodefl_SFP_imp_ID)
+apply (subst SFP_bar)
 apply (rule isodefl_bar)
-apply (rule isodefl_ID_REP)
+apply (rule isodefl_ID_SFP)
 done
 
 lemma baz_map_ID: "baz_map\<cdot>ID = ID"
-apply (rule isodefl_REP_imp_ID)
-apply (subst REP_baz)
+apply (rule isodefl_SFP_imp_ID)
+apply (subst SFP_baz)
 apply (rule isodefl_baz)
-apply (rule isodefl_ID_REP)
+apply (rule isodefl_ID_SFP)
 done
 
 (********************************************************************)
--- a/src/HOLCF/ex/Powerdomain_ex.thy	Tue Oct 05 17:53:00 2010 -0700
+++ b/src/HOLCF/ex/Powerdomain_ex.thy	Wed Oct 06 10:49:27 2010 -0700
@@ -8,7 +8,7 @@
 imports HOLCF
 begin
 
-default_sort bifinite
+default_sort sfp
 
 subsection {* Monadic sorting example *}