src/HOL/Unix/Unix.thy
author wenzelm
Tue, 30 Aug 2011 17:36:12 +0200
changeset 45472 04f64e602fa6
parent 45107 b73b7832b384
child 45761 22f665a2e91c
permissions -rw-r--r--
tuned document;
wenzelm@10966
     1
(*  Title:      HOL/Unix/Unix.thy
wenzelm@10966
     2
    Author:     Markus Wenzel, TU Muenchen
wenzelm@10966
     3
*)
wenzelm@10966
     4
wenzelm@10966
     5
header {* Unix file-systems \label{sec:unix-file-system} *}
wenzelm@10966
     6
wenzelm@17455
     7
theory Unix
wenzelm@41661
     8
imports
wenzelm@45107
     9
  Nested_Environment
wenzelm@41661
    10
  "~~/src/HOL/Library/List_Prefix"
wenzelm@17455
    11
begin
wenzelm@10966
    12
wenzelm@10966
    13
text {*
wenzelm@10966
    14
  We give a simple mathematical model of the basic structures
wenzelm@10966
    15
  underlying the Unix file-system, together with a few fundamental
wenzelm@10966
    16
  operations that could be imagined to be performed internally by the
wenzelm@10966
    17
  Unix kernel.  This forms the basis for the set of Unix system-calls
wenzelm@10966
    18
  to be introduced later (see \secref{sec:unix-trans}), which are the
wenzelm@10966
    19
  actual interface offered to processes running in user-space.
wenzelm@10966
    20
wenzelm@10966
    21
  \medskip Basically, any Unix file is either a \emph{plain file} or a
wenzelm@10966
    22
  \emph{directory}, consisting of some \emph{content} plus
wenzelm@10966
    23
  \emph{attributes}.  The content of a plain file is plain text.  The
wenzelm@10966
    24
  content of a directory is a mapping from names to further
wenzelm@10966
    25
  files.\footnote{In fact, this is the only way that names get
wenzelm@10966
    26
  associated with files.  In Unix files do \emph{not} have a name in
wenzelm@10966
    27
  itself.  Even more, any number of names may be associated with the
wenzelm@10966
    28
  very same file due to \emph{hard links} (although this is excluded
wenzelm@10966
    29
  from our model).}  Attributes include information to control various
wenzelm@10966
    30
  ways to access the file (read, write etc.).
wenzelm@10966
    31
wenzelm@10966
    32
  Our model will be quite liberal in omitting excessive detail that is
wenzelm@10966
    33
  easily seen to be ``irrelevant'' for the aspects of Unix
wenzelm@10966
    34
  file-systems to be discussed here.  First of all, we ignore
wenzelm@10966
    35
  character and block special files, pipes, sockets, hard links,
wenzelm@10966
    36
  symbolic links, and mount points.
wenzelm@10966
    37
*}
wenzelm@10966
    38
wenzelm@10966
    39
wenzelm@10966
    40
subsection {* Names *}
wenzelm@10966
    41
wenzelm@10966
    42
text {*
wenzelm@10966
    43
  User ids and file name components shall be represented by natural
wenzelm@10966
    44
  numbers (without loss of generality).  We do not bother about
wenzelm@10966
    45
  encoding of actual names (e.g.\ strings), nor a mapping between user
wenzelm@10966
    46
  names and user ids as would be present in a reality.
wenzelm@10966
    47
*}
wenzelm@10966
    48
wenzelm@41827
    49
type_synonym uid = nat
wenzelm@41827
    50
type_synonym name = nat
wenzelm@41827
    51
type_synonym path = "name list"
wenzelm@10966
    52
wenzelm@10966
    53
wenzelm@10966
    54
subsection {* Attributes *}
wenzelm@10966
    55
wenzelm@10966
    56
text {*
wenzelm@10966
    57
  Unix file attributes mainly consist of \emph{owner} information and
wenzelm@10966
    58
  a number of \emph{permission} bits which control access for
wenzelm@10966
    59
  ``user'', ``group'', and ``others'' (see the Unix man pages @{text
wenzelm@10966
    60
  "chmod(2)"} and @{text "stat(2)"} for more details).
wenzelm@10966
    61
wenzelm@10966
    62
  \medskip Our model of file permissions only considers the ``others''
wenzelm@10966
    63
  part.  The ``user'' field may be omitted without loss of overall
wenzelm@10966
    64
  generality, since the owner is usually able to change it anyway by
wenzelm@10966
    65
  performing @{text chmod}.\footnote{The inclined Unix expert may try
wenzelm@10966
    66
  to figure out some exotic arrangements of a real-world Unix
wenzelm@10966
    67
  file-system such that the owner of a file is unable to apply the
wenzelm@10966
    68
  @{text chmod} system call.}  We omit ``group'' permissions as a
wenzelm@10966
    69
  genuine simplification as we just do not intend to discuss a model
wenzelm@10966
    70
  of multiple groups and group membership, but pretend that everyone
wenzelm@10966
    71
  is member of a single global group.\footnote{A general HOL model of
wenzelm@10966
    72
  user group structures and related issues is given in
wenzelm@10966
    73
  \cite{Naraschewski:2001}.}
wenzelm@10966
    74
*}
wenzelm@10966
    75
wenzelm@10966
    76
datatype perm =
wenzelm@10966
    77
    Readable
wenzelm@10966
    78
  | Writable
wenzelm@10966
    79
  | Executable    -- "(ignored)"
wenzelm@10966
    80
wenzelm@41827
    81
type_synonym perms = "perm set"
wenzelm@10966
    82
wenzelm@10966
    83
record att =
wenzelm@10966
    84
  owner :: uid
wenzelm@10966
    85
  others :: perms
wenzelm@10966
    86
wenzelm@10966
    87
text {*
wenzelm@10966
    88
  For plain files @{term Readable} and @{term Writable} specify read
wenzelm@10966
    89
  and write access to the actual content, i.e.\ the string of text
wenzelm@10966
    90
  stored here.  For directories @{term Readable} determines if the set
wenzelm@10966
    91
  of entry names may be accessed, and @{term Writable} controls the
wenzelm@10966
    92
  ability to create or delete any entries (both plain files or
wenzelm@10966
    93
  sub-directories).
wenzelm@10966
    94
wenzelm@10966
    95
  As another simplification, we ignore the @{term Executable}
wenzelm@10966
    96
  permission altogether.  In reality it would indicate executable
wenzelm@10966
    97
  plain files (also known as ``binaries''), or control actual lookup
wenzelm@10966
    98
  of directory entries (recall that mere directory browsing is
wenzelm@10966
    99
  controlled via @{term Readable}).  Note that the latter means that
wenzelm@10966
   100
  in order to perform any file-system operation whatsoever, all
wenzelm@10966
   101
  directories encountered on the path would have to grant @{term
wenzelm@10966
   102
  Executable}.  We ignore this detail and pretend that all directories
wenzelm@10966
   103
  give @{term Executable} permission to anybody.
wenzelm@10966
   104
*}
wenzelm@10966
   105
wenzelm@10966
   106
wenzelm@10966
   107
subsection {* Files *}
wenzelm@10966
   108
wenzelm@10966
   109
text {*
wenzelm@10966
   110
  In order to model the general tree structure of a Unix file-system
wenzelm@10966
   111
  we use the arbitrarily branching datatype @{typ "('a, 'b, 'c) env"}
wenzelm@10966
   112
  from the standard library of Isabelle/HOL
wenzelm@13380
   113
  \cite{Bauer-et-al:2002:HOL-Library}.  This type provides
wenzelm@10966
   114
  constructors @{term Val} and @{term Env} as follows:
wenzelm@10966
   115
wenzelm@10966
   116
  \medskip
wenzelm@10966
   117
  {\def\isastyleminor{\isastyle}
wenzelm@10966
   118
  \begin{tabular}{l}
wenzelm@10966
   119
  @{term [source] "Val :: 'a \<Rightarrow> ('a, 'b, 'c) env"} \\
wenzelm@10966
   120
  @{term [source] "Env :: 'b \<Rightarrow> ('c \<Rightarrow> ('a, 'b, 'c) env option) \<Rightarrow> ('a, 'b, 'c) env"} \\
wenzelm@10966
   121
  \end{tabular}}
wenzelm@10966
   122
  \medskip
wenzelm@10966
   123
wenzelm@10966
   124
  Here the parameter @{typ 'a} refers to plain values occurring at
wenzelm@10966
   125
  leaf positions, parameter @{typ 'b} to information kept with inner
wenzelm@10966
   126
  branch nodes, and parameter @{typ 'c} to the branching type of the
wenzelm@10966
   127
  tree structure.  For our purpose we use the type instance with @{typ
wenzelm@10966
   128
  "att \<times> string"} (representing plain files), @{typ att} (for
wenzelm@10966
   129
  attributes of directory nodes), and @{typ name} (for the index type
wenzelm@10966
   130
  of directory nodes).
wenzelm@10966
   131
*}
wenzelm@10966
   132
wenzelm@41827
   133
type_synonym "file" = "(att \<times> string, att, name) env"
wenzelm@10966
   134
wenzelm@10966
   135
text {*
wenzelm@10966
   136
  \medskip The HOL library also provides @{term lookup} and @{term
wenzelm@10966
   137
  update} operations for general tree structures with the subsequent
wenzelm@10966
   138
  primitive recursive characterizations.
wenzelm@10966
   139
wenzelm@10966
   140
  \medskip
wenzelm@10966
   141
  {\def\isastyleminor{\isastyle}
wenzelm@10966
   142
  \begin{tabular}{l}
wenzelm@10966
   143
  @{term [source] "lookup :: ('a, 'b, 'c) env \<Rightarrow> 'c list \<Rightarrow> ('a, 'b, 'c) env option"} \\
wenzelm@10966
   144
  @{term [source]
wenzelm@10966
   145
  "update :: 'c list \<Rightarrow> ('a, 'b, 'c) env option \<Rightarrow> ('a, 'b, 'c) env \<Rightarrow> ('a, 'b, 'c) env"} \\
wenzelm@10966
   146
  \end{tabular}}
wenzelm@10966
   147
wenzelm@10966
   148
  @{thm [display, indent = 2, eta_contract = false] lookup_eq [no_vars]}
wenzelm@10966
   149
  @{thm [display, indent = 2, eta_contract = false] update_eq [no_vars]}
wenzelm@10966
   150
wenzelm@10966
   151
  Several further properties of these operations are proven in
wenzelm@13380
   152
  \cite{Bauer-et-al:2002:HOL-Library}.  These will be routinely used
wenzelm@10966
   153
  later on without further notice.
wenzelm@10966
   154
berghofe@17146
   155
  \bigskip Apparently, the elements of type @{typ "file"} contain an
wenzelm@10966
   156
  @{typ att} component in either case.  We now define a few auxiliary
wenzelm@10966
   157
  operations to manipulate this field uniformly, following the
wenzelm@10966
   158
  conventions for record types in Isabelle/HOL
wenzelm@10966
   159
  \cite{Nipkow-et-al:2000:HOL}.
wenzelm@10966
   160
*}
wenzelm@10966
   161
wenzelm@19086
   162
definition
wenzelm@19086
   163
  "attributes file =
wenzelm@10966
   164
    (case file of
wenzelm@10966
   165
      Val (att, text) \<Rightarrow> att
wenzelm@10966
   166
    | Env att dir \<Rightarrow> att)"
wenzelm@10966
   167
wenzelm@21404
   168
definition
wenzelm@21001
   169
  "map_attributes f file =
wenzelm@10966
   170
    (case file of
wenzelm@21001
   171
      Val (att, text) \<Rightarrow> Val (f att, text)
wenzelm@21001
   172
    | Env att dir \<Rightarrow> Env (f att) dir)"
wenzelm@10966
   173
wenzelm@10966
   174
lemma [simp]: "attributes (Val (att, text)) = att"
wenzelm@10966
   175
  by (simp add: attributes_def)
wenzelm@10966
   176
wenzelm@10966
   177
lemma [simp]: "attributes (Env att dir) = att"
wenzelm@10966
   178
  by (simp add: attributes_def)
wenzelm@10966
   179
wenzelm@21001
   180
lemma [simp]: "attributes (map_attributes f file) = f (attributes file)"
wenzelm@21001
   181
  by (cases "file") (simp_all add: attributes_def map_attributes_def
wenzelm@10966
   182
    split_tupled_all)
wenzelm@10966
   183
wenzelm@21001
   184
lemma [simp]: "map_attributes f (Val (att, text)) = Val (f att, text)"
wenzelm@21001
   185
  by (simp add: map_attributes_def)
wenzelm@10966
   186
wenzelm@21001
   187
lemma [simp]: "map_attributes f (Env att dir) = Env (f att) dir"
wenzelm@21001
   188
  by (simp add: map_attributes_def)
wenzelm@10966
   189
wenzelm@10966
   190
wenzelm@10966
   191
subsection {* Initial file-systems *}
wenzelm@10966
   192
wenzelm@10966
   193
text {*
wenzelm@10966
   194
  Given a set of \emph{known users} a file-system shall be initialized
wenzelm@10966
   195
  by providing an empty home directory for each user, with read-only
wenzelm@10966
   196
  access for everyone else.  (Note that we may directly use the user
wenzelm@10966
   197
  id as home directory name, since both types have been identified.)
wenzelm@10966
   198
  Certainly, the very root directory is owned by the super user (who
wenzelm@10966
   199
  has user id 0).
wenzelm@10966
   200
*}
wenzelm@10966
   201
wenzelm@19086
   202
definition
wenzelm@19086
   203
  "init users =
wenzelm@10966
   204
    Env \<lparr>owner = 0, others = {Readable}\<rparr>
wenzelm@10966
   205
      (\<lambda>u. if u \<in> users then Some (Env \<lparr>owner = u, others = {Readable}\<rparr> empty)
wenzelm@10966
   206
        else None)"
wenzelm@10966
   207
wenzelm@10966
   208
wenzelm@10966
   209
subsection {* Accessing file-systems *}
wenzelm@10966
   210
wenzelm@10966
   211
text {*
wenzelm@10966
   212
  The main internal file-system operation is access of a file by a
wenzelm@10966
   213
  user, requesting a certain set of permissions.  The resulting @{typ
wenzelm@10966
   214
  "file option"} indicates if a file had been present at the
wenzelm@10966
   215
  corresponding @{typ path} and if access was granted according to the
wenzelm@10966
   216
  permissions recorded within the file-system.
wenzelm@10966
   217
wenzelm@10966
   218
  Note that by the rules of Unix file-system security (e.g.\
wenzelm@10966
   219
  \cite{Tanenbaum:1992}) both the super-user and owner may always
wenzelm@10966
   220
  access a file unconditionally (in our simplified model).
wenzelm@10966
   221
*}
wenzelm@10966
   222
wenzelm@19086
   223
definition
wenzelm@20321
   224
  "access root path uid perms =
wenzelm@10966
   225
    (case lookup root path of
wenzelm@10966
   226
      None \<Rightarrow> None
wenzelm@10966
   227
    | Some file \<Rightarrow>
wenzelm@10966
   228
        if uid = 0
wenzelm@10966
   229
          \<or> uid = owner (attributes file)
wenzelm@10966
   230
          \<or> perms \<subseteq> others (attributes file)
wenzelm@10966
   231
        then Some file
wenzelm@10966
   232
        else None)"
wenzelm@10966
   233
wenzelm@10966
   234
text {*
wenzelm@10966
   235
  \medskip Successful access to a certain file is the main
wenzelm@10966
   236
  prerequisite for system-calls to be applicable (cf.\
wenzelm@10966
   237
  \secref{sec:unix-trans}).  Any modification of the file-system is
wenzelm@10966
   238
  then performed using the basic @{term update} operation.
wenzelm@10966
   239
wenzelm@10966
   240
  \medskip We see that @{term access} is just a wrapper for the basic
wenzelm@10966
   241
  @{term lookup} function, with additional checking of
wenzelm@10966
   242
  attributes. Subsequently we establish a few auxiliary facts that
wenzelm@10966
   243
  stem from the primitive @{term lookup} used within @{term access}.
wenzelm@10966
   244
*}
wenzelm@10966
   245
wenzelm@10966
   246
lemma access_empty_lookup: "access root path uid {} = lookup root path"
wenzelm@10966
   247
  by (simp add: access_def split: option.splits)
wenzelm@10966
   248
wenzelm@10966
   249
lemma access_some_lookup:
wenzelm@10966
   250
  "access root path uid perms = Some file \<Longrightarrow>
wenzelm@10966
   251
    lookup root path = Some file"
wenzelm@10966
   252
  by (simp add: access_def split: option.splits if_splits)
wenzelm@10966
   253
wenzelm@18372
   254
lemma access_update_other:
wenzelm@18372
   255
  assumes parallel: "path' \<parallel> path"
wenzelm@18372
   256
  shows "access (update path' opt root) path uid perms = access root path uid perms"
wenzelm@10966
   257
proof -
wenzelm@18372
   258
  from parallel obtain y z xs ys zs where
wenzelm@11072
   259
      "y \<noteq> z" and "path' = xs @ y # ys" and "path = xs @ z # zs"
wenzelm@10966
   260
    by (blast dest: parallel_decomp)
wenzelm@18372
   261
  then have "lookup (update path' opt root) path = lookup root path"
wenzelm@10966
   262
    by (blast intro: lookup_update_other)
wenzelm@18372
   263
  then show ?thesis by (simp only: access_def)
wenzelm@10966
   264
qed
wenzelm@10966
   265
wenzelm@10966
   266
wenzelm@10966
   267
section {* File-system transitions \label{sec:unix-trans} *}
wenzelm@10966
   268
wenzelm@10966
   269
subsection {* Unix system calls \label{sec:unix-syscall} *}
wenzelm@10966
   270
wenzelm@10966
   271
text {*
wenzelm@10966
   272
  According to established operating system design (cf.\
wenzelm@10966
   273
  \cite{Tanenbaum:1992}) user space processes may only initiate system
wenzelm@10966
   274
  operations by a fixed set of \emph{system-calls}.  This enables the
wenzelm@10966
   275
  kernel to enforce certain security policies in the first
wenzelm@10966
   276
  place.\footnote{Incidently, this is the very same principle employed
wenzelm@10966
   277
  by any ``LCF-style'' theorem proving system according to Milner's
wenzelm@10966
   278
  principle of ``correctness by construction'', such as Isabelle/HOL
wenzelm@10966
   279
  itself.}
wenzelm@10966
   280
wenzelm@10966
   281
  \medskip In our model of Unix we give a fixed datatype @{text
wenzelm@10966
   282
  operation} for the syntax of system-calls, together with an
wenzelm@10966
   283
  inductive definition of file-system state transitions of the form
wenzelm@10966
   284
  @{text "root \<midarrow>x\<rightarrow> root'"} for the operational semantics.
wenzelm@10966
   285
*}
wenzelm@10966
   286
wenzelm@10966
   287
datatype operation =
wenzelm@10966
   288
    Read uid string path
wenzelm@10966
   289
  | Write uid string path
wenzelm@10966
   290
  | Chmod uid perms path
wenzelm@10966
   291
  | Creat uid perms path
wenzelm@10966
   292
  | Unlink uid path
wenzelm@10966
   293
  | Mkdir uid perms path
wenzelm@10966
   294
  | Rmdir uid path
wenzelm@10966
   295
  | Readdir uid "name set" path
wenzelm@10966
   296
wenzelm@10966
   297
text {*
wenzelm@10966
   298
  The @{typ uid} field of an operation corresponds to the
wenzelm@10966
   299
  \emph{effective user id} of the underlying process, although our
wenzelm@10966
   300
  model never mentions processes explicitly.  The other parameters are
wenzelm@10966
   301
  provided as arguments by the caller; the @{term path} one is common
wenzelm@10966
   302
  to all kinds of system-calls.
wenzelm@10966
   303
*}
wenzelm@10966
   304
wenzelm@45472
   305
primrec uid_of :: "operation \<Rightarrow> uid"
wenzelm@25974
   306
where
wenzelm@25974
   307
    "uid_of (Read uid text path) = uid"
wenzelm@25974
   308
  | "uid_of (Write uid text path) = uid"
wenzelm@25974
   309
  | "uid_of (Chmod uid perms path) = uid"
wenzelm@25974
   310
  | "uid_of (Creat uid perms path) = uid"
wenzelm@25974
   311
  | "uid_of (Unlink uid path) = uid"
wenzelm@25974
   312
  | "uid_of (Mkdir uid path perms) = uid"
wenzelm@25974
   313
  | "uid_of (Rmdir uid path) = uid"
wenzelm@25974
   314
  | "uid_of (Readdir uid names path) = uid"
wenzelm@25974
   315
wenzelm@45472
   316
primrec path_of :: "operation \<Rightarrow> path"
wenzelm@25974
   317
where
wenzelm@25974
   318
    "path_of (Read uid text path) = path"
wenzelm@25974
   319
  | "path_of (Write uid text path) = path"
wenzelm@25974
   320
  | "path_of (Chmod uid perms path) = path"
wenzelm@25974
   321
  | "path_of (Creat uid perms path) = path"
wenzelm@25974
   322
  | "path_of (Unlink uid path) = path"
wenzelm@25974
   323
  | "path_of (Mkdir uid perms path) = path"
wenzelm@25974
   324
  | "path_of (Rmdir uid path) = path"
wenzelm@25974
   325
  | "path_of (Readdir uid names path) = path"
wenzelm@10966
   326
wenzelm@10966
   327
text {*
wenzelm@10966
   328
  \medskip Note that we have omitted explicit @{text Open} and @{text
wenzelm@10966
   329
  Close} operations, pretending that @{term Read} and @{term Write}
wenzelm@10966
   330
  would already take care of this behind the scenes.  Thus we have
wenzelm@10966
   331
  basically treated actual sequences of real system-calls @{text
wenzelm@13380
   332
  "open"}--@{text read}/@{text write}--@{text close} as atomic.
wenzelm@10966
   333
wenzelm@10966
   334
  In principle, this could make big a difference in a model with
wenzelm@10966
   335
  explicit concurrent processes.  On the other hand, even on a real
wenzelm@13380
   336
  Unix system the exact scheduling of concurrent @{text "open"} and
wenzelm@10966
   337
  @{text close} operations does \emph{not} directly affect the success
wenzelm@10966
   338
  of corresponding @{text read} or @{text write}.  Unix allows several
wenzelm@10966
   339
  processes to have files opened at the same time, even for writing!
wenzelm@10966
   340
  Certainly, the result from reading the contents later may be hard to
wenzelm@10966
   341
  predict, but the system-calls involved here will succeed in any
wenzelm@10966
   342
  case.
wenzelm@10966
   343
wenzelm@10966
   344
  \bigskip The operational semantics of system calls is now specified
wenzelm@10966
   345
  via transitions of the file-system configuration.  This is expressed
wenzelm@10966
   346
  as an inductive relation (although there is no actual recursion
wenzelm@10966
   347
  involved here).
wenzelm@10966
   348
*}
wenzelm@10966
   349
wenzelm@45472
   350
inductive transition :: "file \<Rightarrow> operation \<Rightarrow> file \<Rightarrow> bool"
wenzelm@45472
   351
  ("_ \<midarrow>_\<rightarrow> _" [90, 1000, 90] 100)
wenzelm@21372
   352
where
wenzelm@10966
   353
  read:
wenzelm@10966
   354
    "access root path uid {Readable} = Some (Val (att, text)) \<Longrightarrow>
wenzelm@21372
   355
      root \<midarrow>(Read uid text path)\<rightarrow> root" |
wenzelm@36520
   356
  "write":
wenzelm@10966
   357
    "access root path uid {Writable} = Some (Val (att, text')) \<Longrightarrow>
wenzelm@21372
   358
      root \<midarrow>(Write uid text path)\<rightarrow> update path (Some (Val (att, text))) root" |
wenzelm@10966
   359
wenzelm@10966
   360
  chmod:
wenzelm@10966
   361
    "access root path uid {} = Some file \<Longrightarrow>
wenzelm@10966
   362
      uid = 0 \<or> uid = owner (attributes file) \<Longrightarrow>
wenzelm@10966
   363
      root \<midarrow>(Chmod uid perms path)\<rightarrow> update path
schirmer@25706
   364
        (Some (map_attributes (others_update (\<lambda>_. perms)) file)) root" |
wenzelm@10966
   365
wenzelm@10966
   366
  creat:
wenzelm@10966
   367
    "path = parent_path @ [name] \<Longrightarrow>
wenzelm@10966
   368
      access root parent_path uid {Writable} = Some (Env att parent) \<Longrightarrow>
wenzelm@10966
   369
      access root path uid {} = None \<Longrightarrow>
wenzelm@10966
   370
      root \<midarrow>(Creat uid perms path)\<rightarrow> update path
wenzelm@21372
   371
        (Some (Val (\<lparr>owner = uid, others = perms\<rparr>, []))) root" |
wenzelm@10966
   372
  unlink:
wenzelm@10966
   373
    "path = parent_path @ [name] \<Longrightarrow>
wenzelm@10966
   374
      access root parent_path uid {Writable} = Some (Env att parent) \<Longrightarrow>
wenzelm@10966
   375
      access root path uid {} = Some (Val plain) \<Longrightarrow>
wenzelm@21372
   376
      root \<midarrow>(Unlink uid path)\<rightarrow> update path None root" |
wenzelm@10966
   377
wenzelm@10966
   378
  mkdir:
wenzelm@10966
   379
    "path = parent_path @ [name] \<Longrightarrow>
wenzelm@10966
   380
      access root parent_path uid {Writable} = Some (Env att parent) \<Longrightarrow>
wenzelm@10966
   381
      access root path uid {} = None \<Longrightarrow>
wenzelm@10966
   382
      root \<midarrow>(Mkdir uid perms path)\<rightarrow> update path
wenzelm@21372
   383
        (Some (Env \<lparr>owner = uid, others = perms\<rparr> empty)) root" |
wenzelm@10966
   384
  rmdir:
wenzelm@10966
   385
    "path = parent_path @ [name] \<Longrightarrow>
wenzelm@10966
   386
      access root parent_path uid {Writable} = Some (Env att parent) \<Longrightarrow>
wenzelm@10966
   387
      access root path uid {} = Some (Env att' empty) \<Longrightarrow>
wenzelm@21372
   388
      root \<midarrow>(Rmdir uid path)\<rightarrow> update path None root" |
wenzelm@10966
   389
wenzelm@10966
   390
  readdir:
wenzelm@10966
   391
    "access root path uid {Readable} = Some (Env att dir) \<Longrightarrow>
wenzelm@10966
   392
      names = dom dir \<Longrightarrow>
wenzelm@10966
   393
      root \<midarrow>(Readdir uid names path)\<rightarrow> root"
wenzelm@10966
   394
wenzelm@10966
   395
text {*
wenzelm@10966
   396
  \medskip Certainly, the above specification is central to the whole
wenzelm@10966
   397
  formal development.  Any of the results to be established later on
wenzelm@10966
   398
  are only meaningful to the outside world if this transition system
wenzelm@10966
   399
  provides an adequate model of real Unix systems.  This kind of
wenzelm@10966
   400
  ``reality-check'' of a formal model is the well-known problem of
wenzelm@10966
   401
  \emph{validation}.
wenzelm@10966
   402
wenzelm@10966
   403
  If in doubt, one may consider to compare our definition with the
wenzelm@10966
   404
  informal specifications given the corresponding Unix man pages, or
wenzelm@10966
   405
  even peek at an actual implementation such as
wenzelm@10966
   406
  \cite{Torvalds-et-al:Linux}.  Another common way to gain confidence
wenzelm@10966
   407
  into the formal model is to run simple simulations (see
wenzelm@10966
   408
  \secref{sec:unix-examples}), and check the results with that of
wenzelm@10966
   409
  experiments performed on a real Unix system.
wenzelm@10966
   410
*}
wenzelm@10966
   411
wenzelm@10966
   412
wenzelm@10966
   413
subsection {* Basic properties of single transitions \label{sec:unix-single-trans} *}
wenzelm@10966
   414
wenzelm@10966
   415
text {*
wenzelm@10966
   416
  The transition system @{text "root \<midarrow>x\<rightarrow> root'"} defined above
wenzelm@10966
   417
  determines a unique result @{term root'} from given @{term root} and
wenzelm@10966
   418
  @{term x} (this is holds rather trivially, since there is even only
wenzelm@10966
   419
  one clause for each operation).  This uniqueness statement will
wenzelm@10966
   420
  simplify our subsequent development to some extent, since we only
wenzelm@10966
   421
  have to reason about a partial function rather than a general
wenzelm@10966
   422
  relation.
wenzelm@10966
   423
*}
wenzelm@10966
   424
wenzelm@18372
   425
theorem transition_uniq:
wenzelm@18372
   426
  assumes root': "root \<midarrow>x\<rightarrow> root'"
wenzelm@18372
   427
    and root'': "root \<midarrow>x\<rightarrow> root''"
wenzelm@18372
   428
  shows "root' = root''"
wenzelm@18372
   429
  using root''
wenzelm@18372
   430
proof cases
wenzelm@18372
   431
  case read
wenzelm@18372
   432
  with root' show ?thesis by cases auto
wenzelm@18372
   433
next
wenzelm@36520
   434
  case "write"
wenzelm@18372
   435
  with root' show ?thesis by cases auto
wenzelm@18372
   436
next
wenzelm@18372
   437
  case chmod
wenzelm@18372
   438
  with root' show ?thesis by cases auto
wenzelm@18372
   439
next
wenzelm@18372
   440
  case creat
wenzelm@18372
   441
  with root' show ?thesis by cases auto
wenzelm@18372
   442
next
wenzelm@18372
   443
  case unlink
wenzelm@18372
   444
  with root' show ?thesis by cases auto
wenzelm@18372
   445
next
wenzelm@18372
   446
  case mkdir
wenzelm@18372
   447
  with root' show ?thesis by cases auto
wenzelm@18372
   448
next
wenzelm@18372
   449
  case rmdir
wenzelm@18372
   450
  with root' show ?thesis by cases auto
wenzelm@18372
   451
next
wenzelm@18372
   452
  case readdir
wenzelm@18372
   453
  with root' show ?thesis by cases fastsimp+
wenzelm@10966
   454
qed
wenzelm@10966
   455
wenzelm@10966
   456
text {*
wenzelm@10966
   457
  \medskip Apparently, file-system transitions are \emph{type-safe} in
wenzelm@10966
   458
  the sense that the result of transforming an actual directory yields
wenzelm@10966
   459
  again a directory.
wenzelm@10966
   460
*}
wenzelm@10966
   461
wenzelm@10966
   462
theorem transition_type_safe:
wenzelm@18372
   463
  assumes tr: "root \<midarrow>x\<rightarrow> root'"
wenzelm@18372
   464
    and inv: "\<exists>att dir. root = Env att dir"
wenzelm@18372
   465
  shows "\<exists>att dir. root' = Env att dir"
wenzelm@18372
   466
proof (cases "path_of x")
wenzelm@18372
   467
  case Nil
wenzelm@18372
   468
  with tr inv show ?thesis
wenzelm@18372
   469
    by cases (auto simp add: access_def split: if_splits)
wenzelm@18372
   470
next
wenzelm@18372
   471
  case Cons
wenzelm@18372
   472
  from tr obtain opt where
wenzelm@18372
   473
      "root' = root \<or> root' = update (path_of x) opt root"
wenzelm@18372
   474
    by cases auto
wenzelm@18372
   475
  with inv Cons show ?thesis
wenzelm@18372
   476
    by (auto simp add: update_eq split: list.splits)
wenzelm@10966
   477
qed
wenzelm@10966
   478
wenzelm@10966
   479
text {*
wenzelm@10966
   480
  The previous result may be seen as the most basic invariant on the
wenzelm@10966
   481
  file-system state that is enforced by any proper kernel
wenzelm@10966
   482
  implementation.  So user processes --- being bound to the
wenzelm@10966
   483
  system-call interface --- may never mess up a file-system such that
wenzelm@10966
   484
  the root becomes a plain file instead of a directory, which would be
wenzelm@10966
   485
  a strange situation indeed.
wenzelm@10966
   486
*}
wenzelm@10966
   487
wenzelm@10966
   488
wenzelm@10966
   489
subsection {* Iterated transitions *}
wenzelm@10966
   490
wenzelm@10966
   491
text {*
wenzelm@10966
   492
  Iterated system transitions via finite sequences of system
wenzelm@10966
   493
  operations are modeled inductively as follows.  In a sense, this
wenzelm@10966
   494
  relation describes the cumulative effect of the sequence of
wenzelm@10966
   495
  system-calls issued by a number of running processes over a finite
wenzelm@10966
   496
  amount of time.
wenzelm@10966
   497
*}
wenzelm@10966
   498
wenzelm@45472
   499
inductive transitions :: "file \<Rightarrow> operation list \<Rightarrow> file \<Rightarrow> bool"
wenzelm@45472
   500
  ("_ =_\<Rightarrow> _" [90, 1000, 90] 100)
wenzelm@21372
   501
where
wenzelm@10966
   502
    nil: "root =[]\<Rightarrow> root"
wenzelm@21372
   503
  | cons: "root \<midarrow>x\<rightarrow> root' \<Longrightarrow> root' =xs\<Rightarrow> root'' \<Longrightarrow> root =(x # xs)\<Rightarrow> root''"
wenzelm@10966
   504
wenzelm@10966
   505
text {*
wenzelm@10966
   506
  \medskip We establish a few basic facts relating iterated
wenzelm@10966
   507
  transitions with single ones, according to the recursive structure
wenzelm@10966
   508
  of lists.
wenzelm@10966
   509
*}
wenzelm@10966
   510
wenzelm@10966
   511
lemma transitions_nil_eq: "root =[]\<Rightarrow> root' = (root = root')"
wenzelm@10966
   512
proof
wenzelm@10966
   513
  assume "root =[]\<Rightarrow> root'"
wenzelm@18372
   514
  then show "root = root'" by cases simp_all
wenzelm@10966
   515
next
wenzelm@10966
   516
  assume "root = root'"
wenzelm@18372
   517
  then show "root =[]\<Rightarrow> root'" by (simp only: transitions.nil)
wenzelm@10966
   518
qed
wenzelm@10966
   519
wenzelm@10966
   520
lemma transitions_cons_eq:
wenzelm@10966
   521
  "root =(x # xs)\<Rightarrow> root'' = (\<exists>root'. root \<midarrow>x\<rightarrow> root' \<and> root' =xs\<Rightarrow> root'')"
wenzelm@10966
   522
proof
wenzelm@10966
   523
  assume "root =(x # xs)\<Rightarrow> root''"
wenzelm@18372
   524
  then show "\<exists>root'. root \<midarrow>x\<rightarrow> root' \<and> root' =xs\<Rightarrow> root''"
wenzelm@10966
   525
    by cases auto
wenzelm@10966
   526
next
wenzelm@10966
   527
  assume "\<exists>root'. root \<midarrow>x\<rightarrow> root' \<and> root' =xs\<Rightarrow> root''"
wenzelm@18372
   528
  then show "root =(x # xs)\<Rightarrow> root''"
wenzelm@10966
   529
    by (blast intro: transitions.cons)
wenzelm@10966
   530
qed
wenzelm@10966
   531
wenzelm@10966
   532
text {*
wenzelm@10966
   533
  The next two rules show how to ``destruct'' known transition
wenzelm@10966
   534
  sequences.  Note that the second one actually relies on the
wenzelm@10966
   535
  uniqueness property of the basic transition system (see
wenzelm@10966
   536
  \secref{sec:unix-single-trans}).
wenzelm@10966
   537
*}
wenzelm@10966
   538
wenzelm@10966
   539
lemma transitions_nilD: "root =[]\<Rightarrow> root' \<Longrightarrow> root' = root"
wenzelm@10966
   540
  by (simp add: transitions_nil_eq)
wenzelm@10966
   541
wenzelm@10966
   542
lemma transitions_consD:
wenzelm@18372
   543
  assumes list: "root =(x # xs)\<Rightarrow> root''"
wenzelm@18372
   544
    and hd: "root \<midarrow>x\<rightarrow> root'"
wenzelm@18372
   545
  shows "root' =xs\<Rightarrow> root''"
wenzelm@10966
   546
proof -
wenzelm@18372
   547
  from list obtain r' where r': "root \<midarrow>x\<rightarrow> r'" and root'': "r' =xs\<Rightarrow> root''"
wenzelm@10966
   548
    by cases simp_all
wenzelm@18372
   549
  from r' hd have "r' = root'" by (rule transition_uniq)
wenzelm@10966
   550
  with root'' show "root' =xs\<Rightarrow> root''" by simp
wenzelm@10966
   551
qed
wenzelm@10966
   552
wenzelm@10966
   553
text {*
wenzelm@10966
   554
  \medskip The following fact shows how an invariant @{term Q} of
wenzelm@10966
   555
  single transitions with property @{term P} may be transferred to
wenzelm@10966
   556
  iterated transitions.  The proof is rather obvious by rule induction
wenzelm@10966
   557
  over the definition of @{term "root =xs\<Rightarrow> root'"}.
wenzelm@10966
   558
*}
wenzelm@10966
   559
wenzelm@10966
   560
lemma transitions_invariant:
wenzelm@18372
   561
  assumes r: "\<And>r x r'. r \<midarrow>x\<rightarrow> r' \<Longrightarrow> Q r \<Longrightarrow>  P x \<Longrightarrow> Q r'"
wenzelm@18372
   562
    and trans: "root =xs\<Rightarrow> root'"
wenzelm@18372
   563
  shows "Q root \<Longrightarrow> \<forall>x \<in> set xs. P x \<Longrightarrow> Q root'"
wenzelm@18372
   564
  using trans
wenzelm@18372
   565
proof induct
wenzelm@18372
   566
  case nil
wenzelm@23463
   567
  show ?case by (rule nil.prems)
wenzelm@18372
   568
next
wenzelm@21372
   569
  case (cons root x root' xs root'')
wenzelm@18372
   570
  note P = `\<forall>x \<in> set (x # xs). P x`
wenzelm@18372
   571
  then have "P x" by simp
wenzelm@18372
   572
  with `root \<midarrow>x\<rightarrow> root'` and `Q root` have Q': "Q root'" by (rule r)
wenzelm@18372
   573
  from P have "\<forall>x \<in> set xs. P x" by simp
wenzelm@18372
   574
  with Q' show "Q root''" by (rule cons.hyps)
wenzelm@10966
   575
qed
wenzelm@10966
   576
wenzelm@10966
   577
text {*
wenzelm@10966
   578
  As an example of applying the previous result, we transfer the basic
wenzelm@10966
   579
  type-safety property (see \secref{sec:unix-single-trans}) from
wenzelm@10966
   580
  single transitions to iterated ones, which is a rather obvious
wenzelm@10966
   581
  result indeed.
wenzelm@10966
   582
*}
wenzelm@10966
   583
wenzelm@10966
   584
theorem transitions_type_safe:
wenzelm@16670
   585
  assumes "root =xs\<Rightarrow> root'"
wenzelm@16670
   586
    and "\<exists>att dir. root = Env att dir"
wenzelm@16670
   587
  shows "\<exists>att dir. root' = Env att dir"
wenzelm@23394
   588
  using transition_type_safe and assms
wenzelm@16670
   589
proof (rule transitions_invariant)
wenzelm@16670
   590
  show "\<forall>x \<in> set xs. True" by blast
wenzelm@10966
   591
qed
wenzelm@10966
   592
wenzelm@10966
   593
wenzelm@10966
   594
section {* Executable sequences *}
wenzelm@10966
   595
wenzelm@10966
   596
text {*
wenzelm@17455
   597
  An inductively defined relation such as the one of @{text "root \<midarrow>x\<rightarrow>
wenzelm@17455
   598
  root'"} (see \secref{sec:unix-syscall}) has two main aspects.  First
wenzelm@17455
   599
  of all, the resulting system admits a certain set of transition
wenzelm@17455
   600
  rules (introductions) as given in the specification.  Furthermore,
wenzelm@17455
   601
  there is an explicit least-fixed-point construction involved, which
wenzelm@17455
   602
  results in induction (and case-analysis) rules to eliminate known
wenzelm@17455
   603
  transitions in an exhaustive manner.
wenzelm@10966
   604
wenzelm@10966
   605
  \medskip Subsequently, we explore our transition system in an
wenzelm@10966
   606
  experimental style, mainly using the introduction rules with basic
wenzelm@10966
   607
  algebraic properties of the underlying structures.  The technique
wenzelm@10966
   608
  closely resembles that of Prolog combined with functional evaluation
wenzelm@10966
   609
  in a very simple manner.
wenzelm@10966
   610
wenzelm@10966
   611
  Just as the ``closed-world assumption'' is left implicit in Prolog,
wenzelm@10966
   612
  we do not refer to induction over the whole transition system here.
wenzelm@10966
   613
  So this is still purely positive reasoning about possible
wenzelm@10966
   614
  executions; exhaustive reasoning will be employed only later on (see
wenzelm@10966
   615
  \secref{sec:unix-bogosity}), when we shall demonstrate that certain
wenzelm@10966
   616
  behavior is \emph{not} possible.
wenzelm@10966
   617
*}
wenzelm@10966
   618
wenzelm@10966
   619
wenzelm@10966
   620
subsection {* Possible transitions *}
wenzelm@10966
   621
wenzelm@10966
   622
text {*
wenzelm@10966
   623
  Rather obviously, a list of system operations can be executed within
wenzelm@10966
   624
  a certain state if there is a result state reached by an iterated
wenzelm@10966
   625
  transition.
wenzelm@10966
   626
*}
wenzelm@10966
   627
wenzelm@45472
   628
definition "can_exec root xs = (\<exists>root'. root =xs\<Rightarrow> root')"
wenzelm@10966
   629
wenzelm@10966
   630
lemma can_exec_nil: "can_exec root []"
wenzelm@18730
   631
  unfolding can_exec_def by (blast intro: transitions.intros)
wenzelm@10966
   632
wenzelm@10966
   633
lemma can_exec_cons:
wenzelm@10966
   634
    "root \<midarrow>x\<rightarrow> root' \<Longrightarrow> can_exec root' xs \<Longrightarrow> can_exec root (x # xs)"
wenzelm@18730
   635
  unfolding can_exec_def by (blast intro: transitions.intros)
wenzelm@10966
   636
wenzelm@10966
   637
text {*
wenzelm@10966
   638
  \medskip In case that we already know that a certain sequence can be
wenzelm@10966
   639
  executed we may destruct it backwardly into individual transitions.
wenzelm@10966
   640
*}
wenzelm@10966
   641
wenzelm@18372
   642
lemma can_exec_snocD: "can_exec root (xs @ [y])
wenzelm@10966
   643
    \<Longrightarrow> \<exists>root' root''. root =xs\<Rightarrow> root' \<and> root' \<midarrow>y\<rightarrow> root''"
wenzelm@20503
   644
proof (induct xs arbitrary: root)
wenzelm@18372
   645
  case Nil
wenzelm@18372
   646
  then show ?case
wenzelm@18372
   647
    by (simp add: can_exec_def transitions_nil_eq transitions_cons_eq)
wenzelm@18372
   648
next
wenzelm@18372
   649
  case (Cons x xs)
wenzelm@18372
   650
  from `can_exec root ((x # xs) @ [y])` obtain r root'' where
wenzelm@18372
   651
      x: "root \<midarrow>x\<rightarrow> r" and
wenzelm@18372
   652
      xs_y: "r =(xs @ [y])\<Rightarrow> root''"
wenzelm@18372
   653
    by (auto simp add: can_exec_def transitions_nil_eq transitions_cons_eq)
wenzelm@18372
   654
  from xs_y Cons.hyps obtain root' r' where xs: "r =xs\<Rightarrow> root'" and y: "root' \<midarrow>y\<rightarrow> r'"
wenzelm@18730
   655
    unfolding can_exec_def by blast
wenzelm@18372
   656
  from x xs have "root =(x # xs)\<Rightarrow> root'"
wenzelm@18372
   657
    by (rule transitions.cons)
wenzelm@18372
   658
  with y show ?case by blast
wenzelm@10966
   659
qed
wenzelm@10966
   660
wenzelm@10966
   661
wenzelm@10966
   662
subsection {* Example executions \label{sec:unix-examples} *}
wenzelm@10966
   663
wenzelm@10966
   664
text {*
wenzelm@10966
   665
  We are now ready to perform a few experiments within our formal
wenzelm@10966
   666
  model of Unix system-calls.  The common technique is to alternate
wenzelm@10966
   667
  introduction rules of the transition system (see
wenzelm@10966
   668
  \secref{sec:unix-trans}), and steps to solve any emerging side
wenzelm@10966
   669
  conditions using algebraic properties of the underlying file-system
wenzelm@10966
   670
  structures (see \secref{sec:unix-file-system}).
wenzelm@10966
   671
*}
wenzelm@10966
   672
wenzelm@10966
   673
lemmas eval = access_def init_def
wenzelm@10966
   674
wenzelm@10966
   675
theorem "u \<in> users \<Longrightarrow> can_exec (init users)
wenzelm@10966
   676
    [Mkdir u perms [u, name]]"
wenzelm@10966
   677
  apply (rule can_exec_cons)
wenzelm@10966
   678
    -- {* back-chain @{term can_exec} (of @{term [source] Cons}) *}
wenzelm@10966
   679
  apply (rule mkdir)
wenzelm@10966
   680
    -- {* back-chain @{term Mkdir} *}
wenzelm@10966
   681
  apply (force simp add: eval)+
wenzelm@10966
   682
    -- {* solve preconditions of @{term Mkdir} *}
wenzelm@10966
   683
  apply (simp add: eval)
wenzelm@10966
   684
    -- {* peek at resulting dir (optional) *}
wenzelm@10966
   685
  txt {* @{subgoals [display]} *}
wenzelm@10966
   686
  apply (rule can_exec_nil)
wenzelm@10966
   687
    -- {* back-chain @{term can_exec} (of @{term [source] Nil}) *}
wenzelm@10966
   688
  done
wenzelm@10966
   689
wenzelm@10966
   690
text {*
wenzelm@10966
   691
  By inspecting the result shown just before the final step above, we
wenzelm@10966
   692
  may gain confidence that our specification of Unix system-calls
wenzelm@10966
   693
  actually makes sense.  Further common errors are usually exhibited
wenzelm@10966
   694
  when preconditions of transition rules fail unexpectedly.
wenzelm@10966
   695
wenzelm@10966
   696
  \medskip Here are a few further experiments, using the same
wenzelm@10966
   697
  techniques as before.
wenzelm@10966
   698
*}
wenzelm@10966
   699
wenzelm@10966
   700
theorem "u \<in> users \<Longrightarrow> can_exec (init users)
wenzelm@10966
   701
   [Creat u perms [u, name],
wenzelm@10966
   702
    Unlink u [u, name]]"
wenzelm@10966
   703
  apply (rule can_exec_cons)
wenzelm@10966
   704
  apply (rule creat)
wenzelm@10966
   705
  apply (force simp add: eval)+
wenzelm@10966
   706
  apply (simp add: eval)
wenzelm@10966
   707
  apply (rule can_exec_cons)
wenzelm@10966
   708
  apply (rule unlink)
wenzelm@10966
   709
  apply (force simp add: eval)+
wenzelm@10966
   710
  apply (simp add: eval)
wenzelm@10966
   711
  txt {* peek at result: @{subgoals [display]} *}
wenzelm@10966
   712
  apply (rule can_exec_nil)
wenzelm@10966
   713
  done
wenzelm@10966
   714
wenzelm@17455
   715
theorem "u \<in> users \<Longrightarrow> Writable \<in> perms\<^isub>1 \<Longrightarrow>
wenzelm@17455
   716
  Readable \<in> perms\<^isub>2 \<Longrightarrow> name\<^isub>3 \<noteq> name\<^isub>4 \<Longrightarrow>
wenzelm@10966
   717
  can_exec (init users)
wenzelm@17455
   718
   [Mkdir u perms\<^isub>1 [u, name\<^isub>1],
wenzelm@17455
   719
    Mkdir u' perms\<^isub>2 [u, name\<^isub>1, name\<^isub>2],
wenzelm@17455
   720
    Creat u' perms\<^isub>3 [u, name\<^isub>1, name\<^isub>2, name\<^isub>3],
wenzelm@17455
   721
    Creat u' perms\<^isub>3 [u, name\<^isub>1, name\<^isub>2, name\<^isub>4],
wenzelm@17455
   722
    Readdir u {name\<^isub>3, name\<^isub>4} [u, name\<^isub>1, name\<^isub>2]]"
wenzelm@10966
   723
  apply (rule can_exec_cons, rule transition.intros,
wenzelm@10966
   724
    (force simp add: eval)+, (simp add: eval)?)+
wenzelm@10966
   725
  txt {* peek at result: @{subgoals [display]} *}
wenzelm@10966
   726
  apply (rule can_exec_nil)
wenzelm@10966
   727
  done
wenzelm@10966
   728
wenzelm@17455
   729
theorem "u \<in> users \<Longrightarrow> Writable \<in> perms\<^isub>1 \<Longrightarrow> Readable \<in> perms\<^isub>3 \<Longrightarrow>
wenzelm@10966
   730
  can_exec (init users)
wenzelm@17455
   731
   [Mkdir u perms\<^isub>1 [u, name\<^isub>1],
wenzelm@17455
   732
    Mkdir u' perms\<^isub>2 [u, name\<^isub>1, name\<^isub>2],
wenzelm@17455
   733
    Creat u' perms\<^isub>3 [u, name\<^isub>1, name\<^isub>2, name\<^isub>3],
wenzelm@17455
   734
    Write u' ''foo'' [u, name\<^isub>1, name\<^isub>2, name\<^isub>3],
wenzelm@17455
   735
    Read u ''foo'' [u, name\<^isub>1, name\<^isub>2, name\<^isub>3]]"
wenzelm@10966
   736
  apply (rule can_exec_cons, rule transition.intros,
wenzelm@10966
   737
    (force simp add: eval)+, (simp add: eval)?)+
wenzelm@10966
   738
  txt {* peek at result: @{subgoals [display]} *}
wenzelm@10966
   739
  apply (rule can_exec_nil)
wenzelm@10966
   740
  done
wenzelm@10966
   741
wenzelm@10966
   742
wenzelm@10966
   743
section {* Odd effects --- treated formally \label{sec:unix-bogosity} *}
wenzelm@10966
   744
wenzelm@10966
   745
text {*
wenzelm@10966
   746
  We are now ready to give a completely formal treatment of the
wenzelm@10966
   747
  slightly odd situation discussed in the introduction (see
wenzelm@10966
   748
  \secref{sec:unix-intro}): the file-system can easily reach a state
wenzelm@10966
   749
  where a user is unable to remove his very own directory, because it
wenzelm@10966
   750
  is still populated by items placed there by another user in an
wenzelm@10966
   751
  uncouth manner.
wenzelm@10966
   752
*}
wenzelm@10966
   753
wenzelm@10966
   754
subsection {* The general procedure \label{sec:unix-inv-procedure} *}
wenzelm@10966
   755
wenzelm@10966
   756
text {*
wenzelm@10966
   757
  The following theorem expresses the general procedure we are
wenzelm@10966
   758
  following to achieve the main result.
wenzelm@10966
   759
*}
wenzelm@10966
   760
wenzelm@10966
   761
theorem general_procedure:
wenzelm@18372
   762
  assumes cannot_y: "\<And>r r'. Q r \<Longrightarrow> r \<midarrow>y\<rightarrow> r' \<Longrightarrow> False"
wenzelm@18372
   763
    and init_inv: "\<And>root. init users =bs\<Rightarrow> root \<Longrightarrow> Q root"
wenzelm@18372
   764
    and preserve_inv: "\<And>r x r'. r \<midarrow>x\<rightarrow> r' \<Longrightarrow> Q r \<Longrightarrow> P x \<Longrightarrow> Q r'"
wenzelm@18372
   765
    and init_result: "init users =bs\<Rightarrow> root"
wenzelm@18372
   766
  shows "\<not> (\<exists>xs. (\<forall>x \<in> set xs. P x) \<and> can_exec root (xs @ [y]))"
wenzelm@10966
   767
proof -
wenzelm@10966
   768
  {
wenzelm@10966
   769
    fix xs
wenzelm@10966
   770
    assume Ps: "\<forall>x \<in> set xs. P x"
wenzelm@10966
   771
    assume can_exec: "can_exec root (xs @ [y])"
wenzelm@10966
   772
    then obtain root' root'' where
wenzelm@10966
   773
        xs: "root =xs\<Rightarrow> root'" and y: "root' \<midarrow>y\<rightarrow> root''"
wenzelm@10966
   774
      by (blast dest: can_exec_snocD)
wenzelm@10966
   775
    from init_result have "Q root" by (rule init_inv)
wenzelm@10966
   776
    from preserve_inv xs this Ps have "Q root'"
wenzelm@10966
   777
      by (rule transitions_invariant)
wenzelm@10966
   778
    from this y have False by (rule cannot_y)
wenzelm@10966
   779
  }
wenzelm@18372
   780
  then show ?thesis by blast
wenzelm@10966
   781
qed
wenzelm@10966
   782
wenzelm@10966
   783
text {*
wenzelm@10966
   784
  Here @{prop "P x"} refers to the restriction on file-system
wenzelm@10966
   785
  operations that are admitted after having reached the critical
wenzelm@10966
   786
  configuration; according to the problem specification this will
wenzelm@17455
   787
  become @{prop "uid_of x = user\<^isub>1"} later on.  Furthermore,
wenzelm@17455
   788
  @{term y} refers to the operations we claim to be impossible to
wenzelm@17455
   789
  perform afterwards, we will take @{term Rmdir} later.  Moreover
wenzelm@17455
   790
  @{term Q} is a suitable (auxiliary) invariant over the file-system;
wenzelm@17455
   791
  choosing @{term Q} adequately is very important to make the proof
wenzelm@17455
   792
  work (see \secref{sec:unix-inv-lemmas}).
wenzelm@10966
   793
*}
wenzelm@10966
   794
wenzelm@10966
   795
wenzelm@12079
   796
subsection {* The particular situation *}
wenzelm@10966
   797
wenzelm@10966
   798
text {*
wenzelm@10966
   799
  We introduce a few global declarations and axioms to describe our
wenzelm@12079
   800
  particular situation considered here.  Thus we avoid excessive use
wenzelm@12079
   801
  of local parameters in the subsequent development.
wenzelm@10966
   802
*}
wenzelm@10966
   803
wenzelm@12079
   804
locale situation =
wenzelm@12079
   805
  fixes users :: "uid set"
wenzelm@17455
   806
    and user\<^isub>1 :: uid
wenzelm@17455
   807
    and user\<^isub>2 :: uid
wenzelm@17455
   808
    and name\<^isub>1 :: name
wenzelm@17455
   809
    and name\<^isub>2 :: name
wenzelm@17455
   810
    and name\<^isub>3 :: name
wenzelm@17455
   811
    and perms\<^isub>1 :: perms
wenzelm@17455
   812
    and perms\<^isub>2 :: perms
wenzelm@17455
   813
  assumes user\<^isub>1_known: "user\<^isub>1 \<in> users"
wenzelm@17455
   814
    and user\<^isub>1_not_root: "user\<^isub>1 \<noteq> 0"
wenzelm@17455
   815
    and users_neq: "user\<^isub>1 \<noteq> user\<^isub>2"
wenzelm@17455
   816
    and perms\<^isub>1_writable: "Writable \<in> perms\<^isub>1"
wenzelm@17455
   817
    and perms\<^isub>2_not_writable: "Writable \<notin> perms\<^isub>2"
wenzelm@17455
   818
  notes facts = user\<^isub>1_known user\<^isub>1_not_root users_neq
wenzelm@17455
   819
    perms\<^isub>1_writable perms\<^isub>2_not_writable
wenzelm@21029
   820
begin
wenzelm@10966
   821
wenzelm@21029
   822
definition
wenzelm@19086
   823
  "bogus =
wenzelm@17455
   824
     [Mkdir user\<^isub>1 perms\<^isub>1 [user\<^isub>1, name\<^isub>1],
wenzelm@17455
   825
      Mkdir user\<^isub>2 perms\<^isub>2 [user\<^isub>1, name\<^isub>1, name\<^isub>2],
wenzelm@17455
   826
      Creat user\<^isub>2 perms\<^isub>2 [user\<^isub>1, name\<^isub>1, name\<^isub>2, name\<^isub>3]]"
wenzelm@21404
   827
definition
wenzelm@19086
   828
  "bogus_path = [user\<^isub>1, name\<^isub>1, name\<^isub>2]"
wenzelm@10966
   829
wenzelm@10966
   830
text {*
wenzelm@12079
   831
  The @{term bogus} operations are the ones that lead into the uncouth
wenzelm@12079
   832
  situation; @{term bogus_path} is the key position within the
wenzelm@12079
   833
  file-system where things go awry.
wenzelm@10966
   834
*}
wenzelm@10966
   835
wenzelm@10966
   836
wenzelm@10966
   837
subsection {* Invariance lemmas \label{sec:unix-inv-lemmas} *}
wenzelm@10966
   838
wenzelm@10966
   839
text {*
wenzelm@10966
   840
  The following invariant over the root file-system describes the
wenzelm@10966
   841
  bogus situation in an abstract manner: located at a certain @{term
wenzelm@10966
   842
  path} within the file-system is a non-empty directory that is
wenzelm@44305
   843
  neither owned nor writable by @{term user\<^isub>1}.
wenzelm@10966
   844
*}
wenzelm@10966
   845
wenzelm@21029
   846
definition
wenzelm@20321
   847
  "invariant root path =
wenzelm@10966
   848
    (\<exists>att dir.
wenzelm@17455
   849
      access root path user\<^isub>1 {} = Some (Env att dir) \<and> dir \<noteq> empty \<and>
wenzelm@17455
   850
      user\<^isub>1 \<noteq> owner att \<and>
wenzelm@17455
   851
      access root path user\<^isub>1 {Writable} = None)"
wenzelm@10966
   852
wenzelm@10966
   853
text {*
wenzelm@10966
   854
  Following the general procedure (see
wenzelm@17455
   855
  \secref{sec:unix-inv-procedure}) we will now establish the three key
wenzelm@17455
   856
  lemmas required to yield the final result.
wenzelm@10966
   857
wenzelm@10966
   858
  \begin{enumerate}
wenzelm@10966
   859
wenzelm@10966
   860
  \item The invariant is sufficiently strong to entail the
wenzelm@17455
   861
  pathological case that @{term user\<^isub>1} is unable to remove the
wenzelm@17455
   862
  (owned) directory at @{term "[user\<^isub>1, name\<^isub>1]"}.
wenzelm@10966
   863
wenzelm@10966
   864
  \item The invariant does hold after having executed the @{term
wenzelm@10966
   865
  bogus} list of operations (starting with an initial file-system
wenzelm@10966
   866
  configuration).
wenzelm@10966
   867
wenzelm@10966
   868
  \item The invariant is preserved by any file-system operation
wenzelm@17455
   869
  performed by @{term user\<^isub>1} alone, without any help by other
wenzelm@17455
   870
  users.
wenzelm@10966
   871
wenzelm@10966
   872
  \end{enumerate}
wenzelm@10966
   873
wenzelm@10966
   874
  As the invariant appears both as assumptions and conclusions in the
wenzelm@10966
   875
  course of proof, its formulation is rather critical for the whole
wenzelm@10966
   876
  development to work out properly.  In particular, the third step is
wenzelm@10966
   877
  very sensitive to the invariant being either too strong or too weak.
wenzelm@10966
   878
  Moreover the invariant has to be sufficiently abstract, lest the
wenzelm@10966
   879
  proof become cluttered by confusing detail.
wenzelm@10966
   880
wenzelm@10966
   881
  \medskip The first two lemmas are technically straight forward ---
wenzelm@10966
   882
  we just have to inspect rather special cases.
wenzelm@10966
   883
*}
wenzelm@10966
   884
wenzelm@21029
   885
lemma cannot_rmdir:
wenzelm@18372
   886
  assumes inv: "invariant root bogus_path"
wenzelm@18372
   887
    and rmdir: "root \<midarrow>(Rmdir user\<^isub>1 [user\<^isub>1, name\<^isub>1])\<rightarrow> root'"
wenzelm@18372
   888
  shows False
wenzelm@10966
   889
proof -
wenzelm@18372
   890
  from inv obtain "file" where "access root bogus_path user\<^isub>1 {} = Some file"
wenzelm@18730
   891
    unfolding invariant_def by blast
wenzelm@10966
   892
  moreover
wenzelm@18372
   893
  from rmdir obtain att where
wenzelm@17455
   894
      "access root [user\<^isub>1, name\<^isub>1] user\<^isub>1 {} = Some (Env att empty)"
wenzelm@10966
   895
    by cases auto
wenzelm@18372
   896
  then have "access root ([user\<^isub>1, name\<^isub>1] @ [name\<^isub>2]) user\<^isub>1 {} = empty name\<^isub>2"
wenzelm@10966
   897
    by (simp only: access_empty_lookup lookup_append_some) simp
wenzelm@10966
   898
  ultimately show False by (simp add: bogus_path_def)
wenzelm@10966
   899
qed
wenzelm@10966
   900
wenzelm@21029
   901
lemma
wenzelm@18372
   902
  assumes init: "init users =bogus\<Rightarrow> root"
wenzelm@18372
   903
  notes eval = facts access_def init_def
wenzelm@18372
   904
  shows init_invariant: "invariant root bogus_path"
wenzelm@18372
   905
  using init
wenzelm@18372
   906
  apply (unfold bogus_def bogus_path_def)
wenzelm@18372
   907
  apply (drule transitions_consD, rule transition.intros,
wenzelm@10966
   908
      (force simp add: eval)+, (simp add: eval)?)+
wenzelm@18372
   909
    -- "evaluate all operations"
wenzelm@18372
   910
  apply (drule transitions_nilD)
wenzelm@18372
   911
    -- "reach final result"
wenzelm@18372
   912
  apply (simp add: invariant_def eval)
wenzelm@18372
   913
    -- "check the invariant"
wenzelm@18372
   914
  done
wenzelm@10966
   915
wenzelm@10966
   916
text {*
wenzelm@10966
   917
  \medskip At last we are left with the main effort to show that the
wenzelm@10966
   918
  ``bogosity'' invariant is preserved by any file-system operation
wenzelm@17455
   919
  performed by @{term user\<^isub>1} alone.  Note that this holds for
wenzelm@17455
   920
  any @{term path} given, the particular @{term bogus_path} is not
wenzelm@10966
   921
  required here.
wenzelm@11004
   922
*}
wenzelm@10966
   923
wenzelm@21029
   924
lemma preserve_invariant:
wenzelm@18372
   925
  assumes tr: "root \<midarrow>x\<rightarrow> root'"
wenzelm@18372
   926
    and inv: "invariant root path"
wenzelm@18372
   927
    and uid: "uid_of x = user\<^isub>1"
wenzelm@18372
   928
  shows "invariant root' path"
wenzelm@10966
   929
proof -
wenzelm@10966
   930
  from inv obtain att dir where
wenzelm@17455
   931
      inv1: "access root path user\<^isub>1 {} = Some (Env att dir)" and
wenzelm@10966
   932
      inv2: "dir \<noteq> empty" and
wenzelm@17455
   933
      inv3: "user\<^isub>1 \<noteq> owner att" and
wenzelm@17455
   934
      inv4: "access root path user\<^isub>1 {Writable} = None"
wenzelm@10966
   935
    by (auto simp add: invariant_def)
wenzelm@10966
   936
wenzelm@10966
   937
  from inv1 have lookup: "lookup root path = Some (Env att dir)"
wenzelm@10966
   938
    by (simp only: access_empty_lookup)
wenzelm@10966
   939
wenzelm@17455
   940
  from inv1 inv3 inv4 and user\<^isub>1_not_root
wenzelm@10966
   941
  have not_writable: "Writable \<notin> others att"
wenzelm@39471
   942
    by (auto simp add: access_def split: option.splits)
wenzelm@10966
   943
wenzelm@10966
   944
  show ?thesis
wenzelm@10966
   945
  proof cases
wenzelm@10966
   946
    assume "root' = root"
wenzelm@10966
   947
    with inv show "invariant root' path" by (simp only:)
wenzelm@10966
   948
  next
wenzelm@10966
   949
    assume changed: "root' \<noteq> root"
wenzelm@10966
   950
    with tr obtain opt where root': "root' = update (path_of x) opt root"
wenzelm@10966
   951
      by cases auto
wenzelm@10966
   952
    show ?thesis
wenzelm@10966
   953
    proof (rule prefix_cases)
wenzelm@10966
   954
      assume "path_of x \<parallel> path"
wenzelm@10966
   955
      with inv root'
wenzelm@17455
   956
      have "\<And>perms. access root' path user\<^isub>1 perms = access root path user\<^isub>1 perms"
wenzelm@10966
   957
        by (simp only: access_update_other)
wenzelm@10966
   958
      with inv show "invariant root' path"
wenzelm@10966
   959
        by (simp only: invariant_def)
wenzelm@10966
   960
    next
wenzelm@10966
   961
      assume "path_of x \<le> path"
wenzelm@10966
   962
      then obtain ys where path: "path = path_of x @ ys" ..
wenzelm@10966
   963
wenzelm@10966
   964
      show ?thesis
wenzelm@10966
   965
      proof (cases ys)
wenzelm@10966
   966
        assume "ys = []"
wenzelm@17455
   967
        with tr uid inv2 inv3 lookup changed path and user\<^isub>1_not_root
wenzelm@10966
   968
        have False
wenzelm@10966
   969
          by cases (auto simp add: access_empty_lookup dest: access_some_lookup)
wenzelm@18372
   970
        then show ?thesis ..
wenzelm@10966
   971
      next
wenzelm@10966
   972
        fix z zs assume ys: "ys = z # zs"
wenzelm@10966
   973
        have "lookup root' path = lookup root path"
wenzelm@10966
   974
        proof -
wenzelm@10966
   975
          from inv2 lookup path ys
wenzelm@10966
   976
          have look: "lookup root (path_of x @ z # zs) = Some (Env att dir)"
wenzelm@10966
   977
            by (simp only:)
wenzelm@10966
   978
          then obtain att' dir' file' where
wenzelm@10966
   979
              look': "lookup root (path_of x) = Some (Env att' dir')" and
wenzelm@10966
   980
              dir': "dir' z = Some file'" and
wenzelm@10966
   981
              file': "lookup file' zs = Some (Env att dir)"
wenzelm@10966
   982
            by (blast dest: lookup_some_upper)
wenzelm@10966
   983
wenzelm@10966
   984
          from tr uid changed look' dir' obtain att'' where
wenzelm@10966
   985
              look'': "lookup root' (path_of x) = Some (Env att'' dir')"
wenzelm@10966
   986
            by cases (auto simp add: access_empty_lookup lookup_update_some
wenzelm@10966
   987
              dest: access_some_lookup)
wenzelm@10966
   988
          with dir' file' have "lookup root' (path_of x @ z # zs) =
wenzelm@10966
   989
              Some (Env att dir)"
wenzelm@10966
   990
            by (simp add: lookup_append_some)
wenzelm@10966
   991
          with look path ys show ?thesis
wenzelm@10966
   992
            by simp
wenzelm@10966
   993
        qed
wenzelm@10966
   994
        with inv show "invariant root' path"
wenzelm@10966
   995
          by (simp only: invariant_def access_def)
wenzelm@10966
   996
      qed
wenzelm@10966
   997
    next
wenzelm@10966
   998
      assume "path < path_of x"
wenzelm@10966
   999
      then obtain y ys where path: "path_of x = path @ y # ys" ..
wenzelm@10966
  1000
wenzelm@10966
  1001
      obtain dir' where
wenzelm@10966
  1002
        lookup': "lookup root' path = Some (Env att dir')" and
wenzelm@10966
  1003
        inv2': "dir' \<noteq> empty"
wenzelm@10966
  1004
      proof (cases ys)
wenzelm@10966
  1005
        assume "ys = []"
wenzelm@10966
  1006
        with path have parent: "path_of x = path @ [y]" by simp
berghofe@17146
  1007
        with tr uid inv4 changed obtain "file" where
wenzelm@10966
  1008
            "root' = update (path_of x) (Some file) root"
wenzelm@10966
  1009
          by cases auto
wenzelm@10979
  1010
        with lookup parent have "lookup root' path = Some (Env att (dir(y\<mapsto>file)))"
wenzelm@10966
  1011
          by (simp only: update_append_some update_cons_nil_env)
wenzelm@10966
  1012
        moreover have "dir(y\<mapsto>file) \<noteq> empty" by simp
wenzelm@10966
  1013
        ultimately show ?thesis ..
wenzelm@10966
  1014
      next
wenzelm@10966
  1015
        fix z zs assume ys: "ys = z # zs"
wenzelm@10966
  1016
        with lookup root' path
wenzelm@10966
  1017
        have "lookup root' path = Some (update (y # ys) opt (Env att dir))"
wenzelm@10966
  1018
          by (simp only: update_append_some)
wenzelm@10966
  1019
        also obtain file' where
wenzelm@10966
  1020
          "update (y # ys) opt (Env att dir) = Env att (dir(y\<mapsto>file'))"
wenzelm@10966
  1021
        proof -
wenzelm@10966
  1022
          have "dir y \<noteq> None"
wenzelm@10966
  1023
          proof -
wenzelm@10966
  1024
            have "dir y = lookup (Env att dir) [y]"
wenzelm@10966
  1025
              by (simp split: option.splits)
wenzelm@10966
  1026
            also from lookup have "\<dots> = lookup root (path @ [y])"
wenzelm@10966
  1027
              by (simp only: lookup_append_some)
wenzelm@10966
  1028
            also have "\<dots> \<noteq> None"
wenzelm@10966
  1029
            proof -
wenzelm@10966
  1030
              from ys obtain us u where rev_ys: "ys = us @ [u]"
berghofe@13601
  1031
                by (cases ys rule: rev_cases) fastsimp+
wenzelm@10966
  1032
              with tr path
wenzelm@10966
  1033
              have "lookup root ((path @ [y]) @ (us @ [u])) \<noteq> None \<or>
wenzelm@10966
  1034
                  lookup root ((path @ [y]) @ us) \<noteq> None"
wenzelm@10966
  1035
                by cases (auto dest: access_some_lookup)
paulson@18447
  1036
              then show ?thesis 
nipkow@18576
  1037
                by (simp, blast dest!: lookup_some_append)
wenzelm@10966
  1038
            qed
wenzelm@10966
  1039
            finally show ?thesis .
wenzelm@10966
  1040
          qed
wenzelm@39471
  1041
          with ys show ?thesis using that by auto
wenzelm@10966
  1042
        qed
wenzelm@10966
  1043
        also have "dir(y\<mapsto>file') \<noteq> empty" by simp
wenzelm@10966
  1044
        ultimately show ?thesis ..
wenzelm@10966
  1045
      qed
wenzelm@10966
  1046
wenzelm@17455
  1047
      from lookup' have inv1': "access root' path user\<^isub>1 {} = Some (Env att dir')"
wenzelm@10966
  1048
        by (simp only: access_empty_lookup)
wenzelm@10966
  1049
wenzelm@17455
  1050
      from inv3 lookup' and not_writable user\<^isub>1_not_root
wenzelm@17455
  1051
      have "access root' path user\<^isub>1 {Writable} = None"
wenzelm@10966
  1052
        by (simp add: access_def)
wenzelm@18730
  1053
      with inv1' inv2' inv3 show ?thesis unfolding invariant_def by blast
wenzelm@10966
  1054
    qed
wenzelm@10966
  1055
  qed
wenzelm@10966
  1056
qed
wenzelm@10966
  1057
wenzelm@10966
  1058
wenzelm@10966
  1059
subsection {* Putting it all together \label{sec:unix-main-result} *}
wenzelm@10966
  1060
wenzelm@10966
  1061
text {*
wenzelm@10966
  1062
  The main result is now imminent, just by composing the three
wenzelm@10966
  1063
  invariance lemmas (see \secref{sec:unix-inv-lemmas}) according the the
wenzelm@10966
  1064
  overall procedure (see \secref{sec:unix-inv-procedure}).
wenzelm@10966
  1065
*}
wenzelm@10966
  1066
wenzelm@21029
  1067
corollary
wenzelm@13380
  1068
  assumes bogus: "init users =bogus\<Rightarrow> root"
wenzelm@17455
  1069
  shows "\<not> (\<exists>xs. (\<forall>x \<in> set xs. uid_of x = user\<^isub>1) \<and>
wenzelm@17455
  1070
    can_exec root (xs @ [Rmdir user\<^isub>1 [user\<^isub>1, name\<^isub>1]]))"
wenzelm@10966
  1071
proof -
wenzelm@13380
  1072
  from cannot_rmdir init_invariant preserve_invariant
wenzelm@13380
  1073
    and bogus show ?thesis by (rule general_procedure)
wenzelm@10966
  1074
qed
wenzelm@10966
  1075
wenzelm@10966
  1076
end
wenzelm@21029
  1077
wenzelm@21029
  1078
end