(* * Copyright 2022, Proofcraft Pty Ltd * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: GPL-2.0-only *) chapter "Decoding Architecture-specific System Calls" theory ArchDecode_A imports Interrupt_A InvocationLabels_A "ExecSpec.InvocationLabels_H" begin context Arch begin global_naming AARCH64_A section "Helper definitions" definition check_vp_alignment :: "vmpage_size \ machine_word \ (unit,'z::state_ext) se_monad" where "check_vp_alignment sz vptr \ unlessE (is_aligned vptr (pageBitsForSize sz)) $ throwError AlignmentError" definition page_base :: "vspace_ref \ vmpage_size \ vspace_ref" where "page_base vaddr vmsize \ vaddr && ~~ mask (pageBitsForSize vmsize)" section "Architecture-specific Decode Functions" definition arch_decode_irq_control_invocation :: "data \ data list \ cslot_ptr \ cap list \ (arch_irq_control_invocation,'z::state_ext) se_monad" where "arch_decode_irq_control_invocation label args src_slot cps \ (if invocation_type label = ArchInvocationLabel ARMIRQIssueIRQHandlerTrigger then if length args \ 4 \ length cps \ 1 then let irq_word = args ! 0; trigger = args ! 1; index = args ! 2; depth = args ! 3; cnode = cps ! 0; irq = ucast irq_word in doE arch_check_irq irq_word; irq_active \ liftE $ is_irq_active irq; whenE irq_active $ throwError RevokeFirst; dest_slot \ lookup_target_slot cnode (data_to_cptr index) (unat depth); ensure_empty dest_slot; returnOk $ ARMIRQControlInvocation irq dest_slot src_slot (trigger \ 0) odE else throwError TruncatedMessage else throwError IllegalOperation)" definition attribs_from_word :: "machine_word \ vm_attributes" where "attribs_from_word w \ {attr. \w!!0 \ attr = Device \ \w !! 2 \ attr = Execute}" definition make_user_pte :: "paddr \ vm_attributes \ vm_rights \ vmpage_size \ pte" where "make_user_pte addr attr rights vm_size \ PagePTE addr (vm_size = ARMSmallPage) (attr - {Global}) rights" definition check_vspace_root :: "cap \ nat \ (obj_ref \ asid, 'z) se_monad" where "check_vspace_root cap arg_no \ case cap of ArchObjectCap (PageTableCap pt VSRootPT_T (Some (asid, _))) \ returnOk (pt, asid) | _ \ throwError $ InvalidCapability arg_no" type_synonym 'z arch_decoder = "data \ data list \ cslot_ptr \ arch_cap \ (cap \ cslot_ptr) list \ (arch_invocation,'z) se_monad" definition decode_fr_inv_map :: "'z::state_ext arch_decoder" where "decode_fr_inv_map label args cte cap extra_caps \ case cap of FrameCap p R pgsz dev mapped_address \ if length args > 2 \ length extra_caps > 0 then let vaddr = args ! 0; rights_mask = args ! 1; attr = args ! 2; vspace_cap = fst (extra_caps ! 0) in doE (pt, asid) \ check_vspace_root vspace_cap 1; pt' \ lookup_error_on_failure False $ find_vspace_for_asid asid; whenE (pt' \ pt) $ throwError $ InvalidCapability 1; check_vp_alignment pgsz vaddr; pg_bits \ returnOk $ pageBitsForSize pgsz; case mapped_address of Some (asid', vaddr') \ doE whenE (asid' \ asid) $ throwError $ InvalidCapability 1; whenE (vaddr' \ vaddr) $ throwError $ InvalidArgument 0 odE | None \ doE vtop \ returnOk $ vaddr + mask (pageBitsForSize pgsz); whenE (vtop > user_vtop) $ throwError $ InvalidArgument 0 odE; (level, slot) \ liftE $ gets_the $ pt_lookup_slot pt vaddr \ ptes_of; unlessE (pt_bits_left level = pg_bits) $ throwError $ FailedLookup False $ MissingCapability $ pt_bits_left level; vm_rights \ returnOk $ mask_vm_rights R (data_to_rights rights_mask); attribs \ returnOk $ attribs_from_word attr; pte \ returnOk $ make_user_pte (addrFromPPtr p) attribs vm_rights pgsz; returnOk $ InvokePage $ PageMap (FrameCap p R pgsz dev (Some (asid,vaddr))) cte (pte,slot,level) odE else throwError TruncatedMessage | _ \ fail" definition label_to_flush_type :: "data \ flush_type" where "label_to_flush_type label \ case invocation_type label of ArchInvocationLabel ARMVSpaceClean_Data \ Clean | ArchInvocationLabel ARMPageClean_Data \ Clean | ArchInvocationLabel ARMVSpaceInvalidate_Data \ Invalidate | ArchInvocationLabel ARMPageInvalidate_Data \ Invalidate | ArchInvocationLabel ARMVSpaceCleanInvalidate_Data \ CleanInvalidate | ArchInvocationLabel ARMPageCleanInvalidate_Data \ CleanInvalidate | ArchInvocationLabel ARMVSpaceUnify_Instruction \ Unify | ArchInvocationLabel ARMPageUnify_Instruction \ Unify" definition decode_fr_inv_flush :: "'z::state_ext arch_decoder" where "decode_fr_inv_flush label args cte cap extra_caps \ case cap of FrameCap p R pgsz dev mapped_address \ if length args > 1 then let start = args ! 0; end = args ! 1 in doE (asid, vaddr) \ case mapped_address of Some a \ returnOk a | _ \ throwError IllegalOperation; vspace \ lookup_error_on_failure False $ find_vspace_for_asid asid; whenE (end \ start) $ throwError $ InvalidArgument 1; page_size \ returnOk $ 1 << pageBitsForSize pgsz; whenE (start \ page_size \ end > page_size) $ throwError $ InvalidArgument 0; pstart \ returnOk $ addrFromPPtr p + start; \ \flush only inside the kernel window:\ whenE (pstart < paddrBase \ ((end - start) + pstart > paddrTop)) $ throwError IllegalOperation; returnOk $ InvokePage $ PageFlush (label_to_flush_type label) (vaddr + start) (vaddr + end - 1) pstart vspace asid odE else throwError TruncatedMessage | _ \ fail" definition decode_frame_invocation :: "'z::state_ext arch_decoder" where "decode_frame_invocation label args cte cap extra_caps \ if invocation_type label = ArchInvocationLabel ARMPageMap then decode_fr_inv_map label args cte cap extra_caps else if invocation_type label = ArchInvocationLabel ARMPageUnmap then returnOk $ InvokePage $ PageUnmap cap cte else if invocation_type label = ArchInvocationLabel ARMPageGetAddress then returnOk $ InvokePage $ PageGetAddr (acap_obj cap) else if isPageFlushLabel (invocation_type label) then decode_fr_inv_flush label args cte cap extra_caps else throwError IllegalOperation" definition decode_pt_inv_map :: "'z::state_ext arch_decoder" where "decode_pt_inv_map label args cte cap extra_caps \ case cap of PageTableCap p t mapped_address \ if length args > 1 \ length extra_caps > 0 then let vaddr = args ! 0; attr = args ! 1; vspace_cap = fst (extra_caps ! 0) in doE whenE (mapped_address \ None) $ throwError $ InvalidCapability 0; (pt, asid) \ check_vspace_root vspace_cap 1; whenE (user_vtop < vaddr) $ throwError $ InvalidArgument 0; pt' \ lookup_error_on_failure False $ find_vspace_for_asid asid; whenE (pt' \ pt) $ throwError $ InvalidCapability 1; (level, slot) \ liftE $ gets_the $ pt_lookup_slot pt vaddr \ ptes_of; old_pte \ liftE $ get_pte (level_type level) slot; whenE (pt_bits_left level = pageBits \ old_pte \ InvalidPTE) $ throwError DeleteFirst; pte \ returnOk $ PageTablePTE (ppn_from_pptr p); cap' <- returnOk $ PageTableCap p t $ Some (asid, vaddr && ~~mask (pt_bits_left level)); returnOk $ InvokePageTable $ PageTableMap cap' cte pte slot level odE else throwError TruncatedMessage | _ \ fail" definition decode_page_table_invocation :: "'z::state_ext arch_decoder" where "decode_page_table_invocation label args cte cap extra_caps \ if invocation_type label = ArchInvocationLabel ARMPageTableMap then decode_pt_inv_map label args cte cap extra_caps else if invocation_type label = ArchInvocationLabel ARMPageTableUnmap then doE final \ liftE $ is_final_cap (ArchObjectCap cap); unlessE final $ throwError RevokeFirst; returnOk $ InvokePageTable $ PageTableUnmap cap cte odE else throwError IllegalOperation" definition level_of_vmsize :: "vmpage_size \ vm_level" where "level_of_vmsize vmsz \ case vmsz of ARMSmallPage \ 0 | ARMLargePage \ 1 | ARMHugePage \ 2" definition vmsize_of_level :: "vm_level \ vmpage_size" where "vmsize_of_level level \ if level = 0 then ARMSmallPage else if level = 1 then ARMLargePage else ARMHugePage" definition lookup_frame :: "obj_ref \ vspace_ref \ ptes_of \ (vmpage_size \ paddr) option" where "lookup_frame vspace vaddr = do { (level, slot) \ pt_lookup_slot vspace vaddr; pte \ oapply2 (level_type level) slot; oassert (is_PagePTE pte); oassert (level \ 2); oreturn (vmsize_of_level level, pte_base_addr pte) }" definition decode_vs_inv_flush :: "'z::state_ext arch_decoder" where "decode_vs_inv_flush label args cte cap extra_caps \ case cap of PageTableCap p VSRootPT_T mapped_address \ if length args > 1 then let start = args ! 0; end = args ! 1 in doE whenE (end \ start) $ throwError $ InvalidArgument 1; whenE (end > pptrUserTop) $ throwError $ IllegalOperation; (vspace, asid) \ check_vspace_root (ArchObjectCap cap) 0; vspace' \ lookup_error_on_failure False $ find_vspace_for_asid asid; whenE (vspace' \ vspace) $ throwError $ InvalidCapability 0; frame_info \ liftE $ gets $ lookup_frame p start \ ptes_of; case frame_info of None \ returnOk $ InvokeVSpace VSpaceNothing | Some (pgsz, paddr) \ doE bits_left \ returnOk $ pt_bits_left (level_of_vmsize pgsz); base_start \ returnOk $ page_base start pgsz; base_end \ returnOk $ page_base (end - 1) pgsz; whenE (base_start \ base_end) $ throwError $ RangeError start (base_start + mask bits_left); pstart \ returnOk $ paddr + (start && mask bits_left); returnOk $ InvokeVSpace $ VSpaceFlush (label_to_flush_type label) start (end - 1) pstart vspace asid odE odE else throwError TruncatedMessage | _ \ fail" definition decode_vspace_invocation :: "'z::state_ext arch_decoder" where "decode_vspace_invocation label args cte cap extra_caps \ if isVSpaceFlushLabel (invocation_type label) then decode_vs_inv_flush label args cte cap extra_caps else throwError IllegalOperation" definition decode_asid_control_invocation :: "'z::state_ext arch_decoder" where "decode_asid_control_invocation label args cte cap extra_caps \ if invocation_type label = ArchInvocationLabel ARMASIDControlMakePool then if length args > 1 \ length extra_caps > 1 then let index = args ! 0; depth = args ! 1; (untyped, parent_slot) = extra_caps ! 0; root = fst (extra_caps ! 1) in doE asid_table \ liftE $ gets asid_table; free_set \ returnOk (- dom asid_table); whenE (free_set = {}) $ throwError DeleteFirst; free \ liftE $ select_ext (\_. free_asid_select asid_table) free_set; base \ returnOk (ucast free << asid_low_bits); (p,n) \ case untyped of UntypedCap False p n _ \ returnOk (p,n) | _ \ throwError $ InvalidCapability 1; frame \ if n = pageBits then doE ensure_no_children parent_slot; returnOk p odE else throwError $ InvalidCapability 1; dest_slot \ lookup_target_slot root (to_bl index) (unat depth); ensure_empty dest_slot; returnOk $ InvokeASIDControl $ MakePool frame dest_slot parent_slot base odE else throwError TruncatedMessage else throwError IllegalOperation" definition decode_asid_pool_invocation :: "'z::state_ext arch_decoder" where "decode_asid_pool_invocation label args cte cap extra_caps \ if invocation_type label = ArchInvocationLabel ARMASIDPoolAssign then if length extra_caps > 0 then let (pt_cap, pt_cap_slot) = extra_caps ! 0; p = acap_obj cap; base = acap_asid_base cap in case pt_cap of ArchObjectCap (PageTableCap _ VSRootPT_T None) \ doE asid_table \ liftE $ gets asid_table; pool_ptr \ returnOk (asid_table (asid_high_bits_of base)); whenE (pool_ptr = None) $ throwError $ FailedLookup False InvalidRoot; whenE (p \ the pool_ptr) $ throwError $ InvalidCapability 0; pool \ liftE $ get_asid_pool p; free_set \ returnOk (- dom pool \ {x. ucast x + base \ 0}); whenE (free_set = {}) $ throwError DeleteFirst; offset \ liftE $ select_ext (\_. free_asid_pool_select pool base) free_set; returnOk $ InvokeASIDPool $ Assign (ucast offset + base) p pt_cap_slot odE | _ \ throwError $ InvalidCapability 1 else throwError TruncatedMessage else throwError IllegalOperation" definition arch_decode_invocation :: "data \ data list \ cap_ref \ cslot_ptr \ arch_cap \ (cap \ cslot_ptr) list \ (arch_invocation,'z::state_ext) se_monad" where "arch_decode_invocation label args x_slot cte cap extra_caps \ case cap of PageTableCap _ VSRootPT_T _ \ decode_vspace_invocation label args cte cap extra_caps | PageTableCap _ NormalPT_T _ \ decode_page_table_invocation label args cte cap extra_caps | FrameCap _ _ _ _ _ \ decode_frame_invocation label args cte cap extra_caps | ASIDControlCap \ decode_asid_control_invocation label args cte cap extra_caps | ASIDPoolCap _ _ \ decode_asid_pool_invocation label args cte cap extra_caps | VCPUCap _ \ decode_vcpu_invocation label args cap extra_caps" section "Interface Functions used in Decode" definition arch_data_to_obj_type :: "nat \ aobject_type option" where "arch_data_to_obj_type n \ if n = 0 then Some HugePageObj else if n = 1 then Some VSpaceObj else if n = 2 then Some SmallPageObj else if n = 3 then Some LargePageObj else if n = 4 then Some PageTableObj else if n = 5 then Some VCPUObj else None" end end