author | paulson |
Mon, 24 May 1999 15:49:44 +0200 | |
changeset 6711 | d45359e7dcbf |
parent 6295 | 351b3c2b0d83 |
child 6835 | 588f791ee737 |
permissions | -rw-r--r-- |
5899 | 1 |
(* Title: HOL/UNITY/PPROD.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5972
diff
changeset
|
6 |
General products of programs (Pi operation), for replicating components. |
5899 | 7 |
*) |
8 |
||
5972 | 9 |
PPROD = Union + Comp + |
5899 | 10 |
|
11 |
constdefs |
|
12 |
||
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6020
diff
changeset
|
13 |
(**possible to force all acts to be included in common initial state |
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6020
diff
changeset
|
14 |
(by intersection) ??*) |
5899 | 15 |
|
16 |
lift_act :: "['a, ('b*'b) set] => (('a=>'b) * ('a=>'b)) set" |
|
17 |
"lift_act i act == {(f,f'). EX s'. f' = f(i:=s') & (f i, s') : act}" |
|
18 |
||
5972 | 19 |
drop_act :: "['a, (('a=>'b) * ('a=>'b)) set] => ('b*'b) set" |
20 |
"drop_act i act == (%(f,f'). (f i, f' i)) `` act" |
|
21 |
||
6020 | 22 |
lift_set :: "['a, 'b set] => ('a => 'b) set" |
23 |
"lift_set i A == {f. f i : A}" |
|
24 |
||
5899 | 25 |
lift_prog :: "['a, 'b program] => ('a => 'b) program" |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5972
diff
changeset
|
26 |
"lift_prog i F == |
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6020
diff
changeset
|
27 |
mk_program (lift_set i (Init F), |
6020 | 28 |
lift_act i `` Acts F)" |
5899 | 29 |
|
30 |
(*products of programs*) |
|
5972 | 31 |
PPROD :: ['a set, 'a => 'b program] => ('a => 'b) program |
32 |
"PPROD I F == JN i:I. lift_prog i (F i)" |
|
5899 | 33 |
|
34 |
syntax |
|
35 |
"@PPROD" :: [pttrn, 'a set, 'b set] => ('a => 'b) set ("(3PPI _:_./ _)" 10) |
|
36 |
||
37 |
translations |
|
38 |
"PPI x:A. B" == "PPROD A (%x. B)" |
|
39 |
||
40 |
end |