src/HOLCF/Bifinite.thy
changeset 39989 ad60d7311f43
parent 39987 8c2f449af35a
child 40002 c5b5f7a3a3b1
--- a/src/HOLCF/Bifinite.thy	Mon Oct 11 07:09:42 2010 -0700
+++ b/src/HOLCF/Bifinite.thy	Mon Oct 11 08:32:09 2010 -0700
@@ -18,12 +18,12 @@
 class bifinite = pcpo +
   fixes emb :: "'a::pcpo \<rightarrow> udom"
   fixes prj :: "udom \<rightarrow> 'a::pcpo"
-  fixes sfp :: "'a itself \<Rightarrow> sfp"
+  fixes defl :: "'a itself \<Rightarrow> defl"
   assumes ep_pair_emb_prj: "ep_pair emb prj"
-  assumes cast_SFP: "cast\<cdot>(sfp TYPE('a)) = emb oo prj"
+  assumes cast_DEFL: "cast\<cdot>(defl TYPE('a)) = emb oo prj"
 
-syntax "_SFP" :: "type \<Rightarrow> sfp"  ("(1SFP/(1'(_')))")
-translations "SFP('t)" \<rightleftharpoons> "CONST sfp TYPE('t)"
+syntax "_DEFL" :: "type \<Rightarrow> defl"  ("(1DEFL/(1'(_')))")
+translations "DEFL('t)" \<rightleftharpoons> "CONST defl TYPE('t)"
 
 interpretation bifinite:
   pcpo_ep_pair "emb :: 'a::bifinite \<rightarrow> udom" "prj :: udom \<rightarrow> 'a::bifinite"
@@ -47,24 +47,24 @@
   ideal_completion below Rep_compact_basis "approximants::'a::bifinite \<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
+  and DEFL: "DEFL('a) = (\<Squnion>i. defl_principal (Y i))"
+    by (rule defl.obtain_principal_chain)
+  def approx \<equiv> "\<lambda>i. (prj oo cast\<cdot>(defl_principal (Y i)) oo emb) :: 'a \<rightarrow> 'a"
+  interpret defl_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)
+      by (simp add: lub_distribs Y DEFL [symmetric] cast_DEFL expand_cfun_eq)
     show "\<And>i. finite_deflation (approx i)"
       unfolding approx_def
       apply (rule bifinite.finite_deflation_p_d_e)
       apply (rule finite_deflation_cast)
-      apply (rule sfp.compact_principal)
+      apply (rule defl.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)
+      apply (simp add: lub_distribs Y DEFL [symmetric] cast_DEFL)
       done
   qed
   (* FIXME: why does show ?thesis fail here? *)
@@ -74,30 +74,30 @@
 subsection {* Type combinators *}
 
 definition
-  sfp_fun1 ::
-    "(nat \<Rightarrow> 'a \<rightarrow> 'a) \<Rightarrow> ((udom \<rightarrow> udom) \<rightarrow> ('a \<rightarrow> 'a)) \<Rightarrow> (sfp \<rightarrow> sfp)"
+  defl_fun1 ::
+    "(nat \<Rightarrow> 'a \<rightarrow> 'a) \<Rightarrow> ((udom \<rightarrow> udom) \<rightarrow> ('a \<rightarrow> 'a)) \<Rightarrow> (defl \<rightarrow> defl)"
 where
-  "sfp_fun1 approx f =
-    sfp.basis_fun (\<lambda>a.
-      sfp_principal (Abs_fin_defl
+  "defl_fun1 approx f =
+    defl.basis_fun (\<lambda>a.
+      defl_principal (Abs_fin_defl
         (udom_emb approx oo f\<cdot>(Rep_fin_defl a) oo udom_prj approx)))"
 
 definition
-  sfp_fun2 ::
+  defl_fun2 ::
     "(nat \<Rightarrow> 'a \<rightarrow> 'a) \<Rightarrow> ((udom \<rightarrow> udom) \<rightarrow> (udom \<rightarrow> udom) \<rightarrow> ('a \<rightarrow> 'a))
-      \<Rightarrow> (sfp \<rightarrow> sfp \<rightarrow> sfp)"
+      \<Rightarrow> (defl \<rightarrow> defl \<rightarrow> defl)"
 where
-  "sfp_fun2 approx f =
-    sfp.basis_fun (\<lambda>a.
-      sfp.basis_fun (\<lambda>b.
-        sfp_principal (Abs_fin_defl
+  "defl_fun2 approx f =
+    defl.basis_fun (\<lambda>a.
+      defl.basis_fun (\<lambda>b.
+        defl_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:
+lemma cast_defl_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"
+  shows "cast\<cdot>(defl_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)"
@@ -106,23 +106,23 @@
     apply (rule f, rule finite_deflation_Rep_fin_defl)
     done
   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
+    by (induct A rule: defl.principal_induct, simp)
+       (simp only: defl_fun1_def
+                   defl.basis_fun_principal
+                   defl.basis_fun_mono
+                   defl.principal_mono
                    Abs_fin_defl_mono [OF 1 1]
                    monofun_cfun below_refl
                    Rep_fin_defl_mono
-                   cast_sfp_principal
+                   cast_defl_principal
                    Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1])
 qed
 
-lemma cast_sfp_fun2:
+lemma cast_defl_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) =
+  shows "cast\<cdot>(defl_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
@@ -132,15 +132,15 @@
     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
+    by (induct A B rule: defl.principal_induct2, simp, simp)
+       (simp only: defl_fun2_def
+                   defl.basis_fun_principal
+                   defl.basis_fun_mono
+                   defl.principal_mono
                    Abs_fin_defl_mono [OF 1 1]
                    monofun_cfun below_refl
                    Rep_fin_defl_mono
-                   cast_sfp_principal
+                   cast_defl_principal
                    Abs_fin_defl_inverse [unfolded mem_Collect_eq, OF 1])
 qed
 
@@ -156,22 +156,22 @@
   "prj = (ID :: udom \<rightarrow> udom)"
 
 definition
-  "sfp (t::udom itself) = (\<Squnion>i. sfp_principal (Abs_fin_defl (udom_approx i)))"
+  "defl (t::udom itself) = (\<Squnion>i. defl_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
+  show "cast\<cdot>DEFL(udom) = emb oo (prj :: udom \<rightarrow> udom)"
+    unfolding defl_udom_def
     apply (subst contlub_cfun_arg)
     apply (rule chainI)
-    apply (rule sfp.principal_mono)
+    apply (rule defl.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 (subst cast_defl_principal)
     apply (simp add: Abs_fin_defl_inverse finite_deflation_udom_approx)
     done
 qed
@@ -197,14 +197,14 @@
     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"
+definition cfun_defl :: "defl \<rightarrow> defl \<rightarrow> defl"
+where "cfun_defl = defl_fun2 cfun_approx cfun_map"
 
-lemma cast_cfun_sfp:
-  "cast\<cdot>(cfun_sfp\<cdot>A\<cdot>B) =
+lemma cast_cfun_defl:
+  "cast\<cdot>(cfun_defl\<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])
+unfolding cfun_defl_def
+apply (rule cast_defl_fun2 [OF cfun_approx])
 apply (erule (1) finite_deflation_cfun_map)
 done
 
@@ -218,7 +218,7 @@
   "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)"
+  "defl (t::('a \<rightarrow> 'b) itself) = cfun_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
 
 instance proof
   show "ep_pair emb (prj :: udom \<rightarrow> 'a \<rightarrow> 'b)"
@@ -226,16 +226,16 @@
     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)
+  show "cast\<cdot>DEFL('a \<rightarrow> 'b) = emb oo (prj :: udom \<rightarrow> 'a \<rightarrow> 'b)"
+    unfolding emb_cfun_def prj_cfun_def defl_cfun_def cast_cfun_defl
+    by (simp add: cast_DEFL oo_def expand_cfun_eq cfun_map_map)
 qed
 
 end
 
-lemma SFP_cfun:
-  "SFP('a::bifinite \<rightarrow> 'b::bifinite) = cfun_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
-by (rule sfp_cfun_def)
+lemma DEFL_cfun:
+  "DEFL('a::bifinite \<rightarrow> 'b::bifinite) = cfun_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
+by (rule defl_cfun_def)
 
 subsection {* Cartesian product is a bifinite domain *}
 
@@ -256,14 +256,14 @@
     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"
+definition prod_defl :: "defl \<rightarrow> defl \<rightarrow> defl"
+where "prod_defl = defl_fun2 prod_approx cprod_map"
 
-lemma cast_prod_sfp:
-  "cast\<cdot>(prod_sfp\<cdot>A\<cdot>B) = udom_emb prod_approx oo
+lemma cast_prod_defl:
+  "cast\<cdot>(prod_defl\<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])
+unfolding prod_defl_def
+apply (rule cast_defl_fun2 [OF prod_approx])
 apply (erule (1) finite_deflation_cprod_map)
 done
 
@@ -277,7 +277,7 @@
   "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)"
+  "defl (t::('a \<times> 'b) itself) = prod_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
 
 instance proof
   show "ep_pair emb (prj :: udom \<rightarrow> 'a \<times> 'b)"
@@ -285,16 +285,16 @@
     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)
+  show "cast\<cdot>DEFL('a \<times> 'b) = emb oo (prj :: udom \<rightarrow> 'a \<times> 'b)"
+    unfolding emb_prod_def prj_prod_def defl_prod_def cast_prod_defl
+    by (simp add: cast_DEFL oo_def expand_cfun_eq cprod_map_map)
 qed
 
 end
 
-lemma SFP_prod:
-  "SFP('a::bifinite \<times> 'b::bifinite) = prod_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
-by (rule sfp_prod_def)
+lemma DEFL_prod:
+  "DEFL('a::bifinite \<times> 'b::bifinite) = prod_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
+by (rule defl_prod_def)
 
 subsection {* Strict product is a bifinite domain *}
 
@@ -315,16 +315,16 @@
     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"
+definition sprod_defl :: "defl \<rightarrow> defl \<rightarrow> defl"
+where "sprod_defl = defl_fun2 sprod_approx sprod_map"
 
-lemma cast_sprod_sfp:
-  "cast\<cdot>(sprod_sfp\<cdot>A\<cdot>B) =
+lemma cast_sprod_defl:
+  "cast\<cdot>(sprod_defl\<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])
+unfolding sprod_defl_def
+apply (rule cast_defl_fun2 [OF sprod_approx])
 apply (erule (1) finite_deflation_sprod_map)
 done
 
@@ -338,7 +338,7 @@
   "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)"
+  "defl (t::('a \<otimes> 'b) itself) = sprod_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
 
 instance proof
   show "ep_pair emb (prj :: udom \<rightarrow> 'a \<otimes> 'b)"
@@ -346,16 +346,16 @@
     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)
+  show "cast\<cdot>DEFL('a \<otimes> 'b) = emb oo (prj :: udom \<rightarrow> 'a \<otimes> 'b)"
+    unfolding emb_sprod_def prj_sprod_def defl_sprod_def cast_sprod_defl
+    by (simp add: cast_DEFL oo_def expand_cfun_eq sprod_map_map)
 qed
 
 end
 
-lemma SFP_sprod:
-  "SFP('a::bifinite \<otimes> 'b::bifinite) = sprod_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
-by (rule sfp_sprod_def)
+lemma DEFL_sprod:
+  "DEFL('a::bifinite \<otimes> 'b::bifinite) = sprod_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
+by (rule defl_sprod_def)
 
 subsection {* Lifted cpo is a bifinite domain *}
 
@@ -374,14 +374,14 @@
     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"
+definition u_defl :: "defl \<rightarrow> defl"
+where "u_defl = defl_fun1 u_approx u_map"
 
-lemma cast_u_sfp:
-  "cast\<cdot>(u_sfp\<cdot>A) =
+lemma cast_u_defl:
+  "cast\<cdot>(u_defl\<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])
+unfolding u_defl_def
+apply (rule cast_defl_fun1 [OF u_approx])
 apply (erule finite_deflation_u_map)
 done
 
@@ -395,7 +395,7 @@
   "prj = u_map\<cdot>prj oo udom_prj u_approx"
 
 definition
-  "sfp (t::'a u itself) = u_sfp\<cdot>SFP('a)"
+  "defl (t::'a u itself) = u_defl\<cdot>DEFL('a)"
 
 instance proof
   show "ep_pair emb (prj :: udom \<rightarrow> 'a u)"
@@ -403,15 +403,15 @@
     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)
+  show "cast\<cdot>DEFL('a u) = emb oo (prj :: udom \<rightarrow> 'a u)"
+    unfolding emb_u_def prj_u_def defl_u_def cast_u_defl
+    by (simp add: cast_DEFL oo_def expand_cfun_eq u_map_map)
 qed
 
 end
 
-lemma SFP_u: "SFP('a::bifinite u) = u_sfp\<cdot>SFP('a)"
-by (rule sfp_u_def)
+lemma DEFL_u: "DEFL('a::bifinite u) = u_defl\<cdot>DEFL('a)"
+by (rule defl_u_def)
 
 subsection {* Lifted countable types are bifinite domains *}
 
@@ -472,25 +472,25 @@
   "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)))"
+  "defl (t::'a lift itself) =
+    (\<Squnion>i. defl_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
+  show "cast\<cdot>DEFL('a lift) = emb oo (prj :: udom \<rightarrow> 'a lift)"
+    unfolding defl_lift_def
     apply (subst contlub_cfun_arg)
     apply (rule chainI)
-    apply (rule sfp.principal_mono)
+    apply (rule defl.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 (subst cast_defl_principal)
     apply (simp add: Abs_fin_defl_inverse finite_deflation_lift_approx
                      ep_pair.finite_deflation_e_d_p [OF ep] lub_distribs)
     done
@@ -517,14 +517,14 @@
     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"
+definition ssum_defl :: "defl \<rightarrow> defl \<rightarrow> defl"
+where "ssum_defl = defl_fun2 ssum_approx ssum_map"
 
-lemma cast_ssum_sfp:
-  "cast\<cdot>(ssum_sfp\<cdot>A\<cdot>B) =
+lemma cast_ssum_defl:
+  "cast\<cdot>(ssum_defl\<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])
+unfolding ssum_defl_def
+apply (rule cast_defl_fun2 [OF ssum_approx])
 apply (erule (1) finite_deflation_ssum_map)
 done
 
@@ -538,7 +538,7 @@
   "prj = ssum_map\<cdot>prj\<cdot>prj oo udom_prj ssum_approx"
 
 definition
-  "sfp (t::('a \<oplus> 'b) itself) = ssum_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
+  "defl (t::('a \<oplus> 'b) itself) = ssum_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
 
 instance proof
   show "ep_pair emb (prj :: udom \<rightarrow> 'a \<oplus> 'b)"
@@ -546,15 +546,15 @@
     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)
+  show "cast\<cdot>DEFL('a \<oplus> 'b) = emb oo (prj :: udom \<rightarrow> 'a \<oplus> 'b)"
+    unfolding emb_ssum_def prj_ssum_def defl_ssum_def cast_ssum_defl
+    by (simp add: cast_DEFL oo_def expand_cfun_eq ssum_map_map)
 qed
 
 end
 
-lemma SFP_ssum:
-  "SFP('a::bifinite \<oplus> 'b::bifinite) = ssum_sfp\<cdot>SFP('a)\<cdot>SFP('b)"
-by (rule sfp_ssum_def)
+lemma DEFL_ssum:
+  "DEFL('a::bifinite \<oplus> 'b::bifinite) = ssum_defl\<cdot>DEFL('a)\<cdot>DEFL('b)"
+by (rule defl_ssum_def)
 
 end