# HG changeset patch # User chunhan # Date 1372035702 -28800 # Node ID 259a50be438128fa2cb28f22c0ee16903a88c588 # Parent 566b0d1c3669002e01901938a4f91c75afc7af00 wrong of info-flow-shm, it is a inductive(transitive) notion, not a simple relation just between 2 nodes, more information, see 5.7 of ideas_of_selinux.txt 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" diff -r 566b0d1c3669 -r 259a50be4381 Flask.thy --- a/Flask.thy Sun Jun 16 08:05:37 2013 +0800 +++ b/Flask.thy Mon Jun 24 09:01:42 2013 +0800 @@ -473,10 +473,17 @@ (\ h. (procs_of_shm s h) - {(p, flag) | flag. (p, flag) \ procs_of_shm s h})" | "procs_of_shm (e # \) = procs_of_shm \" +inductive info_flow_shm :: "t_state \ t_process \ t_process \ bool" +where + "p \ current_procs s \ info_flow_shm s p p" +| "\info_flow_shm s p p'; (p', SHM_RDWR) \ procs_of_shm s h; (p'', flag) \ procs_of_shm s h; p' \ p''\ + \ info_flow_shm s p p''" +(* definition info_flow_shm :: "t_state \ t_process \ t_process \ bool" where "info_flow_shm \ from to \ (from = to \ from \ current_procs \) \ (\ h toflag. (((from, SHM_RDWR) \ procs_of_shm \ h) \ ((to, toflag) \ procs_of_shm \ h)))" +*) fun current_msgqs :: "t_state \ t_msg set" where diff -r 566b0d1c3669 -r 259a50be4381 Static.thy --- a/Static.thy Sun Jun 16 08:05:37 2013 +0800 +++ b/Static.thy Mon Jun 24 09:01:42 2013 +0800 @@ -463,9 +463,16 @@ where "info_flow_sproc_sshms shms shms' \ (\ sh flag'. (sh, SHM_RDWR) \ shms \ (sh, flag') \ shms')" +(* fun info_flow_sshm :: "t_sproc \ t_sproc \ bool" where "info_flow_sshm (pi,sec,fds,shms) (pi',sec',fds',shms') = info_flow_sproc_sshms shms shms'" +*) +inductive info_flow_sshm :: "t_sproc \ t_sproc \ bool" +where + "info_flow_sshm sp sp" +| "\info_flow_sshm sp (pi,sec,fds,shms); info_flow_sproc_sshms shms shms'\ + \ info_flow_sshm sp (pi',sec',fds',shms')" definition update_ss_shms :: "t_static_state \ t_sproc \ bool \ t_static_state" where diff -r 566b0d1c3669 -r 259a50be4381 Tainted_prop.thy --- a/Tainted_prop.thy Sun Jun 16 08:05:37 2013 +0800 +++ b/Tainted_prop.thy Mon Jun 24 09:01:42 2013 +0800 @@ -95,6 +95,13 @@ (*?? need simpset for tainted *) sorry +lemma info_flow_shm_intro: + "\(p, SHM_RDWR) \ procs_of_shm s h; (p', flag) \ procs_of_shm s h; valid s\ + \ info_flow_shm s p p'" +apply (rule_tac p = p and p' = p in info_flow_shm.intros(2)) +apply (auto intro!:info_flow_shm.intros(1) procs_of_shm_prop2) +done + lemma info_flow_shm_Tainted: "\O_proc p \ Tainted s; info_flow_shm s p p'; valid s\ \ O_proc p' \ Tainted s" proof (induct s arbitrary:p p') @@ -109,7 +116,7 @@ have p4': "\ p p' h flag. \O_proc p \ Tainted s; (p, SHM_RDWR) \ procs_of_shm s h; (p', flag) \ procs_of_shm s h\ \ O_proc p' \ Tainted s" - by (rule p4, auto simp:info_flow_shm_def) + by (rule p4, auto intro!:info_flow_shm_intro p5) from p2 p3 have p7: "p \ current_procs (e # s)" and p8: "p' \ current_procs (e # s)" by (auto dest:info_shm_flow_in_procs) show ?case @@ -120,8 +127,12 @@ with p1 p2 p5 p6 p7 p8 p3 show ?thesis apply (case_tac e) prefer 7 - apply (simp add:info_flow_shm_def split:if_splits option.splits) + apply (simp split:if_splits option.splits) apply (rule allI|rule impI|rule conjI)+ + + + apply (auto dest:p4' procs_of_shm_prop2 Tainted_in_current split:if_splits option.splits)[1] + apply (erule disjE, drule_tac p = p and p' = p' in p4', simp+) apply (erule disjE, rule disjI2, rule disjI2, rule_tac x = h in exI, simp, rule_tac x= toflag in exI, simp) apply ((erule exE|erule conjE)+)