(* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: GPL-2.0-only *) (* * Operations on thread control blocks. *) theory Tcb_D imports Invocations_D CSpace_D begin definition cdl_update_cnode_cap_data :: "cdl_cap \ word32 \ cdl_cap" where "cdl_update_cnode_cap_data cap data \ case cap of cdl_cap.CNodeCap oid _ _ sz \ if data\0 then (let reserved_bits = 3; guard_bits = 18; guard_size_bits = 5; new_guard_size = unat ((data >> reserved_bits) && mask guard_size_bits); new_guard = (data >> reserved_bits + guard_size_bits) && mask (min (unat ((data >> reserved_bits) && mask guard_size_bits)) guard_bits) in CNodeCap oid new_guard new_guard_size sz) else cap | _ \ cap" definition cdl_same_arch_obj_as :: "cdl_cap \ cdl_cap \ bool" where "cdl_same_arch_obj_as capa capb \ case capa of AsidPoolCap x _ \ ( case capb of AsidPoolCap y _ \ y = x | _ \ False) | AsidControlCap \ ( case capb of AsidControlCap \ True | _ \ False) | FrameCap dev ra _ sa _ _ \ ( case capb of FrameCap dev' rb _ sb _ _ \ rb = ra \ sb = sa \ dev = dev' | _ \ False) | cdl_cap.PageTableCap a _ _ \ ( case capb of cdl_cap.PageTableCap b _ _ \ b = a | _ \ False) | cdl_cap.PageDirectoryCap a _ _ \ ( case capb of cdl_cap.PageDirectoryCap b _ _ \ b = a | _ \ False) | _ \ False" definition decode_tcb_invocation :: "cdl_cap \ cdl_cap_ref \ (cdl_cap \ cdl_cap_ref) list \ cdl_tcb_intent \ cdl_tcb_invocation except_monad" where "decode_tcb_invocation target slot caps intent \ case intent of \ \Read another thread's registers.\ TcbReadRegistersIntent suspend flags count \ returnOk (ReadRegisters (cap_object target) suspend 0 0) \ throw \ \Write another thread's registers.\ | TcbWriteRegistersIntent resume flags count regs \ returnOk (WriteRegisters (cap_object target) resume [0] 0) \ throw \ \Copy registers from one thread to another.\ | TcbCopyRegistersIntent suspend_source resume_target f1 f2 f3 \ doE (source_cap, _) \ throw_on_none $ get_index caps 0; source_tcb \ ( case source_cap of TcbCap x \ returnOk x | _ \ throw); target_tcb \ returnOk $ cap_object target; returnOk (CopyRegisters target_tcb source_tcb suspend_source resume_target f1 f2 0) odE \ throw \ \Suspend the target thread.\ | TcbSuspendIntent \ returnOk (Suspend (cap_object target)) \ throw \ \Resume the target thread.\ | TcbResumeIntent \ returnOk (Resume (cap_object target)) \ throw \ \Configure: target, fault_ep, mcp, priority, cspace_root_data, vspace_root_data, buffer\ | TcbConfigureIntent fault_ep cspace_root_data vspace_root_data buffer \ doE cspace_root \ throw_on_none $ get_index caps 0; vspace_root \ throw_on_none $ get_index caps 1; buffer_frame \ throw_on_none $ get_index caps 2; cspace_root_cap_ref \ returnOk $ (cdl_update_cnode_cap_data (fst cspace_root) cspace_root_data,snd cspace_root); vspace_root_cap_ref \ returnOk $ vspace_root; buffer_frame_opt \ returnOk $ (if (buffer \ 0) then Some (reset_mem_mapping (fst buffer_frame), snd buffer_frame) else None); returnOk (ThreadControl (cap_object target) slot (Some fault_ep) (Some cspace_root_cap_ref) (Some vspace_root_cap_ref) (buffer_frame_opt)) odE \ throw \ \Modify a thread's maximum control priority.\ | TcbSetMCPriorityIntent mcp \ doE auth_cap \ throw_on_none $ get_index caps 0; returnOk (ThreadControl (cap_object target) slot None None None None) odE \ throw \ \Modify a thread's priority.\ | TcbSetPriorityIntent priority \ doE auth_cap \ throw_on_none $ get_index caps 0; returnOk (ThreadControl (cap_object target) slot None None None None) odE \ throw \ \Modify a thread's mcp and priority at the same time.\ | TcbSetSchedParamsIntent mcp priority \ doE auth_cap \ throw_on_none $ get_index caps 0; returnOk (ThreadControl (cap_object target) slot None None None None) odE \ throw \ \Modify a thread's IPC buffer.\ | TcbSetIPCBufferIntent buffer \ doE buffer_frame \ throw_on_none $ get_index caps 0; buffer_frame_opt \ returnOk $ (if (buffer \ 0) then Some (reset_mem_mapping (fst buffer_frame), snd buffer_frame) else None); returnOk (ThreadControl (cap_object target) slot None None None buffer_frame_opt) odE \ throw \ \Update the various spaces (CSpace/VSpace) of a thread.\ | TcbSetSpaceIntent fault_ep cspace_root_data vspace_root_data \ doE cspace_root \ throw_on_none $ get_index caps 0; vspace_root \ throw_on_none $ get_index caps 1; cspace_root_cap_ref \ returnOk $ (cdl_update_cnode_cap_data (fst cspace_root) cspace_root_data,snd cspace_root); vspace_root_cap_ref \ returnOk $ vspace_root; returnOk (ThreadControl (cap_object target) slot (Some fault_ep) (Some cspace_root_cap_ref) (Some vspace_root_cap_ref) None) odE \ throw | TcbBindNTFNIntent \ doE (ntfn_cap, _) \ throw_on_none $ get_index caps 0; returnOk (NotificationControl (cap_object target) (Some (cap_object ntfn_cap))) odE \ throw | TcbUnbindNTFNIntent \ returnOk (NotificationControl (cap_object target) None) \ throw | TCBSetTLSBaseIntent \ returnOk (SetTLSBase (cap_object target)) \ throw " (* Delete the given slot of a TCB. *) definition tcb_empty_thread_slot :: "cdl_object_id \ cdl_cnode_index \ unit preempt_monad" where "tcb_empty_thread_slot target_tcb target_slot \ doE cap \ liftE $ get_cap (target_tcb,target_slot); whenE (cap \ NullCap) $ delete_cap (target_tcb, target_slot) odE" (* Update the given slot of a TCB with a new cap, delete the previous * capability that was in the slot. *) definition tcb_update_thread_slot :: "cdl_object_id \ cdl_cap_ref \ cdl_cnode_index \ (cdl_cap \ cdl_cap_ref) \ unit preempt_monad" where "tcb_update_thread_slot target_tcb tcb_cap_slot target_slot pcap \ liftE (do thread_cap \ get_cap tcb_cap_slot; when (thread_cap = TcbCap target_tcb) (insert_cap_child (fst pcap) (snd pcap) (target_tcb, target_slot) \ insert_cap_sibling (fst pcap) (snd pcap) (target_tcb,target_slot)) od)" (* Update a thread's CSpace root. *) definition tcb_update_cspace_root :: "cdl_object_id \ cdl_cap_ref \ cdl_cap \ cdl_cap_ref \ unit preempt_monad" where "tcb_update_cspace_root target_tcb tcb_cap_ref croot \ doE tcb_empty_thread_slot target_tcb tcb_cspace_slot; src_cap \ liftE $ get_cap (snd croot); whenE (is_cnode_cap src_cap \ (cap_object src_cap = cap_object (fst croot))) $ tcb_update_thread_slot target_tcb tcb_cap_ref tcb_cspace_slot croot odE" (* Update a thread's VSpace root. *) definition tcb_update_vspace_root :: "cdl_object_id \ cdl_cap_ref \ (cdl_cap \ cdl_cap_ref) \ unit preempt_monad" where "tcb_update_vspace_root target_tcb tcb_cap_ref vroot \ doE tcb_empty_thread_slot target_tcb tcb_vspace_slot; src_cap \ liftE $ get_cap (snd vroot); whenE (cdl_same_arch_obj_as (fst vroot) src_cap) $ tcb_update_thread_slot target_tcb tcb_cap_ref tcb_vspace_slot vroot odE" (* Modify the TCB's intent to indicate an error during decode. *) definition mark_tcb_intent_error :: "cdl_object_id \ bool \ unit k_monad" where "mark_tcb_intent_error target_tcb has_error \ update_thread target_tcb (\t. (t\cdl_tcb_intent := (cdl_tcb_intent t)\cdl_intent_error := has_error\\))" (* Update a thread's IPC buffer. *) definition tcb_update_ipc_buffer :: "cdl_object_id \ cdl_cap_ref \ (cdl_cap \ cdl_cap_ref) \ unit preempt_monad" where "tcb_update_ipc_buffer target_tcb tcb_cap_ref ipc_buffer \ doE tcb_empty_thread_slot target_tcb tcb_ipcbuffer_slot; liftE $ corrupt_tcb_intent target_tcb; src_cap \ liftE $ get_cap (snd ipc_buffer); whenE (cdl_same_arch_obj_as (fst ipc_buffer) src_cap) $ tcb_update_thread_slot target_tcb tcb_cap_ref tcb_ipcbuffer_slot ipc_buffer odE " (* Resume a thread, aborting any pending operation, and revoking * any incoming reply caps. *) definition restart :: "cdl_object_id \ unit k_monad" where "restart target_tcb \ do cap \ KHeap_D.get_cap (target_tcb,tcb_pending_op_slot); when (cap \ RestartCap \ cap\ RunningCap) (do CSpace_D.cancel_ipc target_tcb; KHeap_D.set_cap (target_tcb,tcb_replycap_slot) (cdl_cap.MasterReplyCap target_tcb); KHeap_D.set_cap (target_tcb,tcb_pending_op_slot) (cdl_cap.RestartCap) od) od" (* Suspend a thread, aborting any pending operation, and revoking * any incoming reply caps. *) definition suspend :: "cdl_object_id \ unit k_monad" where "suspend target_tcb \ CSpace_D.cancel_ipc target_tcb >>= K (KHeap_D.set_cap (target_tcb,tcb_pending_op_slot) cdl_cap.NullCap)" definition bind_notification :: "cdl_object_id \ cdl_object_id \ unit k_monad" where "bind_notification tcb_id ntfn_id \ set_cap (tcb_id, tcb_boundntfn_slot) (BoundNotificationCap ntfn_id)" definition invoke_tcb :: "cdl_tcb_invocation \ unit preempt_monad" where "invoke_tcb params \ case params of \ \Modify a thread's registers.\ WriteRegisters target_tcb resume _ _ \ liftE $ do corrupt_tcb_intent target_tcb; when resume $ restart target_tcb od \ \Read a thread's registers.\ | ReadRegisters src_tcb _ _ _ \ liftE $ suspend src_tcb \ return () \ \Copy registers from one thread to another\ | CopyRegisters target_tcb source_tcb _ _ _ _ _ \ liftE $ do suspend source_tcb \ return (); restart target_tcb \ return (); corrupt_tcb_intent target_tcb od \ \Suspend this thread.\ | Suspend target_tcb \ liftE $ suspend target_tcb \ return () \ \Resume this thread.\ | Resume target_tcb \ liftE $ restart target_tcb \ \Update a thread's options.\ | ThreadControl target_tcb tcb_cap_slot faultep croot vroot ipc_buffer \ doE case faultep of Some x \ liftE $ update_thread target_tcb (\tcb. tcb\cdl_tcb_fault_endpoint := x\) | None \ returnOk (); \ \Possibly update CSpace\ case croot of Some x \ tcb_update_cspace_root target_tcb tcb_cap_slot x | None \ returnOk (); \ \Possibly update VSpace\ case vroot of Some x \ tcb_update_vspace_root target_tcb tcb_cap_slot x | None \ returnOk (); \ \Possibly update Ipc Buffer\ case ipc_buffer of Some x \ tcb_update_ipc_buffer target_tcb tcb_cap_slot x | None \ (returnOk () \ (doE tcb_empty_thread_slot target_tcb tcb_ipcbuffer_slot; liftE $ corrupt_tcb_intent target_tcb odE)) odE | NotificationControl tcb ntfn \ liftE $ (case ntfn of Some ntfn_id \ bind_notification tcb ntfn_id | None \ unbind_notification tcb) | SetTLSBase tcb \ liftE $ corrupt_tcb_intent tcb" definition decode_domain_invocation :: "(cdl_cap \ cdl_cap_ref) list \ cdl_domain_intent \ cdl_domain_invocation except_monad" where "decode_domain_invocation caps intent \ case intent of DomainSetIntent d \ returnOk (SetDomain (cap_object (fst (hd caps))) d) \ throw" definition set_domain :: "cdl_object_id \ word8 \ unit k_monad" where "set_domain tcb d \ update_thread tcb (\t. (t\cdl_tcb_domain := d \))" definition invoke_domain :: "cdl_domain_invocation \ unit preempt_monad" where "invoke_domain params \ case params of SetDomain tcb d \ liftE $ set_domain tcb d" end