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