| author | wenzelm |
| Thu, 21 Apr 2005 22:10:12 +0200 | |
| changeset 15805 | 1e8017f1e971 |
| parent 14981 | e73f8140af78 |
| child 17233 | 41eee2e7b465 |
| permissions | -rw-r--r-- |
| 4559 | 1 |
(* Title: HOLCF/IOA/meta_theory/TLS.thy |
2 |
ID: $Id$ |
|
| 12218 | 3 |
Author: Olaf Müller |
| 4559 | 4 |
|
| 12218 | 5 |
A General Temporal Logic. |
| 4559 | 6 |
*) |
| 12218 | 7 |
|
| 4559 | 8 |
TL = Pred + Sequence + |
9 |
||
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12218
diff
changeset
|
10 |
default type |
| 4559 | 11 |
|
12 |
types |
|
13 |
||
14 |
'a temporal = 'a Seq predicate |
|
15 |
||
16 |
||
17 |
consts |
|
18 |
||
19 |
||
20 |
suffix :: "'a Seq => 'a Seq => bool" |
|
21 |
tsuffix :: "'a Seq => 'a Seq => bool" |
|
22 |
||
23 |
validT :: "'a Seq predicate => bool" |
|
24 |
||
25 |
unlift :: "'a lift => 'a" |
|
26 |
||
27 |
Init ::"'a predicate => 'a temporal" ("<_>" [0] 1000)
|
|
28 |
||
29 |
Box ::"'a temporal => 'a temporal" ("[] (_)" [80] 80)
|
|
30 |
Diamond ::"'a temporal => 'a temporal" ("<> (_)" [80] 80)
|
|
31 |
Next ::"'a temporal => 'a temporal" |
|
32 |
Leadsto ::"'a temporal => 'a temporal => 'a temporal" (infixr "~>" 22) |
|
33 |
||
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
12028
diff
changeset
|
34 |
syntax (xsymbols) |
| 4559 | 35 |
"Box" ::"'a temporal => 'a temporal" ("\\<box> (_)" [80] 80)
|
36 |
"Diamond" ::"'a temporal => 'a temporal" ("\\<diamond> (_)" [80] 80)
|
|
37 |
"Leadsto" ::"'a temporal => 'a temporal => 'a temporal" (infixr "\\<leadsto>" 22) |
|
38 |
||
39 |
defs |
|
40 |
||
41 |
||
42 |
unlift_def |
|
43 |
"unlift x == (case x of |
|
| 12028 | 44 |
UU => arbitrary |
| 4559 | 45 |
| Def y => y)" |
46 |
||
47 |
(* this means that for nil and UU the effect is unpredictable *) |
|
48 |
Init_def |
|
| 10835 | 49 |
"Init P s == (P (unlift (HD$s)))" |
| 4559 | 50 |
|
51 |
suffix_def |
|
52 |
"suffix s2 s == ? s1. (Finite s1 & s = s1 @@ s2)" |
|
53 |
||
54 |
tsuffix_def |
|
55 |
"tsuffix s2 s == s2 ~= nil & s2 ~= UU & suffix s2 s" |
|
56 |
||
57 |
Box_def |
|
58 |
"([] P) s == ! s2. tsuffix s2 s --> P s2" |
|
59 |
||
60 |
Next_def |
|
| 10835 | 61 |
"(Next P) s == if (TL$s=UU | TL$s=nil) then (P s) else P (TL$s)" |
| 4559 | 62 |
|
63 |
Diamond_def |
|
64 |
"<> P == .~ ([] (.~ P))" |
|
65 |
||
66 |
Leadsto_def |
|
67 |
"P ~> Q == ([] (P .--> (<> Q)))" |
|
68 |
||
69 |
validT_def |
|
70 |
"validT P == ! s. s~=UU & s~=nil --> (s |= P)" |
|
71 |
||
72 |
end |