Theory Representable

Up to index of Isabelle/HOLCF

theory Representable
imports Algebraic Universal Domain_Aux
uses (Tools/repdef.ML) (Tools/Domain/domain_isomorphism.ML)

(*  Title:      HOLCF/Representable.thy
Author: Brian Huffman
*)


header {* Representable Types *}

theory Representable
imports Algebraic Universal Ssum Sprod 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 -> udom"
fixes prj :: "udom -> 'a::pcpo"
assumes ep_pair_emb_prj: "ep_pair emb prj"


interpretation rep!:
pcpo_ep_pair
"emb :: 'a::rep -> udom"
"prj :: udom -> '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

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 => TypeRep"
where
"Rep_of TYPE('a::rep) =
(\<Squnion>i. alg_defl_principal (Abs_fin_defl
(emb oo (approx i :: 'a -> 'a) oo prj)))"


syntax "_REP" :: "type => TypeRep" ("(1REP/(1'(_')))")
translations "REP('t)" \<rightleftharpoons> "CONST Rep_of TYPE('t)"

lemma cast_REP:
"cast·REP('a::rep) = (emb::'a -> udom) oo (prj::udom -> 'a)"

proof -
let ?a = "λi. emb oo approx i oo (prj::udom -> 'a)"
have a: "!!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·((prj·x)::'a::rep) = cast·REP('a)·x"
by (simp add: cast_REP)

lemma in_REP_iff:
"x ::: REP('a::rep) <-> emb·((prj·x)::'a) = x"

by (simp add: in_deflation_def cast_REP)

lemma prj_inverse:
"x ::: REP('a::rep) ==> emb·((prj·x)::'a) = x"

by (simp only: in_REP_iff)

lemma emb_in_REP [simp]:
"emb·(x::'a::rep) ::: REP('a)"

by (simp add: in_REP_iff)

subsection {* Coerce operator *}

definition coerce :: "'a -> 'b"
where "coerce = prj oo emb"


lemma beta_coerce: "coerce·x = prj·(emb·x)"
by (simp add: coerce_def)

lemma prj_emb: "prj·(emb·x) = coerce·x"
by (simp add: coerce_def)

lemma coerce_strict [simp]: "coerce·⊥ = ⊥"
by (simp add: coerce_def)

lemma coerce_eq_ID [simp]: "(coerce :: 'a -> 'a) = ID"
by (rule ext_cfun, simp add: beta_coerce)

lemma emb_coerce:
"REP('a) \<sqsubseteq> REP('b)
==> emb·((coerce::'a -> 'b)·x) = emb·x"

apply (simp add: beta_coerce)
apply (rule prj_inverse)
apply (erule subdeflationD)
apply (rule emb_in_REP)
done

lemma coerce_prj:
"REP('a) \<sqsubseteq> REP('b)
==> (coerce::'b -> 'a)·(prj·x) = prj·x"

apply (simp add: coerce_def)
apply (rule emb_eq_iff [THEN iffD1])
apply (simp only: emb_prj)
apply (rule deflation_below_comp1)
apply (rule deflation_cast)
apply (rule deflation_cast)
apply (erule monofun_cfun_arg)
done

lemma coerce_coerce [simp]:
"REP('a) \<sqsubseteq> REP('b)
==> coerce·((coerce::'a -> 'b)·x) = coerce·x"

by (simp add: beta_coerce prj_inverse subdeflationD)

lemma coerce_inverse:
"emb·(x::'a) ::: REP('b) ==> coerce·(coerce·x :: 'b) = x"

by (simp only: beta_coerce prj_inverse emb_inverse)

lemma coerce_type:
"REP('a) \<sqsubseteq> REP('b)
==> emb·((coerce::'a -> 'b)·x) ::: REP('a)"

by (simp add: beta_coerce prj_inverse subdeflationD)

lemma ep_pair_coerce:
"REP('a) \<sqsubseteq> REP('b)
==> ep_pair (coerce::'a -> 'b) (coerce::'b -> 'a)"

apply (rule ep_pair.intro)
apply simp
apply (simp only: beta_coerce)
apply (rule below_trans)
apply (rule monofun_cfun_arg)
apply (rule emb_prj_below)
apply simp
done

text {* Isomorphism lemmas used internally by the domain package: *}

lemma domain_abs_iso:
fixes abs and rep
assumes REP: "REP('b) = REP('a)"
assumes abs_def: "abs ≡ (coerce :: 'a -> 'b)"
assumes rep_def: "rep ≡ (coerce :: 'b -> 'a)"
shows "rep·(abs·x) = x"

unfolding abs_def rep_def by (simp add: REP)

lemma domain_rep_iso:
fixes abs and rep
assumes REP: "REP('b) = REP('a)"
assumes abs_def: "abs ≡ (coerce :: 'a -> 'b)"
assumes rep_def: "rep ≡ (coerce :: 'b -> 'a)"
shows "abs·(rep·x) = x"

unfolding abs_def rep_def by (simp add: REP [symmetric])


subsection {* Proving a subtype is representable *}

text {*
Temporarily relax type constraints for @{term "approx"},
@{term emb}, and @{term prj}.
*}


setup {* Sign.add_const_constraint
(@{const_name "approx"}, SOME @{typ "nat => 'a::cpo -> 'a"}) *}


setup {* Sign.add_const_constraint
(@{const_name emb}, SOME @{typ "'a::pcpo -> udom"}) *}


setup {* Sign.add_const_constraint
(@{const_name prj}, SOME @{typ "udom -> 'a::pcpo"}) *}


definition
repdef_approx ::
"('a::pcpo => udom) => (udom => 'a) => udom alg_defl => nat => 'a -> 'a"
where
"repdef_approx Rep Abs t = (λi. Λ x. Abs (cast·(approx i·t)·(Rep x)))"


lemma typedef_rep_class:
fixes Rep :: "'a::pcpo => udom"
fixes Abs :: "udom => 'a::pcpo"
fixes t :: TypeRep
assumes type: "type_definition Rep Abs {x. x ::: t}"
assumes below: "op \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> Rep y"
assumes emb: "emb ≡ (Λ x. Rep x)"
assumes prj: "prj ≡ (Λ x. Abs (cast·t·x))"
assumes approx: "(approx :: nat => 'a -> 'a) ≡ repdef_approx Rep Abs t"
shows "OFCLASS('a, rep_class)"

proof
have adm: "adm (λx. x ∈ {x. x ::: t})"
by (simp add: adm_in_deflation)
have emb_beta: "!!x. emb·x = Rep x"
unfolding emb
apply (rule beta_cfun)
apply (rule typedef_cont_Rep [OF type below adm])
done
have prj_beta: "!!y. prj·y = Abs (cast·t·y)"
unfolding prj
apply (rule beta_cfun)
apply (rule typedef_cont_Abs [OF type below adm])
apply simp_all
done
have cast_cast_approx:
"!!i x. cast·t·(cast·(approx i·t)·x) = cast·(approx i·t)·x"

apply (rule cast_fixed)
apply (rule subdeflationD)
apply (rule approx.below)
apply (rule cast_in_deflation)
done
have approx':
"approx = (λi. Λ(x::'a). prj·(cast·(approx i·t)·(emb·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: "!!x::'a. emb·x ::: t"
using type_definition.Rep [OF type]
by (simp add: emb_beta)
have prj_emb: "!!x::'a. prj·(emb·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: "!!y. emb·(prj·y :: 'a) = cast·t·y"
unfolding prj_beta emb_beta
by (simp add: type_definition.Abs_inverse [OF type])
show "ep_pair (emb :: 'a -> udom) prj"
apply default
apply (simp add: prj_emb)
apply (simp add: emb_prj cast.below)
done
show "chain (approx :: nat => 'a -> 'a)"
unfolding approx' by simp
show "!!x::'a. (\<Squnion>i. approx i·x) = x"
unfolding approx'
apply (simp add: lub_distribs)
apply (subst cast_fixed [OF emb_in_deflation])
apply (rule prj_emb)
done
show "!!(i::nat) (x::'a). approx i·(approx i·x) = approx i·x"
unfolding approx'
apply simp
apply (simp add: emb_prj)
apply (simp add: cast_cast_approx)
done
show "!!i::nat. finite {x::'a. approx i·x = x}"
apply (rule_tac B="(λx. prj·x) ` {x. cast·(approx i·t)·x = x}"
in finite_subset)

apply (clarsimp simp add: approx')
apply (drule_tac f="λx. emb·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 => 'a::profinite -> 'a"}) *}


setup {* Sign.add_const_constraint
(@{const_name emb}, SOME @{typ "'a::rep -> udom"}) *}


setup {* Sign.add_const_constraint
(@{const_name prj}, SOME @{typ "udom -> 'a::rep"}) *}


lemma typedef_REP:
fixes Rep :: "'a::rep => udom"
fixes Abs :: "udom => 'a::rep"
fixes t :: TypeRep
assumes type: "type_definition Rep Abs {x. x ::: t}"
assumes below: "op \<sqsubseteq> ≡ λx y. Rep x \<sqsubseteq> Rep y"
assumes emb: "emb ≡ (Λ x. Rep x)"
assumes prj: "prj ≡ (Λ x. Abs (cast·t·x))"
shows "REP('a) = t"

proof -
have adm: "adm (λx. x ∈ {x. x ::: t})"
by (simp add: adm_in_deflation)
have emb_beta: "!!x. emb·x = Rep x"
unfolding emb
apply (rule beta_cfun)
apply (rule typedef_cont_Rep [OF type below adm])
done
have prj_beta: "!!y. prj·y = Abs (cast·t·y)"
unfolding prj
apply (rule beta_cfun)
apply (rule typedef_cont_Abs [OF type below adm])
apply simp_all
done
have emb_in_deflation: "!!x::'a. emb·x ::: t"
using type_definition.Rep [OF type]
by (simp add: emb_beta)
have prj_emb: "!!x::'a. prj·(emb·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: "!!y. emb·(prj·y :: 'a) = cast·t·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)
done
qed

lemma adm_mem_Collect_in_deflation: "adm (λx. x ∈ {x. x ::: A})"
unfolding mem_Collect_eq by (rule adm_in_deflation)

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 -> udom)"
definition prj_udom_def [simp]: "prj = (ID :: udom -> 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 (∃x::'a::countable. n = to_nat x)
then Def (THE x::'a. n = to_nat x) else ⊥) 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·prj·emb"
definition prj_cfun_def: "prj = cfun_map·emb·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·emb·emb"
definition prj_sprod_def: "prj = sprod_map·prj·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·emb·emb"
definition prj_ssum_def: "prj = ssum_map·prj·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·emb"
definition prj_u_def: "prj = u_map·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 "*" :: (rep, rep) rep
begin


definition emb_cprod_def: "emb = udom_emb oo cprod_map·emb·emb"
definition prj_cprod_def: "prj = cprod_map·prj·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 -> udom) -> ('a -> 'a))
=> (TypeRep -> TypeRep)"

where
"TypeRep_fun1 f =
alg_defl.basis_fun (λa.
alg_defl_principal (
Abs_fin_defl (udom_emb oo f·(Rep_fin_defl a) oo udom_prj)))"


definition
TypeRep_fun2 ::
"((udom -> udom) -> (udom -> udom) -> ('a -> 'a))
=> (TypeRep -> TypeRep -> TypeRep)"

where
"TypeRep_fun2 f =
alg_defl.basis_fun (λa.
alg_defl.basis_fun (λb.
alg_defl_principal (
Abs_fin_defl (udom_emb oo
f·(Rep_fin_defl a)·(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 ==> Rep_fin_defl a \<sqsubseteq> Rep_fin_defl b"
unfolding below_fin_defl_def .

lemma cast_TypeRep_fun1:
assumes f: "!!a. finite_deflation a ==> finite_deflation (f·a)"
shows "cast·(TypeRep_fun1 f·A) = udom_emb oo f·(cast·A) oo udom_prj"

proof -
have 1: "!!a. finite_deflation (udom_emb oo f·(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: "!!a b. finite_deflation a ==> finite_deflation b ==>
finite_deflation (f·a·b)"

shows "cast·(TypeRep_fun2 f·A·B) =
udom_emb oo f·(cast·A)·(cast·B) oo udom_prj"

proof -
have 1: "!!a b. finite_deflation
(udom_emb oo f·(Rep_fin_defl a)·(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·(cfun_defl·A·B) = udom_emb oo cfun_map·(cast·A)·(cast·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·(ssum_defl·A·B) = udom_emb oo ssum_map·(cast·A)·(cast·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·(sprod_defl·A·B) = udom_emb oo sprod_map·(cast·A)·(cast·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·(cprod_defl·A·B) = udom_emb oo cprod_map·(cast·A)·(cast·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·(u_defl·A) = udom_emb oo u_map·(cast·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 -> 'b) = cfun_defl·REP('a)·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 ⊕ 'b) = ssum_defl·REP('a)·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 ⊗ 'b) = sprod_defl·REP('a)·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 × 'b) = cprod_defl·REP('a)·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·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 -> 'a) => udom alg_defl => bool"
where
"isodefl d t <-> cast·t = emb oo d oo prj"


lemma isodeflI: "(!!x. cast·t·x = emb·(d·(prj·x))) ==> isodefl d t"
unfolding isodefl_def by (simp add: ext_cfun)

lemma cast_isodefl: "isodefl d t ==> cast·t = (Λ x. emb·(d·(prj·x)))"
unfolding isodefl_def by (simp add: ext_cfun)

lemma isodefl_strict: "isodefl d t ==> d·⊥ = ⊥"
unfolding isodefl_def
by (drule cfun_fun_cong [where x="⊥"], simp)

lemma isodefl_imp_deflation:
fixes d :: "'a::rep -> 'a"
assumes "isodefl d t" shows "deflation d"

proof
note prems [unfolded isodefl_def, simp]
fix x :: 'a
show "d·(d·x) = d·x"
using cast.idem [of t "emb·x"] by simp
show "d·x \<sqsubseteq> x"
using cast.below [of t "emb·x"] by simp
qed

lemma isodefl_ID_REP: "isodefl (ID :: 'a -> 'a) REP('a)"
unfolding isodefl_def by (simp add: cast_REP)

lemma isodefl_REP_imp_ID: "isodefl (d :: 'a -> 'a) REP('a) ==> d = ID"
unfolding isodefl_def
apply (simp add: cast_REP)
apply (simp add: expand_cfun_eq)
apply (rule allI)
apply (drule_tac x="emb·x" in spec)
apply simp
done

lemma isodefl_bottom: "isodefl ⊥ ⊥"
unfolding isodefl_def by (simp add: expand_cfun_eq)

lemma adm_isodefl:
"cont f ==> cont g ==> adm (λx. isodefl (f x) (g x))"

unfolding isodefl_def by simp

lemma isodefl_lub:
assumes "chain d" and "chain t"
assumes "!!i. isodefl (d i) (t i)"
shows "isodefl (\<Squnion>i. d i) (\<Squnion>i. t i)"

using prems unfolding isodefl_def
by (simp add: contlub_cfun_arg contlub_cfun_fun)

lemma isodefl_fix:
assumes "!!d t. isodefl d t ==> isodefl (f·d) (g·t)"
shows "isodefl (fix·f) (fix·g)"

unfolding fix_def2
apply (rule isodefl_lub, simp, simp)
apply (induct_tac i)
apply (simp add: isodefl_bottom)
apply (simp add: prems)
done

lemma isodefl_coerce:
fixes d :: "'a -> 'a"
assumes REP: "REP('b) = REP('a)"
shows "isodefl d t ==> isodefl (coerce oo d oo coerce :: 'b -> 'b) t"

unfolding isodefl_def
apply (simp add: expand_cfun_eq)
apply (simp add: emb_coerce coerce_prj REP)
done

lemma isodefl_abs_rep:
fixes abs and rep and d
assumes REP: "REP('b) = REP('a)"
assumes abs_def: "abs ≡ (coerce :: 'a -> 'b)"
assumes rep_def: "rep ≡ (coerce :: 'b -> 'a)"
shows "isodefl d t ==> isodefl (abs oo d oo rep) t"

unfolding abs_def rep_def using REP by (rule isodefl_coerce)

lemma isodefl_cfun:
"isodefl d1 t1 ==> isodefl d2 t2 ==>
isodefl (cfun_map·d1·d2) (cfun_defl·t1·t2)"

apply (rule isodeflI)
apply (simp add: cast_cfun_defl 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 ==> isodefl d2 t2 ==>
isodefl (ssum_map·d1·d2) (ssum_defl·t1·t2)"

apply (rule isodeflI)
apply (simp add: cast_ssum_defl 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 ==> isodefl d2 t2 ==>
isodefl (sprod_map·d1·d2) (sprod_defl·t1·t2)"

apply (rule isodeflI)
apply (simp add: cast_sprod_defl 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 ==> isodefl d2 t2 ==>
isodefl (cprod_map·d1·d2) (cprod_defl·t1·t2)"

apply (rule isodeflI)
apply (simp add: cast_cprod_defl cast_isodefl)
apply (simp add: emb_cprod_def prj_cprod_def)
apply (simp add: cprod_map_map cfcomp1)
done

lemma isodefl_u:
"isodefl d t ==> isodefl (u_map·d) (u_defl·t)"

apply (rule isodeflI)
apply (simp add: cast_u_defl cast_isodefl)
apply (simp add: emb_u_def prj_u_def)
apply (simp add: u_map_map)
done

subsection {* Constructing Domain Isomorphisms *}

use "Tools/Domain/domain_isomorphism.ML"

setup {*
fold Domain_Isomorphism.add_type_constructor
[(@{type_name cfun}, @{term cfun_defl}, @{const_name cfun_map}, @{thm REP_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},
@{thm isodefl_ssum}, @{thm ssum_map_ID}, @{thm deflation_ssum_map}),

(@{type_name sprod}, @{term sprod_defl}, @{const_name sprod_map}, @{thm REP_sprod},
@{thm isodefl_sprod}, @{thm sprod_map_ID}, @{thm deflation_sprod_map}),

(@{type_name "*"}, @{term cprod_defl}, @{const_name cprod_map}, @{thm REP_cprod},
@{thm isodefl_cprod}, @{thm cprod_map_ID}, @{thm deflation_cprod_map}),

(@{type_name "u"}, @{term u_defl}, @{const_name u_map}, @{thm REP_up},
@{thm isodefl_u}, @{thm u_map_ID}, @{thm deflation_u_map})]
*}


end