refactor(library/algebra/{field,ordered_field}, library/*): more renaming, setting implicit arguments

Many theorems for division rings and fields have stronger versions for discrete fields, where we
assume x / 0 = 0. Before, we used primes to distinguish the versions, but that has the downside
that e.g. for rat and real, all the theorems are equally present. Now, I qualified the weaker ones
with division_ring.foo or field.foo. Maybe that is not ideal, but let's try it.

I also set implicit arguments with the following convention: an argument to a theorem should be
explicit unless it can be inferred from the other arguments and hypotheses.
This commit is contained in:
Jeremy Avigad 2015-08-27 13:29:19 -04:00 committed by Leonardo de Moura
parent 93c2d1e9d0
commit 7d72c9b6b5
10 changed files with 247 additions and 219 deletions

View file

@ -6,7 +6,6 @@ Authors: Robert Lewis
Structures with multiplicative and additive components, including division rings and fields.
The development is modeled after Isabelle's library.
-/
----------------------------------------------------------------------------------------------------
import logic.eq logic.connectives data.unit data.sigma data.prod
import algebra.binary algebra.group algebra.ring
open eq eq.ops
@ -37,23 +36,11 @@ section division_ring
theorem inv_mul_cancel (H : a ≠ 0) : a⁻¹ * a = 1 :=
division_ring.inv_mul_cancel H
theorem inv_eq_one_div : a⁻¹ = 1 / a := !one_mul⁻¹
theorem inv_eq_one_div (a : A) : a⁻¹ = 1 / a := !one_mul⁻¹
-- the following are only theorems if we assume inv_zero here
/- theorem inv_zero : 0⁻¹ = 0 := !division_ring.inv_zero
theorem one_div_zero : 1 / 0 = 0 :=
calc
1 / 0 = 1 * 0⁻¹ : refl
... = 1 * 0 : division_ring.inv_zero A
... = 0 : mul_zero
-/
theorem div_eq_mul_one_div : a / b = a * (1 / b) :=
theorem div_eq_mul_one_div (a b : A) : a / b = a * (1 / b) :=
by rewrite [↑divide, one_mul]
-- theorem div_zero : a / 0 = 0 := by rewrite [div_eq_mul_one_div, one_div_zero, mul_zero]
theorem mul_one_div_cancel (H : a ≠ 0) : a * (1 / a) = 1 :=
by rewrite [-inv_eq_one_div, (mul_inv_cancel H)]
@ -64,36 +51,32 @@ section division_ring
theorem one_div_one : 1 / 1 = (1:A) := div_self (ne.symm zero_ne_one)
theorem mul_div_assoc : (a * b) / c = a * (b / c) := !mul.assoc
theorem mul_div_assoc (a b : A) : (a * b) / c = a * (b / c) := !mul.assoc
theorem one_div_ne_zero (H : a ≠ 0) : 1 / a ≠ 0 :=
assume H2 : 1 / a = 0,
have C1 : 0 = (1:A), from symm (by rewrite [-(mul_one_div_cancel H), H2, mul_zero]),
absurd C1 zero_ne_one
-- theorem ne_zero_of_one_div_ne_zero (H : 1 / a ≠ 0) : a ≠ 0 :=
-- assume Ha : a = 0, absurd (Ha⁻¹ ▸ one_div_zero) H
theorem one_inv_eq : 1⁻¹ = (1:A) :=
by rewrite [-mul_one, inv_mul_cancel (ne.symm (@zero_ne_one A _))]
theorem div_one : a / 1 = a :=
theorem div_one (a : A) : a / 1 = a :=
by rewrite [↑divide, one_inv_eq, mul_one]
theorem zero_div : 0 / a = 0 := !zero_mul
theorem zero_div (a : A) : 0 / a = 0 := !zero_mul
-- note: integral domain has a "mul_ne_zero". Discrete fields are int domains.
theorem mul_ne_zero' (Ha : a ≠ 0) (Hb : b ≠ 0) : a * b ≠ 0 :=
-- note: integral domain has a "mul_ne_zero". A commutative division ring is an integral
-- domain, but let's not define that class for now.
theorem division_ring.mul_ne_zero (Ha : a ≠ 0) (Hb : b ≠ 0) : a * b ≠ 0 :=
assume H : a * b = 0,
have C1 : a = 0, by rewrite [-mul_one, -(mul_one_div_cancel Hb), -mul.assoc, H, zero_mul],
absurd C1 Ha
theorem mul_ne_zero_comm (H : a * b ≠ 0) : b * a ≠ 0 :=
have H2 : a ≠ 0 ∧ b ≠ 0, from ne_zero_and_ne_zero_of_mul_ne_zero H,
mul_ne_zero' (and.right H2) (and.left H2)
division_ring.mul_ne_zero (and.right H2) (and.left H2)
-- make "left" and "right" versions?
theorem eq_one_div_of_mul_eq_one (H : a * b = 1) : b = 1 / a :=
have a ≠ 0, from
(suppose a = 0,
@ -106,7 +89,6 @@ section division_ring
... = 1 * b : one_div_mul_cancel this
... = b : one_mul)
-- which one is left and which is right?
theorem eq_one_div_of_mul_eq_one_left (H : b * a = 1) : b = 1 / a :=
have a ≠ 0, from
(suppose a = 0,
@ -122,16 +104,18 @@ section division_ring
... = b * 1 : mul_one_div_cancel this
... = b : mul_one)
theorem one_div_mul_one_div (Ha : a ≠ 0) (Hb : b ≠ 0) : (1 / a) * (1 / b) = 1 / (b * a) :=
theorem division_ring.one_div_mul_one_div (Ha : a ≠ 0) (Hb : b ≠ 0) :
(1 / a) * (1 / b) = 1 / (b * a) :=
have (b * a) * ((1 / a) * (1 / b)) = 1, by
rewrite [mul.assoc, -(mul.assoc a), (mul_one_div_cancel Ha), one_mul, (mul_one_div_cancel Hb)],
rewrite [mul.assoc, -(mul.assoc a), (mul_one_div_cancel Ha), one_mul,
(mul_one_div_cancel Hb)],
eq_one_div_of_mul_eq_one this
theorem one_div_neg_one_eq_neg_one : (1:A) / (-1) = -1 :=
have (-1) * (-1) = (1:A), by rewrite [-neg_eq_neg_one_mul, neg_neg],
symm (eq_one_div_of_mul_eq_one this)
theorem one_div_neg_eq_neg_one_div (H : a ≠ 0) : 1 / (- a) = - (1 / a) :=
theorem division_ring.one_div_neg_eq_neg_one_div (H : a ≠ 0) : 1 / (- a) = - (1 / a) :=
have -1 ≠ 0, from
(suppose -1 = 0, absurd (symm (calc
1 = -(-1) : neg_neg
@ -139,86 +123,87 @@ section division_ring
... = (0:A) : neg_zero)) zero_ne_one),
calc
1 / (- a) = 1 / ((-1) * a) : neg_eq_neg_one_mul
... = (1 / a) * (1 / (- 1)) : one_div_mul_one_div H this
... = (1 / a) * (1 / (- 1)) : division_ring.one_div_mul_one_div H this
... = (1 / a) * (-1) : one_div_neg_one_eq_neg_one
... = - (1 / a) : mul_neg_one_eq_neg
theorem div_neg_eq_neg_div (Ha : a ≠ 0) : b / (- a) = - (b / a) :=
theorem div_neg_eq_neg_div (b : A) (Ha : a ≠ 0) : b / (- a) = - (b / a) :=
calc
b / (- a) = b * (1 / (- a)) : inv_eq_one_div
... = b * -(1 / a) : one_div_neg_eq_neg_one_div Ha
... = b * -(1 / a) : division_ring.one_div_neg_eq_neg_one_div Ha
... = -(b * (1 / a)) : neg_mul_eq_mul_neg
... = - (b * a⁻¹) : inv_eq_one_div
theorem neg_div : (-b) / a = - (b / a) :=
theorem neg_div (a b : A) : (-b) / a = - (b / a) :=
by rewrite [neg_eq_neg_one_mul, mul_div_assoc, -neg_eq_neg_one_mul]
theorem neg_div_neg_eq_div (Hb : b ≠ 0) : (-a) / (-b) = a / b :=
by rewrite [(div_neg_eq_neg_div Hb), neg_div, neg_neg]
theorem division_ring.neg_div_neg_eq (a : A) {b : A} (Hb : b ≠ 0) : (-a) / (-b) = a / b :=
by rewrite [(div_neg_eq_neg_div _ Hb), neg_div, neg_neg]
theorem div_div (H : a ≠ 0) : 1 / (1 / a) = a :=
theorem division_ring.one_div_one_div (H : a ≠ 0) : 1 / (1 / a) = a :=
symm (eq_one_div_of_mul_eq_one_left (mul_one_div_cancel H))
theorem eq_of_invs_eq (Ha : a ≠ 0) (Hb : b ≠ 0) (H : 1 / a = 1 / b) : a = b :=
by rewrite [-(div_div Ha), H, (div_div Hb)]
theorem division_ring.eq_of_one_div_eq_one_div (Ha : a ≠ 0) (Hb : b ≠ 0) (H : 1 / a = 1 / b) :
a = b :=
by rewrite [-(division_ring.one_div_one_div Ha), H, (division_ring.one_div_one_div Hb)]
theorem mul_inv_eq (Ha : a ≠ 0) (Hb : b ≠ 0) : (b * a)⁻¹ = a⁻¹ * b⁻¹ :=
eq.symm (calc
a⁻¹ * b⁻¹ = (1 / a) * b⁻¹ : inv_eq_one_div
... = (1 / a) * (1 / b) : inv_eq_one_div
... = (1 / (b * a)) : one_div_mul_one_div Ha Hb
... = (1 / (b * a)) : division_ring.one_div_mul_one_div Ha Hb
... = (b * a)⁻¹ : inv_eq_one_div)
theorem mul_div_cancel (Hb : b ≠ 0) : a * b / b = a :=
theorem mul_div_cancel (a : A) {b : A} (Hb : b ≠ 0) : a * b / b = a :=
by rewrite [↑divide, mul.assoc, (mul_inv_cancel Hb), mul_one]
theorem div_mul_cancel (Hb : b ≠ 0) : a / b * b = a :=
theorem div_mul_cancel (a : A) {b : A} (Hb : b ≠ 0) : a / b * b = a :=
by rewrite [↑divide, mul.assoc, (inv_mul_cancel Hb), mul_one]
theorem div_add_div_same : a / c + b / c = (a + b) / c := !right_distrib⁻¹
theorem div_add_div_same (a b c : A) : a / c + b / c = (a + b) / c := !right_distrib⁻¹
theorem div_sub_div_same : (a / c) - (b / c) = (a - b) / c :=
theorem div_sub_div_same (a b c : A) : (a / c) - (b / c) = (a - b) / c :=
by rewrite [sub_eq_add_neg, -neg_div, div_add_div_same]
theorem inv_mul_add_mul_inv_eq_inv_add_inv (Ha : a ≠ 0) (Hb : b ≠ 0) :
theorem one_div_mul_add_mul_one_div_eq_one_div_add_one_div (Ha : a ≠ 0) (Hb : b ≠ 0) :
(1 / a) * (a + b) * (1 / b) = 1 / a + 1 / b :=
by rewrite [(left_distrib (1 / a)), (one_div_mul_cancel Ha), right_distrib, one_mul,
mul.assoc, (mul_one_div_cancel Hb), mul_one, add.comm]
theorem inv_mul_sub_mul_inv_eq_inv_add_inv (Ha : a ≠ 0) (Hb : b ≠ 0) :
theorem one_div_mul_sub_mul_one_div_eq_one_div_add_one_div (Ha : a ≠ 0) (Hb : b ≠ 0) :
(1 / a) * (b - a) * (1 / b) = 1 / a - 1 / b :=
by rewrite [(mul_sub_left_distrib (1 / a)), (one_div_mul_cancel Ha), mul_sub_right_distrib,
one_mul, mul.assoc, (mul_one_div_cancel Hb), mul_one, one_mul]
theorem div_eq_one_iff_eq (Hb : b ≠ 0) : a / b = 1 ↔ a = b :=
theorem div_eq_one_iff_eq (a : A) {b : A} (Hb : b ≠ 0) : a / b = 1 ↔ a = b :=
iff.intro
(suppose a / b = 1, symm (calc
b = 1 * b : one_mul
... = a / b * b : this
... = a : div_mul_cancel Hb))
... = a : div_mul_cancel _ Hb))
(suppose a = b, calc
a / b = b / b : this
... = 1 : div_self Hb)
theorem eq_of_div_eq_one (Hb : b ≠ 0) : a / b = 1 → a = b :=
iff.mp (div_eq_one_iff_eq Hb)
theorem eq_of_div_eq_one (a : A) {b : A} (Hb : b ≠ 0) : a / b = 1 → a = b :=
iff.mp (!div_eq_one_iff_eq Hb)
theorem eq_div_iff_mul_eq (Hc : c ≠ 0) : a = b / c ↔ a * c = b :=
theorem eq_div_iff_mul_eq (a : A) {b : A} (Hc : c ≠ 0) : a = b / c ↔ a * c = b :=
iff.intro
(suppose a = b / c, by rewrite [this, (div_mul_cancel Hc)])
(suppose a * c = b, by rewrite [-(mul_div_cancel Hc), this])
(suppose a = b / c, by rewrite [this, (!div_mul_cancel Hc)])
(suppose a * c = b, by rewrite [-(!mul_div_cancel Hc), this])
theorem eq_div_of_mul_eq (Hc : c ≠ 0) : a * c = b → a = b / c :=
iff.mpr (eq_div_iff_mul_eq Hc)
theorem eq_div_of_mul_eq (a b : A) {c : A} (Hc : c ≠ 0) : a * c = b → a = b / c :=
iff.mpr (!eq_div_iff_mul_eq Hc)
theorem mul_eq_of_eq_div (Hc : c ≠ 0) : a = b / c → a * c = b :=
iff.mp (eq_div_iff_mul_eq Hc)
theorem mul_eq_of_eq_div (a b: A) {c : A} (Hc : c ≠ 0) : a = b / c → a * c = b :=
iff.mp (!eq_div_iff_mul_eq Hc)
theorem add_div_eq_mul_add_div (Hc : c ≠ 0) : a + b / c = (a * c + b) / c :=
have (a + b / c) * c = a * c + b, by rewrite [right_distrib, (div_mul_cancel Hc)],
(iff.elim_right (eq_div_iff_mul_eq Hc)) this
theorem add_div_eq_mul_add_div (a b : A) {c : A} (Hc : c ≠ 0) : a + b / c = (a * c + b) / c :=
have (a + b / c) * c = a * c + b, by rewrite [right_distrib, (!div_mul_cancel Hc)],
(iff.elim_right (!eq_div_iff_mul_eq Hc)) this
theorem mul_mul_div (Hc : c ≠ 0) : a = a * c * (1 / c) :=
theorem mul_mul_div (a : A) {c : A} (Hc : c ≠ 0) : a = a * c * (1 / c) :=
calc
a = a * 1 : mul_one
... = a * (c * (1 / c)) : mul_one_div_cancel Hc
@ -226,7 +211,6 @@ section division_ring
-- There are many similar rules to these last two in the Isabelle library
-- that haven't been ported yet. Do as necessary.
end division_ring
structure field [class] (A : Type) extends division_ring A, comm_ring A
@ -236,83 +220,92 @@ section field
include s
local attribute divide [reducible]
theorem one_div_mul_one_div' (Ha : a ≠ 0) (Hb : b ≠ 0) : (1 / a) * (1 / b) = 1 / (a * b) :=
by rewrite [(one_div_mul_one_div Ha Hb), mul.comm b]
theorem field.one_div_mul_one_div (Ha : a ≠ 0) (Hb : b ≠ 0) : (1 / a) * (1 / b) = 1 / (a * b) :=
by rewrite [(division_ring.one_div_mul_one_div Ha Hb), mul.comm b]
theorem div_mul_right (Hb : b ≠ 0) (H : a * b ≠ 0) : a / (a * b) = 1 / b :=
theorem field.div_mul_right (Hb : b ≠ 0) (H : a * b ≠ 0) : a / (a * b) = 1 / b :=
have a ≠ 0, from and.left (ne_zero_and_ne_zero_of_mul_ne_zero H),
symm (calc
1 / b = 1 * (1 / b) : one_mul
... = (a * a⁻¹) * (1 / b) : mul_inv_cancel this
... = a * (a⁻¹ * (1 / b)) : mul.assoc
... = a * ((1 / a) * (1 / b)) :inv_eq_one_div
... = a * (1 / (b * a)) : one_div_mul_one_div this Hb
... = a * ((1 / a) * (1 / b)) : inv_eq_one_div
... = a * (1 / (b * a)) : division_ring.one_div_mul_one_div this Hb
... = a * (1 / (a * b)) : mul.comm
... = a * (a * b)⁻¹ : inv_eq_one_div)
theorem div_mul_left (Ha : a ≠ 0) (H : a * b ≠ 0) : b / (a * b) = 1 / a :=
theorem field.div_mul_left (Ha : a ≠ 0) (H : a * b ≠ 0) : b / (a * b) = 1 / a :=
let H1 : b * a ≠ 0 := mul_ne_zero_comm H in
by rewrite [mul.comm a, (div_mul_right Ha H1)]
by rewrite [mul.comm a, (field.div_mul_right Ha H1)]
theorem mul_div_cancel_left (Ha : a ≠ 0) : a * b / a = b :=
by rewrite [mul.comm a, (mul_div_cancel Ha)]
by rewrite [mul.comm a, (!mul_div_cancel Ha)]
theorem mul_div_cancel' (Hb : b ≠ 0) : b * (a / b) = a :=
by rewrite [mul.comm, (div_mul_cancel Hb)]
by rewrite [mul.comm, (!div_mul_cancel Hb)]
theorem one_div_add_one_div (Ha : a ≠ 0) (Hb : b ≠ 0) : 1 / a + 1 / b = (a + b) / (a * b) :=
assert a * b ≠ 0, from (mul_ne_zero' Ha Hb),
by rewrite [add.comm, -(div_mul_left Ha this), -(div_mul_right Hb this), ↑divide, -right_distrib]
assert a * b ≠ 0, from (division_ring.mul_ne_zero Ha Hb),
by rewrite [add.comm, -(field.div_mul_left Ha this), -(field.div_mul_right Hb this), ↑divide,
-right_distrib]
theorem div_mul_div (Hb : b ≠ 0) (Hd : d ≠ 0) : (a / b) * (c / d) = (a * c) / (b * d) :=
theorem field.div_mul_div (a : A) {b : A} (c : A) {d : A} (Hb : b ≠ 0) (Hd : d ≠ 0) :
(a / b) * (c / d) = (a * c) / (b * d) :=
by rewrite [↑divide, 2 mul.assoc, (mul.comm b⁻¹), mul.assoc, (mul_inv_eq Hd Hb)]
theorem mul_div_mul_left (Hb : b ≠ 0) (Hc : c ≠ 0) : (c * a) / (c * b) = a / b :=
by rewrite [-(div_mul_div Hc Hb), (div_self Hc), one_mul]
theorem mul_div_mul_left (a : A) {b c : A} (Hb : b ≠ 0) (Hc : c ≠ 0) :
(c * a) / (c * b) = a / b :=
by rewrite [-(!field.div_mul_div Hc Hb), (div_self Hc), one_mul]
theorem mul_div_mul_right (Hb : b ≠ 0) (Hc : c ≠ 0) : (a * c) / (b * c) = a / b :=
by rewrite [(mul.comm a), (mul.comm b), (mul_div_mul_left Hb Hc)]
theorem mul_div_mul_right (a : A) {b c : A} (Hb : b ≠ 0) (Hc : c ≠ 0) :
(a * c) / (b * c) = a / b :=
by rewrite [(mul.comm a), (mul.comm b), (!mul_div_mul_left Hb Hc)]
theorem div_mul_eq_mul_div : (b / c) * a = (b * a) / c :=
theorem div_mul_eq_mul_div (a b c : A) : (b / c) * a = (b * a) / c :=
by rewrite [↑divide, mul.assoc, (mul.comm c⁻¹), -mul.assoc]
-- this one is odd -- I am not sure what to call it, but again, the prefix is right
theorem div_mul_eq_mul_div_comm (Hc : c ≠ 0) : (b / c) * a = b * (a / c) :=
by rewrite [(div_mul_eq_mul_div), -(one_mul c), -(div_mul_div (ne.symm zero_ne_one) Hc), div_one, one_mul]
theorem field.div_mul_eq_mul_div_comm (a b : A) {c : A} (Hc : c ≠ 0) :
(b / c) * a = b * (a / c) :=
by rewrite [(div_mul_eq_mul_div), -(one_mul c), -(!field.div_mul_div (ne.symm zero_ne_one) Hc),
div_one, one_mul]
theorem div_add_div (Hb : b ≠ 0) (Hd : d ≠ 0) :
theorem div_add_div (a : A) {b : A} (c : A) {d : A} (Hb : b ≠ 0) (Hd : d ≠ 0) :
(a / b) + (c / d) = ((a * d) + (b * c)) / (b * d) :=
by rewrite [-(mul_div_mul_right Hb Hd), -(mul_div_mul_left Hd Hb), div_add_div_same]
by rewrite [-(!mul_div_mul_right Hb Hd), -(!mul_div_mul_left Hd Hb), div_add_div_same]
theorem div_sub_div (Hb : b ≠ 0) (Hd : d ≠ 0) :
theorem div_sub_div (a : A) {b : A} (c : A) {d : A} (Hb : b ≠ 0) (Hd : d ≠ 0) :
(a / b) - (c / d) = ((a * d) - (b * c)) / (b * d) :=
by rewrite [↑sub, neg_eq_neg_one_mul, -mul_div_assoc, (div_add_div Hb Hd),
by rewrite [↑sub, neg_eq_neg_one_mul, -mul_div_assoc, (!div_add_div Hb Hd),
-mul.assoc, (mul.comm b), mul.assoc, -neg_eq_neg_one_mul]
theorem mul_eq_mul_of_div_eq_div (Hb : b ≠ 0) (Hd : d ≠ 0) (H : a / b = c / d) : a * d = c * b :=
theorem mul_eq_mul_of_div_eq_div (a : A) {b : A} (c : A) {d : A} (Hb : b ≠ 0)
(Hd : d ≠ 0) (H : a / b = c / d) : a * d = c * b :=
by rewrite [-mul_one, mul.assoc, (mul.comm d), -mul.assoc, -(div_self Hb),
-(div_mul_eq_mul_div_comm Hb), H, (div_mul_eq_mul_div), (div_mul_cancel Hd)]
-(!field.div_mul_eq_mul_div_comm Hb), H, (div_mul_eq_mul_div), (!div_mul_cancel Hd)]
theorem one_div_div (Ha : a ≠ 0) (Hb : b ≠ 0) : 1 / (a / b) = b / a :=
theorem field.one_div_div (Ha : a ≠ 0) (Hb : b ≠ 0) : 1 / (a / b) = b / a :=
have (a / b) * (b / a) = 1, from calc
(a / b) * (b / a) = (a * b) / (b * a) : div_mul_div Hb Ha
(a / b) * (b / a) = (a * b) / (b * a) : !field.div_mul_div Hb Ha
... = (a * b) / (a * b) : mul.comm
... = 1 : div_self (mul_ne_zero' Ha Hb),
... = 1 : div_self (division_ring.mul_ne_zero Ha Hb),
symm (eq_one_div_of_mul_eq_one this)
theorem div_div_eq_mul_div (Hb : b ≠ 0) (Hc : c ≠ 0) : a / (b / c) = (a * c) / b :=
by rewrite [div_eq_mul_one_div, (one_div_div Hb Hc), -mul_div_assoc]
theorem field.div_div_eq_mul_div (a : A) {b c : A} (Hb : b ≠ 0) (Hc : c ≠ 0) :
a / (b / c) = (a * c) / b :=
by rewrite [div_eq_mul_one_div, (field.one_div_div Hb Hc), -mul_div_assoc]
theorem div_div_eq_div_mul (Hb : b ≠ 0) (Hc : c ≠ 0) : (a / b) / c = a / (b * c) :=
by rewrite [div_eq_mul_one_div, (div_mul_div Hb Hc), mul_one]
theorem div_div_div_div (Hb : b ≠ 0) (Hc : c ≠ 0) (Hd : d ≠ 0) : (a / b) / (c / d) = (a * d) / (b * c) :=
by rewrite [(div_div_eq_mul_div Hc Hd), (div_mul_eq_mul_div), (div_div_eq_div_mul Hb Hc)]
theorem div_mul_eq_div_mul_one_div (Hb : b ≠ 0) (Hc : c ≠ 0) : a / (b * c) = (a / b) * (1 / c) :=
by rewrite [-div_div_eq_div_mul Hb Hc, -div_eq_mul_one_div]
theorem field.div_div_eq_div_mul (a : A) {b c : A} (Hb : b ≠ 0) (Hc : c ≠ 0) :
(a / b) / c = a / (b * c) :=
by rewrite [div_eq_mul_one_div, (!field.div_mul_div Hb Hc), mul_one]
theorem field.div_div_div_div_eq (a : A) {b c d : A} (Hb : b ≠ 0) (Hc : c ≠ 0) (Hd : d ≠ 0) :
(a / b) / (c / d) = (a * d) / (b * c) :=
by rewrite [(!field.div_div_eq_mul_div Hc Hd), (div_mul_eq_mul_div),
(!field.div_div_eq_div_mul Hb Hc)]
theorem field.div_mul_eq_div_mul_one_div (a : A) {b c : A} (Hb : b ≠ 0) (Hc : c ≠ 0) :
a / (b * c) = (a / b) * (1 / c) :=
by rewrite [-!field.div_div_eq_div_mul Hb Hc, -div_eq_mul_one_div]
end field
structure discrete_field [class] (A : Type) extends field A :=
@ -328,7 +321,6 @@ section discrete_field
-- many of the theorems in discrete_field are the same as theorems in field or division ring,
-- but with fewer hypotheses since 0⁻¹ = 0 and equality is decidable.
-- they are named with '. Is there a better convention?
theorem discrete_field.eq_zero_or_eq_zero_of_mul_eq_zero
(x y : A) (H : x * y = 0) : x = 0 y = 0 :=
@ -350,18 +342,18 @@ section discrete_field
... = 1 * 0 : discrete_field.inv_zero A
... = 0 : mul_zero
theorem div_zero : a / 0 = 0 := by rewrite [div_eq_mul_one_div, one_div_zero, mul_zero]
theorem div_zero (a : A) : a / 0 = 0 := by rewrite [div_eq_mul_one_div, one_div_zero, mul_zero]
theorem ne_zero_of_one_div_ne_zero (H : 1 / a ≠ 0) : a ≠ 0 :=
assume Ha : a = 0, absurd (Ha⁻¹ ▸ one_div_zero) H
theorem inv_zero_imp_zero (H : 1 / a = 0) : a = 0 :=
theorem eq_zero_of_one_div_eq_zero (H : 1 / a = 0) : a = 0 :=
decidable.by_cases
(assume Ha, Ha)
(assume Ha, false.elim ((one_div_ne_zero Ha) H))
-- the following could all go in "discrete_division_ring"
theorem one_div_mul_one_div'' : (1 / a) * (1 / b) = 1 / (b * a) :=
variables (a b)
theorem one_div_mul_one_div' : (1 / a) * (1 / b) = 1 / (b * a) :=
decidable.by_cases
(suppose a = 0,
by rewrite [this, div_zero, zero_mul, -(@div_zero A s 1), mul_zero b])
@ -369,32 +361,34 @@ section discrete_field
decidable.by_cases
(suppose b = 0,
by rewrite [this, div_zero, mul_zero, -(@div_zero A s 1), zero_mul a])
(suppose b ≠ 0, one_div_mul_one_div Ha this))
(suppose b ≠ 0, division_ring.one_div_mul_one_div Ha this))
theorem one_div_neg_eq_neg_one_div' : 1 / (- a) = - (1 / a) :=
theorem one_div_neg_eq_neg_one_div : 1 / (- a) = - (1 / a) :=
decidable.by_cases
(suppose a = 0, by rewrite [this, neg_zero, 2 div_zero, neg_zero])
(suppose a ≠ 0, one_div_neg_eq_neg_one_div this)
(suppose a ≠ 0, division_ring.one_div_neg_eq_neg_one_div this)
theorem neg_div_neg_eq_div' : (-a) / (-b) = a / b :=
theorem neg_div_neg_eq : (-a) / (-b) = a / b :=
decidable.by_cases
(assume Hb : b = 0, by rewrite [Hb, neg_zero, 2 div_zero])
(assume Hb : b ≠ 0, neg_div_neg_eq_div Hb)
(assume Hb : b ≠ 0, !division_ring.neg_div_neg_eq Hb)
theorem div_div' : 1 / (1 / a) = a :=
theorem one_div_one_div : 1 / (1 / a) = a :=
decidable.by_cases
(assume Ha : a = 0, by rewrite [Ha, 2 div_zero])
(assume Ha : a ≠ 0, div_div Ha)
(assume Ha : a ≠ 0, division_ring.one_div_one_div Ha)
theorem eq_of_invs_eq' (H : 1 / a = 1 / b) : a = b :=
variables {a b}
theorem eq_of_one_div_eq_one_div (H : 1 / a = 1 / b) : a = b :=
decidable.by_cases
(assume Ha : a = 0,
have Hb : b = 0, from inv_zero_imp_zero (by rewrite [-H, Ha, div_zero]),
have Hb : b = 0, from eq_zero_of_one_div_eq_zero (by rewrite [-H, Ha, div_zero]),
Hb⁻¹ ▸ Ha)
(assume Ha : a ≠ 0,
have Hb : b ≠ 0, from ne_zero_of_one_div_ne_zero (H ▸ (one_div_ne_zero Ha)),
eq_of_invs_eq Ha Hb H)
division_ring.eq_of_one_div_eq_one_div Ha Hb H)
variables (a b)
theorem mul_inv' : (b * a)⁻¹ = a⁻¹ * b⁻¹ :=
decidable.by_cases
(assume Ha : a = 0, by rewrite [Ha, mul_zero, 2 inv_zero, zero_mul])
@ -404,62 +398,68 @@ section discrete_field
(assume Hb : b ≠ 0, mul_inv_eq Ha Hb))
-- the following are specifically for fields
theorem one_div_mul_one_div''' : (1 / a) * (1 / b) = 1 / (a * b) :=
by rewrite [(one_div_mul_one_div''), mul.comm b]
theorem one_div_mul_one_div : (1 / a) * (1 / b) = 1 / (a * b) :=
by rewrite [one_div_mul_one_div', mul.comm b]
theorem div_mul_right' (Ha : a ≠ 0) : a / (a * b) = 1 / b :=
variable {a}
theorem div_mul_right (Ha : a ≠ 0) : a / (a * b) = 1 / b :=
decidable.by_cases
(assume Hb : b = 0, by rewrite [Hb, mul_zero, 2 div_zero])
(assume Hb : b ≠ 0, div_mul_right Hb (mul_ne_zero Ha Hb))
(assume Hb : b ≠ 0, field.div_mul_right Hb (mul_ne_zero Ha Hb))
theorem div_mul_left' (Hb : b ≠ 0) : b / (a * b) = 1 / a :=
by rewrite [mul.comm a, div_mul_right' Hb]
variables (a) {b}
theorem div_mul_left (Hb : b ≠ 0) : b / (a * b) = 1 / a :=
by rewrite [mul.comm a, div_mul_right _ Hb]
theorem div_mul_div' : (a / b) * (c / d) = (a * c) / (b * d) :=
variables (a b c)
theorem div_mul_div : (a / b) * (c / d) = (a * c) / (b * d) :=
decidable.by_cases
(assume Hb : b = 0, by rewrite [Hb, div_zero, zero_mul, -(@div_zero A s (a * c)), zero_mul])
(assume Hb : b ≠ 0,
decidable.by_cases
(assume Hd : d = 0, by rewrite [Hd, div_zero, mul_zero, -(@div_zero A s (a * c)), mul_zero])
(assume Hd : d ≠ 0, div_mul_div Hb Hd))
(assume Hd : d = 0, by rewrite [Hd, div_zero, mul_zero, -(@div_zero A s (a * c)),
mul_zero])
(assume Hd : d ≠ 0, !field.div_mul_div Hb Hd))
variable {c}
theorem mul_div_mul_left' (Hc : c ≠ 0) : (c * a) / (c * b) = a / b :=
decidable.by_cases
(assume Hb : b = 0, by rewrite [Hb, mul_zero, 2 div_zero])
(assume Hb : b ≠ 0, mul_div_mul_left Hb Hc)
(assume Hb : b ≠ 0, !mul_div_mul_left Hb Hc)
theorem mul_div_mul_right' (Hc : c ≠ 0) : (a * c) / (b * c) = a / b :=
by rewrite [(mul.comm a), (mul.comm b), (mul_div_mul_left' Hc)]
by rewrite [(mul.comm a), (mul.comm b), (!mul_div_mul_left' Hc)]
-- this one is odd -- I am not sure what to call it, but again, the prefix is right
theorem div_mul_eq_mul_div_comm' : (b / c) * a = b * (a / c) :=
variables (a b c d)
theorem div_mul_eq_mul_div_comm : (b / c) * a = b * (a / c) :=
decidable.by_cases
(assume Hc : c = 0, by rewrite [Hc, div_zero, zero_mul, -(mul_zero b), -(@div_zero A s a)])
(assume Hc : c ≠ 0, div_mul_eq_mul_div_comm Hc)
(assume Hc : c ≠ 0, !field.div_mul_eq_mul_div_comm Hc)
theorem one_div_div' : 1 / (a / b) = b / a :=
theorem one_div_div : 1 / (a / b) = b / a :=
decidable.by_cases
(assume Ha : a = 0, by rewrite [Ha, zero_div, 2 div_zero])
(assume Ha : a ≠ 0,
decidable.by_cases
(assume Hb : b = 0, by rewrite [Hb, 2 div_zero, zero_div])
(assume Hb : b ≠ 0, one_div_div Ha Hb))
(assume Hb : b ≠ 0, field.one_div_div Ha Hb))
theorem div_div_eq_mul_div' : a / (b / c) = (a * c) / b :=
by rewrite [div_eq_mul_one_div, one_div_div', -mul_div_assoc]
theorem div_div_eq_mul_div : a / (b / c) = (a * c) / b :=
by rewrite [div_eq_mul_one_div, one_div_div, -mul_div_assoc]
theorem div_div_eq_div_mul' : (a / b) / c = a / (b * c) :=
by rewrite [div_eq_mul_one_div, div_mul_div', mul_one]
theorem div_div_eq_div_mul : (a / b) / c = a / (b * c) :=
by rewrite [div_eq_mul_one_div, div_mul_div, mul_one]
theorem div_div_div_div' : (a / b) / (c / d) = (a * d) / (b * c) :=
by rewrite [div_div_eq_mul_div', div_mul_eq_mul_div, div_div_eq_div_mul']
theorem div_div_div_div_eq : (a / b) / (c / d) = (a * d) / (b * c) :=
by rewrite [div_div_eq_mul_div, div_mul_eq_mul_div, div_div_eq_div_mul]
variable {a}
theorem div_helper (H : a ≠ 0) : (1 / (a * b)) * a = 1 / b :=
by rewrite [div_mul_eq_mul_div, one_mul, (div_mul_right' H)]
theorem div_mul_eq_div_mul_one_div' : a / (b * c) = (a / b) * (1 / c) :=
by rewrite [-div_div_eq_div_mul', -div_eq_mul_one_div]
by rewrite [div_mul_eq_mul_div, one_mul, !div_mul_right H]
variable (a)
theorem div_mul_eq_div_mul_one_div : a / (b * c) = (a / b) * (1 / c) :=
by rewrite [-div_div_eq_div_mul, -div_eq_mul_one_div]
end discrete_field
end algebra

View file

@ -2,9 +2,7 @@
Copyright (c) 2014 Robert Lewis. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Robert Lewis
-/
import algebra.ordered_ring algebra.field
open eq eq.ops
@ -46,7 +44,7 @@ section linear_ordered_field
theorem lt_mul_of_gt_one_right (Hb : b > 0) (H : a > 1) : b < b * a :=
mul_one _ ▸ (mul_lt_mul_of_pos_left H Hb)
theorem one_le_div_iff_le (Hb : b > 0) : 1 ≤ a / b ↔ b ≤ a :=
theorem one_le_div_iff_le (a : A) {b : A} (Hb : b > 0) : 1 ≤ a / b ↔ b ≤ a :=
have Hb' : b ≠ 0, from ne.symm (ne_of_lt Hb),
iff.intro
(assume H : 1 ≤ a / b,
@ -61,12 +59,12 @@ section linear_ordered_field
... = a / b : div_eq_mul_one_div)
theorem le_of_one_le_div (Hb : b > 0) (H : 1 ≤ a / b) : b ≤ a :=
(iff.mp (one_le_div_iff_le Hb)) H
(iff.mp (!one_le_div_iff_le Hb)) H
theorem one_le_div_of_le (Hb : b > 0) (H : b ≤ a) : 1 ≤ a / b :=
(iff.mpr (one_le_div_iff_le Hb)) H
(iff.mpr (!one_le_div_iff_le Hb)) H
theorem one_lt_div_iff_lt (Hb : b > 0) : 1 < a / b ↔ b < a :=
theorem one_lt_div_iff_lt (a : A) {b : A} (Hb : b > 0) : 1 < a / b ↔ b < a :=
have Hb' : b ≠ 0, from ne.symm (ne_of_lt Hb),
iff.intro
(assume H : 1 < a / b,
@ -80,60 +78,57 @@ section linear_ordered_field
... = a / b : div_eq_mul_one_div)
theorem lt_of_one_lt_div (Hb : b > 0) (H : 1 < a / b) : b < a :=
(iff.mp (one_lt_div_iff_lt Hb)) H
(iff.mp (!one_lt_div_iff_lt Hb)) H
theorem one_lt_div_of_lt (Hb : b > 0) (H : b < a) : 1 < a / b :=
(iff.mpr (one_lt_div_iff_lt Hb)) H
(iff.mpr (!one_lt_div_iff_lt Hb)) H
theorem exists_lt : ∃ x, x < a :=
theorem exists_lt (a : A) : ∃ x, x < a :=
have H : a - 1 < a, from add_lt_of_le_of_neg (le.refl _) zero_gt_neg_one,
exists.intro _ H
theorem exists_gt : ∃ x, x > a :=
theorem exists_gt (a : A) : ∃ x, x > a :=
have H : a + 1 > a, from lt_add_of_le_of_pos (le.refl _) zero_lt_one,
exists.intro _ H
-- the following theorems amount to four iffs, for <, ≤, ≥, >.
theorem mul_le_of_le_div (Hc : 0 < c) (H : a ≤ b / c) : a * c ≤ b :=
div_mul_cancel (ne.symm (ne_of_lt Hc)) ▸ mul_le_mul_of_nonneg_right H (le_of_lt Hc)
!div_mul_cancel (ne.symm (ne_of_lt Hc)) ▸ mul_le_mul_of_nonneg_right H (le_of_lt Hc)
theorem le_div_of_mul_le (Hc : 0 < c) (H : a * c ≤ b) : a ≤ b / c :=
calc
a = a * c * (1 / c) : mul_mul_div (ne.symm (ne_of_lt Hc))
a = a * c * (1 / c) : !mul_mul_div (ne.symm (ne_of_lt Hc))
... ≤ b * (1 / c) : mul_le_mul_of_nonneg_right H (le_of_lt (one_div_pos_of_pos Hc))
... = b / c : div_eq_mul_one_div
theorem mul_lt_of_lt_div (Hc : 0 < c) (H : a < b / c) : a * c < b :=
div_mul_cancel (ne.symm (ne_of_lt Hc)) ▸ mul_lt_mul_of_pos_right H Hc
!div_mul_cancel (ne.symm (ne_of_lt Hc)) ▸ mul_lt_mul_of_pos_right H Hc
theorem lt_div_of_mul_lt (Hc : 0 < c) (H : a * c < b) : a < b / c :=
calc
a = a * c * (1 / c) : mul_mul_div (ne.symm (ne_of_lt Hc))
a = a * c * (1 / c) : !mul_mul_div (ne.symm (ne_of_lt Hc))
... < b * (1 / c) : mul_lt_mul_of_pos_right H (one_div_pos_of_pos Hc)
... = b / c : div_eq_mul_one_div
theorem mul_le_of_div_le_of_neg (Hc : c < 0) (H : b / c ≤ a) : a * c ≤ b :=
div_mul_cancel (ne_of_lt Hc) ▸ mul_le_mul_of_nonpos_right H (le_of_lt Hc)
!div_mul_cancel (ne_of_lt Hc) ▸ mul_le_mul_of_nonpos_right H (le_of_lt Hc)
theorem div_le_of_mul_le_of_neg (Hc : c < 0) (H : a * c ≤ b) : b / c ≤ a :=
calc
a = a * c * (1 / c) : mul_mul_div (ne_of_lt Hc)
a = a * c * (1 / c) : !mul_mul_div (ne_of_lt Hc)
... ≥ b * (1 / c) : mul_le_mul_of_nonpos_right H (le_of_lt (one_div_neg_of_neg Hc))
... = b / c : div_eq_mul_one_div
theorem mul_lt_of_gt_div_of_neg (Hc : c < 0) (H : a > b / c) : a * c < b :=
div_mul_cancel (ne_of_lt Hc) ▸ mul_lt_mul_of_neg_right H Hc
!div_mul_cancel (ne_of_lt Hc) ▸ mul_lt_mul_of_neg_right H Hc
theorem gt_div_of_mul_gt_of_neg (Hc : c < 0) (H : a * c < b) : a > b / c :=
theorem div_lt_of_mul_gt_of_neg (Hc : c < 0) (H : a * c < b) : b / c < a :=
calc
a = a * c * (1 / c) : mul_mul_div (ne_of_lt Hc)
a = a * c * (1 / c) : !mul_mul_div (ne_of_lt Hc)
... > b * (1 / c) : mul_lt_mul_of_neg_right H (one_div_neg_of_neg Hc)
... = b / c : div_eq_mul_one_div
-----
theorem div_le_of_le_mul (Hb : b > 0) (H : a ≤ b * c) : a / b ≤ c :=
calc
a / b = a * (1 / b) : div_eq_mul_one_div
@ -143,7 +138,7 @@ section linear_ordered_field
theorem le_mul_of_div_le (Hc : c > 0) (H : a / c ≤ b) : a ≤ b * c :=
calc
a = a / c * c : div_mul_cancel (ne.symm (ne_of_lt Hc))
a = a / c * c : !div_mul_cancel (ne.symm (ne_of_lt Hc))
... ≤ b * c : mul_le_mul_of_nonneg_right H (le_of_lt Hc)
-- following these in the isabelle file, there are 8 biconditionals for the above with - signs
@ -156,7 +151,7 @@ section linear_ordered_field
... = 0 : sub_self,
calc
0 > a / c - b / d : H1
... = (a * d - c * b) / (c * d) : div_sub_div Hc Hd
... = (a * d - c * b) / (c * d) : !div_sub_div Hc Hd
... = (a * d - b * c) / (c * d) : mul.comm
theorem mul_sub_mul_div_mul_nonpos (Hc : c ≠ 0) (Hd : d ≠ 0) (H : a / c ≤ b / d) :
@ -166,20 +161,20 @@ section linear_ordered_field
... = 0 : sub_self,
calc
0 ≥ a / c - b / d : H1
... = (a * d - c * b) / (c * d) : div_sub_div Hc Hd
... = (a * d - c * b) / (c * d) : !div_sub_div Hc Hd
... = (a * d - b * c) / (c * d) : mul.comm
theorem div_lt_div_of_mul_sub_mul_div_neg (Hc : c ≠ 0) (Hd : d ≠ 0)
(H : (a * d - b * c) / (c * d) < 0) : a / c < b / d :=
assert H1 : (a * d - c * b) / (c * d) < 0, by rewrite [mul.comm c b]; exact H,
assert H2 : a / c - b / d < 0, by rewrite [div_sub_div Hc Hd]; exact H1,
assert H2 : a / c - b / d < 0, by rewrite [!div_sub_div Hc Hd]; exact H1,
assert H3 : a / c - b / d + b / d < 0 + b / d, from add_lt_add_right H2 _,
begin rewrite [zero_add at H3, neg_add_cancel_right at H3], exact H3 end
theorem div_le_div_of_mul_sub_mul_div_nonpos (Hc : c ≠ 0) (Hd : d ≠ 0)
(H : (a * d - b * c) / (c * d) ≤ 0) : a / c ≤ b / d :=
assert H1 : (a * d - c * b) / (c * d) ≤ 0, by rewrite [mul.comm c b]; exact H,
assert H2 : a / c - b / d ≤ 0, by rewrite [div_sub_div Hc Hd]; exact H1,
assert H2 : a / c - b / d ≤ 0, by rewrite [!div_sub_div Hc Hd]; exact H1,
assert H3 : a / c - b / d + b / d ≤ 0 + b / d, from add_le_add_right H2 _,
begin rewrite [zero_add at H3, neg_add_cancel_right at H3], exact H3 end
@ -249,7 +244,6 @@ section linear_ordered_field
exact Hb
end
theorem div_nonneg_of_nonpos_of_neg (Ha : a ≤ 0) (Hb : b < 0) : 0 ≤ a / b :=
begin
rewrite div_eq_mul_one_div,
@ -266,7 +260,6 @@ section linear_ordered_field
exact mul_lt_mul_of_pos_right H (one_div_pos_of_pos Hc)
end
theorem div_le_div_of_le_of_pos (H : a ≤ b) (Hc : 0 < c) : a / c ≤ b / c :=
begin
rewrite [{a/c}div_eq_mul_one_div, {b/c}div_eq_mul_one_div],
@ -293,17 +286,17 @@ section linear_ordered_field
notation 2 := 1 + 1
theorem add_halves : a / 2 + a / 2 = a :=
theorem add_halves (a : A) : a / 2 + a / 2 = a :=
calc
a / 2 + a / 2 = (a + a) / 2 : by rewrite div_add_div_same
... = (a * 1 + a * 1) / 2 : by rewrite mul_one
... = (a * 2) / 2 : by rewrite left_distrib
... = a : by rewrite [@mul_div_cancel A _ _ _ two_ne_zero]
theorem sub_self_div_two : a - a / 2 = a / 2 :=
theorem sub_self_div_two (a : A) : a - a / 2 = a / 2 :=
by rewrite [-{a}add_halves at {1}, add_sub_cancel]
theorem div_two_sub_self : a / 2 - a = - (a / 2) :=
theorem div_two_sub_self (a : A) : a / 2 - a = - (a / 2) :=
by rewrite [-{a}add_halves at {2}, sub_add_eq_sub_sub, sub_self, zero_sub]
theorem nonneg_le_nonneg_of_squares_le (Ha : a ≥ 0) (Hb : b ≥ 0) (H : a * a ≤ b * b) : a ≤ b :=
@ -317,8 +310,8 @@ theorem nonneg_le_nonneg_of_squares_le (Ha : a ≥ 0) (Hb : b ≥ 0) (H : a * a
apply (not_le_of_gt H') H
end
theorem add_self_div_two : (a + a) / 2 = a :=
symm (iff.mpr (eq_div_iff_mul_eq (ne_of_gt (add_pos zero_lt_one zero_lt_one)))
theorem add_self_div_two (a : A) : (a + a) / 2 = a :=
symm (iff.mpr (!eq_div_iff_mul_eq (ne_of_gt (add_pos zero_lt_one zero_lt_one)))
(by rewrite [left_distrib, *mul_one]))
theorem two_ge_one : (2 : A) ≥ 1 :=
@ -340,7 +333,8 @@ theorem nonneg_le_nonneg_of_squares_le (Ha : a ≥ 0) (Hb : b ≥ 0) (H : a * a
theorem div_mul_le_div_mul_of_div_le_div_pos {e : A} (Hb : b ≠ 0) (Hd : d ≠ 0) (H : a / b ≤ c / d)
(He : e > 0) : a / (b * e) ≤ c / (d * e) :=
begin
rewrite [div_mul_eq_div_mul_one_div Hb (ne_of_gt He), div_mul_eq_div_mul_one_div Hd (ne_of_gt He)],
rewrite [!field.div_mul_eq_div_mul_one_div Hb (ne_of_gt He),
!field.div_mul_eq_div_mul_one_div Hd (ne_of_gt He)],
apply mul_le_mul_of_nonneg_right H,
apply le_of_lt,
apply one_div_pos_of_pos He
@ -355,7 +349,8 @@ theorem nonneg_le_nonneg_of_squares_le (Ha : a ≥ 0) (Hb : b ≥ 0) (H : a * a
... = (b + b) + (a - b) : add_sub,
assert H3 : (a + a) / (1 + 1) > ((b + b) + (a - b)) / (1 + 1),
from div_lt_div_of_lt_of_pos H2 two_pos,
by rewrite [add_self_div_two at H3, -div_add_div_same at H3, add_self_div_two at H3]; exact H3)
by rewrite [add_self_div_two at H3, -div_add_div_same at H3, add_self_div_two at H3];
exact H3)
(div_pos_of_pos_of_pos (iff.mpr !sub_pos_iff_lt H) two_pos))
end linear_ordered_field
@ -388,14 +383,14 @@ section discrete_linear_ordered_field
have H1 : 0 < 1 / (1 / a), from one_div_pos_of_pos H,
have H2 : 1 / a ≠ 0, from
(assume H3 : 1 / a = 0,
have H4 : 1 / (1 / a) = 0, from H3⁻¹ ▸ div_zero,
have H4 : 1 / (1 / a) = 0, from H3⁻¹ ▸ !div_zero,
absurd H4 (ne.symm (ne_of_lt H1))),
(div_div (ne_zero_of_one_div_ne_zero H2)) ▸ H1
(division_ring.one_div_one_div (ne_zero_of_one_div_ne_zero H2)) ▸ H1
theorem neg_of_one_div_neg (H : 1 / a < 0) : a < 0 :=
have H1 : 0 < - (1 / a), from neg_pos_of_neg H,
have Ha : a ≠ 0, from ne_zero_of_one_div_ne_zero (ne_of_lt H),
have H2 : 0 < 1 / (-a), from (one_div_neg_eq_neg_one_div Ha)⁻¹ ▸ H1,
have H2 : 0 < 1 / (-a), from (division_ring.one_div_neg_eq_neg_one_div Ha)⁻¹ ▸ H1,
have H3 : 0 < -a, from pos_of_one_div_pos H2,
neg_of_neg_pos H3
@ -417,9 +412,9 @@ section discrete_linear_ordered_field
have H' : -b > 0, from neg_pos_of_neg H,
have Hl' : - (1 / b) ≤ - (1 / a), from neg_le_neg Hl,
have Hl'' : 1 / - b ≤ 1 / - a, from calc
1 / -b = - (1 / b) : by rewrite [one_div_neg_eq_neg_one_div (ne_of_lt H)]
1 / -b = - (1 / b) : by rewrite [division_ring.one_div_neg_eq_neg_one_div (ne_of_lt H)]
... ≤ - (1 / a) : Hl'
... = 1 / -a : by rewrite [one_div_neg_eq_neg_one_div Ha],
... = 1 / -a : by rewrite [division_ring.one_div_neg_eq_neg_one_div Ha],
le_of_neg_le_neg (le_of_one_div_le_one_div H' Hl'')
theorem lt_of_one_div_lt_one_div (H : 0 < a) (Hl : 1 / a < 1 / b) : b < a :=
@ -487,19 +482,19 @@ section discrete_linear_ordered_field
theorem div_mul_le_div_mul_of_div_le_div_pos' {d e : A} (H : a / b ≤ c / d)
(He : e > 0) : a / (b * e) ≤ c / (d * e) :=
begin
rewrite [2 div_mul_eq_div_mul_one_div'],
rewrite [2 div_mul_eq_div_mul_one_div],
apply mul_le_mul_of_nonneg_right H,
apply le_of_lt,
apply one_div_pos_of_pos He
end
theorem abs_one_div : abs (1 / a) = 1 / abs a :=
theorem abs_one_div (a : A) : abs (1 / a) = 1 / abs a :=
if H : a > 0 then
by rewrite [abs_of_pos H, abs_of_pos (one_div_pos_of_pos H)]
else
(if H' : a < 0 then
by rewrite [abs_of_neg H', abs_of_neg (one_div_neg_of_neg H'),
-(one_div_neg_eq_neg_one_div (ne_of_lt H'))]
-(division_ring.one_div_neg_eq_neg_one_div (ne_of_lt H'))]
else
assert Heq : a = 0, from eq_of_le_of_ge (le_of_not_gt H) (le_of_not_gt H'),
by rewrite [Heq, div_zero, *abs_zero, div_zero])
@ -517,7 +512,7 @@ section discrete_linear_ordered_field
theorem sub_le_of_abs_sub_le_right (H : abs (a - b) ≤ c) : a - c ≤ b :=
sub_le_of_abs_sub_le_left (!abs_sub ▸ H)
theorem abs_sub_square : abs (a - b) * abs (a - b) = a * a + b * b - 2 * a * b :=
theorem abs_sub_square (a b : A) : abs (a - b) * abs (a - b) = a * a + b * b - 2 * a * b :=
by rewrite [abs_mul_abs_self, *mul_sub_left_distrib, *mul_sub_right_distrib,
sub_add_eq_sub_sub, sub_neg_eq_add, *right_distrib, sub_add_eq_sub_sub, *one_mul,
*add.assoc, {_ + b * b}add.comm, {_ + (b * b + _)}add.comm, mul.comm b a, *add.assoc]
@ -541,7 +536,7 @@ section discrete_linear_ordered_field
(suppose a = 0, by subst a; rewrite [zero_div, sign_zero])
(suppose a ≠ 0,
have abs a ≠ 0, from assume H, this (eq_zero_of_abs_eq_zero H),
eq_div_of_mul_eq this !eq_sign_mul_abs⁻¹)
!eq_div_of_mul_eq this !eq_sign_mul_abs⁻¹)
end discrete_linear_ordered_field
end algebra

View file

@ -52,8 +52,33 @@ or.elim (eq_zero_or_pos m)
obtain m' (h₂ : m = succ m'), from exists_eq_succ_of_pos `m > 0`,
show a = 0, by rewrite h₂ at H; apply h₁ m' H)
theorem pow_ne_zero_of_ne_zero {a : A} {m : } (H : a ≠ 0) : a^m ≠ 0 :=
assume H', H (eq_zero_of_pow_eq_zero H')
end integral_domain
section division_ring
variable [s : division_ring A]
include s
theorem division_ring.pow_ne_zero_of_ne_zero {a : A} {m : } (H : a ≠ 0) : a^m ≠ 0 :=
or.elim (eq_zero_or_pos m)
(suppose m = 0,
by rewrite [`m = 0`, pow_zero]; exact (ne.symm zero_ne_one))
(suppose m > 0,
have h₁ : ∀ m, a^succ m ≠ 0,
begin
intro m,
induction m with m ih,
{rewrite pow_one; assumption},
rewrite pow_succ,
apply division_ring.mul_ne_zero H ih
end,
obtain m' (h₂ : m = succ m'), from exists_eq_succ_of_pos `m > 0`,
show a^m ≠ 0, by rewrite h₂; apply h₁ m')
end division_ring
section linear_ordered_semiring
variable [s : linear_ordered_semiring A]
include s
@ -111,16 +136,29 @@ end
end decidable_linear_ordered_comm_ring
section field
variable [s : field A]
include s
theorem field.div_pow (a : A) {b : A} {n : } (bnz : b ≠ 0) : (a / b)^n = a^n / b^n :=
begin
induction n with n ih,
rewrite [*pow_zero, div_one],
have bnnz : b^n ≠ 0, from division_ring.pow_ne_zero_of_ne_zero bnz,
rewrite [*pow_succ, ih, !field.div_mul_div bnz bnnz]
end
end field
section discrete_field
variable [s : discrete_field A]
include s
theorem div_pow (a : A) {b : A} {n : } (bnz : b ≠ 0) : (a / b)^n = a^n / b^n :=
theorem div_pow (a : A) {b : A} {n : } : (a / b)^n = a^n / b^n :=
begin
induction n with n ih,
rewrite [*pow_zero, div_one],
have bnnz : b^n ≠ 0, from suppose b^n = 0, bnz (eq_zero_of_pow_eq_zero this),
rewrite [*pow_succ, ih, div_mul_div bnz bnnz]
rewrite [*pow_succ, ih, div_mul_div]
end
end discrete_field

View file

@ -185,7 +185,7 @@ theorem two_mul (p : +) : rat_of_pnat (2 * p) = (1 + 1) * rat_of_pnat p :=
theorem add_halves (p : +) : (2 * p)⁻¹ + (2 * p)⁻¹ = p⁻¹ :=
begin
rewrite [↑inv, -(@add_halves (1 / (rat_of_pnat p))), *div_div_eq_div_mul'],
rewrite [↑inv, -(@add_halves (1 / (rat_of_pnat p))), rat.div_div_eq_div_mul],
have H : rat_of_pnat (2 * p) = rat_of_pnat p * (1 + 1), by rewrite [rat.mul.comm, two_mul],
rewrite *H
end
@ -197,7 +197,7 @@ theorem add_halves_double (m n : +) :
by rewrite [-add_halves m, -add_halves n, hsimp]
theorem inv_mul_eq_mul_inv {p q : +} : (p * q)⁻¹ = p⁻¹ * q⁻¹ :=
by rewrite [↑inv, pnat_to_rat_mul, one_div_mul_one_div''']
by rewrite [↑inv, pnat_to_rat_mul, one_div_mul_one_div]
theorem inv_mul_le_inv (p q : +) : (p * q)⁻¹ ≤ q⁻¹ :=
begin
@ -277,10 +277,9 @@ theorem pceil_helper {a : } {n : +} (H : pceil a ≤ n) (Ha : a > 0) : n
rat.le.trans (inv_ge_of_le H) (div_le_div_of_le Ha (ubound_ge a))
theorem inv_pceil_div (a b : ) (Ha : a > 0) (Hb : b > 0) : (pceil (a / b))⁻¹ ≤ b / a :=
div_div' ▸ div_le_div_of_le
!one_div_one_div ▸ div_le_div_of_le
(one_div_pos_of_pos (div_pos_of_pos_of_pos Hb Ha))
((div_div_eq_mul_div (ne_of_gt Hb) (ne_of_gt Ha))⁻¹ ▸
!rat.one_mul⁻¹ ▸ !ubound_ge)
(!div_div_eq_mul_div⁻¹ ▸ !rat.one_mul⁻¹ ▸ !ubound_ge)
theorem sep_by_inv {a b : } (H : a > b) : ∃ N : +, a > (b + N⁻¹ + N⁻¹) :=
begin
@ -311,6 +310,6 @@ theorem nonneg_of_ge_neg_invs (a : ) (H : ∀ n : +, -n⁻¹ ≤ a) : 0
: !inv_pceil_div dec_trivial H2
... < -a / 1
: div_lt_div_of_pos_of_lt_of_pos dec_trivial dec_trivial H2
... = -a : div_one)))
... = -a : !div_one)))
end pnat

View file

@ -544,7 +544,7 @@ end migrate_algebra
theorem eq_num_div_denom (a : ) : a = num a / denom a :=
have H : of_int (denom a) ≠ 0, from assume H', ne_of_gt (denom_pos a) (of_int.inj H'),
iff.mpr (eq_div_iff_mul_eq H) (mul_denom a)
iff.mpr (!eq_div_iff_mul_eq H) (mul_denom a)
theorem of_int_div {a b : } (H : b a) : of_int (a div b) = of_int a / of_int b :=
decidable.by_cases
@ -556,7 +556,7 @@ decidable.by_cases
int.dvd.elim H
(take c, assume Hc : a = b * c,
by rewrite [Hc, !int.mul_div_cancel_left bnz, mul.comm]),
iff.mpr (eq_div_iff_mul_eq bnz') H')
iff.mpr (!eq_div_iff_mul_eq bnz') H')
theorem of_int_pow (a : ) (n : ) : of_int (a^n) = (of_int a)^n :=
begin

View file

@ -8,13 +8,12 @@ This construction follows Bishop and Bridges (1985).
To do:
o Rename things and possibly make theorems private
-/
import data.nat data.rat.order data.pnat
open nat eq eq.ops pnat
open -[coercions] rat
local notation 0 := rat.of_num 0
local notation 1 := rat.of_num 1
----------------------------------------------------------------------------------------------------
-- small helper lemmas
theorem s_mul_assoc_lemma_3 (a b n : +) (p : ) :
@ -39,7 +38,7 @@ theorem find_thirds (a b : ) (H : b > 0) : ∃ n : +, a + n⁻¹ + n⁻¹
... ≤ of_nat 4 * (b / of_nat 4)
: rat.mul_le_mul_of_nonneg_left (!inv_pceil_div dec_trivial H) !of_nat_nonneg
... = b / of_nat 4 * of_nat 4 : rat.mul.comm
... = b : rat.div_mul_cancel dec_trivial,
... = b : !rat.div_mul_cancel dec_trivial,
exists.intro n (calc
a + n⁻¹ + n⁻¹ + n⁻¹ = a + (1 + 1 + 1) * n⁻¹ : by rewrite[*rat.right_distrib,*rat.one_mul,-*rat.add.assoc]
... = a + of_nat 3 * n⁻¹ : {show 1+1+1=of_nat 3, from dec_trivial}
@ -203,7 +202,7 @@ theorem eq_of_bdd_var {s t : seq} (Hs : regular s) (Ht : regular t)
theorem pnat_bound {ε : } (Hε : ε > 0) : ∃ p : +, p⁻¹ ≤ ε :=
begin
existsi (pceil (1 / ε)),
rewrite -(rat.div_div (rat.ne_of_gt Hε)) at {2},
rewrite -(rat.one_div_one_div ε) at {2},
apply pceil_helper,
apply le.refl,
apply one_div_pos_of_pos Hε
@ -622,7 +621,7 @@ theorem mul_bound_helper {s t : seq} (Hs : regular s) (Ht : regular t) (a b c :
apply rat.ne_of_gt,
repeat (apply rat.mul_pos | apply rat.add_pos | apply rat_of_pnat_is_pos | apply inv_pos),
end,
rewrite (rat.div_helper H),
rewrite (!rat.div_helper H),
apply rat.le.refl
end,
apply rat.add_le_add,

View file

@ -738,8 +738,7 @@ theorem width (n : ) : over_seq n - under_seq n = (over - under) / (rat.pow 2
let Hou := calc
(over_seq a) / 2 - (under_seq a) / 2 = ((over - under) / rat.pow 2 a) / 2 :
by rewrite [rat.div_sub_div_same, Ha]
... = (over - under) / (rat.pow 2 a * 2) :
rat.div_div_eq_div_mul (rat.ne_of_gt (rat.pow_pos dec_trivial _)) dec_trivial
... = (over - under) / (rat.pow 2 a * 2) : rat.div_div_eq_div_mul
... = (over - under) / rat.pow 2 (a + 1) : by rewrite rat.pow_add,
cases em (ub (avg_seq a)),
rewrite [*if_pos a_1, -add_one, -Hou, ↑avg_seq, ↑avg, rat.add.assoc, rat.div_two_sub_self],

View file

@ -8,7 +8,6 @@ This construction follows Bishop and Bridges (1985).
At this point, we no longer proceed constructively: this file makes heavy use of decidability
and excluded middle.
-/
import data.real.basic data.real.order data.rat data.nat
open -[coercions] rat
open -[coercions] nat
@ -196,11 +195,11 @@ theorem reg_inv_reg {s : seq} (Hs : regular s) (Hsep : sep s zero) : regular (s_
apply add_invs_nonneg,
rewrite [(s_inv_of_sep_lt_p Hs Hsep Hmlt),
(s_inv_of_sep_gt_p Hs Hsep (le_of_not_gt Hnlt))],
rewrite [(div_sub_div Hsp Hspn), div_eq_mul_one_div, *abs_mul, *mul_one, *one_mul],
rewrite [(!div_sub_div Hsp Hspn), div_eq_mul_one_div, *abs_mul, *mul_one, *one_mul],
apply rat.le.trans,
apply rat.mul_le_mul,
apply Hs,
rewrite [-(mul_one 1), -(div_mul_div Hsp Hspn), abs_mul],
rewrite [-(mul_one 1), -(!field.div_mul_div Hsp Hspn), abs_mul],
apply rat.mul_le_mul,
rewrite -(s_inv_of_sep_lt_p Hs Hsep Hmlt),
apply le_ps Hs Hsep,
@ -218,11 +217,11 @@ theorem reg_inv_reg {s : seq} (Hs : regular s) (Hsep : sep s zero) : regular (s_
cases em (n < ps Hs Hsep) with [Hnlt, Hnlt],
rewrite [(s_inv_of_sep_lt_p Hs Hsep Hnlt),
(s_inv_of_sep_gt_p Hs Hsep (le_of_not_gt Hmlt))],
rewrite [(div_sub_div Hspm Hsp), div_eq_mul_one_div, *abs_mul, *mul_one, *one_mul],
rewrite [(!div_sub_div Hspm Hsp), div_eq_mul_one_div, *abs_mul, *mul_one, *one_mul],
apply rat.le.trans,
apply rat.mul_le_mul,
apply Hs,
rewrite [-(mul_one 1), -(div_mul_div Hspm Hsp), abs_mul],
rewrite [-(mul_one 1), -(!field.div_mul_div Hspm Hsp), abs_mul],
apply rat.mul_le_mul,
rewrite -(s_inv_of_sep_gt_p Hs Hsep (le_of_not_gt Hmlt)),
apply le_ps Hs Hsep,
@ -239,11 +238,11 @@ theorem reg_inv_reg {s : seq} (Hs : regular s) (Hsep : sep s zero) : regular (s_
apply Hnlt,
rewrite [(s_inv_of_sep_gt_p Hs Hsep (le_of_not_gt Hnlt)),
(s_inv_of_sep_gt_p Hs Hsep (le_of_not_gt Hmlt))],
rewrite [(div_sub_div Hspm Hspn), div_eq_mul_one_div, abs_mul, *one_mul, *mul_one],
rewrite [(!div_sub_div Hspm Hspn), div_eq_mul_one_div, abs_mul, *one_mul, *mul_one],
apply rat.le.trans,
apply rat.mul_le_mul,
apply Hs,
rewrite [-(mul_one 1), -(div_mul_div Hspm Hspn), abs_mul],
rewrite [-(mul_one 1), -(!field.div_mul_div Hspm Hspn), abs_mul],
apply rat.mul_le_mul,
rewrite -(s_inv_of_sep_gt_p Hs Hsep (le_of_not_gt Hmlt)),
apply le_ps Hs Hsep,
@ -303,7 +302,7 @@ theorem mul_inv {s : seq} (Hs : regular s) (Hsep : sep s zero) : smul s (s_inv H
s_ne_zero_of_ge_p Hs Hsep
(show ps Hs Hsep ≤ ((ps Hs Hsep) * (ps Hs Hsep)) * ((K₂ s (s_inv Hs)) * 2 * n),
by rewrite *pnat.mul.assoc; apply pnat.mul_le_mul_right),
rewrite [(s_inv_of_sep_gt_p Hs Hsep Hp), (div_div Hnz')],
rewrite [(s_inv_of_sep_gt_p Hs Hsep Hp), (division_ring.one_div_one_div Hnz')],
apply rat.le.trans,
apply rat.mul_le_mul_of_nonneg_left,
apply Hs,

View file

@ -8,7 +8,6 @@ This construction follows Bishop and Bridges (1985).
To do:
o Rename things and possibly make theorems private
-/
import data.real.basic data.rat data.nat
open -[coercions] rat
open -[coercions] nat
@ -881,9 +880,9 @@ theorem nat_inv_lt_rat {a : } (H : a > 0) : ∃ n : +, n⁻¹ < a :=
apply lt_of_le_of_lt,
rotate 1,
apply div_two_lt_of_pos H,
rewrite -(@div_div' (a / (1 + 1))),
rewrite -(one_div_one_div (a / (1 + 1))),
apply pceil_helper,
rewrite div_div',
rewrite one_div_one_div,
apply pnat.le.refl,
apply one_div_pos_of_pos,
apply div_pos_of_pos_of_pos H dec_trivial

View file

@ -73,8 +73,8 @@ section
have b ≠ 0, from ne_of_gt (denom_pos q),
have bnz : b ≠ (0 : ), from assume H, `b ≠ 0` (of_int.inj H),
have bnnz : (#rat b^n ≠ 0), from assume bneqz, bnz (eq_zero_of_pow_eq_zero bneqz),
have a^n / b^n = c, using bnz, by rewrite [*of_int_pow, -(!div_pow bnz), -eq_num_div_denom, -H],
have a^n = c * b^n, from eq.symm (mul_eq_of_eq_div bnnz this⁻¹),
have a^n / b^n = c, using bnz, by rewrite [*of_int_pow, -div_pow, -eq_num_div_denom, -H],
have a^n = c * b^n, from eq.symm (!mul_eq_of_eq_div bnnz this⁻¹),
have a^n = c * b^n, -- int version
using this, by rewrite [-of_int_pow at this, -of_int_mul at this]; exact of_int.inj this,
have (abs a)^n = abs c * (abs b)^n,