-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathIntents_D.thy
205 lines (177 loc) · 7.21 KB
/
Intents_D.thy
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
(*
* 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)
*)
(*
* This file contains user "intents".
*
* Such intents attempt to capture the semantics of an operation the
* user is attempting to perform, without having to worry about how the
* operation is actually encoded within their message registers.
*
* There is a one-to-one mapping between the following intents and the
* invocations made available to userspace. There is not quite
* a one-to-one mapping between these intents and the invocations listed
* in Invocations_D, as some of these intents are multiplexed onto
* a single invocation when being validated.
*
* Caps required by the intents are not stored in the intent themselves,
* but passed seperately in when required. In some sense, the Intent
* is the "data" part of an invocation, but not the "caps" part of it.
*)
theory Intents_D
imports
"../../lib/$L4V_ARCH/WordSetup"
"../abstract/CapRights_A"
begin
(*
* Entities in seL4 have particular rights to kernel objects, which
* affects how entities can interact with those particular objects.
*)
type_synonym cdl_right = rights
(* A user cap pointer. *)
type_synonym cdl_cptr = word32
abbreviation (input) Read ::rights
where "Read \<equiv> AllowRead"
abbreviation (input) Write::rights
where "Write \<equiv> AllowWrite"
abbreviation (input) Grant::rights
where "Grant \<equiv> AllowGrant"
(* Capability data, such as guard information. *)
type_synonym cdl_raw_capdata = word32
(* VM Attributes, such as page cache attributes. *)
type_synonym cdl_raw_vmattrs = word32
(* TCB context, for operations such as write to a thread's registers. *)
type_synonym cdl_raw_usercontext = "word32 list"
(* Kernel objects types. *)
datatype cdl_object_type =
EndpointType
| NotificationType
| TcbType
| CNodeType
| IRQNodeType
| UntypedType
| AsidPoolType
| PageTableType
| PageDirectoryType
| FrameType nat (* size in bits of desired page *)
datatype cdl_cnode_intent =
(* Copy: (target), dest_index, dest_depth, (src_root), src_index, src_depth, rights *)
CNodeCopyIntent word32 word32 word32 word32 "cdl_right set"
(* Mint: (target), dest_index, dest_depth, (src_root), src_index, src_depth, rights, badge *)
| CNodeMintIntent word32 word32 word32 word32 "cdl_right set" cdl_raw_capdata
(* Move: (target), dest_index, dest_depth, (src_root), src_index, src_depth *)
| CNodeMoveIntent word32 word32 word32 word32
(* Mutate: (target), dest_index, dest_depth, (src_root), src_index, src_depth, badge *)
| CNodeMutateIntent word32 word32 word32 word32 cdl_raw_capdata
(* Revoke: (target), index, depth *)
| CNodeRevokeIntent word32 word32
(* Delete: (target), index, depth *)
| CNodeDeleteIntent word32 word32
(* SaveCaller: (target), index, depth *)
| CNodeSaveCallerIntent word32 word32
(* CancelBadgedSends: (target), index, depth *)
| CNodeCancelBadgedSendsIntent word32 word32
(* Rotate: (target), dest_index, dest_depth, (pivot_root), pivot_index, pivot_depth, pivot_badge, (src_root), src_index, src_depth, src_badge *)
| CNodeRotateIntent word32 word32 word32 word32 cdl_raw_capdata word32 word32 cdl_raw_capdata
datatype cdl_tcb_intent =
(* ReadRegisters: (target), suspend_source, arch_flags, count *)
TcbReadRegistersIntent bool word8 word32
(* WriteRegisters: (target), resume_target, arch_flags, count, regs *)
| TcbWriteRegistersIntent bool word8 word32 cdl_raw_usercontext
(* CopyRegisters: (target), (source), suspend_source, resume_target, transfer_frame, transfer_integer, arch_flags *)
| TcbCopyRegistersIntent bool bool bool bool word8
(* Suspend: (target) *)
| TcbSuspendIntent
(* Resume: (target) *)
| TcbResumeIntent
(* Configure: (target), fault_ep, (mcp, priority), (cspace_root), cspace_root_data, (vspace_root), vspace_root_data, buffer, (bufferFrame) *)
| TcbConfigureIntent cdl_cptr "word8 \<times> word8" cdl_raw_capdata cdl_raw_capdata word32
(* SetMCPriority: (target), mcp *)
| TcbSetMCPriorityIntent word8
(* SetPriority: (target), priority *)
| TcbSetPriorityIntent word8
(* SetIPCBuffer: (target), buffer, (bufferFrame) *)
| TcbSetIPCBufferIntent word32
(* SetSpace: (target), fault_ep, (cspace_root), cspace_root_data, (vspace_root), vspace_root_data *)
| TcbSetSpaceIntent word32 cdl_raw_capdata cdl_raw_capdata
(* BindNTFN: (target), (ntfn) *)
| TcbBindNTFNIntent
(* UnbindNTFN: (target) *)
| TcbUnbindNTFNIntent
datatype cdl_untyped_intent =
(* Retype: (target), (do_reset), type, size_bits, (root), node_index, node_depth, node_offset, node_window, has_children *)
UntypedRetypeIntent cdl_object_type word32 word32 word32 word32 word32
datatype cdl_irq_handler_intent =
(* Ack: (target) *)
IrqHandlerAckIntent
(* SetEndpoint: (target), (endpoint) *)
| IrqHandlerSetEndpointIntent
(* Clear: (target) *)
| IrqHandlerClearIntent
datatype cdl_irq_control_intent =
(* IssueIrqHandler: (target), irq, (root), index, depth *)
IrqControlIssueIrqHandlerIntent "10 word" word32 word32
(* InterruptControl *)
| IrqControlArchIrqControlIntent
datatype cdl_page_table_intent =
(* Map: (target), (pd), vaddr, attr *)
PageTableMapIntent word32 cdl_raw_vmattrs
| PageTableUnmapIntent
datatype cdl_page_intent =
(* Map: (target), (pd), vaddr, rights, attr *)
PageMapIntent word32 "cdl_right set" cdl_raw_vmattrs
(*
* Remap: (target), (pd), rights, attr
*
* Note that Remap differs from the abstract spec. This is to
* prevent Remap() from taking place after all rights to the PD
* are long-gone.
*)
| PageRemapIntent"cdl_right set" cdl_raw_vmattrs
(* Unmap: (target) *)
| PageUnmapIntent
(* FlushCaches: (target) *)
| PageFlushCachesIntent
(* GetAddress *)
| PageGetAddressIntent
datatype cdl_page_directory_intent =
PageDirectoryFlushIntent
| PageDirectoryNothingIntent
datatype cdl_asid_control_intent =
(* MakePool: (target), (untyped), (root), index, depth *)
AsidControlMakePoolIntent word32 word32
datatype cdl_asid_pool_intent =
(* Assign: (target), (vroot) *)
AsidPoolAssignIntent
datatype cdl_notification_intent =
SendSignalIntent word32
datatype cdl_endpoint_intent =
SendMessageIntent "cdl_cptr list"
datatype cdl_domain_intent = DomainSetIntent word8
datatype cdl_intent =
CNodeIntent cdl_cnode_intent
| TcbIntent cdl_tcb_intent
| UntypedIntent cdl_untyped_intent
| IrqHandlerIntent cdl_irq_handler_intent
| IrqControlIntent cdl_irq_control_intent
| PageTableIntent cdl_page_table_intent
| PageIntent cdl_page_intent
| PageDirectoryIntent cdl_page_directory_intent
| AsidControlIntent cdl_asid_control_intent
| AsidPoolIntent cdl_asid_pool_intent
| NotificationIntent cdl_notification_intent
| EndpointIntent cdl_endpoint_intent
| DomainIntent cdl_domain_intent
record cdl_full_intent =
cdl_intent_op :: "cdl_intent option"
cdl_intent_error :: bool
cdl_intent_cap :: cdl_cptr
cdl_intent_extras :: "cdl_cptr list"
cdl_intent_recv_slot :: "(cdl_cptr \<times> word32 \<times> nat) option"
end