Theory Shallow_Micro_Rust.Global_Store

(* Copyright Amazon.com, Inc. or its affiliates. All Rights Reserved.
   SPDX-License-Identifier: MIT *)

(*<*)
theory Global_Store
  imports Lenses_And_Other_Optics.Lenses_And_Other_Optics Misc.Array Misc.Vector
      Core_Expression Result_Type (* Range_Type *)
begin
(*>*)

section‹References, and the global store›

subsection‹References›

text‹References ‹references› are basically pointers that carry around an address that points-into a
global store, such as physical memory or a logical runtime heap.
References of this sort are familiar from functional programming languages like Standard ML and OCaml,
and allow us to introduce a structured form of mutability into our Micro Rust language without having
to deal with how values are encoded into a byte-addressable memory.

References come in two flavors: untyped and typed. ‹Untyped references› (or ‹raw references›) are essentially pointers
into a monomorphic global store, without information of how the global value behind the reference is to be
interpreted. For example, a raw reference to a memory-backed global store would dereference to a list
of raw bytes.

‹Typed references› are constructed as "focused raw references": They attach to a raw reference a focus
specifying how the raw global value behind the reference can be interpreted as a concrete value type.
For example, the focus attached to a typed reference into a memory-backed global store encapsulates
the byte-level presentation of the reference's value type.

Below, the fixed type typ'b› represents the "global" values that the global store holds, while the
typ'v› represents the type of the value actually stored by this reference. When creating a typed
reference one must provide a focus between typ'b and typ'v.›

text‹Untyped/raw references are just wrappers around raw addresses. The global value type is
encoded in the global reference's type.›
datatype_record ('a, 'b) gref =
  gref_address :: 'a

text‹A wrapper around a reference which only allows dereference, not update.›
datatype ('a, 'b) ro_gref =
  ROGRef (unsafe_gref_from_ro_gref: ('a, 'b) gref)

text‹Typed references are references paired with a focus onto a subquotient of the global value
type.›
type_synonym ('a, 'b, 'v) ref = (('a, 'b) gref, 'b, 'v) focused
type_synonym ('a, 'b, 'v) ro_ref = (('a, 'b) ro_gref, 'b, 'v) focused

abbreviation untype_ref where untype_ref  unwrap_focused
abbreviation untype_roref where untype_roref  unwrap_focused

translations
  "CONST unwrap_focused"  "CONST untype_ref"
  "CONST unwrap_focused"  "CONST untype_roref"                                                        

definition ro_ref_from_ref :: ('a, 'b, 'v) ref  ('a, 'b, 'v) ro_ref where
  ro_ref_from_ref  update_unwrap_focused ROGRef

definition unsafe_ref_from_ro_ref :: ('a, 'b, 'v) ro_ref  ('a, 'b, 'v) ref where
  unsafe_ref_from_ro_ref  update_unwrap_focused unsafe_gref_from_ro_gref

abbreviation ref_address :: ('a, 'b, 'v) ref  'a where
  ref_address r  gref_address (unwrap_focused r)

abbreviation ro_address :: ('a,'b,'v) ro_ref  'a where
  ro_address r  gref_address (unsafe_gref_from_ro_gref (unwrap_focused r))

consts
  address :: 'a  'b
adhoc_overloading
  address  gref_address ref_address ro_address

definition make_untyped_ref :: 'a  ('a, 'v) gref where
  make_untyped_ref addr  make_gref addr

abbreviation make_ref_typed_from_untyped ::
    ('a, 'v) gref  ('v,'w) focus  ('a, 'v, 'w) ref where
  make_ref_typed_from_untyped  make_focused

definition make_ro_ref_typed_from_untyped ::
    ('a, 'v) ro_gref  ('v,'w) focus  ('a, 'v, 'w) ro_ref where
  make_ro_ref_typed_from_untyped  make_focused

subsection ‹Using optics to focus references and values›

abbreviation focus_reference :: ('v, 'w) focus  ('a, 'b, 'v) ref  ('a, 'b, 'w) ref where
  focus_reference  focus_focused

abbreviation focus_reference_via_lens l  focus_reference (lens_to_focus l)

lemmas focus_reference_def = focus_focused_def

―‹This empty type class serves as an annotation for types which can appear as the global type
  of a lens. The main example are records, and the main non-example are references. In fact, the
  sole purpose of this class is to help disambiguate the overloaded ‹.›-operator in uRust in case
  the LHS is a reference. In this case, it is a priori not clear if we're focusing the LHS with respect
  to a lens between reference types, or with respect to the lifting of a lens on the value type of the ref.
  We always want the latter, and the ‹localizable› typeclass helps to enforce it.›
class localizable
abbreviation (input) focus_lens_value :: ('v::localizable, 'w) lens  'v  'w where
  focus_lens_value  lens_view

consts
  focus_lens_const :: ('v, 'w) lens  'vp  'wp

adhoc_overloading
  focus_lens_const  focus_reference_via_lens focus_lens_value

subsection‹Focusing references inside structured values›

abbreviation (input) focus_option :: ('a,'b,'v option) ref  ('a,'b,'v) ref where
  focus_option  focus_focused option_focus

abbreviation (input) focus_result_ok :: ('a,'b, ('v,'e) result) ref  ('a,'b,'v) ref where
  focus_result_ok  focus_focused result_ok_focus

abbreviation (input) focus_result_err :: ('a,'b, ('v,'e) result) ref  ('a,'b,'e) ref where
  focus_result_err  focus_focused result_err_focus

abbreviation (input) focus_nth :: nat  ('a,'b,'v list) ref  ('a,'b,'v) ref
  where focus_nth n  focus_focused (nth_focus n)

abbreviation (input) focus_nth_array :: nat  ('a,'b,('v, 'l::len) array) ref  ('a,'b,'v) ref
  where focus_nth_array n  focus_focused (nth_focus_array n)

abbreviation (input) focus_nth_vector :: nat  ('a,'b,('v, 'l::len) vector) ref  ('a,'b,'v) ref
  where focus_nth_vector n  focus_focused (nth_focus_vector n)

subsection‹States with a global store›

named_theorems global_store_simps

text‹The following ‹locale› describes a class of abstract machine states that have a ‹global
store›, basically an abstraction of physical memory, which maps memory addresses, of type typ'a
to values of type typ'b.  As part of this abstraction, we also need an ‹allocation function›
which identifies memory addresses currently unoccupied by any value.  Note that this model is an
abstraction of physical memory in two senses:
\begin{enumerate*}
\item
Our global store contains values stored at addresses, rather than byte-representations of values,
\item
Allocation of a fresh address never fails, so in a sense we adopt an infinite memory.
\end{enumerate*}

TODO: Should the functions fixed below have type typefunction_body?  The
typefunction_body monad may be a misnomer, but functionally it offers what we need. We're
currently lifting those manually to typefunction_body later on, and then have to invent a mapping
from termNone to typeabort.›
locale global_store =
  fixes
    ― ‹Read a value from the global store, at a given address.  This may fail.›
    read_store :: 's  'a  'b and
    ― ‹Write a value to an address in the store.›
    write_store :: 's  'a  'b  's option and
    ― ‹Allocate a fresh address for a new slot in the global store.›
    allocate_store :: 's  ('a × 's) option
  assumes
    ― ‹Reads from an address, immediately after a write, succeed with the value written.›
    ― ‹TODO: What about locations we can write but not read from For such, the read-after-write
              rule below would be violated.›
    write_store_read_store [global_store_simps]: write_store σ k v = Some σ'  read_store σ' k = Some v and
    ― ‹Write commute with each other if they are to diffebrent addresses.›
    write_store_rearrange [global_store_simps]:
        k  k'  Option.bind (write_store σ k v) (λσ'. write_store σ' k' v') =
                     Option.bind (write_store σ k' v') (λσ'. write_store σ' k v) and
    ― ‹Writes to the same address collapse, with only the latter write kept.›
    write_store_write_store [global_store_simps]:
      write_store σ k v = Some σ'  write_store σ' k v' = write_store σ k v' and
    allocate_write_succeeds: allocate_store σ = Some (f, σ')  σ''. write_store σ' f v = Some σ''
begin

text‹The following definition ‹takes a reference to› an existing value.  As a side-effect, it
‹allocates› a fresh location in the state's global store, writes the value to that slot, and then
returns a reference to that slot. As part of the reference definition, one must also give
functions for injecting values of typ'v› into typ'b› and for extracting in the other direction.›
definition reference_raw :: 'b  ('s, ('a, 'b) gref, 't, 'abort, 'i, 'o) expression where
  reference_raw g  Expression (λσ.
     case allocate_store σ of
       None  Abort AssertionFailed σ
     | Some (addr,σ') 
        (case write_store σ' addr g of
          None  Abort AssertionFailed σ'
         | Some σ''  Success (make_gref addr) σ''))

definition reference_raw_fun :: 'b  ('s, ('a, 'b) gref, 'abort, 'i, 'o) function_body where
  reference_raw_fun g  FunctionBody (reference_raw g)

text‹The following definition ‹modifies› the value pointed-to by a reference by applying the
given function. Note that this method will panic if an attempt is made to write to an unallocated
reference.›
definition modify_raw :: ('a, 'b) gref  ('b  'b)  ('s, unit, 'c, 'abort, 'i, 'o) expression where
  modify_raw ref f  bind (get read_store) (λstore.
     case store (address ref) of
       None  abort DanglingPointer
     | Some old  put_assert (λσ. write_store σ (address ref) (f old)) DanglingPointer)

definition modify_raw_fun :: ('a, 'b) gref  ('b  'b)  ('s, unit, 'abort, 'i, 'o) function_body where
  modify_raw_fun ref f  FunctionBody (modify_raw ref f)

text‹The following definition ‹updates› the value pointed-to by a reference to a new value. Note
that this method will panic if an attempt is made to write to an unallocated reference, or
if the value stored in this reference is of the wrong type.›

definition update_raw :: ('a, 'b) gref  'b  ('s, unit, 'c, 'abort, 'i, 'o) expression where
  update_raw ref v  modify_raw ref (λ_. v)

definition update_raw_fun :: ('a, 'b) gref  'b  ('s, unit, 'abort, 'i, 'o) function_body where
  update_raw_fun ref v  FunctionBody (update_raw ref v)

definition dereference_by_value_raw :: ('a, 'b) gref  ('s, 'b, 'c, 'abort, 'i, 'o) expression where
  dereference_by_value_raw ref 
     bind (get (λσ. read_store σ)) (λstore.
          case store (address ref) of
            None    abort DanglingPointer
          | Some b  literal b)

definition dereference_by_value_raw_fun :: ('a, 'b) gref  ('s, 'b, 'abort, 'i, 'o) function_body where
  dereference_by_value_raw_fun ref  FunctionBody (dereference_by_value_raw ref)

end

(*<*)
end
(*>*)