# HG changeset patch # User huffman # Date 1289397757 28800 # Node ID c45a3f8a86ead7ad94d7846d935e8f18e3b08b8e # Parent e380754e8a09d9dd639ccdeb19be20884e7a22c1 instance prod :: (predomain, predomain) predomain diff -r e380754e8a09 -r c45a3f8a86ea src/HOLCF/Bifinite.thy --- a/src/HOLCF/Bifinite.thy Tue Nov 09 19:37:11 2010 -0800 +++ b/src/HOLCF/Bifinite.thy Wed Nov 10 06:02:37 2010 -0800 @@ -37,8 +37,7 @@ syntax "_DEFL" :: "type \ defl" ("(1DEFL/(1'(_')))") translations "DEFL('t)" \ "CONST defl TYPE('t)" -interpretation bifinite: - pcpo_ep_pair "emb :: 'a::bifinite \ udom" "prj :: udom \ 'a::bifinite" +interpretation bifinite: pcpo_ep_pair emb prj unfolding pcpo_ep_pair_def by (rule ep_pair_emb_prj) @@ -430,6 +429,57 @@ subsection {* Cartesian product is a bifinite domain *} +text {* + Types @{typ "('a * 'b) u"} and @{typ "'a u \ 'b u"} are isomorphic. +*} + +definition + "encode_prod_u = (\(up\(x, y)). (:up\x, up\y:))" + +definition + "decode_prod_u = (\(:up\x, up\y:). up\(x, y))" + +lemma decode_encode_prod_u [simp]: "decode_prod_u\(encode_prod_u\x) = x" +unfolding encode_prod_u_def decode_prod_u_def +by (case_tac x, simp, rename_tac y, case_tac y, simp) + +lemma encode_decode_prod_u [simp]: "encode_prod_u\(decode_prod_u\y) = y" +unfolding encode_prod_u_def decode_prod_u_def +apply (case_tac y, simp, rename_tac a b) +apply (case_tac a, simp, case_tac b, simp, simp) +done + +instantiation prod :: (predomain, predomain) predomain +begin + +definition + "liftemb = + (udom_emb sprod_approx oo sprod_map\emb\emb) oo encode_prod_u" + +definition + "liftprj = + decode_prod_u oo (sprod_map\prj\prj oo udom_prj sprod_approx)" + +definition + "liftdefl (t::('a \ 'b) itself) = sprod_defl\DEFL('a u)\DEFL('b u)" + +instance proof + have "ep_pair encode_prod_u decode_prod_u" + by (rule ep_pair.intro, simp_all) + thus "ep_pair liftemb (liftprj :: udom \ ('a \ 'b) u)" + unfolding liftemb_prod_def liftprj_prod_def + apply (rule ep_pair_comp) + apply (rule ep_pair_comp) + apply (intro ep_pair_sprod_map ep_pair_emb_prj) + apply (rule ep_pair_udom [OF sprod_approx]) + done + show "cast\LIFTDEFL('a \ 'b) = liftemb oo (liftprj :: udom \ ('a \ 'b) u)" + unfolding liftemb_prod_def liftprj_prod_def liftdefl_prod_def + by (simp add: cast_sprod_defl cast_DEFL cfcomp1 sprod_map_map) +qed + +end + instantiation prod :: (bifinite, bifinite) bifinite begin @@ -442,18 +492,7 @@ definition "defl (t::('a \ 'b) itself) = prod_defl\DEFL('a)\DEFL('b)" -definition - "(liftemb :: ('a \ 'b) u \ udom) = udom_emb u_approx oo u_map\emb" - -definition - "(liftprj :: udom \ ('a \ 'b) u) = u_map\prj oo udom_prj u_approx" - -definition - "liftdefl (t::('a \ 'b) itself) = u_defl\DEFL('a \ 'b)" - -instance -using liftemb_prod_def liftprj_prod_def liftdefl_prod_def -proof (rule bifinite_class_intro) +instance proof show "ep_pair emb (prj :: udom \ 'a \ 'b)" unfolding emb_prod_def prj_prod_def using ep_pair_udom [OF prod_approx] @@ -471,7 +510,7 @@ by (rule defl_prod_def) lemma LIFTDEFL_prod: - "LIFTDEFL('a::bifinite \ 'b::bifinite) = u_defl\DEFL('a \ 'b)" + "LIFTDEFL('a::predomain \ 'b::predomain) = sprod_defl\DEFL('a u)\DEFL('b u)" by (rule liftdefl_prod_def) subsection {* Strict product is a bifinite domain *} diff -r e380754e8a09 -r c45a3f8a86ea src/HOLCF/Representable.thy --- a/src/HOLCF/Representable.thy Tue Nov 09 19:37:11 2010 -0800 +++ b/src/HOLCF/Representable.thy Wed Nov 10 06:02:37 2010 -0800 @@ -298,11 +298,22 @@ using liftemb_cfun_def liftprj_cfun_def isodefl_cfun [OF assms] by (rule isodefl_u) +lemma encode_prod_u_map: + "encode_prod_u\(u_map\(cprod_map\f\g)\(decode_prod_u\x)) + = sprod_map\(u_map\f)\(u_map\g)\x" +unfolding encode_prod_u_def decode_prod_u_def +apply (case_tac x, simp, rename_tac a b) +apply (case_tac a, simp, case_tac b, simp, simp) +done + lemma isodefl_cprod_u: - assumes "isodefl d1 t1" and "isodefl d2 t2" - shows "isodefl (u_map\(cprod_map\d1\d2)) (u_defl\(prod_defl\t1\t2))" -using liftemb_prod_def liftprj_prod_def isodefl_cprod [OF assms] -by (rule isodefl_u) + assumes "isodefl (u_map\d1) t1" and "isodefl (u_map\d2) t2" + shows "isodefl (u_map\(cprod_map\d1\d2)) (sprod_defl\t1\t2)" +using assms unfolding isodefl_def +apply (simp add: emb_u_def prj_u_def liftemb_prod_def liftprj_prod_def) +apply (simp add: emb_u_def [symmetric] prj_u_def [symmetric]) +apply (simp add: cfcomp1 encode_prod_u_map cast_sprod_defl sprod_map_map) +done lemma isodefl_sprod_u: assumes "isodefl d1 t1" and "isodefl d2 t2"