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