# HG changeset patch # User huffman # Date 1257862949 28800 # Node ID ea9becc59636306a8e45e2dd301a78d6e4ec73f6 # Parent 54f98d22516398f664a8e7f2e4c3cb3a42e22294 theory of representable cpos diff -r 54f98d225163 -r ea9becc59636 src/HOLCF/HOLCF.thy --- a/src/HOLCF/HOLCF.thy Mon Nov 09 15:51:32 2009 -0800 +++ b/src/HOLCF/HOLCF.thy Tue Nov 10 06:22:29 2009 -0800 @@ -7,6 +7,7 @@ theory HOLCF imports Domain ConvexPD Algebraic Universal Sum_Cpo Main + Representable uses "holcf_logic.ML" "Tools/adm_tac.ML" diff -r 54f98d225163 -r ea9becc59636 src/HOLCF/Representable.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/Representable.thy Tue Nov 10 06:22:29 2009 -0800 @@ -0,0 +1,991 @@ +header {* Representable Types *} + +theory Representable +imports Algebraic Universal Ssum Sprod One ConvexPD +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 @{term rep} the default class *} + +text {* + From now on, free type variables are assumed to be in class + @{term rep}, unless specified otherwise. +*} + +defaultsort rep + +subsection {* Representations of types *} + +text "A TypeRep is an algebraic deflation over the universe of values." + +types TypeRep = "udom alg_defl" +translations "TypeRep" \ (type) "udom alg_defl" + +definition + Rep_of :: "'a::rep itself \ TypeRep" +where + "Rep_of TYPE('a::rep) = + (\i. alg_defl_principal (Abs_fin_defl + (emb oo (approx i :: 'a \ 'a) oo prj)))" + +syntax "_REP" :: "type \ TypeRep" ("(1REP/(1'(_')))") +translations "REP(t)" \ "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) \ 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) \ 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) \ 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) \ REP('b) + \ emb\((coerce::'a \ 'b)\x) ::: REP('a)" +by (simp add: beta_coerce prj_inverse subdeflationD) + +lemma ep_pair_coerce: + "REP('a) \ 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 + +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"}) *} + +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 \ \ \x y. Rep x \ Rep y" + assumes emb: "emb = (\ x. Rep x)" + assumes prj: "prj = (\ x. Abs (cast\t\x))" + assumes approx: + "(approx :: nat \ 'a \ 'a) = (\i. prj oo cast\(approx i\t) oo emb)" + 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 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. (\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 + 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 + 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 \ \ \x y. Rep x \ 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 + + +subsection {* Instances of class @{text 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 "->" :: (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 "**" :: (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 "++" :: (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 + +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\emb" +definition prj_upper_pd_def: "prj = upper_map\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\emb" +definition prj_lower_pd_def: "prj = lower_map\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 +begin + +definition emb_convex_pd_def: "emb = udom_emb oo convex_map\emb" +definition prj_convex_pd_def: "prj = convex_map\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 \ 'b::cpo" + assumes "finite ((\x. d\x) ` {x. compact x})" + shows "finite (range (\x. d\x))" +proof (rule finite_subset [OF _ prems]) + { + fix x :: 'a + have "range (\i. d\(approx i\x)) \ (\x. d\x) ` {x. compact x}" + by auto + hence "finite (range (\i. d\(approx i\x)))" + using prems by (rule finite_subset) + hence "finite_chain (\i. d\(approx i\x))" + by (simp add: finite_range_imp_finch) + hence "\i. (\i. d\(approx i\x)) = d\(approx i\x)" + by (simp add: finite_chain_def maxinch_is_thelub) + hence "\i. d\x = d\(approx i\x)" + by (simp add: lub_distribs) + hence "d\x \ (\x. d\x) ` {x. compact x}" + by auto + } + thus "range (\x. d\x) \ (\x. d\x) ` {x. compact x}" + by clarsimp +qed + +lemma finite_deflation_upper_map: + assumes "finite_deflation d" shows "finite_deflation (upper_map\d)" +proof (intro finite_deflation.intro finite_deflation_axioms.intro) + interpret d: finite_deflation d by fact + have "deflation d" by fact + thus "deflation (upper_map\d)" by (rule deflation_upper_map) + have "finite (range (\x. d\x))" by (rule d.finite_range) + hence "finite (Rep_compact_basis -` range (\x. d\x))" + by (rule finite_vimageI, simp add: inj_on_def Rep_compact_basis_inject) + hence "finite (Pow (Rep_compact_basis -` range (\x. d\x)))" by simp + hence "finite (Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\x. d\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 (\x. d\x))))" by simp + hence "finite ((\xs. upper_map\d\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 "\b. d\(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 ((\xs. upper_map\d\xs) ` {xs::'a upper_pd. compact xs})" + by simp + hence "finite (range (\xs. upper_map\d\xs))" + by (rule finite_compact_range_imp_finite_range) + thus "finite {xs. upper_map\d\xs = xs}" + by (rule finite_range_imp_finite_fixes) +qed + +lemma finite_deflation_lower_map: + assumes "finite_deflation d" shows "finite_deflation (lower_map\d)" +proof (intro finite_deflation.intro finite_deflation_axioms.intro) + interpret d: finite_deflation d by fact + have "deflation d" by fact + thus "deflation (lower_map\d)" by (rule deflation_lower_map) + have "finite (range (\x. d\x))" by (rule d.finite_range) + hence "finite (Rep_compact_basis -` range (\x. d\x))" + by (rule finite_vimageI, simp add: inj_on_def Rep_compact_basis_inject) + hence "finite (Pow (Rep_compact_basis -` range (\x. d\x)))" by simp + hence "finite (Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\x. d\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 (\x. d\x))))" by simp + hence "finite ((\xs. lower_map\d\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 "\b. d\(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 ((\xs. lower_map\d\xs) ` {xs::'a lower_pd. compact xs})" + by simp + hence "finite (range (\xs. lower_map\d\xs))" + by (rule finite_compact_range_imp_finite_range) + thus "finite {xs. lower_map\d\xs = xs}" + by (rule finite_range_imp_finite_fixes) +qed + +lemma finite_deflation_convex_map: + assumes "finite_deflation d" shows "finite_deflation (convex_map\d)" +proof (intro finite_deflation.intro finite_deflation_axioms.intro) + interpret d: finite_deflation d by fact + have "deflation d" by fact + thus "deflation (convex_map\d)" by (rule deflation_convex_map) + have "finite (range (\x. d\x))" by (rule d.finite_range) + hence "finite (Rep_compact_basis -` range (\x. d\x))" + by (rule finite_vimageI, simp add: inj_on_def Rep_compact_basis_inject) + hence "finite (Pow (Rep_compact_basis -` range (\x. d\x)))" by simp + hence "finite (Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\x. d\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 (\x. d\x))))" by simp + hence "finite ((\xs. convex_map\d\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 "\b. d\(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 ((\xs. convex_map\d\xs) ` {xs::'a convex_pd. compact xs})" + by simp + hence "finite (range (\xs. convex_map\d\xs))" + by (rule finite_compact_range_imp_finite_range) + thus "finite {xs. convex_map\d\xs = xs}" + by (rule finite_range_imp_finite_fixes) +qed + +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 "one_typ = REP(one)" +definition "tr_typ = REP(tr)" +definition "cfun_typ = TypeRep_fun2 cfun_map" +definition "ssum_typ = TypeRep_fun2 ssum_map" +definition "sprod_typ = TypeRep_fun2 sprod_map" +definition "cprod_typ = TypeRep_fun2 cprod_map" +definition "u_typ = TypeRep_fun1 u_map" +definition "upper_typ = TypeRep_fun1 upper_map" +definition "lower_typ = TypeRep_fun1 lower_map" +definition "convex_typ = TypeRep_fun1 convex_map" + +lemma Rep_fin_defl_mono: "a \ b \ Rep_fin_defl a \ 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_typ: + "cast\(cfun_typ\A\B) = udom_emb oo cfun_map\(cast\A)\(cast\B) oo udom_prj" +unfolding cfun_typ_def +apply (rule cast_TypeRep_fun2) +apply (erule (1) finite_deflation_cfun_map) +done + +lemma cast_ssum_typ: + "cast\(ssum_typ\A\B) = udom_emb oo ssum_map\(cast\A)\(cast\B) oo udom_prj" +unfolding ssum_typ_def +apply (rule cast_TypeRep_fun2) +apply (erule (1) finite_deflation_ssum_map) +done + +lemma cast_sprod_typ: + "cast\(sprod_typ\A\B) = udom_emb oo sprod_map\(cast\A)\(cast\B) oo udom_prj" +unfolding sprod_typ_def +apply (rule cast_TypeRep_fun2) +apply (erule (1) finite_deflation_sprod_map) +done + +lemma cast_cprod_typ: + "cast\(cprod_typ\A\B) = udom_emb oo cprod_map\(cast\A)\(cast\B) oo udom_prj" +unfolding cprod_typ_def +apply (rule cast_TypeRep_fun2) +apply (erule (1) finite_deflation_cprod_map) +done + +lemma cast_u_typ: + "cast\(u_typ\A) = udom_emb oo u_map\(cast\A) oo udom_prj" +unfolding u_typ_def +apply (rule cast_TypeRep_fun1) +apply (erule finite_deflation_u_map) +done + +lemma cast_upper_typ: + "cast\(upper_typ\A) = udom_emb oo upper_map\(cast\A) oo udom_prj" +unfolding upper_typ_def +apply (rule cast_TypeRep_fun1) +apply (erule finite_deflation_upper_map) +done + +lemma cast_lower_typ: + "cast\(lower_typ\A) = udom_emb oo lower_map\(cast\A) oo udom_prj" +unfolding lower_typ_def +apply (rule cast_TypeRep_fun1) +apply (erule finite_deflation_lower_map) +done + +lemma cast_convex_typ: + "cast\(convex_typ\A) = udom_emb oo convex_map\(cast\A) oo udom_prj" +unfolding convex_typ_def +apply (rule cast_TypeRep_fun1) +apply (erule finite_deflation_convex_map) +done + +text {* REP of type constructor = type combinator *} + +lemma REP_one: "REP(one) = one_typ" +by (simp only: one_typ_def) + +lemma REP_tr: "REP(tr) = tr_typ" +by (simp only: tr_typ_def) + +lemma REP_cfun: "REP('a \ 'b) = cfun_typ\REP('a)\REP('b)" +apply (rule cast_eq_imp_eq, rule ext_cfun) +apply (simp add: cast_REP cast_cfun_typ) +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_typ\REP('a)\REP('b)" +apply (rule cast_eq_imp_eq, rule ext_cfun) +apply (simp add: cast_REP cast_ssum_typ) +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_typ\REP('a)\REP('b)" +apply (rule cast_eq_imp_eq, rule ext_cfun) +apply (simp add: cast_REP cast_sprod_typ) +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_typ\REP('a)\REP('b)" +apply (rule cast_eq_imp_eq, rule ext_cfun) +apply (simp add: cast_REP cast_cprod_typ) +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_typ\REP('a)" +apply (rule cast_eq_imp_eq, rule ext_cfun) +apply (simp add: cast_REP cast_u_typ) +apply (simp add: prj_u_def) +apply (simp add: emb_u_def) +apply (simp add: u_map_map cfcomp1) +done + +lemma REP_upper: "REP('a upper_pd) = upper_typ\REP('a)" +apply (rule cast_eq_imp_eq, rule ext_cfun) +apply (simp add: cast_REP cast_upper_typ) +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_typ\REP('a)" +apply (rule cast_eq_imp_eq, rule ext_cfun) +apply (simp add: cast_REP cast_lower_typ) +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_typ\REP('a)" +apply (rule cast_eq_imp_eq, rule ext_cfun) +apply (simp add: cast_REP cast_convex_typ) +apply (simp add: prj_convex_pd_def) +apply (simp add: emb_convex_pd_def) +apply (simp add: convex_map_map cfcomp1) +done + +lemmas REP_simps = + REP_one + REP_tr + REP_cfun + REP_ssum + REP_sprod + REP_cprod + REP_up + REP_upper + REP_lower + REP_convex + +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 \ 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 (\i. d i) (\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_cfun: + "isodefl d1 t1 \ isodefl d2 t2 \ + isodefl (cfun_map\d1\d2) (cfun_typ\t1\t2)" +apply (rule isodeflI) +apply (simp add: cast_cfun_typ 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_typ\t1\t2)" +apply (rule isodeflI) +apply (simp add: cast_ssum_typ 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_typ\t1\t2)" +apply (rule isodeflI) +apply (simp add: cast_sprod_typ cast_isodefl) +apply (simp add: emb_sprod_def prj_sprod_def) +apply (simp add: sprod_map_map isodefl_strict) +done + +lemma isodefl_u: + "isodefl d t \ isodefl (u_map\d) (u_typ\t)" +apply (rule isodeflI) +apply (simp add: cast_u_typ cast_isodefl) +apply (simp add: emb_u_def prj_u_def) +apply (simp add: u_map_map) +done + +lemma isodefl_one: "isodefl (ID :: one \ one) one_typ" +unfolding one_typ_def by (rule isodefl_ID_REP) + +lemma isodefl_tr: "isodefl (ID :: tr \ tr) tr_typ" +unfolding tr_typ_def by (rule isodefl_ID_REP) + +lemma isodefl_upper: + "isodefl d t \ isodefl (upper_map\d) (upper_typ\t)" +apply (rule isodeflI) +apply (simp add: cast_upper_typ 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 \ isodefl (lower_map\d) (lower_typ\t)" +apply (rule isodeflI) +apply (simp add: cast_lower_typ 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 \ isodefl (convex_map\d) (convex_typ\t)" +apply (rule isodeflI) +apply (simp add: cast_convex_typ cast_isodefl) +apply (simp add: emb_convex_pd_def prj_convex_pd_def) +apply (simp add: convex_map_map) +done + +end