(* * Copyright 2014, General Dynamics C4 Systems * * SPDX-License-Identifier: GPL-2.0-only *) chapter "Platform Definitions" theory Platform imports "Lib.Defs" "Lib.Lib" "Word_Lib_l4v.WordSetup" Setup_Locale Kernel_Config begin context Arch begin global_naming ARM_HYP text \ This theory lists platform-specific types and basic constants, in particular the types of interrupts and physical addresses, constants for the kernel location, the offsets between physical and virtual kernel addresses, as well as the range of IRQs on the platform. \ type_synonym irq = "10 word" type_synonym paddr = word32 abbreviation (input) "toPAddr \ id" abbreviation (input) "fromPAddr \ id" definition pageColourBits :: nat where "pageColourBits \ 2" definition cacheLineBits :: nat where "cacheLineBits = 6" definition cacheLine :: nat where "cacheLine = 2^cacheLineBits" (* The first virtual address of the kernel's physical memory window *) definition pptrBase :: word32 where "pptrBase \ 0xe0000000" abbreviation (input) "paddrBase \ physBase" definition pptrBaseOffset :: machine_word where "pptrBaseOffset = pptrBase - paddrBase" definition pptrTop :: "32 word" where "pptrTop \ 0xfff00000" definition paddrTop :: "32 word" where "paddrTop \ pptrTop - pptrBaseOffset" definition kernelELFPAddrBase :: word32 where "kernelELFPAddrBase \ physBase" definition kernelELFBase :: word32 where "kernelELFBase \ pptrBase + (kernelELFPAddrBase && mask 22)" definition kernelELFBaseOffset :: word32 where "kernelELFBaseOffset \ kernelELFBase - kernelELFPAddrBase" definition ptrFromPAddr :: "paddr \ word32" where "ptrFromPAddr paddr \ paddr + pptrBaseOffset" definition addrFromPPtr :: "word32 \ paddr" where "addrFromPPtr pptr \ pptr - pptrBaseOffset" definition addrFromKPPtr :: "word32 \ paddr" where "addrFromKPPtr kpptr \ kpptr - kernelELFBaseOffset" definition minIRQ :: "irq" where "minIRQ \ 0" definition maxIRQ :: "irq" where "maxIRQ \ 191" definition irqVGICMaintenance :: "irq" where "irqVGICMaintenance \ 25" definition irqVTimerEvent :: "irq" where "irqVTimerEvent \ 27" end end