Tainted_prop.thy
author chunhan
Sun, 09 Jun 2013 14:28:58 +0800
changeset 23 25e55731ed01
parent 22 f20a798cdf7d
child 24 566b0d1c3669
permissions -rw-r--r--
locale of tainting for seeds when same shm/inode bugs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     1
theory Tainted_prop 
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
     2
imports Main Flask Flask_type Init_prop Current_files_prop Current_sockets_prop Delete_prop Proc_fd_of_file_prop Current_prop
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     3
begin
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     4
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     5
context tainting begin
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     6
22
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
     7
fun Tainted :: "t_state \<Rightarrow> t_object set"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
     8
where
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
     9
  "Tainted [] = seeds"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    10
| "Tainted (Clone p p' fds shms # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    11
     (if (O_proc p) \<in> Tainted s then Tainted s \<union> {O_proc p'} else Tainted s)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    12
| "Tainted (Execve p f fds # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    13
     (if (O_file f) \<in> Tainted s then Tainted s \<union> {O_proc p} else Tainted s)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    14
| "Tainted (Kill p p' # s) = Tainted s - {O_proc p'}"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    15
| "Tainted (Ptrace p p' # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    16
     (if (O_proc p) \<in> Tainted s 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    17
      then Tainted s \<union> {O_proc p'' | p''. info_flow_shm s p' p''}
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    18
      else if (O_proc p') \<in> Tainted s 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    19
           then Tainted s \<union> {O_proc p'' | p''. info_flow_shm s p p''}
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    20
                else Tainted s)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    21
| "Tainted (Exit p # s) = Tainted s - {O_proc p}"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    22
| "Tainted (Open p f flags fd opt # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    23
     (case opt of
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    24
        Some inum \<Rightarrow> (if (O_proc p) \<in> Tainted s
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    25
                      then Tainted s \<union> {O_file f}
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    26
                      else Tainted s)
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    27
      | _         \<Rightarrow> Tainted s)" 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    28
| "Tainted (ReadFile p fd # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    29
     (case (file_of_proc_fd s p fd) of
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    30
        Some f \<Rightarrow> if (O_file f) \<in> Tainted s
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    31
                  then Tainted s \<union> {O_proc p' | p'. info_flow_shm s p p'}
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    32
                  else Tainted s
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    33
      | None   \<Rightarrow> Tainted s)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    34
| "Tainted (WriteFile p fd # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    35
     (case (file_of_proc_fd s p fd) of 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    36
        Some f \<Rightarrow> if (O_proc p) \<in> Tainted s
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    37
                  then Tainted s \<union> {O_file f' | f'. has_same_inode s f f'}
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    38
                  else Tainted s
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    39
      | None   \<Rightarrow> Tainted s)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    40
| "Tainted (CloseFd p fd # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    41
     (case (file_of_proc_fd s p fd) of
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    42
        Some f \<Rightarrow> ( if ((proc_fd_of_file s f = {(p,fd)}) \<and> (f \<in> files_hung_by_del s))
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    43
                    then Tainted s - {O_file f} else Tainted s )
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    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
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    47
| "Tainted (LinkHard p f f' # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    48
     (if (O_file f \<in> Tainted s) then Tainted s \<union> {O_file f'} else Tainted s)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    49
| "Tainted (Truncate p f len # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    50
     (if (len > 0 \<and> O_proc p \<in> Tainted s)
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    51
      then Tainted s \<union> {O_file f' | f'. has_same_inode s f f'}
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    52
      else Tainted s)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    53
| "Tainted (SendMsg p q m # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    54
     (if (O_proc p \<in> Tainted s) then Tainted s \<union> {O_msg q m} else Tainted s)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    55
| "Tainted (RecvMsg p q m # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    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
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    58
      else Tainted s)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    59
| "Tainted (RemoveMsgq p q # s) = Tainted s - {O_msg q m| m. O_msg q m \<in> Tainted s}"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    60
| "Tainted (e # s) = Tainted s"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    61
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    62
lemma valid_Tainted_obj:
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    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)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    64
apply (induct s, simp, drule seeds_in_init, case_tac obj, simp+)
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    65
apply (frule vd_cons, frule vt_grant_os, case_tac a)
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    66
apply (auto split:if_splits option.splits)
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    67
done
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    68
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    69
lemma Tainted_in_current:
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    70
  "\<lbrakk>obj \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> alive s obj"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    71
apply (induct s, simp)
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    72
apply (drule seeds_in_init, case_tac obj, simp_all add:is_file_nil)
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    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
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   104
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   105
lemma Tainted_imp_tainted:
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   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
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   118
apply (case_tac a)
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   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
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   122
pr 25
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   123
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   124
lemma tainted_imp_Tainted:
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   125
  "obj \<in> tainted s \<Longrightarrow> obj \<in> Tainted s"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   126
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   127
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   128
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   129
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   130
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   131
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   132
lemma tainted_in_current:
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   133
  "obj \<in> tainted s \<Longrightarrow> alive s obj"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   134
apply (erule tainted.induct, auto dest:vt_grant_os vd_cons simp:is_file_simps)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   135
apply (drule seeds_in_init, simp add:tobj_in_alive)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   136
apply (erule has_same_inode_prop2, simp, simp add:vd_cons)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   137
apply (frule vt_grant_os, simp)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   138
apply (erule has_same_inode_prop1, simp, simp add:vd_cons)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   139
done
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   140
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   141
lemma tainted_is_valid:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   142
  "obj \<in> tainted s \<Longrightarrow> valid s"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   143
by (erule tainted.induct, auto intro:valid.intros)
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   144
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   145
lemma t_remain_app:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   146
  "\<lbrakk>obj \<in> tainted s; \<not> deleted obj (s' @ s); valid (s' @ s)\<rbrakk> 
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   147
  \<Longrightarrow> obj \<in> tainted (s' @ s)"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   148
apply (induct s', simp)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   149
apply (simp (no_asm) only:cons_app_simp_aux, rule t_remain)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   150
apply (simp_all add:not_deleted_cons_D vd_cons)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   151
apply (drule tainted_in_current, simp add:not_deleted_imp_alive_cons)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   152
done
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   153
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   154
lemma valid_tainted_obj:
22
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   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
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   156
apply (erule tainted.induct)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   157
apply (drule seeds_in_init)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   158
by auto
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   159
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   160
lemma dir_not_tainted: "O_dir f \<in> tainted s \<Longrightarrow> False"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   161
by (auto dest:valid_tainted_obj)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   162
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   163
lemma msgq_not_tainted: "O_msgq q \<in> tainted s \<Longrightarrow> False"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   164
by (auto dest:valid_tainted_obj)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   165
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   166
lemma shm_not_tainted: "O_shm h \<in> tainted s \<Longrightarrow> False"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   167
by (auto dest:valid_tainted_obj)
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   168
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   169
end
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   170
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   171
end