| author | wenzelm |
| Thu, 27 Aug 1998 16:54:55 +0200 | |
| changeset 5393 | 7299e531d481 |
| parent 5253 | 82a5ca6290aa |
| child 5426 | 566f47250bd0 |
| permissions | -rw-r--r-- |
| 4776 | 1 |
(* Title: HOL/UNITY/Reach.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
6 |
Reachability in Directed Graphs. From Chandy and Misra, section 6.4. |
|
7 |
*) |
|
8 |
||
| 5232 | 9 |
Reach = FP + Traces + SubstAx + |
| 4776 | 10 |
|
11 |
types vertex |
|
12 |
state = "vertex=>bool" |
|
13 |
||
14 |
arities vertex :: term |
|
15 |
||
16 |
consts |
|
17 |
init :: "vertex" |
|
18 |
||
19 |
edges :: "(vertex*vertex) set" |
|
20 |
||
21 |
constdefs |
|
22 |
||
23 |
asgt :: "[vertex,vertex] => (state*state) set" |
|
|
5071
548f398d770b
Consequences of the change from [ := ] to ( := ) in theory Update.
nipkow
parents:
4896
diff
changeset
|
24 |
"asgt u v == {(s,s'). s' = s(v:= s u | s v)}"
|
| 4776 | 25 |
|
| 5253 | 26 |
Rprg :: state program |
27 |
"Rprg == (|Init = {%v. v=init},
|
|
28 |
Acts = insert id (UN (u,v): edges. {asgt u v})|)"
|
|
| 4776 | 29 |
|
|
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
30 |
reach_invariant :: state set |
|
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
31 |
"reach_invariant == {s. (ALL v. s v --> (init, v) : edges^*) & s init}"
|
| 4776 | 32 |
|
33 |
fixedpoint :: state set |
|
34 |
"fixedpoint == {s. ALL (u,v): edges. s u --> s v}"
|
|
35 |
||
36 |
metric :: state => nat |
|
37 |
"metric == (%s. card {v. ~ s v})"
|
|
38 |
||
39 |
rules |
|
40 |
||
41 |
(*We assume that the set of vertices is finite*) |
|
42 |
finite_graph "finite (UNIV :: vertex set)" |
|
43 |
||
44 |
end |