(* * Copyright 2014, General Dynamics C4 Systems * * SPDX-License-Identifier: GPL-2.0-only *) (* Functions for fault handling. *) chapter \arch fault related functions\ theory ArchFault_A imports Structures_A Tcb_A begin context Arch begin global_naming ARM_HYP_A fun make_arch_fault_msg :: "arch_fault \ obj_ref \ (data \ data list,'z::state_ext) s_monad" where "make_arch_fault_msg (VMFault vptr archData) thread = do pc \ as_user thread getRestartPC; return (5, pc # vptr # archData) od" | "make_arch_fault_msg (VCPUFault hsr) thread = return (7, [hsr])" | "make_arch_fault_msg (VPPIEvent irq) thread = return (8, [ucast irq])" | "make_arch_fault_msg (VGICMaintenance archData) thread = do msg \ return $ (case archData of None \ [-1] | Some idx \ [idx]); return (6, msg) od" definition handle_arch_fault_reply :: "arch_fault \ obj_ref \ data \ data list \ (bool,'z::state_ext) s_monad" where "handle_arch_fault_reply af thread x y = return True" end end