(* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: GPL-2.0-only *) theory ArchInfoFlow imports "Access.ArchSyscall_AC" "Lib.EquivValid" begin context Arch begin global_naming RISCV64 section \Arch-specific equivalence properties\ subsection \ASID equivalence\ definition equiv_asid :: "asid \ det_ext state \ det_ext state \ bool" where "equiv_asid asid s s' \ ((riscv_asid_table (arch_state s) (asid_high_bits_of asid)) = (riscv_asid_table (arch_state s') (asid_high_bits_of asid))) \ (\pool_ptr. riscv_asid_table (arch_state s) (asid_high_bits_of asid) = Some pool_ptr \ asid_pool_at pool_ptr s = asid_pool_at pool_ptr s' \ (\asid_pool asid_pool'. asid_pools_of s pool_ptr = Some asid_pool \ asid_pools_of s' pool_ptr = Some asid_pool' \ asid_pool (asid_low_bits_of asid) = asid_pool' (asid_low_bits_of asid)))" definition equiv_asid' where "equiv_asid' asid pool_ptr_opt pool_ptr_opt' kh kh' \ (case pool_ptr_opt of None \ pool_ptr_opt' = None | Some pool_ptr \ (case pool_ptr_opt' of None \ False | Some pool_ptr' \ (pool_ptr' = pool_ptr \ ((\asid_pool. kh pool_ptr = Some (ArchObj (ASIDPool asid_pool))) = (\asid_pool'. kh' pool_ptr' = Some (ArchObj (ASIDPool asid_pool')))) \ (\asid_pool asid_pool'. kh pool_ptr = Some (ArchObj (ASIDPool asid_pool)) \ kh' pool_ptr' = Some (ArchObj (ASIDPool asid_pool')) \ asid_pool (asid_low_bits_of asid) = asid_pool' (asid_low_bits_of asid)))))" definition non_asid_pool_kheap_update where "non_asid_pool_kheap_update s kh \ \x. (\asid_pool. kheap s x = Some (ArchObj (ASIDPool asid_pool)) \ kh x = Some (ArchObj (ASIDPool asid_pool))) \ kheap s x = kh x" subsection \Exclusive machine state equivalence\ subsection \Global (Kernel) VSpace equivalence\ (* globals_equiv should be maintained by everything except the scheduler, since nothing else touches the globals frame *) definition arch_globals_equiv :: "obj_ref \ obj_ref \ kheap \ kheap \ arch_state \ arch_state \ machine_state \ machine_state \ bool" where "arch_globals_equiv ct it kh kh' as as' ms ms' \ riscv_global_pt as = riscv_global_pt as' \ kh (riscv_global_pt as) = kh' (riscv_global_pt as)" declare arch_globals_equiv_def[simp] end requalify_consts RISCV64.equiv_asid RISCV64.equiv_asid' RISCV64.arch_globals_equiv RISCV64.non_asid_pool_kheap_update end