src/HOLCF/Sprod.thy
changeset 33587 54f98d225163
parent 33504 b4210cc3ac97
child 33808 31169fdc5ae7
equal deleted inserted replaced
33586:0e745228d605 33587:54f98d225163
   242 unfolding sprod_map_def by simp
   242 unfolding sprod_map_def by simp
   243 
   243 
   244 lemma sprod_map_spair [simp]:
   244 lemma sprod_map_spair [simp]:
   245   "x \<noteq> \<bottom> \<Longrightarrow> y \<noteq> \<bottom> \<Longrightarrow> sprod_map\<cdot>f\<cdot>g\<cdot>(:x, y:) = (:f\<cdot>x, g\<cdot>y:)"
   245   "x \<noteq> \<bottom> \<Longrightarrow> y \<noteq> \<bottom> \<Longrightarrow> sprod_map\<cdot>f\<cdot>g\<cdot>(:x, y:) = (:f\<cdot>x, g\<cdot>y:)"
   246 by (simp add: sprod_map_def)
   246 by (simp add: sprod_map_def)
       
   247 
       
   248 lemma sprod_map_map:
       
   249   "\<lbrakk>f1\<cdot>\<bottom> = \<bottom>; g1\<cdot>\<bottom> = \<bottom>\<rbrakk> \<Longrightarrow>
       
   250     sprod_map\<cdot>f1\<cdot>g1\<cdot>(sprod_map\<cdot>f2\<cdot>g2\<cdot>p) =
       
   251      sprod_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p"
       
   252 apply (induct p, simp)
       
   253 apply (case_tac "f2\<cdot>x = \<bottom>", simp)
       
   254 apply (case_tac "g2\<cdot>y = \<bottom>", simp)
       
   255 apply simp
       
   256 done
   247 
   257 
   248 lemma ep_pair_sprod_map:
   258 lemma ep_pair_sprod_map:
   249   assumes "ep_pair e1 p1" and "ep_pair e2 p2"
   259   assumes "ep_pair e1 p1" and "ep_pair e2 p2"
   250   shows "ep_pair (sprod_map\<cdot>e1\<cdot>e2) (sprod_map\<cdot>p1\<cdot>p2)"
   260   shows "ep_pair (sprod_map\<cdot>e1\<cdot>e2) (sprod_map\<cdot>p1\<cdot>p2)"
   251 proof
   261 proof