src/HOL/Power.thy
author wenzelm
Thu, 15 Nov 2001 18:20:13 +0100
changeset 12207 4dff931b852f
parent 8844 db71c334e854
child 14348 744c868ee0b7
permissions -rw-r--r--
added Induct/Binary_Trees.thy, Induct/Tree_Forest (converted from former ex/TF.ML ex/TF.thy ex/Term.ML ex/Term.thy);

(*  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"      (infixl "choose" 65)

primrec (power)
  "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