src/HOL/UNITY/PPROD.thy
author paulson
Mon, 24 May 1999 15:49:44 +0200
changeset 6711 d45359e7dcbf
parent 6295 351b3c2b0d83
child 6835 588f791ee737
permissions -rw-r--r--
updated for stronger version of psp
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/PPROD.thy
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     2
    ID:         $Id$
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     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
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     7
*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     8
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
     9
PPROD = Union + Comp +
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    10
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    11
constdefs
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    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
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    15
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    16
  lift_act :: "['a, ('b*'b) set] => (('a=>'b) * ('a=>'b)) set"
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    17
    "lift_act i act == {(f,f'). EX s'. f' = f(i:=s') & (f i, s') : act}"
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    18
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    19
  drop_act :: "['a, (('a=>'b) * ('a=>'b)) set] => ('b*'b) set"
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    20
    "drop_act i act == (%(f,f'). (f i, f' i)) `` act"
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    21
6020
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
    22
  lift_set :: "['a, 'b set] => ('a => 'b) set"
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
    23
    "lift_set i A == {f. f i : A}"
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
    24
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    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
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
    28
		   lift_act i `` Acts F)"
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    29
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    30
  (*products of programs*)
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    31
  PPROD  :: ['a set, 'a => 'b program] => ('a => 'b) program
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    32
    "PPROD I F == JN i:I. lift_prog i (F i)"
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    33
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    34
syntax
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    35
  "@PPROD" :: [pttrn, 'a set, 'b set] => ('a => 'b) set ("(3PPI _:_./ _)" 10)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    36
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    37
translations
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    38
  "PPI x:A. B"   == "PPROD A (%x. B)"
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    39
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    40
end