diff -r 3fe2e418a071 -r 7fbebf75b3ef src/HOL/Nat.thy --- a/src/HOL/Nat.thy Fri Apr 24 08:24:54 2009 +0200 +++ b/src/HOL/Nat.thy Fri Apr 24 17:45:15 2009 +0200 @@ -1166,31 +1166,58 @@ subsection {* Natural operation of natural numbers on functions *} -text {* @{text "f o^ n = f o ... o f"}, the n-fold composition of @{text f} *} +text {* + We use the same logical constant for the power operations on + functions and relations, in order to share the same syntax. +*} + +consts compow :: "nat \ ('a \ 'b) \ ('a \ 'b)" + +abbreviation compower :: "('a \ 'b) \ nat \ 'a \ 'b" (infixr "^^" 80) where + "f ^^ n \ compow n f" + +notation (latex output) + compower ("(_\<^bsup>_\<^esup>)" [1000] 1000) + +notation (HTML output) + compower ("(_\<^bsup>_\<^esup>)" [1000] 1000) + +text {* @{text "f ^^ n = f o ... o f"}, the n-fold composition of @{text f} *} + +overloading + funpow == "compow :: nat \ ('a \ 'a) \ ('a \ 'a)" +begin primrec funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where "funpow 0 f = id" | "funpow (Suc n) f = f o funpow n f" -abbreviation funpower :: "('a \ 'a) \ nat \ 'a \ 'a" (infixr "o^" 80) where - "f o^ n \ funpow n f" +end + +text {* for code generation *} + +definition funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where + funpow_code_def [code post]: "funpow = compow" -notation (latex output) - funpower ("(_\<^bsup>_\<^esup>)" [1000] 1000) +lemmas [code inline] = funpow_code_def [symmetric] -notation (HTML output) - funpower ("(_\<^bsup>_\<^esup>)" [1000] 1000) +lemma [code]: + "funpow 0 f = id" + "funpow (Suc n) f = f o funpow n f" + unfolding funpow_code_def by simp_all + +definition "foo = id ^^ (1 + 1)" lemma funpow_add: - "f o^ (m + n) = f o^ m \ f o^ n" + "f ^^ (m + n) = f ^^ m \ f ^^ n" by (induct m) simp_all lemma funpow_swap1: - "f ((f o^ n) x) = (f o^ n) (f x)" + "f ((f ^^ n) x) = (f ^^ n) (f x)" proof - - have "f ((f o^ n) x) = (f o^ (n + 1)) x" by simp - also have "\ = (f o^ n o f o^ 1) x" by (simp only: funpow_add) - also have "\ = (f o^ n) (f x)" by simp + have "f ((f ^^ n) x) = (f ^^ (n + 1)) x" by simp + also have "\ = (f ^^ n o f ^^ 1) x" by (simp only: funpow_add) + also have "\ = (f ^^ n) (f x)" by simp finally show ?thesis . qed