(* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: GPL-2.0-only *) (* * Operations on CNodes. *) theory CNode_D imports Invocations_D CSpace_D begin definition has_cancel_send_rights :: "cdl_cap \ bool" where "has_cancel_send_rights cap \ case cap of EndpointCap _ _ R \ R = UNIV | _ \ False" definition decode_cnode_invocation :: "cdl_cap \ cdl_cap_ref \ (cdl_cap \ cdl_cap_ref) list \ cdl_cnode_intent \ cdl_cnode_invocation except_monad" where "decode_cnode_invocation target target_ref caps intent \ case intent of \ \Copy a cap to anther capslot, without modifying the cap.\ CNodeCopyIntent dest_index dest_depth src_index src_depth rights \ doE (src_root, _) \ throw_on_none $ get_index caps 0; dest_slot \ lookup_slot_for_cnode_op target dest_index (unat dest_depth); ensure_empty dest_slot; src_slot \ lookup_slot_for_cnode_op src_root src_index (unat src_depth); src_cap \ liftE $ get_cap src_slot; new_cap \ returnOk $ update_cap_rights (cap_rights src_cap \ rights) src_cap; cap \ derive_cap src_slot new_cap; whenE (cap = cdl_cap.NullCap) throw; returnOk $ InsertCall cap src_slot dest_slot odE \ \Copy a cap to another capslot, possibly modifying the cap.\ | CNodeMintIntent dest_index dest_depth src_index src_depth rights badge \ doE (src_root, _) \ throw_on_none $ get_index caps 0; dest_slot \ lookup_slot_for_cnode_op target dest_index (unat dest_depth); ensure_empty dest_slot; src_slot \ lookup_slot_for_cnode_op src_root src_index (unat src_depth); src_cap \ liftE $ get_cap src_slot; \ \Munge the caps rights/data.\ new_cap \ returnOk $ update_cap_rights (cap_rights src_cap \ rights) src_cap; new_cap' \ liftE $ update_cap_data False badge new_cap; cap \ derive_cap src_slot new_cap'; whenE (cap = cdl_cap.NullCap) throw; returnOk $ InsertCall cap src_slot dest_slot odE \ \Move a cap to another capslot, without modifying the cap.\ | CNodeMoveIntent dest_index dest_depth src_index src_depth \ doE (src_root, _) \ throw_on_none $ get_index caps 0; dest_slot \ lookup_slot_for_cnode_op target dest_index (unat dest_depth); ensure_empty dest_slot; src_slot \ lookup_slot_for_cnode_op src_root src_index (unat src_depth); src_cap \ liftE $ get_cap src_slot; whenE (src_cap = NullCap) throw; returnOk $ MoveCall src_cap src_slot dest_slot odE \ \Move a cap to another capslot, possibly modifying the cap.\ | CNodeMutateIntent dest_index dest_depth src_index src_depth badge \ doE (src_root, _) \ throw_on_none $ get_index caps 0; dest_slot \ lookup_slot_for_cnode_op target dest_index (unat dest_depth); ensure_empty dest_slot; src_slot \ lookup_slot_for_cnode_op src_root src_index (unat src_depth); src_cap \ liftE $ get_cap src_slot; \ \Munge the caps rights/data.\ cap \ liftE $ update_cap_data True badge src_cap; whenE (cap = NullCap) throw; returnOk $ MoveCall cap src_slot dest_slot odE \ \Revoke all CDT children of the given cap.\ | CNodeRevokeIntent index depth \ doE target_slot \ lookup_slot_for_cnode_op target index (unat depth); returnOk $ RevokeCall target_slot odE \ \Delete the given cap, but not its children.\ | CNodeDeleteIntent index depth \ doE target_slot \ lookup_slot_for_cnode_op target index (unat depth); returnOk $ DeleteCall target_slot odE \ \Save the current thread's reply cap into the target slot.\ | CNodeSaveCallerIntent index depth \ doE target_slot \ lookup_slot_for_cnode_op target index (unat depth); ensure_empty target_slot; returnOk $ SaveCall target_slot odE \ \Recycle the target cap.\ | CNodeCancelBadgedSendsIntent index depth \ doE target_slot \ lookup_slot_for_cnode_op target index (unat depth); cap \ liftE $ get_cap target_slot; unlessE (has_cancel_send_rights cap) throw; returnOk $ CancelBadgedSendsCall cap odE \ \Atomically move several caps.\ | CNodeRotateIntent dest_index dest_depth pivot_index pivot_depth pivot_badge src_index src_depth src_badge \ doE pivot_root \ throw_on_none $ get_index caps 0; src_root \ throw_on_none $ get_index caps 1; dest_root \ returnOk $ target; pivot_root \ returnOk $ fst pivot_root; src_root \ returnOk $ fst src_root; dest_slot \ lookup_slot_for_cnode_op dest_root dest_index (unat dest_depth); src_slot \ lookup_slot_for_cnode_op src_root src_index (unat src_depth); pivot_slot \ lookup_slot_for_cnode_op pivot_root pivot_index (unat pivot_depth); whenE (pivot_slot = src_slot \ pivot_slot = dest_slot) throw; unlessE (src_slot = dest_slot) $ ensure_empty dest_slot; src_cap \ liftE $ get_cap src_slot; whenE (src_cap = NullCap) throw; pivot_cap \ liftE $ get_cap pivot_slot; whenE (pivot_cap = NullCap) throw; \ \Munge caps.\ new_src \ liftE $ update_cap_data True src_badge src_cap; new_pivot \ liftE $ update_cap_data True pivot_badge pivot_cap; whenE (new_src = NullCap) throw; whenE (new_pivot = NullCap) throw; returnOk $ RotateCall new_src new_pivot src_slot pivot_slot dest_slot odE " definition invoke_cnode :: "cdl_cnode_invocation \ unit preempt_monad" where "invoke_cnode params \ case params of \ \Insert a new cap.\ InsertCall cap src_slot dest_slot \ liftE $ insert_cap_sibling cap src_slot dest_slot \ insert_cap_child cap src_slot dest_slot \ \Move a cap, possibly modifying it in the process.\ | MoveCall cap src_slot dest_slot \ liftE $ move_cap cap src_slot dest_slot \ \Revoke a cap.\ | RevokeCall src_slot \ revoke_cap src_slot \ \Delete a cap.\ | DeleteCall src_slot \ delete_cap src_slot \ \Atomically move two capabilities.\ | RotateCall cap1 cap2 slot1 slot2 slot3 \ liftE $ if slot1 = slot3 then swap_cap cap1 slot1 cap2 slot2 else do move_cap cap2 slot2 slot3; move_cap cap1 slot1 slot2 od \ \Save a reply cap from the caller's TCB into this CNode.\ | SaveCall dest_slot \ liftE $ do current \ gets_the cdl_current_thread; replycap \ get_cap (current, tcb_caller_slot); when (replycap \ NullCap) $ move_cap replycap (current, tcb_caller_slot) dest_slot od \ \Reset an object into its original state.\ | CancelBadgedSendsCall (EndpointCap ep b _) \ liftE $ when (b \ 0) $ cancel_badged_sends ep b | CancelBadgedSendsCall _ \ fail " end