author | wenzelm |
Fri, 02 Sep 2005 17:23:59 +0200 | |
changeset 17233 | 41eee2e7b465 |
parent 14981 | e73f8140af78 |
child 19741 | f65265d71426 |
permissions | -rw-r--r-- |
4559 | 1 |
(* Title: HOLCF/IOA/meta_theory/TLS.thy |
2 |
ID: $Id$ |
|
12218 | 3 |
Author: Olaf Müller |
17233 | 4 |
*) |
4559 | 5 |
|
17233 | 6 |
header {* Temporal Logic of Steps -- tailored for I/O automata *} |
4559 | 7 |
|
17233 | 8 |
theory TLS |
9 |
imports IOA TL |
|
10 |
begin |
|
4559 | 11 |
|
17233 | 12 |
defaultsort type |
4559 | 13 |
|
14 |
types |
|
17233 | 15 |
('a, 's) ioa_temp = "('a option,'s)transition temporal" |
16 |
('a, 's) step_pred = "('a option,'s)transition predicate" |
|
17 |
's state_pred = "'s predicate" |
|
4559 | 18 |
|
19 |
consts |
|
20 |
||
21 |
option_lift :: "('a => 'b) => 'b => ('a option => 'b)" |
|
22 |
plift :: "('a => bool) => ('a option => bool)" |
|
17233 | 23 |
|
4559 | 24 |
temp_sat :: "('a,'s)execution => ('a,'s)ioa_temp => bool" (infixr "|==" 22) |
25 |
xt1 :: "'s predicate => ('a,'s)step_pred" |
|
26 |
xt2 :: "'a option predicate => ('a,'s)step_pred" |
|
27 |
||
28 |
validTE :: "('a,'s)ioa_temp => bool" |
|
29 |
validIOA :: "('a,'s)ioa => ('a,'s)ioa_temp => bool" |
|
30 |
||
5677
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
mueller
parents:
4577
diff
changeset
|
31 |
mkfin :: "'a Seq => 'a Seq" |
4559 | 32 |
|
33 |
ex2seq :: "('a,'s)execution => ('a option,'s)transition Seq" |
|
17233 | 34 |
ex2seqC :: "('a,'s)pairs -> ('s => ('a option,'s)transition Seq)" |
4559 | 35 |
|
36 |
||
37 |
defs |
|
38 |
||
17233 | 39 |
mkfin_def: |
5677
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
mueller
parents:
4577
diff
changeset
|
40 |
"mkfin s == if Partial s then @t. Finite t & s = t @@ UU |
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
mueller
parents:
4577
diff
changeset
|
41 |
else s" |
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
mueller
parents:
4577
diff
changeset
|
42 |
|
17233 | 43 |
option_lift_def: |
4559 | 44 |
"option_lift f s y == case y of None => s | Some x => (f x)" |
45 |
||
17233 | 46 |
(* plift is used to determine that None action is always false in |
4559 | 47 |
transition predicates *) |
17233 | 48 |
plift_def: |
4559 | 49 |
"plift P == option_lift P False" |
50 |
||
17233 | 51 |
temp_sat_def: |
4559 | 52 |
"ex |== P == ((ex2seq ex) |= P)" |
53 |
||
17233 | 54 |
xt1_def: |
4559 | 55 |
"xt1 P tr == P (fst tr)" |
56 |
||
17233 | 57 |
xt2_def: |
4559 | 58 |
"xt2 P tr == P (fst (snd tr))" |
59 |
||
17233 | 60 |
ex2seq_def: |
10835 | 61 |
"ex2seq ex == ((ex2seqC $(mkfin (snd ex))) (fst ex))" |
4559 | 62 |
|
17233 | 63 |
ex2seqC_def: |
64 |
"ex2seqC == (fix$(LAM h ex. (%s. case ex of |
|
4559 | 65 |
nil => (s,None,s)>>nil |
66 |
| x##xs => (flift1 (%pr. |
|
17233 | 67 |
(s,Some (fst pr), snd pr)>> (h$xs) (snd pr)) |
10835 | 68 |
$x) |
4559 | 69 |
)))" |
70 |
||
17233 | 71 |
validTE_def: |
4559 | 72 |
"validTE P == ! ex. (ex |== P)" |
73 |
||
17233 | 74 |
validIOA_def: |
4559 | 75 |
"validIOA A P == ! ex : executions A . (ex |== P)" |
76 |
||
5976 | 77 |
|
17233 | 78 |
axioms |
5976 | 79 |
|
17233 | 80 |
mkfin_UU: |
5677
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
mueller
parents:
4577
diff
changeset
|
81 |
"mkfin UU = nil" |
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
mueller
parents:
4577
diff
changeset
|
82 |
|
17233 | 83 |
mkfin_nil: |
5677
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
mueller
parents:
4577
diff
changeset
|
84 |
"mkfin nil =nil" |
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
mueller
parents:
4577
diff
changeset
|
85 |
|
17233 | 86 |
mkfin_cons: |
5677
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
mueller
parents:
4577
diff
changeset
|
87 |
"(mkfin (a>>s)) = (a>>(mkfin s))" |
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
mueller
parents:
4577
diff
changeset
|
88 |
|
17233 | 89 |
ML {* use_legacy_bindings (the_context ()) *} |
4559 | 90 |
|
91 |
end |