diff --git a/library/data/real/basic.lean b/library/data/real/basic.lean index a94495a76..e34e2e166 100644 --- a/library/data/real/basic.lean +++ b/library/data/real/basic.lean @@ -1055,20 +1055,20 @@ notation `ℝ` := real protected definition prio := num.pred rat.prio -definition add (x y : ℝ) : ℝ := +protected definition add (x y : ℝ) : ℝ := (quot.lift_on₂ x y (λ a b, quot.mk (a + b)) (take a b c d : reg_seq, take Hab : requiv a c, take Hcd : requiv b d, quot.sound (radd_well_defined Hab Hcd))) --infix [priority real.prio] + := add -definition mul (x y : ℝ) : ℝ := +protected definition mul (x y : ℝ) : ℝ := (quot.lift_on₂ x y (λ a b, quot.mk (a * b)) (take a b c d : reg_seq, take Hab : requiv a c, take Hcd : requiv b d, quot.sound (rmul_well_defined Hab Hcd))) --infix [priority real.prio] * := mul -definition neg (x : ℝ) : ℝ := +protected definition neg (x : ℝ) : ℝ := (quot.lift_on x (λ a, quot.mk (-a)) (take a b : reg_seq, take Hab : requiv a b, quot.sound (rneg_well_defined Hab))) --prefix [priority real.prio] `-` := neg @@ -1094,10 +1094,10 @@ definition of_int [coercion] (i : ℤ) : ℝ := int.to.real i definition of_nat [coercion] (n : ℕ) : ℝ := nat.to.real n definition of_num [coercion] [reducible] (n : num) : ℝ := of_rat (rat.of_num n) -definition real_has_zero [reducible] [instance] [priority rat.prio] : has_zero real := +definition real_has_zero [reducible] [instance] [priority real.prio] : has_zero real := has_zero.mk (0:rat) -definition real_has_one [reducible] [instance] [priority rat.prio] : has_one real := +definition real_has_one [reducible] [instance] [priority real.prio] : has_one real := has_one.mk (1:rat) theorem real_zero_eq_rat_zero : (0:real) = of_rat (0:rat) := diff --git a/library/data/real/complete.lean b/library/data/real/complete.lean index ab20557b0..eaf0f11ad 100644 --- a/library/data/real/complete.lean +++ b/library/data/real/complete.lean @@ -17,8 +17,6 @@ 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 -local notation 0 := rat.of_num 0 -local notation 1 := rat.of_num 1 local postfix ⁻¹ := pnat.inv open eq.ops pnat classical @@ -116,7 +114,7 @@ theorem equiv_abs_of_ge_zero {s : seq} (Hs : regular s) (Hz : s_le zero s) : s_a apply rat.le.trans, apply add_le_add, repeat (apply inv_ge_of_le; apply Hn), - rewrite pnat.add_halves, + krewrite pnat.add_halves, apply rat.le.refl end @@ -147,7 +145,7 @@ theorem equiv_neg_abs_of_le_zero {s : seq} (Hs : regular s) (Hz : s_le s zero) : apply rat.le.trans, apply add_le_add, repeat (apply inv_ge_of_le; apply Hn), - rewrite pnat.add_halves, + krewrite pnat.add_halves, apply rat.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, @@ -238,7 +236,7 @@ theorem cauchy_with_rate_of_converges_to_with_rate {X : r_seq} {a : ℝ} {N : apply Hn, xrewrite -of_rat_add, apply of_rat_le_of_rat_of_le, - rewrite pnat.add_halves, + krewrite pnat.add_halves, apply rat.le.refl end @@ -274,7 +272,10 @@ private theorem lim_seq_reg_helper {m n : ℕ+} (Hmn : M (2 * n) ≤M (2 * m)) : apply pnat.le.trans, apply Hmn, apply Nb_spec_right, - rewrite [-*of_rat_add, rat.add.assoc, pnat.add_halves], + krewrite [-+of_rat_add], + change of_rat ((2 * m)⁻¹ + (2 * n)⁻¹ + (2 * n)⁻¹) ≤ of_rat (m⁻¹ + n⁻¹), + rewrite [algebra.add.assoc], + krewrite pnat.add_halves, apply of_rat_le_of_rat_of_le, apply add_le_add_right, apply inv_ge_of_le, @@ -292,7 +293,7 @@ theorem lim_seq_reg : rat_seq.regular lim_seq := cases em (M (2 * m) ≥ M (2 * n)) with [Hor1, Hor2], apply lim_seq_reg_helper Hor1, let Hor2' := pnat.le_of_lt (pnat.lt_of_not_le Hor2), - rewrite [abs_sub (X (Nb M n)), abs_sub (X (Nb M m)), abs_sub, + krewrite [abs_sub (X (Nb M n)), abs_sub (X (Nb M m)), abs_sub, rat.add.comm, add_comm_three], apply lim_seq_reg_helper Hor2' end @@ -351,9 +352,10 @@ theorem converges_to_with_rate_of_cauchy_with_rate : converges_to_with_rate X li rewrite ↑lim_seq, apply approx_spec, apply lim_spec, - rewrite [-*of_rat_add], + krewrite [-+of_rat_add], + change of_rat ((2 * k)⁻¹ + (2 * n)⁻¹ + n⁻¹) ≤ of_rat k⁻¹, apply of_rat_le_of_rat_of_le, - apply rat.le.trans, + apply algebra.le.trans, apply add_le_add_three, apply rat.le.refl, apply inv_ge_of_le, @@ -369,7 +371,8 @@ theorem converges_to_with_rate_of_cauchy_with_rate : converges_to_with_rate X li apply Hn, rotate_right 1, apply Nb_spec_left, - rewrite [-*pnat.mul.assoc, pnat.p_add_fractions], + rewrite -*pnat.mul.assoc, + krewrite pnat.p_add_fractions, apply rat.le.refl end @@ -500,7 +503,7 @@ theorem ceil_lt_ceil_succ (x : ℝ) : ceil x < ceil (x + 1) := apply floor_sub_one_lt_floor end open nat -set_option pp.coercions true + theorem archimedean_small {ε : ℝ} (H : ε > 0) : ∃ (n : ℕ), 1 / succ n < ε := let n := int.nat_abs (ceil (2 / ε)) in assert int.of_nat n ≥ ceil (2 / ε), @@ -529,8 +532,6 @@ local postfix `~` := nat_of_pnat -- The top part of this section could be refactored. What is the appropriate place to define -- bounds, supremum, etc? In algebra/ordered_field? They potentially apply to more than just ℝ. - -local notation 2 := (1 : ℚ) + 1 parameter X : ℝ → Prop definition ub (x : ℝ) := ∀ y : ℝ, X y → y ≤ x @@ -781,7 +782,7 @@ private theorem under_seq_mono (i j : ℕ) (H : i ≤ j) : under_seq i ≤ under rewrite -Hk', apply under_seq_mono_helper end -set_option pp.coercions true + 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) diff --git a/library/data/real/division.lean b/library/data/real/division.lean index fd6a8b28e..8865e3fd0 100644 --- a/library/data/real/division.lean +++ b/library/data/real/division.lean @@ -619,15 +619,15 @@ theorem lt_or_eq_of_le (x y : ℝ) : x ≤ y → x < y ∨ x = y := 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) -noncomputable definition dec_lt : decidable_rel lt := +noncomputable definition dec_lt : decidable_rel real.lt := begin rewrite ↑decidable_rel, intros, apply prop_decidable end - protected noncomputable definition discrete_linear_ordered_field [reducible] [trans_instance]: - algebra.discrete_linear_ordered_field ℝ := +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, @@ -640,13 +640,12 @@ noncomputable definition dec_lt : decidable_rel lt := theorem of_rat_zero : of_rat 0 = 0 := rfl -set_option pp.coercions true theorem of_rat_divide (x y : ℚ) : of_rat (x / y) = of_rat x / of_rat y := by_cases (assume yz : y = 0, by rewrite [yz, algebra.div_zero, *of_rat_zero, algebra.div_zero]) (assume ynz : y ≠ 0, have ynz' : of_rat y ≠ 0, from assume yz', ynz (of_rat.inj yz'), - !eq_div_of_mul_eq ynz' (by rewrite [-of_rat_mul, !div_mul_cancel ynz])) + !eq_div_of_mul_eq ynz' (by krewrite [-of_rat_mul, !div_mul_cancel ynz])) open int diff --git a/library/data/real/order.lean b/library/data/real/order.lean index 6b220f495..673c53545 100644 --- a/library/data/real/order.lean +++ b/library/data/real/order.lean @@ -1015,71 +1015,66 @@ open real open [classes] rat_seq namespace real -definition lt (x y : ℝ) := quot.lift_on₂ x y (λ a b, rat_seq.r_lt a b) rat_seq.r_lt_well_defined ---infix [priority real.prio] `<` := lt +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 le (x y : ℝ) := quot.lift_on₂ x y (λ a b, rat_seq.r_le a b) rat_seq.r_le_well_defined ---infix [priority real.prio] `<=` := le ---infix [priority real.prio] `≤` := le +definition real_has_lt [reducible] [instance] [priority real.prio] : has_lt ℝ := +has_lt.mk real.lt -definition gt [reducible] (a b : ℝ) := lt b a -definition ge [reducible] (a b : ℝ) := le b a - ---infix [priority real.prio] >= := real.ge ---infix [priority real.prio] ≥ := real.ge ---infix [priority real.prio] > := real.gt +definition real_has_le [reducible] [instance] [priority real.prio] : has_le ℝ := +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 : ℝ) : le x x := +theorem le.refl (x : ℝ) : x ≤ x := quot.induction_on x (λ t, rat_seq.r_le.refl t) -theorem le.trans {x y z : ℝ} : le x y → le y z → le x z := +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 : ℝ} : le x y → le y x → x = y := +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 : ℝ) : lt x y ↔ le x y ∧ x ≢ y := +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 : ℝ) : le x y → le (z + x) (z + y) := +theorem add_le_add_of_le_right_var (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 : ℝ) : le x y → ∀ z : ℝ, le (z + x) (z + y) := +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 -theorem mul_gt_zero_of_gt_zero (x y : ℝ) : lt zero x → lt zero y → lt zero (x * y) := +theorem mul_gt_zero_of_gt_zero (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 : ℝ) : le zero x → le zero y → le zero (x * y) := +theorem mul_ge_zero_of_ge_zero (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 : ℝ) : ¬ lt x x := +theorem not_lt_self (x : ℝ) : ¬ x < x := quot.induction_on x (λ s, rat_seq.r_not_lt_self s) -theorem le_of_lt {x y : ℝ} : lt x y → le x y := +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 : ℝ} : le x y → lt y z → lt x z := +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 : ℝ} : lt x y → le y z → lt x z := +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 : ℝ) : lt x y → lt (z + x) (z + y) := +theorem add_lt_add_left_var (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 : ℝ) : lt x y → ∀ z : ℝ, lt (z + x) (z + y) := +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 -theorem zero_lt_one : lt (0 : ℝ) (1 : ℝ) := rat_seq.r_zero_lt_one +theorem zero_lt_one : (0 : ℝ) < (1 : ℝ) := rat_seq.r_zero_lt_one -theorem le_of_lt_or_eq (x y : ℝ) : lt x y ∨ x = y → le x y := +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' @@ -1089,21 +1084,21 @@ theorem le_of_lt_or_eq (x y : ℝ) : lt x y ∨ x = y → le x y := apply (or.inr (quot.exact H')) end))) - 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 - ⦄ +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 +⦄ open int theorem of_rat_sub (a b : ℚ) : of_rat (a - b) = of_rat a - of_rat b := rfl @@ -1128,7 +1123,6 @@ theorem lt_of_of_rat_lt_of_rat {a b : ℚ} : of_rat a < of_rat b → a < b := theorem of_rat_lt_of_rat_iff (a b : ℚ) : of_rat a < of_rat b ↔ a < b := iff.intro lt_of_of_rat_lt_of_rat of_rat_lt_of_rat_of_lt -set_option pp.coercions true theorem of_int_le_of_int_iff (a b : ℤ) : of_int a ≤ of_int b ↔ (a ≤ b) := begin rewrite [+of_int_eq, of_rat_le_of_rat_iff], apply rat.of_int_le_of_int_iff end