| author | oheimb |
| Thu, 24 Sep 1998 17:17:56 +0200 | |
| changeset 5554 | 3cae5d6510c2 |
| parent 4577 | 674b0b354feb |
| child 5976 | 44290b71a85f |
| permissions | -rw-r--r-- |
| 4559 | 1 |
(* Title: HOLCF/IOA/meta_theory/TLS.thy |
2 |
ID: $Id$ |
|
3 |
Author: Olaf M"uller |
|
4 |
Copyright 1997 TU Muenchen |
|
5 |
||
6 |
A General Temporal Logic |
|
7 |
||
8 |
Version 2: Interface directly after Sequeces, i.e. predicates and predicate transformers are in HOL |
|
9 |
||
10 |
*) |
|
11 |
||
12 |
||
13 |
||
14 |
TL = Pred + Sequence + |
|
15 |
||
16 |
default term |
|
17 |
||
18 |
types |
|
19 |
||
20 |
'a temporal = 'a Seq predicate |
|
21 |
||
22 |
||
23 |
consts |
|
24 |
||
25 |
||
26 |
suffix :: "'a Seq => 'a Seq => bool" |
|
27 |
tsuffix :: "'a Seq => 'a Seq => bool" |
|
28 |
||
29 |
validT :: "'a Seq predicate => bool" |
|
30 |
||
31 |
unlift :: "'a lift => 'a" |
|
32 |
||
33 |
Init ::"'a predicate => 'a temporal" ("<_>" [0] 1000)
|
|
34 |
||
35 |
Box ::"'a temporal => 'a temporal" ("[] (_)" [80] 80)
|
|
36 |
Diamond ::"'a temporal => 'a temporal" ("<> (_)" [80] 80)
|
|
37 |
Next ::"'a temporal => 'a temporal" |
|
38 |
Leadsto ::"'a temporal => 'a temporal => 'a temporal" (infixr "~>" 22) |
|
39 |
||
40 |
syntax (symbols) |
|
41 |
"Box" ::"'a temporal => 'a temporal" ("\\<box> (_)" [80] 80)
|
|
42 |
"Diamond" ::"'a temporal => 'a temporal" ("\\<diamond> (_)" [80] 80)
|
|
43 |
"Leadsto" ::"'a temporal => 'a temporal => 'a temporal" (infixr "\\<leadsto>" 22) |
|
44 |
||
45 |
defs |
|
46 |
||
47 |
||
48 |
unlift_def |
|
49 |
"unlift x == (case x of |
|
50 |
Undef => arbitrary |
|
51 |
| Def y => y)" |
|
52 |
||
53 |
(* this means that for nil and UU the effect is unpredictable *) |
|
54 |
Init_def |
|
55 |
"Init P s == (P (unlift (HD`s)))" |
|
56 |
||
57 |
suffix_def |
|
58 |
"suffix s2 s == ? s1. (Finite s1 & s = s1 @@ s2)" |
|
59 |
||
60 |
tsuffix_def |
|
61 |
"tsuffix s2 s == s2 ~= nil & s2 ~= UU & suffix s2 s" |
|
62 |
||
63 |
Box_def |
|
64 |
"([] P) s == ! s2. tsuffix s2 s --> P s2" |
|
65 |
||
66 |
Next_def |
|
|
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
67 |
"(Next P) s == if (TL`s=UU | TL`s=nil) then (P s) else P (TL`s)" |
| 4559 | 68 |
|
69 |
Diamond_def |
|
70 |
"<> P == .~ ([] (.~ P))" |
|
71 |
||
72 |
Leadsto_def |
|
73 |
"P ~> Q == ([] (P .--> (<> Q)))" |
|
74 |
||
75 |
validT_def |
|
76 |
"validT P == ! s. s~=UU & s~=nil --> (s |= P)" |
|
77 |
||
78 |
end |