(* * Copyright 2014, General Dynamics C4 Systems * * SPDX-License-Identifier: GPL-2.0-only *) (* Dummy initial kernel state. Real kernel boot up is more complex. *) chapter "An Initial Kernel State" theory Init_A imports Retype_A begin context Arch begin global_naming ARM_A text \ This is not a specification of true kernel initialisation. This theory describes a dummy initial state only, to show that the invariants and refinement relation are consistent. \ (* Moved to Deterministic_A definition idle_thread_ptr :: word32 where "idle_thread_ptr = kernel_base + 0x1000" *) definition init_tcb_ptr :: word32 where "init_tcb_ptr = kernel_base + 0x2000" definition init_irq_node_ptr :: word32 where "init_irq_node_ptr = kernel_base + 0x8000" definition init_globals_frame :: word32 where "init_globals_frame = kernel_base + 0x5000" definition "us_global_pd_ptr = kernel_base + 0x60000" definition "init_arch_state \ \ arm_asid_table = Map.empty, arm_hwasid_table = Map.empty, arm_next_asid = 0, arm_asid_map = Map.empty, arm_current_vcpu = None, arm_gicvcpu_numlistregs = undefined, arm_kernel_vspace = (\ref. if ref \ {kernel_base .. kernel_base + mask 20} then ArmVSpaceKernelWindow else ArmVSpaceInvalidRegion), arm_us_global_pd = us_global_pd_ptr \" definition [simp]: "us_global_pd \ ArchObj (PageDirectory (\_. InvalidPDE))" definition "init_kheap \ (\x. if \irq :: irq. init_irq_node_ptr + (ucast irq << cte_level_bits) = x then Some (CNode 0 (empty_cnode 0)) else None) (idle_thread_ptr \ TCB \ tcb_ctable = NullCap, tcb_vtable = NullCap, tcb_reply = NullCap, tcb_caller = NullCap, tcb_ipcframe = NullCap, tcb_state = IdleThreadState, tcb_fault_handler = replicate word_bits False, tcb_ipc_buffer = 0, tcb_fault = None, tcb_bound_notification = None, tcb_mcpriority = minBound, tcb_arch = init_arch_tcb \, us_global_pd_ptr \ us_global_pd)" definition "init_cdt \ Map.empty" definition "init_ioc \ \(a,b). (\obj. init_kheap a = Some obj \ (\cap. cap_of obj b = Some cap \ cap \ cap.NullCap))" definition "init_A_st \ \ kheap = init_kheap, cdt = init_cdt, is_original_cap = init_ioc, cur_thread = idle_thread_ptr, idle_thread = idle_thread_ptr, machine_state = init_machine_state, interrupt_irq_node = \irq. init_irq_node_ptr + (ucast irq << cte_level_bits), interrupt_states = \_. Structures_A.IRQInactive, arch_state = init_arch_state, exst = ext_init \" end end