@@ -445,74 +445,25 @@ protected theorem sign_eq_ediv_abs (a : ℤ) : sign a = a / |a| :=
445
445
446
446
/-! ### `/` and ordering -/
447
447
448
-
449
- protected theorem ediv_mul_le (a : ℤ) {b : ℤ} (H : b ≠ 0 ) : a / b * b ≤ a :=
450
- le_of_sub_nonneg <| by rw [mul_comm, ← emod_def]; apply emod_nonneg _ H
451
448
#align int.div_mul_le Int.ediv_mul_le
452
-
453
- protected theorem ediv_le_of_le_mul {a b c : ℤ} (H : 0 < c) (H' : a ≤ b * c) : a / c ≤ b :=
454
- le_of_mul_le_mul_right (le_trans (Int.ediv_mul_le _ (ne_of_gt H)) H') H
455
449
#align int.div_le_of_le_mul Int.ediv_le_of_le_mul
456
-
457
- protected theorem mul_lt_of_lt_ediv {a b c : ℤ} (H : 0 < c) (H3 : a < b / c) : a * c < b :=
458
- lt_of_not_ge <| mt (Int.ediv_le_of_le_mul H) (not_le_of_gt H3)
459
450
#align int.mul_lt_of_lt_div Int.mul_lt_of_lt_ediv
460
-
461
- protected theorem mul_le_of_le_ediv {a b c : ℤ} (H1 : 0 < c) (H2 : a ≤ b / c) : a * c ≤ b :=
462
- le_trans (mul_le_mul_of_nonneg_right H2 (le_of_lt H1)) (Int.ediv_mul_le _ (ne_of_gt H1))
463
451
#align int.mul_le_of_le_div Int.mul_le_of_le_ediv
464
-
465
- protected theorem le_ediv_of_mul_le {a b c : ℤ} (H1 : 0 < c) (H2 : a * c ≤ b) : a ≤ b / c :=
466
- le_of_lt_add_one <|
467
- lt_of_mul_lt_mul_right (lt_of_le_of_lt H2 (lt_ediv_add_one_mul_self _ H1)) (le_of_lt H1)
468
452
#align int.le_div_of_mul_le Int.le_ediv_of_mul_le
469
-
470
- protected theorem le_ediv_iff_mul_le {a b c : ℤ} (H : 0 < c) : a ≤ b / c ↔ a * c ≤ b :=
471
- ⟨Int.mul_le_of_le_ediv H, Int.le_ediv_of_mul_le H⟩
472
453
#align int.le_div_iff_mul_le Int.le_ediv_iff_mul_le
473
-
474
- protected theorem ediv_le_ediv {a b c : ℤ} (H : 0 < c) (H' : a ≤ b) : a / c ≤ b / c :=
475
- Int.le_ediv_of_mul_le H (le_trans (Int.ediv_mul_le _ (ne_of_gt H)) H')
476
454
#align int.div_le_div Int.ediv_le_ediv
477
-
478
- protected theorem ediv_lt_of_lt_mul {a b c : ℤ} (H : 0 < c) (H' : a < b * c) : a / c < b :=
479
- lt_of_not_ge <| mt (Int.mul_le_of_le_ediv H) (not_le_of_gt H')
480
455
#align int.div_lt_of_lt_mul Int.ediv_lt_of_lt_mul
481
-
482
- protected theorem lt_mul_of_ediv_lt {a b c : ℤ} (H1 : 0 < c) (H2 : a / c < b) : a < b * c :=
483
- lt_of_not_ge <| mt (Int.le_ediv_of_mul_le H1) (not_le_of_gt H2)
484
456
#align int.lt_mul_of_div_lt Int.lt_mul_of_ediv_lt
485
-
486
- protected theorem ediv_lt_iff_lt_mul {a b c : ℤ} (H : 0 < c) : a / c < b ↔ a < b * c :=
487
- ⟨Int.lt_mul_of_ediv_lt H, Int.ediv_lt_of_lt_mul H⟩
488
457
#align int.div_lt_iff_lt_mul Int.ediv_lt_iff_lt_mul
489
-
490
- protected theorem le_mul_of_ediv_le {a b c : ℤ} (H1 : 0 ≤ b) (H2 : b ∣ a) (H3 : a / b ≤ c) :
491
- a ≤ c * b := by rw [← Int.ediv_mul_cancel H2]; exact mul_le_mul_of_nonneg_right H3 H1
492
458
#align int.le_mul_of_div_le Int.le_mul_of_ediv_le
493
-
494
- protected theorem lt_ediv_of_mul_lt {a b c : ℤ} (H1 : 0 ≤ b) (H2 : b ∣ c) (H3 : a * b < c) :
495
- a < c / b :=
496
- lt_of_not_ge <| mt (Int.le_mul_of_ediv_le H1 H2) (not_le_of_gt H3)
497
459
#align int.lt_div_of_mul_lt Int.lt_ediv_of_mul_lt
498
-
499
- protected theorem lt_ediv_iff_mul_lt {a b : ℤ} (c : ℤ) (H : 0 < c) (H' : c ∣ b) :
500
- a < b / c ↔ a * c < b :=
501
- ⟨Int.mul_lt_of_lt_ediv H, Int.lt_ediv_of_mul_lt (le_of_lt H) H'⟩
502
460
#align int.lt_div_iff_mul_lt Int.lt_ediv_iff_mul_lt
503
-
504
- theorem ediv_pos_of_pos_of_dvd {a b : ℤ} (H1 : 0 < a) (H2 : 0 ≤ b) (H3 : b ∣ a) : 0 < a / b :=
505
- Int.lt_ediv_of_mul_lt H2 H3 (by rwa [zero_mul])
506
461
#align int.div_pos_of_pos_of_dvd Int.ediv_pos_of_pos_of_dvd
507
462
508
463
theorem natAbs_eq_of_dvd_dvd {s t : ℤ} (hst : s ∣ t) (hts : t ∣ s) : natAbs s = natAbs t :=
509
464
Nat.dvd_antisymm (natAbs_dvd_natAbs.mpr hst) (natAbs_dvd_natAbs.mpr hts)
510
465
#align int.nat_abs_eq_of_dvd_dvd Int.natAbs_eq_of_dvd_dvd
511
466
512
- theorem ediv_eq_ediv_of_mul_eq_mul {a b c d : ℤ} (H2 : d ∣ c) (H3 : b ≠ 0 ) (H4 : d ≠ 0 )
513
- (H5 : a * d = b * c) : a / b = c / d :=
514
- Int.ediv_eq_of_eq_mul_right H3 <| by
515
- rw [← Int.mul_ediv_assoc _ H2]; exact (Int.ediv_eq_of_eq_mul_left H4 H5.symm).symm
516
467
#align int.div_eq_div_of_mul_eq_mul Int.ediv_eq_ediv_of_mul_eq_mul
517
468
518
469
theorem ediv_dvd_of_dvd {s t : ℤ} (hst : s ∣ t) : t / s ∣ t := by
0 commit comments