diff -r 566b0d1c3669 -r 259a50be4381 Current_prop.thy --- a/Current_prop.thy Sun Jun 16 08:05:37 2013 +0800 +++ b/Current_prop.thy Mon Jun 24 09:01:42 2013 +0800 @@ -49,7 +49,66 @@ lemma info_shm_flow_in_procs: "\info_flow_shm s p p'; valid s\ \ p \ current_procs s \ p' \ current_procs s" -by (auto simp:info_flow_shm_def intro:procs_of_shm_prop2) +apply (induct rule:info_flow_shm.induct ) +by (auto intro:procs_of_shm_prop2) + +(*********** simpset for info_flow_shm **************) + +lemma info_flow_shm_prop1: + "\info_flow_shm s p p'; p \ p'; valid s\ + \ \ h h' flag. (p, SHM_RDWR) \ procs_of_shm s h \ (p', flag) \ procs_of_shm s h'" +by (induct rule: info_flow_shm.induct, auto) + +lemma info_flow_shm_cases: + "\info_flow_shm \ pa pb; \p s. \s = \ ; pa = p; pb = p; p \ current_procs s\ \ P; + \s p p' h p'' flag. \s = \; pa = p; pb = p''; info_flow_shm s p p'; (p', SHM_RDWR) \ procs_of_shm s h; + (p'', flag) \ procs_of_shm s h\\ P\ + \ P" +by (erule info_flow_shm.cases, auto) + +definition one_flow_shm :: "t_state \ t_process \ t_process \ bool" +where + "one_flow_shm s p p' \ p \ p' \ (\ h flag. (p, SHM_RDWR) \ procs_of_shm s h \ (p', flag) \ procs_of_shm s h)" + +inductive flows_shm :: "t_state \ t_process \ t_process \ bool" +where + "p \ current_procs s \ flows_shm s p p" +| "\flows_shm s p p'; one_flow_shm s p' p''\ \ flows_shm s p p''" + +definition attached_procs :: "t_state \ t_shm \ t_process set" +where + "attached_procs s h \ {p. \ flag. (p, flag) \ procs_of_shm s h}" + +definition flowed_procs:: "t_state \ t_shm \ t_process set" +where + "flowed_procs s h \ {p'. \ p \ attached_procs s h. flows_shm s p p'}" + +fun Info_flow_shm :: "t_state \ t_process \ t_process set" +where + "Info_flow_shm [] = (\ p. {p'. flows_shm [] p p'})" +| "Info_flow_shm (Attach p h flag # s) = (\ p'. if (p' = p) then {p''. \ }" + + +lemma info_flow_shm_attach: + "valid (Attach p h flag # s) \ info_flow_shm (Attach p h flag # s) = (\ pa pb. (info_flow_shm s pa pb) \ + (if (pa = p) + then (if (flag = SHM_RDWR) + then (\ flag. (pb, flag) \ procs_of_shm s h) + else (pb = p)) + else (if (pb = p) + then (pa, SHM_RDWR) \ procs_of_shm s h + else info_flow_shm s pa pb)) )" +apply (frule vd_cons, frule vt_grant_os, rule ext, rule ext) +apply (case_tac "info_flow_shm s pa pb", simp) + +thm info_flow_shm.cases +apply (auto split:if_splits intro:info_flow_shm.intros elim:info_flow_shm_cases) +apply (erule info_flow_shm_cases, simp, simp split:if_splits) +apply (rule_tac p = pa and p' = p' in info_flow_shm.intros(2), simp+) +apply (rule notI, erule info_flow_shm.cases, simp+) +pr 5 + +lemmas info_flow_shm_simps = info_flow_shm_other lemma has_same_inode_in_current: "\has_same_inode s f f'; valid s\ \ f \ current_files s \ f' \ current_files s"