diff --git a/proof/capDL-api/IRQ_DP.thy b/proof/capDL-api/IRQ_DP.thy index 82ae7f64f0f..fee843a28b5 100644 --- a/proof/capDL-api/IRQ_DP.thy +++ b/proof/capDL-api/IRQ_DP.thy @@ -97,6 +97,29 @@ lemma decode_irq_control_issue_irq_rv: apply sep_solve done +lemma invoke_irq_control_issue_sgi_wp: + "\ c cap \* R> \ + invoke_irq_control (ArchIssueIrqHandler (ARMIssueSGISignal irq targets control_slot dest_slot)) + \\_. c SGISignalCap irq targets \* R> \" + by (clarsimp simp: invoke_irq_control_def arch_invoke_irq_control_def cap_type_def | + wp sep_wp: insert_cap_child_wp | sep_solve)+ + +lemma decode_irq_control_issue_sgi_rv: + "\\s. P (ArchIssueIrqHandler + (ARMIssueSGISignal (ucast irq) (ucast cpus) target_ref + (cap_object root_cap, offset index r))) s \ + caps = (root_cap, root_ptr) # xs \ unat depth \ word_bits \ 0 < unat depth \ + <\ (r, (unat depth)) : root_cap index \u cap \* R> s\ + decode_irq_control_invocation target target_ref caps + (ArchIrqControlIssueIrqHandlerIntent + (ARMIssueSGISignalIntent irq cpus index depth)) + \P\, -" + apply (clarsimp simp: decode_irq_control_invocation_def arch_decode_irq_control_invocation_def) + apply (wp lookup_slot_for_cnode_op_rvu'[where cap=cap and r=r] throw_on_none_rv) + apply (clarsimp simp: get_index_def) + apply sep_solve + done + schematic_goal lookup_extra_caps_once_wp: "\?P\ lookup_extra_caps root_tcb_id [endpoint_cptr] \Q\, \Q'\" apply (clarsimp simp: lookup_extra_caps_def mapME_def sequenceE_def, wp) @@ -256,6 +279,140 @@ lemma seL4_IRQHandler_IRQControl_Get: apply (clarsimp simp: is_irqcontrol_cap_def ep_related_cap_def) done +lemma seL4_IssueSGISignal_helper: + assumes unify: "irq_cap = SGISignalCap (ucast irq) (ucast targets) \ + target_index = offset node_index root_size \ + root_index = offset croot root_size \ + control_index = offset control_cap root_size \ + target_ptr = (cap_object root_cap, target_index) \ + control_ptr = (cap_object root_cap, control_index) \ + root_ptr = (cap_object root_cap, root_index) \ + root_tcb = Tcb t" + shows + "\\root_tcb_id \f root_tcb \* (root_tcb_id, tcb_pending_op_slot) \c RunningCap \* + cap_object root_cap \f CNode (empty_cnode root_size) \* + (root_tcb_id, tcb_cspace_slot) \c root_cap \* control_ptr \c c_cap \* + target_ptr \c target_cap \* root_ptr \c root_cap \* R\ and + K (\ ep_related_cap c_cap \ one_lvl_lookup root_cap word_bits root_size \ + one_lvl_lookup root_cap (unat node_depth) root_size \ + guard_equal root_cap node_index (unat node_depth) \ + guard_equal root_cap croot word_bits \ + guard_equal root_cap control_cap word_bits \ + guard_equal root_cap node_index word_bits \ + unat node_depth \ word_bits \ 0 < unat node_depth \ + is_irqcontrol_cap c_cap \ is_cnode_cap root_cap \ is_cnode_cap root_cap)\ + seL4_IssueSGISignal control_cap irq targets croot node_index node_depth + \\fail. \root_tcb_id \f root_tcb \* + (root_tcb_id, tcb_pending_op_slot) \c RunningCap \* + \ \Root CNode\ + cap_object root_cap \f CNode (empty_cnode root_size) \* + \ \Cap to the root CNode\ + (root_tcb_id, tcb_cspace_slot) \c root_cap \* + control_ptr \c c_cap \* target_ptr \c irq_cap \* + root_ptr \c root_cap \* R\\" + apply (simp add: seL4_IssueSGISignal_def sep_state_projection2_def) + apply (wp do_kernel_op_pull_back) + apply (rule call_kernel_with_intent_allow_error_helper[where + check = True and + Perror = \ and + intent_op = "IrqControlIntent + (ArchIrqControlIssueIrqHandlerIntent + (ARMIssueSGISignalIntent irq targets node_index node_depth))" and + tcb = t and + intent_cptr = control_cap and + intent_extra = "[croot]", simplified]) + apply clarsimp + apply (rule set_cap_wp[sep_wand_wp]) + apply (rule mark_tcb_intent_error_sep_inv) + apply (wp corrupt_ipc_buffer_sep_inv)+ + apply (rule_tac P = "iv = InvokeIrqControl + (ArchIssueIrqHandler + (ARMIssueSGISignal (ucast irq) (ucast targets) control_ptr + (cap_object root_cap, + offset node_index root_size)))" + in hoare_gen_asmEx) + apply (clarsimp simp: unify) + apply (wp invoke_irq_control_issue_sgi_wp[sep_wand_wp]) + apply (wp set_cap_wp[sep_wand_wp]) + apply (rule_tac P = "c=IrqControlCap" in hoare_gen_asmEx, simp) + apply (simp add: decode_invocation_def, wp) + apply (rule_tac P = "irq_control_intent = + ArchIrqControlIssueIrqHandlerIntent + (ARMIssueSGISignalIntent irq targets node_index node_depth)" + in hoare_gen_asmE, simp) + apply (wp decode_irq_control_issue_sgi_rv[where root_cap = root_cap and + root_ptr = root_ptr and + xs = "[]" and r = root_size, simplified]) + apply (simp add: throw_opt_def get_irq_control_intent_def) + apply wp + apply (rule lookup_extra_caps_once_wp[where cap'=root_cap and r=root_size, simplified user_pointer_at_def]) + apply (rule lookup_cap_and_slot_rvu[where cap'=c_cap and r=root_size, simplified user_pointer_at_def]) + apply (wp hoare_vcg_imp_lift hoare_vcg_all_lift update_thread_intent_update) + apply clarsimp + apply (rule conjI, fastforce)+ + apply sep_solve + apply (intro impI conjI allI; clarsimp) + apply (intro impI conjI allI; clarsimp simp: unify) + apply sep_solve + apply sep_cancel+ + apply (intro impI conjI allI; sep_solve) + apply fastforce + apply (metis is_cnode_cap_not_memory_cap not_memory_cap_reset_asid') + apply fastforce + apply (clarsimp simp: user_pointer_at_def Let_def) + apply sep_solve + apply (drule (1) reset_cap_asid_control, clarsimp simp: is_irqcontrol_cap_def) + apply (clarsimp simp: user_pointer_at_def Let_def) + apply sep_solve + apply (drule reset_cap_asid_ep_related_cap) + apply clarsimp + apply (clarsimp simp: user_pointer_at_def Let_def unify) + apply sep_solve + apply (clarsimp simp: unify) + apply sep_solve + done + +lemma seL4_IssueSGISignal_Get: + "\\root_tcb_id \f root_tcb \* + (root_tcb_id, tcb_pending_op_slot) \c RunningCap \* + (root_tcb_id, tcb_cspace_slot) \c cnode_cap \* + cnode_id \f CNode (empty_cnode root_size) \* + (cnode_id, control_index) \c IrqControlCap \* + (cnode_id, root_index) \c cnode_cap \* + (cnode_id, target_index) \c target_cap \* R\ + and K (is_tcb root_tcb \ + is_cnode_cap cnode_cap \ + is_cnode_cap cnode_cap \ + cnode_id = cap_object cnode_cap \ + target_index = offset node_index root_size \ + root_index = offset croot root_size \ + control_index = offset control_cap root_size \ + one_lvl_lookup cnode_cap word_bits root_size \ + one_lvl_lookup cnode_cap (unat node_depth) root_size \ + guard_equal cnode_cap node_index (unat node_depth) \ + guard_equal cnode_cap croot word_bits \ + guard_equal cnode_cap control_cap word_bits \ + guard_equal cnode_cap node_index word_bits \ + unat node_depth \ word_bits \ 0 < unat node_depth)\ + seL4_IssueSGISignal control_cap irq targets croot node_index node_depth + \\fail. + \root_tcb_id \f root_tcb \* + (root_tcb_id, tcb_pending_op_slot) \c RunningCap \* + \ \Cap to the root CNode\ + (root_tcb_id, tcb_cspace_slot) \c cnode_cap \* + \ \Root CNode\ + cnode_id \f CNode (empty_cnode root_size) \* + (cnode_id, control_index) \c IrqControlCap \* + (cnode_id, root_index) \c cnode_cap \* + (cnode_id, target_index) \c SGISignalCap (ucast irq) (ucast targets) \* R\\" + apply (rule hoare_gen_asm, elim conjE) + apply (sep_wp seL4_IssueSGISignal_helper [where t="obj_tcb root_tcb"]) + apply fastforce + apply clarsimp + apply (rule conjI, sep_solve) + apply (clarsimp simp: is_irqcontrol_cap_def ep_related_cap_def) + done + lemma seL4_IRQHandler_SetEndpoint_wp_helper: assumes unify: "irq_cap = IrqHandlerCap irq \ offset endpoint_cptr root_size = irq_handler_slot \ diff --git a/proof/capDL-api/Kernel_DP.thy b/proof/capDL-api/Kernel_DP.thy index f914bd1a1d7..39a70111623 100644 --- a/proof/capDL-api/Kernel_DP.thy +++ b/proof/capDL-api/Kernel_DP.thy @@ -178,6 +178,18 @@ where cdl_intent_extras = [croot], cdl_intent_recv_slot = None\ True" +definition seL4_IssueSGISignal :: + "cdl_cptr \ machine_word \ machine_word \ cdl_cptr \ machine_word \ machine_word \ bool u_monad" + where + "seL4_IssueSGISignal control_cap irq targets croot node_index node_depth \ + do_kernel_op $ call_kernel_with_intent + \cdl_intent_op = Some $ IrqControlIntent $ ArchIrqControlIssueIrqHandlerIntent $ + ARMIssueSGISignalIntent irq targets node_index node_depth, + cdl_intent_error = False, + cdl_intent_cap = control_cap, + cdl_intent_extras = [croot], + cdl_intent_recv_slot = None\ True" + definition seL4_IRQHandler_SetEndpoint :: "cdl_cptr \ cdl_cptr \ bool u_monad" where "seL4_IRQHandler_SetEndpoint handler_cap endpoint_cap \