author | wenzelm |
Wed, 03 Feb 1999 16:42:40 +0100 | |
changeset 6188 | c40e5ac04e3e |
parent 6012 | 1894bfc4aee9 |
child 6295 | 351b3c2b0d83 |
permissions | -rw-r--r-- |
5111 | 1 |
(* Title: HOL/UNITY/Traces |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
6 |
Inductive definitions of |
|
7 |
* traces: the possible execution traces |
|
8 |
* reachable: the set of reachable states |
|
9 |
*) |
|
10 |
||
5648 | 11 |
Traces = LessThan + |
4776 | 12 |
|
5111 | 13 |
consts traces :: "['a set, ('a * 'a)set set] => ('a * 'a list) set" |
4776 | 14 |
|
5111 | 15 |
(*Initial states and program => (final state, reversed trace to it)... |
16 |
Arguments MUST be curried in an inductive definition*) |
|
17 |
||
18 |
inductive "traces Init Acts" |
|
4776 | 19 |
intrs |
20 |
(*Initial trace is empty*) |
|
5111 | 21 |
Init "s: Init ==> (s,[]) : traces Init Acts" |
4776 | 22 |
|
5111 | 23 |
Acts "[| act: Acts; (s,evs) : traces Init Acts; (s,s'): act |] |
24 |
==> (s', s#evs) : traces Init Acts" |
|
4776 | 25 |
|
26 |
||
5596 | 27 |
typedef (Program) |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
28 |
'a program = "{(states :: 'a set, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
29 |
init :: 'a set, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
30 |
acts :: ('a * 'a)set set). |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
31 |
init <= states & |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
32 |
acts <= Pow(states Times states) & |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
33 |
diag states : acts}" |
5584 | 34 |
|
35 |
constdefs |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
36 |
restrict_rel :: "['a set, ('a * 'a) set] => ('a * 'a) set" |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
37 |
"restrict_rel A r == (A Times A) Int r" |
5596 | 38 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
39 |
mk_program :: "('a set * 'a set * ('a * 'a)set set) => 'a program" |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
40 |
"mk_program == |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
41 |
%(states, init, acts). |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
42 |
Abs_Program (states, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
43 |
states Int init, |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
44 |
restrict_rel states `` (insert Id acts))" |
5596 | 45 |
|
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
46 |
States :: "'a program => 'a set" |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
47 |
"States F == (%(states, init, acts). states) (Rep_Program F)" |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
48 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
49 |
Init :: "'a program => 'a set" |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
50 |
"Init F == (%(states, init, acts). init) (Rep_Program F)" |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
51 |
|
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
52 |
Acts :: "'a program => ('a * 'a)set set" |
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5648
diff
changeset
|
53 |
"Acts F == (%(states, init, acts). acts) (Rep_Program F)" |
5253 | 54 |
|
55 |
||
56 |
consts reachable :: "'a program => 'a set" |
|
57 |
||
5620 | 58 |
inductive "reachable F" |
5253 | 59 |
intrs |
5620 | 60 |
Init "s: Init F ==> s : reachable F" |
5253 | 61 |
|
5620 | 62 |
Acts "[| act: Acts F; s : reachable F; (s,s'): act |] |
63 |
==> s' : reachable F" |
|
5253 | 64 |
|
4776 | 65 |
end |