@@ -97,6 +97,29 @@ lemma decode_irq_control_issue_irq_rv:
97
97
apply sep_solve
98
98
done
99
99
100
+ lemma invoke_irq_control_issue_sgi_wp :
101
+ "\<lbrace> <dest_slot \<mapsto>c cap \<and>* R> \<rbrace>
102
+ invoke_irq_control (ArchIssueIrqHandler (ARMIssueSGISignal irq targets control_slot dest_slot))
103
+ \<lbrace>\<lambda>_. <dest_slot \<mapsto>c SGISignalCap irq targets \<and>* R> \<rbrace>"
104
+ by ( clarsimp simp : invoke_irq_control_def arch_invoke_irq_control_def cap_type_def |
105
+ wp sep_wp : insert_cap_child_wp | sep_solve )+
106
+
107
+ lemma decode_irq_control_issue_sgi_rv :
108
+ "\<lbrace>\<lambda>s. P (ArchIssueIrqHandler
109
+ (ARMIssueSGISignal (ucast irq) (ucast cpus) target_ref
110
+ (cap_object root_cap, offset index r))) s \<and>
111
+ caps = (root_cap, root_ptr) # xs \<and> unat depth \<le> word_bits \<and> 0 < unat depth \<and>
112
+ <\<box> (r, (unat depth)) : root_cap index \<mapsto>u cap \<and>* R> s\<rbrace>
113
+ decode_irq_control_invocation target target_ref caps
114
+ (ArchIrqControlIssueIrqHandlerIntent
115
+ (ARMIssueSGISignalIntent irq cpus index depth))
116
+ \<lbrace>P\<rbrace>, -"
117
+ apply ( clarsimp simp : decode_irq_control_invocation_def arch_decode_irq_control_invocation_def )
118
+ apply ( wp lookup_slot_for_cnode_op_rvu' [ where cap = cap and r = r ] throw_on_none_rv )
119
+ apply ( clarsimp simp : get_index_def )
120
+ apply sep_solve
121
+ done
122
+
100
123
schematic_goal lookup_extra_caps_once_wp :
101
124
"\<lbrace>?P\<rbrace> lookup_extra_caps root_tcb_id [endpoint_cptr] \<lbrace>Q\<rbrace>, \<lbrace>Q'\<rbrace>"
102
125
apply ( clarsimp simp : lookup_extra_caps_def mapME_def sequenceE_def , wp )
@@ -256,6 +279,140 @@ lemma seL4_IRQHandler_IRQControl_Get:
256
279
apply ( clarsimp simp : is_irqcontrol_cap_def ep_related_cap_def )
257
280
done
258
281
282
+ lemma seL4_IssueSGISignal_helper :
283
+ assumes unify : "irq_cap = SGISignalCap (ucast irq) (ucast targets) \<and>
284
+ target_index = offset node_index root_size \<and>
285
+ root_index = offset croot root_size \<and>
286
+ control_index = offset control_cap root_size \<and>
287
+ target_ptr = (cap_object root_cap, target_index) \<and>
288
+ control_ptr = (cap_object root_cap, control_index) \<and>
289
+ root_ptr = (cap_object root_cap, root_index) \<and>
290
+ root_tcb = Tcb t"
291
+ shows
292
+ "\<lbrace>\<guillemotleft>root_tcb_id \<mapsto>f root_tcb \<and>* (root_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
293
+ cap_object root_cap \<mapsto>f CNode (empty_cnode root_size) \<and>*
294
+ (root_tcb_id, tcb_cspace_slot) \<mapsto>c root_cap \<and>* control_ptr \<mapsto>c c_cap \<and>*
295
+ target_ptr \<mapsto>c target_cap \<and>* root_ptr \<mapsto>c root_cap \<and>* R\<guillemotright> and
296
+ K (\<not> ep_related_cap c_cap \<and> one_lvl_lookup root_cap word_bits root_size \<and>
297
+ one_lvl_lookup root_cap (unat node_depth) root_size \<and>
298
+ guard_equal root_cap node_index (unat node_depth) \<and>
299
+ guard_equal root_cap croot word_bits \<and>
300
+ guard_equal root_cap control_cap word_bits \<and>
301
+ guard_equal root_cap node_index word_bits \<and>
302
+ unat node_depth \<le> word_bits \<and> 0 < unat node_depth \<and>
303
+ is_irqcontrol_cap c_cap \<and> is_cnode_cap root_cap \<and> is_cnode_cap root_cap)\<rbrace>
304
+ seL4_IssueSGISignal control_cap irq targets croot node_index node_depth
305
+ \<lbrace>\<lambda>fail. \<guillemotleft>root_tcb_id \<mapsto>f root_tcb \<and>*
306
+ (root_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
307
+ \<comment> \<open>Root CNode\<close>
308
+ cap_object root_cap \<mapsto>f CNode (empty_cnode root_size) \<and>*
309
+ \<comment> \<open>Cap to the root CNode\<close>
310
+ (root_tcb_id, tcb_cspace_slot) \<mapsto>c root_cap \<and>*
311
+ control_ptr \<mapsto>c c_cap \<and>* target_ptr \<mapsto>c irq_cap \<and>*
312
+ root_ptr \<mapsto>c root_cap \<and>* R\<guillemotright>\<rbrace>"
313
+ apply ( simp add : seL4_IssueSGISignal_def sep_state_projection2_def )
314
+ apply ( wp do_kernel_op_pull_back )
315
+ apply ( rule call_kernel_with_intent_allow_error_helper [ where
316
+ check = True and
317
+ Perror = \<top> and
318
+ intent_op = "IrqControlIntent
319
+ (ArchIrqControlIssueIrqHandlerIntent
320
+ (ARMIssueSGISignalIntent irq targets node_index node_depth))" and
321
+ tcb = t and
322
+ intent_cptr = control_cap and
323
+ intent_extra = "[croot]" , simplified ])
324
+ apply clarsimp
325
+ apply ( rule set_cap_wp [ sep_wand_wp ])
326
+ apply ( rule mark_tcb_intent_error_sep_inv )
327
+ apply ( wp corrupt_ipc_buffer_sep_inv )+
328
+ apply ( rule_tac P = "iv = InvokeIrqControl
329
+ (ArchIssueIrqHandler
330
+ (ARMIssueSGISignal (ucast irq) (ucast targets) control_ptr
331
+ (cap_object root_cap,
332
+ offset node_index root_size)))"
333
+ in hoare_gen_asmEx )
334
+ apply ( clarsimp simp : unify )
335
+ apply ( wp invoke_irq_control_issue_sgi_wp [ sep_wand_wp ])
336
+ apply ( wp set_cap_wp [ sep_wand_wp ])
337
+ apply ( rule_tac P = "c=IrqControlCap" in hoare_gen_asmEx , simp )
338
+ apply ( simp add : decode_invocation_def , wp )
339
+ apply ( rule_tac P = "irq_control_intent =
340
+ ArchIrqControlIssueIrqHandlerIntent
341
+ (ARMIssueSGISignalIntent irq targets node_index node_depth)"
342
+ in hoare_gen_asmE , simp )
343
+ apply ( wp decode_irq_control_issue_sgi_rv [ where root_cap = root_cap and
344
+ root_ptr = root_ptr and
345
+ xs = "[]" and r = root_size , simplified ])
346
+ apply ( simp add : throw_opt_def get_irq_control_intent_def )
347
+ apply wp
348
+ apply ( rule lookup_extra_caps_once_wp [ where cap' = root_cap and r = root_size , simplified user_pointer_at_def ])
349
+ apply ( rule lookup_cap_and_slot_rvu [ where cap' = c_cap and r = root_size , simplified user_pointer_at_def ])
350
+ apply ( wp hoare_vcg_imp_lift hoare_vcg_all_lift update_thread_intent_update )
351
+ apply clarsimp
352
+ apply ( rule conjI , fastforce )+
353
+ apply sep_solve
354
+ apply ( intro impI conjI allI ; clarsimp )
355
+ apply ( intro impI conjI allI ; clarsimp simp : unify )
356
+ apply sep_solve
357
+ apply sep_cancel +
358
+ apply ( intro impI conjI allI ; sep_solve )
359
+ apply fastforce
360
+ apply ( metis is_cnode_cap_not_memory_cap not_memory_cap_reset_asid' )
361
+ apply fastforce
362
+ apply ( clarsimp simp : user_pointer_at_def Let_def )
363
+ apply sep_solve
364
+ apply ( drule ( 1 ) reset_cap_asid_control , clarsimp simp : is_irqcontrol_cap_def )
365
+ apply ( clarsimp simp : user_pointer_at_def Let_def )
366
+ apply sep_solve
367
+ apply ( drule reset_cap_asid_ep_related_cap )
368
+ apply clarsimp
369
+ apply ( clarsimp simp : user_pointer_at_def Let_def unify )
370
+ apply sep_solve
371
+ apply ( clarsimp simp : unify )
372
+ apply sep_solve
373
+ done
374
+
375
+ lemma seL4_IssueSGISignal_Get :
376
+ "\<lbrace>\<guillemotleft>root_tcb_id \<mapsto>f root_tcb \<and>*
377
+ (root_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
378
+ (root_tcb_id, tcb_cspace_slot) \<mapsto>c cnode_cap \<and>*
379
+ cnode_id \<mapsto>f CNode (empty_cnode root_size) \<and>*
380
+ (cnode_id, control_index) \<mapsto>c IrqControlCap \<and>*
381
+ (cnode_id, root_index) \<mapsto>c cnode_cap \<and>*
382
+ (cnode_id, target_index) \<mapsto>c target_cap \<and>* R\<guillemotright>
383
+ and K (is_tcb root_tcb \<and>
384
+ is_cnode_cap cnode_cap \<and>
385
+ is_cnode_cap cnode_cap \<and>
386
+ cnode_id = cap_object cnode_cap \<and>
387
+ target_index = offset node_index root_size \<and>
388
+ root_index = offset croot root_size \<and>
389
+ control_index = offset control_cap root_size \<and>
390
+ one_lvl_lookup cnode_cap word_bits root_size \<and>
391
+ one_lvl_lookup cnode_cap (unat node_depth) root_size \<and>
392
+ guard_equal cnode_cap node_index (unat node_depth) \<and>
393
+ guard_equal cnode_cap croot word_bits \<and>
394
+ guard_equal cnode_cap control_cap word_bits \<and>
395
+ guard_equal cnode_cap node_index word_bits \<and>
396
+ unat node_depth \<le> word_bits \<and> 0 < unat node_depth)\<rbrace>
397
+ seL4_IssueSGISignal control_cap irq targets croot node_index node_depth
398
+ \<lbrace>\<lambda>fail.
399
+ \<guillemotleft>root_tcb_id \<mapsto>f root_tcb \<and>*
400
+ (root_tcb_id, tcb_pending_op_slot) \<mapsto>c RunningCap \<and>*
401
+ \<comment> \<open>Cap to the root CNode\<close>
402
+ (root_tcb_id, tcb_cspace_slot) \<mapsto>c cnode_cap \<and>*
403
+ \<comment> \<open>Root CNode\<close>
404
+ cnode_id \<mapsto>f CNode (empty_cnode root_size) \<and>*
405
+ (cnode_id, control_index) \<mapsto>c IrqControlCap \<and>*
406
+ (cnode_id, root_index) \<mapsto>c cnode_cap \<and>*
407
+ (cnode_id, target_index) \<mapsto>c SGISignalCap (ucast irq) (ucast targets) \<and>* R\<guillemotright>\<rbrace>"
408
+ apply ( rule hoare_gen_asm , elim conjE )
409
+ apply ( sep_wp seL4_IssueSGISignal_helper [ where t = "obj_tcb root_tcb" ])
410
+ apply fastforce
411
+ apply clarsimp
412
+ apply ( rule conjI , sep_solve )
413
+ apply ( clarsimp simp : is_irqcontrol_cap_def ep_related_cap_def )
414
+ done
415
+
259
416
lemma seL4_IRQHandler_SetEndpoint_wp_helper :
260
417
assumes unify : "irq_cap = IrqHandlerCap irq \<and>
261
418
offset endpoint_cptr root_size = irq_handler_slot \<and>
0 commit comments