nipkow@10213: (* Title: HOL/Relation_Power.thy nipkow@10213: ID: $Id$ nipkow@10213: Author: Tobias Nipkow nipkow@10213: Copyright 1996 TU Muenchen nipkow@10213: nipkow@10213: R^n = R O ... O R, the n-fold composition of R nipkow@11306: f^n = f o ... o f, the n-fold composition of f nipkow@11306: nipkow@11306: WARNING: due to the limits of Isabelle's type classes, ^ on functions and nipkow@11306: relations has too general a domain, namely ('a * 'b)set and 'a => 'b. nipkow@11306: This means that it may be necessary to attach explicit type constraints. nipkow@11306: Examples: range(f^n) = A and Range(R^n) = B need constraints. nipkow@10213: *) nipkow@10213: nipkow@15131: theory Relation_Power nipkow@15131: import Nat nipkow@15131: begin nipkow@10213: nipkow@10213: instance nipkow@15112: set :: (type) power .. (* only ('a * 'a) set should be in power! *) nipkow@10213: nipkow@10213: primrec (relpow) nipkow@10213: "R^0 = Id" nipkow@10213: "R^(Suc n) = R O (R^n)" nipkow@10213: nipkow@11305: nipkow@11306: instance nipkow@15112: fun :: (type, type) power .. (* only 'a => 'a should be in power! *) nipkow@11305: nipkow@11305: primrec (funpow) nipkow@11305: "f^0 = id" nipkow@11305: "f^(Suc n) = f o (f^n)" nipkow@11305: nipkow@15112: lemma funpow_add: "f ^ (m+n) = f^m o f^n" nipkow@15112: by(induct m) simp_all nipkow@15112: nipkow@10213: end