diff -r a1a251b297dd -r 1b3780be6cc2 src/HOL/Unix/Unix.thy --- a/src/HOL/Unix/Unix.thy Thu Feb 16 21:11:58 2006 +0100 +++ b/src/HOL/Unix/Unix.thy Thu Feb 16 21:12:00 2006 +0100 @@ -160,15 +160,13 @@ \cite{Nipkow-et-al:2000:HOL}. *} -constdefs - attributes :: "file \ att" - "attributes file \ +definition + "attributes file = (case file of Val (att, text) \ att | Env att dir \ att)" - attributes_update :: "att \ file \ file" - "attributes_update att file \ + "attributes_update att file = (case file of Val (att', text) \ Val (att, text) | Env att' dir \ Env att dir)" @@ -201,9 +199,8 @@ has user id 0). *} -constdefs - init :: "uid set \ file" - "init users \ +definition + "init users = Env \owner = 0, others = {Readable}\ (\u. if u \ users then Some (Env \owner = u, others = {Readable}\ empty) else None)" @@ -223,8 +220,7 @@ access a file unconditionally (in our simplified model). *} -constdefs - access :: "file \ path \ uid \ perms \ file option" +definition "access root path uid perms \ (case lookup root path of None \ None @@ -648,8 +644,7 @@ transition. *} -constdefs - can_exec :: "file \ operation list \ bool" +definition "can_exec root xs \ \root'. root =xs\ root'" lemma can_exec_nil: "can_exec root []" @@ -843,13 +838,12 @@ notes facts = user\<^isub>1_known user\<^isub>1_not_root users_neq perms\<^isub>1_writable perms\<^isub>2_not_writable - fixes bogus :: "operation list" - and bogus_path :: path - defines "bogus \ +definition (in situation) + "bogus = [Mkdir user\<^isub>1 perms\<^isub>1 [user\<^isub>1, name\<^isub>1], Mkdir user\<^isub>2 perms\<^isub>2 [user\<^isub>1, name\<^isub>1, name\<^isub>2], Creat user\<^isub>2 perms\<^isub>2 [user\<^isub>1, name\<^isub>1, name\<^isub>2, name\<^isub>3]]" - and "bogus_path \ [user\<^isub>1, name\<^isub>1, name\<^isub>2]" + "bogus_path = [user\<^isub>1, name\<^isub>1, name\<^isub>2]" text {* The @{term bogus} operations are the ones that lead into the uncouth @@ -867,9 +861,8 @@ neither owned and nor writable by @{term user\<^isub>1}. *} -locale invariant = situation + - fixes invariant :: "file \ path \ bool" - defines "invariant root path \ +definition (in situation) + "invariant root path \ (\att dir. access root path user\<^isub>1 {} = Some (Env att dir) \ dir \ empty \ user\<^isub>1 \ owner att \ @@ -907,7 +900,7 @@ we just have to inspect rather special cases. *} -lemma (in invariant) cannot_rmdir: +lemma (in situation) cannot_rmdir: assumes inv: "invariant root bogus_path" and rmdir: "root \(Rmdir user\<^isub>1 [user\<^isub>1, name\<^isub>1])\ root'" shows False @@ -923,7 +916,7 @@ ultimately show False by (simp add: bogus_path_def) qed -lemma (in invariant) +lemma (in situation) assumes init: "init users =bogus\ root" notes eval = facts access_def init_def shows init_invariant: "invariant root bogus_path" @@ -946,7 +939,7 @@ required here. *} -lemma (in invariant) preserve_invariant: +lemma (in situation) preserve_invariant: assumes tr: "root \x\ root'" and inv: "invariant root path" and uid: "uid_of x = user\<^isub>1" @@ -1090,8 +1083,7 @@ overall procedure (see \secref{sec:unix-inv-procedure}). *} -corollary result: - includes invariant +corollary (in situation) assumes bogus: "init users =bogus\ root" shows "\ (\xs. (\x \ set xs. uid_of x = user\<^isub>1) \ can_exec root (xs @ [Rmdir user\<^isub>1 [user\<^isub>1, name\<^isub>1]]))" @@ -1100,11 +1092,4 @@ and bogus show ?thesis by (rule general_procedure) qed -text {* - So this is our final result: - - @{thm [display, show_question_marks = false] result [OF - invariant.intro, OF situation.intro]} -*} - end