42fbc63bb6
@avigad, @fpvandoorn, @rlewis1988, @dselsam I changed how transitive instances are named. The motivation is to avoid a naming collision problem found by Daniel. Before this commit, we were getting an error on the following file tests/lean/run/collision_bug.lean. Now, transitive instances contain the prefix "_trans_". It makes it clear this is an internal definition and it should not be used by users. This change also demonstrates (again) how the `rewrite` tactic is fragile. The problem is that the matching procedure used by it has very little support for solving matching constraints that involving type class instances. Eventually, we will need to reimplement `rewrite` using the new unification procedure used in blast. In the meantime, the workaround is to use `krewrite` (as usual).
29 lines
775 B
Text
29 lines
775 B
Text
/-
|
||
Copyright (c) 2015 Jeremy Avigad. All rights reserved.
|
||
Released under Apache 2.0 license as described in the file LICENSE.
|
||
Author: Jeremy Avigad
|
||
|
||
The power function on the integers.
|
||
-/
|
||
import data.int.basic data.int.order data.int.div algebra.group_power data.nat.power
|
||
|
||
namespace int
|
||
|
||
definition int_has_pow_nat [reducible] [instance] [priority int.prio] : has_pow_nat int :=
|
||
has_pow_nat.mk has_pow_nat.pow_nat
|
||
|
||
/-
|
||
definition nmul (n : ℕ) (a : ℤ) : ℤ := algebra.nmul n a
|
||
infix [priority int.prio] ⬝ := nmul
|
||
definition imul (i : ℤ) (a : ℤ) : ℤ := algebra.imul i a
|
||
-/
|
||
|
||
open nat
|
||
theorem of_nat_pow (a n : ℕ) : of_nat (a^n) = (of_nat a)^n :=
|
||
begin
|
||
induction n with n ih,
|
||
apply eq.refl,
|
||
krewrite [pow_succ, pow_succ, of_nat_mul, ih]
|
||
end
|
||
|
||
end int
|