(* * Copyright 2014, General Dynamics C4 Systems * * SPDX-License-Identifier: GPL-2.0-only *) (* Arch-specific functions for the abstract model of CSpace. *) chapter "ArchCSpace" theory ArchCSpace_A imports ArchVSpace_A begin context Arch begin global_naming X64_A definition cnode_guard_size_bits :: "nat" where "cnode_guard_size_bits \ 6" definition cnode_padding_bits :: "nat" where "cnode_padding_bits \ 0" text \On a user request to modify a cnode capability, extract new guard bits and guard.\ definition update_cnode_cap_data :: "data \ nat \ data" where "update_cnode_cap_data w \ let guard_bits = 58; guard_size' = unat ((w >> cnode_padding_bits) && mask cnode_guard_size_bits); guard'' = (w >> (cnode_padding_bits + cnode_guard_size_bits)) && mask guard_bits in (guard_size', guard'')" text \For some purposes capabilities to physical objects are treated differently to others.\ definition arch_is_physical :: "arch_cap \ bool" where "arch_is_physical cap \ case cap of ASIDControlCap \ False | IOPortCap _ _ \ False | IOPortControlCap \ False | _ \ True" text \Check whether the second capability is to the same object or an object contained in the region of the first one.\ fun arch_same_region_as :: "arch_cap \ arch_cap \ bool" where "arch_same_region_as (PageCap dev r R t s m) c' = (\ dev' r' R' t' s' m'. c' = PageCap dev' r' R' t' s' m' \ (let topA = r + (1 << pageBitsForSize s) - 1; topB = r' + (1 << pageBitsForSize s') - 1 in r \ r' \ topA \ topB \ r' \ topB))" | "arch_same_region_as (PageTableCap r _) c' = (\r' d'. c' = PageTableCap r' d' \ r = r')" | "arch_same_region_as (PageDirectoryCap r _) c' = (\r' d'. c' = PageDirectoryCap r' d' \ r = r')" | "arch_same_region_as (PDPointerTableCap r _) c' = (\r' d'. c' = PDPointerTableCap r' d' \ r = r')" | "arch_same_region_as (PML4Cap r _) c' = (\r' d'. c' = PML4Cap r' d' \ r = r')" | "arch_same_region_as ASIDControlCap c' = (c' = ASIDControlCap)" | "arch_same_region_as (ASIDPoolCap r _) c' = (\r' d'. c' = ASIDPoolCap r' d' \ r = r')" (* FIXME x64-vtd: | "arch_same_region_as (IOPageTableCap r _ _) c = (is_IOPageTableCap c \ aobj_ref c = Some r)" | "arch_same_region_as (IOSpaceCap d_id pci_d) c = (is_IOSpaceCap c \ cap_io_pci_device c = pci_d)" FIXME x64-vtd: should this also check domain id equality? C kernel does not" *) | "arch_same_region_as (IOPortCap frst lst) c' = (\frst' lst'. c' = IOPortCap frst' lst' \ frst' = frst \ lst' = lst)" | "arch_same_region_as IOPortControlCap c' = (c' = IOPortControlCap \ (\f l. c' = IOPortCap f l))" text \Check whether two arch capabilities are to the same object.\ definition same_aobject_as :: "arch_cap \ arch_cap \ bool" where "same_aobject_as cp cp' \ (case (cp, cp') of (PageCap dev ref _ _ pgsz _, PageCap dev' ref' _ _ pgsz' _) \ (dev, ref, pgsz) = (dev', ref', pgsz') \ ref \ ref + 2 ^ pageBitsForSize pgsz - 1 | (IOPortControlCap, IOPortCap f' l') \ False | _ \ arch_same_region_as cp cp')" declare same_aobject_as_def[simp] definition arch_is_cap_revocable :: "cap \ cap \ bool" where "arch_is_cap_revocable new_cap src_cap \ if \f l. new_cap = ArchObjectCap (IOPortCap f l) then src_cap = ArchObjectCap IOPortControlCap else False" end end