(* * Copyright 2014, General Dynamics C4 Systems * * SPDX-License-Identifier: GPL-2.0-only *) chapter "Deleting Capabilities" theory Delete_H imports CNode_H Interrupt_H Endpoint_H Thread_H begin definition slotsPointed :: "capability \ machine_word set" where "slotsPointed cap \ case cap of CNodeCap ptr a b c \ {ptr} | ThreadCap ptr \ {ptr} | Zombie ptr bits num \ {ptr} | _ \ {}" primrec sethelper :: "bool \ 'a set \ 'a set" where "sethelper True s = {}" | "sethelper False s = s" function finaliseSlot' :: "machine_word \ bool \ (bool * capability) kernel_p" where "finaliseSlot' x b s = (\ finaliseSlot. (\ cteDelete. (\ reduceZombie. #INCLUDE_HASKELL SEL4/Object/CNode.lhs BODY finaliseSlot ) ( #INCLUDE_HASKELL SEL4/Object/CNode.lhs BODY reduceZombie ) ) ( #INCLUDE_HASKELL SEL4/Object/CNode.lhs BODY cteDelete ) ) finaliseSlot' x b s" by auto defs finaliseSlot_def: "finaliseSlot \ finaliseSlot'" #INCLUDE_HASKELL_PREPARSE SEL4/Object/Structures.lhs function cteDeleteOne' :: "machine_word \ unit kernel" where "cteDeleteOne' x s = (\ cteDeleteOne. (\ deletingIRQHandler. (\ cancelIPC. (\ suspend. (\ finaliseCap. #INCLUDE_HASKELL SEL4/Object/CNode.lhs BODY cteDeleteOne ) ( #INCLUDE_HASKELL SEL4/Object/ObjectType.lhs Arch=Arch BODY finaliseCap ) ) ( #INCLUDE_HASKELL SEL4/Kernel/Thread.lhs BODY suspend ) ) ( #INCLUDE_HASKELL SEL4/Object/Endpoint.lhs BODY cancelIPC ) ) ( #INCLUDE_HASKELL SEL4/Object/Interrupt.lhs BODY deletingIRQHandler ) ) cteDeleteOne' x s" by auto defs cteDeleteOne_def1: "cteDeleteOne \ cteDeleteOne'" termination cteDeleteOne' by (rule cteDeleteOne'.termination[OF wf_empty], simp+) lemma cteDeleteOne_def: "cteDeleteOne = ( #INCLUDE_HASKELL SEL4/Object/CNode.lhs BODY cteDeleteOne )" apply (rule ext)+ apply (subst cteDeleteOne_def1) apply (subst cteDeleteOne'.simps) apply (unfold finaliseCap_def suspend_def cancelIPC_def deletingIRQHandler_def cteDeleteOne_def1) apply (rule refl) done lemma card_reduce: "(s :: ('a :: finite) set) \ s' = {} \ card (UNIV - (s \ s')) < card (UNIV - s) = (s' \ {})" apply (case_tac "s' \ s") apply (simp add: Un_absorb2) apply (simp add: Int_absorb1) apply (clarsimp simp: subset_iff) apply (subst psubset_card_mono) apply simp apply blast apply blast done lemma isCapDs: "isUntypedCap cap \ \dev ptr size freeIndex. cap = UntypedCap dev ptr size freeIndex" "isEndpointCap cap \ \ptr bdg cans canr cang cangr. cap = EndpointCap ptr bdg cans canr cang cangr" "isNotificationCap cap \ \ptr bdg cans canr. cap = NotificationCap ptr bdg cans canr" "isCNodeCap cap \ \ptr bits grd gsize. cap = CNodeCap ptr bits grd gsize" "isThreadCap cap \ \ptr. cap = ThreadCap ptr" "isArchObjectCap cap \ \archcap. cap = ArchObjectCap archcap" "isZombie cap \ \ptr bits num. cap = Zombie ptr bits num" apply (case_tac cap, simp_all add: isUntypedCap_def) apply (case_tac cap, simp_all add: isEndpointCap_def) apply (case_tac cap, simp_all add: isNotificationCap_def) apply (case_tac cap, simp_all add: isCNodeCap_def) apply (case_tac cap, simp_all add: isThreadCap_def) apply (case_tac cap, simp_all add: isArchObjectCap_def) apply (case_tac cap, simp_all add: isZombie_def) done end