| author | paulson |
| Fri, 10 Sep 1999 18:40:06 +0200 | |
| changeset 7537 | 875754b599df |
| parent 7188 | 2bc63a44721b |
| child 8251 | 9be357df93d4 |
| permissions | -rw-r--r-- |
(* Title: HOL/UNITY/PPROD.thy ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge Abstraction over replicated components (PLam) General products of programs (Pi operation) *) PPROD = Lift_prog + constdefs PLam :: ['a set, 'a => 'b program] => ('a => 'b) program "PLam I F == JN i:I. lift_prog i (F i)" syntax "@PLam" :: [pttrn, 'a set, 'b set] => ('a => 'b) set ("(3plam _:_./ _)" 10) translations "plam x:A. B" == "PLam A (%x. B)" end