3210
|
1 |
(* Title: HOL/Modelcheck/Example.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Olaf Mueller, Jan Philipps, Robert Sandner
|
|
4 |
Copyright 1997 TU Muenchen
|
|
5 |
*)
|
|
6 |
|
|
7 |
Example = MCSyn +
|
|
8 |
|
|
9 |
|
|
10 |
types
|
|
11 |
state = "bool * bool * bool"
|
|
12 |
|
|
13 |
consts
|
|
14 |
|
|
15 |
INIT :: "state pred"
|
|
16 |
N :: "[state,state] => bool"
|
|
17 |
reach:: "state pred"
|
|
18 |
|
|
19 |
defs
|
|
20 |
|
|
21 |
INIT_def "INIT x == ~(fst x)&~(fst (snd x))&~(snd (snd x))"
|
|
22 |
|
3225
|
23 |
N_def "N == % (x1,x2,x3) (y1,y2,y3).
|
|
24 |
(~x1 & ~x2 & ~x3 & y1 & ~y2 & ~y3) |
|
|
25 |
( x1 & ~x2 & ~x3 & ~y1 & ~y2 & ~y3) |
|
|
26 |
( x1 & ~x2 & ~x3 & y1 & y2 & y3) "
|
3210
|
27 |
|
|
28 |
reach_def "reach == mu (%Q x. INIT x | (? y. Q y & N y x))"
|
|
29 |
|
|
30 |
|
|
31 |
end
|
|
32 |
|