src/HOL/UNITY/PPROD.thy
author paulson
Wed, 18 Nov 1998 15:10:46 +0100
changeset 5931 325300576da7
parent 5899 13d4753079fe
child 5972 2430ccbde87d
permissions -rw-r--r--
Finally removing "Compl" from HOL
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
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     6
General products of programs (Pi operation).
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     7
Also merging of state sets.
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     8
*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     9
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    10
PPROD = Union +
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    11
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    12
constdefs
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    13
  (*Cartesian product of two relations*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    14
  RTimes :: "[('a*'a) set, ('b*'b) set] => (('a*'b) * ('a*'b)) set"
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    15
	("_ RTimes _" [81, 80] 80)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    16
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    17
    "R RTimes S == {((x,y),(x',y')). (x,x'):R & (y,y'):S}"
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    18
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    19
(*FIXME: syntax (symbols) to use <times> ??
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    20
  RTimes :: "[('a*'a) set, ('b*'b) set] => (('a*'b) * ('a*'b)) set"
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    21
    ("_ \\<times> _" [81, 80] 80)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    22
*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    23
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    24
constdefs
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    25
  Lcopy :: "'a program => ('a*'b) program"
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    26
    "Lcopy F == mk_program (Init F Times UNIV,
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    27
			    (%act. act RTimes Id) `` Acts F)"
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    28
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    29
  lift_act :: "['a, ('b*'b) set] => (('a=>'b) * ('a=>'b)) set"
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    30
    "lift_act i act == {(f,f'). EX s'. f' = f(i:=s') & (f i, s') : act}"
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    31
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    32
  lift_prog :: "['a, 'b program] => ('a => 'b) program"
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    33
    "lift_prog i F == mk_program ({f. f i : Init F}, lift_act i `` Acts F)"
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    34
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    35
  (*products of programs*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    36
  PPROD  :: ['a set, 'b program] => ('a => 'b) program
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    37
    "PPROD I F == JN i:I. lift_prog i F"
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    38
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    39
syntax
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    40
  "@PPROD" :: [pttrn, 'a set, 'b set] => ('a => 'b) set ("(3PPI _:_./ _)" 10)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    41
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    42
translations
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    43
  "PPI x:A. B"   == "PPROD A (%x. B)"
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    44
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    45
end