Current_prop.thy
changeset 29 622516c0fe34
parent 28 e298d755bc35
child 30 01274a64aece
equal deleted inserted replaced
28:e298d755bc35 29:622516c0fe34
     1 (*<*)
     1 (*<*)
     2 theory Current_prop
     2 theory Current_prop
     3 imports Main Flask_type Flask My_list_prefix Init_prop Valid_prop Delete_prop
     3 imports Main Flask_type Flask My_list_prefix Init_prop Valid_prop Delete_prop
     4 begin
     4 begin
     5 (*>*)
     5 (*>*)
       
     6 
       
     7 ML {*quick_and_dirty := true*}
     6 
     8 
     7 context flask begin
     9 context flask begin
     8 
    10 
     9 lemma procs_of_shm_prop1: "\<lbrakk> p_flag \<in> procs_of_shm s h; valid s\<rbrakk> \<Longrightarrow> h \<in> current_shms s"
    11 lemma procs_of_shm_prop1: "\<lbrakk> p_flag \<in> procs_of_shm s h; valid s\<rbrakk> \<Longrightarrow> h \<in> current_shms s"
    10 apply (induct s arbitrary:p_flag)
    12 apply (induct s arbitrary:p_flag)
   141 done
   143 done
   142 
   144 
   143 lemmas one_flow_shm_simps = one_flow_shm_other one_flow_shm_attach one_flow_shm_detach one_flow_shm_deleteshm
   145 lemmas one_flow_shm_simps = one_flow_shm_other one_flow_shm_attach one_flow_shm_detach one_flow_shm_deleteshm
   144   one_flow_shm_clone one_flow_shm_execve one_flow_shm_kill one_flow_shm_exit
   146   one_flow_shm_clone one_flow_shm_execve one_flow_shm_kill one_flow_shm_exit
   145 
   147 
   146 
   148 type_synonym t_edge_shm = "t_process \<times> t_shm \<times> t_process"
   147 inductive Info_flow_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> bool"
   149 fun Fst:: "t_edge_shm \<Rightarrow> t_process" where "Fst (a, b, c) = a"
       
   150 fun Snd:: "t_edge_shm \<Rightarrow> t_shm" where "Snd (a, b, c) = b"
       
   151 fun Trd:: "t_edge_shm \<Rightarrow> t_process" where "Trd (a, b, c) = c"
       
   152 
       
   153 fun edge_related:: "t_edge_shm list \<Rightarrow> t_process \<Rightarrow> t_shm \<Rightarrow> bool"
   148 where
   154 where
   149   ifs_self: "p \<in> current_procs s \<Longrightarrow> Info_flow_shm s p p"
   155   "edge_related [] p h = False"
   150 | ifs_flow:"\<lbrakk>Info_flow_shm s p p'; one_flow_shm s h p' p''\<rbrakk> \<Longrightarrow> Info_flow_shm s p p''"
   156 | "edge_related ((from, shm, to) # path) p h = 
   151 
   157      (if (((p = from) \<or> (p = to)) \<and> (h = shm)) then True 
   152 lemma Info_flow_trans_aux:
   158       else edge_related path p h)"
   153   "Info_flow_shm s p' p'' \<Longrightarrow> \<forall>p. Info_flow_shm s p p' \<longrightarrow> Info_flow_shm s p p''"
   159          
   154 apply (erule Info_flow_shm.induct)
   160 inductive path_by_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_edge_shm list \<Rightarrow> t_process \<Rightarrow> bool"
   155 by (auto intro:Info_flow_shm.intros)
   161 where
   156 
   162   pbs1: "p \<in> current_procs s \<Longrightarrow> path_by_shm s p [] p"
   157 lemma Info_flow_trans:
   163 | pbs2: "\<lbrakk>path_by_shm s p path p'; one_flow_shm s h p' p''; p'' \<notin> set (map Fst path)\<rbrakk> 
   158   "\<lbrakk>Info_flow_shm s p p'; Info_flow_shm s p' p''\<rbrakk> \<Longrightarrow> Info_flow_shm s p p''"
   164          \<Longrightarrow> path_by_shm s p ((p', h, p'')# path) p''"
   159 by (auto dest:Info_flow_trans_aux)
   165 
   160 
   166 
   161 lemma one_flow_flows:
   167 lemma one_step_path: "\<lbrakk>one_flow_shm s h p p'; valid s\<rbrakk> \<Longrightarrow> path_by_shm s p [(p, h, p')] p'"
   162   "\<lbrakk>one_flow_shm s h p p'; valid s\<rbrakk> \<Longrightarrow> Info_flow_shm s p p'"
   168 apply (rule_tac path = "[]" and p = p in path_by_shm.intros(2))
   163 apply (rule Info_flow_shm.intros(2), simp_all)
   169 apply (rule path_by_shm.intros(1))
   164 apply (rule Info_flow_shm.intros(1))
       
   165 apply (auto intro:procs_of_shm_prop2 simp:one_flow_shm_def)
   170 apply (auto intro:procs_of_shm_prop2 simp:one_flow_shm_def)
   166 done
   171 done
   167 
   172 
   168 lemma ifs_flow': "\<lbrakk>one_flow_shm s h p p'; Info_flow_shm s p' p''; valid s\<rbrakk> \<Longrightarrow> Info_flow_shm s p p''"
   173 lemma pbs_prop1:
   169 apply (drule one_flow_flows, simp+)
   174   "path_by_shm s p path p' \<Longrightarrow> ((path = []) = (p = p')) \<and> (path \<noteq> [] \<longrightarrow> p \<in> set (map Fst path))"
   170 apply (erule Info_flow_trans, simp+)
   175 apply (erule path_by_shm.induct, simp)
   171 done
   176 apply (auto simp:one_flow_shm_def)
   172 
   177 done
   173 lemma Info_flow_shm_cases1:
   178 
   174   "\<lbrakk>Info_flow_shm s pa pb; 
   179 lemma pbs_prop2:
   175     \<And>p \<tau>. \<lbrakk>\<tau> = s; pa = p; pb = p; p \<in> current_procs \<tau>\<rbrakk> \<Longrightarrow> P;
   180   "path_by_shm s p path p' \<Longrightarrow> (path = []) = (p = p')"
   176     \<And>\<tau> p p' h p''. \<lbrakk>\<tau> = s; pa = p; pb = p''; Info_flow_shm \<tau> p p'; one_flow_shm \<tau> h p' p''\<rbrakk> \<Longrightarrow> P\<rbrakk>
   181 by (simp add:pbs_prop1)
   177    \<Longrightarrow> P"
   182 
   178 by (erule Info_flow_shm.cases, auto)
   183 lemma pbs_prop2':
   179 
   184   "path_by_shm s p path p \<Longrightarrow> path = []"
   180 lemma Info_flow_shm_prop1:
   185 by (simp add:pbs_prop2)
   181   "\<not> Info_flow_shm s p p \<Longrightarrow> p \<notin> current_procs s" 
   186 
   182 by (rule notI, drule Info_flow_shm.intros(1), simp)
   187 lemma pbs_prop3:
   183 
   188   "\<lbrakk>path_by_shm s p path p'; path \<noteq> []\<rbrakk> \<Longrightarrow> p \<in> set (map Fst path)"
   184 lemma Info_flow_shm_intro3:
   189 by (drule pbs_prop1, auto)
   185   "\<lbrakk>Info_flow_shm s p from; (from, SHM_RDWR) \<in> procs_of_shm s h; (to, flag) \<in> procs_of_shm s h\<rbrakk>
   190 
   186    \<Longrightarrow> Info_flow_shm s p to"
   191 lemma pbs_prop4[rule_format]:
   187 apply (case_tac "from = to", simp)
   192   "path_by_shm s p path p'\<Longrightarrow> path \<noteq> [] \<longrightarrow> p' \<in> set (map Trd path)"
   188 apply (erule_tac h = h in Info_flow_shm.intros(2), simp add:one_flow_shm_def)
   193 by (erule path_by_shm.induct, auto)
   189 by (rule_tac x = flag in exI, simp)
   194 
   190 
   195 lemma pbs_prop5[rule_format]:
   191 lemma Info_flow_shm_intro4:
   196   "path_by_shm s p path p' \<Longrightarrow> path \<noteq> [] \<longrightarrow> p' \<notin> set (map Fst path)"
   192   "\<lbrakk>(p, flagb) \<in> procs_of_shm s h; valid s\<rbrakk> \<Longrightarrow> Info_flow_shm s p p"
   197 by (erule path_by_shm.induct, auto simp:one_flow_shm_def)
   193 by (drule procs_of_shm_prop2, simp, simp add:Info_flow_shm.intros)
   198 
   194 
   199 lemma pbs_prop6_aux:
       
   200   "path_by_shm s pa pathac pc \<Longrightarrow> valid s \<longrightarrow> (\<forall> pb \<in> set (map Fst pathac). \<exists> pathab pathbc. path_by_shm s pa pathab pb \<and> path_by_shm s pb pathbc pc \<and> pathac = pathbc @ pathab)"
       
   201 apply (erule path_by_shm.induct)
       
   202 apply simp
       
   203 apply clarify
       
   204 apply (case_tac "pb = p'", simp)
       
   205 apply (rule_tac x = path in exI, simp)
       
   206 apply (erule one_step_path, simp)
       
   207 apply (erule_tac x = pb in ballE, simp_all, clarsimp)
       
   208 apply (rule_tac x = pathab in exI, simp)
       
   209 apply (erule pbs2, auto)
       
   210 done
       
   211 
       
   212 lemma pbs_prop6:
       
   213   "\<lbrakk>path_by_shm s pa pathac pc; pb \<in> set (map Fst pathac); valid s\<rbrakk>
       
   214    \<Longrightarrow> \<exists> pathab pathbc. path_by_shm s pa pathab pb \<and> path_by_shm s pb pathbc pc \<and> pathac = pathbc @ pathab"
       
   215 by (drule pbs_prop6_aux, auto)
       
   216 
       
   217 lemma pbs_prop7_aux:
       
   218   "path_by_shm s pa pathac pc \<Longrightarrow> valid s \<longrightarrow> (\<forall> pb \<in> set (map Trd pathac). \<exists> pathab pathbc. path_by_shm s pa pathab pb \<and> path_by_shm s pb pathbc pc \<and> pathac = pathbc @ pathab)"
       
   219 apply (erule path_by_shm.induct)
       
   220 apply simp
       
   221 apply clarify
       
   222 apply (case_tac "pb = p''", simp)
       
   223 apply (rule_tac x = "(p',h,p'') # path" in exI, simp)
       
   224 apply (rule conjI, erule pbs2, simp+)
       
   225 apply (rule pbs1, clarsimp simp:one_flow_shm_def procs_of_shm_prop2)
       
   226 apply (erule_tac x = pb in ballE, simp_all, clarsimp)
       
   227 apply (rule_tac x = pathab in exI, simp)
       
   228 apply (erule pbs2, auto)
       
   229 done
       
   230 
       
   231 lemma pbs_prop7:
       
   232   "\<lbrakk>path_by_shm s pa pathac pc; pb \<in> set (map Trd pathac); valid s\<rbrakk>
       
   233    \<Longrightarrow> \<exists> pathab pathbc. path_by_shm s pa pathab pb \<and> path_by_shm s pb pathbc pc \<and> pathac = pathbc @ pathab"
       
   234 by (drule pbs_prop7_aux, drule mp, simp, erule_tac x = pb in ballE, auto)
       
   235 
       
   236 lemma pbs_prop8:
       
   237   "path_by_shm s p path p' \<Longrightarrow> (set (map Fst path) - {p}) = (set (map Trd path) - {p'})"
       
   238 proof (induct rule:path_by_shm.induct)
       
   239   case (pbs1 p s)
       
   240   thus ?case by simp
       
   241 next
       
   242   case (pbs2 s p path p' h p'')
       
   243   assume p1:"path_by_shm s p path p'" and p2: "set (map Fst path) - {p} = set (map Trd path) - {p'}"
       
   244     and p3: "one_flow_shm s h p' p''" and p4: "p'' \<notin> set (map Fst path)" 
       
   245   show "set (map Fst ((p', h, p'') # path)) - {p} = set (map Trd ((p', h, p'') # path)) - {p''}"
       
   246     (is "?left = ?right")
       
   247   proof (cases "path = []")
       
   248     case True
       
   249     with p1 have "p = p'" by (drule_tac pbs_prop2, simp)
       
   250     thus ?thesis using True
       
   251       using p2 by (simp)
       
   252   next
       
   253     case False
       
   254     with p1 have a1: "p \<noteq> p'" by (drule_tac pbs_prop2, simp)
       
   255     from p3 have a2: "p' \<noteq> p''" by (simp add:one_flow_shm_def)
       
   256     from p1 False have a3: "p' \<in> set (map Trd path)" by (drule_tac pbs_prop4, simp+)
       
   257     from p4 p1 False have a4: "p \<noteq> p''" by (drule_tac pbs_prop3, auto)
       
   258     with p2 a2 p4 have a5: "p'' \<notin> set (map Trd path)" by auto
       
   259     
       
   260     have "?left = (set (map Fst path) - {p}) \<union> {p'}" using a1 by auto
       
   261     moreover have "... = (set (map Trd path) - {p'}) \<union> {p'}"  
       
   262       using p2 by auto
       
   263     moreover have "... = set (map Trd path)" using a3 by auto
       
   264     moreover have "... = set (map Trd path) - {p''}" using a5 by simp
       
   265     moreover have "... = ?right" by simp
       
   266     ultimately show ?thesis by simp
       
   267   qed
       
   268 qed
       
   269 
       
   270 lemma pbs_prop9_aux[rule_format]:
       
   271   "path_by_shm s p path p' \<Longrightarrow> h \<in> set (map Snd path) \<and> valid s \<longrightarrow> (\<exists> pa pb patha pathb. path_by_shm s p patha pa \<and> path_by_shm s pb pathb p' \<and> one_flow_shm s h pa pb \<and> path = pathb @ [(pa, h, pb)] @ patha \<and> h \<notin> set (map Snd patha))"
       
   272 apply (erule path_by_shm.induct, simp)
       
   273 apply (rule impI, case_tac "h \<in> set (map Snd path)", simp_all)
       
   274 apply (erule exE|erule conjE)+
       
   275 apply (rule_tac x = pa in exI, rule_tac x = pb in exI, rule_tac x = patha in exI, simp)
       
   276 apply (rule pbs2, auto)
       
   277 apply (rule_tac x = p' in exI, rule_tac x = p'' in exI, rule_tac x = path in exI, simp)
       
   278 apply (rule pbs1, clarsimp simp:one_flow_shm_def procs_of_shm_prop2)
       
   279 done
       
   280 
       
   281 lemma pbs_prop9:
       
   282   "\<lbrakk>h \<in> set (map Snd path); path_by_shm s p path p'; valid s\<rbrakk>
       
   283    \<Longrightarrow> \<exists> pa pb patha pathb. path_by_shm s p patha pa \<and> path_by_shm s pb pathb p' \<and> 
       
   284         one_flow_shm s h pa pb \<and> path = pathb @ [(pa, h, pb)] @ patha \<and> h \<notin> set (map Snd patha)"
       
   285 by (rule pbs_prop9_aux, auto)
       
   286 
       
   287 lemma path_by_shm_trans_aux[rule_format]:
       
   288   "path_by_shm s p' path' p'' \<Longrightarrow> valid s \<longrightarrow> (\<forall> p path. path_by_shm s p path p' \<longrightarrow> (\<exists> path''. path_by_shm s p path'' p''))"
       
   289 proof (induct rule:path_by_shm.induct)
       
   290   case (pbs1 p s)
       
   291   thus ?case
       
   292     by (clarify, rule_tac x = path in exI, simp)
       
   293 next
       
   294   case (pbs2 s p path p' h p'')
       
   295   hence p1: "path_by_shm s p path p'" and p2: "one_flow_shm s h p' p''" 
       
   296     and p3: "valid s \<longrightarrow> (\<forall>pa path. path_by_shm s pa path p \<longrightarrow> (\<exists>path''. path_by_shm s pa path'' p'))"
       
   297     and p4: "p'' \<notin> set (map Fst path)" by auto
       
   298   show ?case
       
   299   proof clarify
       
   300     fix pa path'
       
   301     assume p5: "path_by_shm s pa path' p" and p6: "valid s"
       
   302     with p3 obtain path'' where a1: "path_by_shm s pa path'' p'" by auto
       
   303     have p3': "\<forall>pa path. path_by_shm s pa path p \<longrightarrow> (\<exists>path''. path_by_shm s pa path'' p')" 
       
   304       using p3 p6 by simp
       
   305     show "\<exists>path''. path_by_shm s pa path'' p''"
       
   306     proof (cases "p'' \<in> set (map Fst path'')")
       
   307       case True
       
   308       then obtain res where "path_by_shm s pa res p''" using a1 pbs_prop6 p6 by blast
       
   309       thus ?thesis by auto
       
   310     next
       
   311       case False
       
   312       with p2 a1 show ?thesis 
       
   313         apply (rule_tac x = "(p', h, p'') # path''" in exI)
       
   314         apply (rule path_by_shm.intros(2), auto)
       
   315         done
       
   316     qed
       
   317   qed
       
   318 qed
       
   319 
       
   320 lemma path_by_shm_trans:
       
   321   "\<lbrakk>path_by_shm s p path p'; path_by_shm s p' path' p''; valid s\<rbrakk> \<Longrightarrow> \<exists> path''. path_by_shm s p path'' p''"
       
   322 by (drule_tac p' = p' and p'' = p'' in path_by_shm_trans_aux, auto)
       
   323 
       
   324 lemma path_by_shm_intro1_prop:
       
   325   "\<not> path_by_shm s p [] p \<Longrightarrow> p \<notin> current_procs s"
       
   326 by (auto dest:path_by_shm.intros(1))
       
   327 
       
   328 lemma path_by_shm_intro3:
       
   329   "\<lbrakk>path_by_shm s p path from; (from, SHM_RDWR) \<in> procs_of_shm s h; (to, flag) \<in> procs_of_shm s h; 
       
   330     to \<notin> set (map Fst path); from \<noteq> to\<rbrakk>
       
   331    \<Longrightarrow> path_by_shm s p ((from, h, to)#path) to"
       
   332 apply (rule path_by_shm.intros(2), simp_all)
       
   333 by (auto simp:one_flow_shm_def)
       
   334 
       
   335 lemma path_by_shm_intro4:
       
   336   "\<lbrakk>(p, flag) \<in> procs_of_shm s h; valid s\<rbrakk> \<Longrightarrow> path_by_shm s p [] p"
       
   337 by (drule procs_of_shm_prop2, simp, simp add:path_by_shm.intros(1))
       
   338 
       
   339 lemma path_by_shm_intro5:
       
   340   "\<lbrakk>(from, SHM_RDWR) \<in> procs_of_shm s h; (to,flag) \<in> procs_of_shm s h; valid s; from \<noteq> to\<rbrakk>
       
   341    \<Longrightarrow> path_by_shm s from [(from, h, to)] to"
       
   342 apply (rule_tac p' = "from" and h = h in path_by_shm.intros(2))
       
   343 apply (rule path_by_shm.intros(1), simp add:procs_of_shm_prop2)
       
   344 apply (simp add:one_flow_shm_def, rule_tac x = flag in exI, auto)
       
   345 done
       
   346 
       
   347 (* p'' \<notin> set (map Fst path): not duplicated target process;
       
   348  * p1 - ha \<rightarrow> p2; p2 - hb \<rightarrow> p3; p3 - ha \<rightarrow> p4; so path_by_shm p1 [(p3,ha,p4), (p2,hb,p3),(p1,ha,p2)] p4,
       
   349  * but this could be also path_by_shm p1 [(p1,ha,p4)] p4, so the former one is redundant!  *)
       
   350 
       
   351 inductive path_by_shm':: "t_state \<Rightarrow> t_process \<Rightarrow> t_edge_shm list \<Rightarrow> t_process \<Rightarrow> bool"
       
   352 where
       
   353   pbs1': "p \<in> current_procs s \<Longrightarrow> path_by_shm' s p [] p"
       
   354 | pbs2': "\<lbrakk>path_by_shm s p path p'; one_flow_shm s h p' p''; p'' \<notin> set (map Fst path); 
       
   355            h \<notin> set (map Snd path)\<rbrakk> 
       
   356           \<Longrightarrow> path_by_shm' s p ((p', h, p'')# path) p''"
       
   357 
       
   358 lemma pbs_imp_pbs'[rule_format]:
       
   359   "path_by_shm s p path p' \<Longrightarrow> valid s \<longrightarrow> (\<exists> path'. path_by_shm' s p path' p')"
       
   360 apply (erule path_by_shm.induct)
       
   361 apply (rule impI, rule_tac x = "[]" in exI, erule pbs1')
       
   362 apply (rule impI, simp,  erule exE, case_tac "h \<in> set (map Snd path)")
       
   363 apply (drule_tac s = s and p = p and p' = p' in pbs_prop9, simp+) defer
       
   364 apply (rule_tac x = "(p', h, p'') # path" in exI, erule pbs2', simp+) 
       
   365 apply ((erule exE|erule conjE)+)
       
   366 apply (rule_tac x = "(pa, h, p'') # patha" in exI)
       
   367 apply (erule pbs2', auto simp:one_flow_shm_def)
       
   368 done
       
   369 
       
   370 lemma pbs'_imp_pbs[rule_format]:
       
   371   "path_by_shm' s p path p' \<Longrightarrow> valid s \<longrightarrow> (\<exists> path'. path_by_shm s p path' p')"
       
   372 apply (erule path_by_shm'.induct)
       
   373 apply (rule impI, rule_tac x = "[]" in exI, simp add:pbs1)
       
   374 apply (rule impI, rule_tac x = "(p', h, p'') # path" in exI, simp add:pbs2)
       
   375 done
       
   376 
       
   377 definition flow_by_shm :: "t_state \<Rightarrow> t_process \<Rightarrow> t_process \<Rightarrow> bool"
       
   378 where
       
   379   "flow_by_shm s p p' \<equiv> \<exists> path. path_by_shm s p path p'"
       
   380 
       
   381 lemma flow_by_shm_intro':
       
   382   "valid s \<Longrightarrow> flow_by_shm s p p' = (\<exists> path. path_by_shm' s p path p')"
       
   383 by (auto simp:flow_by_shm_def intro: pbs_imp_pbs' pbs'_imp_pbs)
       
   384 
       
   385 lemma one_step_flows: "\<lbrakk>one_flow_shm s h p p'; valid s\<rbrakk> \<Longrightarrow> flow_by_shm s p p'"
       
   386 by (drule one_step_path, auto simp:flow_by_shm_def)
       
   387 
       
   388 lemma flow_by_shm_trans:
       
   389   "\<lbrakk>flow_by_shm s p p'; flow_by_shm s p' p''; valid s\<rbrakk> \<Longrightarrow> flow_by_shm s p p''"
       
   390 by (auto simp:flow_by_shm_def intro!:path_by_shm_trans)
       
   391 
       
   392 lemma flow_by_shm_intro1_prop:
       
   393   "\<not> flow_by_shm s p p \<Longrightarrow> p \<notin> current_procs s"
       
   394 by (auto dest:path_by_shm.intros(1) simp:flow_by_shm_def)
       
   395 
       
   396 lemma flow_by_shm_intro1:
       
   397   "p \<in> current_procs s \<Longrightarrow> flow_by_shm s p p"
       
   398 by (auto dest:path_by_shm.intros(1) simp:flow_by_shm_def)
       
   399 
       
   400 lemma flow_by_shm_intro2:
       
   401   "\<lbrakk>flow_by_shm s p p'; one_flow_shm s h p' p''; valid s\<rbrakk> \<Longrightarrow> flow_by_shm s p p''"
       
   402 by (auto intro:flow_by_shm_trans dest:one_step_flows)
       
   403 
       
   404 lemma flow_by_shm_intro3:
       
   405   "\<lbrakk>flow_by_shm s p from; (from, SHM_RDWR) \<in> procs_of_shm s h; (to, flag) \<in> procs_of_shm s h; from \<noteq> to; valid s\<rbrakk>
       
   406    \<Longrightarrow> flow_by_shm s p to"
       
   407 apply (rule flow_by_shm_intro2)
       
   408 by (auto simp:one_flow_shm_def)
       
   409 
       
   410 lemma flow_by_shm_intro4:
       
   411   "\<lbrakk>(p, flag) \<in> procs_of_shm s h; valid s\<rbrakk> \<Longrightarrow> flow_by_shm s p p"
       
   412 by (drule procs_of_shm_prop2, simp, simp add:flow_by_shm_intro1)
       
   413 
       
   414 lemma flow_by_shm_intro5:
       
   415   "\<lbrakk>(from, SHM_RDWR) \<in> procs_of_shm s h; (to,flag) \<in> procs_of_shm s h; valid s; from \<noteq> to\<rbrakk>
       
   416    \<Longrightarrow> flow_by_shm s from  to"
       
   417 apply (rule_tac p' = "from" and h = h in flow_by_shm_intro2)
       
   418 apply (rule flow_by_shm_intro1, simp add:procs_of_shm_prop2)
       
   419 apply (simp add:one_flow_shm_def, rule_tac x = flag in exI, auto)
       
   420 done
       
   421 
       
   422 lemma flow_by_shm_intro6:
       
   423   "path_by_shm s p path p' \<Longrightarrow> flow_by_shm s p p'"
       
   424 by (auto simp:flow_by_shm_def)
   195 (********* simpset for inductive Info_flow_shm **********)
   425 (********* simpset for inductive Info_flow_shm **********)
   196 
   426 term edge_related
   197 lemma Info_flow_shm_attach1_aux:
   427 lemma path_by_shm_detach1_aux:
   198   "Info_flow_shm s' pa pb \<Longrightarrow> valid s' \<and> (s' = Attach p h flag # s) \<longrightarrow> 
   428   "path_by_shm s' pa path pb \<Longrightarrow> valid (Detach p h # s) \<and> (s' = Detach p h # s) 
   199      (if Info_flow_shm s pa pb then True else
   429      \<longrightarrow> \<not> edge_related path p h \<and> path_by_shm s pa path pb"
   200      (if (pa = p \<and> flag = SHM_RDWR) 
   430 apply (erule path_by_shm.induct, simp)
   201       then (\<exists> p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> Info_flow_shm s p' pb)
   431 apply (rule impI, rule path_by_shm.intros(1), simp+)
   202       else if (pb = p) 
   432 by (auto simp:one_flow_shm_def split:if_splits intro:path_by_shm_intro3)
   203            then (\<exists> p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s pa p')
   433 
   204            else (\<exists> p' flag'. Info_flow_shm s pa p \<and> flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and> 
   434 lemma path_by_shm_detach1:
   205                              Info_flow_shm s p' pb) \<or>
   435   "\<lbrakk>path_by_shm (Detach p h # s) pa path pb; valid (Detach p h # s)\<rbrakk> 
   206                 (\<exists> p'. Info_flow_shm s pa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s p pb)
   436    \<Longrightarrow> \<not> edge_related path p h \<and> path_by_shm s pa path pb"
   207      )  )"
   437 by (auto dest:path_by_shm_detach1_aux)
   208 proof (induct rule:Info_flow_shm.induct)
   438 
   209   case (ifs_self proc \<tau>)
   439 lemma path_by_shm_detach2_aux[rule_format]:
       
   440   "path_by_shm s pa path pb \<Longrightarrow> valid (Detach p h # s) \<and> \<not> edge_related path p h 
       
   441    \<longrightarrow> path_by_shm (Detach p h # s) pa path pb"
       
   442 apply (induct rule:path_by_shm.induct)
       
   443 apply (rule impI, rule path_by_shm.intros(1), simp)
       
   444 apply (rule impI, erule conjE, simp split:if_splits)
       
   445 apply (rule path_by_shm.intros(2), simp)
       
   446 apply (simp add:one_flow_shm_detach)
       
   447 apply (rule impI, simp+)
       
   448 done
       
   449 
       
   450 lemma path_by_shm_detach2:
       
   451   "\<lbrakk>valid (Detach p h # s); \<not> edge_related path p h; path_by_shm s pa path pb\<rbrakk> 
       
   452    \<Longrightarrow> path_by_shm (Detach p h # s) pa path pb"
       
   453 by (auto intro!:path_by_shm_detach2_aux)
       
   454 
       
   455 lemma path_by_shm_detach:
       
   456   "valid (Detach p h # s) \<Longrightarrow>
       
   457    path_by_shm (Detach p h # s) pa path pb = (\<not> edge_related path p h  \<and> path_by_shm s pa path pb)"
       
   458 by (auto dest:path_by_shm_detach1 path_by_shm_detach2)
       
   459 
       
   460 lemma flow_by_shm_detach:
       
   461   "valid (Detach p h # s) \<Longrightarrow> 
       
   462    flow_by_shm (Detach p h # s) pa pb = (\<exists> path. \<not> edge_related path p h \<and> path_by_shm s pa path pb)"
       
   463 by (auto dest:path_by_shm_detach simp:flow_by_shm_def)
       
   464 
       
   465 lemma path_by_shm_attach1_aux:
       
   466   "path_by_shm s' pa path pb \<Longrightarrow> valid s' \<and> (s' = Attach p h flag # s) \<longrightarrow>
       
   467      (path_by_shm s pa path pb) \<or>
       
   468      (\<exists> path1 path2 p'. path_by_shm s pa path1 p' \<and> path_by_shm s p path2 pb \<and> 
       
   469          (p', SHM_RDWR) \<in> procs_of_shm s h \<and> path = path2 @ [(p', h, p)] @ path1 ) \<or>
       
   470      (\<exists> path1 path2 p' flag'. path_by_shm s pa path1 p \<and> path_by_shm s p' path2 pb \<and> 
       
   471          (p', flag') \<in> procs_of_shm s h \<and> path = path2 @ [(p, h, p')] @ path1 \<and> flag = SHM_RDWR)"
       
   472 apply (erule path_by_shm.induct)
       
   473 apply (simp, rule impI, rule pbs1, simp)
       
   474 apply (rule impI, erule impE, clarsimp)
       
   475 apply (erule disjE)
       
   476 apply (clarsimp simp:one_flow_shm_attach split:if_splits)
       
   477 apply (erule disjE, clarsimp)
       
   478 apply (erule_tac x = path in allE, clarsimp)
       
   479 apply (erule impE, rule pbs1, erule procs_of_shm_prop2, erule vd_cons, simp)
       
   480 apply (erule disjE, clarsimp)
       
   481 apply (rule_tac x = path in exI, rule_tac x = "[]" in exI, rule_tac x = p' in exI, simp)
       
   482 apply (rule pbs1, drule vt_grant_os, clarsimp)
       
   483 apply (drule_tac p = pa and p' = p' and p'' = p'' in pbs2, simp+)
       
   484 apply (drule_tac p = pa and p' = p' and p'' = p'' in pbs2, simp+)
       
   485 
       
   486 apply (erule disjE)
       
   487 apply ((erule exE|erule conjE)+, simp split:if_splits add:one_flow_shm_attach)
       
   488 apply (clarsimp simp:one_flow_shm_attach split:if_splits)
       
   489 apply (erule disjE, clarsimp)
       
   490 apply (clarsimp)
       
   491 
       
   492 
       
   493 apply (erule conjE)+
       
   494 
       
   495 
       
   496 
       
   497 apply (erule conjE, clarsimp simp only:one_flow_shm_attach split:if_splits)
       
   498 apply simp
       
   499 
       
   500 
       
   501 
       
   502 lemma path_by_shm_attach1_aux:
       
   503   "path_by_shm s' pa path pb \<Longrightarrow> valid s' \<and> (s' = Attach p h flag # s) \<longrightarrow> 
       
   504      path_by_shm s pa path pb \<or>
       
   505       (if (pa = p \<and> flag = SHM_RDWR)
       
   506        then \<exists> p' flagb path'. (p', flagb) \<in> procs_of_shm s h \<and> 
       
   507                path_by_shm s p' path' pb \<and> path = path' @ [(p, h, p')]
       
   508        else if (pb = p)
       
   509             then \<exists> p' path'. path_by_shm s pa path' p' \<and> path = (p', h, p) # path' \<and> 
       
   510                    (p', SHM_RDWR) \<in> procs_of_shm s h
       
   511             else (\<exists> p' flag' patha pathb. path_by_shm s pa patha p \<and> flag = SHM_RDWR \<and> 
       
   512                    (p', flag') \<in> procs_of_shm s h \<and> path_by_shm s p' pathb pb \<and> 
       
   513                    path = pathb @ [(p, h, p')] @ patha) \<or>
       
   514                  (\<exists> p' patha pathb. path_by_shm s pa patha p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> 
       
   515                    path_by_shm s p pathb pb \<and> path = pathb @ [(p', h, p)] @ patha))"
       
   516 proof (induct rule:path_by_shm.induct)
       
   517   case (pbs1 proc \<tau>)
   210   show ?case
   518   show ?case
   211   proof (rule impI)
   519   proof (rule impI)
   212     assume pre: "valid \<tau> \<and> \<tau> = Attach p h flag # s"
   520     assume pre: "valid \<tau> \<and> \<tau> = Attach p h flag # s"
   213     hence p1: "p \<in> current_procs s" and p2: "valid s" by (auto intro:vd_cons dest:vt_grant_os)
   521     from pbs1 pre have "proc \<in> current_procs s" by simp 
   214     hence p3: "Info_flow_shm s p p" by (auto intro:Info_flow_shm.intros)
   522     thus "path_by_shm s proc [] proc \<or>
   215     from ifs_self pre have "proc \<in> current_procs s" by simp 
   523          (if proc = p \<and> flag = SHM_RDWR
   216     hence p4: "Info_flow_shm s proc proc" by (auto intro:Info_flow_shm.intros)
   524           then \<exists>p' flagb path'.
   217     show "if Info_flow_shm s proc proc then True
   525                   (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' path' proc \<and> [] = path' @ [(p, h, p')]
   218     else if proc = p \<and> flag = SHM_RDWR then \<exists>p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> Info_flow_shm s p' proc
   526           else if proc = p
   219          else if proc = p then \<exists>p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s proc p'
   527                then \<exists>p' path'.
   220               else (\<exists>p' flag'. Info_flow_shm s proc p \<and>
   528                        path_by_shm s proc path' p' \<and> [] = (p', h, p) # path' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h
   221                        flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and> Info_flow_shm s p' proc) \<or>
   529                else (\<exists>p' flag' patha pathb.
   222                    (\<exists>p'. Info_flow_shm s proc p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s p proc)"      using p4 p3 by auto
   530                         path_by_shm s proc patha p \<and>
       
   531                         flag = SHM_RDWR \<and>
       
   532                         (p', flag') \<in> procs_of_shm s h \<and>
       
   533                         path_by_shm s p' pathb proc \<and> [] = pathb @ [(p, h, p')] @ patha) \<or>
       
   534                     (\<exists>p' patha pathb.
       
   535                         path_by_shm s proc patha p' \<and>
       
   536                         (p', SHM_RDWR) \<in> procs_of_shm s h \<and>
       
   537                         path_by_shm s p pathb proc \<and> [] = pathb @ [(p', h, p)] @ patha))"
       
   538       by (auto intro:path_by_shm.intros)
   223   qed
   539   qed
   224 next
   540 next
   225   case (ifs_flow \<tau> pa pb h' pc)
   541   case (pbs2 \<tau> pa path pb h' pc)
   226   thus ?case
   542   thus ?case
   227   proof (rule_tac impI)
   543   proof (rule_tac impI)
   228     assume p1:"Info_flow_shm \<tau> pa pb" and p2: "valid \<tau> \<and> \<tau> = Attach p h flag # s \<longrightarrow>
   544     assume p1:"path_by_shm \<tau> pa path pb" and p2: "valid \<tau> \<and> \<tau> = Attach p h flag # s \<longrightarrow>
   229      (if Info_flow_shm s pa pb then True
   545      path_by_shm s pa path pb \<or>
   230       else if pa = p \<and> flag = SHM_RDWR then \<exists>p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> Info_flow_shm s p' pb
   546      (if pa = p \<and> flag = SHM_RDWR
   231            else if pb = p then \<exists>p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s pa p'
   547       then \<exists>p' flagb path'. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' path' pb \<and> path = path' @ [(p, h, p')]
   232                 else (\<exists>p' flag'. Info_flow_shm s pa p \<and>
   548       else if pb = p
   233                          flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and> Info_flow_shm s p' pb) \<or>
   549            then \<exists>p' path'. path_by_shm s pa path' p' \<and> path = (p', h, p) # path' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h
   234                      (\<exists>p'. Info_flow_shm s pa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s p pb))"
   550            else (\<exists>p' flag' pathaa pathb. path_by_shm s pa pathaa p \<and> flag = SHM_RDWR \<and> 
       
   551                     (p', flag') \<in> procs_of_shm s h \<and> path_by_shm s p' pathb pb \<and> 
       
   552                     path = pathb @ [(p, h, p')] @ pathaa) \<or>
       
   553                 (\<exists>p' pathaa pathb. path_by_shm s pa pathaa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and>
       
   554                     path_by_shm s p pathb pb \<and> path = pathb @ [(p', h, p)] @ pathaa))"
   235       and p3: "one_flow_shm \<tau> h' pb pc" and p4: "valid \<tau> \<and> \<tau> = Attach p h flag # s"
   555       and p3: "one_flow_shm \<tau> h' pb pc" and p4: "valid \<tau> \<and> \<tau> = Attach p h flag # s"
   236 
   556     
   237     from p2 and p4 have p2': "(if Info_flow_shm s pa pb then True
   557     from p2 and p4 have p2': "
   238       else if pa = p \<and> flag = SHM_RDWR then \<exists>p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> Info_flow_shm s p' pb
   558       path_by_shm s pa path pb \<or>
   239            else if pb = p then \<exists>p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s pa p'
   559      (if pa = p \<and> flag = SHM_RDWR
   240                 else (\<exists>p' flag'. Info_flow_shm s pa p \<and>
   560       then \<exists>p' flagb path'. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' path' pb \<and> path = path' @ [(p, h, p')]
   241                          flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and> Info_flow_shm s p' pb) \<or>
   561       else if pb = p
   242                      (\<exists>p'. Info_flow_shm s pa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s p pb))"
   562            then \<exists>p' path'. path_by_shm s pa path' p' \<and> path = (p', h, p) # path' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h
       
   563            else (\<exists>p' flag' pathaa pathb. path_by_shm s pa pathaa p \<and> flag = SHM_RDWR \<and> 
       
   564                     (p', flag') \<in> procs_of_shm s h \<and> path_by_shm s p' pathb pb \<and> 
       
   565                     path = pathb @ [(p, h, p')] @ pathaa) \<or>
       
   566                 (\<exists>p' pathaa pathb. path_by_shm s pa pathaa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and>
       
   567                     path_by_shm s p pathb pb \<and> path = pathb @ [(p', h, p)] @ pathaa))"
   243       by (erule_tac impE, simp)
   568       by (erule_tac impE, simp)
   244     from p4 have p5: "valid s" and p6: "os_grant s (Attach p h flag)" by (auto intro:vd_cons dest:vt_grant_os)
   569     from p4 have p5: "valid s" and p6: "os_grant s (Attach p h flag)" by (auto intro:vd_cons dest:vt_grant_os)
   245     from p6 have "p \<in> current_procs s" by simp hence p7:"Info_flow_shm s p p" by (erule_tac Info_flow_shm.intros)
   570     from p6 have "p \<in> current_procs s" by simp hence p7:"path_by_shm s p [] p" by (erule_tac path_by_shm.intros)
   246     from p3 p4 have p8: "if (h' = h) 
   571     from p3 p4 have p8: "if (h' = h) 
   247      then (pb = p \<and> pc \<noteq> p \<and> flag = SHM_RDWR \<and> (\<exists> flagb. (pc, flagb) \<in> procs_of_shm s h)) \<or>
   572      then (pb = p \<and> pc \<noteq> p \<and> flag = SHM_RDWR \<and> (\<exists> flagb. (pc, flagb) \<in> procs_of_shm s h)) \<or>
   248           (pc = p \<and> pb \<noteq> p \<and> (pb, SHM_RDWR) \<in> procs_of_shm s h) \<or>
   573           (pc = p \<and> pb \<noteq> p \<and> (pb, SHM_RDWR) \<in> procs_of_shm s h) \<or>
   249           (one_flow_shm s h pb pc)               
   574           (one_flow_shm s h pb pc)               
   250      else one_flow_shm s h' pb pc        " by (auto simp add:one_flow_shm_attach) 
   575      else one_flow_shm s h' pb pc" by (auto simp add:one_flow_shm_attach) 
   251 
   576     
       
   577     
       
   578 (*
   252     have "\<And> flagb. (pc, flagb) \<in> procs_of_shm s h 
   579     have "\<And> flagb. (pc, flagb) \<in> procs_of_shm s h 
   253       \<Longrightarrow> \<exists> p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> Info_flow_shm s p' pc"
   580       \<Longrightarrow> \<exists> p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' [] pc"
   254       apply (rule_tac x= pc in exI, rule_tac x = flagb in exI, frule procs_of_shm_prop2)
   581       apply (rule_tac x= pc in exI, rule_tac x = flagb in exI, frule procs_of_shm_prop2)
   255       by (simp add:p5, simp add:Info_flow_shm.intros(1))
   582       by (simp add:p5, simp add:path_by_shm.intros(1))
   256     hence p10: "\<not> Info_flow_shm s p pc \<Longrightarrow> (\<exists>p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> Info_flow_shm s p' pc) \<or>
   583     hence p10: "\<not> path_by_shm s p path pc \<Longrightarrow> (\<exists>p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p'  pc) \<or>
   257       Info_flow_shm s pa pc"
   584       path_by_shm s pa pc"
   258       using p2' p7 p8 p5
   585       using p2' p7 p8 p5
   259       by (auto split:if_splits dest:Info_flow_shm.intros(2))      
   586       by (auto split:if_splits dest:path_by_shm.intros(2))      
   260   (*     apply (rule_tac x = pb in exI, simp add:one_flow_flows, rule_tac x = flagb in exI, simp)+  *)
   587   (*     apply (rule_tac x = pb in exI, simp add:one_flow_flows, rule_tac x = flagb in exI, simp)+  *) *)
   261     moreover have "pc = p \<Longrightarrow> (\<exists> p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s pa p') 
   588 
   262                             \<or> Info_flow_shm s pa pc"
   589     from p1 have a0: "(path = []) = (pa = pb)" using pbs_prop2 by simp
       
   590     have a1:"\<lbrakk>pa = p; flag = SHM_RDWR; \<not> path_by_shm s pa path pb\<rbrakk> \<Longrightarrow> 
       
   591       \<exists>p' flagb path'. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' path' pb \<and> path = path' @ [(p, h, p')]"
       
   592       using p2' by auto
       
   593     have b1: "\<lbrakk>pa = p; flag = SHM_RDWR; \<not> path_by_shm s pa path pc\<rbrakk> \<Longrightarrow> 
       
   594       \<exists>p' flagb path'. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' path' pc \<and>
       
   595         (pb, h', pc) # path = path' @ [(p, h, p')]"
       
   596       
       
   597       
       
   598       using p8 a1 p7 p5 a0 
       
   599       apply (auto split:if_splits elim:path_by_shm_intro4)
       
   600       apply (rule_tac x = pb in exI, rule conjI, rule_tac x = SHM_RDWR in exI, simp)
       
   601       apply (rule_tac x = pc in exI, rule conjI, rule_tac x = flagb in exI, simp)
       
   602       apply (rule_tac x = "[]" in exI, rule conjI)
       
   603 apply (erule path_by_shm_intro4, simp)
       
   604 
       
   605       apply (case_tac "path_by_shm s pa path pb", simp) defer
       
   606       apply (drule a1, simp+, clarsimp)
       
   607       apply (rule conjI, rule_tac x = flagb in exI, simp)
       
   608       apply (rule path_by_shm_
       
   609       using p2' p8 p5
       
   610       apply (auto split:if_splits dest!:pbs_prop2' simp:path_by_shm_intro4)
       
   611       apply (drule pbs_prop2', simp)
       
   612       apply (erule_tac x = pc in allE, simp add:path_by_shm_intro4)
       
   613      
       
   614       apply (drule_tac x = "pc" in allE)
       
   615       
       
   616       apply simp
       
   617 
       
   618       sorry
       
   619     moreover have "pc = p \<Longrightarrow> (\<exists>p' path'. path_by_shm s pa path' p' \<and>
       
   620              (pb, h', pc) # path = path' @ [(p', h, p)] \<and> (p', SHM_RDWR) \<in> procs_of_shm s h) \<or>
       
   621       (path_by_shm s pa path pc \<and> \<not> edge_related path p h)"
   263       using p2' p7 p8 p5
   622       using p2' p7 p8 p5
   264       by (auto split:if_splits intro:Info_flow_shm_intro3 simp:one_flow_shm_def)
   623       sorry (*
   265     moreover have "\<lbrakk>pc \<noteq> p; pa \<noteq> p \<or> flag \<noteq> SHM_RDWR\<rbrakk> \<Longrightarrow> (\<exists>p' flag'. Info_flow_shm s pa p \<and>
   624       apply (auto split:if_splits intro:path_by_shm_intro3 simp:one_flow_shm_def) *)
   266                           flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and> Info_flow_shm s p' pc) \<or>
   625     moreover have "\<lbrakk>pc \<noteq> p; pa \<noteq> p \<or> flag \<noteq> SHM_RDWR\<rbrakk> \<Longrightarrow> 
   267                       (\<exists>p'. Info_flow_shm s pa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s p pc) \<or>
   626       (\<exists>p' flag' pathaa pathb. path_by_shm s pa pathaa p \<and> flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and>
   268                       Info_flow_shm s pa pc"
   627            path_by_shm s p' pathb pc \<and> (pb, h', pc) # path = pathaa @ [(p, h, p')] @ pathb) \<or>
   269       using p2' p7 p8 p5
   628       (\<exists>p' pathaa pathb. path_by_shm s pa pathaa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and>
   270       apply (auto split:if_splits intro:Info_flow_shm_intro3 simp:one_flow_shm_def)
   629            path_by_shm s p pathb pc \<and> (pb, h', pc) # path = pathaa @ [(p', h, p)] @ pathb) \<or>
   271       apply (rule_tac x = pc in exI, simp add:Info_flow_shm_intro4)
   630       (path_by_shm s pa path pc \<and> \<not> edge_related path p h)"
       
   631       using p2' p7 p8 p5 (*
       
   632       apply (auto split:if_splits intro:path_by_shm_intro3 simp:one_flow_shm_def)
       
   633       apply (rule_tac x = pc in exI, simp add:path_by_shm_intro4)
   272       apply (rule_tac x = flagb in exI, simp)      
   634       apply (rule_tac x = flagb in exI, simp)      
   273       done
   635       done *)
   274     ultimately  show "if Info_flow_shm s pa pc then True
   636       sorry
   275        else if pa = p \<and> flag = SHM_RDWR then \<exists>p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> Info_flow_shm s p' pc
   637     ultimately  
   276             else if pc = p then \<exists>p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s pa p'
   638     show "if (pb, h', pc) # path = [] then pa = pc \<and> pa \<in> current_procs s
   277                  else (\<exists>p' flag'. Info_flow_shm s pa p \<and>
   639        else path_by_shm s pa ((pb, h', pc) # path) pc \<and> \<not> edge_related ((pb, h', pc) # path) p h \<or>
   278                           flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and> Info_flow_shm s p' pc) \<or>
   640        (if pa = p \<and> flag = SHM_RDWR
   279                       (\<exists>p'. Info_flow_shm s pa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s p pc)"
   641         then \<exists>p' flagb path'. (p', flagb) \<in> procs_of_shm s h \<and>
   280       using p7 by auto
   642                 path_by_shm s p' path' pc \<and> (pb, h', pc) # path = path' @ [(p, h, p')]
       
   643         else if pc = p
       
   644              then \<exists>p' path'. path_by_shm s pa path' p' \<and>
       
   645                      (pb, h', pc) # path = (p', h, p) # path' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h
       
   646              else (\<exists>p' flag' pathaa pathb. path_by_shm s pa pathaa p \<and> flag = SHM_RDWR \<and>
       
   647                       (p', flag') \<in> procs_of_shm s h \<and>
       
   648                       path_by_shm s p' pathb pc \<and> (pb, h', pc) # path = pathb @ [(p, h, p')] @ pathaa) \<or>
       
   649                   (\<exists>p' pathaa pathb. path_by_shm s pa pathaa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and>
       
   650                       path_by_shm s p pathb pc \<and> (pb, h', pc) # path = pathb @ [(p', h, p)] @ pathaa))"
       
   651         apply (auto split:if_splits)
       
   652         using p7 by auto
   281   qed
   653   qed
   282 qed
   654 qed
   283 
   655 
   284 lemma Info_flow_shm_attach1:
   656 lemma path_by_shm_attach1:
   285   "\<lbrakk>valid (Attach p h flag # s); Info_flow_shm (Attach p h flag # s) pa pb\<rbrakk>
   657   "\<lbrakk>valid (Attach p h flag # s); path_by_shm (Attach p h flag # s) pa pb\<rbrakk>
   286    \<Longrightarrow> (if Info_flow_shm s pa pb then True else
   658    \<Longrightarrow> (if path_by_shm s pa pb then True else
   287      (if (pa = p \<and> flag = SHM_RDWR) 
   659      (if (pa = p \<and> flag = SHM_RDWR) 
   288       then (\<exists> p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> Info_flow_shm s p' pb)
   660       then (\<exists> p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' pb)
   289       else if (pb = p) 
   661       else if (pb = p) 
   290            then (\<exists> p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s pa p')
   662            then (\<exists> p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> path_by_shm s pa p')
   291            else (\<exists> p' flag'. Info_flow_shm s pa p \<and> flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and> 
   663            else (\<exists> p' flag'. path_by_shm s pa p \<and> flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and> 
   292                              Info_flow_shm s p' pb) \<or>
   664                              path_by_shm s p' pb) \<or>
   293                 (\<exists> p'. Info_flow_shm s pa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s p pb)
   665                 (\<exists> p'. path_by_shm s pa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> path_by_shm s p pb)
   294      )  )"
   666      )  )"
   295 apply (drule_tac p = p and h = h and flag = flag in Info_flow_shm_attach1_aux)
   667 apply (drule_tac p = p and h = h and flag = flag in path_by_shm_attach1_aux)
   296 by auto
   668 by auto
   297 
   669 
   298 lemma Info_flow_shm_attach_aux[rule_format]:
   670 lemma path_by_shm_attach_aux[rule_format]:
   299   "Info_flow_shm s pa pb \<Longrightarrow> valid (Attach p h flag # s) \<longrightarrow> Info_flow_shm (Attach p h flag # s) pa pb"
   671   "path_by_shm s pa pb \<Longrightarrow> valid (Attach p h flag # s) \<longrightarrow> path_by_shm (Attach p h flag # s) pa pb"
   300 apply (erule Info_flow_shm.induct)
   672 apply (erule path_by_shm.induct)
   301 apply (rule impI, rule Info_flow_shm.intros(1), simp)
   673 apply (rule impI, rule path_by_shm.intros(1), simp)
   302 apply (rule impI, simp, rule_tac h = ha in Info_flow_shm.intros(2), simp)
   674 apply (rule impI, simp, rule_tac h = ha in path_by_shm.intros(2), simp)
   303 apply (auto simp add:one_flow_shm_simps)
   675 apply (auto simp add:one_flow_shm_simps)
   304 done
   676 done
   305 
   677 
   306 lemma Info_flow_shm_attach2:
   678 lemma path_by_shm_attach2:
   307   "\<lbrakk>valid (Attach p h flag # s); if Info_flow_shm s pa pb then True else
   679   "\<lbrakk>valid (Attach p h flag # s); if path_by_shm s pa pb then True else
   308      (if (pa = p \<and> flag = SHM_RDWR) 
   680      (if (pa = p \<and> flag = SHM_RDWR) 
   309       then (\<exists> p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> Info_flow_shm s p' pb)
   681       then (\<exists> p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' pb)
   310       else if (pb = p) 
   682       else if (pb = p) 
   311            then (\<exists> p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s pa p')
   683            then (\<exists> p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> path_by_shm s pa p')
   312            else (\<exists> p' flag'. Info_flow_shm s pa p \<and> flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and> 
   684            else (\<exists> p' flag'. path_by_shm s pa p \<and> flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and> 
   313                              Info_flow_shm s p' pb) \<or>
   685                              path_by_shm s p' pb) \<or>
   314                 (\<exists> p'. Info_flow_shm s pa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s p pb))\<rbrakk>
   686                 (\<exists> p'. path_by_shm s pa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> path_by_shm s p pb))\<rbrakk>
   315    \<Longrightarrow> Info_flow_shm (Attach p h flag # s) pa pb"
   687    \<Longrightarrow> path_by_shm (Attach p h flag # s) pa pb"
   316 apply (frule vt_grant_os, frule vd_cons)
   688 apply (frule vt_grant_os, frule vd_cons)
   317 apply (auto split:if_splits intro:Info_flow_shm_intro3 simp:one_flow_shm_def intro:Info_flow_shm_attach_aux)
   689 apply (auto split:if_splits intro:path_by_shm_intro3 simp:one_flow_shm_def intro:path_by_shm_attach_aux)
   318 apply (rule_tac p' = p' in Info_flow_trans)
   690 apply (rule_tac p' = p' in Info_flow_trans)
   319 apply (rule_tac p' = p and h = h in Info_flow_shm.intros(2))
   691 apply (rule_tac p' = p and h = h in path_by_shm.intros(2))
   320 apply (rule Info_flow_shm.intros(1), simp)
   692 apply (rule path_by_shm.intros(1), simp)
   321 apply (simp add:one_flow_shm_simps, simp add:one_flow_shm_def)
   693 apply (simp add:one_flow_shm_simps, simp add:one_flow_shm_def)
   322 apply (rule conjI, rule notI, simp, rule_tac x = flagb in exI, simp)
   694 apply (rule conjI, rule notI, simp, rule_tac x = flagb in exI, simp)
   323 apply (simp add:Info_flow_shm_attach_aux)
   695 apply (simp add:path_by_shm_attach_aux)
   324 
   696 
   325 apply (rule_tac p' = p' in Info_flow_trans)
   697 apply (rule_tac p' = p' in Info_flow_trans)
   326 apply (rule_tac p' = p in Info_flow_trans)
   698 apply (rule_tac p' = p in Info_flow_trans)
   327 apply (simp add:Info_flow_shm_attach_aux)
   699 apply (simp add:path_by_shm_attach_aux)
   328 apply (rule_tac p' = p and h = h in Info_flow_shm.intros(2))
   700 apply (rule_tac p' = p and h = h in path_by_shm.intros(2))
   329 apply (rule Info_flow_shm.intros(1), simp)
   701 apply (rule path_by_shm.intros(1), simp)
   330 apply (simp add:one_flow_shm_simps, simp add:one_flow_shm_def)
   702 apply (simp add:one_flow_shm_simps, simp add:one_flow_shm_def)
   331 apply (rule conjI, rule notI, simp, rule_tac x = flag' in exI, simp)
   703 apply (rule conjI, rule notI, simp, rule_tac x = flag' in exI, simp)
   332 apply (simp add:Info_flow_shm_attach_aux)
   704 apply (simp add:path_by_shm_attach_aux)
   333 
   705 
   334 apply (rule_tac p' = p in Info_flow_trans)
   706 apply (rule_tac p' = p in Info_flow_trans)
   335 apply (rule_tac p' = p' in Info_flow_trans)
   707 apply (rule_tac p' = p' in Info_flow_trans)
   336 apply (simp add:Info_flow_shm_attach_aux)
   708 apply (simp add:path_by_shm_attach_aux)
   337 apply (rule_tac p' = p' and h = h in Info_flow_shm.intros(2))
   709 apply (rule_tac p' = p' and h = h in path_by_shm.intros(2))
   338 apply (rule Info_flow_shm.intros(1), simp add:procs_of_shm_prop2)
   710 apply (rule path_by_shm.intros(1), simp add:procs_of_shm_prop2)
   339 apply (simp add:one_flow_shm_simps, simp add:one_flow_shm_def)
   711 apply (simp add:one_flow_shm_simps, simp add:one_flow_shm_def)
   340 apply (rule notI, simp)
   712 apply (rule notI, simp)
   341 apply (simp add:Info_flow_shm_attach_aux)
   713 apply (simp add:path_by_shm_attach_aux)
   342 
   714 
   343 apply (rule_tac p' = p in Info_flow_trans)
   715 apply (rule_tac p' = p in Info_flow_trans)
   344 apply (rule_tac p' = p' in Info_flow_trans)
   716 apply (rule_tac p' = p' in Info_flow_trans)
   345 apply (simp add:Info_flow_shm_attach_aux)
   717 apply (simp add:path_by_shm_attach_aux)
   346 apply (rule_tac p' = p' and h = h in Info_flow_shm.intros(2))
   718 apply (rule_tac p' = p' and h = h in path_by_shm.intros(2))
   347 apply (rule Info_flow_shm.intros(1), simp add:procs_of_shm_prop2)
   719 apply (rule path_by_shm.intros(1), simp add:procs_of_shm_prop2)
   348 apply (simp add:one_flow_shm_simps, simp add:one_flow_shm_def)
   720 apply (simp add:one_flow_shm_simps, simp add:one_flow_shm_def)
   349 apply (rule notI, simp)
   721 apply (rule notI, simp)
   350 apply (simp add:Info_flow_shm_attach_aux)
   722 apply (simp add:path_by_shm_attach_aux)
   351 done
   723 done
   352 
   724 
   353 lemma Info_flow_shm_attach:
   725 lemma path_by_shm_attach:
   354   "valid (Attach p h flag # s) \<Longrightarrow> Info_flow_shm (Attach p h flag # s) = (\<lambda> pa pb. 
   726   "valid (Attach p h flag # s) \<Longrightarrow> path_by_shm (Attach p h flag # s) = (\<lambda> pa pb. 
   355      Info_flow_shm s pa pb \<or>
   727      path_by_shm s pa pb \<or>
   356      (if (pa = p \<and> flag = SHM_RDWR) 
   728      (if (pa = p \<and> flag = SHM_RDWR) 
   357       then (\<exists> p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> Info_flow_shm s p' pb)
   729       then (\<exists> p' flagb. (p', flagb) \<in> procs_of_shm s h \<and> path_by_shm s p' pb)
   358       else if (pb = p) 
   730       else if (pb = p) 
   359            then (\<exists> p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s pa p')
   731            then (\<exists> p'. (p', SHM_RDWR) \<in> procs_of_shm s h \<and> path_by_shm s pa p')
   360            else (\<exists> p' flag'. Info_flow_shm s pa p \<and> flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and> 
   732            else (\<exists> p' flag'. path_by_shm s pa p \<and> flag = SHM_RDWR \<and> (p', flag') \<in> procs_of_shm s h \<and> 
   361                              Info_flow_shm s p' pb) \<or>
   733                              path_by_shm s p' pb) \<or>
   362                 (\<exists> p'. Info_flow_shm s pa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> Info_flow_shm s p pb)
   734                 (\<exists> p'. path_by_shm s pa p' \<and> (p', SHM_RDWR) \<in> procs_of_shm s h \<and> path_by_shm s p pb)
   363      )  )"
   735      )  )"
   364 apply (rule ext, rule ext, rule iffI)
   736 apply (rule ext, rule ext, rule iffI)
   365 apply (drule_tac pa = pa and pb = pb in Info_flow_shm_attach1, simp)
   737 apply (drule_tac pa = pa and pb = pb in path_by_shm_attach1, simp)
   366 apply (auto split:if_splits)[1]
   738 apply (auto split:if_splits)[1]
   367 apply (drule_tac pa = pa and pb = pb in Info_flow_shm_attach2)
   739 apply (drule_tac pa = pa and pb = pb in path_by_shm_attach2)
   368 apply (auto split:if_splits)
   740 apply (auto split:if_splits)
   369 done
   741 done
       
   742 
       
   743 
       
   744 
       
   745 
       
   746 
       
   747 
       
   748 
       
   749 
       
   750 
       
   751 
       
   752 
       
   753 
       
   754 
       
   755 
       
   756 
       
   757 
       
   758 
       
   759 
       
   760 
       
   761 
       
   762 
       
   763 
   370 
   764 
   371 
   765 
   372 lemma info_flow_shm_detach:
   766 lemma info_flow_shm_detach:
   373   "valid (Detach p h # s) \<Longrightarrow> info_flow_shm (Detach p h # s) = (\<lambda> pa pb. 
   767   "valid (Detach p h # s) \<Longrightarrow> info_flow_shm (Detach p h # s) = (\<lambda> pa pb. 
   374      self_shm s pa pb \<or> ((p = pa \<or> p = pb) \<and> (\<exists> h'. h' \<noteq> h \<and> one_flow_shm s h' pa pb)) \<or>
   768      self_shm s pa pb \<or> ((p = pa \<or> p = pb) \<and> (\<exists> h'. h' \<noteq> h \<and> one_flow_shm s h' pa pb)) \<or>
   387      (pa = p' \<and> pb = p') \<or> (pa = p' \<and> pb \<noteq> p' \<and> (\<exists> h \<in> shms. one_flow_shm s h p pb)) \<or> 
   781      (pa = p' \<and> pb = p') \<or> (pa = p' \<and> pb \<noteq> p' \<and> (\<exists> h \<in> shms. one_flow_shm s h p pb)) \<or> 
   388      (pb = p' \<and> pa \<noteq> p' \<and> (\<exists> h \<in> shms. one_flow_shm s h pa p)) \<or> 
   782      (pb = p' \<and> pa \<noteq> p' \<and> (\<exists> h \<in> shms. one_flow_shm s h pa p)) \<or> 
   389      (pa \<noteq> p' \<and> pb \<noteq> p' \<and> info_flow_shm s pa pb))"
   783      (pa \<noteq> p' \<and> pb \<noteq> p' \<and> info_flow_shm s pa pb))"
   390 apply (rule ext, rule ext, frule vt_grant_os, frule vd_cons, clarsimp)
   784 apply (rule ext, rule ext, frule vt_grant_os, frule vd_cons, clarsimp)
   391 apply (frule_tac p = p' in procs_of_shm_prop2', simp)
   785 apply (frule_tac p = p' in procs_of_shm_prop2', simp)
       
   786 sorry (*
   392 apply (auto simp:info_flow_shm_def one_flow_shm_def)
   787 apply (auto simp:info_flow_shm_def one_flow_shm_def)
   393 done
   788 done *)
   394 
   789 
   395 lemma info_flow_shm_execve:
   790 lemma info_flow_shm_execve:
   396   "valid (Execve p f fds # s) \<Longrightarrow> info_flow_shm (Execve p f fds # s) = (\<lambda> pa pb. 
   791   "valid (Execve p f fds # s) \<Longrightarrow> info_flow_shm (Execve p f fds # s) = (\<lambda> pa pb. 
   397      (pa = p \<and> pb = p) \<or> (pa \<noteq> p \<and> pb \<noteq> p \<and> info_flow_shm s pa pb)    )"
   792      (pa = p \<and> pb = p) \<or> (pa \<noteq> p \<and> pb \<noteq> p \<and> info_flow_shm s pa pb)    )"
   398 apply (rule ext, rule ext, frule vt_grant_os, frule vd_cons)
   793 apply (rule ext, rule ext, frule vt_grant_os, frule vd_cons)
   485 apply (erule info_flow_shm_cases, simp, simp split:if_splits)
   880 apply (erule info_flow_shm_cases, simp, simp split:if_splits)
   486 apply (rule_tac p = pa and p' = p' in info_flow_shm.intros(2), simp+)
   881 apply (rule_tac p = pa and p' = p' in info_flow_shm.intros(2), simp+)
   487 apply (rule notI, erule info_flow_shm.cases, simp+)
   882 apply (rule notI, erule info_flow_shm.cases, simp+)
   488 pr 5
   883 pr 5
   489 *)
   884 *)
   490 lemmas info_flow_shm_simps = info_flow_shm_other info_flow_shm_attach info_flow_shm_detach info_flow_shm_deleteshm
   885 lemmas info_flow_shm_simps = info_flow_shm_other (* info_flow_shm_attach *) info_flow_shm_detach info_flow_shm_deleteshm
   491   info_flow_shm_clone info_flow_shm_execve info_flow_shm_kill info_flow_shm_exit
   886   info_flow_shm_clone info_flow_shm_execve info_flow_shm_kill info_flow_shm_exit
   492 
   887 
   493 lemma has_same_inode_in_current:
   888 lemma has_same_inode_in_current:
   494   "\<lbrakk>has_same_inode s f f'; valid s\<rbrakk> \<Longrightarrow> f \<in> current_files s \<and> f' \<in> current_files s"
   889   "\<lbrakk>has_same_inode s f f'; valid s\<rbrakk> \<Longrightarrow> f \<in> current_files s \<and> f' \<in> current_files s"
   495 by (auto simp add:has_same_inode_def current_files_def)
   890 by (auto simp add:has_same_inode_def current_files_def)