diff -r 4eb8061286e5 -r 054153c48bde src/HOL/Unix/Unix.thy --- a/src/HOL/Unix/Unix.thy Tue Nov 06 23:50:24 2001 +0100 +++ b/src/HOL/Unix/Unix.thy Tue Nov 06 23:51:00 2001 +0100 @@ -843,51 +843,45 @@ *} -subsection {* The particular setup *} +subsection {* The particular situation *} text {* We introduce a few global declarations and axioms to describe our - particular setup considered here. Thus we avoid excessive use of - local parameters in the subsequent development. + particular situation considered here. Thus we avoid excessive use + of local parameters in the subsequent development. *} -consts - users :: "uid set" - user1 :: uid - user2 :: uid - name1 :: name - name2 :: name - name3 :: name - perms1 :: perms - perms2 :: perms +locale situation = + fixes users :: "uid set" + and user1 :: uid + and user2 :: uid + and name1 :: name + and name2 :: name + and name3 :: name + and perms1 :: perms + and perms2 :: perms + assumes user1_known: "user1 \ users" + and user1_not_root: "user1 \ 0" + and users_neq: "user1 \ user2" + and perms1_writable: "Writable \ perms1" + and perms2_not_writable: "Writable \ perms2" + notes facts = user1_known user1_not_root users_neq + perms1_writable perms2_not_writable -axioms - user1_known: "user1 \ users" - user1_not_root: "user1 \ 0" - users_neq: "user1 \ user2" - perms1_writable: "Writable \ perms1" - perms2_not_writable: "Writable \ perms2" - -lemmas "setup" = - user1_known user1_not_root users_neq - perms1_writable perms2_not_writable + fixes bogus :: "operation list" + and bogus_path :: path + defines "bogus \ + [Mkdir user1 perms1 [user1, name1], + Mkdir user2 perms2 [user1, name1, name2], + Creat user2 perms2 [user1, name1, name2, name3]]" + and "bogus_path \ [user1, name1, name2]" text {* - \medskip 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. + 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. *} -constdefs - bogus :: "operation list" - "bogus \ - [Mkdir user1 perms1 [user1, name1], - Mkdir user2 perms2 [user1, name1, name2], - Creat user2 perms2 [user1, name1, name2, name3]]" - - bogus_path :: path - "bogus_path \ [user1, name1, name2]" - subsection {* Invariance lemmas \label{sec:unix-inv-lemmas} *} @@ -898,9 +892,9 @@ neither owned and nor writable by @{term user1}. *} -constdefs - invariant :: "file \ path \ bool" - "invariant root path \ +locale invariant = situation + + fixes invariant :: "file \ path \ bool" + defines "invariant root path \ (\att dir. access root path user1 {} = Some (Env att dir) \ dir \ empty \ user1 \ owner att \ @@ -938,8 +932,9 @@ we just have to inspect rather special cases. *} -lemma cannot_rmdir: "invariant root bogus_path \ - root \(Rmdir user1 [user1, name1])\ root' \ False" +lemma (in invariant) + cannot_rmdir: "invariant root bogus_path \ + root \(Rmdir user1 [user1, name1])\ root' \ False" proof - assume "invariant root bogus_path" then obtain file where "access root bogus_path user1 {} = Some file" @@ -954,9 +949,10 @@ ultimately show False by (simp add: bogus_path_def) qed -lemma init_invariant: "init users =bogus\ root \ invariant root bogus_path" +lemma (in invariant) + init_invariant: "init users =bogus\ root \ invariant root bogus_path" proof - - note eval = "setup" access_def init_def + note eval = facts access_def init_def case rule_context thus ?thesis apply (unfold bogus_def bogus_path_def) apply (drule transitions_consD, rule transition.intros, @@ -977,8 +973,9 @@ required here. *} -lemma preserve_invariant: "root \x\ root' \ - invariant root path \ uid_of x = user1 \ invariant root' path" +lemma (in invariant) + preserve_invariant: "root \x\ root' \ + invariant root path \ uid_of x = user1 \ invariant root' path" proof - assume tr: "root \x\ root'" assume inv: "invariant root path" @@ -1121,9 +1118,9 @@ overall procedure (see \secref{sec:unix-inv-procedure}). *} -corollary +corollary (in invariant) (* FIXME in situation!? *) "init users =bogus\ root \ - \ (\xs. (\x \ set xs. uid_of x = user1) \ + \ (\xs. (\x \ set xs. uid_of x = user1) \ can_exec root (xs @ [Rmdir user1 [user1, name1]]))" proof - case rule_context