src/HOL/Power.thy
author huffman
Mon, 23 Feb 2009 16:25:52 -0800
changeset 30016 293b896b9c25
parent 29993 0a35bee25c20
child 30242 aea5d7fa7ef5
permissions -rw-r--r--
make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
     1 (*  Title:      HOL/Power.thy
     2     ID:         $Id$
     3     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     4     Copyright   1997  University of Cambridge
     5 
     6 *)
     7 
     8 header{*Exponentiation*}
     9 
    10 theory Power
    11 imports Nat
    12 begin
    13 
    14 class power =
    15   fixes power :: "'a \<Rightarrow> nat \<Rightarrow> 'a"            (infixr "^" 80)
    16 
    17 subsection{*Powers for Arbitrary Monoids*}
    18 
    19 class recpower = monoid_mult + power +
    20   assumes power_0 [simp]: "a ^ 0       = 1"
    21   assumes power_Suc:      "a ^ Suc n = a * (a ^ n)"
    22 
    23 lemma power_0_Suc [simp]: "(0::'a::{recpower,semiring_0}) ^ (Suc n) = 0"
    24   by (simp add: power_Suc)
    25 
    26 text{*It looks plausible as a simprule, but its effect can be strange.*}
    27 lemma power_0_left: "0^n = (if n=0 then 1 else (0::'a::{recpower,semiring_0}))"
    28   by (induct n) simp_all
    29 
    30 lemma power_one [simp]: "1^n = (1::'a::recpower)"
    31   by (induct n) (simp_all add: power_Suc)
    32 
    33 lemma power_one_right [simp]: "(a::'a::recpower) ^ 1 = a"
    34   unfolding One_nat_def by (simp add: power_Suc)
    35 
    36 lemma power_commutes: "(a::'a::recpower) ^ n * a = a * a ^ n"
    37   by (induct n) (simp_all add: power_Suc mult_assoc)
    38 
    39 lemma power_Suc2: "(a::'a::recpower) ^ Suc n = a ^ n * a"
    40   by (simp add: power_Suc power_commutes)
    41 
    42 lemma power_add: "(a::'a::recpower) ^ (m+n) = (a^m) * (a^n)"
    43   by (induct m) (simp_all add: power_Suc mult_ac)
    44 
    45 lemma power_mult: "(a::'a::recpower) ^ (m*n) = (a^m) ^ n"
    46   by (induct n) (simp_all add: power_Suc power_add)
    47 
    48 lemma power_mult_distrib: "((a::'a::{recpower,comm_monoid_mult}) * b) ^ n = (a^n) * (b^n)"
    49   by (induct n) (simp_all add: power_Suc mult_ac)
    50 
    51 lemma zero_less_power[simp]:
    52      "0 < (a::'a::{ordered_semidom,recpower}) ==> 0 < a^n"
    53 apply (induct "n")
    54 apply (simp_all add: power_Suc zero_less_one mult_pos_pos)
    55 done
    56 
    57 lemma zero_le_power[simp]:
    58      "0 \<le> (a::'a::{ordered_semidom,recpower}) ==> 0 \<le> a^n"
    59 apply (simp add: order_le_less)
    60 apply (erule disjE)
    61 apply (simp_all add: zero_less_one power_0_left)
    62 done
    63 
    64 lemma one_le_power[simp]:
    65      "1 \<le> (a::'a::{ordered_semidom,recpower}) ==> 1 \<le> a^n"
    66 apply (induct "n")
    67 apply (simp_all add: power_Suc)
    68 apply (rule order_trans [OF _ mult_mono [of 1 _ 1]])
    69 apply (simp_all add: zero_le_one order_trans [OF zero_le_one])
    70 done
    71 
    72 lemma gt1_imp_ge0: "1 < a ==> 0 \<le> (a::'a::ordered_semidom)"
    73   by (simp add: order_trans [OF zero_le_one order_less_imp_le])
    74 
    75 lemma power_gt1_lemma:
    76   assumes gt1: "1 < (a::'a::{ordered_semidom,recpower})"
    77   shows "1 < a * a^n"
    78 proof -
    79   have "1*1 < a*1" using gt1 by simp
    80   also have "\<dots> \<le> a * a^n" using gt1
    81     by (simp only: mult_mono gt1_imp_ge0 one_le_power order_less_imp_le
    82         zero_le_one order_refl)
    83   finally show ?thesis by simp
    84 qed
    85 
    86 lemma one_less_power[simp]:
    87   "\<lbrakk>1 < (a::'a::{ordered_semidom,recpower}); 0 < n\<rbrakk> \<Longrightarrow> 1 < a ^ n"
    88 by (cases n, simp_all add: power_gt1_lemma power_Suc)
    89 
    90 lemma power_gt1:
    91      "1 < (a::'a::{ordered_semidom,recpower}) ==> 1 < a ^ (Suc n)"
    92 by (simp add: power_gt1_lemma power_Suc)
    93 
    94 lemma power_le_imp_le_exp:
    95   assumes gt1: "(1::'a::{recpower,ordered_semidom}) < a"
    96   shows "!!n. a^m \<le> a^n ==> m \<le> n"
    97 proof (induct m)
    98   case 0
    99   show ?case by simp
   100 next
   101   case (Suc m)
   102   show ?case
   103   proof (cases n)
   104     case 0
   105     from prems have "a * a^m \<le> 1" by (simp add: power_Suc)
   106     with gt1 show ?thesis
   107       by (force simp only: power_gt1_lemma
   108           linorder_not_less [symmetric])
   109   next
   110     case (Suc n)
   111     from prems show ?thesis
   112       by (force dest: mult_left_le_imp_le
   113           simp add: power_Suc order_less_trans [OF zero_less_one gt1])
   114   qed
   115 qed
   116 
   117 text{*Surely we can strengthen this? It holds for @{text "0<a<1"} too.*}
   118 lemma power_inject_exp [simp]:
   119      "1 < (a::'a::{ordered_semidom,recpower}) ==> (a^m = a^n) = (m=n)"
   120   by (force simp add: order_antisym power_le_imp_le_exp)
   121 
   122 text{*Can relax the first premise to @{term "0<a"} in the case of the
   123 natural numbers.*}
   124 lemma power_less_imp_less_exp:
   125      "[| (1::'a::{recpower,ordered_semidom}) < a; a^m < a^n |] ==> m < n"
   126 by (simp add: order_less_le [of m n] order_less_le [of "a^m" "a^n"]
   127               power_le_imp_le_exp)
   128 
   129 
   130 lemma power_mono:
   131      "[|a \<le> b; (0::'a::{recpower,ordered_semidom}) \<le> a|] ==> a^n \<le> b^n"
   132 apply (induct "n")
   133 apply (simp_all add: power_Suc)
   134 apply (auto intro: mult_mono order_trans [of 0 a b])
   135 done
   136 
   137 lemma power_strict_mono [rule_format]:
   138      "[|a < b; (0::'a::{recpower,ordered_semidom}) \<le> a|]
   139       ==> 0 < n --> a^n < b^n"
   140 apply (induct "n")
   141 apply (auto simp add: mult_strict_mono power_Suc
   142                       order_le_less_trans [of 0 a b])
   143 done
   144 
   145 lemma power_eq_0_iff [simp]:
   146   "(a^n = 0) \<longleftrightarrow>
   147    (a = (0::'a::{mult_zero,zero_neq_one,no_zero_divisors,recpower}) & n\<noteq>0)"
   148 apply (induct "n")
   149 apply (auto simp add: power_Suc zero_neq_one [THEN not_sym] no_zero_divisors)
   150 done
   151 
   152 
   153 lemma field_power_not_zero:
   154   "a \<noteq> (0::'a::{ring_1_no_zero_divisors,recpower}) ==> a^n \<noteq> 0"
   155 by force
   156 
   157 lemma nonzero_power_inverse:
   158   fixes a :: "'a::{division_ring,recpower}"
   159   shows "a \<noteq> 0 ==> inverse (a ^ n) = (inverse a) ^ n"
   160 apply (induct "n")
   161 apply (auto simp add: power_Suc nonzero_inverse_mult_distrib power_commutes)
   162 done (* TODO: reorient or rename to nonzero_inverse_power *)
   163 
   164 text{*Perhaps these should be simprules.*}
   165 lemma power_inverse:
   166   fixes a :: "'a::{division_ring,division_by_zero,recpower}"
   167   shows "inverse (a ^ n) = (inverse a) ^ n"
   168 apply (cases "a = 0")
   169 apply (simp add: power_0_left)
   170 apply (simp add: nonzero_power_inverse)
   171 done (* TODO: reorient or rename to inverse_power *)
   172 
   173 lemma power_one_over: "1 / (a::'a::{field,division_by_zero,recpower})^n = 
   174     (1 / a)^n"
   175 apply (simp add: divide_inverse)
   176 apply (rule power_inverse)
   177 done
   178 
   179 lemma nonzero_power_divide:
   180     "b \<noteq> 0 ==> (a/b) ^ n = ((a::'a::{field,recpower}) ^ n) / (b ^ n)"
   181 by (simp add: divide_inverse power_mult_distrib nonzero_power_inverse)
   182 
   183 lemma power_divide:
   184     "(a/b) ^ n = ((a::'a::{field,division_by_zero,recpower}) ^ n / b ^ n)"
   185 apply (case_tac "b=0", simp add: power_0_left)
   186 apply (rule nonzero_power_divide)
   187 apply assumption
   188 done
   189 
   190 lemma power_abs: "abs(a ^ n) = abs(a::'a::{ordered_idom,recpower}) ^ n"
   191 apply (induct "n")
   192 apply (auto simp add: power_Suc abs_mult)
   193 done
   194 
   195 lemma zero_less_power_abs_iff [simp,noatp]:
   196      "(0 < (abs a)^n) = (a \<noteq> (0::'a::{ordered_idom,recpower}) | n=0)"
   197 proof (induct "n")
   198   case 0
   199     show ?case by (simp add: zero_less_one)
   200 next
   201   case (Suc n)
   202     show ?case by (auto simp add: prems power_Suc zero_less_mult_iff
   203       abs_zero)
   204 qed
   205 
   206 lemma zero_le_power_abs [simp]:
   207      "(0::'a::{ordered_idom,recpower}) \<le> (abs a)^n"
   208 by (rule zero_le_power [OF abs_ge_zero])
   209 
   210 lemma power_minus: "(-a) ^ n = (- 1)^n * (a::'a::{ring_1,recpower}) ^ n"
   211 proof (induct n)
   212   case 0 show ?case by simp
   213 next
   214   case (Suc n) then show ?case
   215     by (simp add: power_Suc2 mult_assoc)
   216 qed
   217 
   218 text{*Lemma for @{text power_strict_decreasing}*}
   219 lemma power_Suc_less:
   220      "[|(0::'a::{ordered_semidom,recpower}) < a; a < 1|]
   221       ==> a * a^n < a^n"
   222 apply (induct n)
   223 apply (auto simp add: power_Suc mult_strict_left_mono)
   224 done
   225 
   226 lemma power_strict_decreasing:
   227      "[|n < N; 0 < a; a < (1::'a::{ordered_semidom,recpower})|]
   228       ==> a^N < a^n"
   229 apply (erule rev_mp)
   230 apply (induct "N")
   231 apply (auto simp add: power_Suc power_Suc_less less_Suc_eq)
   232 apply (rename_tac m)
   233 apply (subgoal_tac "a * a^m < 1 * a^n", simp)
   234 apply (rule mult_strict_mono)
   235 apply (auto simp add: zero_less_one order_less_imp_le)
   236 done
   237 
   238 text{*Proof resembles that of @{text power_strict_decreasing}*}
   239 lemma power_decreasing:
   240      "[|n \<le> N; 0 \<le> a; a \<le> (1::'a::{ordered_semidom,recpower})|]
   241       ==> a^N \<le> a^n"
   242 apply (erule rev_mp)
   243 apply (induct "N")
   244 apply (auto simp add: power_Suc  le_Suc_eq)
   245 apply (rename_tac m)
   246 apply (subgoal_tac "a * a^m \<le> 1 * a^n", simp)
   247 apply (rule mult_mono)
   248 apply (auto simp add: zero_le_one)
   249 done
   250 
   251 lemma power_Suc_less_one:
   252      "[| 0 < a; a < (1::'a::{ordered_semidom,recpower}) |] ==> a ^ Suc n < 1"
   253 apply (insert power_strict_decreasing [of 0 "Suc n" a], simp)
   254 done
   255 
   256 text{*Proof again resembles that of @{text power_strict_decreasing}*}
   257 lemma power_increasing:
   258      "[|n \<le> N; (1::'a::{ordered_semidom,recpower}) \<le> a|] ==> a^n \<le> a^N"
   259 apply (erule rev_mp)
   260 apply (induct "N")
   261 apply (auto simp add: power_Suc le_Suc_eq)
   262 apply (rename_tac m)
   263 apply (subgoal_tac "1 * a^n \<le> a * a^m", simp)
   264 apply (rule mult_mono)
   265 apply (auto simp add: order_trans [OF zero_le_one])
   266 done
   267 
   268 text{*Lemma for @{text power_strict_increasing}*}
   269 lemma power_less_power_Suc:
   270      "(1::'a::{ordered_semidom,recpower}) < a ==> a^n < a * a^n"
   271 apply (induct n)
   272 apply (auto simp add: power_Suc mult_strict_left_mono order_less_trans [OF zero_less_one])
   273 done
   274 
   275 lemma power_strict_increasing:
   276      "[|n < N; (1::'a::{ordered_semidom,recpower}) < a|] ==> a^n < a^N"
   277 apply (erule rev_mp)
   278 apply (induct "N")
   279 apply (auto simp add: power_less_power_Suc power_Suc less_Suc_eq)
   280 apply (rename_tac m)
   281 apply (subgoal_tac "1 * a^n < a * a^m", simp)
   282 apply (rule mult_strict_mono)
   283 apply (auto simp add: order_less_trans [OF zero_less_one] order_less_imp_le)
   284 done
   285 
   286 lemma power_increasing_iff [simp]:
   287   "1 < (b::'a::{ordered_semidom,recpower}) ==> (b ^ x \<le> b ^ y) = (x \<le> y)"
   288 by (blast intro: power_le_imp_le_exp power_increasing order_less_imp_le) 
   289 
   290 lemma power_strict_increasing_iff [simp]:
   291   "1 < (b::'a::{ordered_semidom,recpower}) ==> (b ^ x < b ^ y) = (x < y)"
   292 by (blast intro: power_less_imp_less_exp power_strict_increasing) 
   293 
   294 lemma power_le_imp_le_base:
   295 assumes le: "a ^ Suc n \<le> b ^ Suc n"
   296     and ynonneg: "(0::'a::{ordered_semidom,recpower}) \<le> b"
   297 shows "a \<le> b"
   298 proof (rule ccontr)
   299   assume "~ a \<le> b"
   300   then have "b < a" by (simp only: linorder_not_le)
   301   then have "b ^ Suc n < a ^ Suc n"
   302     by (simp only: prems power_strict_mono)
   303   from le and this show "False"
   304     by (simp add: linorder_not_less [symmetric])
   305 qed
   306 
   307 lemma power_less_imp_less_base:
   308   fixes a b :: "'a::{ordered_semidom,recpower}"
   309   assumes less: "a ^ n < b ^ n"
   310   assumes nonneg: "0 \<le> b"
   311   shows "a < b"
   312 proof (rule contrapos_pp [OF less])
   313   assume "~ a < b"
   314   hence "b \<le> a" by (simp only: linorder_not_less)
   315   hence "b ^ n \<le> a ^ n" using nonneg by (rule power_mono)
   316   thus "~ a ^ n < b ^ n" by (simp only: linorder_not_less)
   317 qed
   318 
   319 lemma power_inject_base:
   320      "[| a ^ Suc n = b ^ Suc n; 0 \<le> a; 0 \<le> b |]
   321       ==> a = (b::'a::{ordered_semidom,recpower})"
   322 by (blast intro: power_le_imp_le_base order_antisym order_eq_refl sym)
   323 
   324 lemma power_eq_imp_eq_base:
   325   fixes a b :: "'a::{ordered_semidom,recpower}"
   326   shows "\<lbrakk>a ^ n = b ^ n; 0 \<le> a; 0 \<le> b; 0 < n\<rbrakk> \<Longrightarrow> a = b"
   327 by (cases n, simp_all, rule power_inject_base)
   328 
   329 text {* The divides relation *}
   330 
   331 lemma le_imp_power_dvd:
   332   fixes a :: "'a::{comm_semiring_1,recpower}"
   333   assumes "m \<le> n" shows "a^m dvd a^n"
   334 proof
   335   have "a^n = a^(m + (n - m))"
   336     using `m \<le> n` by simp
   337   also have "\<dots> = a^m * a^(n - m)"
   338     by (rule power_add)
   339   finally show "a^n = a^m * a^(n - m)" .
   340 qed
   341 
   342 lemma power_le_dvd:
   343   fixes a b :: "'a::{comm_semiring_1,recpower}"
   344   shows "a^n dvd b \<Longrightarrow> m \<le> n \<Longrightarrow> a^m dvd b"
   345   by (rule dvd_trans [OF le_imp_power_dvd])
   346 
   347 
   348 subsection{*Exponentiation for the Natural Numbers*}
   349 
   350 instantiation nat :: recpower
   351 begin
   352 
   353 primrec power_nat where
   354   "p ^ 0 = (1\<Colon>nat)"
   355   | "p ^ (Suc n) = (p\<Colon>nat) * (p ^ n)"
   356 
   357 instance proof
   358   fix z n :: nat
   359   show "z^0 = 1" by simp
   360   show "z^(Suc n) = z * (z^n)" by simp
   361 qed
   362 
   363 end
   364 
   365 lemma of_nat_power:
   366   "of_nat (m ^ n) = (of_nat m::'a::{semiring_1,recpower}) ^ n"
   367 by (induct n, simp_all add: power_Suc of_nat_mult)
   368 
   369 lemma nat_one_le_power [simp]: "Suc 0 \<le> i ==> Suc 0 \<le> i^n"
   370 by (rule one_le_power [of i n, unfolded One_nat_def])
   371 
   372 lemma nat_zero_less_power_iff [simp]: "(x^n > 0) = (x > (0::nat) | n=0)"
   373 by (induct "n", auto)
   374 
   375 lemma nat_power_eq_Suc_0_iff [simp]: 
   376   "((x::nat)^m = Suc 0) = (m = 0 | x = Suc 0)"
   377 by (induct_tac m, auto)
   378 
   379 lemma power_Suc_0[simp]: "(Suc 0)^n = Suc 0"
   380 by simp
   381 
   382 text{*Valid for the naturals, but what if @{text"0<i<1"}?
   383 Premises cannot be weakened: consider the case where @{term "i=0"},
   384 @{term "m=1"} and @{term "n=0"}.*}
   385 lemma nat_power_less_imp_less:
   386   assumes nonneg: "0 < (i\<Colon>nat)"
   387   assumes less: "i^m < i^n"
   388   shows "m < n"
   389 proof (cases "i = 1")
   390   case True with less power_one [where 'a = nat] show ?thesis by simp
   391 next
   392   case False with nonneg have "1 < i" by auto
   393   from power_strict_increasing_iff [OF this] less show ?thesis ..
   394 qed
   395 
   396 lemma power_diff:
   397   assumes nz: "a ~= 0"
   398   shows "n <= m ==> (a::'a::{recpower, field}) ^ (m-n) = (a^m) / (a^n)"
   399   by (induct m n rule: diff_induct)
   400     (simp_all add: power_Suc nonzero_mult_divide_cancel_left nz)
   401 
   402 
   403 text{*ML bindings for the general exponentiation theorems*}
   404 ML
   405 {*
   406 val power_0 = thm"power_0";
   407 val power_Suc = thm"power_Suc";
   408 val power_0_Suc = thm"power_0_Suc";
   409 val power_0_left = thm"power_0_left";
   410 val power_one = thm"power_one";
   411 val power_one_right = thm"power_one_right";
   412 val power_add = thm"power_add";
   413 val power_mult = thm"power_mult";
   414 val power_mult_distrib = thm"power_mult_distrib";
   415 val zero_less_power = thm"zero_less_power";
   416 val zero_le_power = thm"zero_le_power";
   417 val one_le_power = thm"one_le_power";
   418 val gt1_imp_ge0 = thm"gt1_imp_ge0";
   419 val power_gt1_lemma = thm"power_gt1_lemma";
   420 val power_gt1 = thm"power_gt1";
   421 val power_le_imp_le_exp = thm"power_le_imp_le_exp";
   422 val power_inject_exp = thm"power_inject_exp";
   423 val power_less_imp_less_exp = thm"power_less_imp_less_exp";
   424 val power_mono = thm"power_mono";
   425 val power_strict_mono = thm"power_strict_mono";
   426 val power_eq_0_iff = thm"power_eq_0_iff";
   427 val field_power_eq_0_iff = thm"power_eq_0_iff";
   428 val field_power_not_zero = thm"field_power_not_zero";
   429 val power_inverse = thm"power_inverse";
   430 val nonzero_power_divide = thm"nonzero_power_divide";
   431 val power_divide = thm"power_divide";
   432 val power_abs = thm"power_abs";
   433 val zero_less_power_abs_iff = thm"zero_less_power_abs_iff";
   434 val zero_le_power_abs = thm "zero_le_power_abs";
   435 val power_minus = thm"power_minus";
   436 val power_Suc_less = thm"power_Suc_less";
   437 val power_strict_decreasing = thm"power_strict_decreasing";
   438 val power_decreasing = thm"power_decreasing";
   439 val power_Suc_less_one = thm"power_Suc_less_one";
   440 val power_increasing = thm"power_increasing";
   441 val power_strict_increasing = thm"power_strict_increasing";
   442 val power_le_imp_le_base = thm"power_le_imp_le_base";
   443 val power_inject_base = thm"power_inject_base";
   444 *}
   445 
   446 text{*ML bindings for the remaining theorems*}
   447 ML
   448 {*
   449 val nat_one_le_power = thm"nat_one_le_power";
   450 val nat_power_less_imp_less = thm"nat_power_less_imp_less";
   451 val nat_zero_less_power_iff = thm"nat_zero_less_power_iff";
   452 *}
   453 
   454 end