6470
|
1 |
(* Title: HOLCF/IOA/TrivEx.thy
|
|
2 |
ID: $Id$
|
12218
|
3 |
Author: Olaf Müller
|
6470
|
4 |
*)
|
|
5 |
|
17244
|
6 |
header {* Trivial Abstraction Example with fairness *}
|
|
7 |
|
|
8 |
theory TrivEx2
|
|
9 |
imports IOA Abstraction
|
|
10 |
begin
|
6470
|
11 |
|
|
12 |
datatype action = INC
|
|
13 |
|
|
14 |
consts
|
|
15 |
|
|
16 |
C_asig :: "action signature"
|
17244
|
17 |
C_trans :: "(action, nat)transition set"
|
|
18 |
C_ioa :: "(action, nat)ioa"
|
|
19 |
C_live_ioa :: "(action, nat)live_ioa"
|
6470
|
20 |
|
|
21 |
A_asig :: "action signature"
|
17244
|
22 |
A_trans :: "(action, bool)transition set"
|
|
23 |
A_ioa :: "(action, bool)ioa"
|
|
24 |
A_live_ioa :: "(action, bool)live_ioa"
|
6470
|
25 |
|
|
26 |
h_abs :: "nat => bool"
|
|
27 |
|
|
28 |
defs
|
|
29 |
|
17244
|
30 |
C_asig_def:
|
6470
|
31 |
"C_asig == ({},{INC},{})"
|
|
32 |
|
17244
|
33 |
C_trans_def: "C_trans ==
|
|
34 |
{tr. let s = fst(tr);
|
|
35 |
t = snd(snd(tr))
|
|
36 |
in case fst(snd(tr))
|
|
37 |
of
|
6470
|
38 |
INC => t = Suc(s)}"
|
|
39 |
|
17244
|
40 |
C_ioa_def: "C_ioa ==
|
6470
|
41 |
(C_asig, {0}, C_trans,{},{})"
|
|
42 |
|
17244
|
43 |
C_live_ioa_def:
|
6470
|
44 |
"C_live_ioa == (C_ioa, WF C_ioa {INC})"
|
|
45 |
|
17244
|
46 |
A_asig_def:
|
6470
|
47 |
"A_asig == ({},{INC},{})"
|
|
48 |
|
17244
|
49 |
A_trans_def: "A_trans ==
|
|
50 |
{tr. let s = fst(tr);
|
|
51 |
t = snd(snd(tr))
|
|
52 |
in case fst(snd(tr))
|
|
53 |
of
|
6470
|
54 |
INC => t = True}"
|
|
55 |
|
17244
|
56 |
A_ioa_def: "A_ioa ==
|
6470
|
57 |
(A_asig, {False}, A_trans,{},{})"
|
|
58 |
|
17244
|
59 |
A_live_ioa_def:
|
6470
|
60 |
"A_live_ioa == (A_ioa, WF A_ioa {INC})"
|
|
61 |
|
|
62 |
|
17244
|
63 |
h_abs_def:
|
6470
|
64 |
"h_abs n == n~=0"
|
|
65 |
|
17244
|
66 |
axioms
|
6470
|
67 |
|
17244
|
68 |
MC_result:
|
6470
|
69 |
"validLIOA (A_ioa,WF A_ioa {INC}) (<>[] <%(b,a,c). b>)"
|
|
70 |
|
17244
|
71 |
ML {* use_legacy_bindings (the_context ()) *}
|
|
72 |
|
|
73 |
end
|