src/HOL/TLA/Stfun.thy
author wenzelm
Wed, 07 Sep 2005 20:22:39 +0200
changeset 17309 c43ed29bd197
parent 12607 16b63730cfbb
child 21624 6f79647cf536
permissions -rw-r--r--
converted to Isar theory format;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
     1
(*
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
     2
    File:        TLA/Stfun.thy
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
     3
    ID:          $Id$
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     4
    Author:      Stephan Merz
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
     5
    Copyright:   1998 University of Munich
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     6
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     7
    Theory Name: Stfun
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     8
    Logic Image: HOL
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     9
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    10
States and state functions for TLA as an "intensional" logic.
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    11
*)
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    12
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    13
theory Stfun
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    14
imports Intensional
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    15
begin
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    16
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    17
typedecl state
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    18
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    19
instance state :: world ..
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    20
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    21
types
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    22
  'a stfun = "state => 'a"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    23
  stpred  = "bool stfun"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    24
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    25
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    26
consts
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    27
  (* Formalizing type "state" would require formulas to be tagged with
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    28
     their underlying state space and would result in a system that is
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    29
     much harder to use. (Unlike Hoare logic or Unity, TLA has quantification
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    30
     over state variables, and therefore one usually works with different
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    31
     state spaces within a single specification.) Instead, "state" is just
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    32
     an anonymous type whose only purpose is to provide "Skolem" constants.
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    33
     Moreover, we do not define a type of state variables separate from that
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    34
     of arbitrary state functions, again in order to simplify the definition
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    35
     of flexible quantification later on. Nevertheless, we need to distinguish
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    36
     state variables, mainly to define the enabledness of actions. The user
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    37
     identifies (tuples of) "base" state variables in a specification via the
12607
16b63730cfbb update by Stephan Merz;
wenzelm
parents: 12338
diff changeset
    38
     "meta predicate" basevars, which is defined here.
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    39
  *)
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    40
  stvars    :: "'a stfun => bool"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    41
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    42
syntax
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    43
  "PRED"    :: "lift => 'a"                          ("PRED _")
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    44
  "_stvars" :: "lift => bool"                        ("basevars _")
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    45
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    46
translations
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    47
  "PRED P"   =>  "(P::state => _)"
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    48
  "_stvars"  ==  "stvars"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    49
12607
16b63730cfbb update by Stephan Merz;
wenzelm
parents: 12338
diff changeset
    50
defs
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    51
  (* Base variables may be assigned arbitrary (type-correct) values.
12607
16b63730cfbb update by Stephan Merz;
wenzelm
parents: 12338
diff changeset
    52
     Note that vs may be a tuple of variables. The correct identification
16b63730cfbb update by Stephan Merz;
wenzelm
parents: 12338
diff changeset
    53
     of base variables is up to the user who must take care not to
16b63730cfbb update by Stephan Merz;
wenzelm
parents: 12338
diff changeset
    54
     introduce an inconsistency. For example, "basevars (x,x)" would
16b63730cfbb update by Stephan Merz;
wenzelm
parents: 12338
diff changeset
    55
     definitely be inconsistent.
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    56
  *)
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    57
  basevars_def:  "stvars vs == range vs = UNIV"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    58
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 12607
diff changeset
    59
ML {* use_legacy_bindings (the_context ()) *}
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    60
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    61
end