src/HOL/TLA/Memory/MemoryParameters.thy
author wenzelm
Mon, 29 Nov 1999 15:52:49 +0100
changeset 8039 a901bafe4578
parent 6255 db63752140c7
child 9517 f58863b1406a
permissions -rw-r--r--
Goal: tuned pris;

(*
    File:        MemoryParameters.thy
    Author:      Stephan Merz
    Copyright:   1997 University of Munich

    Theory Name: MemoryParameters
    Logic Image: TLA

    RPC-Memory example: Memory parameters
*)

MemoryParameters = Datatype + RPCMemoryParams +

(* the memory operations *)
(***
datatype  Rd = read
datatype  Wr = write
***)

datatype memOp = read Locs | write Locs Vals

(***
types
  (* legal arguments for the memory *)
  memArgType = "(Rd * Locs) + (Wr * Locs * Vals)"
***)

consts
  (* memory locations and contents *)
  MemLoc         :: Locs set
  MemVal         :: Vals set

  (* some particular values *)
  OK             :: "Vals"
  BadArg         :: "Vals"
  MemFailure     :: "Vals"
  NotAResult     :: "Vals"  (* defined here for simplicity *)
  
  (* the initial value stored in each memory cell *)
  InitVal        :: "Vals"

rules
  (* basic assumptions about the above constants and predicates *)
  BadArgNoMemVal    "BadArg ~: MemVal"
  MemFailNoMemVal   "MemFailure ~: MemVal"
  InitValMemVal     "InitVal : MemVal"
  NotAResultNotVal  "NotAResult ~: MemVal"
  NotAResultNotOK   "NotAResult ~= OK"
  NotAResultNotBA   "NotAResult ~= BadArg"
  NotAResultNotMF   "NotAResult ~= MemFailure"
end