| 3807 |      1 | (*
 | 
|  |      2 |     File:        MemoryParameters.thy
 | 
| 17309 |      3 |     ID:          $Id$
 | 
| 3807 |      4 |     Author:      Stephan Merz
 | 
|  |      5 |     Copyright:   1997 University of Munich
 | 
| 21624 |      6 | *)
 | 
| 3807 |      7 | 
 | 
| 21624 |      8 | header {* RPC-Memory example: Memory parameters *}
 | 
| 3807 |      9 | 
 | 
| 17309 |     10 | theory MemoryParameters
 | 
|  |     11 | imports RPCMemoryParams
 | 
|  |     12 | begin
 | 
| 3807 |     13 | 
 | 
| 6255 |     14 | (* the memory operations *)
 | 
| 36504 |     15 | datatype memOp = read Locs | "write" Locs Vals
 | 
| 6255 |     16 | 
 | 
| 3807 |     17 | consts
 | 
|  |     18 |   (* memory locations and contents *)
 | 
| 17309 |     19 |   MemLoc         :: "Locs set"
 | 
|  |     20 |   MemVal         :: "Vals set"
 | 
| 3807 |     21 | 
 | 
|  |     22 |   (* some particular values *)
 | 
|  |     23 |   OK             :: "Vals"
 | 
|  |     24 |   BadArg         :: "Vals"
 | 
|  |     25 |   MemFailure     :: "Vals"
 | 
|  |     26 |   NotAResult     :: "Vals"  (* defined here for simplicity *)
 | 
| 17309 |     27 | 
 | 
| 3807 |     28 |   (* the initial value stored in each memory cell *)
 | 
|  |     29 |   InitVal        :: "Vals"
 | 
|  |     30 | 
 | 
| 17309 |     31 | axioms
 | 
| 3807 |     32 |   (* basic assumptions about the above constants and predicates *)
 | 
| 17309 |     33 |   BadArgNoMemVal:    "BadArg ~: MemVal"
 | 
|  |     34 |   MemFailNoMemVal:   "MemFailure ~: MemVal"
 | 
|  |     35 |   InitValMemVal:     "InitVal : MemVal"
 | 
|  |     36 |   NotAResultNotVal:  "NotAResult ~: MemVal"
 | 
|  |     37 |   NotAResultNotOK:   "NotAResult ~= OK"
 | 
|  |     38 |   NotAResultNotBA:   "NotAResult ~= BadArg"
 | 
|  |     39 |   NotAResultNotMF:   "NotAResult ~= MemFailure"
 | 
|  |     40 | 
 | 
| 21624 |     41 | lemmas [simp] =
 | 
|  |     42 |   BadArgNoMemVal MemFailNoMemVal InitValMemVal NotAResultNotVal
 | 
|  |     43 |   NotAResultNotOK NotAResultNotBA NotAResultNotMF
 | 
|  |     44 |   NotAResultNotOK [symmetric] NotAResultNotBA [symmetric] NotAResultNotMF [symmetric]
 | 
|  |     45 | 
 | 
|  |     46 | lemma MemValNotAResultE: "[| x : MemVal; (x ~= NotAResult ==> P) |] ==> P"
 | 
|  |     47 |   using NotAResultNotVal by blast
 | 
| 17309 |     48 | 
 | 
| 3807 |     49 | end
 |