Theory Unix
section ‹Unix file-systems \label{sec:unix-file-system}›
theory Unix
imports
Nested_Environment
"HOL-Library.Sublist"
begin
text ‹
We give a simple mathematical model of the basic structures underlying the
Unix file-system, together with a few fundamental operations that could be
imagined to be performed internally by the Unix kernel. This forms the basis
for the set of Unix system-calls to be introduced later (see
\secref{sec:unix-trans}), which are the actual interface offered to
processes running in user-space.
┉
Basically, any Unix file is either a ∗‹plain file› or a ∗‹directory›,
consisting of some ∗‹content› plus ∗‹attributes›. The content of a plain
file is plain text. The content of a directory is a mapping from names to
further files.⁋‹In fact, this is the only way that names get associated with
files. In Unix files do ∗‹not› have a name in itself. Even more, any number
of names may be associated with the very same file due to ∗‹hard links›
(although this is excluded from our model).› Attributes include information
to control various ways to access the file (read, write etc.).
Our model will be quite liberal in omitting excessive detail that is easily
seen to be ``irrelevant'' for the aspects of Unix file-systems to be
discussed here. First of all, we ignore character and block special files,
pipes, sockets, hard links, symbolic links, and mount points.
›
subsection ‹Names›
text ‹
User ids and file name components shall be represented by natural numbers
(without loss of generality). We do not bother about encoding of actual
names (e.g.\ strings), nor a mapping between user names and user ids as
would be present in a reality.
›
type_synonym uid = nat
type_synonym name = nat
type_synonym path = "name list"
subsection ‹Attributes›
text ‹
Unix file attributes mainly consist of ∗‹owner› information and a number of
∗‹permission› bits which control access for ``user'', ``group'', and
``others'' (see the Unix man pages ‹chmod(2)› and ‹stat(2)› for more
details).
┉
Our model of file permissions only considers the ``others'' part. The
``user'' field may be omitted without loss of overall generality, since the
owner is usually able to change it anyway by performing ‹chmod›.⁋‹The
inclined Unix expert may try to figure out some exotic arrangements of a
real-world Unix file-system such that the owner of a file is unable to apply
the ‹chmod› system call.› We omit ``group'' permissions as a genuine
simplification as we just do not intend to discuss a model of multiple
groups and group membership, but pretend that everyone is member of a single
global group.⁋‹A general HOL model of user group structures and related
issues is given in \<^cite>‹"Naraschewski:2001"›.›
›