@@ -12,6 +12,7 @@ import Mathlib.Tactic.PushNeg
12
12
import Mathlib.Util.AssertExists
13
13
14
14
#align_import data.nat.basic from "leanprover-community/mathlib" @"bd835ef554f37ef9b804f0903089211f89cb370b"
15
+ #align_import data.nat.sqrt from "leanprover-community/mathlib" @"ba2245edf0c8bb155f1569fd9b9492a9b384cde6"
15
16
16
17
/-!
17
18
# Basic operations on the natural numbers
@@ -678,6 +679,22 @@ protected lemma eq_div_of_mul_eq_left (hc : c ≠ 0) (h : a * c = b) : a = b / c
678
679
protected lemma eq_div_of_mul_eq_right (hc : c ≠ 0 ) (h : c * a = b) : a = b / c := by
679
680
rw [← h, Nat.mul_div_cancel_left _ (Nat.pos_iff_ne_zero.2 hc)]
680
681
682
+ protected lemma mul_le_of_le_div (k x y : ℕ) (h : x ≤ y / k) : x * k ≤ y := by
683
+ if hk : k = 0 then
684
+ rw [hk, Nat.mul_zero]; exact zero_le _
685
+ else
686
+ rwa [← le_div_iff_mul_le (Nat.pos_iff_ne_zero.2 hk)]
687
+
688
+ protected lemma div_mul_div_le (a b c d : ℕ) :
689
+ (a / b) * (c / d) ≤ (a * c) / (b * d) := by
690
+ if hb : b = 0 then simp [hb] else
691
+ if hd : d = 0 then simp [hd] else
692
+ have hbd : b * d ≠ 0 := Nat.mul_ne_zero hb hd
693
+ rw [le_div_iff_mul_le (Nat.pos_of_ne_zero hbd)]
694
+ transitivity ((a / b) * b) * ((c / d) * d)
695
+ · apply Nat.le_of_eq; simp only [Nat.mul_assoc, Nat.mul_left_comm]
696
+ · apply Nat.mul_le_mul <;> apply div_mul_le_self
697
+
681
698
/-!
682
699
### `pow`
683
700
@@ -1419,6 +1436,235 @@ lemma div_lt_div_of_lt_of_dvd {a b d : ℕ} (hdb : d ∣ b) (h : a < b) : a / d
1419
1436
exact lt_of_le_of_lt (mul_div_le a d) h
1420
1437
#align nat.div_lt_div_of_lt_of_dvd Nat.div_lt_div_of_lt_of_dvd
1421
1438
1439
+ /-!
1440
+ ### `sqrt`
1441
+
1442
+ See [Wikipedia, *Methods of computing square roots* ]
1443
+ (https://en.wikipedia.org/wiki/Methods_of_computing_square_roots#Binary_numeral_system_(base_2)).
1444
+ -/
1445
+
1446
+ private lemma iter_fp_bound (n k : ℕ) :
1447
+ let iter_next (n guess : ℕ) := (guess + n / guess) / 2 ;
1448
+ sqrt.iter n k ≤ iter_next n (sqrt.iter n k) := by
1449
+ intro iter_next
1450
+ unfold sqrt.iter
1451
+ if h : (k + n / k) / 2 < k then
1452
+ simpa [if_pos h] using iter_fp_bound _ _
1453
+ else
1454
+ simpa [if_neg h] using Nat.le_of_not_lt h
1455
+
1456
+ private lemma AM_GM : {a b : ℕ} → (4 * a * b ≤ (a + b) * (a + b))
1457
+ | 0 , _ => by rw [Nat.mul_zero, Nat.zero_mul]; exact zero_le _
1458
+ | _, 0 => by rw [Nat.mul_zero]; exact zero_le _
1459
+ | a + 1 , b + 1 => by
1460
+ simpa only [Nat.mul_add, Nat.add_mul, show (4 : ℕ) = 1 + 1 + 1 + 1 from rfl, Nat.one_mul,
1461
+ Nat.mul_one, Nat.add_assoc, Nat.add_left_comm, Nat.add_le_add_iff_left]
1462
+ using Nat.add_le_add_right (@AM_GM a b) 4
1463
+
1464
+ -- These two lemmas seem like they belong to `Std.Data.Nat.Basic`.
1465
+
1466
+ lemma sqrt.iter_sq_le (n guess : ℕ) : sqrt.iter n guess * sqrt.iter n guess ≤ n := by
1467
+ unfold sqrt.iter
1468
+ let next := (guess + n / guess) / 2
1469
+ if h : next < guess then
1470
+ simpa only [dif_pos h] using sqrt.iter_sq_le n next
1471
+ else
1472
+ simp only [dif_neg h]
1473
+ apply Nat.mul_le_of_le_div
1474
+ apply Nat.le_of_add_le_add_left (a := guess)
1475
+ rw [← Nat.mul_two, ← le_div_iff_mul_le]
1476
+ · exact Nat.le_of_not_lt h
1477
+ · exact Nat.zero_lt_two
1478
+
1479
+ lemma sqrt.lt_iter_succ_sq (n guess : ℕ) (hn : n < (guess + 1 ) * (guess + 1 )) :
1480
+ n < (sqrt.iter n guess + 1 ) * (sqrt.iter n guess + 1 ) := by
1481
+ unfold sqrt.iter
1482
+ -- m was `next`
1483
+ let m := (guess + n / guess) / 2
1484
+ dsimp
1485
+ split_ifs with h
1486
+ · suffices n < (m + 1 ) * (m + 1 ) by
1487
+ simpa only [dif_pos h] using sqrt.lt_iter_succ_sq n m this
1488
+ refine Nat.lt_of_mul_lt_mul_left ?_ (a := 4 * (guess * guess))
1489
+ apply Nat.lt_of_le_of_lt AM_GM
1490
+ rw [show (4 : ℕ) = 2 * 2 from rfl]
1491
+ rw [Nat.mul_mul_mul_comm 2 , Nat.mul_mul_mul_comm (2 * guess)]
1492
+ refine Nat.mul_self_lt_mul_self (?_ : _ < _ * succ (_ / 2 ))
1493
+ rw [← add_div_right _ (by decide), Nat.mul_comm 2 , Nat.mul_assoc,
1494
+ show guess + n / guess + 2 = (guess + n / guess + 1 ) + 1 from rfl]
1495
+ have aux_lemma {a : ℕ} : a ≤ 2 * ((a + 1 ) / 2 ) := by omega
1496
+ refine lt_of_lt_of_le ?_ (Nat.mul_le_mul_left _ aux_lemma)
1497
+ rw [Nat.add_assoc, Nat.mul_add]
1498
+ exact Nat.add_lt_add_left (lt_mul_div_succ _ (lt_of_le_of_lt (Nat.zero_le m) h)) _
1499
+ · simpa only [dif_neg h] using hn
1500
+
1501
+ #align nat.sqrt Nat.sqrt
1502
+ -- Porting note: the implementation òf `Nat.sqrt` in `Std` no longer needs `sqrt_aux`.
1503
+ #noalign nat.sqrt_aux_dec
1504
+ #noalign nat.sqrt_aux
1505
+ #noalign nat.sqrt_aux_0
1506
+ #noalign nat.sqrt_aux_1
1507
+ #noalign nat.sqrt_aux_2
1508
+ private def IsSqrt (n q : ℕ) : Prop :=
1509
+ q * q ≤ n ∧ n < (q + 1 ) * (q + 1 )
1510
+ -- Porting note: as the definition of square root has changed,
1511
+ -- the proof of `sqrt_isSqrt` is attempted from scratch.
1512
+ /-
1513
+ Sketch of proof:
1514
+ Up to rounding, in terms of the definition of `sqrt.iter`,
1515
+
1516
+ * By AM-GM inequality, `next² ≥ n` giving one of the bounds.
1517
+ * When we terminated, we have `guess ≥ next` from which we deduce the other bound `n ≥ next²`.
1518
+
1519
+ To turn this into a lean proof we need to manipulate, use properties of natural number division etc.
1520
+ -/
1521
+ private lemma sqrt_isSqrt (n : ℕ) : IsSqrt n (sqrt n) := by
1522
+ match n with
1523
+ | 0 => simp [IsSqrt, sqrt]
1524
+ | 1 => simp [IsSqrt, sqrt]
1525
+ | n + 2 =>
1526
+ have h : ¬ (n + 2 ) ≤ 1 := by simp
1527
+ simp only [IsSqrt, sqrt, h, ite_false]
1528
+ refine ⟨sqrt.iter_sq_le _ _, sqrt.lt_iter_succ_sq _ _ ?_⟩
1529
+ simp only [Nat.mul_add, Nat.add_mul, Nat.one_mul, Nat.mul_one, ← Nat.add_assoc]
1530
+ rw [lt_add_one_iff, Nat.add_assoc, ← Nat.mul_two]
1531
+ refine le_trans (Nat.le_of_eq (div_add_mod' (n + 2 ) 2 ).symm) ?_
1532
+ rw [Nat.add_comm, Nat.add_le_add_iff_right, add_mod_right]
1533
+ simp only [Nat.zero_lt_two, add_div_right, succ_mul_succ]
1534
+ refine le_trans (b := 1 ) ?_ ?_
1535
+ · exact (lt_succ.1 <| mod_lt n Nat.zero_lt_two)
1536
+ · exact Nat.le_add_left _ _
1537
+
1538
+ lemma sqrt_le (n : ℕ) : sqrt n * sqrt n ≤ n := (sqrt_isSqrt n).left
1539
+ #align nat.sqrt_le Nat.sqrt_le
1540
+
1541
+ lemma sqrt_le' (n : ℕ) : sqrt n ^ 2 ≤ n := by simpa [Nat.pow_two] using sqrt_le n
1542
+ #align nat.sqrt_le' Nat.sqrt_le'
1543
+
1544
+ lemma lt_succ_sqrt (n : ℕ) : n < succ (sqrt n) * succ (sqrt n) := (sqrt_isSqrt n).right
1545
+ #align nat.lt_succ_sqrt Nat.lt_succ_sqrt
1546
+
1547
+ lemma lt_succ_sqrt' (n : ℕ) : n < succ (sqrt n) ^ 2 := by simpa [Nat.pow_two] using lt_succ_sqrt n
1548
+ #align nat.lt_succ_sqrt' Nat.lt_succ_sqrt'
1549
+
1550
+ lemma sqrt_le_add (n : ℕ) : n ≤ sqrt n * sqrt n + sqrt n + sqrt n := by
1551
+ rw [← succ_mul]; exact le_of_lt_succ (lt_succ_sqrt n)
1552
+ #align nat.sqrt_le_add Nat.sqrt_le_add
1553
+
1554
+ lemma le_sqrt : m ≤ sqrt n ↔ m * m ≤ n :=
1555
+ ⟨fun h ↦ le_trans (mul_self_le_mul_self h) (sqrt_le n),
1556
+ fun h ↦ le_of_lt_succ <| Nat.mul_self_lt_mul_self_iff.1 <| lt_of_le_of_lt h (lt_succ_sqrt n)⟩
1557
+ #align nat.le_sqrt Nat.le_sqrt
1558
+
1559
+ lemma le_sqrt' : m ≤ sqrt n ↔ m ^ 2 ≤ n := by simpa only [Nat.pow_two] using le_sqrt
1560
+ #align nat.le_sqrt' Nat.le_sqrt'
1561
+
1562
+ lemma sqrt_lt : sqrt m < n ↔ m < n * n := by simp only [← not_le, le_sqrt]
1563
+ #align nat.sqrt_lt Nat.sqrt_lt
1564
+
1565
+ lemma sqrt_lt' : sqrt m < n ↔ m < n ^ 2 := by simp only [← not_le, le_sqrt']
1566
+ #align nat.sqrt_lt' Nat.sqrt_lt'
1567
+
1568
+ lemma sqrt_le_self (n : ℕ) : sqrt n ≤ n := le_trans (le_mul_self _) (sqrt_le n)
1569
+ #align nat.sqrt_le_self Nat.sqrt_le_self
1570
+
1571
+ lemma sqrt_le_sqrt (h : m ≤ n) : sqrt m ≤ sqrt n := le_sqrt.2 (le_trans (sqrt_le _) h)
1572
+ #align nat.sqrt_le_sqrt Nat.sqrt_le_sqrt
1573
+
1574
+ @[simp] lemma sqrt_zero : sqrt 0 = 0 := rfl
1575
+ #align nat.sqrt_zero Nat.sqrt_zero
1576
+
1577
+ @[simp] lemma sqrt_one : sqrt 1 = 1 := rfl
1578
+ #align nat.sqrt_one Nat.sqrt_one
1579
+
1580
+ lemma sqrt_eq_zero : sqrt n = 0 ↔ n = 0 :=
1581
+ ⟨fun h ↦
1582
+ Nat.eq_zero_of_le_zero <| le_of_lt_succ <| (@sqrt_lt n 1 ).1 <| by rw [h]; decide,
1583
+ by rintro rfl; simp⟩
1584
+ #align nat.sqrt_eq_zero Nat.sqrt_eq_zero
1585
+
1586
+ lemma eq_sqrt : a = sqrt n ↔ a * a ≤ n ∧ n < (a + 1 ) * (a + 1 ) :=
1587
+ ⟨fun e ↦ e.symm ▸ sqrt_isSqrt n,
1588
+ fun ⟨h₁, h₂⟩ ↦ le_antisymm (le_sqrt.2 h₁) (le_of_lt_succ <| sqrt_lt.2 h₂)⟩
1589
+ #align nat.eq_sqrt Nat.eq_sqrt
1590
+
1591
+ lemma eq_sqrt' : a = sqrt n ↔ a ^ 2 ≤ n ∧ n < (a + 1 ) ^ 2 := by
1592
+ simpa only [Nat.pow_two] using eq_sqrt
1593
+ #align nat.eq_sqrt' Nat.eq_sqrt'
1594
+
1595
+ lemma le_three_of_sqrt_eq_one (h : sqrt n = 1 ) : n ≤ 3 :=
1596
+ le_of_lt_succ <| (@sqrt_lt n 2 ).1 <| by rw [h]; decide
1597
+ #align nat.le_three_of_sqrt_eq_one Nat.le_three_of_sqrt_eq_one
1598
+
1599
+ lemma sqrt_lt_self (h : 1 < n) : sqrt n < n :=
1600
+ sqrt_lt.2 <| by have := Nat.mul_lt_mul_of_pos_left h (lt_of_succ_lt h); rwa [Nat.mul_one] at this
1601
+ #align nat.sqrt_lt_self Nat.sqrt_lt_self
1602
+
1603
+ lemma sqrt_pos : 0 < sqrt n ↔ 0 < n :=
1604
+ le_sqrt
1605
+ #align nat.sqrt_pos Nat.sqrt_pos
1606
+
1607
+ lemma sqrt_add_eq (n : ℕ) (h : a ≤ n + n) : sqrt (n * n + a) = n :=
1608
+ le_antisymm
1609
+ (le_of_lt_succ <|
1610
+ sqrt_lt.2 <| by
1611
+ rw [succ_mul, mul_succ, add_succ, Nat.add_assoc];
1612
+ exact lt_succ_of_le (Nat.add_le_add_left h _))
1613
+ (le_sqrt.2 <| Nat.le_add_right _ _)
1614
+ #align nat.sqrt_add_eq Nat.sqrt_add_eq
1615
+
1616
+ lemma sqrt_add_eq' (n : ℕ) (h : a ≤ n + n) : sqrt (n ^ 2 + a) = n := by
1617
+ simpa [Nat.pow_two] using sqrt_add_eq n h
1618
+ #align nat.sqrt_add_eq' Nat.sqrt_add_eq'
1619
+
1620
+ lemma sqrt_eq (n : ℕ) : sqrt (n * n) = n := sqrt_add_eq n (zero_le _)
1621
+ #align nat.sqrt_eq Nat.sqrt_eq
1622
+
1623
+ lemma sqrt_eq' (n : ℕ) : sqrt (n ^ 2 ) = n := sqrt_add_eq' n (zero_le _)
1624
+ #align nat.sqrt_eq' Nat.sqrt_eq'
1625
+
1626
+ lemma sqrt_succ_le_succ_sqrt (n : ℕ) : sqrt n.succ ≤ n.sqrt.succ :=
1627
+ le_of_lt_succ <| sqrt_lt.2 <| lt_succ_of_le <|
1628
+ succ_le_succ <| le_trans (sqrt_le_add n) <| Nat.add_le_add_right
1629
+ (by refine' add_le_add (Nat.mul_le_mul_right _ _) _ <;> exact Nat.le_add_right _ 2 ) _
1630
+ #align nat.sqrt_succ_le_succ_sqrt Nat.sqrt_succ_le_succ_sqrt
1631
+
1632
+ lemma exists_mul_self (x : ℕ) : (∃ n, n * n = x) ↔ sqrt x * sqrt x = x :=
1633
+ ⟨fun ⟨n, hn⟩ ↦ by rw [← hn, sqrt_eq], fun h ↦ ⟨sqrt x, h⟩⟩
1634
+ #align nat.exists_mul_self Nat.exists_mul_self
1635
+
1636
+ lemma exists_mul_self' (x : ℕ) : (∃ n, n ^ 2 = x) ↔ sqrt x ^ 2 = x := by
1637
+ simpa only [Nat.pow_two] using exists_mul_self x
1638
+ #align nat.exists_mul_self' Nat.exists_mul_self'
1639
+
1640
+ lemma sqrt_mul_sqrt_lt_succ (n : ℕ) : sqrt n * sqrt n < n + 1 :=
1641
+ Nat.lt_succ_iff.mpr (sqrt_le _)
1642
+ #align nat.sqrt_mul_sqrt_lt_succ Nat.sqrt_mul_sqrt_lt_succ
1643
+
1644
+ lemma sqrt_mul_sqrt_lt_succ' (n : ℕ) : sqrt n ^ 2 < n + 1 :=
1645
+ Nat.lt_succ_iff.mpr (sqrt_le' _)
1646
+ #align nat.sqrt_mul_sqrt_lt_succ' Nat.sqrt_mul_sqrt_lt_succ'
1647
+
1648
+ lemma succ_le_succ_sqrt (n : ℕ) : n + 1 ≤ (sqrt n + 1 ) * (sqrt n + 1 ) :=
1649
+ le_of_pred_lt (lt_succ_sqrt _)
1650
+ #align nat.succ_le_succ_sqrt Nat.succ_le_succ_sqrt
1651
+
1652
+ lemma succ_le_succ_sqrt' (n : ℕ) : n + 1 ≤ (sqrt n + 1 ) ^ 2 :=
1653
+ le_of_pred_lt (lt_succ_sqrt' _)
1654
+ #align nat.succ_le_succ_sqrt' Nat.succ_le_succ_sqrt'
1655
+
1656
+ /-- There are no perfect squares strictly between m² and (m+1)² -/
1657
+ lemma not_exists_sq (hl : m * m < n) (hr : n < (m + 1 ) * (m + 1 )) : ¬∃ t, t * t = n := by
1658
+ rintro ⟨t, rfl⟩
1659
+ have h1 : m < t := Nat.mul_self_lt_mul_self_iff.1 hl
1660
+ have h2 : t < m + 1 := Nat.mul_self_lt_mul_self_iff.1 hr
1661
+ exact (not_lt_of_ge <| le_of_lt_succ h2) h1
1662
+ #align nat.not_exists_sq Nat.not_exists_sq
1663
+
1664
+ lemma not_exists_sq' : m ^ 2 < n → n < (m + 1 ) ^ 2 → ¬∃ t, t ^ 2 = n := by
1665
+ simpa only [Nat.pow_two] using not_exists_sq
1666
+ #align nat.not_exists_sq' Nat.not_exists_sq'
1667
+
1422
1668
/-! ### `find` -/
1423
1669
1424
1670
section Find
0 commit comments