simple_selinux/Static_type.thy
author chunhan
Tue, 07 Jan 2014 22:04:06 +0800
changeset 90 003cac7b8bf5
parent 74 271e9818b6f6
permissions -rwxr-xr-x
enrich_msgq
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 Static_type 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     2
imports Main OS_type_def Flask_type OS_type_def
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
(* option: if some \<rightarrow> initial else none \<rightarrow> new created 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     6
 * the constructor each is one-to-one according to security_class_t
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     7
 * but dynamic-special objects are not included, such as:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     8
 *  1. file-descriptor: C_fd
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
     9
 *  2. network node   : C_node
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    10
 *  3. 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    11
 *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    12
(* Init \<rightarrow> exists in the initial state; Only: the only static object(cannot be cloned, e.g.)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    13
 * Many \<rightarrow> Many instantiation of this static object in the dynamic world 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    14
 * tainted under:
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    15
 *   Init \<rightarrow> this init obj is tainted
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    16
 *   Only \<rightarrow> this new created is tainted
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    17
 *   Many \<rightarrow> at least one of them is tainted or all of them can be tainted
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    18
 *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    19
datatype 'a Instance =
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    20
  Init 'a 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    21
| Created (* new created *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    22
(*
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    23
| Only  (* for process, cannot clone(that means no childprocesses in ss!); 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    24
           for file, renamed init-file(still keeps Init fname!)) *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    25
| Many
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
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    28
(* (fi, file sectxt, parent file's sectxt, ancestral files' sectxts) *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    29
type_synonym t_sfile = "(t_file Instance) \<times> security_context_t \<times> (security_context_t option) \<times> (security_context_t set)" 
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
type_synonym t_sfn = "(t_fname Instance) \<times> security_context_t"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    33
type_synonym t_sfile = "t_sfn list" (* apparently, "sfile set" is not a finite set! *)
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
(*
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    36
(* Sroot si sectxt : root sfile 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    37
 * Sfile si sectxt Sdir : normal sfile *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    38
datatype t_sfile = 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    39
  Sroot "t_file Instance" "security_context_t"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    40
| Sfile "t_file Instance" "security_context_t" "t_sfile"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    41
*)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    42
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    43
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    44
(* This is not finite set! think it as nature numbers 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    45
lemma finite_sfile_set: "finite (UNIV::(t_sfile set))"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    46
apply (rule finite.intros)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    47
unfolding UNIV_def
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    48
apply (induct x rule:t_sfile.induct)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    49
apply (simp add:UNIV_def)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    50
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    51
done *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    52
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    53
(* the flags tells the sfile is a file or dir: NO!
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    54
 * here sfile means sfd is opened to a sfile that has that "type", it is associated with all dynamic
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    55
 * files that according to this sfile. When we need to give such a file instance, we can give any dynamic
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    56
 * file in that state which according to this sfd's "creating static_state", so we might record ss in fds??
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    57
 * that means our static is a "(static-state list) set"?? *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    58
type_synonym t_sfd = "(security_context_t \<times> t_open_flags \<times> t_sfile)"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    59
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    60
(*
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    61
type_synonym t_sshm = "(t_shm Instance \<times> security_context_t)"
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
type_synonym t_sproc_sshm = "(t_sshm \<times> t_shm_attach_flag)"
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
(* (si, sectxt, sectxts of fds, sectxt of attached shms) *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    67
type_synonym t_sproc = 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    68
  "t_process Instance \<times> security_context_t \<times> (t_sfd set)"
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
(* here is a exception, the tainted-flag is in this type, case, msgq will not be-tainted *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    71
type_synonym t_smsg = "t_msg Instance \<times> security_context_t \<times> bool"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    72
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    73
(* (qmi, sectxt, sectxt of queue) *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    74
type_synonym t_smsgq = "t_msgq Instance \<times> security_context_t \<times> (t_smsg list)"
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
(* O_fd, O_*_sock, O_shm, O_msgq are all that can not be tainted *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    77
datatype t_sobject =
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    78
  S_proc "t_sproc"     "bool"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    79
| S_file "t_sfile"     "bool"  (* here sfile set is for "linking", tainted is the inode *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    80
| S_dir  "t_sfile"     
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    81
| S_msgq "t_smsgq"
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
| S_shm  "t_sshm"  
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    84
| S_msg  "t_smsgq"    "t_smsg"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    85
*)
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
datatype t_tainted_sobj = 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    88
  TS_proc "t_sproc"     "bool"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    89
| TS_file "t_sfile set" "bool"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    90
| TS_msg  "t_smsgq"     "t_smsg"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    91
*)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    92
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    93
(* the static state is the current static objects 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    94
 * But, how to record the static fds ??? fd is a dynamic concept !!! 
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    95
 *)
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    96
type_synonym t_static_state = "t_sobject set"
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    97
271e9818b6f6 remove shm and linkhard, make a simplified version of selinux
chunhan
parents:
diff changeset
    98
end