lh-l4v/spec/abstract/Invocations_A.thy

73 lines
2.1 KiB
Plaintext
Raw Normal View History

2014-07-14 19:32:44 +00:00
(*
* Copyright 2014, General Dynamics C4 Systems
*
* 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(GD_GPL)
*)
(*
Data types for syscall invocations
*)
2015-05-16 09:52:49 +00:00
chapter "Kernel Object Invocations"
2014-07-14 19:32:44 +00:00
theory Invocations_A
imports ArchInvocation_A
begin
text {* These datatypes encode the arguments to the available system calls. *}
datatype cnode_invocation =
InsertCall cap cslot_ptr cslot_ptr
| MoveCall cap cslot_ptr cslot_ptr
| RevokeCall cslot_ptr
| DeleteCall cslot_ptr
| RotateCall cap cap cslot_ptr cslot_ptr cslot_ptr
| SaveCall cslot_ptr
| RecycleCall cslot_ptr
datatype untyped_invocation =
Retype cslot_ptr obj_ref obj_ref apiobject_type nat "cslot_ptr list"
datatype arm_copy_register_sets =
ARMNoExtraRegisters
datatype tcb_invocation =
WriteRegisters word32 bool "word32 list" arm_copy_register_sets
| ReadRegisters word32 bool word32 arm_copy_register_sets
| CopyRegisters word32 word32 bool bool bool bool arm_copy_register_sets
2015-05-16 09:52:49 +00:00
| ThreadControl word32 cslot_ptr
(tc_new_fault_ep: "cap_ref option") (tc_new_priority: "word8 option")
(tc_new_croot: "(cap * cslot_ptr) option") (tc_new_vroot: "(cap * cslot_ptr) option")
(tc_new_buffer: "(vspace_ref * (cap * cslot_ptr) option) option")
2014-07-14 19:32:44 +00:00
| Suspend "word32"
| Resume "word32"
2015-09-02 05:43:39 +00:00
| AsyncEndpointControl "obj_ref" "obj_ref option"
2014-07-14 19:32:44 +00:00
datatype irq_control_invocation =
IRQControl irq cslot_ptr cslot_ptr
| InterruptControl arch_interrupt_control
datatype irq_handler_invocation =
ACKIrq irq
| SetIRQHandler irq cap cslot_ptr
| ClearIRQHandler irq
2014-08-22 06:24:40 +00:00
| SetMode irq bool bool
2014-07-14 19:32:44 +00:00
datatype invocation =
InvokeUntyped untyped_invocation
| InvokeEndpoint obj_ref word32 bool
2015-09-02 05:43:39 +00:00
| InvokeAsyncEndpoint obj_ref word32
2014-07-14 19:32:44 +00:00
| InvokeReply obj_ref cslot_ptr
| InvokeTCB tcb_invocation
| InvokeDomain obj_ref word8
| InvokeCNode cnode_invocation
| InvokeIRQControl irq_control_invocation
| InvokeIRQHandler irq_handler_invocation
| InvokeArchObject arch_invocation
end