(* * Copyright 2014, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) theory DataMap imports "~~/src/HOL/Main" begin type_synonym ('k, 'a) map = "'k \ 'a" definition data_map_empty :: "('k, 'a) map" where data_map_empty_def[simp]: "data_map_empty \ Map.empty" definition data_map_insert :: "'k \ 'a \ ('k, 'a) map \ ('k, 'a) map" where data_map_insert_def[simp]: "data_map_insert x y m \ (m (x \ y))" definition lookupBefore :: "('k :: {linorder, finite}) \ ('k, 'a) map \ ('k \ 'a) option" where "lookupBefore x m \ let Ks = {k \ dom m. k \ x} in if Ks = {} then None else Some (Max Ks, the (m (Max Ks)))" definition lookupAround :: "('k :: {linorder, finite}) \ ('k, 'a) map \ (('k \ 'a) option \ 'a option \ ('k \ 'a) option)" where "lookupAround x m \ let Bs = {k \ dom m. k < x}; As = {k \ dom m. k > x} in (if Bs = {} then None else Some (Max Bs, the (m (Max Bs))), m x, if As = {} then None else Some (Min As, the (m (Min Bs))))" definition data_map_filterWithKey :: "('k \ 'a \ bool) \ ('k, 'a) map \ ('k, 'a) map" where "data_map_filterWithKey f m \ \x. case m x of None \ None | Some y \ if f x y then Some y else None" end