refactor(library/real/*): protect theorems
This commit is contained in:
parent
9e5a27dc77
commit
b1719c3823
4 changed files with 152 additions and 136 deletions
|
@ -33,7 +33,8 @@ private theorem s_mul_assoc_lemma_3 (a b n : ℕ+) (p : ℚ) :
|
|||
p * ((a * n)⁻¹ + (b * n)⁻¹) = p * (a⁻¹ + b⁻¹) * n⁻¹ :=
|
||||
by rewrite [rat.mul_assoc, right_distrib, *pnat.inv_mul_eq_mul_inv]
|
||||
|
||||
private theorem s_mul_assoc_lemma_4 {n : ℕ+} {ε q : ℚ} (Hε : ε > 0) (Hq : q > 0) (H : n ≥ pceil (q / ε)) :
|
||||
private theorem s_mul_assoc_lemma_4 {n : ℕ+} {ε q : ℚ} (Hε : ε > 0) (Hq : q > 0)
|
||||
(H : n ≥ pceil (q / ε)) :
|
||||
q * n⁻¹ ≤ ε :=
|
||||
begin
|
||||
let H2 := pceil_helper H (div_pos_of_pos_of_pos Hq Hε),
|
||||
|
@ -73,7 +74,7 @@ by rewrite [algebra.mul_sub_left_distrib, algebra.mul_sub_right_distrib, add_sub
|
|||
|
||||
private theorem rewrite_helper3 (a b c d e f g: ℚ) : a * (b + c) - (d * e + f * g) =
|
||||
(a * b - d * e) + (a * c - f * g) :=
|
||||
by rewrite [rat.left_distrib, add_sub_comm]
|
||||
by rewrite [left_distrib, add_sub_comm]
|
||||
|
||||
private theorem rewrite_helper4 (a b c d : ℚ) : a * b - c * d = (a * b - a * d) + (a * d - c * d) :=
|
||||
by rewrite[add_sub, algebra.sub_add_cancel]
|
||||
|
@ -127,7 +128,7 @@ calc (rat_of_pnat k) * a + b * (rat_of_pnat k)
|
|||
... ≤ (rat_of_pnat k) * (k⁻¹ * (m⁻¹ + n⁻¹)) :
|
||||
iff.mp (!algebra.le_iff_mul_le_mul_left !rat_of_pnat_is_pos) this
|
||||
... = m⁻¹ + n⁻¹ :
|
||||
by rewrite[-rat.mul_assoc, pnat.inv_cancel_left, rat.one_mul]
|
||||
by rewrite[-mul.assoc, pnat.inv_cancel_left, one_mul]
|
||||
|
||||
private theorem factor_lemma (a b c d e : ℚ) : abs (a + b + c - (d + (b + e))) = abs ((a - d) + (c - e)) :=
|
||||
!congr_arg (calc
|
||||
|
@ -187,14 +188,14 @@ theorem eq_of_bdd {s t : seq} (Hs : regular s) (Ht : regular t)
|
|||
have Hj : (∀ j : ℕ+, abs (s n - t n) ≤ n⁻¹ + n⁻¹ + j⁻¹ + j⁻¹ + j⁻¹), begin
|
||||
intros,
|
||||
cases H j with [Nj, HNj],
|
||||
rewrite [-(sub_add_cancel (s n) (s (max j Nj))), +sub_eq_add_neg, algebra.add.assoc (s n + -s (max j Nj)),
|
||||
↑regular at *],
|
||||
rewrite [-(sub_add_cancel (s n) (s (max j Nj))), +sub_eq_add_neg,
|
||||
add.assoc (s n + -s (max j Nj)), ↑regular at *],
|
||||
apply rat.le_trans,
|
||||
apply abs_add_le_abs_add_abs,
|
||||
apply rat.le_trans,
|
||||
apply add_le_add,
|
||||
apply Hs,
|
||||
rewrite [-(sub_add_cancel (s (max j Nj)) (t (max j Nj))), rat.add_assoc],
|
||||
rewrite [-(sub_add_cancel (s (max j Nj)) (t (max j Nj))), add.assoc],
|
||||
apply abs_add_le_abs_add_abs,
|
||||
apply rat.le_trans,
|
||||
apply rat.add_le_add_left,
|
||||
|
@ -203,10 +204,11 @@ theorem eq_of_bdd {s t : seq} (Hs : regular s) (Ht : regular t)
|
|||
apply Ht,
|
||||
have hsimp : ∀ m : ℕ+, n⁻¹ + m⁻¹ + (j⁻¹ + (m⁻¹ + n⁻¹)) = n⁻¹ + n⁻¹ + j⁻¹ + (m⁻¹ + m⁻¹),
|
||||
from λm, calc
|
||||
n⁻¹ + m⁻¹ + (j⁻¹ + (m⁻¹ + n⁻¹)) = n⁻¹ + (j⁻¹ + (m⁻¹ + n⁻¹)) + m⁻¹ : algebra.add.right_comm
|
||||
... = n⁻¹ + (j⁻¹ + m⁻¹ + n⁻¹) + m⁻¹ : rat.add_assoc
|
||||
... = n⁻¹ + (n⁻¹ + (j⁻¹ + m⁻¹)) + m⁻¹ : rat.add_comm
|
||||
... = n⁻¹ + n⁻¹ + j⁻¹ + (m⁻¹ + m⁻¹) : by rewrite[-*rat.add_assoc],
|
||||
n⁻¹ + m⁻¹ + (j⁻¹ + (m⁻¹ + n⁻¹)) = n⁻¹ + (j⁻¹ + (m⁻¹ + n⁻¹)) + m⁻¹ : add.right_comm
|
||||
... = n⁻¹ + (j⁻¹ + m⁻¹ + n⁻¹) + m⁻¹ : add.assoc
|
||||
... = n⁻¹ + (n⁻¹ + (j⁻¹ + m⁻¹)) + m⁻¹ : add.comm
|
||||
... = n⁻¹ + n⁻¹ + j⁻¹ + (m⁻¹ + m⁻¹) :
|
||||
by rewrite[-*add.assoc],
|
||||
rewrite hsimp,
|
||||
have Hms : (max j Nj)⁻¹ + (max j Nj)⁻¹ ≤ j⁻¹ + j⁻¹, begin
|
||||
apply add_le_add,
|
||||
|
@ -392,7 +394,7 @@ theorem s_add_comm (s t : seq) : sadd s t ≡ sadd t s :=
|
|||
end
|
||||
|
||||
theorem s_add_assoc (s t u : seq) (Hs : regular s) (Hu : regular u) :
|
||||
sadd (sadd s t) u ≡ sadd s (sadd t u) :=
|
||||
sadd (sadd s t) u ≡ sadd s (sadd t u) :=
|
||||
begin
|
||||
rewrite [↑sadd, ↑equiv, ↑regular at *],
|
||||
intros,
|
||||
|
@ -426,7 +428,7 @@ private theorem TK_rewrite (s t u : seq) :
|
|||
(DK (λ (n : ℕ+), s (mul (DK s t) n) * t (mul (DK s t) n)) u) = TK s t u := rfl
|
||||
|
||||
private theorem s_mul_assoc_lemma (s t u : seq) (a b c d : ℕ+) :
|
||||
abs (s a * t a * u b - s c * t d * u d) ≤ abs (t a) * abs (u b) * abs (s a - s c) +
|
||||
abs (s a * t a * u b - s c * t d * u d) ≤ abs (t a) * abs (u b) * abs (s a - s c) +
|
||||
abs (s c) * abs (t a) * abs (u b - u d) + abs (s c) * abs (u d) * abs (t a - t d) :=
|
||||
begin
|
||||
rewrite (rewrite_helper7 _ _ _ _ (s c)),
|
||||
|
@ -435,11 +437,11 @@ private theorem s_mul_assoc_lemma (s t u : seq) (a b c d : ℕ+) :
|
|||
rewrite rat.add_assoc,
|
||||
apply add_le_add,
|
||||
rewrite 2 abs_mul,
|
||||
apply rat.le_refl,
|
||||
apply le.refl,
|
||||
rewrite [*rat.mul_assoc, -algebra.mul_sub_left_distrib, -left_distrib, abs_mul],
|
||||
apply mul_le_mul_of_nonneg_left,
|
||||
rewrite rewrite_helper,
|
||||
apply rat.le_trans,
|
||||
apply le.trans,
|
||||
apply abs_add_le_abs_add_abs,
|
||||
apply add_le_add,
|
||||
rewrite abs_mul, apply rat.le_refl,
|
||||
|
@ -451,15 +453,15 @@ private definition Kq (s : seq) := rat_of_pnat (K s) + 1
|
|||
private theorem Kq_bound {s : seq} (H : regular s) : ∀ n, abs (s n) ≤ Kq s :=
|
||||
begin
|
||||
intros,
|
||||
apply rat.le_of_lt,
|
||||
apply rat.lt_of_le_of_lt,
|
||||
apply le_of_lt,
|
||||
apply lt_of_le_of_lt,
|
||||
apply canon_bound H,
|
||||
apply algebra.lt_add_of_pos_right,
|
||||
apply rat.zero_lt_one
|
||||
apply lt_add_of_pos_right,
|
||||
apply zero_lt_one
|
||||
end
|
||||
|
||||
private theorem Kq_bound_nonneg {s : seq} (H : regular s) : 0 ≤ Kq s :=
|
||||
rat.le_trans !abs_nonneg (Kq_bound H 2)
|
||||
le.trans !abs_nonneg (Kq_bound H 2)
|
||||
|
||||
private theorem Kq_bound_pos {s : seq} (H : regular s) : 0 < Kq s :=
|
||||
have H1 : 0 ≤ rat_of_pnat (K s), from rat.le_trans (!abs_nonneg) (canon_bound H 2),
|
||||
|
@ -480,8 +482,8 @@ private theorem s_mul_assoc_lemma_5 {s t u : seq} (Hs : regular s) (Ht : regular
|
|||
apply Kq_bound_nonneg Hu,
|
||||
end
|
||||
|
||||
private theorem s_mul_assoc_lemma_2 {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
||||
(a b c d : ℕ+) :
|
||||
private theorem s_mul_assoc_lemma_2 {s t u : seq} (Hs : regular s) (Ht : regular t)
|
||||
(Hu : regular u) (a b c d : ℕ+) :
|
||||
abs (t a) * abs (u b) * abs (s a - s c) + abs (s c) * abs (t a) * abs (u b - u d)
|
||||
+ abs (s c) * abs (u d) * abs (t a - t d) ≤
|
||||
(Kq t) * (Kq u) * (a⁻¹ + c⁻¹) + (Kq s) * (Kq t) * (b⁻¹ + d⁻¹) + (Kq s) * (Kq u) * (a⁻¹ + d⁻¹) :=
|
||||
|
@ -612,8 +614,9 @@ theorem add_well_defined {s t u v : seq} (Hs : regular s) (Ht : regular t) (Hu :
|
|||
end
|
||||
|
||||
set_option tactic.goal_names false
|
||||
private theorem mul_bound_helper {s t : seq} (Hs : regular s) (Ht : regular t) (a b c : ℕ+) (j : ℕ+) :
|
||||
∃ N : ℕ+, ∀ n : ℕ+, n ≥ N → abs (s (a * n) * t (b * n) - s (c * n) * t (c * n)) ≤ j⁻¹ :=
|
||||
private theorem mul_bound_helper {s t : seq} (Hs : regular s) (Ht : regular t) (a b c : ℕ+)
|
||||
(j : ℕ+) :
|
||||
∃ N : ℕ+, ∀ n : ℕ+, n ≥ N → abs (s (a * n) * t (b * n) - s (c * n) * t (c * n)) ≤ j⁻¹ :=
|
||||
begin
|
||||
existsi pceil (((rat_of_pnat (K s)) * (b⁻¹ + c⁻¹) + (a⁻¹ + c⁻¹) *
|
||||
(rat_of_pnat (K t))) * (rat_of_pnat j)),
|
||||
|
@ -645,7 +648,7 @@ private theorem mul_bound_helper {s t : seq} (Hs : regular s) (Ht : regular t) (
|
|||
apply rat_of_pnat_is_pos,
|
||||
have H : (rat_of_pnat (K s) * (b⁻¹ + c⁻¹) + (a⁻¹ + c⁻¹) * rat_of_pnat (K t)) ≠ 0, begin
|
||||
apply ne_of_gt,
|
||||
repeat (apply algebra.mul_pos | apply algebra.add_pos | apply rat_of_pnat_is_pos | apply pnat.inv_pos),
|
||||
repeat (apply mul_pos | apply add_pos | apply rat_of_pnat_is_pos | apply pnat.inv_pos),
|
||||
end,
|
||||
rewrite (!div_helper H),
|
||||
apply rat.le_refl
|
||||
|
@ -730,8 +733,8 @@ theorem mul_zero_equiv_zero {s t : seq} (Hs : regular s) (Ht : regular t) (Htz :
|
|||
existsi N,
|
||||
intro n Hn,
|
||||
rewrite [↑equiv at Htz, ↑zero at *, algebra.sub_zero, ↑smul, abs_mul],
|
||||
apply rat.le_trans,
|
||||
apply algebra.mul_le_mul,
|
||||
apply le.trans,
|
||||
apply mul_le_mul,
|
||||
apply Kq_bound Hs,
|
||||
have HN' : ∀ (n : ℕ+), N ≤ n → abs (t n) ≤ ε / Kq s,
|
||||
from λ n, (eq.subst (sub_zero (t n)) (HN n)),
|
||||
|
@ -757,7 +760,7 @@ theorem mul_neg_equiv_neg_mul {s t : seq} : smul s (sneg t) ≡ sneg (smul s t)
|
|||
rewrite [↑equiv, ↑smul],
|
||||
intros,
|
||||
rewrite [↑sneg, *sub_neg_eq_add, -neg_mul_eq_mul_neg, add.comm, *sneg_def,
|
||||
*neg_bound2_eq_bound2, algebra.add.right_inv, abs_zero],
|
||||
*neg_bound2_eq_bound2, add.right_inv, abs_zero],
|
||||
apply add_invs_nonneg
|
||||
end
|
||||
|
||||
|
@ -766,8 +769,8 @@ theorem equiv_of_diff_equiv_zero {s t : seq} (Hs : regular s) (Ht : regular t)
|
|||
begin
|
||||
have hsimp : ∀ a b c d e : ℚ, a + b + c + (d + e) = b + d + a + e + c, from
|
||||
λ a b c d e, calc
|
||||
a + b + c + (d + e) = a + b + (d + e) + c : algebra.add.right_comm
|
||||
... = a + (b + d) + e + c : by rewrite[-*rat.add_assoc]
|
||||
a + b + c + (d + e) = a + b + (d + e) + c : add.right_comm
|
||||
... = a + (b + d) + e + c : by rewrite[-*add.assoc]
|
||||
... = b + d + a + e + c : add.comm,
|
||||
apply eq_of_bdd Hs Ht,
|
||||
intros,
|
||||
|
@ -812,8 +815,8 @@ theorem diff_equiv_zero_of_equiv {s t : seq} (Hs : regular s) (Ht : regular t) (
|
|||
repeat (assumption | apply reg_add_reg | apply reg_neg_reg)
|
||||
end
|
||||
|
||||
private theorem mul_well_defined_half1 {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
||||
(Etu : t ≡ u) : smul s t ≡ smul s u :=
|
||||
private theorem mul_well_defined_half1 {s t u : seq} (Hs : regular s) (Ht : regular t)
|
||||
(Hu : regular u) (Etu : t ≡ u) : smul s t ≡ smul s u :=
|
||||
begin
|
||||
apply equiv_of_diff_equiv_zero,
|
||||
rotate 2,
|
||||
|
@ -836,8 +839,8 @@ private theorem mul_well_defined_half1 {s t u : seq} (Hs : regular s) (Ht : regu
|
|||
apply zero_is_reg)
|
||||
end
|
||||
|
||||
private theorem mul_well_defined_half2 {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
||||
(Est : s ≡ t) : smul s u ≡ smul t u :=
|
||||
private theorem mul_well_defined_half2 {s t u : seq} (Hs : regular s) (Ht : regular t)
|
||||
(Hu : regular u) (Est : s ≡ t) : smul s u ≡ smul t u :=
|
||||
begin
|
||||
apply equiv.trans,
|
||||
rotate 3,
|
||||
|
@ -853,7 +856,7 @@ private theorem mul_well_defined_half2 {s t u : seq} (Hs : regular s) (Ht : regu
|
|||
end
|
||||
|
||||
theorem mul_well_defined {s t u v : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
||||
(Hv : regular v) (Esu : s ≡ u) (Etv : t ≡ v) : smul s t ≡ smul u v :=
|
||||
(Hv : regular v) (Esu : s ≡ u) (Etv : t ≡ v) : smul s t ≡ smul u v :=
|
||||
begin
|
||||
apply equiv.trans,
|
||||
exact reg_mul_reg Hs Ht,
|
||||
|
@ -1106,40 +1109,40 @@ rfl
|
|||
theorem real_one_eq_rat_one : (1:real) = of_rat (1:rat) :=
|
||||
rfl
|
||||
|
||||
theorem add_comm (x y : ℝ) : x + y = y + x :=
|
||||
protected theorem add_comm (x y : ℝ) : x + y = y + x :=
|
||||
quot.induction_on₂ x y (λ s t, quot.sound (r_add_comm s t))
|
||||
|
||||
theorem add_assoc (x y z : ℝ) : x + y + z = x + (y + z) :=
|
||||
protected theorem add_assoc (x y z : ℝ) : x + y + z = x + (y + z) :=
|
||||
quot.induction_on₃ x y z (λ s t u, quot.sound (r_add_assoc s t u))
|
||||
|
||||
theorem zero_add (x : ℝ) : 0 + x = x :=
|
||||
protected theorem zero_add (x : ℝ) : 0 + x = x :=
|
||||
quot.induction_on x (λ s, quot.sound (r_zero_add s))
|
||||
|
||||
theorem add_zero (x : ℝ) : x + 0 = x :=
|
||||
protected theorem add_zero (x : ℝ) : x + 0 = x :=
|
||||
quot.induction_on x (λ s, quot.sound (r_add_zero s))
|
||||
|
||||
theorem neg_cancel (x : ℝ) : -x + x = 0 :=
|
||||
protected theorem neg_cancel (x : ℝ) : -x + x = 0 :=
|
||||
quot.induction_on x (λ s, quot.sound (r_neg_cancel s))
|
||||
|
||||
theorem mul_assoc (x y z : ℝ) : x * y * z = x * (y * z) :=
|
||||
protected theorem mul_assoc (x y z : ℝ) : x * y * z = x * (y * z) :=
|
||||
quot.induction_on₃ x y z (λ s t u, quot.sound (r_mul_assoc s t u))
|
||||
|
||||
theorem mul_comm (x y : ℝ) : x * y = y * x :=
|
||||
protected theorem mul_comm (x y : ℝ) : x * y = y * x :=
|
||||
quot.induction_on₂ x y (λ s t, quot.sound (r_mul_comm s t))
|
||||
|
||||
theorem one_mul (x : ℝ) : 1 * x = x :=
|
||||
protected theorem one_mul (x : ℝ) : 1 * x = x :=
|
||||
quot.induction_on x (λ s, quot.sound (r_one_mul s))
|
||||
|
||||
theorem mul_one (x : ℝ) : x * 1 = x :=
|
||||
protected theorem mul_one (x : ℝ) : x * 1 = x :=
|
||||
quot.induction_on x (λ s, quot.sound (r_mul_one s))
|
||||
|
||||
theorem left_distrib (x y z : ℝ) : x * (y + z) = x * y + x * z :=
|
||||
protected theorem left_distrib (x y z : ℝ) : x * (y + z) = x * y + x * z :=
|
||||
quot.induction_on₃ x y z (λ s t u, quot.sound (r_distrib s t u))
|
||||
|
||||
theorem right_distrib (x y z : ℝ) : (x + y) * z = x * z + y * z :=
|
||||
by rewrite [mul_comm, left_distrib, {x * _}mul_comm, {y * _}mul_comm] -- this shouldn't be necessary
|
||||
protected theorem right_distrib (x y z : ℝ) : (x + y) * z = x * z + y * z :=
|
||||
by rewrite [real.mul_comm, real.left_distrib, {x * _}real.mul_comm, {y * _}real.mul_comm]
|
||||
|
||||
theorem zero_ne_one : ¬ (0 : ℝ) = 1 :=
|
||||
protected theorem zero_ne_one : ¬ (0 : ℝ) = 1 :=
|
||||
take H : 0 = 1,
|
||||
absurd (quot.exact H) (r_zero_nequiv_one)
|
||||
|
||||
|
@ -1147,21 +1150,21 @@ protected definition comm_ring [reducible] : algebra.comm_ring ℝ :=
|
|||
begin
|
||||
fapply algebra.comm_ring.mk,
|
||||
exact add,
|
||||
exact add_assoc,
|
||||
exact real.add_assoc,
|
||||
exact of_num 0,
|
||||
exact zero_add,
|
||||
exact add_zero,
|
||||
exact real.zero_add,
|
||||
exact real.add_zero,
|
||||
exact neg,
|
||||
exact neg_cancel,
|
||||
exact add_comm,
|
||||
exact real.neg_cancel,
|
||||
exact real.add_comm,
|
||||
exact mul,
|
||||
exact mul_assoc,
|
||||
exact real.mul_assoc,
|
||||
apply of_num 1,
|
||||
apply one_mul,
|
||||
apply mul_one,
|
||||
apply left_distrib,
|
||||
apply right_distrib,
|
||||
apply mul_comm
|
||||
apply real.one_mul,
|
||||
apply real.mul_one,
|
||||
apply real.left_distrib,
|
||||
apply real.right_distrib,
|
||||
apply real.mul_comm
|
||||
end
|
||||
|
||||
theorem of_int_eq (a : ℤ) : of_int a = of_rat (rat.of_int a) := rfl
|
||||
|
|
|
@ -16,7 +16,7 @@ are independent of each other.
|
|||
-/
|
||||
|
||||
import data.real.basic data.real.order data.real.division data.rat data.nat data.pnat
|
||||
open rat algebra -- nat
|
||||
open rat algebra
|
||||
local postfix ⁻¹ := pnat.inv
|
||||
open eq.ops pnat classical
|
||||
|
||||
|
@ -110,7 +110,7 @@ theorem equiv_abs_of_ge_zero {s : seq} (Hs : regular s) (Hz : s_le zero s) : s_a
|
|||
apply le.trans,
|
||||
apply add_le_add,
|
||||
repeat (apply neg_le_neg; apply Hz'),
|
||||
rewrite algebra.neg_neg,
|
||||
rewrite neg_neg,
|
||||
apply le.trans,
|
||||
apply add_le_add,
|
||||
repeat (apply inv_ge_of_le; apply Hn),
|
||||
|
@ -174,7 +174,8 @@ private theorem rewrite_helper10 (a b c d : ℝ) : c - d = (c - a) + (a - b) + (
|
|||
noncomputable definition rep (x : ℝ) : rat_seq.reg_seq := some (quot.exists_rep x)
|
||||
|
||||
definition re_abs (x : ℝ) : ℝ :=
|
||||
quot.lift_on x (λ a, quot.mk (rat_seq.r_abs a)) (take a b Hab, quot.sound (rat_seq.r_abs_well_defined Hab))
|
||||
quot.lift_on x (λ a, quot.mk (rat_seq.r_abs a))
|
||||
(take a b Hab, quot.sound (rat_seq.r_abs_well_defined Hab))
|
||||
|
||||
theorem r_abs_nonneg {x : ℝ} : zero ≤ x → re_abs x = x :=
|
||||
quot.induction_on x (λ a Ha, quot.sound (rat_seq.r_equiv_abs_of_ge_zero Ha))
|
||||
|
@ -182,7 +183,8 @@ theorem r_abs_nonneg {x : ℝ} : zero ≤ x → re_abs x = x :=
|
|||
theorem r_abs_nonpos {x : ℝ} : x ≤ zero → re_abs x = -x :=
|
||||
quot.induction_on x (λ a Ha, quot.sound (rat_seq.r_equiv_neg_abs_of_le_zero Ha))
|
||||
|
||||
private theorem abs_const' (a : ℚ) : of_rat (abs a) = re_abs (of_rat a) := quot.sound (rat_seq.r_abs_const a)
|
||||
private theorem abs_const' (a : ℚ) : of_rat (abs a) = re_abs (of_rat a) :=
|
||||
quot.sound (rat_seq.r_abs_const a)
|
||||
|
||||
private theorem re_abs_is_abs : re_abs = abs := funext
|
||||
(begin
|
||||
|
|
|
@ -274,7 +274,8 @@ theorem s_inv_ne_zero {s : seq} (Hs : regular s) (Hsep : sep s zero) (n : ℕ+)
|
|||
apply pnat.mul_le_mul_left
|
||||
end)
|
||||
|
||||
theorem mul_inv {s : seq} (Hs : regular s) (Hsep : sep s zero) : smul s (s_inv Hs) ≡ one :=
|
||||
protected theorem mul_inv {s : seq} (Hs : regular s) (Hsep : sep s zero) :
|
||||
smul s (s_inv Hs) ≡ one :=
|
||||
begin
|
||||
let Rsi := reg_inv_reg Hs Hsep,
|
||||
let Rssi := reg_mul_reg Hs Rsi,
|
||||
|
@ -329,12 +330,13 @@ theorem mul_inv {s : seq} (Hs : regular s) (Hsep : sep s zero) : smul s (s_inv H
|
|||
apply pnat.mul_le_mul_right
|
||||
end
|
||||
|
||||
theorem inv_mul {s : seq} (Hs : regular s) (Hsep : sep s zero) : smul (s_inv Hs) s ≡ one :=
|
||||
protected theorem inv_mul {s : seq} (Hs : regular s) (Hsep : sep s zero) :
|
||||
smul (s_inv Hs) s ≡ one :=
|
||||
begin
|
||||
apply equiv.trans,
|
||||
rotate 3,
|
||||
apply s_mul_comm,
|
||||
apply mul_inv,
|
||||
apply rat_seq.mul_inv,
|
||||
repeat (assumption | apply reg_mul_reg | apply reg_inv_reg | apply zero_is_reg)
|
||||
end
|
||||
|
||||
|
@ -391,7 +393,7 @@ theorem inv_unique {s t : seq} (Hs : regular s) (Ht : regular t) (Hsep : sep s z
|
|||
rotate 3,
|
||||
apply mul_well_defined,
|
||||
rotate 4,
|
||||
apply inv_mul,
|
||||
apply rat_seq.inv_mul,
|
||||
rotate 1,
|
||||
apply equiv.refl,
|
||||
apply s_one_mul,
|
||||
|
@ -415,7 +417,7 @@ theorem inv_well_defined {s t : seq} (Hs : regular s) (Ht : regular t) (Heq : s
|
|||
apply equiv.trans,
|
||||
rotate 3,
|
||||
apply Hm,
|
||||
apply mul_inv,
|
||||
apply rat_seq.mul_inv,
|
||||
repeat (assumption | apply reg_inv_reg | apply reg_mul_reg),
|
||||
apply one_is_reg
|
||||
end)
|
||||
|
@ -543,12 +545,12 @@ theorem s_le_of_equiv_le_right {s t u : seq} (Hs : regular s) (Ht : regular t) (
|
|||
|
||||
noncomputable definition r_inv (s : reg_seq) : reg_seq := reg_seq.mk (s_inv (reg_seq.is_reg s))
|
||||
(if H : sep (reg_seq.sq s) zero then reg_inv_reg (reg_seq.is_reg s) H else
|
||||
assert Hz : s_inv (reg_seq.is_reg s) = zero, from funext (λ n, dif_neg H), by rewrite Hz; apply zero_is_reg)
|
||||
assert Hz : s_inv (reg_seq.is_reg s) = zero, from funext (λ n, dif_neg H),
|
||||
by rewrite Hz; apply zero_is_reg)
|
||||
|
||||
theorem r_inv_zero : requiv (r_inv r_zero) r_zero :=
|
||||
s_zero_inv_equiv_zero
|
||||
|
||||
|
||||
theorem r_inv_well_defined {s t : reg_seq} (H : requiv s t) : requiv (r_inv s) (r_inv t) :=
|
||||
inv_well_defined (reg_seq.is_reg s) (reg_seq.is_reg t) H
|
||||
|
||||
|
@ -556,7 +558,7 @@ theorem r_le_total (s t : reg_seq) : r_le s t ∨ r_le t s :=
|
|||
s_le_total (reg_seq.is_reg s) (reg_seq.is_reg t)
|
||||
|
||||
theorem r_mul_inv (s : reg_seq) (Hsep : r_sep s r_zero) : requiv (s * (r_inv s)) r_one :=
|
||||
mul_inv (reg_seq.is_reg s) Hsep
|
||||
rat_seq.mul_inv (reg_seq.is_reg s) Hsep
|
||||
|
||||
theorem r_sep_of_nequiv (s t : reg_seq) (Hneq : ¬ requiv s t) : r_sep s t :=
|
||||
sep_of_nequiv (reg_seq.is_reg s) (reg_seq.is_reg t) Hneq
|
||||
|
@ -581,26 +583,28 @@ end rat_seq
|
|||
namespace real
|
||||
open [classes] rat_seq
|
||||
|
||||
noncomputable protected definition inv (x : ℝ) : ℝ := quot.lift_on x (λ a, quot.mk (rat_seq.r_inv a))
|
||||
noncomputable protected definition inv (x : ℝ) : ℝ :=
|
||||
quot.lift_on x (λ a, quot.mk (rat_seq.r_inv a))
|
||||
(λ a b H, quot.sound (rat_seq.r_inv_well_defined H))
|
||||
|
||||
noncomputable definition real_has_inv [instance] [reducible] [priority real.prio] : has_inv real :=
|
||||
has_inv.mk real.inv
|
||||
has_inv.mk real.inv
|
||||
|
||||
noncomputable protected definition division (x y : ℝ) : ℝ :=
|
||||
x * y⁻¹
|
||||
x * y⁻¹
|
||||
|
||||
noncomputable definition real_has_division [instance] [reducible] [priority real.prio] : has_division real :=
|
||||
noncomputable definition real_has_division [instance] [reducible] [priority real.prio] :
|
||||
has_division real :=
|
||||
has_division.mk real.division
|
||||
|
||||
theorem le_total (x y : ℝ) : x ≤ y ∨ y ≤ x :=
|
||||
protected theorem le_total (x y : ℝ) : x ≤ y ∨ y ≤ x :=
|
||||
quot.induction_on₂ x y (λ s t, rat_seq.r_le_total s t)
|
||||
|
||||
theorem mul_inv' (x : ℝ) : x ≢ 0 → x * x⁻¹ = 1 :=
|
||||
protected theorem mul_inv_cancel' (x : ℝ) : x ≢ 0 → x * x⁻¹ = 1 :=
|
||||
quot.induction_on x (λ s H, quot.sound (rat_seq.r_mul_inv s H))
|
||||
|
||||
theorem inv_mul' (x : ℝ) : x ≢ 0 → x⁻¹ * x = 1 :=
|
||||
by rewrite real.mul_comm; apply mul_inv'
|
||||
protected theorem inv_mul_cancel' (x : ℝ) : x ≢ 0 → x⁻¹ * x = 1 :=
|
||||
by rewrite real.mul_comm; apply real.mul_inv_cancel'
|
||||
|
||||
theorem neq_of_sep {x y : ℝ} (H : x ≢ y) : ¬ x = y :=
|
||||
assume Heq, !not_sep_self (Heq ▸ H)
|
||||
|
@ -611,19 +615,21 @@ theorem sep_of_neq {x y : ℝ} : ¬ x = y → x ≢ y :=
|
|||
theorem sep_is_neq (x y : ℝ) : (x ≢ y) = (¬ x = y) :=
|
||||
propext (iff.intro neq_of_sep sep_of_neq)
|
||||
|
||||
theorem mul_inv (x : ℝ) : x ≠ 0 → x * x⁻¹ = 1 := !sep_is_neq ▸ !mul_inv'
|
||||
protected theorem mul_inv_cancel (x : ℝ) : x ≠ 0 → x * x⁻¹ = 1 :=
|
||||
!sep_is_neq ▸ !real.mul_inv_cancel'
|
||||
|
||||
theorem inv_mul (x : ℝ) : x ≠ 0 → x⁻¹ * x = 1 := !sep_is_neq ▸ !inv_mul'
|
||||
protected theorem inv_mul_cancel (x : ℝ) : x ≠ 0 → x⁻¹ * x = 1 :=
|
||||
!sep_is_neq ▸ !real.inv_mul_cancel'
|
||||
|
||||
theorem inv_zero : (0 : ℝ)⁻¹ = 0 := quot.sound (rat_seq.r_inv_zero)
|
||||
protected theorem inv_zero : (0 : ℝ)⁻¹ = 0 := quot.sound (rat_seq.r_inv_zero)
|
||||
|
||||
theorem lt_or_eq_of_le (x y : ℝ) : x ≤ y → x < y ∨ x = y :=
|
||||
protected theorem lt_or_eq_of_le (x y : ℝ) : x ≤ y → x < y ∨ x = y :=
|
||||
quot.induction_on₂ x y (λ s t H, or.elim (rat_seq.r_lt_or_equiv_of_le s t H)
|
||||
(assume H1, or.inl H1)
|
||||
(assume H2, or.inr (quot.sound H2)))
|
||||
|
||||
theorem le_iff_lt_or_eq (x y : ℝ) : x ≤ y ↔ x < y ∨ x = y :=
|
||||
iff.intro (lt_or_eq_of_le x y) (le_of_lt_or_eq x y)
|
||||
protected theorem le_iff_lt_or_eq (x y : ℝ) : x ≤ y ↔ x < y ∨ x = y :=
|
||||
iff.intro (real.lt_or_eq_of_le x y) (real.le_of_lt_or_eq x y)
|
||||
|
||||
noncomputable definition dec_lt : decidable_rel real.lt :=
|
||||
begin
|
||||
|
@ -635,13 +641,13 @@ noncomputable definition dec_lt : decidable_rel real.lt :=
|
|||
protected noncomputable definition discrete_linear_ordered_field [reducible] [trans_instance]:
|
||||
algebra.discrete_linear_ordered_field ℝ :=
|
||||
⦃ algebra.discrete_linear_ordered_field, real.comm_ring, real.ordered_ring,
|
||||
le_total := le_total,
|
||||
mul_inv_cancel := mul_inv,
|
||||
inv_mul_cancel := inv_mul,
|
||||
zero_lt_one := zero_lt_one,
|
||||
inv_zero := inv_zero,
|
||||
le_iff_lt_or_eq := le_iff_lt_or_eq,
|
||||
decidable_lt := dec_lt
|
||||
le_total := real.le_total,
|
||||
mul_inv_cancel := real.mul_inv_cancel,
|
||||
inv_mul_cancel := real.inv_mul_cancel,
|
||||
zero_lt_one := real.zero_lt_one,
|
||||
inv_zero := real.inv_zero,
|
||||
le_iff_lt_or_eq := real.le_iff_lt_or_eq,
|
||||
decidable_lt := dec_lt
|
||||
⦄
|
||||
|
||||
theorem of_rat_zero : of_rat (0:rat) = (0:real) := rfl
|
||||
|
|
|
@ -271,8 +271,8 @@ theorem equiv_cancel_middle {s t u : seq} (Hs : regular s) (Ht : regular t)
|
|||
repeat (apply reg_add_reg | apply reg_neg_reg | assumption)
|
||||
end
|
||||
|
||||
theorem add_le_add_of_le_right {s t : seq} (Hs : regular s) (Ht : regular t) (Lst : s_le s t) :
|
||||
∀ u : seq, regular u → s_le (sadd u s) (sadd u t) :=
|
||||
protected theorem add_le_add_of_le_right {s t : seq} (Hs : regular s) (Ht : regular t)
|
||||
(Lst : s_le s t) : ∀ u : seq, regular u → s_le (sadd u s) (sadd u t) :=
|
||||
begin
|
||||
intro u Hu,
|
||||
rewrite [↑s_le at *],
|
||||
|
@ -294,7 +294,8 @@ theorem s_add_lt_add_left {s t : seq} (Hs : regular s) (Ht : regular t) (Hst : s
|
|||
repeat (apply reg_add_reg | apply reg_neg_reg | assumption)
|
||||
end
|
||||
|
||||
theorem add_nonneg_of_nonneg {s t : seq} (Hs : nonneg s) (Ht : nonneg t) : nonneg (sadd s t) :=
|
||||
protected theorem add_nonneg_of_nonneg {s t : seq} (Hs : nonneg s) (Ht : nonneg t) :
|
||||
nonneg (sadd s t) :=
|
||||
begin
|
||||
intros,
|
||||
rewrite [-pnat.add_halves, neg_add],
|
||||
|
@ -303,13 +304,13 @@ theorem add_nonneg_of_nonneg {s t : seq} (Hs : nonneg s) (Ht : nonneg t) : nonne
|
|||
apply Ht
|
||||
end
|
||||
|
||||
protected theorem le_trans {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u) (Lst : s_le s t)
|
||||
(Ltu : s_le t u) : s_le s u :=
|
||||
protected theorem le_trans {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
||||
(Lst : s_le s t) (Ltu : s_le t u) : s_le s u :=
|
||||
begin
|
||||
rewrite ↑s_le at *,
|
||||
let Rz := zero_is_reg,
|
||||
have Hsum : nonneg (sadd (sadd u (sneg t)) (sadd t (sneg s))),
|
||||
from add_nonneg_of_nonneg Ltu Lst,
|
||||
from rat_seq.add_nonneg_of_nonneg Ltu Lst,
|
||||
have H' : nonneg (sadd (sadd u (sadd (sneg t) t)) (sneg s)), begin
|
||||
apply nonneg_of_nonneg_equiv,
|
||||
rotate 2,
|
||||
|
@ -610,7 +611,7 @@ theorem s_mul_nonneg_of_nonneg {s t : seq} (Hs : regular s) (Ht : regular t)
|
|||
rotate 1,
|
||||
rewrite -neg_mul_eq_mul_neg,
|
||||
apply neg_le_neg,
|
||||
rewrite [*pnat.mul_assoc, pnat.inv_mul_eq_mul_inv, -mul.assoc, pnat.inv_cancel_left, rat.one_mul],
|
||||
rewrite [*pnat.mul_assoc, pnat.inv_mul_eq_mul_inv, -mul.assoc, pnat.inv_cancel_left, one_mul],
|
||||
apply inv_ge_of_le,
|
||||
apply pnat.mul_le_mul_left,
|
||||
intro Hsn,
|
||||
|
@ -633,11 +634,12 @@ theorem s_mul_nonneg_of_nonneg {s t : seq} (Hs : regular s) (Ht : regular t)
|
|||
rotate 1,
|
||||
rewrite -neg_mul_eq_neg_mul,
|
||||
apply neg_le_neg,
|
||||
rewrite [+pnat.mul_assoc, pnat.inv_mul_eq_mul_inv, mul.comm, -mul.assoc, pnat.inv_cancel_left, one_mul],
|
||||
rewrite [+pnat.mul_assoc, pnat.inv_mul_eq_mul_inv, mul.comm, -mul.assoc, pnat.inv_cancel_left,
|
||||
one_mul],
|
||||
apply inv_ge_of_le,
|
||||
apply pnat.mul_le_mul_left,
|
||||
intro Htn,
|
||||
apply rat.le_trans,
|
||||
apply le.trans,
|
||||
rotate 1,
|
||||
apply mul_nonneg_of_nonpos_of_nonpos,
|
||||
apply Hsn,
|
||||
|
@ -659,7 +661,7 @@ theorem s_mul_ge_zero_of_ge_zero {s t : seq} (Hs : regular s) (Ht : regular t)
|
|||
apply reg_mul_reg Hs Ht
|
||||
end
|
||||
|
||||
theorem not_lt_self (s : seq) : ¬ s_lt s s :=
|
||||
protected theorem not_lt_self (s : seq) : ¬ s_lt s s :=
|
||||
begin
|
||||
intro Hlt,
|
||||
rewrite [↑s_lt at Hlt, ↑pos at Hlt],
|
||||
|
@ -674,7 +676,7 @@ theorem not_sep_self (s : seq) : ¬ s ≢ s :=
|
|||
intro Hsep,
|
||||
rewrite ↑sep at Hsep,
|
||||
let Hsep' := (iff.mp !or_self) Hsep,
|
||||
apply absurd Hsep' (!not_lt_self)
|
||||
apply absurd Hsep' (!rat_seq.not_lt_self)
|
||||
end
|
||||
|
||||
theorem le_well_defined {s t u v : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
||||
|
@ -951,7 +953,7 @@ theorem r_lt_iff_le_and_sep (s t : reg_seq) : r_lt s t ↔ r_le s t ∧ r_sep s
|
|||
|
||||
theorem r_add_le_add_of_le_right {s t : reg_seq} (H : r_le s t) (u : reg_seq) :
|
||||
r_le (u + s) (u + t) :=
|
||||
add_le_add_of_le_right (reg_seq.is_reg s) (reg_seq.is_reg t) H
|
||||
rat_seq.add_le_add_of_le_right (reg_seq.is_reg s) (reg_seq.is_reg t) H
|
||||
(reg_seq.sq u) (reg_seq.is_reg u)
|
||||
|
||||
theorem r_add_le_add_of_le_right_var (s t u : reg_seq) (H : r_le s t) :
|
||||
|
@ -966,7 +968,7 @@ theorem r_mul_nonneg_of_nonneg {s t : reg_seq} (Hs : r_le r_zero s) (Ht : r_le r
|
|||
s_mul_ge_zero_of_ge_zero (reg_seq.is_reg s) (reg_seq.is_reg t) Hs Ht
|
||||
|
||||
theorem r_not_lt_self (s : reg_seq) : ¬ r_lt s s :=
|
||||
not_lt_self (reg_seq.sq s)
|
||||
rat_seq.not_lt_self (reg_seq.sq s)
|
||||
|
||||
theorem r_not_sep_self (s : reg_seq) : ¬ r_sep s s :=
|
||||
not_sep_self (reg_seq.sq s)
|
||||
|
@ -1006,7 +1008,8 @@ theorem r_lt_of_const_lt_const {a b : ℚ} (H : r_lt (r_const a) (r_const b)) :
|
|||
theorem r_le_of_le_reprs (s t : reg_seq) (Hle : ∀ n : ℕ+, r_le s (r_const (reg_seq.sq t n))) : r_le s t :=
|
||||
le_of_le_reprs (reg_seq.is_reg s) (reg_seq.is_reg t) Hle
|
||||
|
||||
theorem r_le_of_reprs_le (s t : reg_seq) (Hle : ∀ n : ℕ+, r_le (r_const (reg_seq.sq t n)) s) : r_le t s :=
|
||||
theorem r_le_of_reprs_le (s t : reg_seq) (Hle : ∀ n : ℕ+, r_le (r_const (reg_seq.sq t n)) s) :
|
||||
r_le t s :=
|
||||
le_of_reprs_le (reg_seq.is_reg s) (reg_seq.is_reg t) Hle
|
||||
|
||||
end rat_seq
|
||||
|
@ -1015,8 +1018,10 @@ open real
|
|||
open [classes] rat_seq
|
||||
namespace real
|
||||
|
||||
protected definition lt (x y : ℝ) := quot.lift_on₂ x y (λ a b, rat_seq.r_lt a b) rat_seq.r_lt_well_defined
|
||||
protected definition le (x y : ℝ) := quot.lift_on₂ x y (λ a b, rat_seq.r_le a b) rat_seq.r_le_well_defined
|
||||
protected definition lt (x y : ℝ) :=
|
||||
quot.lift_on₂ x y (λ a b, rat_seq.r_lt a b) rat_seq.r_lt_well_defined
|
||||
protected definition le (x y : ℝ) :=
|
||||
quot.lift_on₂ x y (λ a b, rat_seq.r_le a b) rat_seq.r_le_well_defined
|
||||
|
||||
definition real_has_lt [reducible] [instance] [priority real.prio] : has_lt ℝ :=
|
||||
has_lt.mk real.lt
|
||||
|
@ -1027,54 +1032,54 @@ has_le.mk real.le
|
|||
definition sep (x y : ℝ) := quot.lift_on₂ x y (λ a b, rat_seq.r_sep a b) rat_seq.r_sep_well_defined
|
||||
infix `≢` : 50 := sep
|
||||
|
||||
theorem le.refl (x : ℝ) : x ≤ x :=
|
||||
protected theorem le_refl (x : ℝ) : x ≤ x :=
|
||||
quot.induction_on x (λ t, rat_seq.r_le.refl t)
|
||||
|
||||
theorem le.trans {x y z : ℝ} : x ≤ y → y ≤ z → x ≤ z :=
|
||||
protected theorem le_trans {x y z : ℝ} : x ≤ y → y ≤ z → x ≤ z :=
|
||||
quot.induction_on₃ x y z (λ s t u, rat_seq.r_le.trans)
|
||||
|
||||
theorem eq_of_le_of_ge {x y : ℝ} : x ≤ y → y ≤ x → x = y :=
|
||||
protected theorem eq_of_le_of_ge {x y : ℝ} : x ≤ y → y ≤ x → x = y :=
|
||||
quot.induction_on₂ x y (λ s t Hst Hts, quot.sound (rat_seq.r_equiv_of_le_of_ge Hst Hts))
|
||||
|
||||
theorem lt_iff_le_and_sep (x y : ℝ) : x < y ↔ x ≤ y ∧ x ≢ y :=
|
||||
quot.induction_on₂ x y (λ s t, rat_seq.r_lt_iff_le_and_sep s t)
|
||||
|
||||
theorem add_le_add_of_le_right_var (x y z : ℝ) : x ≤ y → z + x ≤ z + y :=
|
||||
protected theorem add_le_add_left' (x y z : ℝ) : x ≤ y → z + x ≤ z + y :=
|
||||
quot.induction_on₃ x y z (λ s t u, rat_seq.r_add_le_add_of_le_right_var s t u)
|
||||
|
||||
theorem add_le_add_of_le_right (x y : ℝ) : x ≤ y → ∀ z : ℝ, z + x ≤ z + y :=
|
||||
take H z, add_le_add_of_le_right_var x y z H
|
||||
protected theorem add_le_add_left (x y : ℝ) : x ≤ y → ∀ z : ℝ, z + x ≤ z + y :=
|
||||
take H z, real.add_le_add_left' x y z H
|
||||
|
||||
theorem mul_gt_zero_of_gt_zero (x y : ℝ) : 0 < x → 0 < y → 0 < x * y :=
|
||||
protected theorem mul_pos (x y : ℝ) : 0 < x → 0 < y → 0 < x * y :=
|
||||
quot.induction_on₂ x y (λ s t, rat_seq.r_mul_pos_of_pos)
|
||||
|
||||
theorem mul_ge_zero_of_ge_zero (x y : ℝ) : 0 ≤ x → 0 ≤ y → 0 ≤ x * y :=
|
||||
protected theorem mul_nonneg (x y : ℝ) : 0 ≤ x → 0 ≤ y → 0 ≤ x * y :=
|
||||
quot.induction_on₂ x y (λ s t, rat_seq.r_mul_nonneg_of_nonneg)
|
||||
|
||||
theorem not_sep_self (x : ℝ) : ¬ x ≢ x :=
|
||||
quot.induction_on x (λ s, rat_seq.r_not_sep_self s)
|
||||
|
||||
theorem not_lt_self (x : ℝ) : ¬ x < x :=
|
||||
protected theorem lt_irrefl (x : ℝ) : ¬ x < x :=
|
||||
quot.induction_on x (λ s, rat_seq.r_not_lt_self s)
|
||||
|
||||
theorem le_of_lt {x y : ℝ} : x < y → x ≤ y :=
|
||||
protected theorem le_of_lt {x y : ℝ} : x < y → x ≤ y :=
|
||||
quot.induction_on₂ x y (λ s t H', rat_seq.r_le_of_lt H')
|
||||
|
||||
theorem lt_of_le_of_lt {x y z : ℝ} : x ≤ y → y < z → x < z :=
|
||||
protected theorem lt_of_le_of_lt {x y z : ℝ} : x ≤ y → y < z → x < z :=
|
||||
quot.induction_on₃ x y z (λ s t u H H', rat_seq.r_lt_of_le_of_lt H H')
|
||||
|
||||
theorem lt_of_lt_of_le {x y z : ℝ} : x < y → y ≤ z → x < z :=
|
||||
protected theorem lt_of_lt_of_le {x y z : ℝ} : x < y → y ≤ z → x < z :=
|
||||
quot.induction_on₃ x y z (λ s t u H H', rat_seq.r_lt_of_lt_of_le H H')
|
||||
|
||||
theorem add_lt_add_left_var (x y z : ℝ) : x < y → z + x < z + y :=
|
||||
protected theorem add_lt_add_left' (x y z : ℝ) : x < y → z + x < z + y :=
|
||||
quot.induction_on₃ x y z (λ s t u, rat_seq.r_add_lt_add_left_var s t u)
|
||||
|
||||
theorem add_lt_add_left (x y : ℝ) : x < y → ∀ z : ℝ, z + x < z + y :=
|
||||
take H z, add_lt_add_left_var x y z H
|
||||
protected theorem add_lt_add_left (x y : ℝ) : x < y → ∀ z : ℝ, z + x < z + y :=
|
||||
take H z, real.add_lt_add_left' x y z H
|
||||
|
||||
theorem zero_lt_one : (0 : ℝ) < (1 : ℝ) := rat_seq.r_zero_lt_one
|
||||
protected theorem zero_lt_one : (0 : ℝ) < (1 : ℝ) := rat_seq.r_zero_lt_one
|
||||
|
||||
theorem le_of_lt_or_eq (x y : ℝ) : x < y ∨ x = y → x ≤ y :=
|
||||
protected theorem le_of_lt_or_eq (x y : ℝ) : x < y ∨ x = y → x ≤ y :=
|
||||
(quot.induction_on₂ x y (λ s t H, or.elim H (take H', begin
|
||||
apply rat_seq.r_le_of_lt_or_eq,
|
||||
apply or.inl H'
|
||||
|
@ -1086,18 +1091,18 @@ theorem le_of_lt_or_eq (x y : ℝ) : x < y ∨ x = y → x ≤ y :=
|
|||
|
||||
definition ordered_ring [reducible] [instance] : algebra.ordered_ring ℝ :=
|
||||
⦃ algebra.ordered_ring, real.comm_ring,
|
||||
le_refl := le.refl,
|
||||
le_trans := @le.trans,
|
||||
mul_pos := mul_gt_zero_of_gt_zero,
|
||||
mul_nonneg := mul_ge_zero_of_ge_zero,
|
||||
zero_ne_one := zero_ne_one,
|
||||
add_le_add_left := add_le_add_of_le_right,
|
||||
le_antisymm := @eq_of_le_of_ge,
|
||||
lt_irrefl := not_lt_self,
|
||||
lt_of_le_of_lt := @lt_of_le_of_lt,
|
||||
lt_of_lt_of_le := @lt_of_lt_of_le,
|
||||
le_of_lt := @le_of_lt,
|
||||
add_lt_add_left := add_lt_add_left
|
||||
le_refl := real.le_refl,
|
||||
le_trans := @real.le_trans,
|
||||
mul_pos := real.mul_pos,
|
||||
mul_nonneg := real.mul_nonneg,
|
||||
zero_ne_one := real.zero_ne_one,
|
||||
add_le_add_left := real.add_le_add_left,
|
||||
le_antisymm := @real.eq_of_le_of_ge,
|
||||
lt_irrefl := real.lt_irrefl,
|
||||
lt_of_le_of_lt := @real.lt_of_le_of_lt,
|
||||
lt_of_lt_of_le := @real.lt_of_lt_of_le,
|
||||
le_of_lt := @real.le_of_lt,
|
||||
add_lt_add_left := real.add_lt_add_left
|
||||
⦄
|
||||
|
||||
open int
|
||||
|
|
Loading…
Add table
Reference in a new issue