(* * Copyright 2014, General Dynamics C4 Systems * Copyright 2022, Proofcraft Pty Ltd * * SPDX-License-Identifier: GPL-2.0-only *) chapter \VCPU\ theory VCPU_A imports TcbAcc_A InvocationLabels_A begin text \This is used by some decode functions. VCPU decode functions are the first that need to bounds check IRQs from the user.\ definition arch_check_irq :: "data \ (unit,'z::state_ext) se_monad" where "arch_check_irq irq \ whenE (irq > ucast maxIRQ \ irq < ucast minIRQ) $ throwError (RangeError (ucast minIRQ) (ucast maxIRQ))" context Arch begin global_naming AARCH64_A section "VCPU" subsection "VCPU: Set TCB" definition decode_vcpu_set_tcb :: "arch_cap \ (cap \ cslot_ptr) list \ (arch_invocation,'z::state_ext) se_monad" where "decode_vcpu_set_tcb cap extras \ case (cap, extras) of (VCPUCap v, fs#_) \ (case fs of (ThreadCap t, _) \ returnOk $ InvokeVCPU $ VCPUSetTCB v t | _ \ throwError IllegalOperation) | (VCPUCap v, _) \ throwError TruncatedMessage | _ \ throwError IllegalOperation" text \VCPU objects can be associated with and dissociated from TCBs. It is not possible to dissociate a VCPU and a TCB by using setTCB. Final outcome has to be an associated TCB and VCPU. The only way to get lasting dissociation is to delete the TCB or the VCPU. See ArchVSpace\_A.\ subsection "VCPU: Read/Write Registers" definition read_vcpu_register :: "obj_ref \ vcpureg \ (machine_word,'z::state_ext) s_monad" where "read_vcpu_register vcpu_ptr reg \ do cur_vcpu \ gets (arm_current_vcpu \ arch_state); (on_cur_vcpu, active) \ return (case cur_vcpu of Some (vcpu_ptr', a) \ (vcpu_ptr' = vcpu_ptr, a) | _ \ (False, False)); if on_cur_vcpu then if vcpuRegSavedWhenDisabled reg \ \active then vcpu_read_reg vcpu_ptr reg else do_machine_op $ readVCPUHardwareReg reg else vcpu_read_reg vcpu_ptr reg od" definition write_vcpu_register :: "obj_ref \ vcpureg \ machine_word \ (unit,'z::state_ext) s_monad" where "write_vcpu_register vcpu_ptr reg val \ do cur_vcpu \ gets (arm_current_vcpu o arch_state); (on_cur_vcpu, active) \ return (case cur_vcpu of Some (cv, a) \ (cv = vcpu_ptr, a) | _ \ (False, False)); if on_cur_vcpu then if vcpuRegSavedWhenDisabled reg \ \active then vcpu_write_reg vcpu_ptr reg val else do_machine_op $ writeVCPUHardwareReg reg val else vcpu_write_reg vcpu_ptr reg val od" definition decode_vcpu_read_register :: "machine_word list \ arch_cap \ (arch_invocation,'z::state_ext) se_monad" where "decode_vcpu_read_register args cap \ case (args, cap) of (reg#_, VCPUCap p) \ if fromEnum (maxBound::vcpureg) < unat reg then throwError (InvalidArgument 1) else returnOk $ InvokeVCPU $ VCPUReadRegister p $ toEnum (unat reg) | (_, _) \ throwError TruncatedMessage" definition decode_vcpu_write_register :: "machine_word list \ arch_cap \ (arch_invocation,'z::state_ext) se_monad" where "decode_vcpu_write_register args cap \ case (args, cap) of (reg#val#_, VCPUCap p) \ if fromEnum (maxBound::vcpureg) < unat reg then throwError (InvalidArgument 1) else returnOk $ InvokeVCPU $ VCPUWriteRegister p (toEnum (unat reg)) val | (_, _) \ throwError TruncatedMessage" definition invoke_vcpu_read_register :: "obj_ref \ vcpureg \ (data list, 'z::state_ext) s_monad" where "invoke_vcpu_read_register v reg \ do val \ read_vcpu_register v reg; return [val] od" definition invoke_vcpu_write_register :: "obj_ref \ vcpureg \ machine_word \ (unit,'z::state_ext) s_monad" where "invoke_vcpu_write_register v reg val \ write_vcpu_register v reg val" text \VCPU : inject IRQ\ (* This following function does not correspond to exactly what the C does, but it is the value that is stored inside of lr in the vgic *) definition make_virq :: "obj_ref \ obj_ref \ obj_ref \ virq" where "make_virq grp prio irq \ let groupShift = 30; prioShift = 23; irqPending = 1 << 28; eoiirqen = 1 << 19 in ((grp && 1) << groupShift) || ((prio && 0x1F) << prioShift) || (irq && 0x3FF) || irqPending || eoiirqen" definition virq_type :: "virq \ nat" where "virq_type virq \ unat ((virq >> 28) && 3)" definition is_virq_active :: "virq \ bool" where "is_virq_active virq \ virq_type virq = 2" definition decode_vcpu_inject_irq :: "obj_ref list \ arch_cap \ (arch_invocation,'z::state_ext) se_monad" where "decode_vcpu_inject_irq ptrs cap \ case (ptrs, cap) of (mr0 # _, VCPUCap p) \ doE vid \ returnOk (mr0 && 0xFFFF); priority \ returnOk ((mr0 >> 16) && 0xFF); group \ returnOk ((mr0 >> 24) && 0xFF); index \ returnOk ((mr0 >> 32) && 0xFF); range_check vid 0 ((1 << 10) - 1); range_check priority 0 31; range_check group 0 1; num_list_regs \ liftE $ gets (arm_gicvcpu_numlistregs \ arch_state); whenE (index \ of_nat num_list_regs) $ (throwError $ RangeError 0 (of_nat num_list_regs - 1)); vcpu \ liftE $ get_vcpu p; vcpuLR \ returnOk (vgic_lr $ vcpu_vgic $ vcpu); whenE (is_virq_active (vcpuLR (unat index))) $ throwError DeleteFirst; virq \ returnOk (make_virq group priority vid); returnOk $ InvokeVCPU $ VCPUInjectIRQ p (unat index) virq odE | _ \ throwError TruncatedMessage" definition invoke_vcpu_inject_irq :: "obj_ref \ nat \ virq \ (unit,'z::state_ext) s_monad" where "invoke_vcpu_inject_irq vr index virq \ do cur_v \ gets (arm_current_vcpu \ arch_state); if (cur_v \ None \ fst (the cur_v) = vr) then do_machine_op $ set_gic_vcpu_ctrl_lr (of_nat index) virq else vgic_update_lr vr index virq od" text \VCPU : acknowledge VPPI\ definition decode_vcpu_ack_vppi :: "obj_ref list \ arch_cap \ (arch_invocation,'z::state_ext) se_monad" where "decode_vcpu_ack_vppi mrs cap \ case (mrs, cap) of (mr0 # _, VCPUCap vcpu_ptr) \ doE arch_check_irq mr0; (case irq_vppi_event_index (ucast mr0) of None \ throwError $ InvalidArgument 0 | Some vppi \ returnOk $ InvokeVCPU $ VCPUAckVPPI vcpu_ptr vppi) odE | _ \ throwError TruncatedMessage" definition invoke_vcpu_ack_vppi :: "obj_ref \ vppievent_irq \ (unit,'z::state_ext) s_monad" where "invoke_vcpu_ack_vppi vcpu_ptr vppi = vcpu_update vcpu_ptr (\vcpu. vcpu\ vcpu_vppi_masked := (vcpu_vppi_masked vcpu)(vppi := False) \)" text \VCPU perform and decode main functions\ definition perform_vcpu_invocation :: "vcpu_invocation \ (data list,'z::state_ext) s_monad" where "perform_vcpu_invocation iv \ case iv of VCPUSetTCB vcpu tcb \ do associate_vcpu_tcb vcpu tcb; return [] od | VCPUReadRegister vcpu reg \ invoke_vcpu_read_register vcpu reg | VCPUWriteRegister vcpu reg val \ do invoke_vcpu_write_register vcpu reg val; return [] od | VCPUInjectIRQ vcpu index vir \ do invoke_vcpu_inject_irq vcpu index vir; return [] od | VCPUAckVPPI vcpu vppi \ do invoke_vcpu_ack_vppi vcpu vppi; return [] od" definition decode_vcpu_invocation :: "machine_word \ machine_word list \ arch_cap \ (cap \ cslot_ptr) list \ (arch_invocation,'z::state_ext) se_monad" where "decode_vcpu_invocation label args cap extras \ case cap of VCPUCap _ \ (case invocation_type label of ArchInvocationLabel ARMVCPUSetTCB \ decode_vcpu_set_tcb cap extras | ArchInvocationLabel ARMVCPUReadReg \ decode_vcpu_read_register args cap | ArchInvocationLabel ARMVCPUWriteReg \ decode_vcpu_write_register args cap | ArchInvocationLabel ARMVCPUInjectIRQ \ decode_vcpu_inject_irq args cap | ArchInvocationLabel ARMVCPUAckVPPI \ decode_vcpu_ack_vppi args cap | _ \ throwError IllegalOperation) | _ \ throwError IllegalOperation" end end