diff -r de8733706a06 -r f20a798cdf7d Tainted_prop.thy --- a/Tainted_prop.thy Thu Jun 06 14:50:52 2013 +0800 +++ b/Tainted_prop.thy Sat Jun 08 09:59:33 2013 +0800 @@ -4,6 +4,93 @@ context tainting begin +fun Tainted :: "t_state \ t_object set" +where + "Tainted [] = seeds" +| "Tainted (Clone p p' fds shms # s) = + (if (O_proc p) \ Tainted s then Tainted s \ {O_proc p'} else Tainted s)" +| "Tainted (Execve p f fds # s) = + (if (O_file f) \ Tainted s then Tainted s \ {O_proc p} else Tainted s)" +| "Tainted (Kill p p' # s) = Tainted s - {O_proc p'}" +| "Tainted (Ptrace p p' # s) = + (if (O_proc p) \ Tainted s + then Tainted s \ {O_proc p'' | p''. info_flow_shm s p' p''} + else if (O_proc p') \ Tainted s + then Tainted s \ {O_proc p'' | p''. info_flow_shm s p p''} + else Tainted s)" +| "Tainted (Exit p # s) = Tainted s - {O_proc p}" +| "Tainted (Open p f flags fd opt # s) = + (case opt of + Some inum \ (if (O_proc p) \ Tainted s + then Tainted s \ {O_file f} + else Tainted s) + | _ \ Tainted s)" +| "Tainted (ReadFile p fd # s) = + (case (file_of_proc_fd s p fd) of + Some f \ if (O_file f) \ Tainted s + then Tainted s \ {O_proc p' | p'. info_flow_shm s p p'} + else Tainted s + | None \ Tainted s)" +| "Tainted (WriteFile p fd # s) = + (case (file_of_proc_fd s p fd) of + Some f \ if (O_proc p) \ Tainted s + then Tainted s \ {O_file f' | f'. has_same_inode s f f'} + else Tainted s + | None \ Tainted s)" +| "Tainted (CloseFd p fd # s) = + (case (file_of_proc_fd s p fd) of + Some f \ ( if ((proc_fd_of_file s f = {(p,fd)}) \ (f \ files_hung_by_del s)) + then Tainted s - {O_file f} else Tainted s ) + | _ \ Tainted s)" +| "Tainted (UnLink p f # s) = Tainted s - {O_file f}" +| "Tainted (LinkHard p f f' # s) = + (if (O_file f \ Tainted s) then Tainted s \ {O_file f'} else Tainted s)" +| "Tainted (Truncate p f len # s) = + (if (len > 0 \ O_proc p \ Tainted s) + then Tainted s \ {O_file f' | f'. has_same_inode s f f'} + else Tainted s)" +| "Tainted (SendMsg p q m # s) = + (if (O_proc p \ Tainted s) then Tainted s \ {O_msg q m} else Tainted s)" +| "Tainted (RecvMsg p q m # s) = + (if (O_msg q m \ Tainted s) + then Tainted s \ {O_proc p' | p'. info_flow_shm s p p'} - {O_msg q m} + else Tainted s)" +| "Tainted (RemoveMsgq p q # s) = Tainted s - {O_msg q m| m. O_msg q m \ Tainted s}" +| "Tainted (e # s) = Tainted s" + +lemma valid_Tainted_obj: + "\obj \ Tainted s; valid s\ \ (\ f. obj \ O_dir f) \ (\ q. obj \ O_msgq q) \ (\ h. obj \ O_shm h) \ (\ p fd. obj \ O_fd p fd) \ (\ s. obj \ O_tcp_sock s) \ (\ s. obj \ O_udp_sock s)" +apply (induct s, simp, drule seeds_in_init, case_tac obj, simp+) +apply (frule vd_cons, frule vt_grant_os, case_tac a) +apply (auto split:if_splits option.splits) +done + +lemma Tainted_in_current: + "\obj \ Tainted s; valid s\ \ alive s obj" +apply (induct s, simp) +apply (drule seeds_in_init, case_tac obj, simp_all add:is_file_nil) +apply (frule vd_cons, frule valid_Tainted_obj, simp, frule vt_grant_os, case_tac a) +apply (auto simp:alive_simps split:if_splits option.splits t_object.splits) +apply (auto intro:has_same_inode_prop2 has_same_inode_prop1) +apply + +lemma Tainted_imp_tainted: + "\obj \ Tainted s; valid s\ \ obj \ tainted s" +apply (induct s arbitrary:obj, rule t_init, simp) +apply (frule vd_cons, frule vt_grant_os) +apply (case_tac a) +apply (auto intro!:t_init t_clone t_execve t_cfile t_read t_write t_link t_trunc t_sendmsg t_recvmsg + split:if_splits option.splits) +pr 25 + +lemma tainted_imp_Tainted: + "obj \ tainted s \ obj \ Tainted s" + + + + + + lemma tainted_in_current: "obj \ tainted s \ alive s obj" apply (erule tainted.induct, auto dest:vt_grant_os vd_cons simp:is_file_simps) @@ -27,7 +114,7 @@ done lemma valid_tainted_obj: - "obj \ tainted s \ (\ f. obj \ O_dir f) \ (\ q. obj \ O_msgq q) \ (\ h. obj \ O_shm h)" + "obj \ tainted s \ (\ f. obj \ O_dir f) \ (\ q. obj \ O_msgq q) \ (\ h. obj \ O_shm h) \ (\ p fd. obj \ O_fd p fd) \ (\ s. obj \ O_tcp_sock s) \ (\ s. obj \ O_udp_sock s)" apply (erule tainted.induct) apply (drule seeds_in_init) by auto