fix(real/order): remove sorry

This commit is contained in:
Rob Lewis 2015-10-13 16:14:59 -04:00 committed by Leonardo de Moura
parent 06c1a97259
commit 670ac6ae19

View file

@ -872,18 +872,17 @@ theorem le_of_const_le_const {a b : } (H : s_le (const a) (const b)) : a ≤
end end
theorem nat_inv_lt_rat {a : } (H : a > 0) : ∃ n : +, n⁻¹ < a := theorem nat_inv_lt_rat {a : } (H : a > 0) : ∃ n : +, n⁻¹ < a :=
/-begin begin
existsi (pceil (1 / (a / (2)))), existsi (pceil (1 / (a / (2)))),
apply lt_of_le_of_lt, apply lt_of_le_of_lt,
rotate 1, rotate 1,
apply div_two_lt_of_pos H, apply div_two_lt_of_pos H,
rewrite -(one_div_one_div (a / (1 + 1))), rewrite -(one_div_one_div (a / (1 + 1))),
apply pceil_helper, apply pceil_helper,
rewrite one_div_one_div,
apply pnat.le.refl, apply pnat.le.refl,
apply one_div_pos_of_pos, apply one_div_pos_of_pos,
apply div_pos_of_pos_of_pos H dec_trivial apply div_pos_of_pos_of_pos H dec_trivial
end-/sorry end
theorem const_lt_const_of_lt {a b : } (H : a < b) : s_lt (const a) (const b) := theorem const_lt_const_of_lt {a b : } (H : a < b) : s_lt (const a) (const b) :=