Paper.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Tue, 30 Apr 2013 14:46:18 +0100
changeset 4 8b6ba7168f2d
parent 3 4d25a9919688
child 8 2dab4bbb6684
permissions -rw-r--r--
pictures
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1
dcde836219bc add thy files
chunhan
parents:
diff changeset
     1
(*<*)
dcde836219bc add thy files
chunhan
parents:
diff changeset
     2
theory Paper
dcde836219bc add thy files
chunhan
parents:
diff changeset
     3
imports rc_theory final_theorems rc_theory os_rc 
dcde836219bc add thy files
chunhan
parents:
diff changeset
     4
begin
dcde836219bc add thy files
chunhan
parents:
diff changeset
     5
dcde836219bc add thy files
chunhan
parents:
diff changeset
     6
(* THEOREMS *)
dcde836219bc add thy files
chunhan
parents:
diff changeset
     7
dcde836219bc add thy files
chunhan
parents:
diff changeset
     8
dcde836219bc add thy files
chunhan
parents:
diff changeset
     9
notation (Rule output)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    10
  "==>"  ("\<^raw:\mbox{}\inferrule{\mbox{>_\<^raw:}}>\<^raw:{\mbox{>_\<^raw:}}>")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    11
dcde836219bc add thy files
chunhan
parents:
diff changeset
    12
syntax (Rule output)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    13
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    14
  ("\<^raw:\mbox{}\inferrule{>_\<^raw:}>\<^raw:{\mbox{>_\<^raw:}}>")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    15
dcde836219bc add thy files
chunhan
parents:
diff changeset
    16
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    17
  ("\<^raw:\mbox{>_\<^raw:}\\>/ _")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    18
dcde836219bc add thy files
chunhan
parents:
diff changeset
    19
  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    20
dcde836219bc add thy files
chunhan
parents:
diff changeset
    21
notation (Axiom output)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    22
  "Trueprop"  ("\<^raw:\mbox{}\inferrule{\mbox{}}{\mbox{>_\<^raw:}}>")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    23
dcde836219bc add thy files
chunhan
parents:
diff changeset
    24
notation (IfThen output)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    25
  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    26
syntax (IfThen output)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    27
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    28
  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    29
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}> /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    30
  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    31
dcde836219bc add thy files
chunhan
parents:
diff changeset
    32
notation (IfThenNoBox output)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    33
  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    34
syntax (IfThenNoBox output)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    35
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    36
  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    37
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    38
  "_asm" :: "prop \<Rightarrow> asms" ("_")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    39
dcde836219bc add thy files
chunhan
parents:
diff changeset
    40
(* insert *)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    41
notation (latex)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    42
  "Set.empty" ("\<emptyset>")
dcde836219bc add thy files
chunhan
parents:
diff changeset
    43
dcde836219bc add thy files
chunhan
parents:
diff changeset
    44
translations 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    45
  "{x} \<union> A" <= "CONST insert x A"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    46
  "{x,y}" <= "{x} \<union> {y}"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    47
  "{x,y} \<union> A" <= "{x} \<union> ({y} \<union> A)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    48
  "{x}" <= "{x} \<union> \<emptyset>"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    49
dcde836219bc add thy files
chunhan
parents:
diff changeset
    50
lemma impeq:
dcde836219bc add thy files
chunhan
parents:
diff changeset
    51
  "A = B \<Longrightarrow> (B \<Longrightarrow> A)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    52
by auto
dcde836219bc add thy files
chunhan
parents:
diff changeset
    53
dcde836219bc add thy files
chunhan
parents:
diff changeset
    54
dcde836219bc add thy files
chunhan
parents:
diff changeset
    55
dcde836219bc add thy files
chunhan
parents:
diff changeset
    56
dcde836219bc add thy files
chunhan
parents:
diff changeset
    57
consts DUMMY::'a
dcde836219bc add thy files
chunhan
parents:
diff changeset
    58
dcde836219bc add thy files
chunhan
parents:
diff changeset
    59
abbreviation
dcde836219bc add thy files
chunhan
parents:
diff changeset
    60
  "is_parent f pf \<equiv> (parent f = Some pf)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
    61
dcde836219bc add thy files
chunhan
parents:
diff changeset
    62
context tainting_s_sound begin 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    63
dcde836219bc add thy files
chunhan
parents:
diff changeset
    64
notation (latex output)
dcde836219bc add thy files
chunhan
parents:
diff changeset
    65
  source_dir ("anchor") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    66
  SProc ("P_\<^bsup>_\<^esup>") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    67
  SFile ("F_\<^bsup>_\<^esup>") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    68
  SIPC ("I'(_')\<^bsup>_\<^esup>") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    69
  READ ("Read") and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    70
  WRITE ("Write") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    71
  EXECUTE ("Execute") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    72
  CHANGE_OWNER ("ChangeOwner") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    73
  CREATE ("Create") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    74
  SEND ("Send") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    75
  RECEIVE ("Receive") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    76
  DELETE ("Delete") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    77
  compatible ("permissions") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    78
  comproles ("compatible") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    79
  DUMMY  ("\<^raw:\mbox{$\_$}>") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    80
  Cons ("_::_" [78,77] 79) and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    81
  Proc ("") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    82
  File ("") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    83
  File_type ("") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    84
  Proc_type ("") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    85
  IPC ("") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    86
  init_processes ("init'_procs") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    87
  os_grant ("admissible") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    88
  rc_grant ("granted") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    89
  exists ("alive") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    90
  default_fd_create_type ("default'_type") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    91
  InheritParent_file_type ("InheritPatentType") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    92
  NormalFile_type ("NormalFileType") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    93
  deleted ("deleted _ _" [50, 100] 100) and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    94
  taintable_s ("taintable\<^isup>s") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    95
  tainted_s ("tainted\<^isup>s") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    96
  all_sobjs ("reachable\<^isup>s") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    97
  init_obj2sobj ("\<lbrakk>_\<rbrakk>") and
dcde836219bc add thy files
chunhan
parents:
diff changeset
    98
  erole_functor ("erole'_aux") --"I have a erole_functor and etype_aux to handle 
dcde836219bc add thy files
chunhan
parents:
diff changeset
    99
                                 efficient, but their name not same, so ..., but don't work"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   100
dcde836219bc add thy files
chunhan
parents:
diff changeset
   101
dcde836219bc add thy files
chunhan
parents:
diff changeset
   102
abbreviation
dcde836219bc add thy files
chunhan
parents:
diff changeset
   103
  "is_process_type s p t \<equiv> (type_of_process s p = Some t)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   104
dcde836219bc add thy files
chunhan
parents:
diff changeset
   105
abbreviation
dcde836219bc add thy files
chunhan
parents:
diff changeset
   106
  "is_current_role s p r \<equiv> (currentrole s p = Some r)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   107
dcde836219bc add thy files
chunhan
parents:
diff changeset
   108
abbreviation
dcde836219bc add thy files
chunhan
parents:
diff changeset
   109
  "is_file_type s f t \<equiv> (etype_of_file s f = Some t)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   110
dcde836219bc add thy files
chunhan
parents:
diff changeset
   111
lemma osgrant2:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   112
  "\<lbrakk>p \<in> current_procs \<tau>; f \<notin> current_files \<tau>; parent f = Some pf; pf \<in> current_files \<tau>\<rbrakk> \<Longrightarrow>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   113
   os_grant \<tau> (CreateFile p f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   114
by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   115
dcde836219bc add thy files
chunhan
parents:
diff changeset
   116
lemma osgrant6:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   117
 "\<lbrakk>p \<in> current_procs \<tau>; u \<in> init_users\<rbrakk> \<Longrightarrow> os_grant \<tau> (ChangeOwner p u)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   118
by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   119
dcde836219bc add thy files
chunhan
parents:
diff changeset
   120
lemma osgrant10:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   121
  "\<lbrakk>p \<in> current_procs \<tau>; p' = new_proc \<tau>\<rbrakk> \<Longrightarrow> os_grant \<tau> (Clone p p')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   122
by simp 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   123
dcde836219bc add thy files
chunhan
parents:
diff changeset
   124
dcde836219bc add thy files
chunhan
parents:
diff changeset
   125
lemma rcgrant1:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   126
  "\<lbrakk>is_parent f pf; is_file_type s pf t; is_current_role s p r;
dcde836219bc add thy files
chunhan
parents:
diff changeset
   127
    default_fd_create_type r = InheritParent_file_type; 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   128
    (r, File_type t, WRITE) \<in> compatible\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   129
   \<Longrightarrow> rc_grant s (CreateFile p f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   130
by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   131
dcde836219bc add thy files
chunhan
parents:
diff changeset
   132
lemma rcgrant1':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   133
  "\<lbrakk>is_parent f pf; is_file_type s pf t; is_current_role s p r; 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   134
    default_fd_create_type r = NormalFile_type t'; 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   135
    (r, File_type t, WRITE) \<in> compatible;
dcde836219bc add thy files
chunhan
parents:
diff changeset
   136
    (r, File_type t', CREATE) \<in> compatible\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   137
   \<Longrightarrow> rc_grant s (CreateFile p f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   138
by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   139
dcde836219bc add thy files
chunhan
parents:
diff changeset
   140
lemma rcgrant4:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   141
  "\<lbrakk>is_current_role s p r; is_file_type s f t; (r, File_type t, EXECUTE) \<in> compatible\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   142
   \<Longrightarrow> rc_grant s (Execute p f)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   143
by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   144
dcde836219bc add thy files
chunhan
parents:
diff changeset
   145
lemma rcgrant7: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   146
  "\<lbrakk>is_current_role s p r; r' \<in> comproles r\<rbrakk> \<Longrightarrow> rc_grant s (ChangeRole p r')"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   147
by simp
dcde836219bc add thy files
chunhan
parents:
diff changeset
   148
dcde836219bc add thy files
chunhan
parents:
diff changeset
   149
lemma rcgrant_CHO:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   150
"\<lbrakk>is_current_role s p r; 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   151
  type_of_process s p = Some t;
dcde836219bc add thy files
chunhan
parents:
diff changeset
   152
  (r, Proc_type t, CHANGE_OWNER) \<in> compatible\<rbrakk> \<Longrightarrow> rc_grant s (ChangeOwner p u)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   153
by(simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   154
dcde836219bc add thy files
chunhan
parents:
diff changeset
   155
lemma pf_in_current_paper:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   156
  "\<lbrakk>is_parent f pf; f \<in> current_files s; valid s\<rbrakk> \<Longrightarrow> pf \<in> current_files s"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   157
by (simp add:parent_file_in_current)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   158
dcde836219bc add thy files
chunhan
parents:
diff changeset
   159
lemma dels:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   160
  shows "deleted (Proc p') ((Kill p p')#s)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   161
  and   "deleted (File f) ((DeleteFile p f)#s)" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   162
  and   "deleted (IPC i) ((DeleteIPC p i)#s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   163
  and   "deleted obj s \<Longrightarrow> deleted obj (e#s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   164
apply simp_all
dcde836219bc add thy files
chunhan
parents:
diff changeset
   165
apply(case_tac e)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   166
apply(simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   167
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   168
dcde836219bc add thy files
chunhan
parents:
diff changeset
   169
lemma tainted_10:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   170
  "\<lbrakk>(File f) \<in> tainted s; valid (e # s); f \<in> current_files (e # s)\<rbrakk> 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   171
  \<Longrightarrow> (File f) \<in> tainted (e # s)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   172
apply(rule tainted.intros)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   173
apply(assumption)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   174
apply(assumption)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   175
apply(simp only: exists.simps)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   176
done
dcde836219bc add thy files
chunhan
parents:
diff changeset
   177
dcde836219bc add thy files
chunhan
parents:
diff changeset
   178
definition
dcde836219bc add thy files
chunhan
parents:
diff changeset
   179
  Init ("init _")
dcde836219bc add thy files
chunhan
parents:
diff changeset
   180
where
dcde836219bc add thy files
chunhan
parents:
diff changeset
   181
  "Init obj \<equiv> exists [] obj" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   182
dcde836219bc add thy files
chunhan
parents:
diff changeset
   183
lemma Init_rhs:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   184
  shows "Init (File f) = (f \<in> init_files)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   185
  and   "Init (Proc p) = (p \<in> init_processes)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   186
  and   "Init (IPC i) = (i \<in> init_ipcs)"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   187
unfolding Init_def
dcde836219bc add thy files
chunhan
parents:
diff changeset
   188
by(simp_all)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   189
dcde836219bc add thy files
chunhan
parents:
diff changeset
   190
notation (latex output)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   191
  Init ("_ \<in> init")
dcde836219bc add thy files
chunhan
parents:
diff changeset
   192
dcde836219bc add thy files
chunhan
parents:
diff changeset
   193
lemma af_init':
dcde836219bc add thy files
chunhan
parents:
diff changeset
   194
  "\<lbrakk>f \<in> init_files; is_file_type [] f t\<rbrakk>
dcde836219bc add thy files
chunhan
parents:
diff changeset
   195
  \<Longrightarrow> SFile (t, f) (Some f) \<in> all_sobjs"
dcde836219bc add thy files
chunhan
parents:
diff changeset
   196
apply(rule af_init)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   197
apply(simp)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   198
by (simp add:etype_of_file_def)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   199
dcde836219bc add thy files
chunhan
parents:
diff changeset
   200
declare [[show_question_marks = false]]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   201
dcde836219bc add thy files
chunhan
parents:
diff changeset
   202
dcde836219bc add thy files
chunhan
parents:
diff changeset
   203
(*>*)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   204
dcde836219bc add thy files
chunhan
parents:
diff changeset
   205
section {* Introduction *}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   206
dcde836219bc add thy files
chunhan
parents:
diff changeset
   207
text {*
dcde836219bc add thy files
chunhan
parents:
diff changeset
   208
  Role-based access control models are used in many operating systems
dcde836219bc add thy files
chunhan
parents:
diff changeset
   209
  for enforcing security properties.  The
dcde836219bc add thy files
chunhan
parents:
diff changeset
   210
  \emph{Role-Compatibility Model} (RC-Model), introduced by Ott
dcde836219bc add thy files
chunhan
parents:
diff changeset
   211
  \cite{ottrc,ottthesis}, is one such role-based access control
dcde836219bc add thy files
chunhan
parents:
diff changeset
   212
  model. It defines \emph{roles}, which are associated with processes,
dcde836219bc add thy files
chunhan
parents:
diff changeset
   213
  and defines \emph{types}, which are associated with system
dcde836219bc add thy files
chunhan
parents:
diff changeset
   214
  resources, such as files and directories. The RC-Model also includes
dcde836219bc add thy files
chunhan
parents:
diff changeset
   215
  types for interprocess communication, that is message queues,
dcde836219bc add thy files
chunhan
parents:
diff changeset
   216
  sockets and shared memory. A policy in the RC-Model gives every user
dcde836219bc add thy files
chunhan
parents:
diff changeset
   217
  a default role, and also specifies how roles can be
dcde836219bc add thy files
chunhan
parents:
diff changeset
   218
  changed. Moreover, it specifies which types of resources a role has
dcde836219bc add thy files
chunhan
parents:
diff changeset
   219
  permission to access, and also the \emph{mode} with which the role
dcde836219bc add thy files
chunhan
parents:
diff changeset
   220
  can access the resources, for example read, write, send, receive and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   221
  so on.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   222
   
dcde836219bc add thy files
chunhan
parents:
diff changeset
   223
  The RC-Model is built on top of a collection of system calls
dcde836219bc add thy files
chunhan
parents:
diff changeset
   224
  provided by the operating system, for instance system calls for
dcde836219bc add thy files
chunhan
parents:
diff changeset
   225
  reading and writing files, cloning and killing of processes, and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   226
  sending and receiving messages. The purpose of the RC-Model is to
dcde836219bc add thy files
chunhan
parents:
diff changeset
   227
  restrict access to these system calls and thereby enforce security
dcde836219bc add thy files
chunhan
parents:
diff changeset
   228
  properties of the system. A problem with the RC-Model and role-based
dcde836219bc add thy files
chunhan
parents:
diff changeset
   229
  access control models in general is that a system administrator has
dcde836219bc add thy files
chunhan
parents:
diff changeset
   230
  to specify an appropriate access control policy. The difficulty with
dcde836219bc add thy files
chunhan
parents:
diff changeset
   231
  this is that \emph{``what you specify is what you get but not
dcde836219bc add thy files
chunhan
parents:
diff changeset
   232
  necessarily what you want''} \cite[Page 242]{Jha08}. To overcome
dcde836219bc add thy files
chunhan
parents:
diff changeset
   233
  this difficulty, a system administrator needs some kind of sanity
dcde836219bc add thy files
chunhan
parents:
diff changeset
   234
  check for whether an access control policy is really securing
dcde836219bc add thy files
chunhan
parents:
diff changeset
   235
  resources.  Existing works, for example \cite{sanity01,sanity02},
dcde836219bc add thy files
chunhan
parents:
diff changeset
   236
  provide sanity checks for policies by specifying properties and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   237
  using model checking techniques to ensure a policy at hand satisfies
dcde836219bc add thy files
chunhan
parents:
diff changeset
   238
  these properties. However, these checks only address the problem on
dcde836219bc add thy files
chunhan
parents:
diff changeset
   239
  the level of policies---they can only check ``on the surface''
dcde836219bc add thy files
chunhan
parents:
diff changeset
   240
  whether the policy reflects the intentions of the system
dcde836219bc add thy files
chunhan
parents:
diff changeset
   241
  administrator---these checks are not justified by the actual
dcde836219bc add thy files
chunhan
parents:
diff changeset
   242
  behaviour of the operating system. The main problem this paper addresses is to check
dcde836219bc add thy files
chunhan
parents:
diff changeset
   243
  when a policy matches the intentions of a system administrator
dcde836219bc add thy files
chunhan
parents:
diff changeset
   244
  \emph{and} given such a policy, the operating system actually
dcde836219bc add thy files
chunhan
parents:
diff changeset
   245
  enforces this policy.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   246
 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   247
  Our work is related to the preliminary work by Archer et al
dcde836219bc add thy files
chunhan
parents:
diff changeset
   248
  \cite{Archer03} about the security model of SELinux. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   249
  They also give a dynamic model of system calls on which the access
dcde836219bc add thy files
chunhan
parents:
diff changeset
   250
  controls are implemented. Their dynamic model is defined in terms of
dcde836219bc add thy files
chunhan
parents:
diff changeset
   251
  IO automata and mechanised in the PVS theorem prover. For specifying
dcde836219bc add thy files
chunhan
parents:
diff changeset
   252
  and reasoning about automata they use the TAME tool in PVS. Their work checks
dcde836219bc add thy files
chunhan
parents:
diff changeset
   253
  well-formedness properties of access policies by type-checking 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   254
  generated definitions in PVS. They can also ensure some ``\emph{simple
dcde836219bc add thy files
chunhan
parents:
diff changeset
   255
  properties}'' (their terminology), for example whether a process
dcde836219bc add thy files
chunhan
parents:
diff changeset
   256
  with a particular PID is present in every reachable state from 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   257
  an initial state. They also consider ``\emph{deeper properties}'', for
dcde836219bc add thy files
chunhan
parents:
diff changeset
   258
  example whether only a process with root-permissions 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   259
  or one of its descendents ever gets permission to write to kernel
dcde836219bc add thy files
chunhan
parents:
diff changeset
   260
  log files. They write that they can state such deeper 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   261
  properties about access policies, but about checking such properties
dcde836219bc add thy files
chunhan
parents:
diff changeset
   262
  they write that ``\emph{the feasibility of doing 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   263
  so is currently an open question}'' \cite[Page 167]{Archer03}.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   264
  We improve upon their results by using our sound and complete 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   265
  static policy check to make this feasible. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   266
dcde836219bc add thy files
chunhan
parents:
diff changeset
   267
  Our formal models and correctness proofs are mechanised in the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   268
  interactive theorem prover Isabelle/HOL. The mechanisation of the models is a
dcde836219bc add thy files
chunhan
parents:
diff changeset
   269
  prerequisite for any correctness proof about the RC-Model, since it
dcde836219bc add thy files
chunhan
parents:
diff changeset
   270
  includes a large number of interdependent concepts and very complex
dcde836219bc add thy files
chunhan
parents:
diff changeset
   271
  operations that determine roles and types. In our opinion it is
dcde836219bc add thy files
chunhan
parents:
diff changeset
   272
  futile to attempt to reason about them by just using ``pencil-and-paper''.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   273
  Following good experience in earlier mechanisation work
dcde836219bc add thy files
chunhan
parents:
diff changeset
   274
  \cite{ZhangUrbanWu12}, we use Paulson's inductive method for
dcde836219bc add thy files
chunhan
parents:
diff changeset
   275
  reasoning about sequences of events \cite{Paulson98}.  For example
dcde836219bc add thy files
chunhan
parents:
diff changeset
   276
  we model system calls as events and reason about an inductive
dcde836219bc add thy files
chunhan
parents:
diff changeset
   277
  definition of valid traces, that is lists of events.  Central to
dcde836219bc add thy files
chunhan
parents:
diff changeset
   278
  this paper is a notion of a resource being \emph{tainted}, which for
dcde836219bc add thy files
chunhan
parents:
diff changeset
   279
  example means it contains a virus or a back door.  We use our model
dcde836219bc add thy files
chunhan
parents:
diff changeset
   280
  of system calls in order to characterise how such a tainted object
dcde836219bc add thy files
chunhan
parents:
diff changeset
   281
  can ``spread'' through the system.  For a system administrator the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   282
  important question is whether such a tainted file, possibly
dcde836219bc add thy files
chunhan
parents:
diff changeset
   283
  introduced by a user, can affect core system files and render the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   284
  whole system insecure, or whether it can be contained by the access
dcde836219bc add thy files
chunhan
parents:
diff changeset
   285
  policy. Our results show that a corresponding check can be performed
dcde836219bc add thy files
chunhan
parents:
diff changeset
   286
  statically by analysing the initial state of the system and the access policy.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   287
  \smallskip
dcde836219bc add thy files
chunhan
parents:
diff changeset
   288
 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   289
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   290
  {\bf Contributions:} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   291
  We give a complete formalisation of the RC-Model in the interactive
dcde836219bc add thy files
chunhan
parents:
diff changeset
   292
  theorem prover Isabelle/HOL.  We also give a dynamic model of the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   293
  operating system by formalising all security related events that can
dcde836219bc add thy files
chunhan
parents:
diff changeset
   294
  happen while the system is running.  As far as we are aware, we are
dcde836219bc add thy files
chunhan
parents:
diff changeset
   295
  the first ones who formally prove that if a policy in the RC-Model
dcde836219bc add thy files
chunhan
parents:
diff changeset
   296
  satisfies an access property, then there is no sequence of events
dcde836219bc add thy files
chunhan
parents:
diff changeset
   297
  (system calls) that can violate this access property.  We also prove
dcde836219bc add thy files
chunhan
parents:
diff changeset
   298
  the opposite: if a policy does not meet an access property, then
dcde836219bc add thy files
chunhan
parents:
diff changeset
   299
  there is a sequence of events that will violate this property in our
dcde836219bc add thy files
chunhan
parents:
diff changeset
   300
  model of the operating system.  With these two results in place we
dcde836219bc add thy files
chunhan
parents:
diff changeset
   301
  can show that a static policy check is sufficient in order to
dcde836219bc add thy files
chunhan
parents:
diff changeset
   302
  guarantee the access properties before running the system. Again as
dcde836219bc add thy files
chunhan
parents:
diff changeset
   303
  far as we know, no such check that is the operating
dcde836219bc add thy files
chunhan
parents:
diff changeset
   304
  system behaviour has been designed before.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   305
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   306
dcde836219bc add thy files
chunhan
parents:
diff changeset
   307
  %Specified dynamic behaviour of the system; 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   308
  %we specified a static AC model; designed a tainted relation for 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   309
  %the system; proved that they coincide.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   310
  %In our paper ....
dcde836219bc add thy files
chunhan
parents:
diff changeset
   311
   
dcde836219bc add thy files
chunhan
parents:
diff changeset
   312
*}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   313
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   314
section {* Preliminaries about the RC-Model *}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   315
dcde836219bc add thy files
chunhan
parents:
diff changeset
   316
dcde836219bc add thy files
chunhan
parents:
diff changeset
   317
text {*
dcde836219bc add thy files
chunhan
parents:
diff changeset
   318
  The Role-Compatibility Model (RC-Model) is a role-based access
dcde836219bc add thy files
chunhan
parents:
diff changeset
   319
  control model. It has been introduced by Ott \cite{ottrc} and is
dcde836219bc add thy files
chunhan
parents:
diff changeset
   320
  used in running systems for example to secure Apache servers. It
dcde836219bc add thy files
chunhan
parents:
diff changeset
   321
  provides a more fine-grained control over access permissions than
dcde836219bc add thy files
chunhan
parents:
diff changeset
   322
  simple Unix-style access control models. This more fine-grained
dcde836219bc add thy files
chunhan
parents:
diff changeset
   323
  control solves the problem of server processes running as root with
dcde836219bc add thy files
chunhan
parents:
diff changeset
   324
  too many access permissions in order to accomplish a task at
dcde836219bc add thy files
chunhan
parents:
diff changeset
   325
  hand. In the RC-Model, system administrators are able to restrict
dcde836219bc add thy files
chunhan
parents:
diff changeset
   326
  what the role of server is allowed to do and in doing so reduce the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   327
  attack surface of a system.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   328
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   329
  Policies in the RC-Model talk about \emph{users}, \emph{roles},
dcde836219bc add thy files
chunhan
parents:
diff changeset
   330
  \emph{types} and \emph{objects}.  Objects are processes, files or
dcde836219bc add thy files
chunhan
parents:
diff changeset
   331
  IPCs (interprocess communication objects---such as message queues,
dcde836219bc add thy files
chunhan
parents:
diff changeset
   332
  sockets and shared memory). Objects are the resources of a system an
dcde836219bc add thy files
chunhan
parents:
diff changeset
   333
  RC-policy can restrict access to.  In what follows we use the letter
dcde836219bc add thy files
chunhan
parents:
diff changeset
   334
  @{term u} to stand for users, @{text r} for roles, @{term p} for
dcde836219bc add thy files
chunhan
parents:
diff changeset
   335
  processes, @{term f} for files and @{term i} for IPCs. We also
dcde836219bc add thy files
chunhan
parents:
diff changeset
   336
  use @{text obj} as a generic variable for objects.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   337
  The RC-Model has the following eight kinds of access modes to objects:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   338
dcde836219bc add thy files
chunhan
parents:
diff changeset
   339
  \begin{isabelle}\ \ \ \ \ %%%
dcde836219bc add thy files
chunhan
parents:
diff changeset
   340
  \begin{tabular}{@ {}l}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   341
  @{term READ}, @{term WRITE}, @{term EXECUTE}, @{term "CHANGE_OWNER"}, 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   342
  @{term CREATE}, @{term SEND}, @{term RECEIVE} and @{term DELETE}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   343
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   344
  \end{isabelle}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   345
dcde836219bc add thy files
chunhan
parents:
diff changeset
   346
  In the RC-Model, roles group users according to tasks they need to 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   347
  accomplish. Users have a default role specified by the policy,
dcde836219bc add thy files
chunhan
parents:
diff changeset
   348
  which is the role they start with whenever they log into the system.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   349
  A process contains the information about its owner (a user), its
dcde836219bc add thy files
chunhan
parents:
diff changeset
   350
  role and its type, whereby a type in the RC-Model allows system
dcde836219bc add thy files
chunhan
parents:
diff changeset
   351
  administrators to group resources according to a common criteria.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   352
  Such detailed information is needed in the RC-Model, for example, in
dcde836219bc add thy files
chunhan
parents:
diff changeset
   353
  order to allow a process to change its ownership. For this the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   354
  RC-Model checks the role of the process and its type: if the access
dcde836219bc add thy files
chunhan
parents:
diff changeset
   355
  control policy states that the role has @{term CHANGE_OWNER} access mode for
dcde836219bc add thy files
chunhan
parents:
diff changeset
   356
  processes of that type, then the process is permitted to assume a
dcde836219bc add thy files
chunhan
parents:
diff changeset
   357
  new owner.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   358
dcde836219bc add thy files
chunhan
parents:
diff changeset
   359
  Files in the RC-Model contain the information about their types.  A
dcde836219bc add thy files
chunhan
parents:
diff changeset
   360
  policy then specifies whether a process with a given role can access
dcde836219bc add thy files
chunhan
parents:
diff changeset
   361
  a file under a certain access mode. Files, however, also
dcde836219bc add thy files
chunhan
parents:
diff changeset
   362
  include in the RC-Model information about roles. This information is
dcde836219bc add thy files
chunhan
parents:
diff changeset
   363
  used when a process is permitted to execute a file. By doing so it
dcde836219bc add thy files
chunhan
parents:
diff changeset
   364
  might change its role. This is often used in the context of
dcde836219bc add thy files
chunhan
parents:
diff changeset
   365
  web-servers when a cgi-script is uploaded and then executed by the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   366
  server.  The resulting process should have much more restricted
dcde836219bc add thy files
chunhan
parents:
diff changeset
   367
  access permissions.  This kind of behaviour when executing a file can
dcde836219bc add thy files
chunhan
parents:
diff changeset
   368
  be specified in an RC-policy in several ways: first, the role of the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   369
  process does not change when executing a file; second, the process
dcde836219bc add thy files
chunhan
parents:
diff changeset
   370
  takes on the role specified with the file; or third, use the role of
dcde836219bc add thy files
chunhan
parents:
diff changeset
   371
  the owner, who currently owns this process. The RC-Model also makes
dcde836219bc add thy files
chunhan
parents:
diff changeset
   372
  assumptions on how types can change.  For example for files and IPCs
dcde836219bc add thy files
chunhan
parents:
diff changeset
   373
  the type can never change once they are created. But processes can
dcde836219bc add thy files
chunhan
parents:
diff changeset
   374
  change their types according to the roles they have.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   375
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   376
  As can be seen, the information contained in a policy in the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   377
  RC-Model can be rather complex: Roles and types, for example, are
dcde836219bc add thy files
chunhan
parents:
diff changeset
   378
  policy-dependent, meaning each policy needs to define a set of roles and a
dcde836219bc add thy files
chunhan
parents:
diff changeset
   379
  set of types. Apart from recording for each role the information
dcde836219bc add thy files
chunhan
parents:
diff changeset
   380
  which type of resource it can access and under which access-mode, it
dcde836219bc add thy files
chunhan
parents:
diff changeset
   381
  also needs to include a role compatibility set. This set specifies how one
dcde836219bc add thy files
chunhan
parents:
diff changeset
   382
  role can change into another role. Moreover it needs to include default
dcde836219bc add thy files
chunhan
parents:
diff changeset
   383
  information for cases when new processes or files are created.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   384
  For example, when a process clones itself, the type of the new
dcde836219bc add thy files
chunhan
parents:
diff changeset
   385
  process is determined as follows: the policy might specify a default
dcde836219bc add thy files
chunhan
parents:
diff changeset
   386
  type whenever a process with a certain role is cloned, or the policy
dcde836219bc add thy files
chunhan
parents:
diff changeset
   387
  might specify that the cloned process inherits the type of the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   388
  parent process.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   389
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   390
  Ott implemented the RC-Model on top of Linux, but only specified it
dcde836219bc add thy files
chunhan
parents:
diff changeset
   391
  as a set of informal rules, partially given as logic formulas,
dcde836219bc add thy files
chunhan
parents:
diff changeset
   392
  partially given as rules in ``English''. Unfortunately, some
dcde836219bc add thy files
chunhan
parents:
diff changeset
   393
  presentations about the RC-Model give conflicting definitions for
dcde836219bc add thy files
chunhan
parents:
diff changeset
   394
  some concepts---for example when defining the semantics of the special role
dcde836219bc add thy files
chunhan
parents:
diff changeset
   395
  ``inherit parent''. In \cite{ottrc} it means inherit the initial role
dcde836219bc add thy files
chunhan
parents:
diff changeset
   396
  of the parent directory, but in \cite{ottweb} it means inherit
dcde836219bc add thy files
chunhan
parents:
diff changeset
   397
  the role of the parent process.  In our formalisation we mainly follow the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   398
  version given in \cite{ottrc}. In the next section we give a mechanised
dcde836219bc add thy files
chunhan
parents:
diff changeset
   399
  model of the system calls on which the RC-Model is implemented.  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   400
*}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   401
dcde836219bc add thy files
chunhan
parents:
diff changeset
   402
dcde836219bc add thy files
chunhan
parents:
diff changeset
   403
dcde836219bc add thy files
chunhan
parents:
diff changeset
   404
section {* Dynamic Model of System Calls *}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   405
dcde836219bc add thy files
chunhan
parents:
diff changeset
   406
text {*
dcde836219bc add thy files
chunhan
parents:
diff changeset
   407
  Central to the RC-Model are processes, since they initiate any action
dcde836219bc add thy files
chunhan
parents:
diff changeset
   408
  involving resources and access control. We use natural numbers to stand for process IDs,
dcde836219bc add thy files
chunhan
parents:
diff changeset
   409
  but do not model the fact that the number of processes in any practical 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   410
  system is limited. Similarly, IPCs and users are represented by natural 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   411
  numbers. The thirteen actions a process can perform are represented by
dcde836219bc add thy files
chunhan
parents:
diff changeset
   412
  the following datatype of \emph{events}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   413
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   414
  \begin{isabelle}\ \ \ \ \ %%%
dcde836219bc add thy files
chunhan
parents:
diff changeset
   415
  \mbox{
dcde836219bc add thy files
chunhan
parents:
diff changeset
   416
  \begin{tabular}{r@ {\hspace{1.5mm}}c@ {\hspace{1.5mm}}l@ {\hspace{3mm}}l@ 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   417
                     {\hspace{1.5mm}}l@ {\hspace{3mm}}l@ {\hspace{1.5mm}}l@ 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   418
                     {\hspace{3mm}}l@ {\hspace{1.5mm}}l}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   419
  event 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   420
  & @{text "::="} & @{term "CreateFile p f"} & @{text "|"} & @{term "ReadFile p f"} & @{text "|"} & @{term "Send p i"}   & @{text "|"} & @{term "Kill p p'"}  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   421
  & @{text "|"} & @{term "WriteFile p f"}    & @{text "|"} & @{term "Execute p f"} & @{text "|"} & @{term "Recv p i"}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   422
  & @{text "|"} & @{term "DeleteFile p f"}   & @{text "|"} & @{term "Clone p p'"} & @{text "|"} & @{term "CreateIPC p i"}  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   423
  & @{text "|"} & @{term "ChangeOwner p u"}  & @{text "|"} & @{term "ChangeRole p r"} & @{text "|"} & @{term "DeleteIPC p i"}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   424
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   425
  \end{isabelle}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   426
dcde836219bc add thy files
chunhan
parents:
diff changeset
   427
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   428
  with the idea that for example in @{term Clone} a process @{term p} is cloned
dcde836219bc add thy files
chunhan
parents:
diff changeset
   429
  and the new process has the ID @{term "p'"}; with @{term Kill} the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   430
  intention is that the process @{term p} kills another process with
dcde836219bc add thy files
chunhan
parents:
diff changeset
   431
  ID @{term p'}. We will later give the definition what the role
dcde836219bc add thy files
chunhan
parents:
diff changeset
   432
  @{term r} can stand for in the constructor @{term ChangeRole}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   433
  (namely \emph{normal roles} only). As is custom in Unix, there is no
dcde836219bc add thy files
chunhan
parents:
diff changeset
   434
  difference between a directory and a file. The files @{term f} in
dcde836219bc add thy files
chunhan
parents:
diff changeset
   435
  the definition above are simply lists of strings.  For example, the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   436
  file @{text "/usr/bin/make"} is represented by the list @{text
dcde836219bc add thy files
chunhan
parents:
diff changeset
   437
  "[make, bin, usr]"} and the @{text root}-directory is the @{text
dcde836219bc add thy files
chunhan
parents:
diff changeset
   438
  Nil}-list. Following the presentation in \cite{ottrc}, our model of
dcde836219bc add thy files
chunhan
parents:
diff changeset
   439
  IPCs is rather simple-minded: we only have events for creation and deletion Of IPCs,
dcde836219bc add thy files
chunhan
parents:
diff changeset
   440
  as well as sending and receiving messages.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   441
dcde836219bc add thy files
chunhan
parents:
diff changeset
   442
  Events essentially transform one state of the system into
dcde836219bc add thy files
chunhan
parents:
diff changeset
   443
  another. The system starts with an initial state determining which 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   444
  processes, files and IPCs are active at the start of the system. We assume the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   445
  users of the system are fixed in the initial state; we also assume
dcde836219bc add thy files
chunhan
parents:
diff changeset
   446
  that the policy does not change while the system is running. We have 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   447
  three sets, namely
dcde836219bc add thy files
chunhan
parents:
diff changeset
   448
  @{term init_processes},
dcde836219bc add thy files
chunhan
parents:
diff changeset
   449
  @{term init_files} and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   450
  @{term init_ipcs}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   451
  specifying the processes, files and IPCs present in the initial state.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   452
  We will often use the abbreviation
dcde836219bc add thy files
chunhan
parents:
diff changeset
   453
dcde836219bc add thy files
chunhan
parents:
diff changeset
   454
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   455
  @{thm (lhs) Init_def} @{text "\<equiv>"} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   456
  @{thm (rhs) Init_rhs(1)[where f=obj]} @{text "\<or>"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   457
  @{thm (rhs) Init_rhs(2)[where p=obj]} @{text "\<or>"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   458
  @{thm (rhs) Init_rhs(3)[where i=obj]}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   459
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   460
dcde836219bc add thy files
chunhan
parents:
diff changeset
   461
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   462
  There are some assumptions we make about the files present in the initial state: we always
dcde836219bc add thy files
chunhan
parents:
diff changeset
   463
  require that the @{text "root"}-directory @{term "[]"} is part of the initial state
dcde836219bc add thy files
chunhan
parents:
diff changeset
   464
  and for every file in the initial state (excluding @{term "[]"}) we require that its 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   465
  parent is also part of the 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   466
  initial state.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   467
  After the initial state, the next states are determined
dcde836219bc add thy files
chunhan
parents:
diff changeset
   468
  by a list of events, called the \emph{trace}.  We need to define
dcde836219bc add thy files
chunhan
parents:
diff changeset
   469
  functions that allow us to make some observations about traces. One 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   470
  such function is called @{term "current_procs"} and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   471
  calculates the set of ``alive'' processes in a state:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   472
dcde836219bc add thy files
chunhan
parents:
diff changeset
   473
  %initial state:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   474
  %We make assumptions about the initial state, they're:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   475
  %1. there exists a set of processes, files, IPCs and users already in the initial state,
dcde836219bc add thy files
chunhan
parents:
diff changeset
   476
  %users are not changed in system's running, we regards users adding and deleting a
dcde836219bc add thy files
chunhan
parents:
diff changeset
   477
  %administration task, not the issue for our policy checker; 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   478
  %2. every object in the initial state have got already roles/types/owner ... information assigned;
dcde836219bc add thy files
chunhan
parents:
diff changeset
   479
  %3. all the policy information are already preloaded in the initial state, including:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   480
  %a compatible type table, @{term compatible};
dcde836219bc add thy files
chunhan
parents:
diff changeset
   481
  %a mapping function from a role to its compatible role set, @{term comproles};
dcde836219bc add thy files
chunhan
parents:
diff changeset
   482
  %every role's default values is pre-set, e.g. default process create type and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   483
  %and default file/directory create type.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   484
dcde836219bc add thy files
chunhan
parents:
diff changeset
   485
  \begin{isabelle}\ \ \ \ \ %%%
dcde836219bc add thy files
chunhan
parents:
diff changeset
   486
  \mbox{\begin{tabular}{l@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   487
  @{thm (lhs) current_procs.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) current_procs.simps(1)}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   488
  @{thm (lhs) current_procs.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) current_procs.simps(2)}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   489
  @{thm (lhs) current_procs.simps(3)} & @{text "\<equiv>"} & @{thm (rhs) current_procs.simps(3)}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   490
  @{term "current_procs (DUMMY#s)"} & @{text "\<equiv>"} & @{term "current_procs s"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   491
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   492
  \end{isabelle}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   493
    
dcde836219bc add thy files
chunhan
parents:
diff changeset
   494
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   495
  The first clause states that in the empty trace, that is initial
dcde836219bc add thy files
chunhan
parents:
diff changeset
   496
  state, the processes are given by @{text "init_processes"}. The 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   497
  events for cloning a process, respectively killing a process, update this
dcde836219bc add thy files
chunhan
parents:
diff changeset
   498
  set of processes appropriately. Otherwise the set of live
dcde836219bc add thy files
chunhan
parents:
diff changeset
   499
  processes is unchanged. We have similar functions for alive files and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   500
  IPCs, called @{term "current_files"} and @{term "current_ipcs"}. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   501
dcde836219bc add thy files
chunhan
parents:
diff changeset
   502
  We can use these function in order to formally model which events are
dcde836219bc add thy files
chunhan
parents:
diff changeset
   503
  \emph{admissible} by the operating system in each state. We show just three 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   504
  rules that give the gist of this definition. First the rule for changing 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   505
  an owner of a process:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   506
dcde836219bc add thy files
chunhan
parents:
diff changeset
   507
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   508
  @{thm[mode=Rule] osgrant6}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   509
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   510
dcde836219bc add thy files
chunhan
parents:
diff changeset
   511
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   512
  We require that the process @{text p} is alive in the state @{text s}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   513
  (first premise) and that the new owner is a user that existed in the initial state
dcde836219bc add thy files
chunhan
parents:
diff changeset
   514
  (second premise).
dcde836219bc add thy files
chunhan
parents:
diff changeset
   515
  Next the rule for creating a new file:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   516
dcde836219bc add thy files
chunhan
parents:
diff changeset
   517
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   518
  @{thm[mode=Rule] osgrant2}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   519
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   520
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   521
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   522
  It states that
dcde836219bc add thy files
chunhan
parents:
diff changeset
   523
  a file @{text f} can be created by a process @{text p} being alive in the state @{text s},
dcde836219bc add thy files
chunhan
parents:
diff changeset
   524
  the new file does not exist already in this state and there exists
dcde836219bc add thy files
chunhan
parents:
diff changeset
   525
  a parent file @{text "pf"} for the new file. The parent file is just
dcde836219bc add thy files
chunhan
parents:
diff changeset
   526
  the tail of the list representing @{text f}. % if it exists 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   527
  %(@{text "Some"}-case) or @{text None} if it does not. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   528
  Finally, the rule for cloning a process:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   529
dcde836219bc add thy files
chunhan
parents:
diff changeset
   530
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   531
  @{thm[mode=Rule] osgrant10}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   532
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   533
dcde836219bc add thy files
chunhan
parents:
diff changeset
   534
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   535
  Clearly the operating system should only allow to clone a process @{text p} if the 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   536
  process is currently alive. The cloned process will get the process
dcde836219bc add thy files
chunhan
parents:
diff changeset
   537
  ID generated by the function @{term new_proc}. This process ID should
dcde836219bc add thy files
chunhan
parents:
diff changeset
   538
  not already exist. Therefore we define @{term new_proc} as
dcde836219bc add thy files
chunhan
parents:
diff changeset
   539
dcde836219bc add thy files
chunhan
parents:
diff changeset
   540
  \begin{isabelle}\ \ \ \ \ %%%
dcde836219bc add thy files
chunhan
parents:
diff changeset
   541
  \mbox{\begin{tabular}{l@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   542
  @{term "new_proc s"} & @{text "\<equiv>"} & @{term "Max (current_procs s) + 1"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   543
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   544
  \end{isabelle}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   545
dcde836219bc add thy files
chunhan
parents:
diff changeset
   546
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   547
  namely the highest ID currently in existence increased by one. The 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   548
  admissibility rules for the other events impose similar conditions.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   549
dcde836219bc add thy files
chunhan
parents:
diff changeset
   550
  However, the admissibility check by the operating system is only one
dcde836219bc add thy files
chunhan
parents:
diff changeset
   551
  ``side'' of the constraints the RC-Model imposes. We also need to
dcde836219bc add thy files
chunhan
parents:
diff changeset
   552
  model the constraints of the access policy. For this we introduce
dcde836219bc add thy files
chunhan
parents:
diff changeset
   553
  separate @{text granted}-rules involving the sets @{text
dcde836219bc add thy files
chunhan
parents:
diff changeset
   554
  permissions} and @{text "compatible r"}: the former contains triples
dcde836219bc add thy files
chunhan
parents:
diff changeset
   555
  describing access control rules; the latter specifies for each role @{text r} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   556
  which roles are compatible with @{text r}. These sets are used in the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   557
  RC-Model when a process having a role @{text r} takes on a new role
dcde836219bc add thy files
chunhan
parents:
diff changeset
   558
  @{text r'}. For example, a login-process might belong to root;
dcde836219bc add thy files
chunhan
parents:
diff changeset
   559
  once the user logs in, however, the role of the process should change to
dcde836219bc add thy files
chunhan
parents:
diff changeset
   560
  the user's default role. The corresponding @{text "granted"}-rule is
dcde836219bc add thy files
chunhan
parents:
diff changeset
   561
  as follows
dcde836219bc add thy files
chunhan
parents:
diff changeset
   562
dcde836219bc add thy files
chunhan
parents:
diff changeset
   563
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   564
  @{thm[mode=Rule] rcgrant7}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   565
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   566
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   567
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   568
  where we check whether the process @{text p} has currently role @{text r} and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   569
  whether the RC-policy states that @{text r'} is in the role compatibility
dcde836219bc add thy files
chunhan
parents:
diff changeset
   570
  set of @{text r}. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   571
dcde836219bc add thy files
chunhan
parents:
diff changeset
   572
  The complication in the RC-Model arises from the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   573
  way how the current role of a process in a state @{text s} is
dcde836219bc add thy files
chunhan
parents:
diff changeset
   574
  calculated---represented by the predicate @{term is_current_role} in our formalisation.  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   575
  For defining this predicate we need to trace the role of a process from
dcde836219bc add thy files
chunhan
parents:
diff changeset
   576
  the initial state to the current state. In the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   577
  initial state all processes have the role given by the function
dcde836219bc add thy files
chunhan
parents:
diff changeset
   578
  @{term "init_current_role"}.  If a @{term Clone} event happens then
dcde836219bc add thy files
chunhan
parents:
diff changeset
   579
  the new process will inherit the role from the parent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   580
  process. Similarly, if a @{term ChangeRole} event happens, then 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   581
  as seen in the rule above we just change the role accordingly. More interesting 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   582
  is an @{term Execute} event in the RC-Model. For this event we have 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   583
  to check the role attached to the file to be executed. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   584
  There are a number of cases: If the role of the file is a 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   585
  \emph{normal} role, then the process will just take on this role 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   586
  when executing the file (this is like the setuid mechanism in Unix). But 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   587
  there are also four \emph{special} roles in the RC-Model:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   588
  @{term "InheritProcessRole"}, @{term "InheritUserRole"}, 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   589
   @{term "InheritParentRole"} and @{term
dcde836219bc add thy files
chunhan
parents:
diff changeset
   590
  InheritUpMixed}. For example, if a file to be executed has
dcde836219bc add thy files
chunhan
parents:
diff changeset
   591
  @{term "InheritProcessRole"} attached to it, then the process 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   592
  that executes this file keeps its role regardless of the information 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   593
  attached to the file. In this way programs can be can quarantined;
dcde836219bc add thy files
chunhan
parents:
diff changeset
   594
  @{term "InheritUserRole"} can be used for login shells
dcde836219bc add thy files
chunhan
parents:
diff changeset
   595
  to make sure they run with the user's default role.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   596
  The purpose of the other special roles is to determine the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   597
  role of a process according to the directory in which the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   598
  files are stored.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   599
dcde836219bc add thy files
chunhan
parents:
diff changeset
   600
  Having the notion of current role in place, we can define the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   601
  granted rule for the @{term Execute}-event: Suppose a process @{term
dcde836219bc add thy files
chunhan
parents:
diff changeset
   602
  p} wants to execute a file @{text f}. The RC-Model first fetches the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   603
  role @{text r} of this process (in the current state @{text s}) and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   604
  the type @{text t} of the file. It then checks if the tuple @{term
dcde836219bc add thy files
chunhan
parents:
diff changeset
   605
  "(r, t, EXECUTE)"} is part of the policy, that is in our
dcde836219bc add thy files
chunhan
parents:
diff changeset
   606
  formalisation being an element in the set @{term compatible}. The
dcde836219bc add thy files
chunhan
parents:
diff changeset
   607
  corresponding rule is as follows
dcde836219bc add thy files
chunhan
parents:
diff changeset
   608
dcde836219bc add thy files
chunhan
parents:
diff changeset
   609
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   610
  @{thm[mode=Rule] rcgrant4}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   611
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   612
dcde836219bc add thy files
chunhan
parents:
diff changeset
   613
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   614
  The next @{text granted}-rule concerns the @{term CreateFile} event.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   615
  If this event occurs, then we have two rules in our RC-Model
dcde836219bc add thy files
chunhan
parents:
diff changeset
   616
  depending on how the type of the created file is derived. If the type is inherited
dcde836219bc add thy files
chunhan
parents:
diff changeset
   617
  from the parent directory @{text pf}, then the @{term granted}-rule is as follows:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   618
dcde836219bc add thy files
chunhan
parents:
diff changeset
   619
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   620
  @{thm[mode=Rule] rcgrant1} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   621
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   622
dcde836219bc add thy files
chunhan
parents:
diff changeset
   623
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   624
  We check whether @{term pf} is the parent file (directory) of @{text f} and check
dcde836219bc add thy files
chunhan
parents:
diff changeset
   625
  whether the type of @{term pf} is @{term t}. We also need to fetch the 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   626
  the role @{text r} of the process that seeks to get permission for creating 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   627
  the file. If the default type of this role @{text r} states that the 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   628
  type of the newly created file will be inherited from the parent file
dcde836219bc add thy files
chunhan
parents:
diff changeset
   629
  type, then we only need to check that the policy states that @{text r}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   630
  has permission to write into the directory @{text pf}.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   631
dcde836219bc add thy files
chunhan
parents:
diff changeset
   632
  The situation is different if the default type of role @{text r} is
dcde836219bc add thy files
chunhan
parents:
diff changeset
   633
  some \emph{normal} type, like text-file or executable. In such cases we want
dcde836219bc add thy files
chunhan
parents:
diff changeset
   634
  that the process creates some predetermined type of files. Therefore in the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   635
  rule we have to check whether the role is allowed to create a file of that
dcde836219bc add thy files
chunhan
parents:
diff changeset
   636
  type, and also check whether the role is allowed to write any new
dcde836219bc add thy files
chunhan
parents:
diff changeset
   637
  file into the parent file (directory). The corresponding rule is
dcde836219bc add thy files
chunhan
parents:
diff changeset
   638
  as follows.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   639
dcde836219bc add thy files
chunhan
parents:
diff changeset
   640
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   641
  @{thm[mode=Rule] rcgrant1'}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   642
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   643
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   644
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   645
  Interestingly, the type-information in the RC-model is also used for
dcde836219bc add thy files
chunhan
parents:
diff changeset
   646
  processes, for example when they need to change their owner. For
dcde836219bc add thy files
chunhan
parents:
diff changeset
   647
  this we have the rule
dcde836219bc add thy files
chunhan
parents:
diff changeset
   648
dcde836219bc add thy files
chunhan
parents:
diff changeset
   649
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   650
  @{thm[mode=Rule] rcgrant_CHO}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   651
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   652
dcde836219bc add thy files
chunhan
parents:
diff changeset
   653
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   654
  whereby we have to obtain both the role and type of the process @{term p}, and then check
dcde836219bc add thy files
chunhan
parents:
diff changeset
   655
  whether the policy allows a @{term ChangeOwner}-event for that role and type.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   656
dcde836219bc add thy files
chunhan
parents:
diff changeset
   657
  Overall we have 13 rules for the admissibility check by the operating system and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   658
  14 rules for the granted check by the RC-Model.  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   659
  They are used to characterise when an event @{text e} is \emph{valid} to 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   660
  occur in a state @{text s}. This can be inductively defined as the set of valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   661
  states.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   662
dcde836219bc add thy files
chunhan
parents:
diff changeset
   663
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   664
  \begin{tabular}{@ {}c@ {}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   665
  \mbox{@{thm [mode=Axiom] valid.intros(1)}}\hspace{5mm}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   666
  \mbox{@{thm [mode=Rule] valid.intros(2)}} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   667
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   668
  \end{center}  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   669
dcde836219bc add thy files
chunhan
parents:
diff changeset
   670
  The novel notion we introduce in this paper is the @{text tainted} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   671
  relation. It characterises how a system can become infected when 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   672
  a file in the system contains, for example, a virus. We assume
dcde836219bc add thy files
chunhan
parents:
diff changeset
   673
  that the initial state contains some tainted
dcde836219bc add thy files
chunhan
parents:
diff changeset
   674
  objects (we call them @{term "seeds"}). Therefore in the initial state @{term "[]"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   675
  an object is tainted, if it is an element in @{text "seeds"}.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   676
dcde836219bc add thy files
chunhan
parents:
diff changeset
   677
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   678
  \mbox{@{thm [mode=Rule] tainted.intros(1)}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   679
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   680
dcde836219bc add thy files
chunhan
parents:
diff changeset
   681
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   682
  Let us first assume such a tainted object is a file @{text f}.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   683
  If a process reads or executes a tainted file, then this process becomes
dcde836219bc add thy files
chunhan
parents:
diff changeset
   684
  tainted (in the state where the corresponding event occurs).
dcde836219bc add thy files
chunhan
parents:
diff changeset
   685
dcde836219bc add thy files
chunhan
parents:
diff changeset
   686
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   687
  \mbox{@{thm [mode=Rule] tainted.intros(3)}}\hspace{3mm}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   688
  \mbox{@{thm [mode=Rule] tainted.intros(6)}} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   689
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   690
dcde836219bc add thy files
chunhan
parents:
diff changeset
   691
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   692
  We have a similar rule for a tainted IPC, namely
dcde836219bc add thy files
chunhan
parents:
diff changeset
   693
dcde836219bc add thy files
chunhan
parents:
diff changeset
   694
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   695
  \mbox{@{thm [mode=Rule] tainted.intros(9)}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   696
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   697
dcde836219bc add thy files
chunhan
parents:
diff changeset
   698
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   699
  which means if we receive anything from a tainted IPC, then
dcde836219bc add thy files
chunhan
parents:
diff changeset
   700
  the process becomes tainted. A process is also tainted 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   701
  when it is a produced by a @{text Clone}-event.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   702
dcde836219bc add thy files
chunhan
parents:
diff changeset
   703
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   704
  \mbox{@{thm [mode=Rule] tainted.intros(2)}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   705
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   706
dcde836219bc add thy files
chunhan
parents:
diff changeset
   707
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   708
  However, the tainting relationship must also work in the 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   709
  ``other'' direction, meaning if a process is tainted, then
dcde836219bc add thy files
chunhan
parents:
diff changeset
   710
  every file that is written or created will be tainted. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   711
  This is captured by the four rules:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   712
dcde836219bc add thy files
chunhan
parents:
diff changeset
   713
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   714
  \begin{tabular}{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   715
  \mbox{@{thm [mode=Rule] tainted.intros(4)}} \hspace{3mm}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   716
  \mbox{@{thm [mode=Rule] tainted.intros(7)}} \medskip\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   717
  \mbox{@{thm [mode=Rule] tainted.intros(5)}} \hspace{3mm}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   718
  \mbox{@{thm [mode=Rule] tainted.intros(8)}} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   719
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   720
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   721
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   722
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   723
  Finally, we have three rules that state whenever an object is tainted
dcde836219bc add thy files
chunhan
parents:
diff changeset
   724
  in a state @{text s}, then it will be still tainted in the 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   725
  next state @{term "e#s"}, provided the object is still \emph{alive}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   726
  in that state. We have such a rule for each kind of objects, for
dcde836219bc add thy files
chunhan
parents:
diff changeset
   727
  example for files the rule is:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   728
dcde836219bc add thy files
chunhan
parents:
diff changeset
   729
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   730
  \mbox{@{thm [mode=Rule] tainted_10}} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   731
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   732
dcde836219bc add thy files
chunhan
parents:
diff changeset
   733
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   734
  Similarly for alive processes and IPCs (then respectively with premises 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   735
  @{term "p \<in> current_procs (e#s)"} and @{term "i \<in> current_ipcs (e#s)"}). 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   736
  When an object present in the initial state can be tainted in 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   737
  \emph{some} state (system run), we say it is @{text "taintable"}:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   738
dcde836219bc add thy files
chunhan
parents:
diff changeset
   739
  \begin{isabelle}\ \ \ \ \ %%%
dcde836219bc add thy files
chunhan
parents:
diff changeset
   740
  \mbox{\begin{tabular}{lcl}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   741
  @{thm (lhs) taintable_def} & @{text "\<equiv>"} & @{term "init obj"} @{text "\<and>"} @{thm (rhs) taintable_def}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   742
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   743
  \end{isabelle}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   744
dcde836219bc add thy files
chunhan
parents:
diff changeset
   745
  Before we can describe our static check deciding when a file is taintable, we
dcde836219bc add thy files
chunhan
parents:
diff changeset
   746
  need to describe the notions @{term deleted} and @{term undeletable}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   747
  for objects. The former characterises whether there is an event that deletes  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   748
  these objects (files, processes or IPCs). For this we have the following 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   749
  four rules:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   750
dcde836219bc add thy files
chunhan
parents:
diff changeset
   751
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   752
  \begin{tabular}{@ {}c@ {\hspace{10mm}}c@ {}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   753
  \begin{tabular}{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   754
  @{thm [mode=Axiom] dels(1)}\\[-2mm]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   755
  @{thm [mode=Axiom] dels(2)}\\[-2mm]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   756
  @{thm [mode=Axiom] dels(3)}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   757
  \end{tabular} &
dcde836219bc add thy files
chunhan
parents:
diff changeset
   758
  @{thm [mode=Rule] dels(4)}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   759
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   760
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   761
dcde836219bc add thy files
chunhan
parents:
diff changeset
   762
dcde836219bc add thy files
chunhan
parents:
diff changeset
   763
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   764
  Note that an object cannot be deleted in the initial state @{text
dcde836219bc add thy files
chunhan
parents:
diff changeset
   765
  "[]"}.  An object is then said to be @{text "undeletable"} provided
dcde836219bc add thy files
chunhan
parents:
diff changeset
   766
  it did exist in the initial state and there does not exists a valid
dcde836219bc add thy files
chunhan
parents:
diff changeset
   767
  state in which the object is deleted:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   768
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   769
  \begin{isabelle}\ \ \ \ \ %%%
dcde836219bc add thy files
chunhan
parents:
diff changeset
   770
  \mbox{\begin{tabular}{l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   771
  @{thm (lhs) undeletable_def} & @{text "\<equiv>"} & %%@{thm (rhs) undeletable_def}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   772
  @{term "init obj \<and> \<not>(\<exists> s. (valid s \<and> deleted obj s))"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   773
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   774
  \end{isabelle}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   775
dcde836219bc add thy files
chunhan
parents:
diff changeset
   776
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   777
  The point of this definition is that our static taintable check will only be
dcde836219bc add thy files
chunhan
parents:
diff changeset
   778
  complete for undeletable objects. But these are
dcde836219bc add thy files
chunhan
parents:
diff changeset
   779
  the ones system administrators are typically interested in (for
dcde836219bc add thy files
chunhan
parents:
diff changeset
   780
  example system files).  It should be clear, however, that we cannot
dcde836219bc add thy files
chunhan
parents:
diff changeset
   781
  hope for a meaningful check by just trying out all possible
dcde836219bc add thy files
chunhan
parents:
diff changeset
   782
  valid states in our dynamic model. The reason is that there are
dcde836219bc add thy files
chunhan
parents:
diff changeset
   783
  potentially infinitely many of them and therefore the search space would be
dcde836219bc add thy files
chunhan
parents:
diff changeset
   784
  infinite.  For example staring from an
dcde836219bc add thy files
chunhan
parents:
diff changeset
   785
  initial state containing a process @{text p} and a file @{text pf}, 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   786
  we can create files @{text "f\<^isub>1"}, @{text "f\<^isub>2"}, @{text "..."} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   787
  via @{text "CreateFile"}-events. This can be pictured roughly as follows:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   788
dcde836219bc add thy files
chunhan
parents:
diff changeset
   789
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   790
  \begin{tabular}[t]{c@ {\hspace{-8mm}}c@ {\hspace{-8mm}}c@ {\hspace{-8mm}}c@ {\hspace{-8mm}}cc}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   791
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   792
  Initial state:\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   793
  @{term "{p, pf}"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   794
  \end{tabular} &
dcde836219bc add thy files
chunhan
parents:
diff changeset
   795
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   796
  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   797
  @{text "\<Longrightarrow>"}\\[2mm]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   798
  {\small@{term "CreateFile p (f\<^isub>1#pf)"}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   799
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   800
  &
dcde836219bc add thy files
chunhan
parents:
diff changeset
   801
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   802
  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   803
  @{term "{p, pf, f\<^isub>1#pf}"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   804
  \end{tabular} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   805
  &
dcde836219bc add thy files
chunhan
parents:
diff changeset
   806
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   807
  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   808
  @{text "\<Longrightarrow>"}\\[2mm]
dcde836219bc add thy files
chunhan
parents:
diff changeset
   809
  {\small@{term "CreateFile p (f\<^isub>2#f\<^isub>1#pf)"}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   810
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   811
  &
dcde836219bc add thy files
chunhan
parents:
diff changeset
   812
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   813
  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   814
  @{term "{p, pf, f\<^isub>1#pf, f\<^isub>2#f\<^isub>1#pf}"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   815
  \end{tabular} &
dcde836219bc add thy files
chunhan
parents:
diff changeset
   816
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   817
  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   818
  @{text "..."}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   819
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   820
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   821
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   822
dcde836219bc add thy files
chunhan
parents:
diff changeset
   823
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   824
  Instead, the idea of our static check is to use
dcde836219bc add thy files
chunhan
parents:
diff changeset
   825
  the policies of the RC-model for generating an answer, since they
dcde836219bc add thy files
chunhan
parents:
diff changeset
   826
  provide always a finite ``description of the system''. As we 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   827
  will see in the next section, this needs some care, however.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   828
*}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   829
dcde836219bc add thy files
chunhan
parents:
diff changeset
   830
section {* Our Static Check *}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   831
dcde836219bc add thy files
chunhan
parents:
diff changeset
   832
text {*
dcde836219bc add thy files
chunhan
parents:
diff changeset
   833
  Assume there is a tainted file in the system and suppose we face the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   834
  problem of finding out whether this file can affect other files,
dcde836219bc add thy files
chunhan
parents:
diff changeset
   835
  IPCs or processes? One idea is to work on the level of policies only, and
dcde836219bc add thy files
chunhan
parents:
diff changeset
   836
  check which operations are permitted by the role and type of this
dcde836219bc add thy files
chunhan
parents:
diff changeset
   837
  file. Then one builds the ``transitive closure'' of this information
dcde836219bc add thy files
chunhan
parents:
diff changeset
   838
  and checks for example whether the role @{text root} has become
dcde836219bc add thy files
chunhan
parents:
diff changeset
   839
  affected, in which case the whole system is compromised.  This is indeed the solution investigated
dcde836219bc add thy files
chunhan
parents:
diff changeset
   840
  in~\cite{guttman2005verifying} in the context of information flow
dcde836219bc add thy files
chunhan
parents:
diff changeset
   841
  and SELinux.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   842
dcde836219bc add thy files
chunhan
parents:
diff changeset
   843
  Unfortunately, restricting the calculations to only use policies is
dcde836219bc add thy files
chunhan
parents:
diff changeset
   844
  too simplistic for obtaining a check that is sound and complete---it
dcde836219bc add thy files
chunhan
parents:
diff changeset
   845
  over-approximates the dynamic tainted relation defined in the previous
dcde836219bc add thy files
chunhan
parents:
diff changeset
   846
  section. To see the problem consider
dcde836219bc add thy files
chunhan
parents:
diff changeset
   847
  the case where the tainted file has, say, the type @{text bin}. If
dcde836219bc add thy files
chunhan
parents:
diff changeset
   848
  the RC-policy contains a role @{text r} that can both read and write
dcde836219bc add thy files
chunhan
parents:
diff changeset
   849
  @{text bin}-files, we would conclude that all @{text bin}-files can potentially 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   850
  be tainted. That
dcde836219bc add thy files
chunhan
parents:
diff changeset
   851
  is indeed the case, \emph{if} there is a process having this role @{text
dcde836219bc add thy files
chunhan
parents:
diff changeset
   852
  r} running in the system. But if there is \emph{not}, then the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   853
  tainted file cannot ``spread''.  A similar problem arises in case there
dcde836219bc add thy files
chunhan
parents:
diff changeset
   854
  are two processes having the same role @{text r},  and this role is
dcde836219bc add thy files
chunhan
parents:
diff changeset
   855
  restricted to read files only.  Now if one of the processes is tainted, then
dcde836219bc add thy files
chunhan
parents:
diff changeset
   856
  the simple check involving only policies would incorrectly infer
dcde836219bc add thy files
chunhan
parents:
diff changeset
   857
  that all processes involving that role are tainted. But since the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   858
  policy for @{text r} is restricted to be read-only, there is in fact
dcde836219bc add thy files
chunhan
parents:
diff changeset
   859
  no danger that both processes can become tainted.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   860
dcde836219bc add thy files
chunhan
parents:
diff changeset
   861
  The main idea of our sound and complete check is to find a ``middle'' ground between
dcde836219bc add thy files
chunhan
parents:
diff changeset
   862
  the potentially infinite dynamic model and the too coarse
dcde836219bc add thy files
chunhan
parents:
diff changeset
   863
  information contained in the RC-policies. Our solution is to
dcde836219bc add thy files
chunhan
parents:
diff changeset
   864
  define a ``static'' version of the tainted relation, called @{term
dcde836219bc add thy files
chunhan
parents:
diff changeset
   865
  "tainted_s"}, that records relatively precisely the information
dcde836219bc add thy files
chunhan
parents:
diff changeset
   866
  about the initial state of the system (the one in which an object
dcde836219bc add thy files
chunhan
parents:
diff changeset
   867
  might be a @{term seed} and therefore tainted). However,
dcde836219bc add thy files
chunhan
parents:
diff changeset
   868
  we are less precise about the objects created in every subsequent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   869
  state. The result is that we can avoid the potential infinity of
dcde836219bc add thy files
chunhan
parents:
diff changeset
   870
  the dynamic model. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   871
  For the @{term tainted_s}-relation we will consider the following 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   872
  three kinds of \emph{items} recording the information we need about
dcde836219bc add thy files
chunhan
parents:
diff changeset
   873
  processes, files and IPCs, respectively:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   874
dcde836219bc add thy files
chunhan
parents:
diff changeset
   875
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   876
  \begin{tabular}{l@ {\hspace{5mm}}l}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   877
  & Recorded information:\smallskip\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   878
  Processes: & @{term "SProc (r, dr, t, u) po"}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   879
  Files: & @{term "SFile (t, a) fo"}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   880
  IPCs: & @{term "SIPC (t) io"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   881
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   882
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   883
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   884
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   885
  For a process we record its role @{text r}, its default role @{text dr} (used to determine
dcde836219bc add thy files
chunhan
parents:
diff changeset
   886
  the role when executing a file or changing the owner of a process), its type @{text t} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   887
  and its owner @{text u}. For a file we record
dcde836219bc add thy files
chunhan
parents:
diff changeset
   888
  just the type @{text t} and its @{term "source_dir"} @{text a} (we define this
dcde836219bc add thy files
chunhan
parents:
diff changeset
   889
  notion shortly). For IPCs we only record its type @{text t}. Note the superscripts
dcde836219bc add thy files
chunhan
parents:
diff changeset
   890
  @{text po}, @{text fo} and @{text io} in each item. They are optional arguments and depend on
dcde836219bc add thy files
chunhan
parents:
diff changeset
   891
  whether the corresponding object is present in the initial state or not.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   892
  If it \emph{is}, then for processes and IPCs we will record @{term "Some(id)"},
dcde836219bc add thy files
chunhan
parents:
diff changeset
   893
  where @{text id} is the natural number that uniquely identifies a process or IPC;
dcde836219bc add thy files
chunhan
parents:
diff changeset
   894
  for files we just record their path @{term "Some(f)"}. If the object is
dcde836219bc add thy files
chunhan
parents:
diff changeset
   895
  \emph{not} present in the initial state, that is newly created, then we just have
dcde836219bc add thy files
chunhan
parents:
diff changeset
   896
  @{term None} as superscript. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   897
  Let us illustrate the different superscripts with the following example
dcde836219bc add thy files
chunhan
parents:
diff changeset
   898
  where the initial state contains a process @{term p} and a file (directory)
dcde836219bc add thy files
chunhan
parents:
diff changeset
   899
  @{term pf}. Then this
dcde836219bc add thy files
chunhan
parents:
diff changeset
   900
  process creates a file via a @{term "CreateFile"}-event and after that reads
dcde836219bc add thy files
chunhan
parents:
diff changeset
   901
  the created file via a @{term Read}-event:
dcde836219bc add thy files
chunhan
parents:
diff changeset
   902
dcde836219bc add thy files
chunhan
parents:
diff changeset
   903
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   904
  \begin{tabular}[t]{ccccc}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   905
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   906
  Initial state:\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   907
  @{term "{p, pf}"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   908
  \end{tabular} &
dcde836219bc add thy files
chunhan
parents:
diff changeset
   909
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   910
  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   911
  @{text "\<Longrightarrow>"}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   912
  {\small@{term "CreateFile p (f#pf)"}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   913
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   914
  &
dcde836219bc add thy files
chunhan
parents:
diff changeset
   915
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   916
  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   917
  @{term "{p, pf, f#pf}"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   918
  \end{tabular} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   919
  &
dcde836219bc add thy files
chunhan
parents:
diff changeset
   920
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   921
  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   922
  @{text "\<Longrightarrow>"}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   923
  {\small@{term "ReadFile p (f#pf)"}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   924
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   925
  &
dcde836219bc add thy files
chunhan
parents:
diff changeset
   926
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   927
  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   928
  @{term "{p, pf, f#pf}"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   929
  \end{tabular} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   930
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   931
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   932
dcde836219bc add thy files
chunhan
parents:
diff changeset
   933
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   934
  For the two objects in the initial state our static check records
dcde836219bc add thy files
chunhan
parents:
diff changeset
   935
  the information @{term "SProc (r, dr, t, u) (Some(p))"} and @{term
dcde836219bc add thy files
chunhan
parents:
diff changeset
   936
  "SFile (t', a) (Some(pf))"} (assuming @{text "r"}, @{text t} and so
dcde836219bc add thy files
chunhan
parents:
diff changeset
   937
  on are the corresponding roles, types etc).  In both cases we have
dcde836219bc add thy files
chunhan
parents:
diff changeset
   938
  the superscript @{text "Some(...)"}  since they are objects present
dcde836219bc add thy files
chunhan
parents:
diff changeset
   939
  in the initial state. For the file @{term "f#pf"} created by the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   940
  @{term "CreateFile"}-event, we record @{term "SFile (t', a')
dcde836219bc add thy files
chunhan
parents:
diff changeset
   941
  (None)"}, since it is a newly created file.  The @{text
dcde836219bc add thy files
chunhan
parents:
diff changeset
   942
  "ReadFile"}-event does not change the set of objects, therefore no
dcde836219bc add thy files
chunhan
parents:
diff changeset
   943
  new information needs to be recorded. The problem we are avoiding
dcde836219bc add thy files
chunhan
parents:
diff changeset
   944
  with this setup of recording the precise information for the initial
dcde836219bc add thy files
chunhan
parents:
diff changeset
   945
  state is where two processes have the same role and type
dcde836219bc add thy files
chunhan
parents:
diff changeset
   946
  information, but only one is tainted in the initial state, but the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   947
  other is not. The recorded unique process IDs allows us to
dcde836219bc add thy files
chunhan
parents:
diff changeset
   948
  distinguish between both processes.  For all newly created objects,
dcde836219bc add thy files
chunhan
parents:
diff changeset
   949
  on the other hand, we do not care.  This is crucial, because
dcde836219bc add thy files
chunhan
parents:
diff changeset
   950
  otherwise exploring all possible ``reachable'' objects can lead to
dcde836219bc add thy files
chunhan
parents:
diff changeset
   951
  the potential infinity like in the dynamic model.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   952
dcde836219bc add thy files
chunhan
parents:
diff changeset
   953
  An @{term source_dir} for a file is the ``nearest'' directory that
dcde836219bc add thy files
chunhan
parents:
diff changeset
   954
  is present in the initial state and has not been deleted in a state
dcde836219bc add thy files
chunhan
parents:
diff changeset
   955
  @{text s}. Its definition is the recursive function
dcde836219bc add thy files
chunhan
parents:
diff changeset
   956
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   957
  \begin{isabelle}\ \ \ \ \ %%%
dcde836219bc add thy files
chunhan
parents:
diff changeset
   958
  \mbox{\begin{tabular}{lcl}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   959
  @{thm (lhs) source_dir.simps(1)} & @{text "\<equiv>"} \;\; &
dcde836219bc add thy files
chunhan
parents:
diff changeset
   960
  @{text "if"} @{text "\<not> deleted [] s"} @{text "then"} @{term "Some []"} @{text "else"} @{term "None"}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   961
  @{thm (lhs) source_dir.simps(2)} & @{text "\<equiv>"} & 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   962
  @{text "if"} @{term "(f#pf) \<in> init_files \<and> \<not>(deleted (File (f#pf)) s)"}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   963
  & & @{text "then"} @{term "Some (f#pf)"} @{text "else"} @{term "source_dir s pf"}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
   964
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   965
  \end{isabelle}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   966
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
   967
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
   968
  generating an optional value.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   969
  The first clause states that the anchor of the @{text
dcde836219bc add thy files
chunhan
parents:
diff changeset
   970
  root}-directory is always its own anchor unless it has been
dcde836219bc add thy files
chunhan
parents:
diff changeset
   971
  deleted. If a file is present in the initial state and not deleted
dcde836219bc add thy files
chunhan
parents:
diff changeset
   972
  in @{text s}, then it is also its own anchor, otherwise the anchor
dcde836219bc add thy files
chunhan
parents:
diff changeset
   973
  will be the anchor of the parent directory.  For example if we have
dcde836219bc add thy files
chunhan
parents:
diff changeset
   974
  a directory @{text pf} in the initial state, then its anchor is @{text "Some pf"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   975
  (assuming it is not deleted). If we create a new file in this directory,
dcde836219bc add thy files
chunhan
parents:
diff changeset
   976
  say @{term "f#pf"}, then its anchor will also be @{text "Some pf"}.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   977
  The purpose of @{term source_dir} is to determine the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   978
  role information when a file is executed, because the role of the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   979
  corresponding process, according to the RC-model, is determined by the role information of the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   980
  anchor of the file to be executed.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   981
dcde836219bc add thy files
chunhan
parents:
diff changeset
   982
  There is one last problem we have to solve before we can give the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   983
  rules of our @{term "tainted_s"}-check. Suppose an RC-policy 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   984
  includes the rule @{text "(r, foo, Write) \<in> permissions"}, that is
dcde836219bc add thy files
chunhan
parents:
diff changeset
   985
  a process of role @{text "r"} is allowed to write files of type @{text "foo"}.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   986
  If there is a tainted process with this role, we would conclude that
dcde836219bc add thy files
chunhan
parents:
diff changeset
   987
  also every file of that type can potentially become tainted. However, that
dcde836219bc add thy files
chunhan
parents:
diff changeset
   988
  is not the case if the initial state does not contain any file
dcde836219bc add thy files
chunhan
parents:
diff changeset
   989
  with type @{text foo} and the RC-policy does not allow the
dcde836219bc add thy files
chunhan
parents:
diff changeset
   990
  creation of such files, that is does not contain an access rule
dcde836219bc add thy files
chunhan
parents:
diff changeset
   991
  @{text "(r, foo, Create) \<in> permissions"}. In a sense the original
dcde836219bc add thy files
chunhan
parents:
diff changeset
   992
  @{text "(r, foo, Write)"} is ``useless'' and should not contribute
dcde836219bc add thy files
chunhan
parents:
diff changeset
   993
  to the relation characterising the objects that are tainted.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   994
  To exclude such ``useless'' access rules, we define
dcde836219bc add thy files
chunhan
parents:
diff changeset
   995
  a relation @{term "all_sobjs"} restricting our search space 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   996
  to only configurations that correspond to states in our dynamic model.
dcde836219bc add thy files
chunhan
parents:
diff changeset
   997
  We first have a rule for reachable items of the form  @{text "F(t, f)\<^bsup>Some f\<^esup>"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
   998
  where the file @{text f} with type @{text t} is present in 
dcde836219bc add thy files
chunhan
parents:
diff changeset
   999
  the initial state. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1000
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1001
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1002
  @{thm [mode=Rule] af_init'}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1003
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1004
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1005
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1006
  We have similar reachability rules for processes and IPCs that are part of the 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1007
  initial state. Next is the reachability rule in case a file is created
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1008
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1009
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1010
  @{thm [mode=Rule] af_cfd[where sd=a and sf="fo" and sp="po" and fr="dr"]}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1011
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1012
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1013
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1014
  where we require that we have a reachable parent directory, recorded
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1015
  as @{text "F(t, a)\<^bsup>fo\<^esup>"}, and also a
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1016
  process that can create the file, recorded as @{text "P(r, dr, pt,
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1017
  u)\<^bsup>po\<^esup>"}. As can be seen, we also require that we have both @{text "(r, t,
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1018
  Write)"} and \mbox{@{text "(r, t', Create)"}} in the @{text permissions} set 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1019
  for this rule to apply. If we did \emph{not} impose this requirement
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1020
  about the RC-policy, then there would be no way to create a file
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1021
  with @{term "NormalFileType t'"} according to our ``dynamic'' model. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1022
  However in case we want to create a 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1023
  file of type @{term InheritPatentType}, then we only need the access-rule
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1024
  @{text "(r, t, Write)"}: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1025
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1026
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1027
  @{thm [mode=Rule] af_cfd'[where sd=a and sf="fo" and sp="po" and fr="dr"]}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1028
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1029
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1030
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1031
  We also have reachability rules for processes executing files, and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1032
  for changing their roles and owners, for example
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1033
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1034
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1035
  @{thm [mode=Rule] ap_crole[where sp="po" and fr="dr"]}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1036
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1037
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1038
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1039
  which states that when we have a process with role @{text r}, and the role
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1040
  @{text "r'"} is in the corresponding role-compatibility set, then also
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1041
  a process with role @{text "r'"} is reachable.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1042
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1043
  The crucial difference between between the ``dynamic'' notion of validity 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1044
  and the ``static'' notion of @{term "all_sobjs"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1045
  is that there can be infinitely many valid states, but assuming the initial
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1046
  state contains only finitely many objects, then also @{term "all_sobjs"} will 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1047
  be finite. To see the difference, consider the infinite ``chain'' of events 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1048
  just cloning a process @{text "p\<^isub>0"}:
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1049
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1050
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1051
  \begin{tabular}[t]{c@ {\hspace{-2mm}}c@ {\hspace{-2mm}}c@ {\hspace{-2mm}}c@ {\hspace{-2mm}}cc}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1052
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1053
  Initial state:\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1054
  @{term "{p\<^isub>0}"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1055
  \end{tabular} &
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1056
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1057
  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1058
  @{text "\<Longrightarrow>"}\\[2mm]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1059
  {\small@{term "Clone p\<^isub>0 p\<^isub>1"}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1060
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1061
  &
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1062
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1063
  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1064
  @{term "{p\<^isub>0, p\<^isub>1}"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1065
  \end{tabular} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1066
  &
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1067
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1068
  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1069
  @{text "\<Longrightarrow>"}\\[2mm]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1070
  {\small@{term "Clone p\<^isub>0 p\<^isub>2"}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1071
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1072
  &
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1073
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1074
  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1075
  @{term "{p\<^isub>0, p\<^isub>1, p\<^isub>2}"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1076
  \end{tabular} &
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1077
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1078
  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1079
  @{text "..."}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1080
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1081
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1082
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1083
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1084
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1085
  The corresponding reachable objects are
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1086
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1087
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1088
  \begin{tabular}[t]{cccc}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1089
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1090
  @{text "{P(r, dr, t, u)\<^bsup>Some (p\<^isub>0)\<^esup>}"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1091
  \end{tabular} &
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1092
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1093
  @{text "\<Longrightarrow>"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1094
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1095
  &
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1096
  \begin{tabular}[t]{c}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1097
  @{text "{P(r, dr, t, u)\<^bsup>Some (p\<^isub>0)\<^esup>, P(r, dr, t, u)\<^bsup>None\<^esup>}"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1098
  \end{tabular} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1099
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1100
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1101
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1102
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1103
  where no further progress can be made because the information
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1104
  recorded about @{text "p\<^isub>2"}, @{text "p\<^isub>3"} and so on is just the same
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1105
  as for @{text "p\<^isub>1"}, namely @{text "P(r, dr, t, u)\<^bsup>None\<^esup>"}.  Indeed we
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1106
  can prove the lemma:
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1107
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1108
  \begin{lemma}\label{finite}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1109
  If @{text "finite init"}, then @{term "finite all_sobjs"}.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1110
  \end{lemma}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1111
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1112
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1113
  This fact of @{term all_sobjs} being finite enables us to design a
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1114
  decidable tainted-check. For this we introduce inductive rules defining the
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1115
  set @{term "tainted_s"}. Like in the ``dynamic'' version of tainted,
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1116
  if an object is element of @{text seeds}, then it is @{term "tainted_s"}.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1117
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1118
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1119
  @{thm [mode=Rule] ts_init}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1120
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1121
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1122
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1123
  The function @{text "\<lbrakk>_\<rbrakk>"} extracts the static information from an object.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1124
  For example for a process it extracts the role, default role, type and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1125
  user; for a file the type and the anchor. If a process in tainted and creates
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1126
  a file with a normal type @{text "t'"} then also the created file
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1127
  is tainted. The corresponding rule is
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1128
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1129
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1130
  @{thm [mode=Rule] ts_cfd[where sd=a and sf="fo" and sp="po" and fr="dr"]}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1131
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1132
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1133
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1134
  If a tainted process  creates a file that inherits the type of the directory,
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1135
  then the file will also be tainted:
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1136
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1137
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1138
  @{thm [mode=Rule] ts_cfd'[where sd=a and sf="fo" and sp="po" and fr="dr"]}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1139
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1140
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1141
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1142
  If a tainted process changes its role, then also with this changed role
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1143
  it will be tainted:
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1144
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1145
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1146
  @{thm [mode=Rule] ts_crole[where pt=t and sp="po" and fr="dr"]}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1147
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1148
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1149
  \noindent 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1150
  Similarly when a process changes its owner. If a file is tainted, and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1151
  a process has read-permission to that type of files, then the 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1152
  process becomes tainted. The corresponding rule is
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1153
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1154
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1155
  @{thm [mode=Rule] ts_read[where sd=a and sf="fo" and sp="po" and fr="dr"]}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1156
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1157
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1158
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1159
  If a process is tainted and it has write-permission for files of type @{text t}, 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1160
  then these files will be tainted:
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1161
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1162
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1163
  @{thm [mode=Rule] ts_write[where sd=a and sf="fo" and sp="po" and fr="dr"]}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1164
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1165
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1166
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1167
  We omit the remaining rules for executing a file, cloning a process and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1168
  rules involving IPCs, which are similar. A simple consequence of our definitions 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1169
  is that every tainted object is also reachable:
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1170
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1171
  \begin{lemma}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1172
  @{text "tainted\<^isup>s \<subseteq> reachable\<^isup>s"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1173
  \end{lemma}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1174
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1175
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1176
  which in turn means that the set of @{term "tainted_s"} items is finite by Lemma~\ref{finite}.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1177
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1178
  Returning to our original question about whether tainted objects can spread
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1179
  in the system. To answer this question, we take these tainted objects as
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1180
  seeds and calculate the set of items that are @{term "tainted_s"}. We proved this
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1181
  set is finite and can be enumerated using the rules for @{term tainted_s}.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1182
  However, this set is about items, not about whether objects are tainted or not. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1183
  Assuming an item in @{term tainted_s} arises from an object present in the initial
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1184
  state, we have recorded enough information to translate items back into objects 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1185
  via the function @{text "|_|"}:
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1186
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1187
  \begin{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1188
  \begin{tabular}{lcl}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1189
  @{text "|P(r, dr, t, u)\<^bsup>po\<^esup>|"} & @{text "\<equiv>"} & @{text po}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1190
  @{text "|F(t, a)\<^bsup>fo\<^esup>|"} & @{text "\<equiv>"} & @{text fo}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1191
  @{text "|I(t\<^bsup>\<^esup>)\<^bsup>io\<^esup>|"} & @{text "\<equiv>"} & @{text io}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1192
  \end{tabular}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1193
  \end{center}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1194
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1195
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1196
  Using this function, we can define when an object is @{term taintable_s} in terms of
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1197
  an item being @{term tainted_s}, namely
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1198
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1199
  \begin{isabelle}\ \ \ \ \ %%%
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1200
  \mbox{\begin{tabular}{lcl}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1201
  @{thm (lhs) taintable_s_def} & @{text "\<equiv>"} & @{text "\<exists>item. item \<in> tainted\<^isup>s \<and> |item| = Some obj"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1202
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1203
  \end{isabelle}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1204
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1205
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1206
  Note that @{term taintable_s} is only about objects that are present in
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1207
  the initial state, because for all other items @{text "|_|"} returns @{text None}. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1208
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1209
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1210
  With these definitions in place, we can state our theorem about the soundness of our 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1211
  static @{term taintable_s}-check for objects.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1212
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1213
  \begin{theorem}[Soundness]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1214
  @{thm [mode=IfThen] static_sound}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1215
  \end{theorem}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1216
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1217
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1218
  The proof of this theorem generates for every object that is ``flagged'' as 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1219
  @{term taintable_s} by our check, a sequence of events which shows how the
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1220
  object can become tainted in the dynamic model.  We can also state a completeness
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1221
  theorem for our @{term taintable_s}-check.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1222
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1223
  \begin{theorem}[Completeness]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1224
  @{thm [mode=IfThen] static_complete}  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1225
  \end{theorem}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1226
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1227
  \noindent
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1228
  This completeness theorem however needs to be restricted to
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1229
  undeletebale objects. The reason is that a tainted process can be
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1230
  killed by another process, and after that can be ``recreated'' by a
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1231
  cloning event from an untainted process---remember we have no control
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1232
  over which process ID a process will be assigned with.  Clearly, in
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1233
  this case the cloned process should be considered untainted, and
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1234
  indeed our dynamic tainted relation is defined in this way. The
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1235
  problem is that a static test cannot know about a process being
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1236
  killed and then recreated. Therefore the static test will not be
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1237
  able to ``detect'' the difference.  Therefore we solve this problem
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1238
  by considering only objects that are present in the initial state
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1239
  and cannot be deleted. By the latter we mean that the RC-policy
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1240
  stipulates an object cannot be deleted (for example it has been created
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1241
  by @{term root} in single-user mode, but in the everyday running
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1242
  of the system the RC-policy forbids to delete an object belonging to
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1243
  @{term root}). Like @{term "taintable_s"}, we also have a static check 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1244
  for when a file is undeletable according to an RC-policy.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1245
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1246
  This restriction to undeletable objects might be seen as a great
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1247
  weakness of our result, but in practice this seems to cover the
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1248
  interesting scenarios encountered by system administrators. They
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1249
  want to know whether a virus-infected file introduced by a user can
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1250
  affect the core system files. Our test allows the system
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1251
  administrator to find this out provided the RC-policy makes the core
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1252
  system files undeletable. We assume that this provisio is already part 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1253
  of best practice rule for running a system.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1254
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1255
  We envisage our test to be useful in two kind of situations: First, if
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1256
  there was a break-in into a system, then, clearly, the system
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1257
  administrator can find out whether the existing access policy was
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1258
  strong enough to contain the break-in, or whether core system files
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1259
  could have been affected.  In the first case, the system
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1260
  administrator can just plug the hole and forget about the break-in;
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1261
  in the other case the system administrator is wise to completely
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1262
  reinstall the system.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1263
  Second, the system administrator can proactively check whether an
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1264
  RC-policy is strong enough to withstand serious break-ins. To do so
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1265
  one has to identify the set of ``core'' system files that the policy
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1266
  should protect and mark every possible entry point for an attacker
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1267
  as tainted (they are the seeds of the @{term "tainted_s"} relation).  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1268
  Then the test will reveal
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1269
  whether the policy is strong enough or needs to be redesigned. For
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1270
  this redesign, the sequence of events our check generates should be
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1271
  informative.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1272
*}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1273
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1274
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1275
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1276
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1277
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1278
section {*Conclusion and Related Works*}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1279
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1280
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1281
text {*
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1282
  We have presented the first completely formalised dynamic model of
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1283
  the Role-Compa\-tibility Model. This is a framework, introduced by Ott
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1284
  \cite{ottrc}, in which role-based access control policies
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1285
  can be formulated and is used in practice, for example, for securing Apache 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1286
  servers. Previously, the RC-Model was presented as a
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1287
  collection of rules partly given in ``English'', partly given as formulas.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1288
  During the formalisation we uncovered an inconsistency in the
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1289
  semantics of the special role @{term "InheritParentRole"} in
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1290
  the existing works about the RC-Model \cite{ottrc,ottweb}.  By proving
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1291
  the soundness and completeness of our static @{term
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1292
  "taintable_s"}-check, we have formally related the dynamic behaviour
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1293
  of the operating system implementing access control and the static 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1294
  behaviour of the access policies of the RC-Model. The
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1295
  crucial idea in our static check is to record precisely the
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1296
  information available about the initial state (in which some resources might be
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1297
  tainted), but be less precise
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1298
  about the subsequent states. The former fact essentially gives us
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1299
  the soundness of our check, while the latter results in a finite
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1300
  search space. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1301
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1302
  The two most closely related works are by Archer et al and by Guttman et al
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1303
  \cite{Archer03,guttman2005verifying}. The first describes a
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1304
  formalisation of the dynamic behaviour of SELinux carried out in the
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1305
  theorem prover PVS. However, they cannot use their formalisation in
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1306
  order to prove any ``deep'' properties about access control rules
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1307
  \cite[Page 167]{Archer03}.  The second analyses access control
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1308
  policies in the context of information flow. Since this work 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1309
  is completely on the level of policies, it does
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1310
  not lead to a sound and complete check for files being taintable (a dynamic notion
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1311
  defined in terms of operations performed by the operating system).
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1312
  While our results concern the RC-Model, we expect that they 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1313
  equally apply to the access control model of SELinux. In fact,
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1314
  we expect that the formalisation is simpler for SELinux, since 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1315
  its rules governing roles are much simpler than in the RC-Model.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1316
  The definition of our admissibility rules can be copied verbatim for SELinux;
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1317
  we would need to modify our granted rules and slightly adapt our
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1318
  static check. We leave this as future work.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1319
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1320
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1321
  Our formalisation is carried out in the Isabelle/HOL theorem prover.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1322
  It uses Paulson's inductive method for
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1323
  reasoning about sequences of events \cite{Paulson98}.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1324
  We have approximately 1000 lines of code for definitions and 6000 lines of
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1325
  code for proofs. Our formalisation is available from the
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1326
  Mercurial repository at \url{http://www.dcs.kcl.ac.uk/staff/urbanc/cgi-bin/repos.cgi/rc/}.\\[-12mm]
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1327
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1328
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1329
%In a word, what the manager need is that given the
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1330
%initial state of the system, a policy checker that make sure the under the policy
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1331
%he made, this important file cannot: 1. be deleted; 2. be tainted. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1332
%Formally speaking, this policy-checker @{text "PC"} (a function that given the 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1333
%initial state of system, a policy and an object, it tells whether this object 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1334
%will be fully protected or not) should satisfy this criteria: 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1335
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1336
%  @{text "(PC init policy obj) \<and> (exists init obj) \<longrightarrow> \<not> taintable obj"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1337
%If the @{text obj} exists in the initial-state, and @{text "PC"} justify the safety
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1338
%of this object under @{text "policy"}, this object should not be @{text taintable}.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1339
%We call this criteria the \emph{completeness} of @{text "PC"}. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1340
%And there is the \emph{soundness} criteria of @{text "PC"} too, otherwise a "NO-to-ALL"
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1341
%answer always satisfy the \emph{completeness}. \emph{soundness} formally is:
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1342
%  @{text "PC init policy obj \<longrightarrow> taintable obj"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1343
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1344
%This policy-checker should satisfy other properties:
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1345
% 1. fully statical, that means this policy-checker should not rely on the system 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1346
%running information, like new created files/process, and most importantly the 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1347
%trace of system running. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1348
% 2. decidable, that means this policy-checker should always terminate.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1349
3
chunhan
parents: 1
diff changeset
  1350
*}
1
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1351
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1352
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1353
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1354
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1355
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1356
(*<*)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1357
end
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1358
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1359
end
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1360
(*>*)
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1361
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1362
(*
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1363
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1364
  Central to RC-Model is the roles and types. We start with do formalisation on
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1365
  types first.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1366
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1367
  \begin{isabelle}\ \ \ \ \ %%%
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1368
  \mbox{
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1369
  \begin{tabular}{r@ {\hspace{1mm}}l@ {\hspace{2mm}}l}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1370
  @{text t_client} & @{text "="} & @{text "Christian"} \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1371
                   & @{text "|"} & @{text "Chunhan"}   \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1372
                   & @{text "|"} & @{text " ... "}     \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1373
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1374
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1375
  \mbox{
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1376
  \begin{tabular}{r@ {\hspace{1mm}}l@ {\hspace{2mm}}l@ {\hspace{2mm}}l}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1377
  @{text t_normal_file_type}  & @{text "="} & @{text "WebServerLog_file"} &             \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1378
                              & @{text "|"} & @{text "WebData_file"} & @{text t_client} \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1379
                              & @{text "|"} & @{text "CGI_file"}     & @{text t_client} \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1380
                              & @{text "|"} & @{text "Private_file"} & @{text t_client} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1381
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1382
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1383
  \mbox{
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1384
  \begin{tabular} {r@ {\hspace{1mm}}l@ {\hspace{5mm}}l}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1385
  @{text t_rc_file_type} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1386
     & @{text "="} & @{term "InheritParent_file_type"}   \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1387
     & @{text "|"} & @{term "NormalFile_type t_normal_file_type"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1388
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1389
  \end{isabelle}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1390
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1391
  @{term "type_of_file"} function calculates the current type for the files:
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1392
  \begin{isabelle}\ \ \ \ \ %%%
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1393
  \mbox{\begin{tabular}{lcl}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1394
  @{thm (lhs) type_of_file.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) type_of_file.simps(1)}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1395
  @{thm (lhs) type_of_file.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) type_of_file.simps(2)}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1396
  @{term "type_of_file (DUMMY#s)"} & @{text "\<equiv>"} & @{term "type_of_file s"}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1397
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1398
  \end{isabelle}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1399
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1400
  Note that this @{term "type_of_file"} is not the function @{term "etype_of_file"} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1401
  that we call in the grant check of RC-Model, @{term "rc_grant"}. The reason is
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1402
  that file's type can be set to a special type of @{term "InheritParent_file_type"}, 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1403
  means that the ``efficient'' type of this file is the efficient type of its directory. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1404
  \mbox{\begin{tabular}{lcl}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1405
  @{thm (lhs) etype_aux.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) etype_aux.simps(1)}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1406
  @{thm (lhs) etype_aux.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) etype_aux.simps(2)}\smallskip\\ 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1407
  @{thm (lhs) etype_of_file_def} & @{text "\<equiv>"} & @{thm (rhs) etype_of_file_def}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1408
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1409
  Here @{term etype_aux} is an auxiliary function which do recursion
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1410
  on the pathname of files. By the way, in our proofs, we do proved
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1411
  that functions like @{term "etype_of_file"} will always return
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1412
  ``normal'' values.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1413
  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1414
 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1415
  We have similar observation functions calculating the current type for processes
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1416
  and IPCs too, only diffence here is that there is no ``effcient'' type here for 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1417
  processes and IPCs, all types that calculated by @{term "type_of_process"} and 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1418
  @{term "type_of_ipc"} are alrealdy efficient types. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1419
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1420
*}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1421
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1422
text {*
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1423
  \begin{isabelle}\ \ \ \ \ %%%
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1424
  \mbox{
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1425
  \begin{tabular}{r@ {\hspace{1mm}}l@ {\hspace{2mm}}l@ {\hspace{2mm}}l}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1426
  @{text t_normal_role}  & @{text "="} & @{text "WebServer"} &                  \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1427
                         & @{text "|"} & @{text "WS_client"} & @{text t_client} \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1428
                         & @{text "|"} & @{text "UpLoader"}  & @{text t_client} \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1429
                         & @{text "|"} & @{text "CGI   "}    & @{text t_client} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1430
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1431
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1432
  \mbox{
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1433
  \begin{tabular} {r@ {\hspace{1mm}}l@ {\hspace{2mm}}l@ {\hspace{5mm}}l}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1434
  @{text t_role} 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1435
     & @{text "="} & @{term "InheritParentRole"}  & ``for file's initial/forced role, 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1436
                                                         meaning using parent directory's 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1437
                                                         role instead'' \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1438
     & @{text "|"} & @{term "UseForcedRole"}      & ``for file's initial role'' \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1439
     & @{text "|"} & @{term "InheritProcessRole"} & ``using process' current role''\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1440
     & @{text "|"} & @{term "InheritUserRole"}    & ``using owner's default role''\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1441
     & @{text "|"} &   ...                        & \\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1442
     & @{text "|"} & @{term "NormalRole t_normal_role"} & ``user-defined 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1443
                                                           policy roles" 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1444
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1445
  \end{isabelle}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1446
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1447
  @{text "t_normal roles"} are normally user-defined roles in the
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1448
  policy, where @{text "WebServer"} is the role who plays for the
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1449
  server, while @{text "WS_client"} is the role server plays for
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1450
  certain client, so is for @{text "UpLoader"} role. @{text "CGI"} is
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1451
  the role that client's programme scripts play. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1452
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1453
  @{term "currentrole"} function calculates the current role of process, here we 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1454
  only show 3 cases of its definition, it responses to @{term "ChangeOwner"}, 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1455
  @{term "ChangeRole"} events too. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1456
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1457
  \begin{isabelle}\ \ \ \ \ %%%
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1458
  \mbox{\begin{tabular}{lcl}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1459
  @{thm (lhs) currentrole.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) currentrole.simps(1)}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1460
  @{thm (lhs) currentrole.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) currentrole.simps(2)}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1461
  @{thm (lhs) currentrole.simps(3)} & @{text "\<equiv>"} & @{thm (rhs) currentrole.simps(3)}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1462
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1463
  \end{isabelle}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1464
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1465
  If the event trace is @{term "[]"}, means the
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1466
  system state currently is the initial state, then @{term "init_currentrole"} will
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1467
  do. @{term "Execute p f"} event is one complex case, when this event happens, process
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1468
  @{term p}'s role will be changed according to the efficient initial role of the 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1469
  executable file @{term f}, here ``efficient'' is like the file's type too. 
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1470
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1471
  \begin{isabelle}\ \ \ \ \ %%%
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1472
  \mbox{\begin{tabular}{lcl}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1473
  @{thm (lhs) initialrole.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) initialrole.simps(1)}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1474
  @{thm (lhs) initialrole.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) initialrole.simps(2)}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1475
  @{thm (lhs) initialrole.simps(3)} & @{text "\<equiv>"} & @{thm (rhs) initialrole.simps(3)}\medskip\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1476
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1477
  @{thm (lhs) erole_functor.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) erole_functor.simps(1)}\\
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1478
  @{thm (lhs) erole_functor.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) erole_functor.simps(2)}  
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1479
  \end{tabular}}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1480
  \end{isabelle}
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1481
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1482
  If this efficient initial role is normal role, then RC-Model assigns
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1483
  this role to the process after execution finished. Otherwise if this
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1484
  efficient initial role is using special value @{term
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1485
  "UseForcedRole"}, then the new role for the process is then
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1486
  determinated by the efficient forced role of the executable file
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1487
  @{term "forcedrole"}.  When new process is created, this process'
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1488
  role is assigned to its creator's role.
dcde836219bc add thy files
chunhan
parents:
diff changeset
  1489
*)