(* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: GPL-2.0-only *) theory SysInit_SI imports "DSpecProofs.Kernel_DP" "Lib.NonDetMonadLemmaBucket" begin definition parse_bootinfo :: "cdl_bootinfo \ (cdl_cptr list \ cdl_cptr list) u_monad" where "parse_bootinfo bootinfo \ do (untyped_start, untyped_end) \ return $ bi_untypes bootinfo; untyped_list \ return [untyped_start .e. (untyped_end - 1)]; (free_slot_start, free_slot_end) \ return $ bi_free_slots bootinfo; free_slots \ return [free_slot_start .e. (free_slot_end - 1)]; return (untyped_list, free_slots) od" (* Create a new object in the next free cnode slot using a given untyped. Return whether or not the operation fails. *) definition retype_untyped :: "cdl_cptr \ cdl_cptr \ cdl_object_type \ nat \ bool u_monad" where "retype_untyped free_slot free_untyped type size_bits \ seL4_Untyped_Retype free_untyped type (of_nat size_bits) seL4_CapInitThreadCNode 0 0 free_slot 1" definition inc_when :: "bool \ nat \ ('s,nat) nondet_monad" where "inc_when P x \ return (if P then Suc x else x)" lemma inc_when_wp [wp]: "\Q (if B then Suc x else x)\ inc_when B x \Q\" by (unfold inc_when_def, wp) definition update_when :: "bool \ ('a \ 'b option) \ 'a \ 'b \ ('s,('a \ 'b option)) nondet_monad" where "update_when P t a b \ return (if P then t(a \ b) else t)" lemma update_when_wp [wp]: "\Q (if B then t(a \ b) else t)\ update_when B t a b \Q\" by (unfold update_when_def, wp) (* Create the objects in the capDL spec. Caps to new objects are stored in the free slots. *) definition create_objects :: "cdl_state \ cdl_object_id list \ cdl_cptr list \ cdl_cptr list \ ((cdl_object_id \ cdl_cptr option) \ cdl_cptr list) u_monad" where "create_objects spec obj_ids untyped_cptrs free_cptrs \ do (obj_id_index, untyped_index, si_caps) \ whileLoop (\(obj_id_index, untyped_index, si_caps) _. obj_id_index < length [obj\obj_ids. real_object_at obj spec] \ obj_id_index < length free_cptrs \ untyped_index < length untyped_cptrs) (\(obj_id_index, untyped_index, si_caps). do obj_id \ return $ [obj\obj_ids. real_object_at obj spec] ! obj_id_index; free_cptr \ return $ free_cptrs ! obj_id_index; untyped_cptr \ return $ untyped_cptrs ! untyped_index; object \ get_spec_object spec obj_id; object_type \ return $ object_type object; object_size \ return $ object_size_bits object; fail \ retype_untyped free_cptr untyped_cptr object_type object_size; obj_id_index \ inc_when (\ fail) obj_id_index; untyped_index \ inc_when fail untyped_index; si_caps \ update_when (\ fail) si_caps obj_id free_cptr; return (obj_id_index, untyped_index, si_caps) od) (0, 0, Map.empty); assert (untyped_index \ length untyped_cptrs); return (si_caps, drop (length [obj\obj_ids. real_object_at obj spec]) free_cptrs) od" (* This makes the IRQ handler caps. * These caps are then used to set up the IRQs. * We need some predicate to say that they are there. * irq \c IrqHandlerCap cnode_id? *) definition create_irq_cap :: "cdl_state \ (cdl_irq \ cdl_cptr) \ unit u_monad" where "create_irq_cap spec \ \(irq, free_cptr). do control_cap \ return seL4_CapIRQControl; root \ return seL4_CapInitThreadCNode; index \ return $ free_cptr; depth \ return (32::word32); fail \ seL4_IRQControl_Get control_cap irq root index depth; assert (\fail) od" definition used_irq_list_compute :: "cdl_state \ cdl_irq list" where "used_irq_list_compute \ \s. filter (HOL.Not o Option.is_none o cdl_objects s o cdl_irq_node s) [0 .e. maxBound]" definition create_irq_caps :: "cdl_state \ cdl_cptr list \ ((cdl_irq \ cdl_cptr option) \ cdl_cptr list) u_monad" where "create_irq_caps spec free_cptrs \ do irqs \ return $ used_irq_list_compute spec; mapM_x (create_irq_cap spec) (zip irqs free_cptrs); si_irq_caps \ return $ map_of (zip irqs free_cptrs); return (si_irq_caps, drop (length irqs) free_cptrs) od" (* This one installs the caps in the CNodes. * It uses seL4 _IRQHandler_SetEndpoint. * It turns irq_empty to irq_initialised? *) definition init_irq :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ (cdl_irq \ cdl_cptr option) \ cdl_irq \ unit u_monad" where "init_irq spec orig_caps irq_caps irq \ do irq_handler_cap \ assert_opt $ irq_caps irq; irq_slot \ return $ get_irq_slot irq spec; endpoint_cap \ assert_opt $ opt_cap irq_slot spec; endpoint_cptr \ assert_opt $ orig_caps (cap_object endpoint_cap); fail \ seL4_IRQHandler_SetEndpoint irq_handler_cap endpoint_cptr; assert (\fail) od" definition init_irqs :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ (cdl_irq \ cdl_cptr option) \ unit u_monad" where "init_irqs spec orig_caps irq_caps \ do irqs \ return $ bound_irq_list spec; mapM_x (init_irq spec orig_caps irq_caps) irqs od" (* Make a duplicate of a cap in a new free slot. *) definition duplicate_cap :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ (cdl_object_id \ cdl_cptr) \ unit u_monad" where "duplicate_cap spec orig_caps \ \(obj_id, free_slot). do rights \ return $ all_cdl_rights; dest_root \ return seL4_CapInitThreadCNode; dest_index \ return $ free_slot; dest_depth \ return (32::word32); src_root \ return seL4_CapInitThreadCNode; src_index \ assert_opt $ orig_caps obj_id; src_depth \ return (32::word32); fail \ seL4_CNode_Copy dest_root dest_index dest_depth src_root src_index src_depth rights; assert (\fail) od" (* As CNode caps can be moved, referencing the CNodes can be tricky. * Moving capabilities in a particular order would work for all be difficult circular cases, * but just keeping a copy of the caps to all of the CNodes and using the copy is much easier. * * Thread caps are duplicated as the system initialiser needs to have them to start the threads * at the end of initialisation. *) definition duplicate_caps :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ cdl_object_id list \ cdl_cptr list \ (cdl_object_id \ cdl_cptr option) u_monad" where "duplicate_caps spec orig_caps obj_ids free_slots \ do obj_ids' \ return [obj_id \ obj_ids. cnode_or_tcb_at obj_id spec]; assert (length obj_ids' \ length free_slots); mapM_x (duplicate_cap spec orig_caps) (zip obj_ids' free_slots); return $ map_of $ zip obj_ids' free_slots od" (* Initialise a single tcb (cspace, vspace and fault_ep). *) definition init_tcb :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ cdl_object_id \ unit u_monad" where "init_tcb spec orig_caps tcb_id \ do cdl_tcb \ assert_opt $ opt_thread tcb_id spec; cdl_cspace_root \ assert_opt $ opt_cap (tcb_id, tcb_cspace_slot) spec; cdl_vspace_root \ assert_opt $ opt_cap (tcb_id, tcb_vspace_slot) spec; cdl_ipcbuffer \ assert_opt $ opt_cap (tcb_id, tcb_ipcbuffer_slot) spec; ipcbuf_addr \ return $ tcb_ipc_buffer_address cdl_tcb; priority \ return $ tcb_priority cdl_tcb; sel4_tcb \ assert_opt $ orig_caps tcb_id; sel4_cspace_root \ assert_opt $ orig_caps (cap_object cdl_cspace_root); sel4_vspace_root \ assert_opt $ orig_caps (cap_object cdl_vspace_root); sel4_ipcbuffer \ assert_opt $ orig_caps (cap_object cdl_ipcbuffer); sel4_fault_ep \ return $ cdl_tcb_fault_endpoint cdl_tcb; sel4_cspace_root_data \ return $ guard_as_rawdata cdl_cspace_root; sel4_vspace_root_data \ return 0; fail \ seL4_TCB_Configure sel4_tcb sel4_fault_ep sel4_cspace_root sel4_cspace_root_data sel4_vspace_root sel4_vspace_root_data ipcbuf_addr sel4_ipcbuffer; assert (\fail) od" (* Set the registers of a single tcb (cspace, vspace and fault_ep). *) definition configure_tcb :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ cdl_object_id \ unit u_monad" where "configure_tcb spec orig_caps tcb_id \ do cdl_tcb \ assert_opt $ opt_thread tcb_id spec; sel4_tcb \ assert_opt $ orig_caps tcb_id; ip \ return $ tcb_ip cdl_tcb; sp \ return $ tcb_sp cdl_tcb; regs \ return [ip, sp]; fail \ seL4_TCB_WriteRegisters sel4_tcb False 0 2 regs; assert fail od" (* Initialise all the tcbs (cspace, vspace and fault_ep). *) definition init_tcbs :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ cdl_object_id list \ unit u_monad" where "init_tcbs spec orig_caps objects \ do tcb_ids \ return [obj \ objects. tcb_at obj spec]; mapM_x (init_tcb spec orig_caps) tcb_ids; mapM_x (configure_tcb spec orig_caps) tcb_ids od" (************************** * VSpace Initialisation. * **************************) (* Sets the asid for a page directory from the initialiser's asid pool. *) definition set_asid :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ cdl_object_id \ unit u_monad" where "set_asid spec orig_caps page_id \ do \ \Set asid pool for the page directory.\ sel4_asid_pool \ return seL4_CapInitThreadASIDPool; sel4_page \ assert_opt $ orig_caps page_id; fail \ seL4_ASIDPool_Assign sel4_asid_pool sel4_page; assert (\fail) od" (* Set the ASIDs for the page directories. *) definition init_pd_asids :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ cdl_object_id list \ unit u_monad" where "init_pd_asids spec orig_caps obj_ids \ do pd_ids \ return [obj_id \ obj_ids. pd_at obj_id spec]; mapM_x (set_asid spec orig_caps) pd_ids od" (* Map a page into a page table or page directory *) definition map_page :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ cdl_object_id \ cdl_object_id \ cdl_right set \ word32 \ cdl_cptr \ unit u_monad" where "map_page spec orig_caps page_id pd_id rights vaddr free_cptr \ do cdl_page \ assert_opt $ cdl_objects spec page_id; sel4_page \ assert_opt $ orig_caps page_id; sel4_pd \ assert_opt $ orig_caps pd_id; vmattribs \ assert_opt $ opt_vmattribs cdl_page; assert (frame_at page_id spec); \ \Copy the frame cap into a new slot for mapping, this enables shared frames\ duplicate_cap spec orig_caps (page_id, free_cptr); seL4_Page_Map free_cptr sel4_pd vaddr rights vmattribs; return () od" (* Map a page table into a page directory *) definition map_page_table :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ cdl_object_id \ cdl_object_id \ cdl_right set \ word32 \ unit u_monad" where "map_page_table spec orig_caps page_id pd_id rights vaddr \ do cdl_page \ assert_opt $ cdl_objects spec page_id; sel4_page \ assert_opt $ orig_caps page_id; sel4_pd \ assert_opt $ orig_caps pd_id; vmattribs \ assert_opt $ opt_vmattribs cdl_page; assert (pt_at page_id spec); seL4_PageTable_Map sel4_page sel4_pd vaddr vmattribs; return () od" (* Map a page table slot *) definition map_page_table_slot :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ cdl_object_id \ cdl_object_id \ word32 \ (cdl_cap_ref \ cdl_cptr) \ cdl_cnode_index \ unit u_monad" where "map_page_table_slot spec orig_caps pd_id pt_id pt_vaddr free_cptr pt_slot \ do frame_cap \ assert_opt $ opt_cap (pt_id, pt_slot) spec; page \ return $ cap_object frame_cap; \ \The page's virtual address is the page table's virtual address, plus the relative address of the page in the page table. Each page stores 12 bits of memory.\ page_vaddr \ return $ pt_vaddr + (of_nat pt_slot << small_frame_size); page_rights \ return (cap_rights frame_cap); if frame_cap \ NullCap then do cptr \ return $ free_cptr (pt_id, pt_slot); map_page spec orig_caps page pd_id page_rights page_vaddr cptr od else return () od" (* Map a page directory slot (and all its contents if it points to a page table) *) definition map_page_directory_slot :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ cdl_object_id \ (cdl_cap_ref \ cdl_cptr) \ cdl_cnode_index \ unit u_monad" where "map_page_directory_slot spec orig_caps pd_id free_cptrs pd_slot \ do page_cap \ assert_opt $ opt_cap (pd_id, pd_slot) spec; page_id \ return $ cap_object page_cap; \ \The page table's virtual address is given by the number of slots and how much memory each maps. Each page directory slot maps 10+12 bits of memory (by the page table and page respectively).\ page_vaddr \ return $ of_nat pd_slot << (pt_size + small_frame_size); page_rights \ return (cap_rights page_cap); if pt_at page_id spec then do map_page_table spec orig_caps page_id pd_id page_rights page_vaddr; page_slots \ return $ slots_of_list spec page_id; mapM_x (map_page_table_slot spec orig_caps pd_id page_id page_vaddr free_cptrs) [slot <- page_slots. cap_at ((\) NullCap) (page_id, slot) spec] od else if frame_at page_id spec then do cptr \ return $ free_cptrs (pd_id, pd_slot); map_page spec orig_caps page_id pd_id page_rights page_vaddr cptr od else return () od" (* Map a page directory and all its contents *) definition map_page_directory :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ (cdl_cap_ref \ cdl_cptr) \ cdl_object_id \ unit u_monad" where "map_page_directory spec orig_caps free_cptrs pd_id \ do pd_slots \ return $ slots_of_list spec pd_id; mapM_x (map_page_directory_slot spec orig_caps pd_id free_cptrs) [pd_slot <- pd_slots. cap_at (\cap. cap \ NullCap \ frame_at (cap_object cap) spec) (pd_id, pd_slot) spec]; mapM_x (map_page_directory_slot spec orig_caps pd_id free_cptrs) [pd_slot <- pd_slots. cap_at (\cap. cap \ NullCap \ pt_at (cap_object cap) spec) (pd_id, pd_slot) spec] od" (* Get a list of all the (object_id, slot) pairs contained within a page directory and its child page tables which contain frame caps *) definition get_frame_caps :: "cdl_state \ cdl_object_id \ (cdl_object_id \ nat) list" where "get_frame_caps spec pd_id \ do { pd_slot <- slots_of_list spec pd_id; obj_cap <- case_option [] (\x. [x]) $ opt_cap (pd_id, pd_slot) spec; let obj_id = cap_object obj_cap; if frame_at obj_id spec then [(pd_id, pd_slot)] else \ \If the slot contains a page table cap, collect all its slots that contain frames\ map (Pair obj_id) o filter (\slot. cap_at ((\) NullCap) (obj_id, slot) spec) $ slots_of_list spec obj_id }" (* Construct a function to assign each frame cap a free cptr, mapping all other slots to 0 *) definition make_frame_cap_map :: "cdl_object_id list \ word32 list \ cdl_state \ cdl_object_id \ nat \ word32" where "make_frame_cap_map obj_ids free_cptrs spec ref \ case_option 0 id (map_of (zip ([obj_id \ obj_ids. pd_at obj_id spec] \ get_frame_caps spec) free_cptrs) ref)" (* Maps page directories and all their contents. *) definition init_vspace :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ cdl_object_id list \ cdl_cptr list \ unit u_monad" where "init_vspace spec orig_caps obj_ids free_cptrs \ do pd_ids \ return [obj_id \ obj_ids. pd_at obj_id spec]; cptr_map \ return $ make_frame_cap_map obj_ids free_cptrs spec; mapM_x (map_page_directory spec orig_caps cptr_map) pd_ids od" (************************** * CSpace Initialisation. * **************************) datatype init_cnode_mode = Move | Copy (* initialises a CNode slot with the desired cap, either moving them or copying it. *) definition init_cnode_slot :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ (cdl_object_id \ cdl_cptr option) \ (cdl_irq \ cdl_cptr option) \ init_cnode_mode \ cdl_object_id \ cdl_cnode_index \ bool u_monad" where "init_cnode_slot spec orig_caps dup_caps irq_caps mode cnode_id cnode_slot \ do target_cap \ assert_opt (opt_cap (cnode_id, cnode_slot) spec); target_cap_obj \ return (cap_object target_cap); target_cap_irq \ return (cap_irq target_cap); target_cap_rights \ return (cap_rights target_cap); \ \for endpoint this is the badge, for cnodes, this is the (encoded) guard\ target_cap_data \ return (cap_data target_cap); is_ep_cap \ return (ep_related_cap target_cap); is_irqhandler_cap \ return (is_irqhandler_cap target_cap); \ \ny original caps (which includes IRQ Hander caps) are moved, not copied.\ move_cap \ return (original_cap_at (cnode_id, cnode_slot) spec); dest_obj \ get_spec_object spec cnode_id; dest_size \ return (object_size_bits dest_obj); \ \Use a copy of the cap to reference the destination, in case the original has already been moved.\ dest_root \ assert_opt (dup_caps cnode_id); dest_index \ return (of_nat cnode_slot); dest_depth \ return (of_nat dest_size); \ \Use an original cap to reference the object to copy.\ src_root \ return seL4_CapInitThreadCNode; src_index \ assert_opt (if is_irqhandler_cap then irq_caps target_cap_irq else orig_caps target_cap_obj); src_depth \ return (32::word32); \ \If it's a NullCap, then there's no need to do anything. d If it's a cap that wants to be moved, and we want to move the cap, move (mutate) it. If it's a cap that wants to be copied, and we want to copy it, then copy (mint) it. \ if (target_cap = NullCap) then return True else if (mode = Move \ move_cap) then ( if (is_ep_cap \ is_irqhandler_cap) then (seL4_CNode_Move dest_root dest_index dest_depth src_root src_index src_depth) else (seL4_CNode_Mutate dest_root dest_index dest_depth src_root src_index src_depth target_cap_data) ) else if (mode = Copy \ \ move_cap) then seL4_CNode_Mint dest_root dest_index dest_depth src_root src_index src_depth target_cap_rights target_cap_data else return True od" (* Initialises a CNode with it's desired caps, either moving them or copying them. *) definition init_cnode :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ (cdl_object_id \ cdl_cptr option) \ (cdl_irq \ cdl_cptr option) \ init_cnode_mode \ cdl_object_id \ unit u_monad" where "init_cnode spec orig_caps dup_caps irq_caps mode cnode_id \ do cnode_slots \ return $ slots_of_list spec cnode_id; mapM_x (init_cnode_slot spec orig_caps dup_caps irq_caps mode cnode_id) cnode_slots od" (* Initialises CNodes with their desired caps in two passes. First pass copies caps, the second moves them. *) definition init_cspace :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ (cdl_object_id \ cdl_cptr option) \ (cdl_irq \ cdl_cptr option) \ cdl_object_id list \ unit u_monad" where "init_cspace spec orig_caps dup_caps irq_caps obj_ids \ do cnode_ids \ return [obj_id \ obj_ids. cnode_at obj_id spec]; mapM_x (init_cnode spec orig_caps dup_caps irq_caps Copy) cnode_ids; mapM_x (init_cnode spec orig_caps dup_caps irq_caps Move) cnode_ids od" (* Initialise a single tcb (cspace, vspace and fault_ep). *) definition start_thread :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ cdl_object_id \ unit u_monad" where "start_thread spec dup_caps tcb_id \ do sel4_tcb \ assert_opt $ dup_caps tcb_id; fail \ seL4_TCB_Resume sel4_tcb; assert fail od" definition start_threads :: "cdl_state \ (cdl_object_id \ cdl_cptr option) \ cdl_object_id list \ unit u_monad" where "start_threads spec orig_caps obj_ids \ do tcbs \ return [obj_id \ obj_ids. is_waiting_thread_at obj_id spec]; mapM_x (start_thread spec orig_caps) tcbs od" definition init_system :: "cdl_state \ cdl_bootinfo \ cdl_object_id list \ unit u_monad" where "init_system spec bootinfo obj_ids \ do (untyped_cptrs, free_cptrs) \ parse_bootinfo bootinfo; (orig_caps, free_cptrs) \ create_objects spec obj_ids untyped_cptrs free_cptrs; (irq_caps, free_cptrs) \ create_irq_caps spec free_cptrs; dup_caps \ duplicate_caps spec orig_caps obj_ids free_cptrs; init_irqs spec orig_caps irq_caps; init_pd_asids spec orig_caps obj_ids; init_vspace spec orig_caps obj_ids free_cptrs; init_tcbs spec orig_caps obj_ids; init_cspace spec orig_caps dup_caps irq_caps obj_ids; start_threads spec dup_caps obj_ids od" end