src/HOL/TLA/Memory/MemoryParameters.thy
author haftmann
Mon, 04 Nov 2019 20:38:15 +0000
changeset 71042 400e9512f1d3
parent 67613 ce654b0e6d69
permissions -rw-r--r--
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41589
bbd861837ebc tuned headers;
wenzelm
parents: 36504
diff changeset
     1
(*  Title:      HOL/TLA/Memory/MemoryParameters.thy
bbd861837ebc tuned headers;
wenzelm
parents: 36504
diff changeset
     2
    Author:     Stephan Merz, University of Munich
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
     3
*)
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     4
60592
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60588
diff changeset
     5
section \<open>RPC-Memory example: Memory parameters\<close>
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     6
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 9517
diff changeset
     7
theory MemoryParameters
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 9517
diff changeset
     8
imports RPCMemoryParams
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 9517
diff changeset
     9
begin
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    10
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 5184
diff changeset
    11
(* the memory operations *)
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
    12
datatype memOp = read Locs | "write" Locs Vals
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 5184
diff changeset
    13
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    14
consts
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    15
  (* memory locations and contents *)
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 9517
diff changeset
    16
  MemLoc         :: "Locs set"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 9517
diff changeset
    17
  MemVal         :: "Vals set"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    18
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    19
  (* some particular values *)
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    20
  OK             :: "Vals"
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    21
  BadArg         :: "Vals"
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    22
  MemFailure     :: "Vals"
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    23
  NotAResult     :: "Vals"  (* defined here for simplicity *)
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 9517
diff changeset
    24
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    25
  (* the initial value stored in each memory cell *)
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    26
  InitVal        :: "Vals"
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    27
47968
3119ad2b5ad3 eliminated old 'axioms';
wenzelm
parents: 41589
diff changeset
    28
axiomatization where
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    29
  (* basic assumptions about the above constants and predicates *)
60587
0318b43ee95c more symbols;
wenzelm
parents: 58889
diff changeset
    30
  BadArgNoMemVal:    "BadArg \<notin> MemVal" and
0318b43ee95c more symbols;
wenzelm
parents: 58889
diff changeset
    31
  MemFailNoMemVal:   "MemFailure \<notin> MemVal" and
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 60592
diff changeset
    32
  InitValMemVal:     "InitVal \<in> MemVal" and
60587
0318b43ee95c more symbols;
wenzelm
parents: 58889
diff changeset
    33
  NotAResultNotVal:  "NotAResult \<notin> MemVal" and
0318b43ee95c more symbols;
wenzelm
parents: 58889
diff changeset
    34
  NotAResultNotOK:   "NotAResult \<noteq> OK" and
0318b43ee95c more symbols;
wenzelm
parents: 58889
diff changeset
    35
  NotAResultNotBA:   "NotAResult \<noteq> BadArg" and
0318b43ee95c more symbols;
wenzelm
parents: 58889
diff changeset
    36
  NotAResultNotMF:   "NotAResult \<noteq> MemFailure"
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 9517
diff changeset
    37
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    38
lemmas [simp] =
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    39
  BadArgNoMemVal MemFailNoMemVal InitValMemVal NotAResultNotVal
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    40
  NotAResultNotOK NotAResultNotBA NotAResultNotMF
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    41
  NotAResultNotOK [symmetric] NotAResultNotBA [symmetric] NotAResultNotMF [symmetric]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    42
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    43
lemma MemValNotAResultE: "\<lbrakk> x \<in> MemVal; (x \<noteq> NotAResult \<Longrightarrow> P) \<rbrakk> \<Longrightarrow> P"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    44
  using NotAResultNotVal by blast
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 9517
diff changeset
    45
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    46
end