simple_selinux/Temp.thy
author chunhan
Thu, 05 Dec 2013 20:13:30 +0800
changeset 76 f27ba31b7e96
parent 74 271e9818b6f6
permissions -rw-r--r--
no shm version
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 Temp
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     2
imports Static_type OS_type_def Flask_type Flask Static
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
context tainting_s
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
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
definition update_ss :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    10
where
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    11
  "update_ss ss so so' \<equiv> if (is_many so) then ss \<union> {so'} else (ss - {so}) \<union> {so'}"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    12
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    13
definition add_ss :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    14
where
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    15
  "add_ss ss so \<equiv> ss \<union> {so}"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    16
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    17
definition del_ss :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> t_static_state"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    18
where
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    19
  "del_ss ss so \<equiv> if (is_many so) then ss else ss - {so}"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    20
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    21
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    22
definition update_ss_shms :: "t_static_state \<Rightarrow> t_sproc \<Rightarrow> bool \<Rightarrow> t_static_state"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    23
where
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    24
  "update_ss_shms ss spfrom tag \<equiv> {sobj. \<exists> sobj' \<in> ss. 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    25
     (case sobj' of 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    26
        S_proc sp tagp \<Rightarrow> if (info_flow_sshm spfrom sp) 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    27
                          then if (is_many_sproc sp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    28
                               then (sobj = S_proc sp tagp \<or> sobj = S_proc sp (tagp \<or> tag))
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    29
                               else (sobj = S_proc sp (tagp \<or> tag))
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    30
                          else (sobj = S_proc sp tag)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    31
     | _ \<Rightarrow> sobj = sobj')}"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    32
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    33
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    34
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    35
(* all reachable static states(sobjects set) *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    36
inductive_set static :: "t_static_state set"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    37
where
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    38
  s_init:    "init_static_state \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    39
| s_execve:  "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; S_file sfs tagf \<in> ss;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    40
               (fi,fsec,pfsec,asecs) \<in> sfs; npctxt_execve pctxt fsec = Some pctxt';
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    41
               grant_execve pctxt fsec pctxt'; search_check_s pctxt (fi,fsec,pfsec,asecs) True; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    42
               inherit_fds_check_s pctxt' fds'; fds' \<subseteq> fds\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    43
  \<Longrightarrow> (update_ss ss (S_proc (pi, pctxt, fds, shms) tagp) 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    44
                    (S_proc (pi, pctxt', fds', {}) (tagp \<or> tagf))) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    45
| s_clone:   "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    46
               permission_check pctxt pctxt C_process P_fork;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    47
               inherit_fds_check_s pctxt fds'; fds' \<subseteq> fds; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    48
               inherit_shms_check_s pctxt shms'; shms' \<subseteq> shms\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    49
  \<Longrightarrow> (add_ss ss (S_proc (Created, pctxt, fds', shms') tagp)) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    50
| s_kill:    "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    51
               S_proc (pi', pctxt', fds', shms') tagp' \<in> ss; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    52
               permission_check pctxt pctxt' C_process P_sigkill\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    53
  \<Longrightarrow> (del_ss ss (S_proc (pi', pctxt', fds', shms') tagp')) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    54
| s_ptrace:  "\<lbrakk>ss \<in> static; S_proc sp tagp \<in> ss; S_proc sp' tagp' \<in> ss; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    55
               permission_check (sextxt_of_sproc sp) (sectxt_of_sproc sp') C_process P_ptrace\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    56
  \<Longrightarrow> (update_ss_shms (update_ss_shms ss sp tagp) sp' tagp') \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    57
| s_exit:    "\<lbrakk>ss \<in> static; S_proc sp tagp \<in> ss\<rbrakk> \<Longrightarrow> (del_ss ss (S_proc sp tagp)) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    58
| s_open:    "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; S_file sfs tagf \<in> ss; sf \<in> sfs;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    59
              search_check_s pctxt sf True; \<not> is_creat_excl_flag flags; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    60
              oflags_check flags pctxt (sectxt_of_sfile sf); permission_check pctxt pctxt C_fd P_setattr\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    61
  \<Longrightarrow> (update_ss ss (S_proc (pi, pctxt, fds, shms) tagp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    62
                    (S_proc (pi, pctxt, fds \<union> {(pctxt,flags,sf)}, shms) tagp)) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    63
| s_open':   "\<lbrakk>ss \<in> static; S_proc (pi, pctxt, fds, shms) tagp \<in> ss; is_creat_excl_flag flags;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    64
               S_dir (pfi,fsec,pfsec,asecs) \<in> ss; search_check_s pctxt (pfi,fsec,pfsec,asecs) False; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    65
               nfsec = nfctxt_create pctxt fsec C_file; oflags_check flags pctxt nfsec;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    66
               permission_check pctxt fsec C_dir P_add_rename; permission_check pctxt pctxt C_fd P_setattr\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    67
  \<Longrightarrow> (update_ss (add_ss ss (S_file {(Created, nfsec, Some fsec, asecs \<union> {fsec})} tagp))
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    68
         (S_proc (pi, pctxt, fds, shms) tagp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    69
         (S_proc (pi, pctxt, fds \<union> {(pctxt, flags, (Created, nfsec, Some fsec, asecs \<union> {fsec}))}, shms) tagp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    70
      ) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    71
| S_readf:   "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; (fdctxt,flags,sf) \<in> fds; 
76
f27ba31b7e96 no shm version
chunhan
parents: 74
diff changeset
    72
               permission_check pctxt fdctxt C_fd P_setattr; S_file sf tagf \<in> ss; 
f27ba31b7e96 no shm version
chunhan
parents: 74
diff changeset
    73
               permission_check pctxt (sectxt_of_sfile sf) C_file P_read\<rbrakk>
74
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    74
  \<Longrightarrow> (update_ss_shms ss (pi, pctxt,fds,shms) (tagp \<or> tagf)) \<in> static"
76
f27ba31b7e96 no shm version
chunhan
parents: 74
diff changeset
    75
(*add following to locale assumptions:
f27ba31b7e96 no shm version
chunhan
parents: 74
diff changeset
    76
     \<And> pctxt fctxt. permission_check pctxt fctxt C_file P_read/P_write \<Longrightarrow>
f27ba31b7e96 no shm version
chunhan
parents: 74
diff changeset
    77
       permission_check pctxt pctxt C_fd P_setattr \<and>
f27ba31b7e96 no shm version
chunhan
parents: 74
diff changeset
    78
       search_check_s pctxt sf True; not \<forall> sf, should be \<forall> sf \<in> static \<and> sectxt_of_sfile sf = fctxt. 
f27ba31b7e96 no shm version
chunhan
parents: 74
diff changeset
    79
f27ba31b7e96 no shm version
chunhan
parents: 74
diff changeset
    80
 *)
74
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    81
| S_writef:  "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; (fdctxt,flags,sf) \<in> fds; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    82
               permission_check pctxt fdctxt C_fd P_setattr; sf \<in> sfs; S_file sfs tagf \<in> ss; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    83
               permission_check pctxt (sectxt_of_sfile sf) C_file P_write; is_write_flag flags\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    84
  \<Longrightarrow> (update_ss ss (S_file sfs tagf) (S_file sfs (tagp \<or> tagf))) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    85
| S_unlink:  "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_file sfs tagf \<in> ss;  
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    86
               (Init f,fsec,Some pfsec,asecs) \<in> sfs; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    87
               search_check_s pctxt (Init f,fsec,Some pfsec,asecs) True; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    88
               permission_check pctxt fsec C_file P_unlink; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    89
               permission_check pctxt pfsec C_dir P_remove_name\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    90
  \<Longrightarrow> ((ss - {S_file sfs tagf}) \<union> {S_file (sfs - {(Init f,fsec,Some pfsec,asecs)}) tagf}) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    91
| S_rmdir:   "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    92
               S_dir (fi,fsec,Some pfsec,asecs) \<in> ss;  
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    93
               search_check_s pctxt (fi,fsec,Some pfsec,asecs) False; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    94
               permission_check pctxt fsec C_dir P_rmdir;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    95
               permission_check pctxt pfsec C_dir P_remove_name\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    96
  \<Longrightarrow> (del_ss ss (S_dir (fi,fsec,Some pfsec,asecs))) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    97
| S_mkdir:   "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_dir (fi,fsec,pfsec,asecs) \<in> ss;  
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    98
               search_check_s pctxt (fi,fsec,pfsec,asecs) False; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    99
               permission_check pctxt (nfctxt_create pctxt fsec C_dir) C_dir P_create;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   100
               permission_check pctxt fsec C_dir P_add_name\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   101
  \<Longrightarrow> (add_ss ss (S_dir (Created,nfctxt_create pctxt fsec C_dir,Some fsec,asecs \<union> {fsec}))) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   102
| s_link:    "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_dir (pfi,pfsec,ppfsec,asecs) \<in> ss;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   103
               S_file sfs tagf \<in> ss; sf \<in> sfs; nfsec = nfctxt_create pctxt pfsec C_file;  
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   104
               search_check_s pctxt (pfi,pfsec,ppfsec,asecs) False; search_check_s pctxt sf True;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   105
               permission_check pctxt (sectxt_of_sfile sf) C_file P_link; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   106
               permission_check pctxt pfsec C_dir P_add_name\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   107
  \<Longrightarrow> (update_ss ss (S_file sfs tagf) 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   108
                  (S_file (sfs \<union> {(Created,nfsec,Some pfsec, asecs \<union> {pfsec})}) tagf)) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   109
| s_trunc:   "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_file sfs tagf \<in> ss; sf \<in> sfs; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   110
               search_check_s pctxt sf True; permission_check pctxt (sectxt_of_sfile sf) C_file P_setattr\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   111
  \<Longrightarrow> (update_ss ss (S_file sfs tagf) (S_file sfs (tagf \<or> tagp))) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   112
(*
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   113
| s_rename:  "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_file sfs tagf \<in> ss;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   114
               (sf#spf') \<in> sfs; S_dir spf tagpf \<in> ss; \<not>((sf#spf') \<preceq> (sf#spf)); 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   115
               search_check_s pctxt spf False; search_check_s pctxt (sf#spf') True; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   116
               sectxt_of_sfile (sf#spf') = Some fctxt; sectxt_of_sfile spf = Some pfctxt;  
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   117
               permission_check pctxt fctxt C_file P_rename;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   118
               permission_check pctxt pfctxt C_dir P_add_name;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   119
               ss_rename ss (sf#spf') (sf#spf) ss'; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   120
               ss_rename_no_same_fname ss (sf#spf') (sf#spf)\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   121
              \<Longrightarrow> ss' \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   122
| s_rename': "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_dir (sf#spf') tagf \<in> ss;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   123
               S_dir spf tagpf \<in> ss; \<not>((sf#spf') \<preceq> (sf#spf)); 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   124
               search_check_s pctxt spf False; search_check_s pctxt (sf#spf') True; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   125
               sectxt_of_sfile (sf#spf') = Some fctxt; sectxt_of_sfile spf = Some pfctxt;  
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   126
               permission_check pctxt fctxt C_dir P_reparent;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   127
               permission_check pctxt pfctxt C_dir P_add_name;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   128
               ss_rename ss (sf#spf') (sf#spf) ss'; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   129
               ss_rename_no_same_fname ss (sf#spf') (sf#spf)\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   130
              \<Longrightarrow> ss' \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   131
*)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   132
| s_createq: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   133
               permission_check pctxt pctxt C_msgq P_associate;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   134
               permission_check pctxt pctxt C_msgq P_create\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   135
  \<Longrightarrow> (add_ss ss (S_msgq (Created,pctxt,[]))) \<in> static" 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   136
| s_sendmsg: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_msgq (qi,qctxt,sms) \<in> ss;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   137
               permission_check pctxt qctxt C_msgq P_enqueue;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   138
               permission_check pctxt qctxt C_msgq P_write; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   139
               permission_check pctxt pctxt C_msg  P_create\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   140
  \<Longrightarrow> (update_ss ss (S_msgq (qi,qctxt,sms)) 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   141
                    (S_msgq (qi,qctxt,sms @ [(Created, pctxt, tagp)]))) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   142
| s_recvmsg: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   143
               S_msgq (qi,qctxt,(mi,mctxt,tagm)#sms) \<in> ss;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   144
               permission_check pctxt qctxt C_msgq P_read; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   145
               permission_check pctxt mctxt C_msg  P_receive\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   146
  \<Longrightarrow> (update_ss (update_ss_shms ss (pi,pctxt,fds,shms) (tagp \<or> tagm)) 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   147
                 (S_msgq (qi, qctxt, (mi, mctxt, tagm)#sms))
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   148
                 (S_msgq (qi, qctxt, sms))) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   149
| s_removeq: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_msgq (qi,qctxt,sms) \<in> ss;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   150
               permission_check pctxt qctxt C_msgq P_destroy\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   151
  \<Longrightarrow> (del_ss ss (S_msgq (qi,qctxt,sms))) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   152
| s_createh: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   153
               permission_check pctxt pctxt C_shm P_associate; 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   154
               permission_check pctxt pctxt C_shm P_create\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   155
   \<Longrightarrow> (add_ss ss (S_shm (Created, pctxt))) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   156
| s_attach:  "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_shm (hi,hctxt) \<in> ss;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   157
               if flag = SHM_RDONLY then permission_check pctxt hctxt C_shm P_read
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   158
               else (permission_check pctxt hctxt C_shm P_read \<and>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   159
                     permission_check pctxt hctxt C_shm P_write)\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   160
   \<Longrightarrow> (update_ss ss (S_proc (pi,pctxt,fds,shms) tagp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   161
                    (S_proc (pi,pctxt,fds,shms \<union> {((hi,hctxt),flag)}) tagp)) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   162
| s_detach:  "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_shm sh \<in> ss;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   163
               (sh,flag) \<in> shms; \<not> is_many_sshm sh\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   164
   \<Longrightarrow> (update_ss ss (S_proc (pi,pctxt,fds,shms) tagp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   165
                    (S_proc (pi,pctxt,fds,shms - {(sh,flag)}) tagp)) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   166
| s_deleteh: "\<lbrakk>ss \<in> static; S_proc (pi,pctxt,fds,shms) tagp \<in> ss; S_shm (hi,hctxt) \<in> ss;
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   167
               permission_check pctxt hctxt C_shm P_destroy; \<not> is_many_sshm sh\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   168
   \<Longrightarrow> (remove_sproc_sshm (del_ss ss (S_shm (hi,hctxt))) (hi,hctxt)) \<in> static"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   169
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
fun tainted_s :: "t_static_state \<Rightarrow> t_sobject \<Rightarrow> bool"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   172
where
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   173
  "tainted_s ss (S_proc sp  tag) = (S_proc sp tag  \<in> ss \<and> tag = True)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   174
| "tainted_s ss (S_file sfs tag) = (S_file sfs tag \<in> ss \<and> tag = True)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   175
(*
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   176
| "tainted_s ss (S_msg  (qi, sec, sms)  (mi, secm, tag)) = 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   177
     (S_msgq (qi, sec, sms) \<in> ss \<and> (mi,secm,tag) \<in> set sms \<and> tag = True)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   178
*)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   179
| "tainted_s ss _ = False"
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
fun tainted_s :: "t_object \<Rightarrow> t_static_state \<Rightarrow> bool"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   183
where 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   184
  "tainted_s (O_proc p)  ss = (\<exists> sp. S_proc sp True \<in> ss \<and> sproc_related p sp)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   185
| "tainted_s (O_file f)  ss = (\<exists> sfs sf. S_file sfs True \<in> ss \<and> sf \<in> sfs \<and> sfile_related f sf)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   186
| "tainted_s (O_msg q m) ss = (\<exists> sq. S_msgq sq \<in> ss \<and> smsgq_smsg_relatainted q m sq)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   187
| "tainted_s _           ss = False"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   188
*)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   189
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   190
definition taintable_s :: "t_object \<Rightarrow> bool"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   191
where
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   192
  "taintable_s obj \<equiv> \<exists> ss \<in> static. \<exists> sobj. tainted_s ss sobj \<and> init_obj_related sobj obj \<and> init_alive obj"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   193
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   194
(* this definition is wrong, cause process can exited from static-world 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   195
definition deletable_s :: "t_object \<Rightarrow> bool"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   196
where
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   197
  "deletable_s obj \<equiv> init_alive obj \<and> (\<exists> ss \<in> static. \<forall> sobj \<in> ss. \<not> init_obj_related sobj obj)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   198
*)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   199
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   200
fun deleted_s :: "t_sproc \<Rightarrow> t_object \<Rightarrow> bool"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   201
where
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   202
  "deleted_s (pi, pctxt, fds, shms) (O_proc p) = (pi \<noteq> Init p \<and> permission_check pctxt  "
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   203
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   204
definition deletable_s :: "t_object \<Rightarrow> bool"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   205
  "deletable_s obj \<equiv> init_alive obj \<and> (\<exists> ss \<in> static. \<exists> sp tagp. S_proc sp tagp \<in> ss \<and> deleted_s sp obj)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   206
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   207
definition undeletable_s :: "t_object \<Rightarrow> bool"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   208
where
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   209
  "undeletable_s obj \<equiv> init_alive obj \<and> (\<forall> ss \<in> static. \<exists> sobj \<in> ss. init_obj_related sobj obj)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   210
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   211
definition init_ss_eq:: "t_static_state \<Rightarrow> t_static_state \<Rightarrow> bool" (infix "\<doteq>" 100)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   212
where
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   213
  "ss \<doteq> ss' \<equiv> ss \<subseteq> ss' \<and> {sobj. is_init_sobj sobj \<and> sobj \<in> ss'} \<subseteq> ss"
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 [simp]: "ss \<doteq> ss"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   216
by (auto simp:init_ss_eq_def)
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
definition init_ss_in:: "t_static_state \<Rightarrow> t_static_state set \<Rightarrow> bool" (infix "\<propto>" 101)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   219
where
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   220
  "ss \<propto> sss \<equiv> \<exists> ss' \<in> sss. ss \<doteq> ss'"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   221
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   222
lemma s2ss_included_sobj:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   223
  "\<lbrakk>alive s obj; co2sobj s obj= Some sobj\<rbrakk> \<Longrightarrow> sobj \<in> (s2ss s)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   224
by (simp add:s2ss_def, rule_tac x = obj in exI, simp)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   225
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   226
lemma init_ss_in_prop:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   227
  "\<lbrakk>s2ss s \<propto> static; co2sobj s obj = Some sobj; alive s obj; init_obj_related sobj obj\<rbrakk>
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   228
   \<Longrightarrow> \<exists> ss \<in> static. sobj \<in> ss"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   229
apply (simp add:init_ss_in_def init_ss_eq_def)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   230
apply (erule bexE, erule conjE)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   231
apply (rule_tac x = ss' in bexI, auto dest!:s2ss_included_sobj)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   232
done
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   233
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   234
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   235
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   236
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   237
end
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   238
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
   239
end