(* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: GPL-2.0-only *) theory Schedule_D imports KHeap_D begin (* * Collect the set of runnable threads in the system. *) definition all_active_tcbs :: "cdl_state \ cdl_object_id set" where "all_active_tcbs state \ {x \ dom (cdl_objects state). \ a. (cdl_objects state) x = Some (Tcb a) \ ( ((cdl_tcb_caps a) tcb_pending_op_slot) = (Some RunningCap) \ ((cdl_tcb_caps a) tcb_pending_op_slot) = (Some RestartCap))}" definition active_tcbs_in_domain :: "word8 \ cdl_state \ cdl_object_id set" where "active_tcbs_in_domain domain state = {x \ dom (cdl_objects state). \ a. (cdl_objects state) x = Some (Tcb a) \ ( ((cdl_tcb_caps a) tcb_pending_op_slot) = (Some RunningCap) \ ((cdl_tcb_caps a) tcb_pending_op_slot) = (Some RestartCap)) \ cdl_tcb_domain a = domain }" (* * Switch to a new thread. *) definition switch_to_thread :: "cdl_object_id option \ unit k_monad" where "switch_to_thread target \ modify (\ t. t\ cdl_current_thread := target \)" definition change_current_domain :: "unit k_monad" where "change_current_domain = do next_domain \ select UNIV; modify (\s. s\ cdl_current_domain := next_domain \) od" (* * Scheduling is fully nondeterministic at this level. *) definition schedule :: "unit k_monad" where "schedule \ do change_current_domain; next_domain \ gets cdl_current_domain; threads \ gets (active_tcbs_in_domain next_domain); next_thread \ select threads; switch_to_thread (Some next_thread) od \ do change_current_domain; switch_to_thread None od" end