src/HOL/UNITY/PPROD.thy
author wenzelm
Fri, 29 Oct 1999 20:18:34 +0200
changeset 7978 1b99ee57d131
parent 7188 2bc63a44721b
child 8251 9be357df93d4
permissions -rw-r--r--
final update by Gertrud Bauer;
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
7188
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6842
diff changeset
     6
Abstraction over replicated components (PLam)
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6842
diff changeset
     7
General products of programs (Pi operation)
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     8
*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     9
7188
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6842
diff changeset
    10
PPROD = Lift_prog +
5899
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
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    14
  PLam  :: ['a set, 'a => 'b program] => ('a => 'b) program
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    15
    "PLam I F == JN i:I. lift_prog i (F i)"
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    16
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    17
syntax
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    18
  "@PLam" :: [pttrn, 'a set, 'b set] => ('a => 'b) set ("(3plam _:_./ _)" 10)
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    19
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    20
translations
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    21
  "plam x:A. B"   == "PLam A (%x. B)"
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    22
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    23
end