(* * Copyright 2014, NICTA * * 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(NICTA_GPL) *) (* * Operations on page table objects and frames. *) theory Asid_D imports Invocations_D CSpace_D Untyped_D begin definition decode_asid_control_invocation :: "cdl_cap \ cdl_cap_ref \ (cdl_cap \ cdl_cap_ref) list \ cdl_asid_control_intent \ cdl_asid_control_invocation except_monad" where "decode_asid_control_invocation target target_ref caps intent \ case intent of AsidControlMakePoolIntent index depth \ doE base \ liftE $ select {x. x < 2 ^ asid_high_bits}; (* Fetch the untyped item, and ensure it is valid. *) (untyped_cap, untyped_cap_ref) \ throw_on_none $ get_index caps 0; (case untyped_cap of UntypedCap s _ \ returnOk () | _ \ throw); ensure_no_children untyped_cap_ref; (* Fetch the slot we plan to put the generated cap into. *) (cspace_cap, _) \ throw_on_none $ get_index caps 1; target_slot \ lookup_slot_for_cnode_op cspace_cap index (unat depth); ensure_empty target_slot; returnOk $ MakePool (set_available_range untyped_cap {}) untyped_cap_ref (cap_objects untyped_cap) target_slot base odE \ throw" definition decode_asid_pool_invocation :: "cdl_cap \ cdl_cap_ref \ (cdl_cap \ cdl_cap_ref) list \ cdl_asid_pool_intent \ cdl_asid_pool_invocation except_monad" where "decode_asid_pool_invocation target target_ref caps intent \ case intent of AsidPoolAssignIntent \ doE (pd_cap, pd_cap_ref) \ throw_on_none $ get_index caps 0; (case pd_cap of PageDirectoryCap _ _ _ \ returnOk () | _ \ throw); base \ (case target of AsidPoolCap p base \ returnOk $ base | _ \ throw); offset \ liftE $ select {x. x < 2 ^ asid_low_bits}; returnOk $ Assign (base, offset) pd_cap_ref (cap_object target, offset) odE \ throw" definition invoke_asid_control :: "cdl_asid_control_invocation \ unit k_monad" where "invoke_asid_control params \ case params of MakePool untyped_cap untyped_cap_ref untyped_covers target_slot base \ do (* Untype the region. A choice may be made about whether to detype objects with Untyped addresses. *) modify (detype untyped_covers); set_cap untyped_cap_ref untyped_cap; targets \ generate_object_ids 1 AsidPoolType untyped_covers; (* Retype the region. *) retype_region 0 AsidPoolType targets; assert (targets \ []); (* Insert the cap. *) frame \ return $ pick (hd targets); insert_cap_child (AsidPoolCap frame base) untyped_cap_ref target_slot; (* Update the asid table. *) asid_table \ gets cdl_asid_table; asid_table' \ return $ asid_table (base \ AsidPoolCap frame 0); modify (\s. s \cdl_asid_table := asid_table'\) od" definition invoke_asid_pool :: "cdl_asid_pool_invocation \ unit k_monad" where "invoke_asid_pool params \ case params of Assign asid pd_cap_ref ap_target_slot \ do pd_cap \ get_cap pd_cap_ref; case pd_cap of PageDirectoryCap pd_id _ _ \ do set_cap pd_cap_ref (PageDirectoryCap pd_id Real (Some asid)); set_cap ap_target_slot (PageDirectoryCap pd_id Fake None) od | _ \ fail od" end