author | chunhan |
Sun, 09 Jun 2013 14:28:58 +0800 | |
changeset 23 | 25e55731ed01 |
parent 22 | f20a798cdf7d |
child 24 | 566b0d1c3669 |
permissions | -rw-r--r-- |
18 | 1 |
theory Tainted_prop |
19 | 2 |
imports Main Flask Flask_type Init_prop Current_files_prop Current_sockets_prop Delete_prop Proc_fd_of_file_prop Current_prop |
18 | 3 |
begin |
4 |
||
5 |
context tainting begin |
|
6 |
||
22 | 7 |
fun Tainted :: "t_state \<Rightarrow> t_object set" |
8 |
where |
|
9 |
"Tainted [] = seeds" |
|
10 |
| "Tainted (Clone p p' fds shms # s) = |
|
11 |
(if (O_proc p) \<in> Tainted s then Tainted s \<union> {O_proc p'} else Tainted s)" |
|
12 |
| "Tainted (Execve p f fds # s) = |
|
13 |
(if (O_file f) \<in> Tainted s then Tainted s \<union> {O_proc p} else Tainted s)" |
|
14 |
| "Tainted (Kill p p' # s) = Tainted s - {O_proc p'}" |
|
15 |
| "Tainted (Ptrace p p' # s) = |
|
16 |
(if (O_proc p) \<in> Tainted s |
|
17 |
then Tainted s \<union> {O_proc p'' | p''. info_flow_shm s p' p''} |
|
18 |
else if (O_proc p') \<in> Tainted s |
|
19 |
then Tainted s \<union> {O_proc p'' | p''. info_flow_shm s p p''} |
|
20 |
else Tainted s)" |
|
21 |
| "Tainted (Exit p # s) = Tainted s - {O_proc p}" |
|
22 |
| "Tainted (Open p f flags fd opt # s) = |
|
23 |
(case opt of |
|
24 |
Some inum \<Rightarrow> (if (O_proc p) \<in> Tainted s |
|
25 |
then Tainted s \<union> {O_file f} |
|
26 |
else Tainted s) |
|
27 |
| _ \<Rightarrow> Tainted s)" |
|
28 |
| "Tainted (ReadFile p fd # s) = |
|
29 |
(case (file_of_proc_fd s p fd) of |
|
30 |
Some f \<Rightarrow> if (O_file f) \<in> Tainted s |
|
31 |
then Tainted s \<union> {O_proc p' | p'. info_flow_shm s p p'} |
|
32 |
else Tainted s |
|
33 |
| None \<Rightarrow> Tainted s)" |
|
34 |
| "Tainted (WriteFile p fd # s) = |
|
35 |
(case (file_of_proc_fd s p fd) of |
|
36 |
Some f \<Rightarrow> if (O_proc p) \<in> Tainted s |
|
37 |
then Tainted s \<union> {O_file f' | f'. has_same_inode s f f'} |
|
38 |
else Tainted s |
|
39 |
| None \<Rightarrow> Tainted s)" |
|
40 |
| "Tainted (CloseFd p fd # s) = |
|
41 |
(case (file_of_proc_fd s p fd) of |
|
42 |
Some f \<Rightarrow> ( if ((proc_fd_of_file s f = {(p,fd)}) \<and> (f \<in> files_hung_by_del s)) |
|
43 |
then Tainted s - {O_file f} else Tainted s ) |
|
44 |
| _ \<Rightarrow> Tainted s)" |
|
23
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
45 |
| "Tainted (UnLink p f # s) = |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
46 |
(if (proc_fd_of_file s f = {}) then Tainted s - {O_file f} else Tainted s)" |
22 | 47 |
| "Tainted (LinkHard p f f' # s) = |
48 |
(if (O_file f \<in> Tainted s) then Tainted s \<union> {O_file f'} else Tainted s)" |
|
49 |
| "Tainted (Truncate p f len # s) = |
|
50 |
(if (len > 0 \<and> O_proc p \<in> Tainted s) |
|
51 |
then Tainted s \<union> {O_file f' | f'. has_same_inode s f f'} |
|
52 |
else Tainted s)" |
|
53 |
| "Tainted (SendMsg p q m # s) = |
|
54 |
(if (O_proc p \<in> Tainted s) then Tainted s \<union> {O_msg q m} else Tainted s)" |
|
55 |
| "Tainted (RecvMsg p q m # s) = |
|
56 |
(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
|
57 |
then (Tainted s \<union> {O_proc p' | p'. info_flow_shm s p p'}) - {O_msg q m} |
22 | 58 |
else Tainted s)" |
59 |
| "Tainted (RemoveMsgq p q # s) = Tainted s - {O_msg q m| m. O_msg q m \<in> Tainted s}" |
|
60 |
| "Tainted (e # s) = Tainted s" |
|
61 |
||
62 |
lemma valid_Tainted_obj: |
|
63 |
"\<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)" |
|
64 |
apply (induct s, simp, drule seeds_in_init, case_tac obj, simp+) |
|
65 |
apply (frule vd_cons, frule vt_grant_os, case_tac a) |
|
66 |
apply (auto split:if_splits option.splits) |
|
67 |
done |
|
68 |
||
69 |
lemma Tainted_in_current: |
|
70 |
"\<lbrakk>obj \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> alive s obj" |
|
71 |
apply (induct s, simp) |
|
72 |
apply (drule seeds_in_init, case_tac obj, simp_all add:is_file_nil) |
|
73 |
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
|
74 |
apply (auto simp:alive_simps split:if_splits option.splits t_object.splits |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
75 |
intro:has_same_inode_prop2 has_same_inode_prop1 dest:info_shm_flow_in_procs) |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
76 |
done |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
77 |
|
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
78 |
|
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
79 |
lemma has_inode_Tainted_aux: |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
80 |
"\<lbrakk>O_file f \<in> tainted s; has_same_inode s f f'\<rbrakk> \<Longrightarrow> O_file f' \<in> tainted s" |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
81 |
apply (erule tainted.induct) |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
82 |
apply (auto intro:tainted.intros simp:has_same_inode_def) |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
83 |
(*?? need simpset for tainted *) |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
84 |
sorry |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
85 |
|
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
86 |
lemma info_flow_shm_Tainted: |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
87 |
"\<lbrakk>O_proc p \<in> Tainted s; info_flow_shm s p p'; valid s\<rbrakk> \<Longrightarrow> O_proc p' \<in> Tainted s" |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
88 |
proof (induct s) |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
89 |
case Nil |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
90 |
thus ?case |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
91 |
apply simp |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
92 |
apply (induct s) |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
93 |
apply simp defer |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
94 |
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
|
95 |
apply (auto simp:info_flow_shm_def elim!:disjE) |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
96 |
sorry |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
97 |
|
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
98 |
lemma tainted_imp_Tainted: |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
99 |
"obj \<in> tainted s \<Longrightarrow> obj \<in> Tainted s" |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
100 |
apply (induct rule:tainted.induct) |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
101 |
apply (auto intro:info_flow_shm_Tainted dest:vd_cons) |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
102 |
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
|
103 |
done |
22 | 104 |
|
105 |
lemma Tainted_imp_tainted: |
|
106 |
"\<lbrakk>obj \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> obj \<in> tainted s" |
|
23
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
107 |
proof (induct s arbitrary:obj) |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
108 |
case Nil |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
109 |
thus ?case by (auto intro:t_init) |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
110 |
next |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
111 |
case (Cons e s) |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
112 |
hence p1: "\<And> obj. obj \<in> Tainted s \<Longrightarrow> obj \<in> tainted s" |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
113 |
and p2: "obj \<in> Tainted (e # s)" and p3: "valid (e # s)" |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
114 |
and p4: "valid s" and p5: "os_grant s e" |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
115 |
by (auto dest:vd_cons intro:vd_cons vt_grant_os) |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
116 |
from p1 have p6: "" |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
117 |
apply (frule vd_cons, frule vt_grant_os, frule valid_Tainted_obj, simp) |
22 | 118 |
apply (case_tac a) |
119 |
apply (auto intro!:t_init t_clone t_execve t_cfile t_read t_write t_link t_trunc t_sendmsg t_recvmsg |
|
23
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
120 |
intro:t_remain |
25e55731ed01
locale of tainting for seeds when same shm/inode bugs
chunhan
parents:
22
diff
changeset
|
121 |
split:if_splits option.splits dest:Tainted_in_current) |
22 | 122 |
pr 25 |
123 |
||
124 |
lemma tainted_imp_Tainted: |
|
125 |
"obj \<in> tainted s \<Longrightarrow> obj \<in> Tainted s" |
|
126 |
||
127 |
||
128 |
||
129 |
||
130 |
||
131 |
||
18 | 132 |
lemma tainted_in_current: |
19 | 133 |
"obj \<in> tainted s \<Longrightarrow> alive s obj" |
134 |
apply (erule tainted.induct, auto dest:vt_grant_os vd_cons simp:is_file_simps) |
|
135 |
apply (drule seeds_in_init, simp add:tobj_in_alive) |
|
136 |
apply (erule has_same_inode_prop2, simp, simp add:vd_cons) |
|
137 |
apply (frule vt_grant_os, simp) |
|
138 |
apply (erule has_same_inode_prop1, simp, simp add:vd_cons) |
|
139 |
done |
|
140 |
||
141 |
lemma tainted_is_valid: |
|
142 |
"obj \<in> tainted s \<Longrightarrow> valid s" |
|
143 |
by (erule tainted.induct, auto intro:valid.intros) |
|
18 | 144 |
|
19 | 145 |
lemma t_remain_app: |
146 |
"\<lbrakk>obj \<in> tainted s; \<not> deleted obj (s' @ s); valid (s' @ s)\<rbrakk> |
|
147 |
\<Longrightarrow> obj \<in> tainted (s' @ s)" |
|
148 |
apply (induct s', simp) |
|
149 |
apply (simp (no_asm) only:cons_app_simp_aux, rule t_remain) |
|
150 |
apply (simp_all add:not_deleted_cons_D vd_cons) |
|
151 |
apply (drule tainted_in_current, simp add:not_deleted_imp_alive_cons) |
|
152 |
done |
|
18 | 153 |
|
19 | 154 |
lemma valid_tainted_obj: |
22 | 155 |
"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 | 156 |
apply (erule tainted.induct) |
157 |
apply (drule seeds_in_init) |
|
158 |
by auto |
|
18 | 159 |
|
19 | 160 |
lemma dir_not_tainted: "O_dir f \<in> tainted s \<Longrightarrow> False" |
161 |
by (auto dest:valid_tainted_obj) |
|
162 |
||
163 |
lemma msgq_not_tainted: "O_msgq q \<in> tainted s \<Longrightarrow> False" |
|
164 |
by (auto dest:valid_tainted_obj) |
|
165 |
||
166 |
lemma shm_not_tainted: "O_shm h \<in> tainted s \<Longrightarrow> False" |
|
167 |
by (auto dest:valid_tainted_obj) |
|
18 | 168 |
|
169 |
end |
|
170 |
||
19 | 171 |
end |