lean2/library/algebra/ring_power.lean
Jeremy Avigad 4a36f843f7 refactor(library/algebra/group_power,library/*): change definition of pow
I changed the definition of pow so that a^(succ n) reduces to a * a^n rather than a^n * a.

This has the nice effect that on nat and int, where multiplication is defined by recursion on the right,
a^1 reduces to a, and a^2 reduces to a * a.

The change was a pain in the neck, and in retrospect maybe not worth it, but oh, well.
2015-08-14 18:49:57 -07:00

63 lines
1.7 KiB
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
Copyright (c) 2015 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Author: Jeremy Avigad
Properties of the power operation in an ordered ring.
(Right now, this file is just a stub. More soon.)
-/
import .group_power
open nat
namespace algebra
variable {A : Type}
section linear_ordered_semiring
variable [s : linear_ordered_semiring A]
include s
theorem pow_pos_of_pos {x : A} (i : ) (H : x > 0) : x^i > 0 :=
begin
induction i with [j, ih],
{show (1 : A) > 0, from zero_lt_one},
{show x^(succ j) > 0, from mul_pos H ih}
end
theorem pow_nonneg_of_nonneg {x : A} (i : ) (H : x ≥ 0) : x^i ≥ 0 :=
begin
induction i with [j, ih],
{show (1 : A) ≥ 0, from le_of_lt zero_lt_one},
{show x^(succ j) ≥ 0, from mul_nonneg H ih}
end
theorem pow_le_pow_of_le {x y : A} (i : ) (H₁ : 0 ≤ x) (H₂ : x ≤ y) : x^i ≤ y^i :=
begin
induction i with [i, ih],
{rewrite *pow_zero, apply le.refl},
rewrite *pow_succ,
have H : 0 ≤ x^i, from pow_nonneg_of_nonneg i H₁,
apply mul_le_mul H₂ ih H (le.trans H₁ H₂)
end
theorem pow_ge_one {x : A} (i : ) (xge1 : x ≥ 1) : x^i ≥ 1 :=
assert H : x^i ≥ 1^i, from pow_le_pow_of_le i (le_of_lt zero_lt_one) xge1,
by rewrite one_pow at H; exact H
set_option formatter.hide_full_terms false
theorem pow_gt_one {x : A} {i : } (xgt1 : x > 1) (ipos : i > 0) : x^i > 1 :=
assert xpos : x > 0, from lt.trans zero_lt_one xgt1,
begin
induction i with [i, ih],
{exfalso, exact !nat.lt.irrefl ipos},
have xige1 : x^i ≥ 1, from pow_ge_one _ (le_of_lt xgt1),
rewrite [pow_succ, -mul_one 1, ↑has_lt.gt],
apply mul_lt_mul xgt1 xige1 zero_lt_one,
apply le_of_lt xpos
end
end linear_ordered_semiring
end algebra