Tainted_prop.thy
author chunhan
Mon, 23 Dec 2013 19:47:17 +0800
changeset 82 0a68c605ae79
parent 65 6f9a588bcfc4
permissions -rw-r--r--
update
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 
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
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 Alive_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
34
e7f850d1e08e remove O_msg from co2sobj/init_obj2sobj
chunhan
parents: 31
diff changeset
     5
ML {*quick_and_dirty := true*}
e7f850d1e08e remove O_msg from co2sobj/init_obj2sobj
chunhan
parents: 31
diff changeset
     6
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     7
context tainting begin
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
     8
22
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
     9
fun Tainted :: "t_state \<Rightarrow> t_object set"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    10
where
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    11
  "Tainted [] = seeds"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    12
| "Tainted (Clone p p' fds shms # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    13
     (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
    14
| "Tainted (Execve p f fds # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    15
     (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
    16
| "Tainted (Kill p p' # s) = Tainted s - {O_proc p'}"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    17
| "Tainted (Ptrace p p' # s) = 
43
chunhan
parents: 34
diff changeset
    18
     (if (O_proc p \<in> Tainted s)
22
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''}
43
chunhan
parents: 34
diff changeset
    20
      else if (O_proc p' \<in> Tainted s)
22
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    21
           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
    22
                else Tainted s)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    23
| "Tainted (Exit p # s) = Tainted s - {O_proc p}"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    24
| "Tainted (Open p f flags fd opt # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    25
     (case opt of
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    26
        Some inum \<Rightarrow> (if (O_proc p) \<in> Tainted s
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    27
                      then Tainted s \<union> {O_file f}
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    28
                      else Tainted s)
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    29
      | _         \<Rightarrow> Tainted s)" 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    30
| "Tainted (ReadFile p fd # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    31
     (case (file_of_proc_fd s p fd) of
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    32
        Some f \<Rightarrow> if (O_file f) \<in> Tainted s
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    33
                  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
    34
                  else Tainted s
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    35
      | None   \<Rightarrow> Tainted s)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    36
| "Tainted (WriteFile p fd # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    37
     (case (file_of_proc_fd s p fd) of 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    38
        Some f \<Rightarrow> if (O_proc p) \<in> Tainted s
65
chunhan
parents: 56
diff changeset
    39
                  then Tainted s \<union> {O_file f' | f'. f' \<in> same_inode_files s f}
22
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    40
                  else Tainted s
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    41
      | None   \<Rightarrow> Tainted s)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    42
| "Tainted (CloseFd p fd # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    43
     (case (file_of_proc_fd s p fd) of
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    44
        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
    45
                    then Tainted s - {O_file f} else Tainted s )
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    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
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    49
| "Tainted (LinkHard p f f' # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    50
     (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
    51
| "Tainted (Truncate p f len # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    52
     (if (len > 0 \<and> O_proc p \<in> Tainted s)
65
chunhan
parents: 56
diff changeset
    53
      then Tainted s \<union> {O_file f' | f'. f' \<in> same_inode_files s f}
22
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    54
      else Tainted s)"
27
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
    55
| "Tainted (Attach p h flag # s) = 
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
    56
     (if (O_proc p \<in> Tainted s \<and> flag = SHM_RDWR) 
56
chunhan
parents: 43
diff changeset
    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
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
    58
      else if (\<exists> p'. O_proc p' \<in> Tainted s \<and> (p', SHM_RDWR) \<in> procs_of_shm s h)
56
chunhan
parents: 43
diff changeset
    59
           then Tainted s \<union> {O_proc p'| p'. info_flow_shm s p p'}
27
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
    60
           else Tainted s)"
22
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    61
| "Tainted (SendMsg p q m # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    62
     (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
    63
| "Tainted (RecvMsg p q m # s) = 
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    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
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    66
      else Tainted s)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    67
| "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
    68
| "Tainted (e # s) = Tainted s"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    69
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    70
lemma valid_Tainted_obj:
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    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)"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    72
apply (induct s, simp, drule seeds_in_init, case_tac obj, simp+)
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    73
apply (frule vd_cons, frule vt_grant_os, case_tac a)
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    74
apply (auto split:if_splits option.splits)
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    75
done
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    76
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    77
lemma Tainted_in_current:
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    78
  "\<lbrakk>obj \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> alive s obj"
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    79
apply (induct s, simp)
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
    80
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
    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
chunhan
parents: 56
diff changeset
    83
           intro:same_inode_files_prop1 procs_of_shm_prop2 
27
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
    84
            dest:info_shm_flow_in_procs)
65
chunhan
parents: 56
diff changeset
    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
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   107
    by (rule p4, auto simp:info_flow_shm_def one_flow_shm_def procs_of_shm_prop2 p5)    
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   108
  from p2 p3 have p7: "p \<in> current_procs (e # s)" and p8: "p' \<in> current_procs (e # s)" 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   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
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   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
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   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
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   118
    apply (simp add:info_flow_shm_simps split:if_splits option.splits)
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   119
    apply (rule allI|rule impI|rule conjI)+
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   120
    apply simp
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   121
    apply (case_tac "O_proc p \<in> Tainted s", drule_tac p'=p' in p4, simp+)
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   122
    apply simp
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   123
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   124
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   125
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   126
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   127
    apply (auto simp:info_flow_shm_simps one_flow_shm_def dest:Tainted_in_current 
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   128
  intro:p4 p4' split:if_splits option.splits)
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   129
    apply (auto simp:info_flow_shm_def one_flow_shm_def)
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   130
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   131
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   132
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   133
    apply (auto simp:one_flow_shm_def intro:p4 p4' split:if_splits option.splits)
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   134
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   135
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   136
fc749f19b894 Info_flow_shm_attach_prop
chunhan
parents: 25
diff changeset
   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
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   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
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   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
chunhan
parents: 56
diff changeset
   165
lemma has_same_inode_comm:
chunhan
parents: 56
diff changeset
   166
  "has_same_inode s f f' = has_same_inode s f' f"
chunhan
parents: 56
diff changeset
   167
by (auto simp add:has_same_inode_def same_inode_files_def is_file_def)
chunhan
parents: 56
diff changeset
   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
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   171
apply (induct rule:tainted.induct)  (*
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   172
apply (simp_all add:vd_cons) 
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   173
apply (rule impI)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   174
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   175
apply (rule disjI2, rule_tac x = flag' in exI, simp)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   176
apply ((rule impI)+, erule_tac x = p' in allE, simp)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   177
*)
65
chunhan
parents: 56
diff changeset
   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
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   181
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   182
lemma Tainted_imp_tainted_aux_remain:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   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>
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   184
   \<Longrightarrow> obj \<in> tainted (e # s)"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   185
by (rule t_remain, auto)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   186
22
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   187
lemma Tainted_imp_tainted:
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   188
  "\<lbrakk>obj \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> obj \<in> tainted s"
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   189
apply (induct s arbitrary:obj, simp add:t_init)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   190
apply (frule Tainted_in_current, simp+)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   191
apply (frule vt_grant_os, frule vd_cons, case_tac a)
65
chunhan
parents: 56
diff changeset
   192
apply (auto split:if_splits option.splits elim:Tainted_imp_tainted_aux_remain intro:tainted.intros
chunhan
parents: 56
diff changeset
   193
  simp:has_same_inode_def)
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   194
done
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   195
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   196
lemma tainted_eq_Tainted:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   197
  "valid s \<Longrightarrow> (obj \<in> tainted s) = (obj \<in> Tainted s)"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   198
by (rule iffI, auto intro:Tainted_imp_tainted tainted_imp_Tainted)
22
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   199
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   200
lemma info_flow_shm_tainted:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   201
  "\<lbrakk>O_proc p \<in> tainted s; info_flow_shm s p p'; valid s\<rbrakk> \<Longrightarrow> O_proc p' \<in> tainted s"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   202
by (simp only:tainted_eq_Tainted info_flow_shm_Tainted)
22
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   203
65
chunhan
parents: 56
diff changeset
   204
chunhan
parents: 56
diff changeset
   205
lemma same_inode_files_Tainted:
chunhan
parents: 56
diff changeset
   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"
chunhan
parents: 56
diff changeset
   207
apply (induct s arbitrary:f f', simp add:same_inode_in_seeds has_same_inode_def)
chunhan
parents: 56
diff changeset
   208
apply (frule vt_grant_os, frule vd_cons, case_tac a)
chunhan
parents: 56
diff changeset
   209
prefer 6
chunhan
parents: 56
diff changeset
   210
apply (simp split:if_splits option.splits add:same_inode_files_open current_files_simps)
chunhan
parents: 56
diff changeset
   211
prefer 8
chunhan
parents: 56
diff changeset
   212
apply (frule Tainted_in_current, simp, simp add:alive.simps, drule is_file_in_current)
chunhan
parents: 56
diff changeset
   213
apply (auto simp add:same_inode_files_closefd split:option.splits if_splits)[1]
chunhan
parents: 56
diff changeset
   214
prefer 8
chunhan
parents: 56
diff changeset
   215
apply (frule Tainted_in_current, simp, simp add:alive.simps, drule is_file_in_current)
chunhan
parents: 56
diff changeset
   216
apply (auto simp add:same_inode_files_unlink split:option.splits if_splits)[1]
chunhan
parents: 56
diff changeset
   217
prefer 10
chunhan
parents: 56
diff changeset
   218
apply (auto split:if_splits option.splits simp:same_inode_files_linkhard current_files_simps)[1]
chunhan
parents: 56
diff changeset
   219
apply (drule Tainted_in_current, simp, simp add:alive.simps is_file_in_current)
chunhan
parents: 56
diff changeset
   220
apply (drule same_inode_files_prop5, simp)
chunhan
parents: 56
diff changeset
   221
apply (drule same_inode_files_prop5, drule_tac f' = list1 and f'' = f' in same_inode_files_prop4, simp, simp)
chunhan
parents: 56
diff changeset
   222
chunhan
parents: 56
diff changeset
   223
apply (auto simp:same_inode_files_other split:if_splits)
chunhan
parents: 56
diff changeset
   224
apply (drule_tac f'' = f' and f' = f and f = fa in same_inode_files_prop4, simp+)
chunhan
parents: 56
diff changeset
   225
apply (drule_tac f'' = f' and f' = f and f = list in same_inode_files_prop4, simp+)
chunhan
parents: 56
diff changeset
   226
done
chunhan
parents: 56
diff changeset
   227
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   228
lemma has_same_inode_Tainted:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   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
chunhan
parents: 56
diff changeset
   230
by (simp add:has_same_inode_def same_inode_files_Tainted)
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   231
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   232
lemma has_same_inode_tainted:
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   233
  "\<lbrakk>O_file f \<in> tainted s; has_same_inode s f f'; valid s\<rbrakk> \<Longrightarrow> O_file f' \<in> tainted s"
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   234
by (simp add:has_same_inode_Tainted tainted_eq_Tainted)
22
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   235
65
chunhan
parents: 56
diff changeset
   236
lemma same_inodes_Tainted:
chunhan
parents: 56
diff changeset
   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)"
chunhan
parents: 56
diff changeset
   238
apply (frule same_inode_files_prop8, frule same_inode_files_prop7)
chunhan
parents: 56
diff changeset
   239
apply (auto intro:has_same_inode_Tainted)
chunhan
parents: 56
diff changeset
   240
done
22
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   241
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   242
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   243
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   244
lemma tainted_in_current:
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   245
  "obj \<in> tainted s \<Longrightarrow> alive s obj"
31
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   246
apply (erule tainted.induct)
aa1375b6c0eb find bugs in os_grant, case RecvMsg
chunhan
parents: 29
diff changeset
   247
apply (auto dest:vt_grant_os vd_cons info_shm_flow_in_procs procs_of_shm_prop2 simp:is_file_simps)
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   248
apply (drule seeds_in_init, simp add:tobj_in_alive)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   249
apply (erule has_same_inode_prop2, simp, simp add:vd_cons)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   250
apply (frule vt_grant_os, simp)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   251
apply (erule has_same_inode_prop1, simp, simp add:vd_cons)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   252
done
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   253
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   254
lemma tainted_is_valid:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   255
  "obj \<in> tainted s \<Longrightarrow> valid s"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   256
by (erule tainted.induct, auto intro:valid.intros)
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   257
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   258
lemma t_remain_app:
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   259
  "\<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
   260
  \<Longrightarrow> obj \<in> tainted (s' @ s)"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   261
apply (induct s', simp)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   262
apply (simp (no_asm) only:cons_app_simp_aux, rule t_remain)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   263
apply (simp_all add:not_deleted_cons_D vd_cons)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   264
apply (drule tainted_in_current, simp add:not_deleted_imp_alive_cons)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   265
done
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   266
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   267
lemma valid_tainted_obj:
22
f20a798cdf7d info_flow_shm bug & update
chunhan
parents: 19
diff changeset
   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
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   269
apply (erule tainted.induct)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   270
apply (drule seeds_in_init)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   271
by auto
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   272
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   273
lemma dir_not_tainted: "O_dir f \<in> tainted s \<Longrightarrow> False"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   274
by (auto dest:valid_tainted_obj)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   275
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   276
lemma msgq_not_tainted: "O_msgq q \<in> tainted s \<Longrightarrow> False"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   277
by (auto dest:valid_tainted_obj)
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   278
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   279
lemma shm_not_tainted: "O_shm h \<in> tainted s \<Longrightarrow> False"
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   280
by (auto dest:valid_tainted_obj)
18
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   281
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   282
end
9b42765ce554 info_flow did NOT guarantee in current_procs
chunhan
parents:
diff changeset
   283
19
ced0fcfbcf8e reprove the top-level dynamic2static
chunhan
parents: 18
diff changeset
   284
end