(* * 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) *) (* Wellformedness of caps, kernel objects, states on the C level *) theory Wellformed_C imports "CLib.CTranslationNICTA" CLevityCatch "CSpec.Substitute" begin context begin interpretation Arch . (*FIXME: arch_split*) abbreviation cte_Ptr :: "word32 \ cte_C ptr" where "cte_Ptr == Ptr" abbreviation mdb_Ptr :: "word32 \ mdb_node_C ptr" where "mdb_Ptr == Ptr" abbreviation cap_Ptr :: "word32 \ cap_C ptr" where "cap_Ptr == Ptr" abbreviation tcb_Ptr :: "word32 \ tcb_C ptr" where "tcb_Ptr == Ptr" abbreviation atcb_Ptr :: "word32 \ arch_tcb_C ptr" where "atcb_Ptr == Ptr" abbreviation vcpu_Ptr :: "word32 \ vcpu_C ptr" where "vcpu_Ptr == Ptr" abbreviation ep_Ptr :: "word32 \ endpoint_C ptr" where "ep_Ptr == Ptr" abbreviation ntfn_Ptr :: "word32 \ notification_C ptr" where "ntfn_Ptr == Ptr" abbreviation ap_Ptr :: "word32 \ asid_pool_C ptr" where "ap_Ptr == Ptr" abbreviation pte_Ptr :: "word32 \ pte_C ptr" where "pte_Ptr == Ptr" abbreviation pde_Ptr :: "word32 \ pde_C ptr" where "pde_Ptr == Ptr" abbreviation regs_C_Ptr :: "addr \ (machine_word_len word[39]) ptr" where "regs_C_Ptr \ Ptr" abbreviation vgic_lr_C_Ptr :: "addr \ (virq_C[64]) ptr" where "vgic_lr_C_Ptr \ Ptr" abbreviation vgic_C_Ptr :: "addr \ gicVCpuIface_C ptr" where "vgic_C_Ptr \ Ptr" (* FIXME might be useful on other platforms if not existing already *) abbreviation word_Ptr :: "addr \ machine_word ptr" where "word_Ptr \ Ptr" lemma halt_spec: "Gamma \ {} Call halt_'proc {}" apply (rule hoare_complete) apply (simp add: HoarePartialDef.valid_def) done definition isUntypedCap_C :: "cap_CL \ bool" where "isUntypedCap_C c \ case c of Cap_untyped_cap q \ True | _ \ False" definition isNullCap_C :: "cap_CL \ bool" where "isNullCap_C c \ case c of Cap_null_cap \ True | _ \ False" definition isEndpointCap_C :: "cap_CL \ bool" where "isEndpointCap_C v \ case v of Cap_endpoint_cap ec \ True | _ \ False" definition isCNodeCap_C :: "cap_CL \ bool" where "isCNodeCap_C c \ case c of Cap_cnode_cap a \ True | _ \ False" definition isThreadCap_C :: "cap_CL \ bool" where "isThreadCap_C c \ case c of Cap_thread_cap a \ True | _ \ False" definition isVCPUCap_C :: "cap_CL \ bool" where "isVCPUCap_C c \ case c of Cap_vcpu_cap a \ True | _ \ False" definition isIRQControlCap_C :: "cap_CL \ bool" where "isIRQControlCap_C c \ case c of Cap_irq_control_cap \ True | _ \ False" definition isIRQHandlerCap_C :: "cap_CL \ bool" where "isIRQHandlerCap_C c \ case c of Cap_irq_handler_cap a \ True | _ \ False" definition isNotificationCap_C :: "cap_CL \ bool" where "isNotificationCap_C v \ case v of Cap_notification_cap aec \ True | _ \ False" definition ep_at_C' :: "word32 \ heap_raw_state \ bool" where "ep_at_C' p h \ Ptr p \ dom (clift h :: endpoint_C typ_heap)" \ \endpoint_lift is total\ definition ntfn_at_C' :: "word32 \ heap_raw_state \ bool" where \ \notification_lift is total\ "ntfn_at_C' p h \ Ptr p \ dom (clift h :: notification_C typ_heap)" definition tcb_at_C' :: "word32 \ heap_raw_state \ bool" where "tcb_at_C' p h \ Ptr p \ dom (clift h :: tcb_C typ_heap)" definition vcpu_at_C' :: "word32 \ heap_raw_state \ bool" where "vcpu_at_C' p h \ Ptr p \ dom (clift h :: vcpu_C typ_heap)" definition cte_at_C' :: "word32 \ heap_raw_state \ bool" where "cte_at_C' p h \ Ptr p \ dom (clift h :: cte_C typ_heap)" definition ctcb_ptr_to_tcb_ptr :: "tcb_C ptr \ word32" where "ctcb_ptr_to_tcb_ptr p \ ptr_val p - ctcb_offset" definition tcb_ptr_to_ctcb_ptr :: "word32 \ tcb_C ptr" where "tcb_ptr_to_ctcb_ptr p \ Ptr (p + ctcb_offset)" primrec tcb_queue_relation :: "(tcb_C \ tcb_C ptr) \ (tcb_C \ tcb_C ptr) \ (tcb_C ptr \ tcb_C option) \ word32 list \ tcb_C ptr \ tcb_C ptr \ bool" where "tcb_queue_relation getNext getPrev hp [] qprev qhead = (qhead = NULL)" | "tcb_queue_relation getNext getPrev hp (x#xs) qprev qhead = (qhead = tcb_ptr_to_ctcb_ptr x \ (\tcb. (hp qhead = Some tcb \ getPrev tcb = qprev \ tcb_queue_relation getNext getPrev hp xs qhead (getNext tcb))))" abbreviation "ep_queue_relation \ tcb_queue_relation tcbEPNext_C tcbEPPrev_C" abbreviation "sched_queue_relation \ tcb_queue_relation tcbSchedNext_C tcbSchedPrev_C" definition wordSizeCase :: "'a \ 'a \ 'a" where "wordSizeCase a b \ (if bitSize (undefined::word32) = 32 then a else if bitSize (undefined::word32) = 64 then b else error [] )" primrec capBits_C :: "cap_CL \ nat" where "capBits_C Cap_null_cap = 0" | "capBits_C (Cap_untyped_cap uc) = unat (capBlockSize_CL uc)" | "capBits_C (Cap_endpoint_cap ec) = wordSizeCase 4 5" | "capBits_C (Cap_notification_cap aec) = wordSizeCase 4 5" | "capBits_C (Cap_cnode_cap cnc) = wordSizeCase 4 5" | "capBits_C (Cap_thread_cap tc) = 10" | "capBits_C (Cap_zombie_cap zc) = (wordSizeCase 4 5)" | "capBits_C (Cap_vcpu_cap tc) = 12" definition capUntypedPtr_C :: "cap_CL \ word32" where "capUntypedPtr_C cap \ case cap of (Cap_untyped_cap uc) \ (capBlockSize_CL uc) | Cap_endpoint_cap ep \ (capEPPtr_CL ep) | Cap_notification_cap ntfn \ (capNtfnPtr_CL ntfn) | Cap_cnode_cap ccap \ (capCNodePtr_CL ccap) | Cap_reply_cap rc \ (cap_reply_cap_CL.capTCBPtr_CL rc) | Cap_thread_cap tc \ (cap_thread_cap_CL.capTCBPtr_CL tc) | Cap_small_frame_cap sfc \ (cap_small_frame_cap_CL.capFBasePtr_CL sfc) | Cap_frame_cap fc \ (cap_frame_cap_CL.capFBasePtr_CL fc) | Cap_page_table_cap ptc \ (capPTBasePtr_CL ptc) | Cap_page_directory_cap pdc \ (capPDBasePtr_CL pdc) | Cap_vcpu_cap tc \ (cap_vcpu_cap_CL.capVCPUPtr_CL tc) | _ \ error []" definition ZombieTCB_C_def: "ZombieTCB_C \ bit 5" definition isZombieTCB_C :: "word32 \ bool" where "isZombieTCB_C v \ v = ZombieTCB_C" definition vmrights_to_H :: "word32 \ vmrights" where "vmrights_to_H c \ if c = scast Kernel_C.VMNoAccess then VMNoAccess else if c = scast Kernel_C.VMKernelOnly then VMKernelOnly else if c = scast Kernel_C.VMReadOnly then VMReadOnly else VMReadWrite" (* Force clarity over name collisions *) abbreviation ARMSmallPage :: "vmpage_size" where "ARMSmallPage == ARM_HYP.ARMSmallPage" abbreviation ARMLargePage :: "vmpage_size" where "ARMLargePage == ARM_HYP.ARMLargePage" abbreviation ARMSection :: "vmpage_size" where "ARMSection == ARM_HYP.ARMSection" abbreviation ARMSuperSection :: "vmpage_size" where "ARMSuperSection == ARM_HYP.ARMSuperSection" \ \ARMSmallFrame is treated in a separate cap in C, so needs special treatment in ccap_relation\ definition framesize_to_H:: "word32 \ vmpage_size" where "framesize_to_H c \ if c = scast Kernel_C.ARMLargePage then ARMLargePage else if c = scast Kernel_C.ARMSection then ARMSection else ARMSuperSection" \ \Use this for results of generic_frame_cap_get_capFSize\ definition gen_framesize_to_H:: "word32 \ vmpage_size" where "gen_framesize_to_H c \ if c = scast Kernel_C.ARMSmallPage then ARMSmallPage else if c = scast Kernel_C.ARMLargePage then ARMLargePage else if c = scast Kernel_C.ARMSection then ARMSection else ARMSuperSection" end record cte_CL = cap_CL :: cap_CL cteMDBNode_CL :: mdb_node_CL context begin interpretation Arch . (*FIXME: arch_split*) definition cte_lift :: "cte_C \ cte_CL" where "cte_lift c \ case cap_lift (cte_C.cap_C c) of None \ None | Some cap \ Some \ cap_CL = cap, cteMDBNode_CL = mdb_node_lift (cteMDBNode_C c) \" lemma to_bool_false [simp]: "\ to_bool false" by (simp add: to_bool_def false_def) (* this is slightly weird, but the bitfield generator masks everything with the expected bit length. So we do that here too. *) definition to_bool_bf :: "'a::len word \ bool" where "to_bool_bf w \ (w && mask 1) = 1" lemma to_bool_bf_mask1 [simp]: "to_bool_bf (mask (Suc 0))" by (simp add: mask_def to_bool_bf_def) lemma to_bool_bf_0 [simp]: "\to_bool_bf 0" by (simp add: to_bool_bf_def) lemma to_bool_bf_1 [simp]: "to_bool_bf 1" by (simp add: to_bool_bf_def mask_def) lemma to_bool_bf_false [simp]: "\to_bool_bf false" by (simp add: false_def) lemma to_bool_bf_true [simp]: "to_bool_bf true" by (simp add: true_def) lemma to_bool_to_bool_bf: "w = false \ w = true \ to_bool_bf w = to_bool w" by (auto simp: false_def true_def to_bool_def to_bool_bf_def mask_def) lemma to_bool_bf_mask_1 [simp]: "to_bool_bf (w && mask (Suc 0)) = to_bool_bf w" by (simp add: to_bool_bf_def) lemma to_bool_bf_and [simp]: "to_bool_bf (a && b) = (to_bool_bf a \ to_bool_bf (b::word32))" apply (clarsimp simp: to_bool_bf_def) apply (rule iffI) apply (subst (asm) bang_eq) apply (simp add: word_size) apply (rule conjI) apply (rule word_eqI) apply (auto simp add: word_size)[1] apply (rule word_eqI) apply (auto simp add: word_size)[1] apply clarsimp apply (rule word_eqI) apply (subst (asm) bang_eq)+ apply (auto simp add: word_size)[1] done lemma to_bool_bf_to_bool_mask: "w && mask (Suc 0) = w \ to_bool_bf w = to_bool (w::word32)" apply (auto simp add: to_bool_bf_def to_bool_def mask_eq_iff_w2p word_size) apply (auto simp add: mask_def dest: word_less_cases) done definition mdb_node_to_H :: "mdb_node_CL \ mdbnode" where "mdb_node_to_H n \ MDB (mdbNext_CL n) (mdbPrev_CL n) (to_bool (mdbRevocable_CL n)) (to_bool (mdbFirstBadged_CL n))" definition cap_to_H :: "cap_CL \ capability" where "cap_to_H c \ case c of Cap_null_cap \ NullCap | Cap_zombie_cap zc \ (if isZombieTCB_C(capZombieType_CL zc) then (Zombie ((capZombieID_CL zc) && ~~(mask(5))) (ZombieTCB) (unat ((capZombieID_CL zc) && mask(5)))) else let radix = unat (capZombieType_CL zc) in (Zombie ((capZombieID_CL zc) && ~~(mask (radix+1))) (ZombieCNode radix) (unat ((capZombieID_CL zc) && mask(radix+1))))) | Cap_cnode_cap ccap \ CNodeCap (capCNodePtr_CL ccap) (unat (capCNodeRadix_CL ccap)) (capCNodeGuard_CL ccap) (unat (capCNodeGuardSize_CL ccap)) | Cap_untyped_cap uc \ UntypedCap (to_bool(capIsDevice_CL uc)) (capPtr_CL uc) (unat (capBlockSize_CL uc)) (unat (capFreeIndex_CL uc << 4)) | Cap_endpoint_cap ec \ EndpointCap (capEPPtr_CL ec) (capEPBadge_CL ec) (to_bool(capCanSend_CL ec)) (to_bool(capCanReceive_CL ec)) (to_bool(capCanGrant_CL ec)) | Cap_notification_cap ntfn \ NotificationCap (capNtfnPtr_CL ntfn)(capNtfnBadge_CL ntfn)(to_bool(capNtfnCanSend_CL ntfn)) (to_bool(capNtfnCanReceive_CL ntfn)) | Cap_reply_cap rc \ ReplyCap (ctcb_ptr_to_tcb_ptr (Ptr (cap_reply_cap_CL.capTCBPtr_CL rc))) (to_bool (capReplyMaster_CL rc)) | Cap_thread_cap tc \ ThreadCap(ctcb_ptr_to_tcb_ptr (Ptr (cap_thread_cap_CL.capTCBPtr_CL tc))) | Cap_irq_handler_cap ihc \ IRQHandlerCap (ucast(capIRQ_CL ihc)) | Cap_irq_control_cap \ IRQControlCap | Cap_asid_control_cap \ ArchObjectCap ASIDControlCap | Cap_asid_pool_cap apc \ ArchObjectCap (ASIDPoolCap (capASIDPool_CL apc) (capASIDBase_CL apc)) | Cap_small_frame_cap sfc \ ArchObjectCap (PageCap (to_bool(cap_small_frame_cap_CL.capFIsDevice_CL sfc)) (cap_small_frame_cap_CL.capFBasePtr_CL sfc) (vmrights_to_H(cap_small_frame_cap_CL.capFVMRights_CL sfc)) (ARMSmallPage) (if cap_small_frame_cap_CL.capFMappedASIDHigh_CL sfc = 0 \ cap_small_frame_cap_CL.capFMappedASIDLow_CL sfc = 0 then None else Some( (((cap_small_frame_cap_CL.capFMappedASIDHigh_CL sfc)< ArchObjectCap (PageCap (to_bool(capFIsDevice_CL fc)) (capFBasePtr_CL fc) (vmrights_to_H(capFVMRights_CL fc)) (framesize_to_H(capFSize_CL fc)) (if capFMappedASIDHigh_CL fc = 0 \ capFMappedASIDLow_CL fc = 0 then None else Some( (((capFMappedASIDHigh_CL fc)< ArchObjectCap (PageTableCap (capPTBasePtr_CL ptc) (if to_bool (capPTIsMapped_CL ptc) then Some( ((capPTMappedASID_CL ptc)),(capPTMappedAddress_CL ptc)) else None)) | Cap_page_directory_cap pdf \ ArchObjectCap (PageDirectoryCap (capPDBasePtr_CL pdf) (if to_bool (capPDIsMapped_CL pdf) then Some (capPDMappedASID_CL pdf) else None)) | Cap_domain_cap \ DomainCap | Cap_vcpu_cap vcpu \ ArchObjectCap (VCPUCap (capVCPUPtr_CL vcpu))" lemmas cap_to_H_simps = cap_to_H_def[split_simps cap_CL.split] definition cte_to_H :: "cte_CL \ cte" where "cte_to_H cte \ CTE (cap_to_H (cap_CL cte)) (mdb_node_to_H (cteMDBNode_CL cte))" definition cl_valid_cap :: "cap_CL \ bool" where "cl_valid_cap c \ case c of Cap_frame_cap fc \ ((capFSize_CL fc) \ scast Kernel_C.ARMSmallPage) | Cap_irq_handler_cap fc \ ((capIRQ_CL fc) && mask 10 = capIRQ_CL fc) | x \ True" definition c_valid_cap :: "cap_C \ bool" where "c_valid_cap c \ case_option True cl_valid_cap (cap_lift c)" definition cl_valid_cte :: "cte_CL \ bool" where "cl_valid_cte c \ cl_valid_cap (cap_CL c)" definition c_valid_cte :: "cte_C \ bool" where "c_valid_cte c \ c_valid_cap (cte_C.cap_C c)" lemma c_valid_cap_simps [simp]: "cap_get_tag c = scast cap_small_frame_cap \ c_valid_cap c" "cap_get_tag c = scast cap_thread_cap \ c_valid_cap c" "cap_get_tag c = scast cap_notification_cap \ c_valid_cap c" "cap_get_tag c = scast cap_endpoint_cap \ c_valid_cap c" "cap_get_tag c = scast cap_cnode_cap \ c_valid_cap c" "cap_get_tag c = scast cap_page_directory_cap \ c_valid_cap c" "cap_get_tag c = scast cap_asid_control_cap \ c_valid_cap c" "cap_get_tag c = scast cap_irq_control_cap \ c_valid_cap c" "cap_get_tag c = scast cap_page_table_cap \ c_valid_cap c" "cap_get_tag c = scast cap_asid_pool_cap \ c_valid_cap c" "cap_get_tag c = scast cap_untyped_cap \ c_valid_cap c" "cap_get_tag c = scast cap_zombie_cap \ c_valid_cap c" "cap_get_tag c = scast cap_reply_cap \ c_valid_cap c" "cap_get_tag c = scast cap_vcpu_cap \ c_valid_cap c" "cap_get_tag c = scast cap_null_cap \ c_valid_cap c" unfolding c_valid_cap_def cap_lift_def cap_tag_defs by (simp add: cl_valid_cap_def)+ lemma ptr_val_tcb_ptr_mask2: "is_aligned thread tcbBlockSizeBits \ ptr_val (tcb_ptr_to_ctcb_ptr thread) && (~~ mask tcbBlockSizeBits) = thread" apply (clarsimp simp: tcb_ptr_to_ctcb_ptr_def projectKOs) apply (simp add: is_aligned_add_helper ctcb_offset_defs objBits_simps') done lemma maxDom_to_H: "ucast maxDom = maxDomain" by (simp add: maxDomain_def maxDom_def numDomains_def) lemma maxPrio_to_H: "ucast seL4_MaxPrio = maxPriority" by (simp add: maxPriority_def seL4_MaxPrio_def numPriorities_def) (* Input abbreviations for API object types *) (* disambiguates names *) abbreviation(input) NotificationObject :: sword32 where "NotificationObject == seL4_NotificationObject" abbreviation(input) CapTableObject :: sword32 where "CapTableObject == seL4_CapTableObject" abbreviation(input) EndpointObject :: sword32 where "EndpointObject == seL4_EndpointObject" abbreviation(input) LargePageObject :: sword32 where "LargePageObject == seL4_ARM_LargePageObject" abbreviation(input) PageDirectoryObject :: sword32 where "PageDirectoryObject == seL4_ARM_PageDirectoryObject" abbreviation(input) PageTableObject :: sword32 where "PageTableObject == seL4_ARM_PageTableObject" abbreviation(input) SectionObject :: sword32 where "SectionObject == seL4_ARM_SectionObject" abbreviation(input) SmallPageObject :: sword32 where "SmallPageObject == seL4_ARM_SmallPageObject" abbreviation(input) SuperSectionObject :: sword32 where "SuperSectionObject == seL4_ARM_SuperSectionObject" abbreviation(input) TCBObject :: sword32 where "TCBObject == seL4_TCBObject" abbreviation(input) UntypedObject :: sword32 where "UntypedObject == seL4_UntypedObject" abbreviation(input) maxPrio :: sword32 where "maxPrio == seL4_MaxPrio" abbreviation(input) minPrio :: sword32 where "minPrio == seL4_MinPrio" abbreviation(input) nAPIObjects :: sword32 where "nAPIObjects == seL4_NonArchObjectTypeCount" abbreviation(input) nObjects :: sword32 where "nObjects == seL4_ObjectTypeCount" abbreviation(input) prioInvalid :: sword32 where "prioInvalid == seL4_InvalidPrio" end end