2015-05-14 04:48:32 +00:00
|
|
|
|
/-
|
|
|
|
|
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 operation on monoids and groups. We separate this from group, because it depends on
|
|
|
|
|
nat, which in turn depends on other parts of algebra.
|
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
We have "pow a n" for natural number powers, and "gpow a i" for integer powers. The notation
|
|
|
|
|
a^n is used for the first, but users can locally redefine it to gpow when needed.
|
2015-05-14 04:48:32 +00:00
|
|
|
|
|
|
|
|
|
Note: power adopts the convention that 0^0=1.
|
|
|
|
|
-/
|
|
|
|
|
import data.nat.basic data.int.basic
|
|
|
|
|
|
|
|
|
|
variables {A : Type}
|
|
|
|
|
|
2015-10-08 18:01:57 +00:00
|
|
|
|
structure has_pow_nat [class] (A : Type) :=
|
|
|
|
|
(pow_nat : A → nat → A)
|
|
|
|
|
|
|
|
|
|
definition pow_nat {A : Type} [s : has_pow_nat A] : A → nat → A :=
|
|
|
|
|
has_pow_nat.pow_nat
|
|
|
|
|
|
|
|
|
|
infix ` ^ ` := pow_nat
|
|
|
|
|
|
|
|
|
|
structure has_pow_int [class] (A : Type) :=
|
|
|
|
|
(pow_int : A → int → A)
|
|
|
|
|
|
|
|
|
|
definition pow_int {A : Type} [s : has_pow_int A] : A → int → A :=
|
|
|
|
|
has_pow_int.pow_int
|
|
|
|
|
|
2015-10-09 01:35:37 +00:00
|
|
|
|
/- monoid -/
|
2015-05-14 04:48:32 +00:00
|
|
|
|
section monoid
|
|
|
|
|
open nat
|
2015-10-09 01:35:37 +00:00
|
|
|
|
|
2015-05-14 04:48:32 +00:00
|
|
|
|
variable [s : monoid A]
|
|
|
|
|
include s
|
|
|
|
|
|
2015-10-08 18:01:57 +00:00
|
|
|
|
definition monoid.pow (a : A) : ℕ → A
|
2015-05-14 04:48:32 +00:00
|
|
|
|
| 0 := 1
|
2015-10-08 18:01:57 +00:00
|
|
|
|
| (n+1) := a * monoid.pow n
|
2015-05-14 04:48:32 +00:00
|
|
|
|
|
2016-02-25 20:26:20 +00:00
|
|
|
|
definition monoid_has_pow_nat [instance] : has_pow_nat A :=
|
2015-10-08 18:01:57 +00:00
|
|
|
|
has_pow_nat.mk monoid.pow
|
2015-05-14 04:48:32 +00:00
|
|
|
|
|
|
|
|
|
theorem pow_zero (a : A) : a^0 = 1 := rfl
|
2015-08-14 15:00:16 +00:00
|
|
|
|
theorem pow_succ (a : A) (n : ℕ) : a^(succ n) = a * a^n := rfl
|
|
|
|
|
|
|
|
|
|
theorem pow_one (a : A) : a^1 = a := !mul_one
|
|
|
|
|
theorem pow_two (a : A) : a^2 = a * a :=
|
|
|
|
|
calc
|
|
|
|
|
a^2 = a * (a * 1) : rfl
|
|
|
|
|
... = a * a : mul_one
|
|
|
|
|
theorem pow_three (a : A) : a^3 = a * (a * a) :=
|
|
|
|
|
calc
|
|
|
|
|
a^3 = a * (a * (a * 1)) : rfl
|
|
|
|
|
... = a * (a * a) : mul_one
|
|
|
|
|
theorem pow_four (a : A) : a^4 = a * (a * (a * a)) :=
|
|
|
|
|
calc
|
|
|
|
|
a^4 = a * a^3 : rfl
|
|
|
|
|
... = a * (a * (a * a)) : pow_three
|
|
|
|
|
|
|
|
|
|
theorem pow_succ' (a : A) : ∀n, a^(succ n) = a^n * a
|
2015-05-14 04:48:32 +00:00
|
|
|
|
| 0 := by rewrite [pow_succ, *pow_zero, one_mul, mul_one]
|
|
|
|
|
| (succ n) := by rewrite [pow_succ, pow_succ' at {1}, pow_succ, mul.assoc]
|
|
|
|
|
|
2015-05-18 22:45:23 +00:00
|
|
|
|
theorem one_pow : ∀ n : ℕ, 1^n = (1:A)
|
2015-05-14 04:48:32 +00:00
|
|
|
|
| 0 := rfl
|
2015-08-14 15:00:16 +00:00
|
|
|
|
| (succ n) := by rewrite [pow_succ, one_mul, one_pow]
|
2015-05-14 04:48:32 +00:00
|
|
|
|
|
2015-08-14 15:00:16 +00:00
|
|
|
|
theorem pow_add (a : A) (m n : ℕ) : a^(m + n) = a^m * a^n :=
|
|
|
|
|
begin
|
|
|
|
|
induction n with n ih,
|
2016-02-04 21:15:42 +00:00
|
|
|
|
{krewrite [nat.add_zero, pow_zero, mul_one]},
|
2015-08-14 15:00:16 +00:00
|
|
|
|
rewrite [add_succ, *pow_succ', ih, mul.assoc]
|
|
|
|
|
end
|
2015-05-14 04:48:32 +00:00
|
|
|
|
|
|
|
|
|
theorem pow_mul (a : A) (m : ℕ) : ∀ n, a^(m * n) = (a^m)^n
|
|
|
|
|
| 0 := by rewrite [nat.mul_zero, pow_zero]
|
2015-08-14 15:00:16 +00:00
|
|
|
|
| (succ n) := by rewrite [nat.mul_succ, pow_add, pow_succ', pow_mul]
|
2015-05-14 04:48:32 +00:00
|
|
|
|
|
|
|
|
|
theorem pow_comm (a : A) (m n : ℕ) : a^m * a^n = a^n * a^m :=
|
2015-10-14 19:27:09 +00:00
|
|
|
|
by rewrite [-*pow_add, add.comm]
|
2015-05-14 04:48:32 +00:00
|
|
|
|
|
|
|
|
|
end monoid
|
|
|
|
|
|
|
|
|
|
/- commutative monoid -/
|
|
|
|
|
|
|
|
|
|
section comm_monoid
|
|
|
|
|
open nat
|
|
|
|
|
variable [s : comm_monoid A]
|
|
|
|
|
include s
|
|
|
|
|
|
|
|
|
|
theorem mul_pow (a b : A) : ∀ n, (a * b)^n = a^n * b^n
|
|
|
|
|
| 0 := by rewrite [*pow_zero, mul_one]
|
2015-08-14 15:00:16 +00:00
|
|
|
|
| (succ n) := by rewrite [*pow_succ', mul_pow, *mul.assoc, mul.left_comm a]
|
2015-05-14 04:48:32 +00:00
|
|
|
|
|
|
|
|
|
end comm_monoid
|
|
|
|
|
|
|
|
|
|
section group
|
|
|
|
|
variable [s : group A]
|
|
|
|
|
include s
|
|
|
|
|
|
|
|
|
|
section nat
|
|
|
|
|
open nat
|
|
|
|
|
theorem inv_pow (a : A) : ∀n, (a⁻¹)^n = (a^n)⁻¹
|
2015-05-23 04:05:06 +00:00
|
|
|
|
| 0 := by rewrite [*pow_zero, one_inv]
|
|
|
|
|
| (succ n) := by rewrite [pow_succ, pow_succ', inv_pow, mul_inv]
|
2015-05-14 04:48:32 +00:00
|
|
|
|
|
|
|
|
|
theorem pow_sub (a : A) {m n : ℕ} (H : m ≥ n) : a^(m - n) = a^m * (a^n)⁻¹ :=
|
|
|
|
|
assert H1 : m - n + n = m, from nat.sub_add_cancel H,
|
|
|
|
|
have H2 : a^(m - n) * a^n = a^m, by rewrite [-pow_add, H1],
|
|
|
|
|
eq_mul_inv_of_mul_eq H2
|
|
|
|
|
|
|
|
|
|
theorem pow_inv_comm (a : A) : ∀m n, (a⁻¹)^m * a^n = a^n * (a⁻¹)^m
|
|
|
|
|
| 0 n := by rewrite [*pow_zero, one_mul, mul_one]
|
|
|
|
|
| m 0 := by rewrite [*pow_zero, one_mul, mul_one]
|
2015-08-14 15:00:16 +00:00
|
|
|
|
| (succ m) (succ n) := by rewrite [pow_succ' at {1}, pow_succ at {1}, pow_succ', pow_succ,
|
2015-05-14 04:48:32 +00:00
|
|
|
|
*mul.assoc, inv_mul_cancel_left, mul_inv_cancel_left, pow_inv_comm]
|
|
|
|
|
|
|
|
|
|
end nat
|
|
|
|
|
|
2015-10-09 01:35:37 +00:00
|
|
|
|
open int
|
2015-05-14 04:48:32 +00:00
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
definition gpow (a : A) : ℤ → A
|
2015-05-14 04:48:32 +00:00
|
|
|
|
| (of_nat n) := a^n
|
2015-06-29 05:16:58 +00:00
|
|
|
|
| -[1+n] := (a^(nat.succ n))⁻¹
|
2015-05-14 04:48:32 +00:00
|
|
|
|
|
2015-10-09 01:35:37 +00:00
|
|
|
|
open nat
|
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
private lemma gpow_add_aux (a : A) (m n : nat) :
|
|
|
|
|
gpow a ((of_nat m) + -[1+n]) = gpow a (of_nat m) * gpow a (-[1+n]) :=
|
2015-05-14 04:48:32 +00:00
|
|
|
|
or.elim (nat.lt_or_ge m (nat.succ n))
|
2015-10-07 23:44:47 +00:00
|
|
|
|
(assume H : (m < nat.succ n),
|
2015-10-09 01:35:37 +00:00
|
|
|
|
assert H1 : (#nat nat.succ n - m > nat.zero), from nat.sub_pos_of_lt H,
|
2015-05-14 04:48:32 +00:00
|
|
|
|
calc
|
2015-08-04 20:02:45 +00:00
|
|
|
|
gpow a ((of_nat m) + -[1+n]) = gpow a (sub_nat_nat m (nat.succ n)) : rfl
|
2015-10-09 01:35:37 +00:00
|
|
|
|
... = gpow a (-[1+ nat.pred (nat.sub (nat.succ n) m)]) : {sub_nat_nat_of_lt H}
|
|
|
|
|
... = (a ^ (nat.succ (nat.pred (nat.sub (nat.succ n) m))))⁻¹ : rfl
|
|
|
|
|
... = (a ^ (nat.succ n) * (a ^ m)⁻¹)⁻¹ :
|
|
|
|
|
by krewrite [succ_pred_of_pos H1, pow_sub a (nat.le_of_lt H)]
|
|
|
|
|
... = a ^ m * (a ^ (nat.succ n))⁻¹ :
|
2015-05-23 04:05:06 +00:00
|
|
|
|
by rewrite [mul_inv, inv_inv]
|
2015-08-04 20:02:45 +00:00
|
|
|
|
... = gpow a (of_nat m) * gpow a (-[1+n]) : rfl)
|
2015-10-09 01:35:37 +00:00
|
|
|
|
(assume H : (m ≥ nat.succ n),
|
2015-05-14 04:48:32 +00:00
|
|
|
|
calc
|
2015-08-04 20:02:45 +00:00
|
|
|
|
gpow a ((of_nat m) + -[1+n]) = gpow a (sub_nat_nat m (nat.succ n)) : rfl
|
|
|
|
|
... = gpow a (#nat m - nat.succ n) : {sub_nat_nat_of_ge H}
|
2015-10-09 01:35:37 +00:00
|
|
|
|
... = a ^ m * (a ^ (nat.succ n))⁻¹ : pow_sub a H
|
2015-08-04 20:02:45 +00:00
|
|
|
|
... = gpow a (of_nat m) * gpow a (-[1+n]) : rfl)
|
2015-05-14 04:48:32 +00:00
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
theorem gpow_add (a : A) : ∀i j : int, gpow a (i + j) = gpow a i * gpow a j
|
2015-05-14 04:48:32 +00:00
|
|
|
|
| (of_nat m) (of_nat n) := !pow_add
|
2015-08-04 20:02:45 +00:00
|
|
|
|
| (of_nat m) -[1+n] := !gpow_add_aux
|
2015-10-22 20:36:45 +00:00
|
|
|
|
| -[1+m] (of_nat n) := by rewrite [add.comm, gpow_add_aux, ↑gpow, -*inv_pow, pow_inv_comm]
|
2015-06-29 05:16:58 +00:00
|
|
|
|
| -[1+m] -[1+n] :=
|
2015-05-14 04:48:32 +00:00
|
|
|
|
calc
|
2015-08-04 20:02:45 +00:00
|
|
|
|
gpow a (-[1+m] + -[1+n]) = (a^(#nat nat.succ m + nat.succ n))⁻¹ : rfl
|
2015-05-23 04:05:06 +00:00
|
|
|
|
... = (a^(nat.succ m))⁻¹ * (a^(nat.succ n))⁻¹ : by rewrite [pow_add, pow_comm, mul_inv]
|
2015-08-04 20:02:45 +00:00
|
|
|
|
... = gpow a (-[1+m]) * gpow a (-[1+n]) : rfl
|
2015-05-14 04:48:32 +00:00
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
theorem gpow_comm (a : A) (i j : ℤ) : gpow a i * gpow a j = gpow a j * gpow a i :=
|
2015-10-22 20:36:45 +00:00
|
|
|
|
by rewrite [-*gpow_add, add.comm]
|
2015-05-14 04:48:32 +00:00
|
|
|
|
end group
|
2015-10-09 01:35:37 +00:00
|
|
|
|
|
2015-07-28 21:44:56 +00:00
|
|
|
|
section ordered_ring
|
|
|
|
|
open nat
|
|
|
|
|
variable [s : linear_ordered_ring A]
|
|
|
|
|
include s
|
|
|
|
|
|
2015-10-09 01:35:37 +00:00
|
|
|
|
theorem pow_pos {a : A} (H : a > 0) (n : ℕ) : a ^ n > 0 :=
|
2015-07-28 21:44:56 +00:00
|
|
|
|
begin
|
|
|
|
|
induction n,
|
2016-02-04 21:15:42 +00:00
|
|
|
|
krewrite pow_zero,
|
2015-07-28 21:44:56 +00:00
|
|
|
|
apply zero_lt_one,
|
2015-08-14 15:00:16 +00:00
|
|
|
|
rewrite pow_succ',
|
2015-07-28 21:44:56 +00:00
|
|
|
|
apply mul_pos,
|
|
|
|
|
apply v_0, apply H
|
|
|
|
|
end
|
|
|
|
|
|
2015-10-09 01:35:37 +00:00
|
|
|
|
theorem pow_ge_one_of_ge_one {a : A} (H : a ≥ 1) (n : ℕ) : a ^ n ≥ 1 :=
|
2015-07-28 21:44:56 +00:00
|
|
|
|
begin
|
|
|
|
|
induction n,
|
2016-02-04 21:15:42 +00:00
|
|
|
|
krewrite pow_zero,
|
2015-07-28 21:44:56 +00:00
|
|
|
|
apply le.refl,
|
2015-10-09 01:35:37 +00:00
|
|
|
|
rewrite [pow_succ', -mul_one 1],
|
2015-07-28 21:44:56 +00:00
|
|
|
|
apply mul_le_mul v_0 H zero_le_one,
|
|
|
|
|
apply le_of_lt,
|
|
|
|
|
apply pow_pos,
|
|
|
|
|
apply gt_of_ge_of_gt H zero_lt_one
|
|
|
|
|
end
|
|
|
|
|
|
2015-10-12 03:29:31 +00:00
|
|
|
|
theorem pow_two_add (n : ℕ) : (2:A)^n + 2^n = 2^(succ n) :=
|
|
|
|
|
by rewrite [pow_succ', -one_add_one_eq_two, left_distrib, *mul_one]
|
2015-07-28 21:44:56 +00:00
|
|
|
|
|
|
|
|
|
end ordered_ring
|
|
|
|
|
|
2015-08-03 16:25:33 +00:00
|
|
|
|
/- additive monoid -/
|
|
|
|
|
|
|
|
|
|
section add_monoid
|
|
|
|
|
variable [s : add_monoid A]
|
|
|
|
|
include s
|
2015-10-09 20:21:03 +00:00
|
|
|
|
local attribute add_monoid.to_monoid [trans_instance]
|
2015-08-03 16:25:33 +00:00
|
|
|
|
open nat
|
|
|
|
|
|
2015-10-09 01:35:37 +00:00
|
|
|
|
definition nmul : ℕ → A → A := λ n a, a^n
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
infix [priority algebra.prio] `⬝` := nmul
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
theorem zero_nmul (a : A) : (0:ℕ) ⬝ a = 0 := pow_zero a
|
2015-08-14 15:00:16 +00:00
|
|
|
|
theorem succ_nmul (n : ℕ) (a : A) : nmul (succ n) a = a + (nmul n a) := pow_succ a n
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
2015-08-14 15:00:16 +00:00
|
|
|
|
theorem succ_nmul' (n : ℕ) (a : A) : succ n ⬝ a = nmul n a + a := pow_succ' a n
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
theorem nmul_zero (n : ℕ) : n ⬝ 0 = (0:A) := one_pow n
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
theorem one_nmul (a : A) : 1 ⬝ a = a := pow_one a
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
theorem add_nmul (m n : ℕ) (a : A) : (m + n) ⬝ a = (m ⬝ a) + (n ⬝ a) := pow_add a m n
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
2015-10-14 19:27:09 +00:00
|
|
|
|
theorem mul_nmul (m n : ℕ) (a : A) : (m * n) ⬝ a = m ⬝ (n ⬝ a) := eq.subst (mul.comm n m) (pow_mul a n m)
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
theorem nmul_comm (m n : ℕ) (a : A) : (m ⬝ a) + (n ⬝ a) = (n ⬝ a) + (m ⬝ a) := pow_comm a m n
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
|
|
|
|
end add_monoid
|
|
|
|
|
|
|
|
|
|
/- additive commutative monoid -/
|
|
|
|
|
|
|
|
|
|
section add_comm_monoid
|
|
|
|
|
open nat
|
|
|
|
|
variable [s : add_comm_monoid A]
|
|
|
|
|
include s
|
2015-10-09 20:21:03 +00:00
|
|
|
|
local attribute add_comm_monoid.to_comm_monoid [trans_instance]
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
theorem nmul_add (n : ℕ) (a b : A) : n ⬝ (a + b) = (n ⬝ a) + (n ⬝ b) := mul_pow a b n
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
|
|
|
|
end add_comm_monoid
|
|
|
|
|
|
|
|
|
|
section add_group
|
|
|
|
|
variable [s : add_group A]
|
|
|
|
|
include s
|
2015-10-09 20:21:03 +00:00
|
|
|
|
local attribute add_group.to_group [trans_instance]
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
|
|
|
|
section nat
|
|
|
|
|
open nat
|
2015-08-04 20:02:45 +00:00
|
|
|
|
theorem nmul_neg (n : ℕ) (a : A) : n ⬝ (-a) = -(n ⬝ a) := inv_pow a n
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
theorem sub_nmul {m n : ℕ} (a : A) (H : m ≥ n) : (m - n) ⬝ a = (m ⬝ a) + -(n ⬝ a) := pow_sub a H
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
theorem nmul_neg_comm (m n : ℕ) (a : A) : (m ⬝ (-a)) + (n ⬝ a) = (n ⬝ a) + (m ⬝ (-a)) := pow_inv_comm a m n
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
|
|
|
|
end nat
|
|
|
|
|
|
|
|
|
|
open int
|
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
definition imul : ℤ → A → A := λ i a, gpow a i
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
theorem add_imul (i j : ℤ) (a : A) : imul (i + j) a = imul i a + imul j a :=
|
|
|
|
|
gpow_add a i j
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
2015-08-04 20:02:45 +00:00
|
|
|
|
theorem imul_comm (i j : ℤ) (a : A) : imul i a + imul j a = imul j a + imul i a := gpow_comm a i j
|
2015-08-03 16:25:33 +00:00
|
|
|
|
|
|
|
|
|
end add_group
|