(* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * Copyright (c) 2022 Apple Inc. All rights reserved. * * SPDX-License-Identifier: BSD-2-Clause *) (* Structures supporting CTypes. Primarily sets up types, defines pointers and the raw heap view. *) theory CTypesBase imports Addr_Type begin section "Type setup" type_synonym byte = "8 word" type_synonym memory = "addr ⇒ byte" type_synonym 'a mem_upd = "addr ⇒ 'a ⇒ memory ⇒ memory" type_synonym 'a mem_read = "addr ⇒ memory ⇒ 'a" class unit_class = assumes there_is_only_one: "x = y" instantiation unit :: unit_class begin instance by (intro_classes, simp) end subsection "Pointers"