src/HOL/Power.thy
author paulson
Mon, 02 Aug 1999 11:24:30 +0200
changeset 7143 9c02848c5404
parent 5183 89f162de39cf
child 7843 077d305615df
permissions -rw-r--r--
the SVC link-up

(*  Title:      HOL/Power.thy
    ID:         $Id$
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1997  University of Cambridge

The (overloaded) exponentiation operator, ^ :: [nat,nat]=>nat
Also binomial coefficents
*)

Power = Divides + 
consts
  binomial :: "[nat,nat] => nat"      ("'(_ choose _')" [50,50])

primrec
  "p ^ 0 = 1"
  "p ^ (Suc n) = (p::nat) * (p ^ n)"
  
primrec
  binomial_0   "(0     choose k) = (if k = 0 then 1 else 0)"

  binomial_Suc "(Suc n choose k) =
                (if k = 0 then 1 else (n choose (k - 1)) + (n choose k))"

end