src/HOL/TLA/Memory/RPCMemoryParams.thy
author wenzelm
Tue, 10 Jul 2007 23:29:43 +0200
changeset 23719 ccd9cb15c062
parent 21624 6f79647cf536
child 41589 bbd861837ebc
permissions -rw-r--r--
more markup for inner and outer syntax; added enclose;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 16417
diff changeset
     1
(*
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     2
    File:        RPCMemoryParams.thy
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 16417
diff changeset
     3
    ID:          $Id$
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     4
    Author:      Stephan Merz
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     5
    Copyright:   1997 University of Munich
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
     6
*)
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     7
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
     8
header {* Basic declarations for the RPC-memory example *}
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     9
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 16417
diff changeset
    10
theory RPCMemoryParams
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 16417
diff changeset
    11
imports Main
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 16417
diff changeset
    12
begin
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    13
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    14
types
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    15
  bit = "bool"   (* Signal wires for the procedure interface.
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    16
                    Defined as bool for simplicity. All I should really need is
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    17
                    the existence of two distinct values. *)
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
(* all of these are simple (HOL) types *)
11703
wenzelm
parents: 6255
diff changeset
    20
typedecl Locs    (* "syntactic" value type *)
wenzelm
parents: 6255
diff changeset
    21
typedecl Vals    (* "syntactic" value type *)
wenzelm
parents: 6255
diff changeset
    22
typedecl PrIds   (* process id's *)
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    23
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    24
end