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