# HG changeset patch # User wenzelm # Date 1412707398 -7200 # Node ID 4a16f8e41879b6f3913c4b25f305632010b5afce # Parent dbe216a75a4b4ccd4c99fcddf1d287eb49668250 more cartouches; more antiquotations; diff -r dbe216a75a4b -r 4a16f8e41879 src/HOL/Unix/Nested_Environment.thy --- a/src/HOL/Unix/Nested_Environment.thy Tue Oct 07 20:34:17 2014 +0200 +++ b/src/HOL/Unix/Nested_Environment.thy Tue Oct 07 20:43:18 2014 +0200 @@ -2,13 +2,13 @@ Author: Markus Wenzel, TU Muenchen *) -header {* Nested environments *} +header \Nested environments\ theory Nested_Environment imports Main begin -text {* +text \ Consider a partial function @{term [source] "e :: 'a \ 'b option"}; this may be understood as an \emph{environment} mapping indexes @{typ 'a} to optional entry values @{typ 'b} (cf.\ the basic theory @@ -17,31 +17,31 @@ basic values or again proper environments. Then each entry is accessed by a \emph{path}, i.e.\ a list of indexes leading to its position within the structure. -*} +\ datatype (dead 'a, dead 'b, dead 'c) env = Val 'a | Env 'b "'c \ ('a, 'b, 'c) env option" -text {* +text \ \medskip In the type @{typ "('a, 'b, 'c) env"} the parameter @{typ 'a} refers to basic values (occurring in terminal positions), type @{typ 'b} to values associated with proper (inner) environments, and type @{typ 'c} with the index type for branching. Note that there is no restriction on any of these types. In particular, arbitrary branching may yield rather large (transfinite) tree structures. -*} +\ -subsection {* The lookup operation *} +subsection \The lookup operation\ -text {* +text \ Lookup in nested environments works by following a given path of index elements, leading to an optional result (a terminal value or nested environment). A \emph{defined position} within a nested environment is one where @{term lookup} at its path does not yield @{term None}. -*} +\ primrec lookup :: "('a, 'b, 'c) env \ 'c list \ ('a, 'b, 'c) env option" and lookup_option :: "('a, 'b, 'c) env option \ 'c list \ ('a, 'b, 'c) env option" @@ -56,10 +56,10 @@ hide_const lookup_option -text {* +text \ \medskip The characteristic cases of @{term lookup} are expressed by the following equalities. -*} +\ theorem lookup_nil: "lookup e [] = Some e" by (cases e) simp_all @@ -90,12 +90,12 @@ | Some e \ lookup e xs)))" by (simp split: list.split env.split) -text {* +text \ \medskip Displaced @{term lookup} operations, relative to a certain base path prefix, may be reduced as follows. There are two cases, depending whether the environment actually extends far enough to follow the base path. -*} +\ theorem lookup_append_none: assumes "lookup env xs = None" @@ -119,7 +119,7 @@ with Env show ?thesis by simp next case (Some e) - note es = `es x = Some e` + note es = \es x = Some e\ show ?thesis proof (cases "lookup e xs") case None @@ -144,7 +144,7 @@ then show "lookup env ([] @ ys) = lookup e ys" by simp next case (Cons x xs) - note asm = `lookup env (x # xs) = Some e` + note asm = \lookup env (x # xs) = Some e\ show "lookup env ((x # xs) @ ys) = lookup e ys" proof (cases env) case (Val a) @@ -159,7 +159,7 @@ then show ?thesis .. next case (Some e') - note es = `es x = Some e'` + note es = \es x = Some e'\ show ?thesis proof (cases "lookup e' xs") case None @@ -176,12 +176,12 @@ qed qed -text {* +text \ \medskip Successful @{term lookup} deeper down an environment structure means we are able to peek further up as well. Note that this is basically just the contrapositive statement of @{thm [source] lookup_append_none} above. -*} +\ theorem lookup_some_append: assumes "lookup env (xs @ ys) = Some e" @@ -193,11 +193,11 @@ then show ?thesis by (simp) qed -text {* +text \ The subsequent statement describes in more detail how a successful @{term lookup} with a non-empty path results in a certain situation at any upper position. -*} +\ theorem lookup_some_upper: assumes "lookup env (xs @ y # ys) = Some e" @@ -236,14 +236,14 @@ qed -subsection {* The update operation *} +subsection \The update operation\ -text {* +text \ Update at a certain position in a nested environment may either delete an existing entry, or overwrite an existing one. Note that update at undefined positions is simple absorbed, i.e.\ the environment is left unchanged. -*} +\ primrec update :: "'c list \ ('a, 'b, 'c) env option \ ('a, 'b, 'c) env \ ('a, 'b, 'c) env" @@ -264,10 +264,10 @@ hide_const update_option -text {* +text \ \medskip The characteristic cases of @{term update} are expressed by the following equalities. -*} +\ theorem update_nil_none: "update [] None env = env" by (cases env) simp_all @@ -314,11 +314,11 @@ | Some e \ Some (update (y # ys) opt e)))))))" by (simp split: list.split env.split option.split) -text {* +text \ \medskip The most basic correspondence of @{term lookup} and @{term update} states that after @{term update} at a defined position, subsequent @{term lookup} operations would yield the new value. -*} +\ theorem lookup_update_some: assumes "lookup env xs = Some e" @@ -331,7 +331,7 @@ next case (Cons x xs) note hyp = Cons.hyps - and asm = `lookup env (x # xs) = Some e` + and asm = \lookup env (x # xs) = Some e\ show ?case proof (cases env) case (Val a) @@ -346,7 +346,7 @@ then show ?thesis .. next case (Some e') - note es = `es x = Some e'` + note es = \es x = Some e'\ show ?thesis proof (cases xs) case Nil @@ -361,13 +361,13 @@ qed qed -text {* +text \ \medskip The properties of displaced @{term update} operations are analogous to those of @{term lookup} above. There are two cases: below an undefined position @{term update} is absorbed altogether, and below a defined positions @{term update} affects subsequent @{term lookup} operations in the obvious way. -*} +\ theorem update_append_none: assumes "lookup env xs = None" @@ -380,7 +380,7 @@ next case (Cons x xs) note hyp = Cons.hyps - and asm = `lookup env (x # xs) = None` + and asm = \lookup env (x # xs) = None\ show "update ((x # xs) @ y # ys) opt env = env" proof (cases env) case (Val a) @@ -390,12 +390,12 @@ show ?thesis proof (cases "es x") case None - note es = `es x = None` + note es = \es x = None\ show ?thesis by (cases xs) (simp_all add: es Env fun_upd_idem_iff) next case (Some e) - note es = `es x = Some e` + note es = \es x = Some e\ show ?thesis proof (cases xs) case Nil @@ -423,7 +423,7 @@ next case (Cons x xs) note hyp = Cons.hyps - and asm = `lookup env (x # xs) = Some e` + and asm = \lookup env (x # xs) = Some e\ show "lookup (update ((x # xs) @ y # ys) opt env) (x # xs) = Some (update (y # ys) opt e)" proof (cases env) @@ -439,7 +439,7 @@ then show ?thesis .. next case (Some e') - note es = `es x = Some e'` + note es = \es x = Some e'\ show ?thesis proof (cases xs) case Nil @@ -456,12 +456,12 @@ qed qed -text {* +text \ \medskip Apparently, @{term update} does not affect the result of subsequent @{term lookup} operations at independent positions, i.e.\ in case that the paths for @{term update} and @{term lookup} fork at a certain point. -*} +\ theorem lookup_update_other: assumes neq: "y \ (z::'c)" @@ -519,7 +519,7 @@ qed -subsection {* Code generation *} +subsection \Code generation\ lemma [code, code del]: "(HOL.equal :: (_, _, _) env \ _) = HOL.equal" .. @@ -549,7 +549,7 @@ show ?lhs proof fix z - from `?rhs` have "?prop z" .. + from \?rhs\ have "?prop z" .. then show "f z = g z" by (auto split: option.splits) qed qed diff -r dbe216a75a4b -r 4a16f8e41879 src/HOL/Unix/Unix.thy --- a/src/HOL/Unix/Unix.thy Tue Oct 07 20:34:17 2014 +0200 +++ b/src/HOL/Unix/Unix.thy Tue Oct 07 20:43:18 2014 +0200 @@ -2,7 +2,7 @@ Author: Markus Wenzel, TU Muenchen *) -header {* Unix file-systems \label{sec:unix-file-system} *} +header \Unix file-systems \label{sec:unix-file-system}\ theory Unix imports @@ -10,7 +10,7 @@ "~~/src/HOL/Library/Sublist" begin -text {* +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 @@ -34,26 +34,26 @@ 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 *} +subsection \Names\ -text {* +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 *} +subsection \Attributes\ -text {* +text \ Unix file attributes mainly consist of \emph{owner} information and a number of \emph{permission} bits which control access for ``user'', ``group'', and ``others'' (see the Unix man pages @{text @@ -70,8 +70,8 @@ of multiple groups and group membership, but pretend that everyone is member of a single global group.\footnote{A general HOL model of user group structures and related issues is given in - \cite{Naraschewski:2001}.} -*} + @{cite "Naraschewski:2001"}.} +\ datatype perm = Readable @@ -84,7 +84,7 @@ owner :: uid others :: perms -text {* +text \ For plain files @{term Readable} and @{term Writable} specify read and write access to the actual content, i.e.\ the string of text stored here. For directories @{term Readable} determines if the set @@ -101,16 +101,16 @@ directories encountered on the path would have to grant @{term Executable}. We ignore this detail and pretend that all directories give @{term Executable} permission to anybody. -*} +\ -subsection {* Files *} +subsection \Files\ -text {* +text \ In order to model the general tree structure of a Unix file-system we use the arbitrarily branching datatype @{typ "('a, 'b, 'c) env"} from the standard library of Isabelle/HOL - \cite{Bauer-et-al:2002:HOL-Library}. This type provides + @{cite "Bauer-et-al:2002:HOL-Library"}. This type provides constructors @{term Val} and @{term Env} as follows: \medskip @@ -128,11 +128,11 @@ "att \ string"} (representing plain files), @{typ att} (for attributes of directory nodes), and @{typ name} (for the index type of directory nodes). -*} +\ type_synonym "file" = "(att \ string, att, name) env" -text {* +text \ \medskip The HOL library also provides @{term lookup} and @{term update} operations for general tree structures with the subsequent primitive recursive characterizations. @@ -149,15 +149,15 @@ @{thm [display, indent = 2, eta_contract = false] update_eq [no_vars]} Several further properties of these operations are proven in - \cite{Bauer-et-al:2002:HOL-Library}. These will be routinely used + @{cite "Bauer-et-al:2002:HOL-Library"}. These will be routinely used later on without further notice. \bigskip Apparently, the elements of type @{typ "file"} contain an @{typ att} component in either case. We now define a few auxiliary operations to manipulate this field uniformly, following the conventions for record types in Isabelle/HOL - \cite{Nipkow-et-al:2000:HOL}. -*} + @{cite "Nipkow-et-al:2000:HOL"}. +\ definition "attributes file = @@ -188,16 +188,16 @@ by (simp add: map_attributes_def) -subsection {* Initial file-systems *} +subsection \Initial file-systems\ -text {* +text \ Given a set of \emph{known users} a file-system shall be initialized by providing an empty home directory for each user, with read-only access for everyone else. (Note that we may directly use the user id as home directory name, since both types have been identified.) Certainly, the very root directory is owned by the super user (who has user id 0). -*} +\ definition "init users = @@ -206,9 +206,9 @@ else None)" -subsection {* Accessing file-systems *} +subsection \Accessing file-systems\ -text {* +text \ The main internal file-system operation is access of a file by a user, requesting a certain set of permissions. The resulting @{typ "file option"} indicates if a file had been present at the @@ -216,9 +216,9 @@ permissions recorded within the file-system. Note that by the rules of Unix file-system security (e.g.\ - \cite{Tanenbaum:1992}) both the super-user and owner may always + @{cite "Tanenbaum:1992"}) both the super-user and owner may always access a file unconditionally (in our simplified model). -*} +\ definition "access root path uid perms = @@ -231,7 +231,7 @@ then Some file else None)" -text {* +text \ \medskip Successful access to a certain file is the main prerequisite for system-calls to be applicable (cf.\ \secref{sec:unix-trans}). Any modification of the file-system is @@ -241,7 +241,7 @@ @{term lookup} function, with additional checking of attributes. Subsequently we establish a few auxiliary facts that stem from the primitive @{term lookup} used within @{term access}. -*} +\ lemma access_empty_lookup: "access root path uid {} = lookup root path" by (simp add: access_def split: option.splits) @@ -264,13 +264,13 @@ qed -section {* File-system transitions \label{sec:unix-trans} *} +section \File-system transitions \label{sec:unix-trans}\ -subsection {* Unix system calls \label{sec:unix-syscall} *} +subsection \Unix system calls \label{sec:unix-syscall}\ -text {* +text \ According to established operating system design (cf.\ - \cite{Tanenbaum:1992}) user space processes may only initiate system + @{cite "Tanenbaum:1992"}) user space processes may only initiate system operations by a fixed set of \emph{system-calls}. This enables the kernel to enforce certain security policies in the first place.\footnote{Incidently, this is the very same principle employed @@ -282,7 +282,7 @@ operation} for the syntax of system-calls, together with an inductive definition of file-system state transitions of the form @{text "root \x\ root'"} for the operational semantics. -*} +\ datatype operation = Read uid string path @@ -294,13 +294,13 @@ | Rmdir uid path | Readdir uid "name set" path -text {* +text \ The @{typ uid} field of an operation corresponds to the \emph{effective user id} of the underlying process, although our model never mentions processes explicitly. The other parameters are provided as arguments by the caller; the @{term path} one is common to all kinds of system-calls. -*} +\ primrec uid_of :: "operation \ uid" where @@ -324,7 +324,7 @@ | "path_of (Rmdir uid path) = path" | "path_of (Readdir uid names path) = path" -text {* +text \ \medskip Note that we have omitted explicit @{text Open} and @{text Close} operations, pretending that @{term Read} and @{term Write} would already take care of this behind the scenes. Thus we have @@ -345,7 +345,7 @@ via transitions of the file-system configuration. This is expressed as an inductive relation (although there is no actual recursion involved here). -*} +\ inductive transition :: "file \ operation \ file \ bool" ("_ \_\ _" [90, 1000, 90] 100) @@ -392,7 +392,7 @@ names = dom dir \ root \(Readdir uid names path)\ root" -text {* +text \ \medskip Certainly, the above specification is central to the whole formal development. Any of the results to be established later on are only meaningful to the outside world if this transition system @@ -403,16 +403,16 @@ If in doubt, one may consider to compare our definition with the informal specifications given the corresponding Unix man pages, or even peek at an actual implementation such as - \cite{Torvalds-et-al:Linux}. Another common way to gain confidence + @{cite "Torvalds-et-al:Linux"}. Another common way to gain confidence into the formal model is to run simple simulations (see \secref{sec:unix-examples}), and check the results with that of experiments performed on a real Unix system. -*} +\ -subsection {* Basic properties of single transitions \label{sec:unix-single-trans} *} +subsection \Basic properties of single transitions \label{sec:unix-single-trans}\ -text {* +text \ The transition system @{text "root \x\ root'"} defined above determines a unique result @{term root'} from given @{term root} and @{term x} (this is holds rather trivially, since there is even only @@ -420,7 +420,7 @@ simplify our subsequent development to some extent, since we only have to reason about a partial function rather than a general relation. -*} +\ theorem transition_uniq: assumes root': "root \x\ root'" @@ -453,11 +453,11 @@ with root' show ?thesis by cases blast+ qed -text {* +text \ \medskip Apparently, file-system transitions are \emph{type-safe} in the sense that the result of transforming an actual directory yields again a directory. -*} +\ theorem transition_type_safe: assumes tr: "root \x\ root'" @@ -476,25 +476,25 @@ by (auto simp add: update_eq split: list.splits) qed -text {* +text \ The previous result may be seen as the most basic invariant on the file-system state that is enforced by any proper kernel implementation. So user processes --- being bound to the system-call interface --- may never mess up a file-system such that the root becomes a plain file instead of a directory, which would be a strange situation indeed. -*} +\ -subsection {* Iterated transitions *} +subsection \Iterated transitions\ -text {* +text \ Iterated system transitions via finite sequences of system operations are modeled inductively as follows. In a sense, this relation describes the cumulative effect of the sequence of system-calls issued by a number of running processes over a finite amount of time. -*} +\ inductive transitions :: "file \ operation list \ file \ bool" ("_ =_\ _" [90, 1000, 90] 100) @@ -502,11 +502,11 @@ nil: "root =[]\ root" | cons: "root \x\ root' \ root' =xs\ root'' \ root =(x # xs)\ root''" -text {* +text \ \medskip We establish a few basic facts relating iterated transitions with single ones, according to the recursive structure of lists. -*} +\ lemma transitions_nil_eq: "root =[]\ root' = (root = root')" proof @@ -529,12 +529,12 @@ by (blast intro: transitions.cons) qed -text {* +text \ The next two rules show how to ``destruct'' known transition sequences. Note that the second one actually relies on the uniqueness property of the basic transition system (see \secref{sec:unix-single-trans}). -*} +\ lemma transitions_nilD: "root =[]\ root' \ root' = root" by (simp add: transitions_nil_eq) @@ -550,12 +550,12 @@ with root'' show "root' =xs\ root''" by simp qed -text {* +text \ \medskip The following fact shows how an invariant @{term Q} of single transitions with property @{term P} may be transferred to iterated transitions. The proof is rather obvious by rule induction over the definition of @{term "root =xs\ root'"}. -*} +\ lemma transitions_invariant: assumes r: "\r x r'. r \x\ r' \ Q r \ P x \ Q r'" @@ -567,19 +567,19 @@ show ?case by (rule nil.prems) next case (cons root x root' xs root'') - note P = `\x \ set (x # xs). P x` + note P = \\x \ set (x # xs). P x\ then have "P x" by simp - with `root \x\ root'` and `Q root` have Q': "Q root'" by (rule r) + with \root \x\ root'\ and \Q root\ have Q': "Q root'" by (rule r) from P have "\x \ set xs. P x" by simp with Q' show "Q root''" by (rule cons.hyps) qed -text {* +text \ As an example of applying the previous result, we transfer the basic type-safety property (see \secref{sec:unix-single-trans}) from single transitions to iterated ones, which is a rather obvious result indeed. -*} +\ theorem transitions_type_safe: assumes "root =xs\ root'" @@ -591,9 +591,9 @@ qed -section {* Executable sequences *} +section \Executable sequences\ -text {* +text \ An inductively defined relation such as the one of @{text "root \x\ root'"} (see \secref{sec:unix-syscall}) has two main aspects. First of all, the resulting system admits a certain set of transition @@ -614,16 +614,16 @@ executions; exhaustive reasoning will be employed only later on (see \secref{sec:unix-bogosity}), when we shall demonstrate that certain behavior is \emph{not} possible. -*} +\ -subsection {* Possible transitions *} +subsection \Possible transitions\ -text {* +text \ Rather obviously, a list of system operations can be executed within a certain state if there is a result state reached by an iterated transition. -*} +\ definition "can_exec root xs = (\root'. root =xs\ root')" @@ -634,10 +634,10 @@ "root \x\ root' \ can_exec root' xs \ can_exec root (x # xs)" unfolding can_exec_def by (blast intro: transitions.intros) -text {* +text \ \medskip In case that we already know that a certain sequence can be executed we may destruct it backwardly into individual transitions. -*} +\ lemma can_exec_snocD: "can_exec root (xs @ [y]) \ \root' root''. root =xs\ root' \ root' \y\ root''" @@ -647,7 +647,7 @@ by (simp add: can_exec_def transitions_nil_eq transitions_cons_eq) next case (Cons x xs) - from `can_exec root ((x # xs) @ [y])` obtain r root'' where + from \can_exec root ((x # xs) @ [y])\ obtain r root'' where x: "root \x\ r" and xs_y: "r =(xs @ [y])\ root''" by (auto simp add: can_exec_def transitions_nil_eq transitions_cons_eq) @@ -659,35 +659,35 @@ qed -subsection {* Example executions \label{sec:unix-examples} *} +subsection \Example executions \label{sec:unix-examples}\ -text {* +text \ We are now ready to perform a few experiments within our formal model of Unix system-calls. The common technique is to alternate introduction rules of the transition system (see \secref{sec:unix-trans}), and steps to solve any emerging side conditions using algebraic properties of the underlying file-system structures (see \secref{sec:unix-file-system}). -*} +\ lemmas eval = access_def init_def theorem "u \ users \ can_exec (init users) [Mkdir u perms [u, name]]" apply (rule can_exec_cons) - -- {* back-chain @{term can_exec} (of @{term [source] Cons}) *} + -- \back-chain @{term can_exec} (of @{term [source] Cons})\ apply (rule mkdir) - -- {* back-chain @{term Mkdir} *} + -- \back-chain @{term Mkdir}\ apply (force simp add: eval)+ - -- {* solve preconditions of @{term Mkdir} *} + -- \solve preconditions of @{term Mkdir}\ apply (simp add: eval) - -- {* peek at resulting dir (optional) *} - txt {* @{subgoals [display]} *} + -- \peek at resulting dir (optional)\ + txt \@{subgoals [display]}\ apply (rule can_exec_nil) - -- {* back-chain @{term can_exec} (of @{term [source] Nil}) *} + -- \back-chain @{term can_exec} (of @{term [source] Nil})\ done -text {* +text \ By inspecting the result shown just before the final step above, we may gain confidence that our specification of Unix system-calls actually makes sense. Further common errors are usually exhibited @@ -695,7 +695,7 @@ \medskip Here are a few further experiments, using the same techniques as before. -*} +\ theorem "u \ users \ can_exec (init users) [Creat u perms [u, name], @@ -708,7 +708,7 @@ apply (rule unlink) apply (force simp add: eval)+ apply (simp add: eval) - txt {* peek at result: @{subgoals [display]} *} + txt \peek at result: @{subgoals [display]}\ apply (rule can_exec_nil) done @@ -722,7 +722,7 @@ Readdir u {name\<^sub>3, name\<^sub>4} [u, name\<^sub>1, name\<^sub>2]]" apply (rule can_exec_cons, rule transition.intros, (force simp add: eval)+, (simp add: eval)?)+ - txt {* peek at result: @{subgoals [display]} *} + txt \peek at result: @{subgoals [display]}\ apply (rule can_exec_nil) done @@ -735,28 +735,28 @@ Read u ''foo'' [u, name\<^sub>1, name\<^sub>2, name\<^sub>3]]" apply (rule can_exec_cons, rule transition.intros, (force simp add: eval)+, (simp add: eval)?)+ - txt {* peek at result: @{subgoals [display]} *} + txt \peek at result: @{subgoals [display]}\ apply (rule can_exec_nil) done -section {* Odd effects --- treated formally \label{sec:unix-bogosity} *} +section \Odd effects --- treated formally \label{sec:unix-bogosity}\ -text {* +text \ We are now ready to give a completely formal treatment of the slightly odd situation discussed in the introduction (see \secref{sec:unix-intro}): the file-system can easily reach a state where a user is unable to remove his very own directory, because it is still populated by items placed there by another user in an uncouth manner. -*} +\ -subsection {* The general procedure \label{sec:unix-inv-procedure} *} +subsection \The general procedure \label{sec:unix-inv-procedure}\ -text {* +text \ The following theorem expresses the general procedure we are following to achieve the main result. -*} +\ theorem general_procedure: assumes cannot_y: "\r r'. Q r \ r \y\ r' \ False" @@ -780,7 +780,7 @@ then show ?thesis by blast qed -text {* +text \ Here @{prop "P x"} refers to the restriction on file-system operations that are admitted after having reached the critical configuration; according to the problem specification this will @@ -790,16 +790,16 @@ @{term Q} is a suitable (auxiliary) invariant over the file-system; choosing @{term Q} adequately is very important to make the proof work (see \secref{sec:unix-inv-lemmas}). -*} +\ -subsection {* The particular situation *} +subsection \The particular situation\ -text {* +text \ We introduce a few global declarations and axioms to describe our particular situation considered here. Thus we avoid excessive use of local parameters in the subsequent development. -*} +\ locale situation = fixes users :: "uid set" @@ -827,21 +827,21 @@ definition "bogus_path = [user\<^sub>1, name\<^sub>1, name\<^sub>2]" -text {* +text \ The @{term bogus} operations are the ones that lead into the uncouth situation; @{term bogus_path} is the key position within the file-system where things go awry. -*} +\ -subsection {* Invariance lemmas \label{sec:unix-inv-lemmas} *} +subsection \Invariance lemmas \label{sec:unix-inv-lemmas}\ -text {* +text \ The following invariant over the root file-system describes the bogus situation in an abstract manner: located at a certain @{term path} within the file-system is a non-empty directory that is neither owned nor writable by @{term user\<^sub>1}. -*} +\ definition "invariant root path = @@ -850,7 +850,7 @@ user\<^sub>1 \ owner att \ access root path user\<^sub>1 {Writable} = None)" -text {* +text \ Following the general procedure (see \secref{sec:unix-inv-procedure}) we will now establish the three key lemmas required to yield the final result. @@ -880,7 +880,7 @@ \medskip The first two lemmas are technically straight forward --- we just have to inspect rather special cases. -*} +\ lemma cannot_rmdir: assumes inv: "invariant root bogus_path" @@ -913,13 +913,13 @@ -- "check the invariant" done -text {* +text \ \medskip At last we are left with the main effort to show that the ``bogosity'' invariant is preserved by any file-system operation performed by @{term user\<^sub>1} alone. Note that this holds for any @{term path} given, the particular @{term bogus_path} is not required here. -*} +\ lemma preserve_invariant: assumes tr: "root \x\ root'" @@ -1056,13 +1056,13 @@ qed -subsection {* Putting it all together \label{sec:unix-main-result} *} +subsection \Putting it all together \label{sec:unix-main-result}\ -text {* +text \ The main result is now imminent, just by composing the three invariance lemmas (see \secref{sec:unix-inv-lemmas}) according the the overall procedure (see \secref{sec:unix-inv-procedure}). -*} +\ corollary assumes bogus: "init users =bogus\ root"