| author | paulson | 
| Fri, 15 Aug 2003 13:45:39 +0200 | |
| changeset 14151 | b8bb6a6a2c46 | 
| parent 12338 | de0f4a63baa5 | 
| child 15112 | 6f0772a94299 | 
| permissions | -rw-r--r-- | 
| 10213 | 1 | (* Title: HOL/Relation_Power.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow | |
| 4 | Copyright 1996 TU Muenchen | |
| 5 | ||
| 6 | R^n = R O ... O R, the n-fold composition of R | |
| 11306 | 7 | f^n = f o ... o f, the n-fold composition of f | 
| 8 | ||
| 9 | WARNING: due to the limits of Isabelle's type classes, ^ on functions and | |
| 10 | relations has too general a domain, namely ('a * 'b)set and 'a => 'b.
 | |
| 11 | This means that it may be necessary to attach explicit type constraints. | |
| 12 | Examples: range(f^n) = A and Range(R^n) = B need constraints. | |
| 10213 | 13 | *) | 
| 14 | ||
| 15 | Relation_Power = Nat + | |
| 16 | ||
| 17 | instance | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11306diff
changeset | 18 |   set :: (type) power   (* only ('a * 'a) set should be in power! *)
 | 
| 10213 | 19 | |
| 20 | primrec (relpow) | |
| 21 | "R^0 = Id" | |
| 22 | "R^(Suc n) = R O (R^n)" | |
| 23 | ||
| 11305 | 24 | |
| 11306 | 25 | instance | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11306diff
changeset | 26 | fun :: (type, type) power (* only 'a => 'a should be in power! *) | 
| 11305 | 27 | |
| 28 | primrec (funpow) | |
| 29 | "f^0 = id" | |
| 30 | "f^(Suc n) = f o (f^n)" | |
| 31 | ||
| 10213 | 32 | end |