lh-l4v/spec/design/skel/PSpaceFuns_H.thy

42 lines
1.3 KiB
Plaintext
Raw Normal View History

2014-07-14 19:32:44 +00:00
(*
* Copyright 2014, General Dynamics C4 Systems
*
2020-03-09 06:18:30 +00:00
* SPDX-License-Identifier: GPL-2.0-only
2014-07-14 19:32:44 +00:00
*)
2015-05-16 09:52:49 +00:00
chapter "Physical Memory Functions"
2014-07-14 19:32:44 +00:00
theory PSpaceFuns_H
imports
ObjectInstances_H
FaultMonad_H
ArchPSpace_H
"Lib.DataMap"
2014-07-14 19:32:44 +00:00
begin
context begin interpretation Arch .
requalify_consts
fromPPtr
PPtr
freeMemory
storeWord
loadWord
end
requalify_consts (in Arch)
deleteGhost
definition deleteRange :: "( machine_word , 'a ) DataMap.map \<Rightarrow> machine_word \<Rightarrow> nat \<Rightarrow> ( machine_word , 'a ) DataMap.map"
2017-07-12 05:13:51 +00:00
where "deleteRange m ptr bits \<equiv>
let inRange = (\<lambda> x. x && ((- mask bits) - 1) = fromPPtr ptr) in
data_map_filterWithKey (\<lambda> x _. Not (inRange x)) m"
2016-01-18 04:34:34 +00:00
#INCLUDE_HASKELL SEL4/Model/PSpace.lhs decls_only Data.Map=DataMap NOT PSpace ptrBits ptrBitsForSize lookupAround maybeToMonad lookupAround2 typeError alignError alignCheck sizeCheck objBits deleteRange
2014-07-14 19:32:44 +00:00
consts
lookupAround2 :: "('k :: {linorder,finite}) \<Rightarrow> ( 'k , 'a ) DataMap.map \<Rightarrow> (('k * 'a) option * 'k option)"
#INCLUDE_HASKELL SEL4/Model/PSpace.lhs bodies_only Data.Map=DataMap NOT PSpace ptrBits ptrBitsForSize lookupAround maybeToMonad typeError alignError alignCheck sizeCheck objBits deletionIsSafe cNodePartialOverlap pointerInUserData ksASIDMapSafe deleteRange
2014-07-14 19:32:44 +00:00
end