(* * 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) *) theory Decode_D imports Asid_D CNode_D Interrupt_D PageTable_D Tcb_D Untyped_D begin definition get_cnode_intent :: "cdl_intent \ cdl_cnode_intent option" where "get_cnode_intent intent \ case intent of CNodeIntent x \ Some x | _ \ None" definition get_tcb_intent :: "cdl_intent \ cdl_tcb_intent option" where "get_tcb_intent intent \ case intent of TcbIntent x \ Some x | _ \ None" definition get_irq_control_intent :: "cdl_intent \ cdl_irq_control_intent option" where "get_irq_control_intent intent \ case intent of IrqControlIntent x \ Some x | _ \ None" definition get_irq_handler_intent :: "cdl_intent \ cdl_irq_handler_intent option" where "get_irq_handler_intent intent \ case intent of IrqHandlerIntent x \ Some x | _ \ None" definition get_asid_pool_intent :: "cdl_intent \ cdl_asid_pool_intent option" where "get_asid_pool_intent intent \ case intent of AsidPoolIntent x \ Some x | _ \ None" definition get_asid_control_intent :: "cdl_intent \ cdl_asid_control_intent option" where "get_asid_control_intent intent \ case intent of AsidControlIntent x \ Some x | _ \ None" definition get_page_intent :: "cdl_intent \ cdl_page_intent option" where "get_page_intent intent \ case intent of PageIntent x \ Some x | _ \ None" definition get_page_table_intent :: "cdl_intent \ cdl_page_table_intent option" where "get_page_table_intent intent \ case intent of PageTableIntent x \ Some x | _ \ None" definition get_page_directory_intent :: "cdl_intent \ cdl_page_directory_intent option" where "get_page_directory_intent intent \ case intent of PageDirectoryIntent x \ Some x | _ \ None" definition get_untyped_intent :: "cdl_intent \ cdl_untyped_intent option" where "get_untyped_intent intent \ case intent of UntypedIntent x \ Some x | _ \ None" definition get_domain_intent :: "cdl_intent \ cdl_domain_intent option" where "get_domain_intent intent \ case intent of DomainIntent x \ Some x | _ \ None" (* * Decode and validate the given intent, turning it into an * invocation. *) definition decode_invocation :: "cdl_cap \ cdl_cap_ref \ (cdl_cap \ cdl_cap_ref) list \ cdl_intent \ cdl_invocation except_monad" where "decode_invocation invoked_cap invoked_cap_ref caps intent \ case invoked_cap of (* For endpoint-like caps, we always perform an operation, * regardless of the user's actual intent. *) EndpointCap o_id badge rights \ (if Write \ rights then returnOk $ InvokeEndpoint (SyncMessage badge (Grant \ rights) o_id) else throw) | NotificationCap o_id badge rights \ (if Write \ rights then returnOk $ InvokeNotification (Signal badge o_id) else throw) | ReplyCap o_id\ returnOk $ InvokeReply (ReplyMessage o_id invoked_cap_ref) (* * For other operations, we only perform the user's intent * if it matches up with the cap. * * Note that this does not currently match the current * implementation: instead, the user's message will be * decoded into a new (undefined) intent for what the * cap happened to be. I propose modifying labels used to * avoid overlaps between different items so that we can * recognise when the user is invoking the wrong item. *) | CNodeCap _ _ _ _ \ doE cnode_intent \ throw_opt undefined $ get_cnode_intent intent; liftME InvokeCNode $ decode_cnode_invocation invoked_cap invoked_cap_ref caps cnode_intent odE | TcbCap _ \ doE tcb_intent \ throw_opt undefined $ get_tcb_intent intent; liftME InvokeTcb $ decode_tcb_invocation invoked_cap invoked_cap_ref caps tcb_intent odE | IrqControlCap \ doE irq_control_intent \ throw_opt undefined $ get_irq_control_intent intent; liftME InvokeIrqControl $ decode_irq_control_invocation invoked_cap invoked_cap_ref caps irq_control_intent odE | IrqHandlerCap _ \ doE irq_handler_intent \ throw_opt undefined $ get_irq_handler_intent intent; liftME InvokeIrqHandler $ decode_irq_handler_invocation invoked_cap invoked_cap_ref caps irq_handler_intent odE | AsidPoolCap _ _\ doE asid_pool_intent \ throw_opt undefined $ get_asid_pool_intent intent; liftME InvokeAsidPool $ decode_asid_pool_invocation invoked_cap invoked_cap_ref caps asid_pool_intent odE | AsidControlCap \ doE asid_control_intent \ throw_opt undefined $ get_asid_control_intent intent; liftME InvokeAsidControl $ decode_asid_control_invocation invoked_cap invoked_cap_ref caps asid_control_intent odE | UntypedCap _ _ _ \ doE untyped_intent \ throw_opt undefined $ get_untyped_intent intent; liftME InvokeUntyped $ decode_untyped_invocation invoked_cap invoked_cap_ref caps untyped_intent odE | FrameCap _ _ _ _ _ _ \ doE page_intent \ throw_opt undefined $ get_page_intent intent; liftME InvokePage $ decode_page_invocation invoked_cap invoked_cap_ref caps page_intent odE | PageTableCap _ _ _ \ doE page_table_intent \ throw_opt undefined $ get_page_table_intent intent; liftME InvokePageTable $ decode_page_table_invocation invoked_cap invoked_cap_ref caps page_table_intent odE | PageDirectoryCap _ _ _ \ doE page_directory_intent \ throw_opt undefined $ get_page_directory_intent intent; liftME InvokePageDirectory $ decode_page_directory_invocation invoked_cap invoked_cap_ref caps page_directory_intent odE | DomainCap \ doE domain_intent \ throw_opt undefined $ get_domain_intent intent; liftME InvokeDomain $ decode_domain_invocation caps domain_intent odE (* Don't support operations on other types of caps. *) | _ \ throw" end