simple_selinux/Tainted_prop.thy
author chunhan
Mon, 02 Dec 2013 10:52:40 +0800
changeset 74 271e9818b6f6
child 75 99af1986e1e0
permissions -rw-r--r--
remove shm and linkhard, make a simplified version of selinux
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
74
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     1
theory Tainted_prop 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
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
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     3
begin
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     4
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     5
ML {*quick_and_dirty := true*}
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     6
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     7
context tainting begin
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     8
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     9
lemma valid_Tainted_obj:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    10
  "\<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)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    11
apply (induct s, simp, drule seeds_in_init, case_tac obj, simp+)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    12
apply (frule vd_cons, frule vt_grant_os, case_tac a)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    13
apply (auto split:if_splits option.splits)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    14
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    15
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    16
lemma Tainted_in_current:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    17
  "\<lbrakk>obj \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> alive s obj"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    18
apply (induct s, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    19
apply (drule seeds_in_init, case_tac obj, simp_all add:is_file_nil)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    20
apply (frule vd_cons, frule valid_Tainted_obj, simp, frule vt_grant_os, case_tac a)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    21
apply (auto simp:alive_simps split:if_splits option.splits t_object.splits
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    22
           intro:same_inode_files_prop1 procs_of_shm_prop2 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    23
            dest:info_shm_flow_in_procs)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    24
apply (auto simp:same_inode_files_def is_file_def split:if_splits)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    25
done 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    26
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    27
lemma Tainted_proc_in_current:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    28
  "\<lbrakk>O_proc p \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> p \<in> current_procs s"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    29
by (drule Tainted_in_current, simp+)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    30
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    31
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    32
lemma info_flow_shm_Tainted:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    33
  "\<lbrakk>O_proc p \<in> Tainted s; info_flow_shm s p p'; valid s\<rbrakk> \<Longrightarrow> O_proc p' \<in> Tainted s"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    34
proof (induct s arbitrary:p p')
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    35
  case Nil
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    36
  thus ?case by (simp add:flow_shm_in_seeds)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    37
next
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    38
  case (Cons e s)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    39
  hence p1: "O_proc p \<in> Tainted (e # s)" and p2: "info_flow_shm (e # s) p p'" and p3: "valid (e # s)"  
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    40
    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" 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    41
    and p5: "valid s" and p6: "os_grant s e"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    42
    by (auto dest:vd_cons intro:vd_cons vt_grant_os)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    43
  have p4': 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    44
    "\<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> 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    45
                \<Longrightarrow> O_proc p' \<in> Tainted s"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    46
    by (rule p4, auto simp:info_flow_shm_def one_flow_shm_def procs_of_shm_prop2 p5)    
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    47
  from p2 p3 have p7: "p \<in> current_procs (e # s)" and p8: "p' \<in> current_procs (e # s)" 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    48
    by (auto dest:info_shm_flow_in_procs) 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    49
  show ?case
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    50
  proof (cases "self_shm s p p'")
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    51
    case True with p1 show ?thesis by simp
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    52
  next
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    53
    case False
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    54
    with p1 p2 p5 p6 p7 p8 p3 show ?thesis
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    55
    apply (case_tac e)(*
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    56
    prefer 7
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    57
    apply (simp add:info_flow_shm_simps split:if_splits option.splits)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    58
    apply (rule allI|rule impI|rule conjI)+
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    59
    apply simp
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    60
    apply (case_tac "O_proc p \<in> Tainted s", drule_tac p'=p' in p4, simp+)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    61
    apply simp
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    62
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    63
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    64
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    65
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    66
    apply (auto simp:info_flow_shm_simps one_flow_shm_def dest:Tainted_in_current 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    67
  intro:p4 p4' split:if_splits option.splits)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    68
    apply (auto simp:info_flow_shm_def one_flow_shm_def)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    69
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    70
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    71
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    72
    apply (auto simp:one_flow_shm_def intro:p4 p4' split:if_splits option.splits)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    73
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    74
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    75
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    76
    prefer 7
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    77
    apply (simp split:if_splits option.splits)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    78
    apply (rule allI|rule impI|rule conjI)+
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    79
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    80
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    81
    apply (auto dest:p4'   procs_of_shm_prop2 Tainted_in_current split:if_splits option.splits)[1]
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    82
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    83
    apply (erule disjE, drule_tac p = p and p' = p' in p4', simp+)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    84
    apply (erule disjE, rule disjI2, rule disjI2, rule_tac x = h in exI, simp, rule_tac x= toflag in exI, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    85
    apply ((erule exE|erule conjE)+)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    86
    
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    87
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    88
    apply (auto simp:info_flow_shm_def dest:p4'
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    89
           procs_of_shm_prop2 Tainted_in_current split:if_splits option.splits)[1]
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    90
    apply (drule_tac p = p and p' = p' in p4')
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    91
    apply (erule_tac x = ha in allE, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    92
    apply (drule_tac p = "nat1" and p' = p' in p4')
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    93
    apply (auto dest:p4'[where p = nat1 and p' = p'])
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    94
    
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    95
apply (induct s) 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    96
apply simp defer
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    97
apply (frule vd_cons, frule vt_grant_os, case_tac a)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    98
apply (auto simp:info_flow_shm_def elim!:disjE)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    99
sorry *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   100
  sorry
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   101
qed
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   102
qed
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   103
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   104
lemma has_same_inode_comm:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   105
  "has_same_inode s f f' = has_same_inode s f' f"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   106
by (auto simp add:has_same_inode_def same_inode_files_def is_file_def)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   107
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   108
lemma tainted_imp_Tainted:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   109
  "obj \<in> tainted s \<Longrightarrow> obj \<in> Tainted s"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   110
apply (induct rule:tainted.induct)  (*
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   111
apply (simp_all add:vd_cons) 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   112
apply (rule impI)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   113
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   114
apply (rule disjI2, rule_tac x = flag' in exI, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   115
apply ((rule impI)+, erule_tac x = p' in allE, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   116
*)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   117
apply (auto intro:info_flow_shm_Tainted simp:has_same_inode_def dest:vd_cons)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   118
apply (case_tac e, auto split:option.splits if_splits simp:alive_simps)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   119
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   120
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   121
lemma Tainted_imp_tainted_aux_remain:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   122
  "\<lbrakk>obj \<in> Tainted s; valid (e # s); alive (e # s) obj; \<And> obj. obj \<in> Tainted s \<Longrightarrow> obj \<in> tainted s\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   123
   \<Longrightarrow> obj \<in> tainted (e # s)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   124
by (rule t_remain, auto)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   125
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   126
lemma Tainted_imp_tainted:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   127
  "\<lbrakk>obj \<in> Tainted s; valid s\<rbrakk> \<Longrightarrow> obj \<in> tainted s"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   128
apply (induct s arbitrary:obj, simp add:t_init)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   129
apply (frule Tainted_in_current, simp+)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   130
apply (frule vt_grant_os, frule vd_cons, case_tac a)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   131
apply (auto split:if_splits option.splits elim:Tainted_imp_tainted_aux_remain intro:tainted.intros
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   132
  simp:has_same_inode_def)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   133
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   134
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   135
lemma tainted_eq_Tainted:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   136
  "valid s \<Longrightarrow> (obj \<in> tainted s) = (obj \<in> Tainted s)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   137
by (rule iffI, auto intro:Tainted_imp_tainted tainted_imp_Tainted)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   138
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   139
lemma info_flow_shm_tainted:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   140
  "\<lbrakk>O_proc p \<in> tainted s; info_flow_shm s p p'; valid s\<rbrakk> \<Longrightarrow> O_proc p' \<in> tainted s"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   141
by (simp only:tainted_eq_Tainted info_flow_shm_Tainted)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   142
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   143
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   144
lemma same_inode_files_Tainted:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   145
  "\<lbrakk>O_file f \<in> Tainted s; f' \<in> same_inode_files s f; valid s\<rbrakk> \<Longrightarrow> O_file f' \<in> Tainted s"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   146
apply (induct s arbitrary:f f', simp add:same_inode_in_seeds has_same_inode_def)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   147
apply (frule vt_grant_os, frule vd_cons, case_tac a)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   148
prefer 6
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   149
apply (simp split:if_splits option.splits add:same_inode_files_open current_files_simps)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   150
prefer 8
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   151
apply (frule Tainted_in_current, simp, simp add:alive.simps, drule is_file_in_current)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   152
apply (auto simp add:same_inode_files_closefd split:option.splits if_splits)[1]
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   153
prefer 8
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   154
apply (frule Tainted_in_current, simp, simp add:alive.simps, drule is_file_in_current)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   155
apply (auto simp add:same_inode_files_unlink split:option.splits if_splits)[1]
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   156
prefer 10
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   157
apply (auto split:if_splits option.splits simp:same_inode_files_linkhard current_files_simps)[1]
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   158
apply (drule Tainted_in_current, simp, simp add:alive.simps is_file_in_current)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   159
apply (drule same_inode_files_prop5, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   160
apply (drule same_inode_files_prop5, drule_tac f' = list1 and f'' = f' in same_inode_files_prop4, simp, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   161
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   162
apply (auto simp:same_inode_files_other split:if_splits)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   163
apply (drule_tac f'' = f' and f' = f and f = fa in same_inode_files_prop4, simp+)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   164
apply (drule_tac f'' = f' and f' = f and f = list in same_inode_files_prop4, simp+)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   165
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   166
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   167
lemma has_same_inode_Tainted:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   168
  "\<lbrakk>O_file f \<in> Tainted s; has_same_inode s f f'; valid s\<rbrakk> \<Longrightarrow> O_file f' \<in> Tainted s"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   169
by (simp add:has_same_inode_def same_inode_files_Tainted)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   170
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   171
lemma has_same_inode_tainted:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   172
  "\<lbrakk>O_file f \<in> tainted s; has_same_inode s f f'; valid s\<rbrakk> \<Longrightarrow> O_file f' \<in> tainted s"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   173
by (simp add:has_same_inode_Tainted tainted_eq_Tainted)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   174
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   175
lemma same_inodes_Tainted:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   176
  "\<lbrakk>f \<in> same_inode_files s f'; valid s\<rbrakk> \<Longrightarrow> (O_file f \<in> Tainted s) = (O_file f' \<in> Tainted s)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   177
apply (frule same_inode_files_prop8, frule same_inode_files_prop7)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   178
apply (auto intro:has_same_inode_Tainted)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   179
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   180
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   181
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   182
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   183
lemma tainted_in_current:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   184
  "obj \<in> tainted s \<Longrightarrow> alive s obj"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   185
apply (erule tainted.induct)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   186
apply (auto dest:vt_grant_os vd_cons info_shm_flow_in_procs procs_of_shm_prop2 simp:is_file_simps)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   187
apply (drule seeds_in_init, simp add:tobj_in_alive)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   188
apply (erule has_same_inode_prop2, simp, simp add:vd_cons)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   189
apply (frule vt_grant_os, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   190
apply (erule has_same_inode_prop1, simp, simp add:vd_cons)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   191
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   192
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   193
lemma tainted_is_valid:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   194
  "obj \<in> tainted s \<Longrightarrow> valid s"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   195
by (erule tainted.induct, auto intro:valid.intros)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   196
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   197
lemma t_remain_app:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   198
  "\<lbrakk>obj \<in> tainted s; \<not> deleted obj (s' @ s); valid (s' @ s)\<rbrakk> 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   199
  \<Longrightarrow> obj \<in> tainted (s' @ s)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   200
apply (induct s', simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   201
apply (simp (no_asm) only:cons_app_simp_aux, rule t_remain)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   202
apply (simp_all add:not_deleted_cons_D vd_cons)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   203
apply (drule tainted_in_current, simp add:not_deleted_imp_alive_cons)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   204
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   205
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   206
lemma valid_tainted_obj:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   207
  "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)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   208
apply (erule tainted.induct)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   209
apply (drule seeds_in_init)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   210
by auto
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   211
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   212
lemma dir_not_tainted: "O_dir f \<in> tainted s \<Longrightarrow> False"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   213
by (auto dest:valid_tainted_obj)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   214
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   215
lemma msgq_not_tainted: "O_msgq q \<in> tainted s \<Longrightarrow> False"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   216
by (auto dest:valid_tainted_obj)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   217
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   218
lemma shm_not_tainted: "O_shm h \<in> tainted s \<Longrightarrow> False"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   219
by (auto dest:valid_tainted_obj)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   220
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   221
end
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   222
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   223
end