| author | chunhan | 
| Mon, 23 Dec 2013 19:47:17 +0800 | |
| changeset 82 | 0a68c605ae79 | 
| parent 65 | 6f9a588bcfc4 | 
| permissions | -rw-r--r-- | 
| 18 | 1  | 
theory Tainted_prop  | 
| 31 | 2  | 
imports Main Flask Flask_type Init_prop Current_files_prop Current_sockets_prop Delete_prop Proc_fd_of_file_prop Current_prop Alive_prop  | 
| 18 | 3  | 
begin  | 
4  | 
||
| 34 | 5  | 
ML {*quick_and_dirty := true*}
 | 
6  | 
||
| 18 | 7  | 
context tainting begin  | 
8  | 
||
| 22 | 9  | 
fun Tainted :: "t_state \<Rightarrow> t_object set"  | 
10  | 
where  | 
|
11  | 
"Tainted [] = seeds"  | 
|
12  | 
| "Tainted (Clone p p' fds shms # s) =  | 
|
13  | 
     (if (O_proc p) \<in> Tainted s then Tainted s \<union> {O_proc p'} else Tainted s)"
 | 
|
14  | 
| "Tainted (Execve p f fds # s) =  | 
|
15  | 
     (if (O_file f) \<in> Tainted s then Tainted s \<union> {O_proc p} else Tainted s)"
 | 
|
16  | 
| "Tainted (Kill p p' # s) = Tainted s - {O_proc p'}"
 | 
|
17  | 
| "Tainted (Ptrace p p' # s) =  | 
|
| 43 | 18  | 
(if (O_proc p \<in> Tainted s)  | 
| 22 | 19  | 
      then Tainted s \<union> {O_proc p'' | p''. info_flow_shm s p' p''}
 | 
| 43 | 20  | 
else if (O_proc p' \<in> Tainted s)  | 
| 22 | 21  | 
           then Tainted s \<union> {O_proc p'' | p''. info_flow_shm s p p''}
 | 
22  | 
else Tainted s)"  | 
|
23  | 
| "Tainted (Exit p # s) = Tainted s - {O_proc p}"
 | 
|
24  | 
| "Tainted (Open p f flags fd opt # s) =  | 
|
25  | 
(case opt of  | 
|
26  | 
Some inum \<Rightarrow> (if (O_proc p) \<in> Tainted s  | 
|
27  | 
                      then Tainted s \<union> {O_file f}
 | 
|
28  | 
else Tainted s)  | 
|
29  | 
| _ \<Rightarrow> Tainted s)"  | 
|
30  | 
| "Tainted (ReadFile p fd # s) =  | 
|
31  | 
(case (file_of_proc_fd s p fd) of  | 
|
32  | 
Some f \<Rightarrow> if (O_file f) \<in> Tainted s  | 
|
33  | 
                  then Tainted s \<union> {O_proc p' | p'. info_flow_shm s p p'}
 | 
|
34  | 
else Tainted s  | 
|
35  | 
| None \<Rightarrow> Tainted s)"  | 
|
36  | 
| "Tainted (WriteFile p fd # s) =  | 
|
37  | 
(case (file_of_proc_fd s p fd) of  | 
|
38  | 
Some f \<Rightarrow> if (O_proc p) \<in> Tainted s  | 
|
| 65 | 39  | 
                  then Tainted s \<union> {O_file f' | f'. f' \<in> same_inode_files s f}
 | 
| 22 | 40  | 
else Tainted s  | 
41  | 
| None \<Rightarrow> Tainted s)"  | 
|
42  | 
| "Tainted (CloseFd p fd # s) =  | 
|
43  | 
(case (file_of_proc_fd s p fd) of  | 
|
44  | 
        Some f \<Rightarrow> ( if ((proc_fd_of_file s f = {(p,fd)}) \<and> (f \<in> files_hung_by_del s))
 | 
|
45  | 
                    then Tainted s - {O_file f} else Tainted s )
 | 
|
46  | 
| _ \<Rightarrow> Tainted s)"  | 
|
| 
23
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
47  | 
| "Tainted (UnLink p f # s) =  | 
| 
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
48  | 
     (if (proc_fd_of_file s f = {}) then Tainted s - {O_file f} else Tainted s)"
 | 
| 22 | 49  | 
| "Tainted (LinkHard p f f' # s) =  | 
50  | 
     (if (O_file f \<in> Tainted s) then Tainted s \<union> {O_file f'} else Tainted s)"
 | 
|
51  | 
| "Tainted (Truncate p f len # s) =  | 
|
52  | 
(if (len > 0 \<and> O_proc p \<in> Tainted s)  | 
|
| 65 | 53  | 
      then Tainted s \<union> {O_file f' | f'. f' \<in> same_inode_files s f}
 | 
| 22 | 54  | 
else Tainted s)"  | 
| 27 | 55  | 
| "Tainted (Attach p h flag # s) =  | 
56  | 
(if (O_proc p \<in> Tainted s \<and> flag = SHM_RDWR)  | 
|
| 56 | 57  | 
      then Tainted s \<union> {O_proc p'' | p' flag' p''. (p', flag') \<in> procs_of_shm s h \<and> info_flow_shm s p' p''}
 | 
| 27 | 58  | 
else if (\<exists> p'. O_proc p' \<in> Tainted s \<and> (p', SHM_RDWR) \<in> procs_of_shm s h)  | 
| 56 | 59  | 
           then Tainted s \<union> {O_proc p'| p'. info_flow_shm s p p'}
 | 
| 27 | 60  | 
else Tainted s)"  | 
| 22 | 61  | 
| "Tainted (SendMsg p q m # s) =  | 
62  | 
     (if (O_proc p \<in> Tainted s) then Tainted s \<union> {O_msg q m} else Tainted s)"
 | 
|
63  | 
| "Tainted (RecvMsg p q m # s) =  | 
|
64  | 
(if (O_msg q m \<in> Tainted s)  | 
|
| 
23
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
65  | 
      then (Tainted s \<union> {O_proc p' | p'. info_flow_shm s p p'}) - {O_msg q m}
 | 
| 22 | 66  | 
else Tainted s)"  | 
67  | 
| "Tainted (RemoveMsgq p q # s) = Tainted s - {O_msg q m| m. O_msg q m \<in> Tainted s}"
 | 
|
68  | 
| "Tainted (e # s) = Tainted s"  | 
|
69  | 
||
70  | 
lemma valid_Tainted_obj:  | 
|
71  | 
"\<lbrakk>obj \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> (\<forall> f. obj \<noteq> O_dir f) \<and> (\<forall> q. obj \<noteq> O_msgq q) \<and> (\<forall> h. obj \<noteq> O_shm h) \<and> (\<forall> p fd. obj \<noteq> O_fd p fd) \<and> (\<forall> s. obj \<noteq> O_tcp_sock s) \<and> (\<forall> s. obj \<noteq> O_udp_sock s)"  | 
|
72  | 
apply (induct s, simp, drule seeds_in_init, case_tac obj, simp+)  | 
|
73  | 
apply (frule vd_cons, frule vt_grant_os, case_tac a)  | 
|
74  | 
apply (auto split:if_splits option.splits)  | 
|
75  | 
done  | 
|
76  | 
||
77  | 
lemma Tainted_in_current:  | 
|
78  | 
"\<lbrakk>obj \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> alive s obj"  | 
|
79  | 
apply (induct s, simp)  | 
|
80  | 
apply (drule seeds_in_init, case_tac obj, simp_all add:is_file_nil)  | 
|
81  | 
apply (frule vd_cons, frule valid_Tainted_obj, simp, frule vt_grant_os, case_tac a)  | 
|
| 
23
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
82  | 
apply (auto simp:alive_simps split:if_splits option.splits t_object.splits  | 
| 65 | 83  | 
intro:same_inode_files_prop1 procs_of_shm_prop2  | 
| 27 | 84  | 
dest:info_shm_flow_in_procs)  | 
| 65 | 85  | 
apply (auto simp:same_inode_files_def is_file_def split:if_splits)  | 
| 
29
 
622516c0fe34
path_by_shm reconstrain path without duplicated process AND shm
 
chunhan 
parents: 
27 
diff
changeset
 | 
86  | 
done  | 
| 
23
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
87  | 
|
| 
24
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
88  | 
lemma Tainted_proc_in_current:  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
89  | 
"\<lbrakk>O_proc p \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> p \<in> current_procs s"  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
90  | 
by (drule Tainted_in_current, simp+)  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
91  | 
|
| 
25
 
259a50be4381
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
 
chunhan 
parents: 
24 
diff
changeset
 | 
92  | 
|
| 
23
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
93  | 
lemma info_flow_shm_Tainted:  | 
| 
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
94  | 
"\<lbrakk>O_proc p \<in> Tainted s; info_flow_shm s p p'; valid s\<rbrakk> \<Longrightarrow> O_proc p' \<in> Tainted s"  | 
| 
24
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
95  | 
proof (induct s arbitrary:p p')  | 
| 
23
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
96  | 
case Nil  | 
| 
24
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
97  | 
thus ?case by (simp add:flow_shm_in_seeds)  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
98  | 
next  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
99  | 
case (Cons e s)  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
100  | 
hence p1: "O_proc p \<in> Tainted (e # s)" and p2: "info_flow_shm (e # s) p p'" and p3: "valid (e # s)"  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
101  | 
and p4: "\<And> p p'. \<lbrakk>O_proc p \<in> Tainted s; info_flow_shm s p p'\<rbrakk> \<Longrightarrow> O_proc p' \<in> Tainted s"  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
102  | 
and p5: "valid s" and p6: "os_grant s e"  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
103  | 
by (auto dest:vd_cons intro:vd_cons vt_grant_os)  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
104  | 
have p4':  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
105  | 
"\<And> p p' h flag. \<lbrakk>O_proc p \<in> Tainted s; (p, SHM_RDWR) \<in> procs_of_shm s h; (p', flag) \<in> procs_of_shm s h\<rbrakk>  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
106  | 
\<Longrightarrow> O_proc p' \<in> Tainted s"  | 
| 31 | 107  | 
by (rule p4, auto simp:info_flow_shm_def one_flow_shm_def procs_of_shm_prop2 p5)  | 
108  | 
from p2 p3 have p7: "p \<in> current_procs (e # s)" and p8: "p' \<in> current_procs (e # s)"  | 
|
109  | 
by (auto dest:info_shm_flow_in_procs)  | 
|
| 
24
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
110  | 
show ?case  | 
| 27 | 111  | 
proof (cases "self_shm s p p'")  | 
| 
24
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
112  | 
case True with p1 show ?thesis by simp  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
113  | 
next  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
114  | 
case False  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
115  | 
with p1 p2 p5 p6 p7 p8 p3 show ?thesis  | 
| 31 | 116  | 
apply (case_tac e)(*  | 
| 
24
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
117  | 
prefer 7  | 
| 27 | 118  | 
apply (simp add:info_flow_shm_simps split:if_splits option.splits)  | 
119  | 
apply (rule allI|rule impI|rule conjI)+  | 
|
120  | 
apply simp  | 
|
121  | 
apply (case_tac "O_proc p \<in> Tainted s", drule_tac p'=p' in p4, simp+)  | 
|
122  | 
apply simp  | 
|
123  | 
||
124  | 
||
125  | 
||
126  | 
||
127  | 
apply (auto simp:info_flow_shm_simps one_flow_shm_def dest:Tainted_in_current  | 
|
128  | 
intro:p4 p4' split:if_splits option.splits)  | 
|
129  | 
apply (auto simp:info_flow_shm_def one_flow_shm_def)  | 
|
130  | 
||
131  | 
||
132  | 
||
133  | 
apply (auto simp:one_flow_shm_def intro:p4 p4' split:if_splits option.splits)  | 
|
134  | 
||
135  | 
||
136  | 
||
137  | 
prefer 7  | 
|
| 
25
 
259a50be4381
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
 
chunhan 
parents: 
24 
diff
changeset
 | 
138  | 
apply (simp split:if_splits option.splits)  | 
| 
24
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
139  | 
apply (rule allI|rule impI|rule conjI)+  | 
| 
25
 
259a50be4381
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
 
chunhan 
parents: 
24 
diff
changeset
 | 
140  | 
|
| 
 
259a50be4381
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
 
chunhan 
parents: 
24 
diff
changeset
 | 
141  | 
|
| 
 
259a50be4381
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
 
chunhan 
parents: 
24 
diff
changeset
 | 
142  | 
apply (auto dest:p4' procs_of_shm_prop2 Tainted_in_current split:if_splits option.splits)[1]  | 
| 
 
259a50be4381
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
 
chunhan 
parents: 
24 
diff
changeset
 | 
143  | 
|
| 
24
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
144  | 
apply (erule disjE, drule_tac p = p and p' = p' in p4', simp+)  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
145  | 
apply (erule disjE, rule disjI2, rule disjI2, rule_tac x = h in exI, simp, rule_tac x= toflag in exI, simp)  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
146  | 
apply ((erule exE|erule conjE)+)  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
147  | 
|
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
148  | 
|
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
149  | 
apply (auto simp:info_flow_shm_def dest:p4'  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
150  | 
procs_of_shm_prop2 Tainted_in_current split:if_splits option.splits)[1]  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
151  | 
apply (drule_tac p = p and p' = p' in p4')  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
152  | 
apply (erule_tac x = ha in allE, simp)  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
153  | 
apply (drule_tac p = "nat1" and p' = p' in p4')  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
154  | 
apply (auto dest:p4'[where p = nat1 and p' = p'])  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
155  | 
|
| 31 | 156  | 
apply (induct s)  | 
| 
23
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
157  | 
apply simp defer  | 
| 
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
158  | 
apply (frule vd_cons, frule vt_grant_os, case_tac a)  | 
| 
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
159  | 
apply (auto simp:info_flow_shm_def elim!:disjE)  | 
| 
24
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
160  | 
sorry *)  | 
| 
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
161  | 
sorry  | 
| 31 | 162  | 
qed  | 
| 
24
 
566b0d1c3669
info_flow_shm is wrongly defined: p -flow-> p'; p' -flow-> p'' ==> p -flow-> p'', this property cannot be infered by current definition. we should use inductive
 
chunhan 
parents: 
23 
diff
changeset
 | 
163  | 
qed  | 
| 
23
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
164  | 
|
| 65 | 165  | 
lemma has_same_inode_comm:  | 
166  | 
"has_same_inode s f f' = has_same_inode s f' f"  | 
|
167  | 
by (auto simp add:has_same_inode_def same_inode_files_def is_file_def)  | 
|
168  | 
||
| 
23
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
169  | 
lemma tainted_imp_Tainted:  | 
| 
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
170  | 
"obj \<in> tainted s \<Longrightarrow> obj \<in> Tainted s"  | 
| 31 | 171  | 
apply (induct rule:tainted.induct) (*  | 
172  | 
apply (simp_all add:vd_cons)  | 
|
173  | 
apply (rule impI)  | 
|
174  | 
||
175  | 
apply (rule disjI2, rule_tac x = flag' in exI, simp)  | 
|
176  | 
apply ((rule impI)+, erule_tac x = p' in allE, simp)  | 
|
177  | 
*)  | 
|
| 65 | 178  | 
apply (auto intro:info_flow_shm_Tainted simp:has_same_inode_def dest:vd_cons)  | 
| 
23
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
179  | 
apply (case_tac e, auto split:option.splits if_splits simp:alive_simps)  | 
| 
 
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
 
chunhan 
parents: 
22 
diff
changeset
 | 
180  | 
done  | 
| 22 | 181  | 
|
| 31 | 182  | 
lemma Tainted_imp_tainted_aux_remain:  | 
183  | 
"\<lbrakk>obj \<in> Tainted s; valid (e # s); alive (e # s) obj; \<And> obj. obj \<in> Tainted s \<Longrightarrow> obj \<in> tainted s\<rbrakk>  | 
|
184  | 
\<Longrightarrow> obj \<in> tainted (e # s)"  | 
|
185  | 
by (rule t_remain, auto)  | 
|
186  | 
||
| 22 | 187  | 
lemma Tainted_imp_tainted:  | 
188  | 
"\<lbrakk>obj \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> obj \<in> tainted s"  | 
|
| 31 | 189  | 
apply (induct s arbitrary:obj, simp add:t_init)  | 
190  | 
apply (frule Tainted_in_current, simp+)  | 
|
191  | 
apply (frule vt_grant_os, frule vd_cons, case_tac a)  | 
|
| 65 | 192  | 
apply (auto split:if_splits option.splits elim:Tainted_imp_tainted_aux_remain intro:tainted.intros  | 
193  | 
simp:has_same_inode_def)  | 
|
| 31 | 194  | 
done  | 
195  | 
||
196  | 
lemma tainted_eq_Tainted:  | 
|
197  | 
"valid s \<Longrightarrow> (obj \<in> tainted s) = (obj \<in> Tainted s)"  | 
|
198  | 
by (rule iffI, auto intro:Tainted_imp_tainted tainted_imp_Tainted)  | 
|
| 22 | 199  | 
|
| 31 | 200  | 
lemma info_flow_shm_tainted:  | 
201  | 
"\<lbrakk>O_proc p \<in> tainted s; info_flow_shm s p p'; valid s\<rbrakk> \<Longrightarrow> O_proc p' \<in> tainted s"  | 
|
202  | 
by (simp only:tainted_eq_Tainted info_flow_shm_Tainted)  | 
|
| 22 | 203  | 
|
| 65 | 204  | 
|
205  | 
lemma same_inode_files_Tainted:  | 
|
206  | 
"\<lbrakk>O_file f \<in> Tainted s; f' \<in> same_inode_files s f; valid s\<rbrakk> \<Longrightarrow> O_file f' \<in> Tainted s"  | 
|
207  | 
apply (induct s arbitrary:f f', simp add:same_inode_in_seeds has_same_inode_def)  | 
|
208  | 
apply (frule vt_grant_os, frule vd_cons, case_tac a)  | 
|
209  | 
prefer 6  | 
|
210  | 
apply (simp split:if_splits option.splits add:same_inode_files_open current_files_simps)  | 
|
211  | 
prefer 8  | 
|
212  | 
apply (frule Tainted_in_current, simp, simp add:alive.simps, drule is_file_in_current)  | 
|
213  | 
apply (auto simp add:same_inode_files_closefd split:option.splits if_splits)[1]  | 
|
214  | 
prefer 8  | 
|
215  | 
apply (frule Tainted_in_current, simp, simp add:alive.simps, drule is_file_in_current)  | 
|
216  | 
apply (auto simp add:same_inode_files_unlink split:option.splits if_splits)[1]  | 
|
217  | 
prefer 10  | 
|
218  | 
apply (auto split:if_splits option.splits simp:same_inode_files_linkhard current_files_simps)[1]  | 
|
219  | 
apply (drule Tainted_in_current, simp, simp add:alive.simps is_file_in_current)  | 
|
220  | 
apply (drule same_inode_files_prop5, simp)  | 
|
221  | 
apply (drule same_inode_files_prop5, drule_tac f' = list1 and f'' = f' in same_inode_files_prop4, simp, simp)  | 
|
222  | 
||
223  | 
apply (auto simp:same_inode_files_other split:if_splits)  | 
|
224  | 
apply (drule_tac f'' = f' and f' = f and f = fa in same_inode_files_prop4, simp+)  | 
|
225  | 
apply (drule_tac f'' = f' and f' = f and f = list in same_inode_files_prop4, simp+)  | 
|
226  | 
done  | 
|
227  | 
||
| 31 | 228  | 
lemma has_same_inode_Tainted:  | 
229  | 
"\<lbrakk>O_file f \<in> Tainted s; has_same_inode s f f'; valid s\<rbrakk> \<Longrightarrow> O_file f' \<in> Tainted s"  | 
|
| 65 | 230  | 
by (simp add:has_same_inode_def same_inode_files_Tainted)  | 
| 31 | 231  | 
|
232  | 
lemma has_same_inode_tainted:  | 
|
233  | 
"\<lbrakk>O_file f \<in> tainted s; has_same_inode s f f'; valid s\<rbrakk> \<Longrightarrow> O_file f' \<in> tainted s"  | 
|
234  | 
by (simp add:has_same_inode_Tainted tainted_eq_Tainted)  | 
|
| 22 | 235  | 
|
| 65 | 236  | 
lemma same_inodes_Tainted:  | 
237  | 
"\<lbrakk>f \<in> same_inode_files s f'; valid s\<rbrakk> \<Longrightarrow> (O_file f \<in> Tainted s) = (O_file f' \<in> Tainted s)"  | 
|
238  | 
apply (frule same_inode_files_prop8, frule same_inode_files_prop7)  | 
|
239  | 
apply (auto intro:has_same_inode_Tainted)  | 
|
240  | 
done  | 
|
| 22 | 241  | 
|
242  | 
||
243  | 
||
| 18 | 244  | 
lemma tainted_in_current:  | 
| 19 | 245  | 
"obj \<in> tainted s \<Longrightarrow> alive s obj"  | 
| 31 | 246  | 
apply (erule tainted.induct)  | 
247  | 
apply (auto dest:vt_grant_os vd_cons info_shm_flow_in_procs procs_of_shm_prop2 simp:is_file_simps)  | 
|
| 19 | 248  | 
apply (drule seeds_in_init, simp add:tobj_in_alive)  | 
249  | 
apply (erule has_same_inode_prop2, simp, simp add:vd_cons)  | 
|
250  | 
apply (frule vt_grant_os, simp)  | 
|
251  | 
apply (erule has_same_inode_prop1, simp, simp add:vd_cons)  | 
|
252  | 
done  | 
|
253  | 
||
254  | 
lemma tainted_is_valid:  | 
|
255  | 
"obj \<in> tainted s \<Longrightarrow> valid s"  | 
|
256  | 
by (erule tainted.induct, auto intro:valid.intros)  | 
|
| 18 | 257  | 
|
| 19 | 258  | 
lemma t_remain_app:  | 
259  | 
"\<lbrakk>obj \<in> tainted s; \<not> deleted obj (s' @ s); valid (s' @ s)\<rbrakk>  | 
|
260  | 
\<Longrightarrow> obj \<in> tainted (s' @ s)"  | 
|
261  | 
apply (induct s', simp)  | 
|
262  | 
apply (simp (no_asm) only:cons_app_simp_aux, rule t_remain)  | 
|
263  | 
apply (simp_all add:not_deleted_cons_D vd_cons)  | 
|
264  | 
apply (drule tainted_in_current, simp add:not_deleted_imp_alive_cons)  | 
|
265  | 
done  | 
|
| 18 | 266  | 
|
| 19 | 267  | 
lemma valid_tainted_obj:  | 
| 22 | 268  | 
"obj \<in> tainted s \<Longrightarrow> (\<forall> f. obj \<noteq> O_dir f) \<and> (\<forall> q. obj \<noteq> O_msgq q) \<and> (\<forall> h. obj \<noteq> O_shm h) \<and> (\<forall> p fd. obj \<noteq> O_fd p fd) \<and> (\<forall> s. obj \<noteq> O_tcp_sock s) \<and> (\<forall> s. obj \<noteq> O_udp_sock s)"  | 
| 19 | 269  | 
apply (erule tainted.induct)  | 
270  | 
apply (drule seeds_in_init)  | 
|
271  | 
by auto  | 
|
| 18 | 272  | 
|
| 19 | 273  | 
lemma dir_not_tainted: "O_dir f \<in> tainted s \<Longrightarrow> False"  | 
274  | 
by (auto dest:valid_tainted_obj)  | 
|
275  | 
||
276  | 
lemma msgq_not_tainted: "O_msgq q \<in> tainted s \<Longrightarrow> False"  | 
|
277  | 
by (auto dest:valid_tainted_obj)  | 
|
278  | 
||
279  | 
lemma shm_not_tainted: "O_shm h \<in> tainted s \<Longrightarrow> False"  | 
|
280  | 
by (auto dest:valid_tainted_obj)  | 
|
| 18 | 281  | 
|
282  | 
end  | 
|
283  | 
||
| 19 | 284  | 
end  |