refactor(library/{pnat,rat,real}/*): protect theorems in pnat and rat
This commit is contained in:
parent
da5bd03656
commit
9e5a27dc77
8 changed files with 414 additions and 411 deletions
|
@ -51,56 +51,56 @@ has_lt.mk pnat.lt
|
|||
definition pnat_has_one [instance] [reducible] : has_one pnat :=
|
||||
has_one.mk (pos (1:nat) dec_trivial)
|
||||
|
||||
lemma mul.def (p q : ℕ+) : p * q = tag (p~ * q~) (mul_pos (pnat_pos p) (pnat_pos q)) :=
|
||||
protected lemma mul_def (p q : ℕ+) : p * q = tag (p~ * q~) (mul_pos (pnat_pos p) (pnat_pos q)) :=
|
||||
rfl
|
||||
|
||||
lemma le.def (p q : ℕ+) : (p ≤ q) = (p~ ≤ q~) :=
|
||||
protected lemma le_def (p q : ℕ+) : (p ≤ q) = (p~ ≤ q~) :=
|
||||
rfl
|
||||
|
||||
lemma lt.def (p q : ℕ+) : (p < q) = (p~ < q~) :=
|
||||
protected lemma lt_def (p q : ℕ+) : (p < q) = (p~ < q~) :=
|
||||
rfl
|
||||
|
||||
protected theorem pnat.eq {p q : ℕ+} : p~ = q~ → p = q :=
|
||||
subtype.eq
|
||||
|
||||
definition pnat_le_decidable [instance] (p q : ℕ+) : decidable (p ≤ q) :=
|
||||
begin rewrite le.def, exact nat.decidable_le p~ q~ end
|
||||
begin rewrite pnat.le_def, exact nat.decidable_le p~ q~ end
|
||||
|
||||
definition pnat_lt_decidable [instance] {p q : ℕ+} : decidable (p < q) :=
|
||||
begin rewrite lt.def, exact nat.decidable_lt p~ q~ end
|
||||
begin rewrite pnat.lt_def, exact nat.decidable_lt p~ q~ end
|
||||
|
||||
theorem le.trans {p q r : ℕ+} : p ≤ q → q ≤ r → p ≤ r :=
|
||||
begin rewrite +le.def, apply nat.le_trans end
|
||||
protected theorem le_trans {p q r : ℕ+} : p ≤ q → q ≤ r → p ≤ r :=
|
||||
begin rewrite *pnat.le_def, apply nat.le_trans end
|
||||
|
||||
definition max (p q : ℕ+) : ℕ+ :=
|
||||
tag (max p~ q~) (lt_of_lt_of_le (!pnat_pos) (!le_max_right))
|
||||
|
||||
theorem max_right (a b : ℕ+) : max a b ≥ b :=
|
||||
begin change b ≤ max a b, rewrite le.def, apply le_max_right end
|
||||
protected theorem max_right (a b : ℕ+) : max a b ≥ b :=
|
||||
begin change b ≤ max a b, rewrite pnat.le_def, apply le_max_right end
|
||||
|
||||
theorem max_left (a b : ℕ+) : max a b ≥ a :=
|
||||
begin change a ≤ max a b, rewrite le.def, apply le_max_left end
|
||||
protected theorem max_left (a b : ℕ+) : max a b ≥ a :=
|
||||
begin change a ≤ max a b, rewrite pnat.le_def, apply le_max_left end
|
||||
|
||||
theorem max_eq_right {a b : ℕ+} (H : a < b) : max a b = b :=
|
||||
begin rewrite lt.def at H, exact pnat.eq (max_eq_right_of_lt H) end
|
||||
protected theorem max_eq_right {a b : ℕ+} (H : a < b) : max a b = b :=
|
||||
begin rewrite pnat.lt_def at H, exact pnat.eq (max_eq_right_of_lt H) end
|
||||
|
||||
theorem max_eq_left {a b : ℕ+} (H : ¬ a < b) : max a b = a :=
|
||||
begin rewrite lt.def at H, exact pnat.eq (max_eq_left (le_of_not_gt H)) end
|
||||
protected theorem max_eq_left {a b : ℕ+} (H : ¬ a < b) : max a b = a :=
|
||||
begin rewrite pnat.lt_def at H, exact pnat.eq (max_eq_left (le_of_not_gt H)) end
|
||||
|
||||
theorem le_of_lt {a b : ℕ+} : a < b → a ≤ b :=
|
||||
begin rewrite [lt.def, le.def], apply nat.le_of_lt end
|
||||
protected theorem le_of_lt {a b : ℕ+} : a < b → a ≤ b :=
|
||||
begin rewrite [pnat.lt_def, pnat.le_def], apply nat.le_of_lt end
|
||||
|
||||
theorem not_lt_of_ge {a b : ℕ+} : a ≤ b → ¬ (b < a) :=
|
||||
begin rewrite [lt.def, le.def], apply not_lt_of_ge end
|
||||
protected theorem not_lt_of_ge {a b : ℕ+} : a ≤ b → ¬ (b < a) :=
|
||||
begin rewrite [pnat.lt_def, pnat.le_def], apply not_lt_of_ge end
|
||||
|
||||
theorem le_of_not_gt {a b : ℕ+} : ¬ a < b → b ≤ a :=
|
||||
begin rewrite [lt.def, le.def], apply le_of_not_gt end
|
||||
protected theorem le_of_not_gt {a b : ℕ+} : ¬ a < b → b ≤ a :=
|
||||
begin rewrite [pnat.lt_def, pnat.le_def], apply le_of_not_gt end
|
||||
|
||||
theorem eq_of_le_of_ge {a b : ℕ+} : a ≤ b → b ≤ a → a = b :=
|
||||
begin rewrite [+le.def], intros H1 H2, exact pnat.eq (eq_of_le_of_ge H1 H2) end
|
||||
protected theorem eq_of_le_of_ge {a b : ℕ+} : a ≤ b → b ≤ a → a = b :=
|
||||
begin rewrite [+pnat.le_def], intros H1 H2, exact pnat.eq (eq_of_le_of_ge H1 H2) end
|
||||
|
||||
theorem le.refl (a : ℕ+) : a ≤ a :=
|
||||
begin rewrite le.def end
|
||||
protected theorem le_refl (a : ℕ+) : a ≤ a :=
|
||||
begin rewrite pnat.le_def end
|
||||
|
||||
notation 2 := (tag 2 dec_trivial : ℕ+)
|
||||
notation 3 := (tag 3 dec_trivial : ℕ+)
|
||||
|
@ -130,20 +130,20 @@ theorem of_nat_lt_of_nat_of_lt {m n : ℕ} (H : m < n) : of_nat m < of_nat n :=
|
|||
of_nat_lt_of_nat_of_lt H
|
||||
|
||||
theorem rat_of_pnat_le_of_pnat_le {m n : ℕ+} (H : m ≤ n) : rat_of_pnat m ≤ rat_of_pnat n :=
|
||||
begin rewrite le.def at H, exact of_nat_le_of_nat_of_le H end
|
||||
begin rewrite pnat.le_def at H, exact of_nat_le_of_nat_of_le H end
|
||||
|
||||
theorem rat_of_pnat_lt_of_pnat_lt {m n : ℕ+} (H : m < n) : rat_of_pnat m < rat_of_pnat n :=
|
||||
begin rewrite lt.def at H, exact of_nat_lt_of_nat_of_lt H end
|
||||
begin rewrite pnat.lt_def at H, exact of_nat_lt_of_nat_of_lt H end
|
||||
|
||||
theorem pnat_le_of_rat_of_pnat_le {m n : ℕ+} (H : rat_of_pnat m ≤ rat_of_pnat n) : m ≤ n :=
|
||||
begin rewrite le.def, exact le_of_of_nat_le_of_nat H end
|
||||
begin rewrite pnat.le_def, exact le_of_of_nat_le_of_nat H end
|
||||
|
||||
definition inv (n : ℕ+) : ℚ :=
|
||||
(1 : ℚ) / rat_of_pnat n
|
||||
|
||||
local postfix `⁻¹` := inv
|
||||
|
||||
theorem inv_pos (n : ℕ+) : n⁻¹ > 0 := one_div_pos_of_pos !rat_of_pnat_is_pos
|
||||
protected theorem inv_pos (n : ℕ+) : n⁻¹ > 0 := one_div_pos_of_pos !rat_of_pnat_is_pos
|
||||
|
||||
theorem inv_le_one (n : ℕ+) : n⁻¹ ≤ (1 : ℚ) :=
|
||||
begin
|
||||
|
@ -168,25 +168,25 @@ theorem pone_inv : pone⁻¹ = 1 := rfl
|
|||
|
||||
theorem add_invs_nonneg (m n : ℕ+) : 0 ≤ m⁻¹ + n⁻¹ :=
|
||||
begin
|
||||
apply rat.le_of_lt,
|
||||
apply le_of_lt,
|
||||
apply add_pos,
|
||||
repeat apply inv_pos
|
||||
repeat apply pnat.inv_pos
|
||||
end
|
||||
|
||||
theorem one_mul (n : ℕ+) : pone * n = n :=
|
||||
protected theorem one_mul (n : ℕ+) : pone * n = n :=
|
||||
begin
|
||||
apply pnat.eq,
|
||||
unfold pone,
|
||||
rewrite [mul.def, ↑nat_of_pnat, algebra.one_mul]
|
||||
rewrite [pnat.mul_def, ↑nat_of_pnat, algebra.one_mul]
|
||||
end
|
||||
|
||||
theorem pone_le (n : ℕ+) : pone ≤ n :=
|
||||
begin rewrite le.def, exact succ_le_of_lt (pnat_pos n) end
|
||||
begin rewrite pnat.le_def, exact succ_le_of_lt (pnat_pos n) end
|
||||
|
||||
theorem pnat_to_rat_mul (a b : ℕ+) : rat_of_pnat (a * b) = rat_of_pnat a * rat_of_pnat b := rfl
|
||||
|
||||
theorem mul_lt_mul_left {a b c : ℕ+} (H : a < b) : a * c < b * c :=
|
||||
begin rewrite [lt.def at *], exact mul_lt_mul_of_pos_right H !pnat_pos end
|
||||
begin rewrite [pnat.lt_def at *], exact mul_lt_mul_of_pos_right H !pnat_pos end
|
||||
|
||||
theorem one_lt_two : pone < 2 :=
|
||||
!nat.le_refl
|
||||
|
@ -197,8 +197,8 @@ begin
|
|||
apply one_div_lt_one_div_of_lt,
|
||||
apply rat_of_pnat_is_pos,
|
||||
have H : n~ < (2 * n)~, begin
|
||||
rewrite -one_mul at {1},
|
||||
rewrite -lt.def,
|
||||
rewrite -pnat.one_mul at {1},
|
||||
rewrite -pnat.lt_def,
|
||||
apply mul_lt_mul_left,
|
||||
apply one_lt_two
|
||||
end,
|
||||
|
@ -223,65 +223,65 @@ by rewrite pnat_to_rat_mul
|
|||
theorem add_halves (p : ℕ+) : (2 * p)⁻¹ + (2 * p)⁻¹ = p⁻¹ :=
|
||||
begin
|
||||
rewrite [↑inv, -(add_halves (1 / (rat_of_pnat p))), algebra.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],
|
||||
have H : rat_of_pnat (2 * p) = rat_of_pnat p * (1 + 1), by rewrite [rat.mul_comm, two_mul],
|
||||
rewrite *H
|
||||
end
|
||||
|
||||
theorem add_halves_double (m n : ℕ+) :
|
||||
m⁻¹ + n⁻¹ = ((2 * m)⁻¹ + (2 * n)⁻¹) + ((2 * m)⁻¹ + (2 * n)⁻¹) :=
|
||||
have hsimp [visible] : ∀ a b : ℚ, (a + a) + (b + b) = (a + b) + (a + b),
|
||||
by intros; rewrite [rat.add.assoc, -(rat.add.assoc a b b), {_+b}rat.add.comm, -*rat.add.assoc],
|
||||
by intros; rewrite [rat.add_assoc, -(rat.add_assoc a b b), {_+b}rat.add_comm, -*rat.add_assoc],
|
||||
by rewrite [-add_halves m, -add_halves n, hsimp]
|
||||
|
||||
theorem inv_mul_eq_mul_inv {p q : ℕ+} : (p * q)⁻¹ = p⁻¹ * q⁻¹ :=
|
||||
protected theorem inv_mul_eq_mul_inv {p q : ℕ+} : (p * q)⁻¹ = p⁻¹ * q⁻¹ :=
|
||||
begin rewrite [↑inv, pnat_to_rat_mul, algebra.one_div_mul_one_div] end
|
||||
|
||||
theorem inv_mul_le_inv (p q : ℕ+) : (p * q)⁻¹ ≤ q⁻¹ :=
|
||||
protected theorem inv_mul_le_inv (p q : ℕ+) : (p * q)⁻¹ ≤ q⁻¹ :=
|
||||
begin
|
||||
rewrite [inv_mul_eq_mul_inv, -{q⁻¹}rat.one_mul at {2}],
|
||||
rewrite [pnat.inv_mul_eq_mul_inv, -{q⁻¹}rat.one_mul at {2}],
|
||||
apply algebra.mul_le_mul,
|
||||
apply inv_le_one,
|
||||
apply rat.le.refl,
|
||||
apply rat.le_of_lt,
|
||||
apply inv_pos,
|
||||
apply le.refl,
|
||||
apply le_of_lt,
|
||||
apply pnat.inv_pos,
|
||||
apply rat.le_of_lt rat.zero_lt_one
|
||||
end
|
||||
|
||||
theorem pnat_mul_le_mul_left' (a b c : ℕ+) : a ≤ b → c * a ≤ c * b :=
|
||||
begin
|
||||
rewrite +le.def, intro H,
|
||||
rewrite +pnat.le_def, intro H,
|
||||
apply mul_le_mul_of_nonneg_left H,
|
||||
apply algebra.le_of_lt,
|
||||
apply pnat_pos
|
||||
end
|
||||
|
||||
theorem mul.assoc (a b c : ℕ+) : a * b * c = a * (b * c) :=
|
||||
protected theorem mul_assoc (a b c : ℕ+) : a * b * c = a * (b * c) :=
|
||||
pnat.eq !mul.assoc
|
||||
|
||||
theorem mul.comm (a b : ℕ+) : a * b = b * a :=
|
||||
protected theorem mul_comm (a b : ℕ+) : a * b = b * a :=
|
||||
pnat.eq !mul.comm
|
||||
|
||||
theorem add.assoc (a b c : ℕ+) : a + b + c = a + (b + c) :=
|
||||
protected theorem add_assoc (a b c : ℕ+) : a + b + c = a + (b + c) :=
|
||||
pnat.eq !add.assoc
|
||||
|
||||
theorem mul_le_mul_left (p q : ℕ+) : q ≤ p * q :=
|
||||
protected theorem mul_le_mul_left (p q : ℕ+) : q ≤ p * q :=
|
||||
begin
|
||||
rewrite [-one_mul at {1}, mul.comm, mul.comm p],
|
||||
rewrite [-pnat.one_mul at {1}, pnat.mul_comm, pnat.mul_comm p],
|
||||
apply pnat_mul_le_mul_left',
|
||||
apply pone_le
|
||||
end
|
||||
|
||||
theorem mul_le_mul_right (p q : ℕ+) : p ≤ p * q :=
|
||||
by rewrite mul.comm; apply mul_le_mul_left
|
||||
protected theorem mul_le_mul_right (p q : ℕ+) : p ≤ p * q :=
|
||||
by rewrite pnat.mul_comm; apply pnat.mul_le_mul_left
|
||||
|
||||
theorem pnat.lt_of_not_le {p q : ℕ+} : ¬ p ≤ q → q < p :=
|
||||
begin rewrite [le.def, lt.def], apply lt_of_not_ge end
|
||||
begin rewrite [pnat.le_def, pnat.lt_def], apply lt_of_not_ge end
|
||||
|
||||
theorem inv_cancel_left (p : ℕ+) : rat_of_pnat p * p⁻¹ = (1 : ℚ) :=
|
||||
protected theorem inv_cancel_left (p : ℕ+) : rat_of_pnat p * p⁻¹ = (1 : ℚ) :=
|
||||
mul_one_div_cancel (ne.symm (ne_of_lt !rat_of_pnat_is_pos))
|
||||
|
||||
theorem inv_cancel_right (p : ℕ+) : p⁻¹ * rat_of_pnat p = (1 : ℚ) :=
|
||||
by rewrite rat.mul.comm; apply inv_cancel_left
|
||||
protected theorem inv_cancel_right (p : ℕ+) : p⁻¹ * rat_of_pnat p = (1 : ℚ) :=
|
||||
by rewrite rat.mul_comm; apply pnat.inv_cancel_left
|
||||
|
||||
theorem lt_add_left (p q : ℕ+) : p < p + q :=
|
||||
begin
|
||||
|
@ -300,7 +300,7 @@ theorem div_le_pnat (q : ℚ) (n : ℕ+) (H : q ≥ n⁻¹) : 1 / q ≤ rat_of_p
|
|||
begin
|
||||
apply algebra.div_le_of_le_mul,
|
||||
apply algebra.lt_of_lt_of_le,
|
||||
apply inv_pos,
|
||||
apply pnat.inv_pos,
|
||||
rotate 1,
|
||||
apply H,
|
||||
apply le_mul_of_div_le,
|
||||
|
@ -312,10 +312,10 @@ theorem pnat_cancel' (n m : ℕ+) : (n * n * m)⁻¹ * (rat_of_pnat n * rat_of_p
|
|||
assert hsimp : ∀ a b c : ℚ, (a * a * (b * b * c)) = (a * b) * (a * b) * c,
|
||||
begin
|
||||
intro a b c,
|
||||
rewrite[-*rat.mul.assoc],
|
||||
rewrite[-*rat.mul_assoc],
|
||||
exact (!mul.right_comm ▸ rfl),
|
||||
end,
|
||||
by rewrite [rat.mul.comm, *inv_mul_eq_mul_inv, hsimp, *inv_cancel_left, *rat.one_mul]
|
||||
by rewrite [rat.mul_comm, *pnat.inv_mul_eq_mul_inv, hsimp, *pnat.inv_cancel_left, *rat.one_mul]
|
||||
|
||||
definition pceil (a : ℚ) : ℕ+ := tag (ubound a) !ubound_pos
|
||||
|
||||
|
@ -350,7 +350,7 @@ begin
|
|||
apply algebra.lt.trans,
|
||||
rotate 1,
|
||||
apply and.left Hc,
|
||||
rewrite rat.add.assoc,
|
||||
rewrite rat.add_assoc,
|
||||
apply rat.add_lt_add_left,
|
||||
rewrite -(algebra.add_halves c) at {3},
|
||||
apply add_lt_add,
|
||||
|
@ -383,13 +383,13 @@ begin
|
|||
existsi (pceil (1 / ε)),
|
||||
rewrite -(one_div_one_div ε) at {2},
|
||||
apply pceil_helper,
|
||||
apply le.refl,
|
||||
apply pnat.le_refl,
|
||||
apply one_div_pos_of_pos Hε
|
||||
end
|
||||
|
||||
theorem p_add_fractions (n : ℕ+) : (2 * n)⁻¹ + (2 * 3 * n)⁻¹ + (3 * n)⁻¹ = n⁻¹ :=
|
||||
assert T : 2⁻¹ + 2⁻¹ * 3⁻¹ + 3⁻¹ = 1, from dec_trivial,
|
||||
by rewrite[*inv_mul_eq_mul_inv,-*right_distrib,T,rat.one_mul]
|
||||
by rewrite[*pnat.inv_mul_eq_mul_inv,-*right_distrib,T,rat.one_mul]
|
||||
|
||||
theorem rat_power_two_le (k : ℕ+) : rat_of_pnat k ≤ 2^k~ :=
|
||||
!binary_nat_bound
|
||||
|
|
|
@ -444,6 +444,9 @@ prerat.of_int.inj (quot.exact H)
|
|||
theorem eq_of_of_int_eq_of_int {a b : ℤ} (H : of_int a = of_int b) : a = b :=
|
||||
of_int.inj H
|
||||
|
||||
theorem of_int_eq_of_int_iff (a b : ℤ) : of_int a = of_int b ↔ a = b :=
|
||||
iff.intro eq_of_of_int_eq_of_int !congr_arg
|
||||
|
||||
theorem of_nat_eq (a : ℕ) : of_nat a = of_int (int.of_nat a) :=
|
||||
rfl
|
||||
|
||||
|
@ -471,47 +474,47 @@ of_nat.inj H
|
|||
theorem of_nat_eq_of_nat_iff (a b : ℕ) : of_nat a = of_nat b ↔ a = b :=
|
||||
iff.intro of_nat.inj !congr_arg
|
||||
|
||||
theorem add.comm (a b : ℚ) : a + b = b + a :=
|
||||
protected theorem add_comm (a b : ℚ) : a + b = b + a :=
|
||||
quot.induction_on₂ a b (take u v, quot.sound !prerat.add.comm)
|
||||
|
||||
theorem add.assoc (a b c : ℚ) : a + b + c = a + (b + c) :=
|
||||
protected theorem add_assoc (a b c : ℚ) : a + b + c = a + (b + c) :=
|
||||
quot.induction_on₃ a b c (take u v w, quot.sound !prerat.add.assoc)
|
||||
|
||||
theorem add_zero (a : ℚ) : a + 0 = a :=
|
||||
protected theorem add_zero (a : ℚ) : a + 0 = a :=
|
||||
quot.induction_on a (take u, quot.sound !prerat.add_zero)
|
||||
|
||||
theorem zero_add (a : ℚ) : 0 + a = a := !add.comm ▸ !add_zero
|
||||
protected theorem zero_add (a : ℚ) : 0 + a = a := !rat.add_comm ▸ !rat.add_zero
|
||||
|
||||
theorem add.left_inv (a : ℚ) : -a + a = 0 :=
|
||||
protected theorem add_left_inv (a : ℚ) : -a + a = 0 :=
|
||||
quot.induction_on a (take u, quot.sound !prerat.add.left_inv)
|
||||
|
||||
theorem mul.comm (a b : ℚ) : a * b = b * a :=
|
||||
protected theorem mul_comm (a b : ℚ) : a * b = b * a :=
|
||||
quot.induction_on₂ a b (take u v, quot.sound !prerat.mul.comm)
|
||||
|
||||
theorem mul.assoc (a b c : ℚ) : a * b * c = a * (b * c) :=
|
||||
protected theorem mul_assoc (a b c : ℚ) : a * b * c = a * (b * c) :=
|
||||
quot.induction_on₃ a b c (take u v w, quot.sound !prerat.mul.assoc)
|
||||
|
||||
theorem mul_one (a : ℚ) : a * 1 = a :=
|
||||
protected theorem mul_one (a : ℚ) : a * 1 = a :=
|
||||
quot.induction_on a (take u, quot.sound !prerat.mul_one)
|
||||
|
||||
theorem one_mul (a : ℚ) : 1 * a = a := !mul.comm ▸ !mul_one
|
||||
protected theorem one_mul (a : ℚ) : 1 * a = a := !rat.mul_comm ▸ !rat.mul_one
|
||||
|
||||
theorem mul.left_distrib (a b c : ℚ) : a * (b + c) = a * b + a * c :=
|
||||
protected theorem left_distrib (a b c : ℚ) : a * (b + c) = a * b + a * c :=
|
||||
quot.induction_on₃ a b c (take u v w, quot.sound !prerat.mul.left_distrib)
|
||||
|
||||
theorem mul.right_distrib (a b c : ℚ) : (a + b) * c = a * c + b * c :=
|
||||
by rewrite [mul.comm, mul.left_distrib, *mul.comm c]
|
||||
protected theorem right_distrib (a b c : ℚ) : (a + b) * c = a * c + b * c :=
|
||||
by rewrite [rat.mul_comm, rat.left_distrib, *rat.mul_comm c]
|
||||
|
||||
theorem mul_inv_cancel {a : ℚ} : a ≠ 0 → a * a⁻¹ = 1 :=
|
||||
protected theorem mul_inv_cancel {a : ℚ} : a ≠ 0 → a * a⁻¹ = 1 :=
|
||||
quot.induction_on a
|
||||
(take u,
|
||||
assume H,
|
||||
quot.sound (!prerat.mul_inv_cancel (assume H1, H (quot.sound H1))))
|
||||
|
||||
theorem inv_mul_cancel {a : ℚ} (H : a ≠ 0) : a⁻¹ * a = 1 :=
|
||||
!mul.comm ▸ mul_inv_cancel H
|
||||
protected theorem inv_mul_cancel {a : ℚ} (H : a ≠ 0) : a⁻¹ * a = 1 :=
|
||||
!rat.mul_comm ▸ rat.mul_inv_cancel H
|
||||
|
||||
theorem zero_ne_one : (0 : ℚ) ≠ 1 :=
|
||||
protected theorem zero_ne_one : (0 : ℚ) ≠ 1 :=
|
||||
assume H, prerat.zero_not_equiv_one (quot.exact H)
|
||||
|
||||
definition has_decidable_eq [instance] : decidable_eq ℚ :=
|
||||
|
@ -521,7 +524,7 @@ take a b, quot.rec_on_subsingleton₂ a b
|
|||
then decidable.inl (quot.sound H)
|
||||
else decidable.inr (assume H1, H (quot.exact H1)))
|
||||
|
||||
theorem inv_zero : inv 0 = (0 : ℚ) :=
|
||||
protected theorem inv_zero : inv 0 = (0 : ℚ) :=
|
||||
quot.sound (prerat.inv_zero' ▸ !prerat.equiv.refl)
|
||||
|
||||
theorem quot_reduce (a : ℚ) : ⟦reduce a⟧ = a :=
|
||||
|
@ -551,25 +554,25 @@ decidable.by_cases
|
|||
protected definition discrete_field [reducible] [trans_instance] : algebra.discrete_field rat :=
|
||||
⦃algebra.discrete_field,
|
||||
add := rat.add,
|
||||
add_assoc := add.assoc,
|
||||
add_assoc := rat.add_assoc,
|
||||
zero := 0,
|
||||
zero_add := zero_add,
|
||||
add_zero := add_zero,
|
||||
zero_add := rat.zero_add,
|
||||
add_zero := rat.add_zero,
|
||||
neg := rat.neg,
|
||||
add_left_inv := add.left_inv,
|
||||
add_comm := add.comm,
|
||||
add_left_inv := rat.add_left_inv,
|
||||
add_comm := rat.add_comm,
|
||||
mul := rat.mul,
|
||||
mul_assoc := mul.assoc,
|
||||
mul_assoc := rat.mul_assoc,
|
||||
one := 1,
|
||||
one_mul := one_mul,
|
||||
mul_one := mul_one,
|
||||
left_distrib := mul.left_distrib,
|
||||
right_distrib := mul.right_distrib,
|
||||
mul_comm := mul.comm,
|
||||
mul_inv_cancel := @mul_inv_cancel,
|
||||
inv_mul_cancel := @inv_mul_cancel,
|
||||
zero_ne_one := zero_ne_one,
|
||||
inv_zero := inv_zero,
|
||||
one_mul := rat.one_mul,
|
||||
mul_one := rat.mul_one,
|
||||
left_distrib := rat.left_distrib,
|
||||
right_distrib := rat.right_distrib,
|
||||
mul_comm := rat.mul_comm,
|
||||
mul_inv_cancel := @rat.mul_inv_cancel,
|
||||
inv_mul_cancel := @rat.inv_mul_cancel,
|
||||
zero_ne_one := rat.zero_ne_one,
|
||||
inv_zero := rat.inv_zero,
|
||||
has_decidable_eq := has_decidable_eq⦄
|
||||
|
||||
definition rat_has_division [instance] [reducible] [priority rat.prio] : has_division rat :=
|
||||
|
|
|
@ -149,10 +149,10 @@ has_lt.mk rat.lt
|
|||
definition rat_has_le [reducible] [instance] [priority rat.prio] : has_le rat :=
|
||||
has_le.mk rat.le
|
||||
|
||||
lemma lt.def (a b : ℚ) : (a < b) = pos (b - a) :=
|
||||
protected lemma lt_def (a b : ℚ) : (a < b) = pos (b - a) :=
|
||||
rfl
|
||||
|
||||
lemma le.def (a b : ℚ) : (a ≤ b) = nonneg (b - a) :=
|
||||
protected lemma le_def (a b : ℚ) : (a ≤ b) = nonneg (b - a) :=
|
||||
rfl
|
||||
|
||||
theorem of_int_lt_of_int_iff (a b : ℤ) : of_int a < of_int b ↔ a < b :=
|
||||
|
@ -202,10 +202,10 @@ iff.mp !of_nat_le_of_nat_iff H
|
|||
theorem of_nat_nonneg (a : ℕ) : (of_nat a ≥ 0) :=
|
||||
of_nat_le_of_nat_of_le !nat.zero_le
|
||||
|
||||
theorem le.refl (a : ℚ) : a ≤ a :=
|
||||
by rewrite [le.def, sub_self]; apply nonneg_zero
|
||||
protected theorem le_refl (a : ℚ) : a ≤ a :=
|
||||
by rewrite [rat.le_def, sub_self]; apply nonneg_zero
|
||||
|
||||
theorem le.trans (H1 : a ≤ b) (H2 : b ≤ c) : a ≤ c :=
|
||||
protected theorem le_trans (H1 : a ≤ b) (H2 : b ≤ c) : a ≤ c :=
|
||||
assert H3 : nonneg (c - b + (b - a)), from nonneg_add H2 H1,
|
||||
begin
|
||||
revert H3,
|
||||
|
@ -213,20 +213,20 @@ begin
|
|||
intro H3, apply H3
|
||||
end
|
||||
|
||||
theorem le.antisymm (H1 : a ≤ b) (H2 : b ≤ a) : a = b :=
|
||||
protected theorem le_antisymm (H1 : a ≤ b) (H2 : b ≤ a) : a = b :=
|
||||
have H3 : nonneg (-(a - b)), from !neg_sub⁻¹ ▸ H1,
|
||||
have H4 : a - b = 0, from nonneg_antisymm H2 H3,
|
||||
eq_of_sub_eq_zero H4
|
||||
|
||||
theorem le.total (a b : ℚ) : a ≤ b ∨ b ≤ a :=
|
||||
protected theorem le_total (a b : ℚ) : a ≤ b ∨ b ≤ a :=
|
||||
or.elim (nonneg_total (b - a))
|
||||
(assume H, or.inl H)
|
||||
(assume H, or.inr begin rewrite neg_sub at H, exact H end)
|
||||
|
||||
theorem le.by_cases {P : Prop} (a b : ℚ) (H : a ≤ b → P) (H2 : b ≤ a → P) : P :=
|
||||
or.elim (!rat.le.total) H H2
|
||||
protected theorem le_by_cases {P : Prop} (a b : ℚ) (H : a ≤ b → P) (H2 : b ≤ a → P) : P :=
|
||||
or.elim (!rat.le_total) H H2
|
||||
|
||||
theorem lt_iff_le_and_ne (a b : ℚ) : a < b ↔ a ≤ b ∧ a ≠ b :=
|
||||
protected theorem lt_iff_le_and_ne (a b : ℚ) : a < b ↔ a ≤ b ∧ a ≠ b :=
|
||||
iff.intro
|
||||
(assume H : a < b,
|
||||
have b - a ≠ 0, from ne_zero_of_pos H,
|
||||
|
@ -237,90 +237,90 @@ iff.intro
|
|||
have b - a ≠ 0, from (assume H', aneb (eq_of_sub_eq_zero H')⁻¹),
|
||||
pos_of_nonneg_of_ne_zero aleb this)
|
||||
|
||||
theorem le_iff_lt_or_eq (a b : ℚ) : a ≤ b ↔ a < b ∨ a = b :=
|
||||
protected theorem le_iff_lt_or_eq (a b : ℚ) : a ≤ b ↔ a < b ∨ a = b :=
|
||||
iff.intro
|
||||
(assume h : a ≤ b,
|
||||
decidable.by_cases
|
||||
(suppose a = b, or.inr this)
|
||||
(suppose a ≠ b, or.inl (iff.mpr !lt_iff_le_and_ne (and.intro h this))))
|
||||
(suppose a ≠ b, or.inl (iff.mpr !rat.lt_iff_le_and_ne (and.intro h this))))
|
||||
(suppose a < b ∨ a = b,
|
||||
or.elim this
|
||||
(suppose a < b, and.left (iff.mp !lt_iff_le_and_ne this))
|
||||
(suppose a = b, this ▸ !le.refl))
|
||||
(suppose a < b, and.left (iff.mp !rat.lt_iff_le_and_ne this))
|
||||
(suppose a = b, this ▸ !rat.le_refl))
|
||||
|
||||
private theorem to_nonneg : a ≥ 0 → nonneg a :=
|
||||
by intros; rewrite -sub_zero; eassumption
|
||||
|
||||
theorem add_le_add_left (H : a ≤ b) (c : ℚ) : c + a ≤ c + b :=
|
||||
protected theorem add_le_add_left (H : a ≤ b) (c : ℚ) : c + a ≤ c + b :=
|
||||
have c + b - (c + a) = b - a,
|
||||
by rewrite [sub.def, neg_add, -add.assoc, add.comm c, add_neg_cancel_right],
|
||||
show nonneg (c + b - (c + a)), from this⁻¹ ▸ H
|
||||
|
||||
theorem mul_nonneg (H1 : a ≥ (0 : ℚ)) (H2 : b ≥ (0 : ℚ)) : a * b ≥ (0 : ℚ) :=
|
||||
protected theorem mul_nonneg (H1 : a ≥ (0 : ℚ)) (H2 : b ≥ (0 : ℚ)) : a * b ≥ (0 : ℚ) :=
|
||||
assert nonneg (a * b), from nonneg_mul (to_nonneg H1) (to_nonneg H2),
|
||||
begin rewrite -sub_zero at this, exact this end
|
||||
|
||||
private theorem to_pos : a > 0 → pos a :=
|
||||
by intros; rewrite -sub_zero; eassumption
|
||||
|
||||
theorem mul_pos (H1 : a > (0 : ℚ)) (H2 : b > (0 : ℚ)) : a * b > (0 : ℚ) :=
|
||||
protected theorem mul_pos (H1 : a > (0 : ℚ)) (H2 : b > (0 : ℚ)) : a * b > (0 : ℚ) :=
|
||||
assert pos (a * b), from pos_mul (to_pos H1) (to_pos H2),
|
||||
begin rewrite -sub_zero at this, exact this end
|
||||
|
||||
definition decidable_lt [instance] : decidable_rel rat.lt :=
|
||||
take a b, decidable_pos (b - a)
|
||||
|
||||
theorem le_of_lt (H : a < b) : a ≤ b := iff.mpr !le_iff_lt_or_eq (or.inl H)
|
||||
protected theorem le_of_lt (H : a < b) : a ≤ b := iff.mpr !rat.le_iff_lt_or_eq (or.inl H)
|
||||
|
||||
theorem lt_irrefl (a : ℚ) : ¬ a < a :=
|
||||
protected theorem lt_irrefl (a : ℚ) : ¬ a < a :=
|
||||
take Ha,
|
||||
let Hand := (iff.mp !lt_iff_le_and_ne) Ha in
|
||||
let Hand := (iff.mp !rat.lt_iff_le_and_ne) Ha in
|
||||
(and.right Hand) rfl
|
||||
|
||||
theorem not_le_of_gt (H : a < b) : ¬ b ≤ a :=
|
||||
protected theorem not_le_of_gt (H : a < b) : ¬ b ≤ a :=
|
||||
assume Hba,
|
||||
let Heq := le.antisymm (le_of_lt H) Hba in
|
||||
!lt_irrefl (Heq ▸ H)
|
||||
let Heq := rat.le_antisymm (rat.le_of_lt H) Hba in
|
||||
!rat.lt_irrefl (Heq ▸ H)
|
||||
|
||||
theorem lt_of_lt_of_le (Hab : a < b) (Hbc : b ≤ c) : a < c :=
|
||||
let Hab' := le_of_lt Hab in
|
||||
let Hac := le.trans Hab' Hbc in
|
||||
(iff.mpr !lt_iff_le_and_ne) (and.intro Hac
|
||||
(assume Heq, not_le_of_gt (Heq ▸ Hab) Hbc))
|
||||
protected theorem lt_of_lt_of_le (Hab : a < b) (Hbc : b ≤ c) : a < c :=
|
||||
let Hab' := rat.le_of_lt Hab in
|
||||
let Hac := rat.le_trans Hab' Hbc in
|
||||
(iff.mpr !rat.lt_iff_le_and_ne) (and.intro Hac
|
||||
(assume Heq, rat.not_le_of_gt (Heq ▸ Hab) Hbc))
|
||||
|
||||
theorem lt_of_le_of_lt (Hab : a ≤ b) (Hbc : b < c) : a < c :=
|
||||
let Hbc' := le_of_lt Hbc in
|
||||
let Hac := le.trans Hab Hbc' in
|
||||
(iff.mpr !lt_iff_le_and_ne) (and.intro Hac
|
||||
(assume Heq, not_le_of_gt (Heq⁻¹ ▸ Hbc) Hab))
|
||||
protected theorem lt_of_le_of_lt (Hab : a ≤ b) (Hbc : b < c) : a < c :=
|
||||
let Hbc' := rat.le_of_lt Hbc in
|
||||
let Hac := rat.le_trans Hab Hbc' in
|
||||
(iff.mpr !rat.lt_iff_le_and_ne) (and.intro Hac
|
||||
(assume Heq, rat.not_le_of_gt (Heq⁻¹ ▸ Hbc) Hab))
|
||||
|
||||
theorem zero_lt_one : (0 : ℚ) < 1 := trivial
|
||||
protected theorem zero_lt_one : (0 : ℚ) < 1 := trivial
|
||||
|
||||
theorem add_lt_add_left (H : a < b) (c : ℚ) : c + a < c + b :=
|
||||
let H' := le_of_lt H in
|
||||
(iff.mpr (lt_iff_le_and_ne _ _)) (and.intro (add_le_add_left H' _)
|
||||
protected theorem add_lt_add_left (H : a < b) (c : ℚ) : c + a < c + b :=
|
||||
let H' := rat.le_of_lt H in
|
||||
(iff.mpr (rat.lt_iff_le_and_ne _ _)) (and.intro (rat.add_le_add_left H' _)
|
||||
(take Heq, let Heq' := add_left_cancel Heq in
|
||||
!lt_irrefl (Heq' ▸ H)))
|
||||
!rat.lt_irrefl (Heq' ▸ H)))
|
||||
|
||||
protected definition discrete_linear_ordered_field [reducible] [trans_instance] :
|
||||
algebra.discrete_linear_ordered_field rat :=
|
||||
⦃algebra.discrete_linear_ordered_field,
|
||||
rat.discrete_field,
|
||||
le_refl := le.refl,
|
||||
le_trans := @le.trans,
|
||||
le_antisymm := @le.antisymm,
|
||||
le_total := @le.total,
|
||||
le_of_lt := @le_of_lt,
|
||||
lt_irrefl := lt_irrefl,
|
||||
lt_of_lt_of_le := @lt_of_lt_of_le,
|
||||
lt_of_le_of_lt := @lt_of_le_of_lt,
|
||||
le_iff_lt_or_eq := @le_iff_lt_or_eq,
|
||||
add_le_add_left := @add_le_add_left,
|
||||
mul_nonneg := @mul_nonneg,
|
||||
mul_pos := @mul_pos,
|
||||
le_refl := rat.le_refl,
|
||||
le_trans := @rat.le_trans,
|
||||
le_antisymm := @rat.le_antisymm,
|
||||
le_total := @rat.le_total,
|
||||
le_of_lt := @rat.le_of_lt,
|
||||
lt_irrefl := rat.lt_irrefl,
|
||||
lt_of_lt_of_le := @rat.lt_of_lt_of_le,
|
||||
lt_of_le_of_lt := @rat.lt_of_le_of_lt,
|
||||
le_iff_lt_or_eq := @rat.le_iff_lt_or_eq,
|
||||
add_le_add_left := @rat.add_le_add_left,
|
||||
mul_nonneg := @rat.mul_nonneg,
|
||||
mul_pos := @rat.mul_pos,
|
||||
decidable_lt := @decidable_lt,
|
||||
zero_lt_one := zero_lt_one,
|
||||
add_lt_add_left := @add_lt_add_left⦄
|
||||
zero_lt_one := rat.zero_lt_one,
|
||||
add_lt_add_left := @rat.add_lt_add_left⦄
|
||||
|
||||
theorem of_nat_abs (a : ℤ) : abs (of_int a) = of_nat (int.nat_abs a) :=
|
||||
assert ∀ n : ℕ, of_int (int.neg_succ_of_nat n) = - of_nat (nat.succ n), from λ n, rfl,
|
||||
|
@ -364,7 +364,7 @@ section
|
|||
have of_int (num q) ≥ of_int 0,
|
||||
begin
|
||||
rewrite [-mul_denom],
|
||||
apply mul_nonneg H,
|
||||
apply rat.mul_nonneg H,
|
||||
rewrite [-of_int_zero, of_int_le_of_int_iff],
|
||||
exact int.le_of_lt !denom_pos
|
||||
end,
|
||||
|
@ -374,7 +374,7 @@ section
|
|||
have of_int (num q) > of_int 0,
|
||||
begin
|
||||
rewrite [-mul_denom],
|
||||
apply mul_pos H,
|
||||
apply rat.mul_pos H,
|
||||
rewrite [-of_int_zero, of_int_lt_of_int_iff],
|
||||
exact !denom_pos
|
||||
end,
|
||||
|
|
|
@ -31,7 +31,7 @@ local postfix `⁻¹` := pnat.inv
|
|||
|
||||
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, *inv_mul_eq_mul_inv]
|
||||
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 / ε)) :
|
||||
q * n⁻¹ ≤ ε :=
|
||||
|
@ -47,7 +47,7 @@ private theorem find_thirds (a b : ℚ) (H : b > 0) : ∃ n : ℕ+, a + n⁻¹ +
|
|||
let n := pceil (of_nat 4 / b) in
|
||||
have of_nat 3 * n⁻¹ < b, from calc
|
||||
of_nat 3 * n⁻¹ < of_nat 4 * n⁻¹
|
||||
: mul_lt_mul_of_pos_right dec_trivial !inv_pos
|
||||
: mul_lt_mul_of_pos_right dec_trivial !pnat.inv_pos
|
||||
... ≤ of_nat 4 * (b / of_nat 4)
|
||||
: mul_le_mul_of_nonneg_left (!inv_pceil_div dec_trivial H) !of_nat_nonneg
|
||||
... = b / of_nat 4 * of_nat 4 : algebra.mul.comm
|
||||
|
@ -73,7 +73,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.mul.left_distrib, add_sub_comm]
|
||||
by rewrite [rat.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]
|
||||
|
@ -104,8 +104,8 @@ private theorem ineq_helper (a b : ℚ) (k m n : ℕ+) (H : a ≤ (k * 2 * m)⁻
|
|||
(rat_of_pnat k) * a + b * (rat_of_pnat k) ≤ m⁻¹ + n⁻¹ :=
|
||||
assert H3 : (k * 2 * m)⁻¹ + (k * 2 * n)⁻¹ = (2 * k)⁻¹ * (m⁻¹ + n⁻¹),
|
||||
begin
|
||||
rewrite [rat.mul.left_distrib,*inv_mul_eq_mul_inv],
|
||||
rewrite (rat.mul.comm k⁻¹)
|
||||
rewrite [left_distrib, *pnat.inv_mul_eq_mul_inv],
|
||||
rewrite (mul.comm k⁻¹)
|
||||
end,
|
||||
have H' : a ≤ (2 * k)⁻¹ * (m⁻¹ + n⁻¹),
|
||||
begin
|
||||
|
@ -119,7 +119,7 @@ have H2' : b ≤ (2 * k)⁻¹ * (m⁻¹ + n⁻¹),
|
|||
end,
|
||||
have a + b ≤ k⁻¹ * (m⁻¹ + n⁻¹), from calc
|
||||
a + b ≤ (2 * k)⁻¹ * (m⁻¹ + n⁻¹) + (2 * k)⁻¹ * (m⁻¹ + n⁻¹) : algebra.add_le_add H' H2'
|
||||
... = ((2 * k)⁻¹ + (2 * k)⁻¹) * (m⁻¹ + n⁻¹) : by rewrite rat.mul.right_distrib
|
||||
... = ((2 * k)⁻¹ + (2 * k)⁻¹) * (m⁻¹ + n⁻¹) : by rewrite right_distrib
|
||||
... = k⁻¹ * (m⁻¹ + n⁻¹) : by rewrite (pnat.add_halves k),
|
||||
calc (rat_of_pnat k) * a + b * (rat_of_pnat k)
|
||||
= (rat_of_pnat k) * a + (rat_of_pnat k) * b : by rewrite (algebra.mul.comm b)
|
||||
|
@ -127,18 +127,18 @@ 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, inv_cancel_left, rat.one_mul]
|
||||
by rewrite[-rat.mul_assoc, pnat.inv_cancel_left, rat.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
|
||||
a + b + c - (d + (b + e)) = a + b + c - (d + b + e) : rat.add.assoc
|
||||
a + b + c - (d + (b + e)) = a + b + c - (d + b + e) : rat.add_assoc
|
||||
... = a + b - (d + b) + (c - e) : add_sub_comm
|
||||
... = a + b - b - d + (c - e) : sub_add_eq_sub_sub_swap
|
||||
... = a - d + (c - e) : algebra.add_sub_cancel)
|
||||
|
||||
private theorem factor_lemma_2 (a b c d : ℚ) : (a + b) + (c + d) = (a + c) + (d + b) :=
|
||||
begin
|
||||
let H := (binary.comm4 rat.add.comm rat.add.assoc a b c d),
|
||||
let H := (binary.comm4 add.comm add.assoc a b c d),
|
||||
rewrite [algebra.add.comm b d at H],
|
||||
exact H
|
||||
end
|
||||
|
@ -189,34 +189,34 @@ theorem eq_of_bdd {s t : seq} (Hs : regular s) (Ht : regular t)
|
|||
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 *],
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply abs_add_le_abs_add_abs,
|
||||
apply rat.le.trans,
|
||||
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))), rat.add_assoc],
|
||||
apply abs_add_le_abs_add_abs,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply rat.add_le_add_left,
|
||||
apply add_le_add,
|
||||
apply HNj (max j Nj) (max_right j Nj),
|
||||
apply HNj (max j Nj) (pnat.max_right j Nj),
|
||||
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⁻¹ + (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],
|
||||
rewrite hsimp,
|
||||
have Hms : (max j Nj)⁻¹ + (max j Nj)⁻¹ ≤ j⁻¹ + j⁻¹, begin
|
||||
apply add_le_add,
|
||||
apply inv_ge_of_le (max_left j Nj),
|
||||
apply inv_ge_of_le (max_left j Nj),
|
||||
apply inv_ge_of_le (pnat.max_left j Nj),
|
||||
apply inv_ge_of_le (pnat.max_left j Nj),
|
||||
end,
|
||||
apply (calc
|
||||
n⁻¹ + n⁻¹ + j⁻¹ + ((max j Nj)⁻¹ + (max j Nj)⁻¹) ≤ n⁻¹ + n⁻¹ + j⁻¹ + (j⁻¹ + j⁻¹) :
|
||||
rat.add_le_add_left Hms
|
||||
... = n⁻¹ + n⁻¹ + j⁻¹ + j⁻¹ + j⁻¹ : by rewrite *rat.add.assoc)
|
||||
... = n⁻¹ + n⁻¹ + j⁻¹ + j⁻¹ + j⁻¹ : by rewrite *rat.add_assoc)
|
||||
end,
|
||||
apply squeeze Hj
|
||||
end
|
||||
|
@ -228,7 +228,7 @@ theorem eq_of_bdd_var {s t : seq} (Hs : regular s) (Ht : regular t)
|
|||
repeat assumption,
|
||||
intros,
|
||||
apply H,
|
||||
apply inv_pos
|
||||
apply pnat.inv_pos
|
||||
end
|
||||
|
||||
theorem bdd_of_eq_var {s t : seq} (Hs : regular s) (Ht : regular t) (Heq : s ≡ t) :
|
||||
|
@ -238,7 +238,7 @@ theorem bdd_of_eq_var {s t : seq} (Hs : regular s) (Ht : regular t) (Heq : s ≡
|
|||
cases pnat_bound Hε with [N, HN],
|
||||
existsi 2 * N,
|
||||
intro n Hn,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply bdd_of_eq Heq N n Hn,
|
||||
exact HN -- assumption -- TODO: something funny here; what is 11.source.to_has_le_2?
|
||||
end
|
||||
|
@ -251,7 +251,7 @@ theorem equiv.trans (s t u : seq) (Hs : regular s) (Ht : regular t) (Hu : regula
|
|||
existsi 2 * (2 * j),
|
||||
intro n Hn,
|
||||
rewrite [-sub_add_cancel (s n) (t n), *sub_eq_add_neg, algebra.add.assoc],
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply abs_add_le_abs_add_abs,
|
||||
have Hst : abs (s n - t n) ≤ (2 * j)⁻¹, from bdd_of_eq H _ _ Hn,
|
||||
have Htu : abs (t n - u n) ≤ (2 * j)⁻¹, from bdd_of_eq H2 _ _ Hn,
|
||||
|
@ -270,10 +270,10 @@ private theorem canon_bound {s : seq} (Hs : regular s) (n : ℕ+) : abs (s n)
|
|||
abs (s n) = abs (s n - s pone + s pone) : by rewrite algebra.sub_add_cancel
|
||||
... ≤ abs (s n - s pone) + abs (s pone) : abs_add_le_abs_add_abs
|
||||
... ≤ n⁻¹ + pone⁻¹ + abs (s pone) : algebra.add_le_add_right !Hs
|
||||
... = n⁻¹ + (1 + abs (s pone)) : by rewrite [pone_inv, rat.add.assoc]
|
||||
... = n⁻¹ + (1 + abs (s pone)) : by rewrite [pone_inv, rat.add_assoc]
|
||||
... ≤ 1 + (1 + abs (s pone)) : algebra.add_le_add_right (inv_le_one n)
|
||||
... = abs (s pone) + (1 + 1) :
|
||||
by rewrite [add.comm 1 (abs (s pone)), rat.add.comm 1, rat.add.assoc]
|
||||
by rewrite [add.comm 1 (abs (s pone)), add.comm 1, rat.add_assoc]
|
||||
... ≤ of_nat (ubound (abs (s pone))) + (1 + 1) : algebra.add_le_add_right (!ubound_ge)
|
||||
... = of_nat (ubound (abs (s pone)) + (1 + 1)) : of_nat_add
|
||||
... = of_nat (ubound (abs (s pone)) + 1 + 1) : algebra.add.assoc
|
||||
|
@ -283,7 +283,7 @@ theorem bdd_of_regular {s : seq} (H : regular s) : ∃ b : ℚ, ∀ n : ℕ+, s
|
|||
begin
|
||||
existsi rat_of_pnat (K s),
|
||||
intros,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply le_abs_self,
|
||||
apply canon_bound H
|
||||
end
|
||||
|
@ -303,29 +303,29 @@ definition K₂ (s t : seq) := max (K s) (K t)
|
|||
|
||||
private theorem K₂_symm (s t : seq) : K₂ s t = K₂ t s :=
|
||||
if H : K s < K t then
|
||||
(assert H1 : K₂ s t = K t, from max_eq_right H,
|
||||
assert H2 : K₂ t s = K t, from max_eq_left (not_lt_of_ge (le_of_lt H)),
|
||||
(assert H1 : K₂ s t = K t, from pnat.max_eq_right H,
|
||||
assert H2 : K₂ t s = K t, from pnat.max_eq_left (pnat.not_lt_of_ge (pnat.le_of_lt H)),
|
||||
by rewrite [H1, -H2])
|
||||
else
|
||||
(assert H1 : K₂ s t = K s, from max_eq_left H,
|
||||
(assert H1 : K₂ s t = K s, from pnat.max_eq_left H,
|
||||
if J : K t < K s then
|
||||
(assert H2 : K₂ t s = K s, from max_eq_right J, by rewrite [H1, -H2])
|
||||
(assert H2 : K₂ t s = K s, from pnat.max_eq_right J, by rewrite [H1, -H2])
|
||||
else
|
||||
(assert Heq : K t = K s, from
|
||||
eq_of_le_of_ge (le_of_not_gt H) (le_of_not_gt J),
|
||||
pnat.eq_of_le_of_ge (pnat.le_of_not_gt H) (pnat.le_of_not_gt J),
|
||||
by rewrite [↑K₂, Heq]))
|
||||
|
||||
theorem canon_2_bound_left (s t : seq) (Hs : regular s) (n : ℕ+) :
|
||||
abs (s n) ≤ rat_of_pnat (K₂ s t) :=
|
||||
calc
|
||||
abs (s n) ≤ rat_of_pnat (K s) : canon_bound Hs n
|
||||
... ≤ rat_of_pnat (K₂ s t) : rat_of_pnat_le_of_pnat_le (!max_left)
|
||||
... ≤ rat_of_pnat (K₂ s t) : rat_of_pnat_le_of_pnat_le (!pnat.max_left)
|
||||
|
||||
theorem canon_2_bound_right (s t : seq) (Ht : regular t) (n : ℕ+) :
|
||||
abs (t n) ≤ rat_of_pnat (K₂ s t) :=
|
||||
calc
|
||||
abs (t n) ≤ rat_of_pnat (K t) : canon_bound Ht n
|
||||
... ≤ rat_of_pnat (K₂ s t) : rat_of_pnat_le_of_pnat_le (!max_right)
|
||||
... ≤ rat_of_pnat (K₂ s t) : rat_of_pnat_le_of_pnat_le (!pnat.max_right)
|
||||
|
||||
definition sadd (s t : seq) : seq := λ n, (s (2 * n)) + (t (2 * n))
|
||||
|
||||
|
@ -334,7 +334,7 @@ theorem reg_add_reg {s t : seq} (Hs : regular s) (Ht : regular t) : regular (sad
|
|||
rewrite [↑regular at *, ↑sadd],
|
||||
intros,
|
||||
rewrite add_sub_comm,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply abs_add_le_abs_add_abs,
|
||||
rewrite add_halves_double,
|
||||
apply add_le_add,
|
||||
|
@ -349,9 +349,9 @@ theorem reg_mul_reg {s t : seq} (Hs : regular s) (Ht : regular t) : regular (smu
|
|||
rewrite [↑regular at *, ↑smul],
|
||||
intros,
|
||||
rewrite rewrite_helper,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply abs_add_le_abs_add_abs,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply add_le_add,
|
||||
rewrite abs_mul,
|
||||
apply mul_le_mul_of_nonneg_right,
|
||||
|
@ -372,7 +372,7 @@ theorem reg_neg_reg {s : seq} (Hs : regular s) : regular (sneg s) :=
|
|||
begin
|
||||
rewrite [↑regular at *, ↑sneg],
|
||||
intros,
|
||||
rewrite [-abs_neg, neg_sub, sub_neg_eq_add, rat.add.comm],
|
||||
rewrite [-abs_neg, neg_sub, sub_neg_eq_add, add.comm],
|
||||
apply Hs
|
||||
end
|
||||
|
||||
|
@ -397,9 +397,9 @@ theorem s_add_assoc (s t u : seq) (Hs : regular s) (Hu : regular u) :
|
|||
rewrite [↑sadd, ↑equiv, ↑regular at *],
|
||||
intros,
|
||||
rewrite factor_lemma,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply abs_add_le_abs_add_abs,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
rotate 1,
|
||||
apply algebra.add_le_add_right,
|
||||
apply inv_two_mul_le_inv,
|
||||
|
@ -413,7 +413,7 @@ theorem s_mul_comm (s t : seq) : smul s t ≡ smul t s :=
|
|||
begin
|
||||
rewrite ↑smul,
|
||||
intros n,
|
||||
rewrite [*(K₂_symm s t), rat.mul.comm, algebra.sub_self, abs_zero],
|
||||
rewrite [*(K₂_symm s t), rat.mul_comm, algebra.sub_self, abs_zero],
|
||||
apply add_invs_nonneg
|
||||
end
|
||||
|
||||
|
@ -430,20 +430,20 @@ private theorem s_mul_assoc_lemma (s t u : seq) (a b c d : ℕ+) :
|
|||
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)),
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply abs_add_le_abs_add_abs,
|
||||
rewrite rat.add.assoc,
|
||||
rewrite rat.add_assoc,
|
||||
apply add_le_add,
|
||||
rewrite 2 abs_mul,
|
||||
apply rat.le.refl,
|
||||
rewrite [*rat.mul.assoc, -algebra.mul_sub_left_distrib, -left_distrib, abs_mul],
|
||||
apply rat.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 rat.le_trans,
|
||||
apply abs_add_le_abs_add_abs,
|
||||
apply add_le_add,
|
||||
rewrite abs_mul, apply rat.le.refl,
|
||||
rewrite [abs_mul, rat.mul.comm], apply rat.le.refl,
|
||||
rewrite abs_mul, apply rat.le_refl,
|
||||
rewrite [abs_mul, rat.mul_comm], apply rat.le_refl,
|
||||
apply abs_nonneg
|
||||
end
|
||||
|
||||
|
@ -459,10 +459,10 @@ private theorem Kq_bound {s : seq} (H : regular s) : ∀ n, abs (s n) ≤ Kq s :
|
|||
end
|
||||
|
||||
private theorem Kq_bound_nonneg {s : seq} (H : regular s) : 0 ≤ Kq s :=
|
||||
rat.le.trans !abs_nonneg (Kq_bound H 2)
|
||||
rat.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),
|
||||
have H1 : 0 ≤ rat_of_pnat (K s), from rat.le_trans (!abs_nonneg) (canon_bound H 2),
|
||||
add_pos_of_nonneg_of_pos H1 rat.zero_lt_one
|
||||
|
||||
private theorem s_mul_assoc_lemma_5 {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
||||
|
@ -518,10 +518,10 @@ theorem s_mul_assoc {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regula
|
|||
intros,
|
||||
apply exists.intro,
|
||||
intros,
|
||||
rewrite [↑smul, *DK_rewrite, *TK_rewrite, -*pnat.mul.assoc, -*rat.mul.assoc],
|
||||
apply rat.le.trans,
|
||||
rewrite [↑smul, *DK_rewrite, *TK_rewrite, -*pnat.mul_assoc, -*mul.assoc],
|
||||
apply rat.le_trans,
|
||||
apply s_mul_assoc_lemma,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply s_mul_assoc_lemma_2,
|
||||
apply Hs,
|
||||
apply Ht,
|
||||
|
@ -534,17 +534,17 @@ theorem s_mul_assoc {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regula
|
|||
apply Kq_bound_pos Ht,
|
||||
apply Kq_bound_pos Hu,
|
||||
apply add_pos,
|
||||
repeat apply inv_pos,
|
||||
repeat apply pnat.inv_pos,
|
||||
repeat apply rat.mul_pos,
|
||||
apply Kq_bound_pos Hs,
|
||||
apply Kq_bound_pos Ht,
|
||||
apply add_pos,
|
||||
repeat apply inv_pos,
|
||||
repeat apply pnat.inv_pos,
|
||||
repeat apply rat.mul_pos,
|
||||
apply Kq_bound_pos Hs,
|
||||
apply Kq_bound_pos Hu,
|
||||
apply add_pos,
|
||||
repeat apply inv_pos,
|
||||
repeat apply pnat.inv_pos,
|
||||
apply a_1
|
||||
end
|
||||
|
||||
|
@ -561,11 +561,11 @@ theorem s_zero_add (s : seq) (H : regular s) : sadd zero s ≡ s :=
|
|||
rewrite [↑sadd, ↑zero, ↑equiv, ↑regular at H],
|
||||
intros,
|
||||
rewrite [rat.zero_add],
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply H,
|
||||
apply add_le_add,
|
||||
apply inv_two_mul_le_inv,
|
||||
apply rat.le.refl
|
||||
apply rat.le_refl
|
||||
end
|
||||
|
||||
theorem s_add_zero (s : seq) (H : regular s) : sadd s zero ≡ s :=
|
||||
|
@ -573,11 +573,11 @@ theorem s_add_zero (s : seq) (H : regular s) : sadd s zero ≡ s :=
|
|||
rewrite [↑sadd, ↑zero, ↑equiv, ↑regular at H],
|
||||
intros,
|
||||
rewrite [rat.add_zero],
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply H,
|
||||
apply add_le_add,
|
||||
apply inv_two_mul_le_inv,
|
||||
apply rat.le.refl
|
||||
apply rat.le_refl
|
||||
end
|
||||
|
||||
theorem s_neg_cancel (s : seq) (H : regular s) : sadd (sneg s) s ≡ zero :=
|
||||
|
@ -604,7 +604,7 @@ theorem add_well_defined {s t u v : seq} (Hs : regular s) (Ht : regular t) (Hu :
|
|||
rewrite [↑sadd, ↑equiv at *],
|
||||
intros,
|
||||
rewrite [add_sub_comm, add_halves_double],
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply abs_add_le_abs_add_abs,
|
||||
apply add_le_add,
|
||||
apply Esu,
|
||||
|
@ -619,14 +619,14 @@ private theorem mul_bound_helper {s t : seq} (Hs : regular s) (Ht : regular t) (
|
|||
(rat_of_pnat (K t))) * (rat_of_pnat j)),
|
||||
intros n Hn,
|
||||
rewrite rewrite_helper4,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply abs_add_le_abs_add_abs,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
rotate 1,
|
||||
show n⁻¹ * ((rat_of_pnat (K s)) * (b⁻¹ + c⁻¹)) +
|
||||
n⁻¹ * ((a⁻¹ + c⁻¹) * (rat_of_pnat (K t))) ≤ j⁻¹, begin
|
||||
rewrite -left_distrib,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply mul_le_mul_of_nonneg_right,
|
||||
apply pceil_helper Hn,
|
||||
{ repeat (apply algebra.mul_pos | apply algebra.add_pos | apply rat_of_pnat_is_pos |
|
||||
|
@ -648,11 +648,11 @@ private theorem mul_bound_helper {s t : seq} (Hs : regular s) (Ht : regular t) (
|
|||
repeat (apply algebra.mul_pos | apply algebra.add_pos | apply rat_of_pnat_is_pos | apply pnat.inv_pos),
|
||||
end,
|
||||
rewrite (!div_helper H),
|
||||
apply rat.le.refl
|
||||
apply rat.le_refl
|
||||
end,
|
||||
apply add_le_add,
|
||||
rewrite [-algebra.mul_sub_left_distrib, abs_mul],
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply algebra.mul_le_mul,
|
||||
apply canon_bound,
|
||||
apply Hs,
|
||||
|
@ -660,29 +660,29 @@ private theorem mul_bound_helper {s t : seq} (Hs : regular s) (Ht : regular t) (
|
|||
apply abs_nonneg,
|
||||
apply rat.le_of_lt,
|
||||
apply rat_of_pnat_is_pos,
|
||||
rewrite [*inv_mul_eq_mul_inv, -right_distrib, -rat.mul.assoc, rat.mul.comm],
|
||||
rewrite [*pnat.inv_mul_eq_mul_inv, -right_distrib, -rat.mul_assoc, rat.mul_comm],
|
||||
apply mul_le_mul_of_nonneg_left,
|
||||
apply rat.le.refl,
|
||||
apply rat.le_refl,
|
||||
apply rat.le_of_lt,
|
||||
apply inv_pos,
|
||||
apply pnat.inv_pos,
|
||||
rewrite [-algebra.mul_sub_right_distrib, abs_mul],
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply algebra.mul_le_mul,
|
||||
apply Hs,
|
||||
apply canon_bound,
|
||||
apply Ht,
|
||||
apply abs_nonneg,
|
||||
apply add_invs_nonneg,
|
||||
rewrite [*inv_mul_eq_mul_inv, -right_distrib, mul.comm _ n⁻¹, rat.mul.assoc],
|
||||
rewrite [*pnat.inv_mul_eq_mul_inv, -right_distrib, mul.comm _ n⁻¹, rat.mul_assoc],
|
||||
apply algebra.mul_le_mul,
|
||||
repeat apply rat.le.refl,
|
||||
repeat apply rat.le_refl,
|
||||
apply rat.le_of_lt,
|
||||
apply rat.mul_pos,
|
||||
apply add_pos,
|
||||
repeat apply inv_pos,
|
||||
repeat apply pnat.inv_pos,
|
||||
apply rat_of_pnat_is_pos,
|
||||
apply rat.le_of_lt,
|
||||
apply inv_pos
|
||||
apply pnat.inv_pos
|
||||
end
|
||||
|
||||
theorem s_distrib {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u) :
|
||||
|
@ -703,17 +703,17 @@ theorem s_distrib {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular
|
|||
intros N2 HN2,
|
||||
existsi max N1 N2,
|
||||
intros n Hn,
|
||||
rewrite [↑sadd at *, ↑smul, rewrite_helper3, -add_halves j, -*pnat.mul.assoc at *],
|
||||
apply rat.le.trans,
|
||||
rewrite [↑sadd at *, ↑smul, rewrite_helper3, -add_halves j, -*pnat.mul_assoc at *],
|
||||
apply rat.le_trans,
|
||||
apply abs_add_le_abs_add_abs,
|
||||
apply add_le_add,
|
||||
apply HN1,
|
||||
apply pnat.le.trans,
|
||||
apply max_left N1 N2,
|
||||
apply pnat.le_trans,
|
||||
apply pnat.max_left N1 N2,
|
||||
apply Hn,
|
||||
apply HN2,
|
||||
apply pnat.le.trans,
|
||||
apply max_right N1 N2,
|
||||
apply pnat.le_trans,
|
||||
apply pnat.max_right N1 N2,
|
||||
apply Hn
|
||||
end
|
||||
|
||||
|
@ -730,18 +730,18 @@ 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 rat.le_trans,
|
||||
apply algebra.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)),
|
||||
apply HN',
|
||||
apply le.trans Hn,
|
||||
apply pnat.le_trans Hn,
|
||||
apply pnat.mul_le_mul_left,
|
||||
apply abs_nonneg,
|
||||
apply rat.le_of_lt (Kq_bound_pos Hs),
|
||||
apply le_of_lt (Kq_bound_pos Hs),
|
||||
rewrite (mul_div_cancel' (ne.symm (ne_of_lt (Kq_bound_pos Hs)))),
|
||||
apply rat.le.refl
|
||||
apply le.refl
|
||||
end
|
||||
|
||||
private theorem neg_bound_eq_bound (s : seq) : K (sneg s) = K s :=
|
||||
|
@ -756,7 +756,7 @@ theorem mul_neg_equiv_neg_mul {s t : seq} : smul s (sneg t) ≡ sneg (smul s t)
|
|||
begin
|
||||
rewrite [↑equiv, ↑smul],
|
||||
intros,
|
||||
rewrite [↑sneg, *sub_neg_eq_add, -neg_mul_eq_mul_neg, rat.add.comm, *sneg_def,
|
||||
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],
|
||||
apply add_invs_nonneg
|
||||
end
|
||||
|
@ -767,27 +767,27 @@ theorem equiv_of_diff_equiv_zero {s t : seq} (Hs : regular s) (Ht : regular t)
|
|||
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]
|
||||
... = b + d + a + e + c : rat.add.comm,
|
||||
... = a + (b + d) + e + c : by rewrite[-*rat.add_assoc]
|
||||
... = b + d + a + e + c : add.comm,
|
||||
apply eq_of_bdd Hs Ht,
|
||||
intros,
|
||||
let He := bdd_of_eq H,
|
||||
existsi 2 * (2 * (2 * j)),
|
||||
intros n Hn,
|
||||
rewrite (rewrite_helper5 _ _ (s (2 * n)) (t (2 * n))),
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply abs_add_three,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply add_le_add_three,
|
||||
apply Hs,
|
||||
rewrite [↑sadd at He, ↑sneg at He, ↑zero at He],
|
||||
let He' := λ a b c, eq.subst !algebra.sub_zero (He a b c),
|
||||
apply (He' _ _ Hn),
|
||||
apply Ht,
|
||||
rewrite [hsimp, add_halves, -(add_halves j), -(add_halves (2 * j)), -*rat.add.assoc],
|
||||
rewrite [hsimp, add_halves, -(add_halves j), -(add_halves (2 * j)), -*rat.add_assoc],
|
||||
apply algebra.add_le_add_right,
|
||||
apply add_le_add_three,
|
||||
repeat (apply rat.le.trans; apply inv_ge_of_le Hn; apply inv_two_mul_le_inv)
|
||||
repeat (apply rat.le_trans; apply inv_ge_of_le Hn; apply inv_two_mul_le_inv)
|
||||
end
|
||||
|
||||
theorem s_sub_cancel (s : seq) : sadd s (sneg s) ≡ zero :=
|
||||
|
@ -869,7 +869,7 @@ theorem neg_well_defined {s t : seq} (Est : s ≡ t) : sneg s ≡ sneg t :=
|
|||
begin
|
||||
rewrite [↑sneg, ↑equiv at *],
|
||||
intros,
|
||||
rewrite [-abs_neg, neg_sub, sub_neg_eq_add, rat.add.comm],
|
||||
rewrite [-abs_neg, neg_sub, sub_neg_eq_add, add.comm],
|
||||
apply Est
|
||||
end
|
||||
|
||||
|
@ -885,10 +885,10 @@ theorem s_one_mul {s : seq} (H : regular s) : smul one s ≡ s :=
|
|||
begin
|
||||
intros,
|
||||
rewrite [↑smul, ↑one, rat.one_mul],
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply H,
|
||||
apply algebra.add_le_add_right,
|
||||
apply inv_mul_le_inv
|
||||
apply pnat.inv_mul_le_inv
|
||||
end
|
||||
|
||||
theorem s_mul_one {s : seq} (H : regular s) : smul s one ≡ s :=
|
||||
|
@ -913,7 +913,7 @@ theorem zero_nequiv_one : ¬ zero ≡ one :=
|
|||
... = abs 1 : abs_of_pos zero_lt_one
|
||||
... ≤ 2⁻¹ : H,
|
||||
let H'' := ge_of_inv_le H',
|
||||
apply absurd (one_lt_two) (not_lt_of_ge H'')
|
||||
apply absurd (one_lt_two) (pnat.not_lt_of_ge H'')
|
||||
end
|
||||
|
||||
---------------------------------------------
|
||||
|
|
|
@ -29,7 +29,7 @@ theorem rat_approx {s : seq} (H : regular s) :
|
|||
existsi (s (2 * n)),
|
||||
existsi 2 * n,
|
||||
intro m Hm,
|
||||
apply rat.le.trans,
|
||||
apply le.trans,
|
||||
apply H,
|
||||
rewrite -(add_halves n),
|
||||
apply algebra.add_le_add_right,
|
||||
|
@ -51,19 +51,19 @@ theorem rat_approx_seq {s : seq} (H : regular s) :
|
|||
existsi N,
|
||||
intro p Hp,
|
||||
rewrite ↑[sadd, sneg, s_abs, const],
|
||||
apply rat.le.trans,
|
||||
apply le.trans,
|
||||
rotate 1,
|
||||
rewrite -sub_eq_add_neg,
|
||||
apply algebra.sub_le_sub_left,
|
||||
apply HN,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
apply Hp,
|
||||
rewrite -*pnat.mul.assoc,
|
||||
rewrite -*pnat.mul_assoc,
|
||||
apply pnat.mul_le_mul_left,
|
||||
rewrite [algebra.sub_self, -neg_zero],
|
||||
apply neg_le_neg,
|
||||
apply rat.le_of_lt,
|
||||
apply inv_pos
|
||||
apply pnat.inv_pos
|
||||
end
|
||||
|
||||
theorem r_rat_approx (s : reg_seq) :
|
||||
|
@ -77,10 +77,10 @@ theorem const_bound {s : seq} (Hs : regular s) (n : ℕ+) :
|
|||
intro m,
|
||||
rewrite -sub_eq_add_neg,
|
||||
apply iff.mp !le_add_iff_neg_le_sub_left,
|
||||
apply rat.le.trans,
|
||||
apply le.trans,
|
||||
apply Hs,
|
||||
apply algebra.add_le_add_right,
|
||||
rewrite -*pnat.mul.assoc,
|
||||
rewrite -*pnat.mul_assoc,
|
||||
apply inv_ge_of_le,
|
||||
apply pnat.mul_le_mul_left
|
||||
end
|
||||
|
@ -103,19 +103,19 @@ theorem equiv_abs_of_ge_zero {s : seq} (Hs : regular s) (Hz : s_le zero s) : s_a
|
|||
cases em (s n ≥ 0) with [Hpos, Hneg],
|
||||
rewrite [abs_of_nonneg Hpos, algebra.sub_self, abs_zero],
|
||||
apply rat.le_of_lt,
|
||||
apply inv_pos,
|
||||
apply pnat.inv_pos,
|
||||
let Hneg' := lt_of_not_ge Hneg,
|
||||
have Hsn : -s n - s n > 0, from add_pos (neg_pos_of_neg Hneg') (neg_pos_of_neg Hneg'),
|
||||
rewrite [abs_of_neg Hneg', abs_of_pos Hsn],
|
||||
apply rat.le.trans,
|
||||
apply le.trans,
|
||||
apply add_le_add,
|
||||
repeat (apply neg_le_neg; apply Hz'),
|
||||
rewrite algebra.neg_neg,
|
||||
apply rat.le.trans,
|
||||
apply le.trans,
|
||||
apply add_le_add,
|
||||
repeat (apply inv_ge_of_le; apply Hn),
|
||||
krewrite pnat.add_halves,
|
||||
apply rat.le.refl
|
||||
apply le.refl
|
||||
end
|
||||
|
||||
theorem equiv_neg_abs_of_le_zero {s : seq} (Hs : regular s) (Hz : s_le s zero) : s_abs s ≡ sneg s :=
|
||||
|
@ -139,19 +139,19 @@ theorem equiv_neg_abs_of_le_zero {s : seq} (Hs : regular s) (Hz : s_le s zero) :
|
|||
have Hsn : s n + s n ≥ 0, from add_nonneg Hpos Hpos,
|
||||
rewrite [abs_of_nonneg Hpos, ↑sneg, sub_neg_eq_add, abs_of_nonneg Hsn],
|
||||
rewrite [↑nonneg at Hz', ↑sneg at Hz'],
|
||||
apply rat.le.trans,
|
||||
apply le.trans,
|
||||
apply add_le_add,
|
||||
repeat apply (le_of_neg_le_neg !Hz'),
|
||||
apply rat.le.trans,
|
||||
apply le.trans,
|
||||
apply add_le_add,
|
||||
repeat (apply inv_ge_of_le; apply Hn),
|
||||
krewrite pnat.add_halves,
|
||||
apply rat.le.refl,
|
||||
apply le.refl,
|
||||
let Hneg' := lt_of_not_ge Hneg,
|
||||
rewrite [abs_of_neg Hneg', ↑sneg, sub_neg_eq_add, neg_add_eq_sub, algebra.sub_self,
|
||||
abs_zero],
|
||||
apply rat.le_of_lt,
|
||||
apply inv_pos
|
||||
apply pnat.inv_pos
|
||||
end
|
||||
|
||||
theorem r_equiv_abs_of_ge_zero {s : reg_seq} (Hz : r_le r_zero s) : requiv (r_abs s) s :=
|
||||
|
@ -237,14 +237,14 @@ theorem cauchy_with_rate_of_converges_to_with_rate {X : r_seq} {a : ℝ} {N :
|
|||
xrewrite -of_rat_add,
|
||||
apply of_rat_le_of_rat_of_le,
|
||||
krewrite pnat.add_halves,
|
||||
apply rat.le.refl
|
||||
apply rat.le_refl
|
||||
end
|
||||
|
||||
private definition Nb (M : ℕ+ → ℕ+) := λ k, pnat.max (3 * k) (M (2 * k))
|
||||
|
||||
private theorem Nb_spec_right (M : ℕ+ → ℕ+) (k : ℕ+) : M (2 * k) ≤ Nb M k := !max_right
|
||||
private theorem Nb_spec_right (M : ℕ+ → ℕ+) (k : ℕ+) : M (2 * k) ≤ Nb M k := !pnat.max_right
|
||||
|
||||
private theorem Nb_spec_left (M : ℕ+ → ℕ+) (k : ℕ+) : 3 * k ≤ Nb M k := !max_left
|
||||
private theorem Nb_spec_left (M : ℕ+ → ℕ+) (k : ℕ+) : 3 * k ≤ Nb M k := !pnat.max_left
|
||||
|
||||
section lim_seq
|
||||
parameter {X : r_seq}
|
||||
|
@ -269,7 +269,7 @@ private theorem lim_seq_reg_helper {m n : ℕ+} (Hmn : M (2 * n) ≤M (2 * m)) :
|
|||
rotate 1,
|
||||
apply Nb_spec_right,
|
||||
rotate 1,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
apply Hmn,
|
||||
apply Nb_spec_right,
|
||||
krewrite [-+of_rat_add],
|
||||
|
@ -294,7 +294,7 @@ theorem lim_seq_reg : rat_seq.regular lim_seq :=
|
|||
apply lim_seq_reg_helper Hor1,
|
||||
let Hor2' := pnat.le_of_lt (pnat.lt_of_not_le Hor2),
|
||||
krewrite [abs_sub (X (Nb M n)), abs_sub (X (Nb M m)), abs_sub,
|
||||
rat.add.comm, add_comm_three],
|
||||
rat.add_comm, add_comm_three],
|
||||
apply lim_seq_reg_helper Hor2'
|
||||
end
|
||||
|
||||
|
@ -334,18 +334,18 @@ theorem converges_to_with_rate_of_cauchy_with_rate : converges_to_with_rate X li
|
|||
apply algebra.le.trans,
|
||||
apply add_le_add_three,
|
||||
apply Hc,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
rotate 1,
|
||||
apply Hn,
|
||||
rotate_right 1,
|
||||
apply Nb_spec_right,
|
||||
have HMk : M (2 * k) ≤ Nb M n, begin
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
apply Nb_spec_right,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
apply Hn,
|
||||
apply pnat.le.trans,
|
||||
apply mul_le_mul_left 3,
|
||||
apply pnat.le_trans,
|
||||
apply pnat.mul_le_mul_left 3,
|
||||
apply Nb_spec_left
|
||||
end,
|
||||
apply HMk,
|
||||
|
@ -357,23 +357,23 @@ theorem converges_to_with_rate_of_cauchy_with_rate : converges_to_with_rate X li
|
|||
apply of_rat_le_of_rat_of_le,
|
||||
apply algebra.le.trans,
|
||||
apply add_le_add_three,
|
||||
apply rat.le.refl,
|
||||
apply rat.le_refl,
|
||||
apply inv_ge_of_le,
|
||||
apply pnat_mul_le_mul_left',
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
rotate 1,
|
||||
apply Hn,
|
||||
rotate_right 1,
|
||||
apply Nb_spec_left,
|
||||
apply inv_ge_of_le,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
rotate 1,
|
||||
apply Hn,
|
||||
rotate_right 1,
|
||||
apply Nb_spec_left,
|
||||
rewrite -*pnat.mul.assoc,
|
||||
rewrite -*pnat.mul_assoc,
|
||||
krewrite pnat.p_add_fractions,
|
||||
apply rat.le.refl
|
||||
apply rat.le_refl
|
||||
end
|
||||
|
||||
end lim_seq
|
||||
|
@ -394,7 +394,7 @@ theorem archimedean_upper (x : ℝ) : ∃ z : ℤ, x ≤ of_int z :=
|
|||
apply rat_seq.s_le_of_le_pointwise (rat_seq.reg_seq.is_reg s),
|
||||
apply rat_seq.const_reg,
|
||||
intro n,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply Hb,
|
||||
apply ubound_ge
|
||||
end,
|
||||
|
@ -602,7 +602,7 @@ private noncomputable definition over_seq := λ n : ℕ, pr2 (iterate bisect n (
|
|||
private noncomputable definition avg_seq := λ n : ℕ, avg (over_seq n) (under_seq n) -- C
|
||||
|
||||
private theorem avg_symm (n : ℕ) : avg_seq n = avg (under_seq n) (over_seq n) :=
|
||||
by rewrite [↑avg_seq, ↑avg, rat.add.comm]
|
||||
by rewrite [↑avg_seq, ↑avg, add.comm]
|
||||
|
||||
private theorem over_0 : over_seq 0 = over := rfl
|
||||
|
||||
|
@ -758,7 +758,7 @@ private theorem under_seq'_lt_over_seq'_single : ∀ n : ℕ, under_seq' n < ove
|
|||
|
||||
private theorem under_seq_mono_helper (i k : ℕ) : under_seq i ≤ under_seq (i + k) :=
|
||||
(nat.induction_on k
|
||||
(by rewrite nat.add_zero; apply rat.le.refl)
|
||||
(by rewrite nat.add_zero; apply rat.le_refl)
|
||||
(begin
|
||||
intros a Ha,
|
||||
rewrite [add_succ, under_succ],
|
||||
|
@ -766,7 +766,7 @@ private theorem under_seq_mono_helper (i k : ℕ) : under_seq i ≤ under_seq (i
|
|||
rewrite (if_pos Havg),
|
||||
apply Ha,
|
||||
rewrite [if_neg Havg, ↑avg_seq, ↑avg],
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply Ha,
|
||||
rewrite -add_halves at {1},
|
||||
apply add_le_add_right,
|
||||
|
@ -785,13 +785,13 @@ private theorem under_seq_mono (i j : ℕ) (H : i ≤ j) : under_seq i ≤ under
|
|||
|
||||
private theorem over_seq_mono_helper (i k : ℕ) : over_seq (i + k) ≤ over_seq i :=
|
||||
nat.induction_on k
|
||||
(by rewrite nat.add_zero; apply rat.le.refl)
|
||||
(by rewrite nat.add_zero; apply rat.le_refl)
|
||||
(begin
|
||||
intros a Ha,
|
||||
rewrite [add_succ, over_succ],
|
||||
cases em (ub (avg_seq (i + a))) with [Havg, Havg],
|
||||
rewrite [if_pos Havg, ↑avg_seq, ↑avg],
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
rotate 1,
|
||||
apply Ha,
|
||||
rotate 1,
|
||||
|
@ -821,18 +821,18 @@ private theorem regular_lemma_helper {s : seq} {m n : ℕ+} (Hm : m ≤ n)
|
|||
abs (s m - s n) ≤ m⁻¹ + n⁻¹ :=
|
||||
begin
|
||||
cases H m n Hm with [T1under, T1over],
|
||||
cases H m m (!pnat.le.refl) with [T2under, T2over],
|
||||
apply rat.le.trans,
|
||||
cases H m m (!pnat.le_refl) with [T2under, T2over],
|
||||
apply rat.le_trans,
|
||||
apply dist_bdd_within_interval,
|
||||
apply under_seq'_lt_over_seq'_single,
|
||||
rotate 1,
|
||||
repeat assumption,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply width',
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply rat_power_two_inv_ge,
|
||||
apply le_add_of_nonneg_right,
|
||||
apply rat.le_of_lt (!inv_pos)
|
||||
apply rat.le_of_lt (!pnat.inv_pos)
|
||||
end
|
||||
|
||||
private theorem regular_lemma (s : seq) (H : ∀ n i : ℕ+, i ≥ n → under_seq' n~ ≤ s i ∧ s i ≤ over_seq' n~) :
|
||||
|
@ -843,7 +843,7 @@ private theorem regular_lemma (s : seq) (H : ∀ n i : ℕ+, i ≥ n → under_s
|
|||
cases em (m ≤ n) with [Hm, Hn],
|
||||
apply regular_lemma_helper Hm H,
|
||||
let T := regular_lemma_helper (pnat.le_of_lt (pnat.lt_of_not_le Hn)) H,
|
||||
rewrite [abs_sub at T, {n⁻¹ + _}rat.add.comm at T],
|
||||
rewrite [abs_sub at T, {n⁻¹ + _}add.comm at T],
|
||||
exact T
|
||||
end
|
||||
|
||||
|
@ -906,14 +906,14 @@ private theorem under_lowest_bound : ∀ y : ℝ, ub y → sup_under ≤ y :=
|
|||
private theorem under_over_equiv : p_under_seq ≡ p_over_seq :=
|
||||
begin
|
||||
intros,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
have H : p_under_seq n < p_over_seq n, from !under_seq_lt_over_seq,
|
||||
rewrite [abs_of_neg (iff.mpr !sub_neg_iff_lt H), neg_sub],
|
||||
apply width',
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply rat_power_two_inv_ge,
|
||||
apply le_add_of_nonneg_left,
|
||||
apply rat.le_of_lt !inv_pos
|
||||
apply rat.le_of_lt !pnat.inv_pos
|
||||
end
|
||||
|
||||
private theorem under_over_eq : sup_under = sup_over := quot.sound under_over_equiv
|
||||
|
|
|
@ -46,7 +46,7 @@ theorem abs_pos_of_nonzero {s : seq} (Hs : regular s) (Hnz : sep s zero) :
|
|||
cases bdd_away_of_pos (reg_neg_reg Hs) H' with [N, HN],
|
||||
existsi N,
|
||||
intro m Hm,
|
||||
apply rat.le.trans,
|
||||
apply le.trans,
|
||||
apply HN m Hm,
|
||||
rewrite ↑sneg,
|
||||
apply neg_le_abs_self,
|
||||
|
@ -56,7 +56,7 @@ theorem abs_pos_of_nonzero {s : seq} (Hs : regular s) (Hnz : sep s zero) :
|
|||
cases H'' with [N, HN],
|
||||
existsi N,
|
||||
intro m Hm,
|
||||
apply rat.le.trans,
|
||||
apply le.trans,
|
||||
apply HN m Hm,
|
||||
apply le_abs_self
|
||||
end
|
||||
|
@ -67,7 +67,7 @@ theorem abs_well_defined {s t : seq} (Hs : regular s) (Ht : regular t) (Heq : s
|
|||
rewrite [↑equiv at *],
|
||||
intro n,
|
||||
rewrite ↑s_abs,
|
||||
apply rat.le.trans,
|
||||
apply le.trans,
|
||||
apply abs_abs_sub_abs_le_abs_sub,
|
||||
apply Heq
|
||||
end
|
||||
|
@ -144,10 +144,10 @@ private theorem le_ps {s : seq} (Hs : regular s) (Hsep : sep s zero) (n : ℕ+)
|
|||
end)
|
||||
else
|
||||
(begin
|
||||
rewrite [(s_inv_of_sep_gt_p Hs Hsep (le_of_not_gt Hn)), abs_one_div],
|
||||
rewrite [(s_inv_of_sep_gt_p Hs Hsep (pnat.le_of_not_gt Hn)), abs_one_div],
|
||||
apply div_le_pnat,
|
||||
apply ps_spec,
|
||||
rewrite pnat.mul.assoc,
|
||||
rewrite pnat.mul_assoc,
|
||||
apply pnat.mul_le_mul_right
|
||||
end)
|
||||
|
||||
|
@ -173,7 +173,7 @@ private theorem s_ne_zero_of_ge_p {s : seq} (Hs : regular s) (Hsep : sep s zero)
|
|||
apply gt_of_ge_of_gt,
|
||||
apply Hps,
|
||||
apply Hn,
|
||||
apply inv_pos
|
||||
apply pnat.inv_pos
|
||||
end
|
||||
|
||||
theorem reg_inv_reg {s : seq} (Hs : regular s) (Hsep : sep s zero) : regular (s_inv Hs) :=
|
||||
|
@ -181,49 +181,49 @@ theorem reg_inv_reg {s : seq} (Hs : regular s) (Hsep : sep s zero) : regular (s_
|
|||
rewrite ↑regular,
|
||||
intros,
|
||||
have Hsp : s ((ps Hs Hsep) * (ps Hs Hsep) * (ps Hs Hsep)) ≠ 0, from
|
||||
s_ne_zero_of_ge_p Hs Hsep !mul_le_mul_left,
|
||||
s_ne_zero_of_ge_p Hs Hsep !pnat.mul_le_mul_left,
|
||||
have Hspn : s ((ps Hs Hsep) * (ps Hs Hsep) * n) ≠ 0, from
|
||||
s_ne_zero_of_ge_p Hs Hsep (show (ps Hs Hsep) * (ps Hs Hsep) * n ≥ ps Hs Hsep, by
|
||||
rewrite pnat.mul.assoc; apply pnat.mul_le_mul_right),
|
||||
rewrite pnat.mul_assoc; apply pnat.mul_le_mul_right),
|
||||
have Hspm : s ((ps Hs Hsep) * (ps Hs Hsep) * m) ≠ 0, from
|
||||
s_ne_zero_of_ge_p Hs Hsep (show (ps Hs Hsep) * (ps Hs Hsep) * m ≥ ps Hs Hsep, by
|
||||
rewrite pnat.mul.assoc; apply pnat.mul_le_mul_right),
|
||||
rewrite pnat.mul_assoc; apply pnat.mul_le_mul_right),
|
||||
cases em (m < ps Hs Hsep) with [Hmlt, Hmlt],
|
||||
cases em (n < ps Hs Hsep) with [Hnlt, Hnlt],
|
||||
rewrite [(s_inv_of_sep_lt_p Hs Hsep Hmlt), (s_inv_of_sep_lt_p Hs Hsep Hnlt)],
|
||||
rewrite [algebra.sub_self, abs_zero],
|
||||
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))],
|
||||
(s_inv_of_sep_gt_p Hs Hsep (pnat.le_of_not_gt Hnlt))],
|
||||
rewrite [(!div_sub_div Hsp Hspn), div_eq_mul_one_div, *abs_mul, *mul_one, *one_mul],
|
||||
apply rat.le.trans,
|
||||
apply algebra.mul_le_mul,
|
||||
apply le.trans,
|
||||
apply mul_le_mul,
|
||||
apply Hs,
|
||||
rewrite [-(mul_one 1), -(!field.div_mul_div Hsp Hspn), abs_mul],
|
||||
apply algebra.mul_le_mul,
|
||||
rewrite -(s_inv_of_sep_lt_p Hs Hsep Hmlt),
|
||||
apply le_ps Hs Hsep,
|
||||
rewrite -(s_inv_of_sep_gt_p Hs Hsep (le_of_not_gt Hnlt)),
|
||||
rewrite -(s_inv_of_sep_gt_p Hs Hsep (pnat.le_of_not_gt Hnlt)),
|
||||
apply le_ps Hs Hsep,
|
||||
apply abs_nonneg,
|
||||
apply le_of_lt !rat_of_pnat_is_pos,
|
||||
apply abs_nonneg,
|
||||
apply add_invs_nonneg,
|
||||
rewrite [right_distrib, *pnat_cancel', rat.add.comm],
|
||||
rewrite [right_distrib, *pnat_cancel', add.comm],
|
||||
apply algebra.add_le_add_right,
|
||||
apply inv_ge_of_le,
|
||||
apply pnat.le_of_lt,
|
||||
apply Hmlt,
|
||||
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))],
|
||||
(s_inv_of_sep_gt_p Hs Hsep (pnat.le_of_not_gt Hmlt))],
|
||||
rewrite [(!div_sub_div Hspm Hsp), div_eq_mul_one_div, *abs_mul, *mul_one, *one_mul],
|
||||
apply rat.le.trans,
|
||||
apply le.trans,
|
||||
apply algebra.mul_le_mul,
|
||||
apply Hs,
|
||||
rewrite [-(mul_one 1), -(!field.div_mul_div Hspm Hsp), abs_mul],
|
||||
apply algebra.mul_le_mul,
|
||||
rewrite -(s_inv_of_sep_gt_p Hs Hsep (le_of_not_gt Hmlt)),
|
||||
rewrite -(s_inv_of_sep_gt_p Hs Hsep (pnat.le_of_not_gt Hmlt)),
|
||||
apply le_ps Hs Hsep,
|
||||
rewrite -(s_inv_of_sep_lt_p Hs Hsep Hnlt),
|
||||
apply le_ps Hs Hsep,
|
||||
|
@ -231,29 +231,29 @@ theorem reg_inv_reg {s : seq} (Hs : regular s) (Hsep : sep s zero) : regular (s_
|
|||
apply le_of_lt !rat_of_pnat_is_pos,
|
||||
apply abs_nonneg,
|
||||
apply add_invs_nonneg,
|
||||
rewrite [right_distrib, *pnat_cancel', rat.add.comm],
|
||||
rewrite [right_distrib, *pnat_cancel', add.comm],
|
||||
apply rat.add_le_add_left,
|
||||
apply inv_ge_of_le,
|
||||
apply pnat.le_of_lt,
|
||||
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 [(s_inv_of_sep_gt_p Hs Hsep (pnat.le_of_not_gt Hnlt)),
|
||||
(s_inv_of_sep_gt_p Hs Hsep (pnat.le_of_not_gt Hmlt))],
|
||||
rewrite [(!div_sub_div Hspm Hspn), div_eq_mul_one_div, abs_mul, *one_mul, *mul_one],
|
||||
apply rat.le.trans,
|
||||
apply le.trans,
|
||||
apply algebra.mul_le_mul,
|
||||
apply Hs,
|
||||
rewrite [-(mul_one 1), -(!field.div_mul_div Hspm Hspn), abs_mul],
|
||||
apply algebra.mul_le_mul,
|
||||
rewrite -(s_inv_of_sep_gt_p Hs Hsep (le_of_not_gt Hmlt)),
|
||||
rewrite -(s_inv_of_sep_gt_p Hs Hsep (pnat.le_of_not_gt Hmlt)),
|
||||
apply le_ps Hs Hsep,
|
||||
rewrite -(s_inv_of_sep_gt_p Hs Hsep (le_of_not_gt Hnlt)),
|
||||
rewrite -(s_inv_of_sep_gt_p Hs Hsep (pnat.le_of_not_gt Hnlt)),
|
||||
apply le_ps Hs Hsep,
|
||||
apply abs_nonneg,
|
||||
apply le_of_lt !rat_of_pnat_is_pos,
|
||||
apply abs_nonneg,
|
||||
apply add_invs_nonneg,
|
||||
rewrite [right_distrib, *pnat_cancel', rat.add.comm],
|
||||
apply rat.le.refl
|
||||
rewrite [right_distrib, *pnat_cancel', add.comm],
|
||||
apply algebra.le.refl
|
||||
end
|
||||
|
||||
theorem s_inv_ne_zero {s : seq} (Hs : regular s) (Hsep : sep s zero) (n : ℕ+) : s_inv Hs n ≠ 0 :=
|
||||
|
@ -262,7 +262,7 @@ theorem s_inv_ne_zero {s : seq} (Hs : regular s) (Hsep : sep s zero) (n : ℕ+)
|
|||
rewrite (s_inv_of_sep_gt_p Hs Hsep H),
|
||||
apply one_div_ne_zero,
|
||||
apply s_ne_zero_of_ge_p,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
apply H,
|
||||
apply pnat.mul_le_mul_left
|
||||
end)
|
||||
|
@ -283,48 +283,48 @@ theorem mul_inv {s : seq} (Hs : regular s) (Hsep : sep s zero) : smul s (s_inv H
|
|||
existsi max (ps Hs Hsep) j,
|
||||
intro n Hn,
|
||||
have Hnz : s_inv Hs ((K₂ s (s_inv Hs)) * 2 * n) ≠ 0, from s_inv_ne_zero Hs Hsep _,
|
||||
rewrite [↑smul, ↑one, rat.mul.comm, -(mul_one_div_cancel Hnz),
|
||||
rewrite [↑smul, ↑one, mul.comm, -(mul_one_div_cancel Hnz),
|
||||
-algebra.mul_sub_left_distrib, abs_mul],
|
||||
apply rat.le.trans,
|
||||
apply le.trans,
|
||||
apply mul_le_mul_of_nonneg_right,
|
||||
apply canon_2_bound_right s,
|
||||
apply Rsi,
|
||||
apply abs_nonneg,
|
||||
have Hp : (K₂ s (s_inv Hs)) * 2 * n ≥ ps Hs Hsep, begin
|
||||
apply pnat.le.trans,
|
||||
apply max_left,
|
||||
apply pnat.le_trans,
|
||||
apply pnat.max_left,
|
||||
rotate 1,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
apply Hn,
|
||||
apply pnat.mul_le_mul_left
|
||||
end,
|
||||
have Hnz' : s (((ps Hs Hsep) * (ps Hs Hsep)) * ((K₂ s (s_inv Hs)) * 2 * n)) ≠ 0, from
|
||||
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),
|
||||
by rewrite *pnat.mul_assoc; apply pnat.mul_le_mul_right),
|
||||
rewrite [(s_inv_of_sep_gt_p Hs Hsep Hp), (division_ring.one_div_one_div Hnz')],
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply mul_le_mul_of_nonneg_left,
|
||||
apply Hs,
|
||||
apply le_of_lt,
|
||||
apply rat_of_pnat_is_pos,
|
||||
rewrite [rat.mul.left_distrib, mul.comm ((ps Hs Hsep) * (ps Hs Hsep)), *pnat.mul.assoc,
|
||||
*(@inv_mul_eq_mul_inv (K₂ s (s_inv Hs))), -*rat.mul.assoc, *inv_cancel_left,
|
||||
rewrite [left_distrib, pnat.mul_comm ((ps Hs Hsep) * (ps Hs Hsep)), *pnat.mul_assoc,
|
||||
*(@pnat.inv_mul_eq_mul_inv (K₂ s (s_inv Hs))), -*mul.assoc, *pnat.inv_cancel_left,
|
||||
*one_mul, -(add_halves j)],
|
||||
apply add_le_add,
|
||||
apply inv_ge_of_le,
|
||||
apply pnat_mul_le_mul_left',
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
rotate 1,
|
||||
apply Hn,
|
||||
rotate_right 1,
|
||||
apply max_right,
|
||||
apply pnat.max_right,
|
||||
apply inv_ge_of_le,
|
||||
apply pnat_mul_le_mul_left',
|
||||
apply pnat.le.trans,
|
||||
apply max_right,
|
||||
apply pnat.le_trans,
|
||||
apply pnat.max_right,
|
||||
rotate 1,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
apply Hn,
|
||||
apply pnat.mul_le_mul_right
|
||||
end
|
||||
|
|
|
@ -36,7 +36,7 @@ theorem bdd_away_of_pos {s : seq} (Hs : regular s) (H : pos s) :
|
|||
rewrite [sub_neg_eq_add, add.comm, -add.assoc],
|
||||
apply le_of_lt HN
|
||||
end,
|
||||
rewrite rat.add.comm at Habs',
|
||||
rewrite add.comm at Habs',
|
||||
have Hin : s m ≥ N⁻¹, from calc
|
||||
s m ≥ s n - abs (s m - s n) : (iff.mp (le_add_iff_sub_left_le _ _ _)) Habs'
|
||||
... ≥ s n - (m⁻¹ + n⁻¹) : algebra.sub_le_sub_left !Hs
|
||||
|
@ -91,11 +91,11 @@ theorem nonneg_of_bdd_within {s : seq} (Hs : regular s)
|
|||
rotate 1,
|
||||
apply add_le_add,
|
||||
rotate 1,
|
||||
apply HN (max (pceil ((2)/ε)) N) !max_right,
|
||||
apply HN (max (pceil ((2)/ε)) N) !pnat.max_right,
|
||||
rotate_right 1,
|
||||
apply neg_le_neg,
|
||||
apply inv_ge_of_le,
|
||||
apply max_left,
|
||||
apply pnat.max_left,
|
||||
rewrite -neg_add,
|
||||
apply neg_le_neg,
|
||||
apply le.trans,
|
||||
|
@ -105,7 +105,7 @@ theorem nonneg_of_bdd_within {s : seq} (Hs : regular s)
|
|||
repeat apply zero_lt_one;
|
||||
exact Hε),
|
||||
rewrite [algebra.add_halves],
|
||||
apply le.refl
|
||||
apply rat.le_refl
|
||||
end
|
||||
|
||||
theorem pos_of_pos_equiv {s t : seq} (Hs : regular s) (Heq : s ≡ t) (Hp : pos s) : pos t :=
|
||||
|
@ -116,13 +116,13 @@ theorem pos_of_pos_equiv {s t : seq} (Hs : regular s) (Heq : s ≡ t) (Hp : pos
|
|||
rotate 1,
|
||||
apply sub_le_of_abs_sub_le_right,
|
||||
apply Heq,
|
||||
have Hs4 : N⁻¹ ≤ s (2 * 2 * N), from HN _ (!mul_le_mul_left),
|
||||
have Hs4 : N⁻¹ ≤ s (2 * 2 * N), from HN _ (!pnat.mul_le_mul_left),
|
||||
apply lt_of_lt_of_le,
|
||||
rotate 1,
|
||||
rewrite sub_eq_add_neg,
|
||||
apply iff.mpr !add_le_add_right_iff,
|
||||
apply Hs4,
|
||||
rewrite [*pnat.mul.assoc, pnat.add_halves, -(add_halves N), -sub_eq_add_neg, algebra.add_sub_cancel],
|
||||
rewrite [*pnat.mul_assoc, pnat.add_halves, -(add_halves N), -sub_eq_add_neg, algebra.add_sub_cancel],
|
||||
apply inv_two_mul_lt_inv
|
||||
end
|
||||
|
||||
|
@ -143,24 +143,24 @@ theorem nonneg_of_nonneg_equiv {s t : seq} (Hs : regular s) (Ht : regular t) (He
|
|||
rotate 1,
|
||||
apply algebra.sub_le_sub_right,
|
||||
apply HNs,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
rotate 1,
|
||||
apply Hm,
|
||||
rotate_right 1,
|
||||
apply max_left,
|
||||
apply pnat.max_left,
|
||||
have Hms : m⁻¹ ≤ (2 * 2 * n)⁻¹, begin
|
||||
apply inv_ge_of_le,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
rotate 1,
|
||||
apply Hm;
|
||||
apply max_right
|
||||
apply pnat.max_right
|
||||
end,
|
||||
have Hms' : m⁻¹ + m⁻¹ ≤ (2 * 2 * n)⁻¹ + (2 * 2 * n)⁻¹, from add_le_add Hms Hms,
|
||||
apply le.trans,
|
||||
rotate 1,
|
||||
apply algebra.sub_le_sub_left,
|
||||
apply Hms',
|
||||
rewrite [*pnat.mul.assoc, pnat.add_halves, -neg_sub, -add_halves n, sub_neg_eq_add],
|
||||
rewrite [*pnat.mul_assoc, pnat.add_halves, -neg_sub, -add_halves n, sub_neg_eq_add],
|
||||
apply neg_le_neg,
|
||||
apply algebra.add_le_add_left,
|
||||
apply inv_two_mul_le_inv
|
||||
|
@ -173,8 +173,8 @@ theorem zero_nonneg : nonneg zero :=
|
|||
begin
|
||||
intros,
|
||||
apply neg_nonpos_of_nonneg,
|
||||
apply le_of_lt,
|
||||
apply inv_pos
|
||||
apply rat.le_of_lt,
|
||||
apply pnat.inv_pos
|
||||
end
|
||||
|
||||
theorem s_zero_lt_one : s_lt zero one :=
|
||||
|
@ -185,7 +185,7 @@ theorem s_zero_lt_one : s_lt zero one :=
|
|||
apply one_lt_two
|
||||
end
|
||||
|
||||
theorem le.refl {s : seq} (Hs : regular s) : s_le s s :=
|
||||
protected theorem le_refl {s : seq} (Hs : regular s) : s_le s s :=
|
||||
begin
|
||||
apply nonneg_of_nonneg_equiv,
|
||||
rotate 2,
|
||||
|
@ -210,12 +210,12 @@ theorem s_nonneg_of_pos {s : seq} (Hs : regular s) (H : pos s) : nonneg s :=
|
|||
apply Hm,
|
||||
apply le.trans,
|
||||
rotate 1,
|
||||
apply le_of_lt,
|
||||
apply inv_pos,
|
||||
apply rat.le_of_lt,
|
||||
apply pnat.inv_pos,
|
||||
rewrite -neg_zero,
|
||||
apply neg_le_neg,
|
||||
apply le_of_lt,
|
||||
apply inv_pos
|
||||
apply rat.le_of_lt,
|
||||
apply pnat.inv_pos
|
||||
end
|
||||
|
||||
theorem s_le_of_s_lt {s t : seq} (Hs : regular s) (Ht : regular t) (H : s_lt s t) : s_le s t :=
|
||||
|
@ -303,7 +303,7 @@ theorem add_nonneg_of_nonneg {s t : seq} (Hs : nonneg s) (Ht : nonneg t) : nonne
|
|||
apply Ht
|
||||
end
|
||||
|
||||
theorem le.trans {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u) (Lst : s_le s t)
|
||||
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 *,
|
||||
|
@ -358,16 +358,16 @@ theorem equiv_of_le_of_ge {s t : seq} (Hs : regular s) (Ht : regular t) (Lst : s
|
|||
apply abs_le_of_le_of_neg_le,
|
||||
apply le_of_neg_le_neg,
|
||||
rewrite [2 neg_add, neg_neg],
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply algebra.neg_add_neg_le_neg_of_pos,
|
||||
apply inv_pos,
|
||||
apply pnat.inv_pos,
|
||||
rewrite add.comm,
|
||||
apply Lst,
|
||||
apply le_of_neg_le_neg,
|
||||
rewrite [neg_add, neg_neg],
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply algebra.neg_add_neg_le_neg_of_pos,
|
||||
apply inv_pos,
|
||||
apply pnat.inv_pos,
|
||||
apply Lts,
|
||||
repeat assumption
|
||||
end
|
||||
|
@ -389,7 +389,7 @@ theorem le_and_sep_of_lt {s t : seq} (Hs : regular s) (Ht : regular t) (Lst : s_
|
|||
sadd t (sneg s) n ≥ sadd t (sneg s) N - N⁻¹ - n⁻¹ : Habs
|
||||
... ≥ 0 - n⁻¹: begin
|
||||
apply algebra.sub_le_sub_right,
|
||||
apply le_of_lt,
|
||||
apply rat.le_of_lt,
|
||||
apply (iff.mpr (sub_pos_iff_lt _ _)),
|
||||
apply HN
|
||||
end
|
||||
|
@ -488,38 +488,38 @@ theorem s_mul_pos_of_pos {s t : seq} (Hs : regular s) (Ht : regular t) (Hps : po
|
|||
rotate 1,
|
||||
apply algebra.mul_le_mul,
|
||||
apply HNs,
|
||||
apply pnat.le.trans,
|
||||
apply max_left Ns Nt,
|
||||
rewrite -pnat.mul.assoc,
|
||||
apply pnat.le_trans,
|
||||
apply pnat.max_left Ns Nt,
|
||||
rewrite -pnat.mul_assoc,
|
||||
apply pnat.mul_le_mul_left,
|
||||
apply HNt,
|
||||
apply pnat.le.trans,
|
||||
apply max_right Ns Nt,
|
||||
rewrite -pnat.mul.assoc,
|
||||
apply pnat.le_trans,
|
||||
apply pnat.max_right Ns Nt,
|
||||
rewrite -pnat.mul_assoc,
|
||||
apply pnat.mul_le_mul_left,
|
||||
apply le_of_lt,
|
||||
apply inv_pos,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_of_lt,
|
||||
apply pnat.inv_pos,
|
||||
apply rat.le_trans,
|
||||
rotate 1,
|
||||
apply HNs,
|
||||
apply pnat.le.trans,
|
||||
apply max_left Ns Nt,
|
||||
rewrite -pnat.mul.assoc,
|
||||
apply pnat.le_trans,
|
||||
apply pnat.max_left Ns Nt,
|
||||
rewrite -pnat.mul_assoc,
|
||||
apply pnat.mul_le_mul_left,
|
||||
rewrite inv_mul_eq_mul_inv,
|
||||
rewrite pnat.inv_mul_eq_mul_inv,
|
||||
apply algebra.mul_lt_mul,
|
||||
rewrite [inv_mul_eq_mul_inv, -one_mul Ns⁻¹],
|
||||
rewrite [pnat.inv_mul_eq_mul_inv, -one_mul Ns⁻¹],
|
||||
apply algebra.mul_lt_mul,
|
||||
apply inv_lt_one_of_gt,
|
||||
apply dec_trivial,
|
||||
apply inv_ge_of_le,
|
||||
apply max_left,
|
||||
apply inv_pos,
|
||||
apply le_of_lt zero_lt_one,
|
||||
apply pnat.max_left,
|
||||
apply pnat.inv_pos,
|
||||
apply rat.le_of_lt zero_lt_one,
|
||||
apply inv_ge_of_le,
|
||||
apply max_right,
|
||||
apply inv_pos,
|
||||
repeat (apply le_of_lt; apply inv_pos)
|
||||
apply pnat.max_right,
|
||||
apply pnat.inv_pos,
|
||||
repeat (apply le_of_lt; apply pnat.inv_pos)
|
||||
end
|
||||
|
||||
theorem s_mul_gt_zero_of_gt_zero {s t : seq} (Hs : regular s) (Ht : regular t)
|
||||
|
@ -560,7 +560,7 @@ theorem s_zero_mul {s : seq} : smul s zero ≡ zero :=
|
|||
end
|
||||
|
||||
theorem s_mul_nonneg_of_pos_of_zero {s t : seq} (Hs : regular s) (Ht : regular t)
|
||||
(Hps : pos s) (Hpt : zero ≡ t) : nonneg (smul s t) :=
|
||||
(Hps : pos s) (Hpt : zero ≡ t) : nonneg (smul s t) :=
|
||||
begin
|
||||
apply nonneg_of_nonneg_equiv,
|
||||
rotate 2,
|
||||
|
@ -581,27 +581,27 @@ theorem s_mul_nonneg_of_nonneg {s t : seq} (Hs : regular s) (Ht : regular t)
|
|||
begin
|
||||
intro n,
|
||||
rewrite ↑smul,
|
||||
apply rat.le.by_cases 0 (s (((K₂ s t) * 2) * n)),
|
||||
apply rat.le_by_cases 0 (s (((K₂ s t) * 2) * n)),
|
||||
intro Hsp,
|
||||
apply rat.le.by_cases 0 (t (((K₂ s t) * 2) * n)),
|
||||
apply rat.le_by_cases 0 (t (((K₂ s t) * 2) * n)),
|
||||
intro Htp,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
rotate 1,
|
||||
apply rat.mul_nonneg Hsp Htp,
|
||||
rotate_right 1,
|
||||
apply le_of_lt,
|
||||
apply neg_neg_of_pos,
|
||||
apply inv_pos,
|
||||
apply pnat.inv_pos,
|
||||
intro Htn,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
rotate 1,
|
||||
apply algebra.mul_le_mul_of_nonpos_right,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply le_abs_self,
|
||||
apply canon_2_bound_left s t Hs,
|
||||
apply Htn,
|
||||
rotate_right 1,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
rotate 1,
|
||||
apply mul_le_mul_of_nonneg_left,
|
||||
apply Hpt,
|
||||
|
@ -610,21 +610,21 @@ 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, inv_mul_eq_mul_inv, -mul.assoc, inv_cancel_left, one_mul],
|
||||
rewrite [*pnat.mul_assoc, pnat.inv_mul_eq_mul_inv, -mul.assoc, pnat.inv_cancel_left, rat.one_mul],
|
||||
apply inv_ge_of_le,
|
||||
apply pnat.mul_le_mul_left,
|
||||
intro Hsn,
|
||||
apply rat.le.by_cases 0 (t (((K₂ s t) * 2) * n)),
|
||||
apply rat.le_by_cases 0 (t (((K₂ s t) * 2) * n)),
|
||||
intro Htp,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
rotate 1,
|
||||
apply mul_le_mul_of_nonpos_left,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
apply le_abs_self,
|
||||
apply canon_2_bound_right s t Ht,
|
||||
apply Hsn,
|
||||
rotate_right 1,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
rotate 1,
|
||||
apply mul_le_mul_of_nonneg_right,
|
||||
apply Hps,
|
||||
|
@ -633,18 +633,18 @@ 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, inv_mul_eq_mul_inv, mul.comm, -mul.assoc, 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 rat.le_trans,
|
||||
rotate 1,
|
||||
apply mul_nonneg_of_nonpos_of_nonpos,
|
||||
apply Hsn,
|
||||
apply Htn,
|
||||
apply le_of_lt,
|
||||
apply neg_neg_of_pos,
|
||||
apply inv_pos
|
||||
apply pnat.inv_pos
|
||||
end
|
||||
|
||||
theorem s_mul_ge_zero_of_ge_zero {s t : seq} (Hs : regular s) (Ht : regular t)
|
||||
|
@ -666,7 +666,7 @@ theorem not_lt_self (s : seq) : ¬ s_lt s s :=
|
|||
apply exists.elim Hlt,
|
||||
intro n Hn, esimp at Hn,
|
||||
rewrite [↑sadd at Hn,↑sneg at Hn, -sub_eq_add_neg at Hn, algebra.sub_self at Hn],
|
||||
apply absurd Hn (algebra.not_lt_of_ge (rat.le_of_lt !inv_pos))
|
||||
apply absurd Hn (algebra.not_lt_of_ge (rat.le_of_lt !pnat.inv_pos))
|
||||
end
|
||||
|
||||
theorem not_sep_self (s : seq) : ¬ s ≢ s :=
|
||||
|
@ -780,26 +780,26 @@ theorem s_lt_of_lt_of_le {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : r
|
|||
existsi max (2 * Nt) Nu,
|
||||
intro n Hn,
|
||||
rewrite Hcan,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
rotate 1,
|
||||
apply algebra.add_le_add,
|
||||
apply HNt,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
apply pnat.mul_le_mul_left 2,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
rotate 1,
|
||||
apply Hn,
|
||||
rotate_right 1,
|
||||
apply max_left,
|
||||
apply pnat.max_left,
|
||||
apply HNu,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
rotate 1,
|
||||
apply Hn,
|
||||
rotate_right 1,
|
||||
apply max_right,
|
||||
apply pnat.max_right,
|
||||
rewrite [-add_halves Nt, -sub_eq_add_neg, algebra.add_sub_cancel],
|
||||
apply inv_ge_of_le,
|
||||
apply max_left
|
||||
apply pnat.max_left
|
||||
end
|
||||
|
||||
theorem s_lt_of_le_of_lt {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : regular u)
|
||||
|
@ -818,26 +818,26 @@ theorem s_lt_of_le_of_lt {s t u : seq} (Hs : regular s) (Ht : regular t) (Hu : r
|
|||
existsi max (2 * Nu) Nt,
|
||||
intro n Hn,
|
||||
rewrite Hcan,
|
||||
apply rat.le.trans,
|
||||
apply rat.le_trans,
|
||||
rotate 1,
|
||||
apply algebra.add_le_add,
|
||||
apply HNt,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
rotate 1,
|
||||
apply Hn,
|
||||
rotate_right 1,
|
||||
apply max_right,
|
||||
apply pnat.max_right,
|
||||
apply HNu,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
apply pnat.mul_le_mul_left 2,
|
||||
apply pnat.le.trans,
|
||||
apply pnat.le_trans,
|
||||
rotate 1,
|
||||
apply Hn,
|
||||
rotate_right 1,
|
||||
apply max_left,
|
||||
apply pnat.max_left,
|
||||
rewrite [-add_halves Nu, neg_add_cancel_left],
|
||||
apply inv_ge_of_le,
|
||||
apply max_left
|
||||
apply pnat.max_left
|
||||
end
|
||||
|
||||
theorem le_of_le_reprs {s t : seq} (Hs : regular s) (Ht : regular t)
|
||||
|
@ -856,10 +856,10 @@ theorem const_le_const_of_le {a b : ℚ} (H : a ≤ b) : s_le (const a) (const b
|
|||
rewrite [↑s_le, ↑nonneg],
|
||||
intro n,
|
||||
rewrite [↑sadd, ↑sneg, ↑const],
|
||||
apply rat.le.trans,
|
||||
apply algebra.le.trans,
|
||||
apply neg_nonpos_of_nonneg,
|
||||
apply rat.le_of_lt,
|
||||
apply inv_pos,
|
||||
apply pnat.inv_pos,
|
||||
apply iff.mpr !sub_nonneg_iff_le,
|
||||
apply H
|
||||
end
|
||||
|
@ -879,7 +879,7 @@ theorem nat_inv_lt_rat {a : ℚ} (H : a > 0) : ∃ n : ℕ+, n⁻¹ < a :=
|
|||
apply div_two_lt_of_pos H,
|
||||
rewrite -(one_div_one_div (a / (1 + 1))),
|
||||
apply pceil_helper,
|
||||
apply pnat.le.refl,
|
||||
apply pnat.le_refl,
|
||||
apply one_div_pos_of_pos,
|
||||
apply div_pos_of_pos_of_pos H dec_trivial
|
||||
end
|
||||
|
@ -908,10 +908,10 @@ theorem s_le_of_le_pointwise {s t : seq} (Hs : regular s) (Ht : regular t)
|
|||
begin
|
||||
rewrite [↑s_le, ↑nonneg, ↑sadd, ↑sneg],
|
||||
intros,
|
||||
apply rat.le.trans,
|
||||
apply algebra.le.trans,
|
||||
apply iff.mpr !neg_nonpos_iff_nonneg,
|
||||
apply le_of_lt,
|
||||
apply inv_pos,
|
||||
apply pnat.inv_pos,
|
||||
apply iff.mpr !sub_nonneg_iff_le,
|
||||
apply H
|
||||
end
|
||||
|
@ -937,10 +937,10 @@ theorem r_sep_well_defined (s t u v : reg_seq) (Hsu : requiv s u) (Htv : requiv
|
|||
propext (sep_well_defined (reg_seq.is_reg s) (reg_seq.is_reg t) (reg_seq.is_reg u)
|
||||
(reg_seq.is_reg v) Hsu Htv)
|
||||
|
||||
theorem r_le.refl (s : reg_seq) : r_le s s := le.refl (reg_seq.is_reg s)
|
||||
theorem r_le.refl (s : reg_seq) : r_le s s := rat_seq.le_refl (reg_seq.is_reg s)
|
||||
|
||||
theorem r_le.trans {s t u : reg_seq} (Hst : r_le s t) (Htu : r_le t u) : r_le s u :=
|
||||
le.trans (reg_seq.is_reg s) (reg_seq.is_reg t) (reg_seq.is_reg u) Hst Htu
|
||||
rat_seq.le_trans (reg_seq.is_reg s) (reg_seq.is_reg t) (reg_seq.is_reg u) Hst Htu
|
||||
|
||||
theorem r_equiv_of_le_of_ge {s t : reg_seq} (Hs : r_le s t) (Hu : r_le t s) :
|
||||
requiv s t :=
|
||||
|
|
|
@ -123,7 +123,7 @@ section
|
|||
∀ k : ℕ+, ∃ N : ℕ+, ∀ m n : ℕ+,
|
||||
m ≥ N → n ≥ N → abs (X (elt_of m) - X (elt_of n)) ≤ of_rat k⁻¹ :=
|
||||
take k : ℕ+,
|
||||
have H1 : (k⁻¹ > (rat.of_num 0)), from !inv_pos,
|
||||
have H1 : (k⁻¹ > (rat.of_num 0)), from !pnat.inv_pos,
|
||||
have H2 : (of_rat k⁻¹ > of_rat (rat.of_num 0)), from !of_rat_lt_of_rat_of_lt H1,
|
||||
obtain (N : ℕ) (H : ∀ m n, m ≥ N → n ≥ N → abs (X m - X n) < of_rat k⁻¹), from H _ H2,
|
||||
exists.intro (pnat.succ N)
|
||||
|
|
Loading…
Reference in a new issue