Static.thy
changeset 15 4ca824cd0c59
parent 12 47a4b2ae0556
child 19 ced0fcfbcf8e
--- a/Static.thy	Thu May 30 15:28:57 2013 +0800
+++ b/Static.thy	Mon Jun 03 10:34:58 2013 +0800
@@ -58,7 +58,7 @@
 definition init_cph2spshs 
   :: "t_process \<Rightarrow> (t_sshm \<times> t_shm_attach_flag) set"
 where
-  "init_cph2spshs p \<equiv> {(sh, flag). \<exists> h. (p, flag) \<in> init_procs_of_shm h \<and> 
+  "init_cph2spshs p \<equiv> {(sh, flag)| sh flag h. (p, flag) \<in> init_procs_of_shm h \<and> 
                                              init_ch2sshm h = Some sh}"
 
 definition init_cp2sproc :: "t_process \<Rightarrow> t_sproc option"
@@ -723,7 +723,7 @@
 
 definition cph2spshs :: "t_state \<Rightarrow> t_process \<Rightarrow> t_sproc_sshm set"
 where
-  "cph2spshs s p \<equiv> {(sh, flag). \<exists> h. (p, flag) \<in> procs_of_shm s h \<and> ch2sshm s h = Some sh}"
+  "cph2spshs s p \<equiv> {(sh, flag)| sh flag h. (p, flag) \<in> procs_of_shm s h \<and> ch2sshm s h = Some sh}"
 
 definition cp2sproc :: "t_state \<Rightarrow> t_process \<Rightarrow> t_sproc option"
 where