src/HOLCF/Ssum.thy
changeset 40502 8e92772bc0e8
parent 40327 1dfdbd66093a
child 40767 a3e505b236e7
equal deleted inserted replaced
40501:2ed71459136e 40502:8e92772bc0e8
     1 (*  Title:      HOLCF/Ssum.thy
     1 (*  Title:      HOLCF/Ssum.thy
     2     Author:     Franz Regensburger and Brian Huffman
     2     Author:     Franz Regensburger
       
     3     Author:     Brian Huffman
     3 *)
     4 *)
     4 
     5 
     5 header {* The type of strict sums *}
     6 header {* The type of strict sums *}
     6 
     7 
     7 theory Ssum
     8 theory Ssum
   192 apply (case_tac x, simp)
   193 apply (case_tac x, simp)
   193 apply (case_tac y, simp_all add: flat_below_iff)
   194 apply (case_tac y, simp_all add: flat_below_iff)
   194 apply (case_tac y, simp_all add: flat_below_iff)
   195 apply (case_tac y, simp_all add: flat_below_iff)
   195 done
   196 done
   196 
   197 
   197 subsection {* Map function for strict sums *}
       
   198 
       
   199 definition
       
   200   ssum_map :: "('a \<rightarrow> 'b) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> 'a \<oplus> 'c \<rightarrow> 'b \<oplus> 'd"
       
   201 where
       
   202   "ssum_map = (\<Lambda> f g. sscase\<cdot>(sinl oo f)\<cdot>(sinr oo g))"
       
   203 
       
   204 lemma ssum_map_strict [simp]: "ssum_map\<cdot>f\<cdot>g\<cdot>\<bottom> = \<bottom>"
       
   205 unfolding ssum_map_def by simp
       
   206 
       
   207 lemma ssum_map_sinl [simp]: "x \<noteq> \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinl\<cdot>x) = sinl\<cdot>(f\<cdot>x)"
       
   208 unfolding ssum_map_def by simp
       
   209 
       
   210 lemma ssum_map_sinr [simp]: "x \<noteq> \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinr\<cdot>x) = sinr\<cdot>(g\<cdot>x)"
       
   211 unfolding ssum_map_def by simp
       
   212 
       
   213 lemma ssum_map_sinl': "f\<cdot>\<bottom> = \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinl\<cdot>x) = sinl\<cdot>(f\<cdot>x)"
       
   214 by (cases "x = \<bottom>") simp_all
       
   215 
       
   216 lemma ssum_map_sinr': "g\<cdot>\<bottom> = \<bottom> \<Longrightarrow> ssum_map\<cdot>f\<cdot>g\<cdot>(sinr\<cdot>x) = sinr\<cdot>(g\<cdot>x)"
       
   217 by (cases "x = \<bottom>") simp_all
       
   218 
       
   219 lemma ssum_map_ID: "ssum_map\<cdot>ID\<cdot>ID = ID"
       
   220 unfolding ssum_map_def by (simp add: cfun_eq_iff eta_cfun)
       
   221 
       
   222 lemma ssum_map_map:
       
   223   "\<lbrakk>f1\<cdot>\<bottom> = \<bottom>; g1\<cdot>\<bottom> = \<bottom>\<rbrakk> \<Longrightarrow>
       
   224     ssum_map\<cdot>f1\<cdot>g1\<cdot>(ssum_map\<cdot>f2\<cdot>g2\<cdot>p) =
       
   225      ssum_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
       
   226 apply (induct p, simp)
       
   227 apply (case_tac "f2\<cdot>x = \<bottom>", simp, simp)
       
   228 apply (case_tac "g2\<cdot>y = \<bottom>", simp, simp)
       
   229 done
       
   230 
       
   231 lemma ep_pair_ssum_map:
       
   232   assumes "ep_pair e1 p1" and "ep_pair e2 p2"
       
   233   shows "ep_pair (ssum_map\<cdot>e1\<cdot>e2) (ssum_map\<cdot>p1\<cdot>p2)"
       
   234 proof
       
   235   interpret e1p1: pcpo_ep_pair e1 p1 unfolding pcpo_ep_pair_def by fact
       
   236   interpret e2p2: pcpo_ep_pair e2 p2 unfolding pcpo_ep_pair_def by fact
       
   237   fix x show "ssum_map\<cdot>p1\<cdot>p2\<cdot>(ssum_map\<cdot>e1\<cdot>e2\<cdot>x) = x"
       
   238     by (induct x) simp_all
       
   239   fix y show "ssum_map\<cdot>e1\<cdot>e2\<cdot>(ssum_map\<cdot>p1\<cdot>p2\<cdot>y) \<sqsubseteq> y"
       
   240     apply (induct y, simp)
       
   241     apply (case_tac "p1\<cdot>x = \<bottom>", simp, simp add: e1p1.e_p_below)
       
   242     apply (case_tac "p2\<cdot>y = \<bottom>", simp, simp add: e2p2.e_p_below)
       
   243     done
       
   244 qed
       
   245 
       
   246 lemma deflation_ssum_map:
       
   247   assumes "deflation d1" and "deflation d2"
       
   248   shows "deflation (ssum_map\<cdot>d1\<cdot>d2)"
       
   249 proof
       
   250   interpret d1: deflation d1 by fact
       
   251   interpret d2: deflation d2 by fact
       
   252   fix x
       
   253   show "ssum_map\<cdot>d1\<cdot>d2\<cdot>(ssum_map\<cdot>d1\<cdot>d2\<cdot>x) = ssum_map\<cdot>d1\<cdot>d2\<cdot>x"
       
   254     apply (induct x, simp)
       
   255     apply (case_tac "d1\<cdot>x = \<bottom>", simp, simp add: d1.idem)
       
   256     apply (case_tac "d2\<cdot>y = \<bottom>", simp, simp add: d2.idem)
       
   257     done
       
   258   show "ssum_map\<cdot>d1\<cdot>d2\<cdot>x \<sqsubseteq> x"
       
   259     apply (induct x, simp)
       
   260     apply (case_tac "d1\<cdot>x = \<bottom>", simp, simp add: d1.below)
       
   261     apply (case_tac "d2\<cdot>y = \<bottom>", simp, simp add: d2.below)
       
   262     done
       
   263 qed
       
   264 
       
   265 lemma finite_deflation_ssum_map:
       
   266   assumes "finite_deflation d1" and "finite_deflation d2"
       
   267   shows "finite_deflation (ssum_map\<cdot>d1\<cdot>d2)"
       
   268 proof (rule finite_deflation_intro)
       
   269   interpret d1: finite_deflation d1 by fact
       
   270   interpret d2: finite_deflation d2 by fact
       
   271   have "deflation d1" and "deflation d2" by fact+
       
   272   thus "deflation (ssum_map\<cdot>d1\<cdot>d2)" by (rule deflation_ssum_map)
       
   273   have "{x. ssum_map\<cdot>d1\<cdot>d2\<cdot>x = x} \<subseteq>
       
   274         (\<lambda>x. sinl\<cdot>x) ` {x. d1\<cdot>x = x} \<union>
       
   275         (\<lambda>x. sinr\<cdot>x) ` {x. d2\<cdot>x = x} \<union> {\<bottom>}"
       
   276     by (rule subsetI, case_tac x, simp_all)
       
   277   thus "finite {x. ssum_map\<cdot>d1\<cdot>d2\<cdot>x = x}"
       
   278     by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes)
       
   279 qed
       
   280 
       
   281 end
   198 end