diff -r 2af87bb52fca -r 25fd656667a7 PIPDefs.thy --- a/PIPDefs.thy Thu Jan 07 22:10:06 2016 +0800 +++ b/PIPDefs.thy Sat Jan 09 22:19:27 2016 +0800 @@ -37,6 +37,24 @@ V thread cs | -- {* Thread @{text "thread"} releasing critical resource @{text "cs"}. *} Set thread priority -- {* Thread @{text "thread"} resets its priority to @{text "priority"}. *} +fun actor :: "event \ thread" where + "actor (Create th pty) = th" | + "actor (Exit th) = th" | + "actor (P th cs) = th" | + "actor (V th cs) = th" | + "actor (Set th pty) = th" + +fun isCreate :: "event \ bool" where + "isCreate (Create th pty) = True" | + "isCreate _ = False" + +fun isP :: "event \ bool" where + "isP (P th cs) = True" | + "isP _ = False" + +fun isV :: "event \ bool" where + "isV (V th cs) = True" | + "isV _ = False" text {* As mentioned earlier, in Paulson's inductive method, the states of system are represented as lists of events,