Flask_type.thy
author chunhan
Mon, 06 Jan 2014 23:07:51 +0800
changeset 89 ded3f83f6cb9
parent 1 7d9c0ed02b56
permissions -rwxr-xr-x
make clear of Indexing-objects and enrich-objects
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     1
theory Flask_type
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     2
imports Main
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     3
begin
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     4
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     5
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     6
(************* types that for modeling Access-Control part of Selinux *************)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     7
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     8
(* Table 4-5 of "Selinux by example" *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
     9
datatype permission_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    10
(* C_file *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    11
  P_append (* the O_append flag for open *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    12
| P_create (* create new file *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    13
| P_entrypoint (* file can be used as the entry point of new domain via a domain transition *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    14
(* | P_execmod  execute memory-mapped files that have modified *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    15
| P_execute (* corresponds to x access in standard linux *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    16
(* | P_execute_no_trans: execute file in calller's domain, without domain transition *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    17
(* | P_getattr :  stat, ls, ioctls, ... not-related to security *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    18
| P_link    (* create hard link to file *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    19
| P_read    (* corresponds to r access in standard linux *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    20
| P_rename  (* rename a hard link *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    21
| P_setattr (* chmod, ioctls ... *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    22
| P_unlink  (* remove hard link (delete *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    23
| P_write   (* corresponds to w access in standard linux *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    24
(* P_lock; P_mounton; P_quotaon; P_relabelfrom; P_relabelto *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    25
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    26
(* C_dir *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    27
| P_add_name (* add file in directory *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    28
| P_remove_name (* like write for file *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    29
| P_reparent (* like rename for file *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    30
| P_search   
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    31
| P_rmdir    (* like unlink for file *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    32
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    33
(* C_fd *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    34
(* P_create; P_getattr; P_setattr*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    35
| P_inherit (* inherit across execve,
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    36
we simplify that execve would pass all fds to new "process", but flask checks the inherit right, if denied the fd is closed. *) 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    37
(* P_receive: ignored, ? I can find a reasonable explanation *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    38
  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    39
(* C_process *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    40
| P_fork 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    41
| P_ptrace
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    42
| P_transition
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    43
| P_share  (* allow state sharing for execve, clone: fds and ipc  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    44
 * for clone, we might have two flags, share Fd and IPC
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    45
 * for execve, as default, fds are shared, IPCs are detached. 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    46
 *   But fds not passed access-control are closed
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    47
 *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    48
(*| P_siginh:  inheritance of signal state, pending signals, this is not needed, as
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    49
we assume all signal are delivered instantly *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    50
(* P_rlimitnh : inheritance of resource limits *) 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    51
(* P_dyntransition; P_execheap; P_execmem; P_execstack; get_attr : not security related;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    52
 * P_getcap; P_getsched *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    53
| P_sigkill     (* signal *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    54
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    55
(* C_msgq, C_shm *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    56
| P_associate
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    57
| P_destroy 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    58
| P_enqueue
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    59
| P_receive
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    60
| P_send
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    61
(* create, read, write *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    62
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    63
type_synonym access_vector_t = "permission_t set"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    64
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    65
(************ Object Classes *************)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    66
datatype security_class_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    67
  C_process
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    68
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    69
(* file-related object classes, 4.3.1 of Book *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    70
| C_file
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    71
| C_fd 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    72
| C_dir   
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    73
(* C_blk file C_chr_file C_fifo_file C_filesystem C_lnk_file C_sock_file *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    74
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    75
(* network, 4.3.2, using only tcp/udp as a demonstration *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    76
| C_node    (* Host by IP address or range of addresses *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    77
| C_netif   (* network interface, eth0 *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    78
| C_tcp_socket 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    79
| C_udp_socket 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    80
(* netlink_socket unix_socket rawip_socket *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    81
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    82
(* IPCs, 4.3.3 *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    83
| C_msg  (* messages within a message queue *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    84
| C_msgq (* message queues *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    85
| C_shm  (* shared memory segment *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    86
(* C_semaphores: only for sync, not tainting related *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    87
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    88
(*other classes, 4.3.4*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    89
(* C_capability; C_security:security server;  *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    90
(* | C_system  the system itself *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    91
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    92
(********* TE rules **********
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    93
 * 1. type declaration rules
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    94
 * 2. type transition rules
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    95
 * 3. type member rules
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    96
 * 4. type change rules
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    97
 * 5. access vector rules
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    98
 * 6. cloning rules
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
    99
 * 7. access vector assertions 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   100
 *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   101
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   102
(* 1. type declaration rule *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   103
datatype type_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   104
  T_syslogd
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   105
| T_syslogd_exec
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   106
| T_devlog
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   107
  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   108
type_synonym attribute_t = "type_t set"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   109
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   110
(* for the asterisk/wildcard in the rules lists*)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   111
datatype 'a patt = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   112
  Single "'a" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   113
| Set    "'a set"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   114
| Tilde  "'a set"  (* complement *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   115
| Asterisk 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   116
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   117
(* 2. type transition rule, applied only when: execve & createfile
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   118
 * (creating subj_type, related obj_typ, type class, default type) 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   119
 * related obj_typ: proc \<rightarrow> executable; file \<rightarrow> parent-file 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   120
 * if not in the list, use subj_type for proc and parent_type for file
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   121
 * if multiple in the list, use the first one (they use the last one, we can say our list is the rev of their list, for "lookup" is easy)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   122
 * if using attributes, it can be transformed into the simple form of this list,
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   123
 now we use "patt" to deal with attributes
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   124
 *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   125
type_synonym type_transition_rule_list_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   126
  "(type_t patt \<times> type_t patt \<times> security_class_t \<times> type_t) list" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   127
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   128
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   129
(* for type member rules:
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   130
 * ?? I don't know the ideal yet, page 13 bottom of left column 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   131
 *
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   132
type_synonym type_member_rule_list_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   133
  "(type_t patt \<times> type_t patt \<times> security_class_t \<times> type_t) list" 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   134
 *)  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   135
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   136
(* for type change rules (relabeling to diff security context):
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   137
 * ?? I don't know what's its meaning yet, page 13 up of the right column
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   138
 *
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   139
type_synonym type_change_rule_list_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   140
  "(type_t patt \<times> type_t patt \<times> security_class_t \<times> type_t) list"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   141
 *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   142
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   143
datatype 'a target_with_self_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   144
  Not_self 'a
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   145
| Self
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   146
  
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   147
(* for access vector rules(verify a role-changed if allowed or not): 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   148
 * Any role-changed should be verified by this list!
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   149
 * here we only do the allow rule, audit and notify rules are ignored 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   150
 * (source type, target type, class, permissionS)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   151
 * permissionS might be access_vector_t, which is already be model as a "set" of "pattern"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   152
 * Self in the target type means this rule should be applied to the source process itself
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   153
 * if multiple, we search the whole list to mimic their "the union" ideal 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   154
 *) 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   155
type_synonym allowed_rule_list_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   156
  "(type_t patt \<times> (type_t patt) target_with_self_t \<times> security_class_t \<times> permission_t patt) list"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   157
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   158
datatype role_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   159
  R_object
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   160
| R_user
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   161
| R_login
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   162
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   163
(* role declarations:
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   164
 * (r, ts): ts are the types/domains that can be associated with r
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   165
 *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   166
type_synonym role_declarations_list_t =
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   167
  "(role_t \<times> type_t set) list"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   168
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   169
(* role transition rules(only happens when execve call): 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   170
 * (r, t, r'): the process with role r executes a file with type t, its role will be changed to r'
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   171
 *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   172
type_synonym role_transition_rule_list_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   173
  "(role_t \<times> type_t \<times> role_t) list"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   174
 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   175
(* role allow rules:
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   176
 * (from_role, to_role) changes when "Execve" calls!
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   177
 *  What's the difference between this list and the role_transition_rule_list?
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   178
 *  It is: if a process execute a file, the new role will be determined by 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   179
 *   role-transition, but this new role should be in role-allow(if not deny),
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   180
 *   and the new type should be compatible with the new role, which is checked
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   181
 *   by role-declarations(if not deny). 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   182
 *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   183
type_synonym role_allow_rule_list_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   184
  "(role_t set \<times> role_t set) list"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   185
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   186
datatype user_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   187
  U_system
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   188
| U_sds
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   189
| U_pal
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   190
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   191
(* user declarations:
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   192
 * (u, roles) roles are associated with user u
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   193
 * here MLS are not in consideration
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   194
 *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   195
type_synonym user_declarations_list_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   196
  "(user_t \<times> role_t set)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   197
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   198
type_synonym security_context_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   199
  "(user_t \<times> role_t \<times> type_t)"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   200
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   201
(* constrains :
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   202
 * (classes, perms, expression) list
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   203
 * when flask do a permission check, constrains will fetch the 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   204
 * associated two security-contexts (source & target), flask
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   205
 * makes sure that appling expression to these two contexts,
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   206
 * will get True. If False, deny the request. 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   207
 *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   208
type_synonym constrains_list_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   209
  "(security_class_t set \<times> permission_t set \<times> (security_context_t \<Rightarrow> security_context_t \<Rightarrow> bool)) list"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   210
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   211
(* comments: 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   212
 * 1. sid is not needed, it is a run-time index for security-context, 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   213
 *    what we do is policy-analysis, so no this "dynamic" sid as mid-ware;
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   214
 * 2. "seeds" of tainting cannot be solid objects in the initial states, 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   215
      should be extended as a "condition/predicate" to describe the "seeds"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   216
 *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   217
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   218
(* "temporary unused definitions :
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   219
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   220
datatype security_id_t = 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   221
  SID  nat
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   222
| SECSID_NULL
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   223
| SECSID_WILD
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   224
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   225
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   226
(* this list mimic the action of "security_compute_av",
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   227
 * if in the list, means func returns "grant"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   228
 * if not in, means "denied"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   229
 * !!! this list is a "work result" of policy by selinux's 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   230
 *  pre-processing func "security_compute_av"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   231
 *  in my formalisation, this is not needed, cause the specification of 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   232
 *  our "selinux" contains already the policy specification. 
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   233
 *  we can directly get the access-vector for sid/sectxt based on allow-rules & constrains
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   234
 *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   235
type_synonym avc_t = "(security_id_t \<times> security_id_t \<times> security_class_t \<times> access_vector_t) list"
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   236
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   237
 *)
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   238
7d9c0ed02b56 thy files
chunhan
parents:
diff changeset
   239
end