(* * Copyright 2014, General Dynamics C4 Systems * * This software may be distributed and modified according to the terms of * the GNU General Public License version 2. Note that NO WARRANTY is provided. * See "LICENSE_GPLv2.txt" for details. * * @TAG(GD_GPL) *) (* Specification of Inter-Process Communication. *) chapter "IPC" theory Ipc_A imports Tcb_A begin context begin interpretation Arch . requalify_consts set_message_info lookup_ipc_buffer end section {* Getting and setting the message info register. *} definition get_message_info :: "obj_ref \ (message_info,'z::state_ext) s_monad" where "get_message_info thread \ do x \ as_user thread $ get_register msg_info_register; return $ data_to_message_info x od" section {* IPC Capability Transfers *} definition remove_rights :: "cap_rights \ cap \ cap" where "remove_rights rights cap \ cap_rights_update (cap_rights cap - rights) cap" text {* In addition to the data payload a message may also contain capabilities. When a thread requests additional capabilities be transferred the identities of those capabilities are retreived from the thread's IPC buffer. *} definition buffer_cptr_index :: nat where "buffer_cptr_index \ (msg_max_length + 2)" primrec get_extra_cptrs :: "obj_ref option \ message_info \ (cap_ref list,'z::state_ext) s_monad" where "get_extra_cptrs (Some buf) mi = (liftM (map data_to_cptr) $ mapM (load_word_offs buf) [buffer_cptr_index ..< buffer_cptr_index + (unat (mi_extra_caps mi))])" | "get_extra_cptrs None mi = return []" definition get_extra_cptr :: "obj_ref \ nat \ (cap_ref,'z::state_ext) s_monad" where "get_extra_cptr buffer n \ liftM data_to_cptr (load_word_offs buffer (n + buffer_cptr_index))" text {* This function both looks up the addresses of the additional capabilities and retreives them from the sender's CSpace. *} definition lookup_extra_caps :: "obj_ref \ data option \ message_info \ ((cap \ cslot_ptr) list,'z::state_ext) f_monad" where "lookup_extra_caps thread buffer mi \ doE cptrs \ liftE $ get_extra_cptrs buffer mi; mapME (\cptr. cap_fault_on_failure (of_bl cptr) False $ lookup_cap_and_slot thread cptr) cptrs odE" text {* Capability transfers. Capabilities passed along with a message are split into two groups. Capabilities to the same endpoint as the message is passed through are not copied. Their badges are unwrapped and stored in the receiver's message buffer instead. Other capabilities are copied into the given slots. Capability unwrapping allows a client to efficiently demonstrate to a server that it possesses authority to two or more services that server provides. *} definition set_extra_badge :: "obj_ref \ machine_word \ nat \ (unit,'z::state_ext) s_monad" where "set_extra_badge buffer badge n \ store_word_offs buffer (buffer_cptr_index + n) badge" type_synonym transfer_caps_data = "(cap \ cslot_ptr) list \ cslot_ptr list" fun transfer_caps_loop :: "obj_ref option \ obj_ref \ nat \ (cap \ cslot_ptr) list \ cslot_ptr list \ message_info \ (message_info,'z::state_ext) s_monad" where "transfer_caps_loop ep rcv_buffer n [] slots mi = return (MI (mi_length mi) (of_nat n) (mi_caps_unwrapped mi) (mi_label mi))" | "transfer_caps_loop ep rcv_buffer n ((cap, slot) # morecaps) slots mi = const_on_failure (MI (mi_length mi) (of_nat n) (mi_caps_unwrapped mi) (mi_label mi)) (doE transfer_rest \ returnOk $ transfer_caps_loop ep rcv_buffer (n + 1) morecaps; if (is_ep_cap cap \ ep = Some (obj_ref_of cap)) then doE liftE $ set_extra_badge rcv_buffer (cap_ep_badge cap) n; liftE $ transfer_rest slots (MI (mi_length mi) (mi_extra_caps mi) (mi_caps_unwrapped mi || (1 << n)) (mi_label mi)) odE else if slots \ [] then doE cap' \ derive_cap slot cap; whenE (cap' = NullCap) $ throwError undefined; liftE $ cap_insert cap' slot (hd slots); liftE $ transfer_rest (tl slots) mi odE else returnOk (MI (mi_length mi) (of_nat n) (mi_caps_unwrapped mi) (mi_label mi)) odE)" definition transfer_caps :: "message_info \ (cap \ cslot_ptr) list \ obj_ref option \ obj_ref \ obj_ref option \ (message_info,'z::state_ext) s_monad" where "transfer_caps info caps endpoint receiver recv_buffer \ do dest_slots \ get_receive_slots receiver recv_buffer; mi' \ return $ MI (mi_length info) 0 0 (mi_label info); case recv_buffer of None \ return mi' | Some receive_buffer \ transfer_caps_loop endpoint receive_buffer 0 caps dest_slots mi' od" section {* Fault Handling *} text {* Threads fault when they attempt to access services that are not backed by any resources. Such a thread is then blocked and a fault messages is sent to its supervisor. When a reply to that message is sent the thread is reactivated. *} text {* Format a message for a given fault type. *} fun make_fault_msg :: "fault \ obj_ref \ (data \ data list,'z::state_ext) s_monad" where "make_fault_msg (CapFault cptr rp lf) thread = (do pc \ as_user thread getRestartPC; return (1, pc # cptr # (if rp then 1 else 0) # msg_from_lookup_failure lf) od)" | "make_fault_msg (VMFault vptr arch) thread = (do pc \ as_user thread getRestartPC; return (2, pc # vptr # arch) od)" | "make_fault_msg (UnknownSyscallException n) thread = (do msg \ as_user thread $ mapM getRegister syscallMessage; return (3, msg @ [n]) od)" | "make_fault_msg (UserException exception code) thread = (do msg \ as_user thread $ mapM getRegister exceptionMessage; return (4, msg @ [exception, code]) od)" text {* React to a fault reply. The reply message is interpreted in a manner that depends on the type of the original fault. For some fault types a thread reconfiguration is performed. This is done entirely to save the fault message recipient an additional system call. This function returns a boolean indicating whether the thread should now be restarted. *} fun handle_fault_reply :: "fault \ obj_ref \ data \ data list \ (bool,'z::state_ext) s_monad" where "handle_fault_reply (CapFault cptr rp lf) thread x y = return True" | "handle_fault_reply (VMFault vptr arch) thread x y = return True" | "handle_fault_reply (UnknownSyscallException n) thread label msg = do as_user thread $ zipWithM_x (\r v. set_register r $ sanitiseRegister r v) syscallMessage msg; return (label = 0) od" | "handle_fault_reply (UserException exception code) thread label msg = do as_user thread $ zipWithM_x (\r v. set_register r $ sanitiseRegister r v) exceptionMessage msg; return (label = 0) od" text {* Transfer a fault message from a faulting thread to its supervisor. *} definition do_fault_transfer :: "data \ obj_ref \ obj_ref \ obj_ref option \ (unit,'z::state_ext) s_monad" where "do_fault_transfer badge sender receiver buf \ do fault \ thread_get tcb_fault sender; f \ (case fault of Some f \ return f | None \ fail); (label, msg) \ make_fault_msg f sender; sent \ set_mrs receiver buf msg; set_message_info receiver $ MI sent 0 0 label; as_user receiver $ set_register badge_register badge od" section {* Synchronous Message Transfers *} text {* Transfer a non-fault message. *} definition do_normal_transfer :: "obj_ref \ obj_ref option \ obj_ref option \ data \ bool \ obj_ref \ obj_ref option \ (unit,'z::state_ext) s_monad" where "do_normal_transfer sender sbuf endpoint badge grant receiver rbuf \ do mi \ get_message_info sender; caps \ if grant then lookup_extra_caps sender sbuf mi K (return []) else return []; mrs_transferred \ copy_mrs sender sbuf receiver rbuf (mi_length mi); mi' \ transfer_caps mi caps endpoint receiver rbuf; set_message_info receiver $ MI mrs_transferred (mi_extra_caps mi') (mi_caps_unwrapped mi') (mi_label mi); as_user receiver $ set_register badge_register badge od" text {* Transfer a message either involving a fault or not. *} definition do_ipc_transfer :: "obj_ref \ obj_ref option \ badge \ bool \ obj_ref \ (unit,'z::state_ext) s_monad" where "do_ipc_transfer sender ep badge grant receiver \ do recv_buffer \ lookup_ipc_buffer True receiver; fault \ thread_get tcb_fault sender; case fault of None \ do send_buffer \ lookup_ipc_buffer False sender; do_normal_transfer sender send_buffer ep badge grant receiver recv_buffer od | Some f \ do_fault_transfer badge sender receiver recv_buffer od" text {* Transfer a reply message and delete the one-use Reply capability. *} definition do_reply_transfer :: "obj_ref \ obj_ref \ cslot_ptr \ (unit,'z::state_ext) s_monad" where "do_reply_transfer sender receiver slot \ do state \ get_thread_state receiver; assert (state = BlockedOnReply); fault \ thread_get tcb_fault receiver; case fault of None \ do do_ipc_transfer sender None 0 True receiver; cap_delete_one slot; set_thread_state receiver Running; do_extended_op (attempt_switch_to receiver) od | Some f \ do cap_delete_one slot; mi \ get_message_info sender; buf \ lookup_ipc_buffer False sender; mrs \ get_mrs sender buf mi; restart \ handle_fault_reply f receiver (mi_label mi) mrs; thread_set (\tcb. tcb \ tcb_fault := None \) receiver; set_thread_state receiver (if restart then Restart else Inactive); when restart $ do_extended_op (attempt_switch_to receiver) od od" text {* This function transfers a reply message to a thread when that message is generated by a kernel service. *} definition reply_from_kernel :: "obj_ref \ (data \ data list) \ (unit,'z::state_ext) s_monad" where "reply_from_kernel thread x \ do (label, msg) \ return x; buf \ lookup_ipc_buffer True thread; as_user thread $ set_register badge_register 0; len \ set_mrs thread buf msg; set_message_info thread $ MI len 0 0 label od" text {* Install a one-use Reply capability. *} definition setup_caller_cap :: "obj_ref \ obj_ref \ (unit,'z::state_ext) s_monad" where "setup_caller_cap sender receiver \ do set_thread_state sender BlockedOnReply; cap_insert (ReplyCap sender False) (sender, tcb_cnode_index 2) (receiver, tcb_cnode_index 3) od" text {* Handle a message send operation performed on an endpoint by a thread. If a receiver is waiting then transfer the message. If no receiver is available and the thread is willing to block waiting to send then put it in the endpoint sending queue. *} definition send_ipc :: "bool \ bool \ badge \ bool \ obj_ref \ obj_ref \ (unit,'z::state_ext) s_monad" where "send_ipc block call badge can_grant thread epptr \ do ep \ get_endpoint epptr; case (ep, block) of (IdleEP, True) \ do set_thread_state thread (BlockedOnSend epptr \ sender_badge = badge, sender_can_grant = can_grant, sender_is_call = call \); set_endpoint epptr $ SendEP [thread] od | (SendEP queue, True) \ do set_thread_state thread (BlockedOnSend epptr \ sender_badge = badge, sender_can_grant = can_grant, sender_is_call = call\); set_endpoint epptr $ SendEP (queue @ [thread]) od | (IdleEP, False) \ return () | (SendEP queue, False) \ return () | (RecvEP (dest # queue), _) \ do set_endpoint epptr $ (case queue of [] \ IdleEP | _ \ RecvEP queue); recv_state \ get_thread_state dest; case recv_state of (BlockedOnReceive x) \ do_ipc_transfer thread (Some epptr) badge can_grant dest | _ \ fail; set_thread_state dest Running; do_extended_op (attempt_switch_to dest); fault \ thread_get tcb_fault thread; when (call \ fault \ None) $ if can_grant then setup_caller_cap thread dest else set_thread_state thread Inactive od | (RecvEP [], _) \ fail od" text {* Handle a message receive operation performed on an endpoint by a thread. If a sender is waiting then transfer the message, otherwise put the thread in the endpoint receiving queue. *} definition isActive :: "notification \ bool" where "isActive ntfn \ case ntfn_obj ntfn of ActiveNtfn _ \ True | _ \ False" text{* Helper function for performing \emph{signal} when receiving on a normal endpoint *} definition complete_signal :: "obj_ref \ obj_ref \ (unit,'z::state_ext) s_monad" where "complete_signal ntfnptr tcb \ do ntfn \ get_notification ntfnptr; case ntfn_obj ntfn of ActiveNtfn badge \ do as_user tcb $ set_register badge_register badge; set_notification ntfnptr $ ntfn_set_obj ntfn IdleNtfn od | _ \ fail od" definition do_nbrecv_failed_transfer :: "obj_ref \ (unit,'z::state_ext) s_monad" where "do_nbrecv_failed_transfer thread = do as_user thread $ set_register badge_register 0; return () od" definition receive_ipc :: "obj_ref \ cap \ bool \ (unit,'z::state_ext) s_monad" where "receive_ipc thread cap is_blocking \ do (epptr,rights) \ (case cap of EndpointCap ref badge rights \ return (ref,rights) | _ \ fail); ep \ get_endpoint epptr; ntfnptr \ get_bound_notification thread; ntfn \ case_option (return default_notification) get_notification ntfnptr; if (ntfnptr \ None \ isActive ntfn) then complete_signal (the ntfnptr) thread else case ep of IdleEP \ case is_blocking of True \ do set_thread_state thread (BlockedOnReceive epptr); set_endpoint epptr (RecvEP [thread]) od | False \ do_nbrecv_failed_transfer thread | RecvEP queue \ case is_blocking of True \ do set_thread_state thread (BlockedOnReceive epptr); set_endpoint epptr (RecvEP (queue @ [thread])) od | False \ do_nbrecv_failed_transfer thread | SendEP q \ do assert (q \ []); queue \ return $ tl q; sender \ return $ hd q; set_endpoint epptr $ (case queue of [] \ IdleEP | _ \ SendEP queue); sender_state \ get_thread_state sender; data \ (case sender_state of BlockedOnSend ref data \ return data | _ \ fail); do_ipc_transfer sender (Some epptr) (sender_badge data) (sender_can_grant data) thread; fault \ thread_get tcb_fault sender; if ((sender_is_call data) \ (fault \ None)) then if sender_can_grant data then setup_caller_cap sender thread else set_thread_state sender Inactive else do set_thread_state sender Running; do_extended_op (switch_if_required_to sender) od od od" section {* Asynchronous Message Transfers *} text {* Helper function to handle a signal operation in the case where a receiver is waiting. *} definition update_waiting_ntfn :: "obj_ref \ obj_ref list \ obj_ref option \ badge \ (unit,'z::state_ext) s_monad" where "update_waiting_ntfn ntfnptr queue bound_tcb badge \ do assert (queue \ []); (dest,rest) \ return $ (hd queue, tl queue); set_notification ntfnptr $ \ ntfn_obj = (case rest of [] \ IdleNtfn | _ \ WaitingNtfn rest), ntfn_bound_tcb = bound_tcb \; set_thread_state dest Running; as_user dest $ set_register badge_register badge; do_extended_op (switch_if_required_to dest) od" text {* Handle a message send operation performed on a notification object. If a receiver is waiting then transfer the message, otherwise combine the new message with whatever message is currently waiting. *} (* helper function for checking if thread is blocked *) definition receive_blocked :: "thread_state \ bool" where "receive_blocked st \ case st of BlockedOnReceive _ \ True | _ \ False" definition send_signal :: "obj_ref \ badge \ (unit,'z::state_ext) s_monad" where "send_signal ntfnptr badge \ do ntfn \ get_notification ntfnptr; case (ntfn_obj ntfn, ntfn_bound_tcb ntfn) of (IdleNtfn, Some tcb) \ do st \ get_thread_state tcb; if (receive_blocked st) then do cancel_ipc tcb; set_thread_state tcb Running; as_user tcb $ set_register badge_register badge; do_extended_op (switch_if_required_to tcb) od else set_notification ntfnptr $ ntfn_set_obj ntfn (ActiveNtfn badge) od | (IdleNtfn, None) \ set_notification ntfnptr $ ntfn_set_obj ntfn (ActiveNtfn badge) | (WaitingNtfn queue, bound_tcb) \ update_waiting_ntfn ntfnptr queue bound_tcb badge | (ActiveNtfn badge', _) \ set_notification ntfnptr $ ntfn_set_obj ntfn $ ActiveNtfn (combine_ntfn_badges badge badge') od" text {* Handle a receive operation performed on a notification object by a thread. If a message is waiting then perform the transfer, otherwise put the thread in the endpoint's receiving queue. *} definition receive_signal :: "obj_ref \ cap \ bool \ (unit,'z::state_ext) s_monad" where "receive_signal thread cap is_blocking \ do ntfnptr \ case cap of NotificationCap ntfnptr badge rights \ return ntfnptr | _ \ fail; ntfn \ get_notification ntfnptr; case ntfn_obj ntfn of IdleNtfn \ case is_blocking of True \ do set_thread_state thread (BlockedOnNotification ntfnptr); set_notification ntfnptr $ ntfn_set_obj ntfn $ WaitingNtfn ([thread]) od | False \ do_nbrecv_failed_transfer thread | WaitingNtfn queue \ case is_blocking of True \ do set_thread_state thread (BlockedOnNotification ntfnptr); set_notification ntfnptr $ ntfn_set_obj ntfn $ WaitingNtfn (queue @ [thread]) od | False \ do_nbrecv_failed_transfer thread | ActiveNtfn badge \ do as_user thread $ set_register badge_register badge; set_notification ntfnptr $ ntfn_set_obj ntfn IdleNtfn od od" section {* Sending Fault Messages *} text {* When a thread encounters a fault, retreive its fault handler capability and send a fault message. *} definition send_fault_ipc :: "obj_ref \ fault \ (unit,'z::state_ext) f_monad" where "send_fault_ipc tptr fault \ doE handler_cptr \ liftE $ thread_get tcb_fault_handler tptr; handler_cap \ cap_fault_on_failure (of_bl handler_cptr) False $ lookup_cap tptr handler_cptr; let f = CapFault (of_bl handler_cptr) False (MissingCapability 0) in (case handler_cap of EndpointCap ref badge rights \ if AllowSend \ rights \ AllowGrant \ rights then liftE $ (do thread_set (\tcb. tcb \ tcb_fault := Some fault \) tptr; send_ipc True False (cap_ep_badge handler_cap) True tptr (cap_ep_ptr handler_cap) od) else throwError f | _ \ throwError f) odE" text {* If a fault message cannot be sent then leave the thread inactive. *} definition handle_double_fault :: "obj_ref \ fault \ fault \ (unit,'z::state_ext) s_monad" where "handle_double_fault tptr ex1 ex2 \ set_thread_state tptr Inactive" text {* Handle a thread fault by sending a fault message if possible. *} definition handle_fault :: "obj_ref \ fault \ (unit,'z::state_ext) s_monad" where "handle_fault thread ex \ do _ \ gets_the $ get_tcb thread; send_fault_ipc thread ex handle_double_fault thread ex; return () od" end