# HG changeset patch # User huffman # Date 1289411948 28800 # Node ID 71283f31a27ff2378c80a637f8862a4b4d5c379d # Parent 2a92d3f5026c016c5fcb9bec31d0c9d2c4e08e54 instance sum :: (predomain, predomain) predomain diff -r 2a92d3f5026c -r 71283f31a27f src/HOLCF/Library/Sum_Cpo.thy --- a/src/HOLCF/Library/Sum_Cpo.thy Wed Nov 10 09:52:50 2010 -0800 +++ b/src/HOLCF/Library/Sum_Cpo.thy Wed Nov 10 09:59:08 2010 -0800 @@ -240,4 +240,53 @@ (@{const_name Inr}, @{const_name match_Inr}) ] *} +subsection {* Disjoint sum is a predomain *} + +definition + "encode_sum_u = + (\(up\x). case x of Inl a \ sinl\(up\a) | Inr b \ sinr\(up\b))" + +definition + "decode_sum_u = sscase\(\(up\a). up\(Inl a))\(\(up\b). up\(Inr b))" + +lemma decode_encode_sum_u [simp]: "decode_sum_u\(encode_sum_u\x) = x" +unfolding decode_sum_u_def encode_sum_u_def +by (case_tac x, simp, rename_tac y, case_tac y, simp_all) + +lemma encode_decode_sum_u [simp]: "encode_sum_u\(decode_sum_u\x) = x" +unfolding decode_sum_u_def encode_sum_u_def +apply (case_tac x, simp) +apply (rename_tac a, case_tac a, simp, simp) +apply (rename_tac b, case_tac b, simp, simp) +done + +instantiation sum :: (predomain, predomain) predomain +begin + +definition + "liftemb = (udom_emb ssum_approx oo ssum_map\emb\emb) oo encode_sum_u" + +definition + "liftprj = + decode_sum_u oo (ssum_map\prj\prj oo udom_prj ssum_approx)" + +definition + "liftdefl (t::('a + 'b) itself) = ssum_defl\DEFL('a u)\DEFL('b u)" + +instance proof + show "ep_pair liftemb (liftprj :: udom \ ('a + 'b) u)" + unfolding liftemb_sum_def liftprj_sum_def + apply (rule ep_pair_comp) + apply (rule ep_pair.intro, simp, simp) + apply (rule ep_pair_comp) + apply (intro ep_pair_ssum_map ep_pair_emb_prj) + apply (rule ep_pair_udom [OF ssum_approx]) + done + show "cast\LIFTDEFL('a + 'b) = liftemb oo (liftprj :: udom \ ('a + 'b) u)" + unfolding liftemb_sum_def liftprj_sum_def liftdefl_sum_def + by (simp add: cast_ssum_defl cast_DEFL cfcomp1 ssum_map_map) +qed + end + +end