diff -r 4d25a9919688 -r 4294abb1f38c rc_theory.thy --- a/rc_theory.thy Fri Apr 12 18:07:03 2013 +0100 +++ b/rc_theory.thy Thu Jun 13 22:12:45 2013 +0800 @@ -463,21 +463,6 @@ (Some r, Some t) \ (r, Proc_type t, DELETE) \ compatible | _ \ False )" - -(**** OS' job: checking resources existence & grant new resource ****) - -definition next_nat :: "nat set \ nat" -where - "next_nat ps = (Max ps) + 1" - -definition new_proc :: "t_state \ t_process" -where - "new_proc \ = next_nat (current_procs \)" - -definition new_ipc :: "t_state \ t_ipc" -where - "new_ipc \ = next_nat (current_ipcs \)" - (* new file pathname is user-defined, so os just check if its unexistence *) fun os_grant :: "t_state \ t_event \ bool" where @@ -487,10 +472,10 @@ | "os_grant \ (ReadFile p f) = (p \ current_procs \ \ f \ current_files \)" | "os_grant \ (WriteFile p f) = (p \ current_procs \ \ f \ current_files \)" | "os_grant \ (ChangeOwner p u)= (p \ current_procs \ \ u \ init_users)" -| "os_grant \ (CreateIPC p i) = (p \ current_procs \ \ i = new_ipc \)" +| "os_grant \ (CreateIPC p i) = (p \ current_procs \ \ i \ current_ipcs \)" (* i = new_ipc \ *) | "os_grant \ (Send p i) = (p \ current_procs \ \ i \ current_ipcs \)" | "os_grant \ (Recv p i) = (p \ current_procs \ \ i \ current_ipcs \)" -| "os_grant \ (Clone p p') = (p \ current_procs \ \ p' = new_proc \)" +| "os_grant \ (Clone p p') = (p \ current_procs \ \ p' \ current_procs \)" (* p' = new_proc \ *) | "os_grant \ (Kill p p') = (p \ current_procs \ \ p' \ current_procs \)" | "os_grant \ (DeleteFile p f) = (p \ current_procs \ \ f \ current_files \ \ \ (\fn. fn # f \ current_files \))" | "os_grant \ (DeleteIPC p i) = (p \ current_procs \ \ i \ current_ipcs \)" @@ -961,7 +946,19 @@ "initp_intact_but s (SProc sp (Some p)) = initp_intact_butp s p" | "initp_intact_but s _ = initp_intact s" -(*** how to generating new valid pathname for examples of CreateFile ***) +(*** how to generating new valid pathname for examples of new_proc, new_ipc and new_file as CreateFile ***) + +definition next_nat :: "nat set \ nat" +where + "next_nat ps = (Max ps) + 1" + +definition new_proc :: "t_state \ t_process" +where + "new_proc \ = next_nat (current_procs \)" + +definition new_ipc :: "t_state \ t_ipc" +where + "new_ipc \ = next_nat (current_ipcs \)" definition all_fname_under_dir:: "t_file \ t_state \ t_fname set" where