(* * Copyright 2014, General Dynamics C4 Systems * * This software may be distributed and modified according to the terms of * the GNU General Public License version 2. Note that NO WARRANTY is provided. * See "LICENSE_GPLv2.txt" for details. * * @TAG(GD_GPL) *) chapter "Platform Definitions" theory Platform imports "Lib.Defs" "Lib.Lib" "Word_Lib.WordSetup" "../Setup_Locale" 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" definition kernelBase_addr :: word32 where "kernelBase_addr \ 0xe0000000" (* Arch specific kernel base address used for haskell spec *) definition kernelBase :: word32 where "kernelBase \ 0xe0000000" definition physBase :: word32 where "physBase \ 0x80000000" definition pptrTop :: "32 word" where "pptrTop \ 0xfff00000" definition paddrTop :: "32 word" where "paddrTop \ pptrTop - (kernelBase - physBase)" definition physMappingOffset :: word32 where "physMappingOffset \ kernelBase_addr - physBase" definition ptrFromPAddr :: "paddr \ word32" where "ptrFromPAddr paddr \ paddr + physMappingOffset" definition addrFromPPtr :: "word32 \ paddr" where "addrFromPPtr pptr \ pptr - physMappingOffset" 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