analysis.special_functions.log.baseMathlib.Analysis.SpecialFunctions.Log.Base

This file has been ported!

Changes since the initial port

The following section lists changes to this file in mathlib3 and mathlib4 that occured after the initial port. Most recent changes are shown first. Hovering over a commit will show all commits associated with the same mathlib3 commit.

Changes in mathlib3

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(last sync)

feat(analysis/special_functions/log/base): add real.logb_mul_base (#18979)

Add proof that states logb (a * b) c = ((logb a c)⁻¹ + (logb b c)⁻¹)⁻¹.

Co-authored-by: Eric Wieser <wieser.eric@gmail.com>

Diff
@@ -55,6 +55,39 @@ by simp_rw [logb, log_div hx hy, sub_div]
 
 @[simp] lemma logb_inv (x : ℝ) : logb b (x⁻¹) = -logb b x := by simp [logb, neg_div]
 
+lemma inv_logb (a b : ℝ) : (logb a b)⁻¹ = logb b a := by simp_rw [logb, inv_div]
+
+theorem inv_logb_mul_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
+  (logb (a * b) c)⁻¹ = (logb a c)⁻¹ + (logb b c)⁻¹ :=
+by simp_rw inv_logb; exact logb_mul h₁ h₂
+
+theorem inv_logb_div_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
+  (logb (a / b) c)⁻¹ = (logb a c)⁻¹ - (logb b c)⁻¹ :=
+by simp_rw inv_logb; exact logb_div h₁ h₂
+
+theorem logb_mul_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
+  logb (a * b) c = ((logb a c)⁻¹ + (logb b c)⁻¹)⁻¹ :=
+by rw [←inv_logb_mul_base h₁ h₂ c, inv_inv]
+
+theorem logb_div_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
+  logb (a / b) c = ((logb a c)⁻¹ - (logb b c)⁻¹)⁻¹ :=
+by rw [←inv_logb_div_base h₁ h₂ c, inv_inv]
+
+theorem mul_logb {a b c : ℝ} (h₁ : b ≠ 0) (h₂ : b ≠ 1) (h₃ : b ≠ -1) :
+  logb a b * logb b c = logb a c :=
+begin
+  unfold logb,
+  rw [mul_comm, div_mul_div_cancel _ (log_ne_zero.mpr ⟨h₁, h₂, h₃⟩)],
+end
+
+theorem div_logb {a b c : ℝ} (h₁ : c ≠ 0) (h₂ : c ≠ 1) (h₃ : c ≠ -1) :
+  logb a c / logb b c = logb a b :=
+begin
+  unfold logb,
+  -- TODO: div_div_div_cancel_left is missing for `group_with_zero`,
+  rw [div_div_div_eq, mul_comm, mul_div_mul_right _ _ (log_ne_zero.mpr ⟨h₁, h₂, h₃⟩)],
+end
+
 section b_pos_and_ne_one
 
 variable (b_pos : 0 < b)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(no changes)

(first ported)

Changes in mathlib3port

mathlib3
mathlib3port
Diff
@@ -517,8 +517,8 @@ theorem tendsto_logb_atTop_of_base_lt_one : Tendsto (logb b) atTop atBot :=
 
 end BPosAndBLtOne
 
-#print Real.floor_logb_nat_cast /-
-theorem floor_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : ⌊logb b r⌋ = Int.log b r :=
+#print Real.floor_logb_natCast /-
+theorem floor_logb_natCast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : ⌊logb b r⌋ = Int.log b r :=
   by
   obtain rfl | hr := hr.eq_or_lt
   · rw [logb_zero, Int.log_zero_right, Int.floor_zero]
@@ -529,11 +529,11 @@ theorem floor_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : 
     exact rpow_le_rpow_of_exponent_le hb1'.le (Int.floor_le _)
   · rw [Int.le_floor, le_logb_iff_rpow_le hb1' hr, rpow_int_cast]
     exact Int.zpow_log_le_self hb hr
-#align real.floor_logb_nat_cast Real.floor_logb_nat_cast
+#align real.floor_logb_nat_cast Real.floor_logb_natCast
 -/
 
-#print Real.ceil_logb_nat_cast /-
-theorem ceil_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : ⌈logb b r⌉ = Int.clog b r :=
+#print Real.ceil_logb_natCast /-
+theorem ceil_logb_natCast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : ⌈logb b r⌉ = Int.clog b r :=
   by
   obtain rfl | hr := hr.eq_or_lt
   · rw [logb_zero, Int.clog_zero_right, Int.ceil_zero]
@@ -544,7 +544,7 @@ theorem ceil_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : ⌈
   · rw [← Int.le_zpow_iff_clog_le hb hr, ← rpow_int_cast b]
     refine' (rpow_logb (zero_lt_one.trans hb1') hb1'.ne' hr).symm.trans_le _
     exact rpow_le_rpow_of_exponent_le hb1'.le (Int.le_ceil _)
-#align real.ceil_logb_nat_cast Real.ceil_logb_nat_cast
+#align real.ceil_logb_nat_cast Real.ceil_logb_natCast
 -/
 
 #print Real.logb_eq_zero /-
Diff
@@ -565,7 +565,7 @@ theorem logb_prod {α : Type _} (s : Finset α) (f : α → ℝ) (hf : ∀ x ∈
   classical
   induction' s using Finset.induction_on with a s ha ih
   · simp
-  simp only [Finset.mem_insert, forall_eq_or_imp] at hf 
+  simp only [Finset.mem_insert, forall_eq_or_imp] at hf
   simp [ha, ih hf.2, logb_mul hf.1 (Finset.prod_ne_zero_iff.2 hf.2)]
 #align real.logb_prod Real.logb_prod
 -/
Diff
@@ -561,7 +561,12 @@ open scoped BigOperators
 
 #print Real.logb_prod /-
 theorem logb_prod {α : Type _} (s : Finset α) (f : α → ℝ) (hf : ∀ x ∈ s, f x ≠ 0) :
-    logb b (∏ i in s, f i) = ∑ i in s, logb b (f i) := by classical
+    logb b (∏ i in s, f i) = ∑ i in s, logb b (f i) := by
+  classical
+  induction' s using Finset.induction_on with a s ha ih
+  · simp
+  simp only [Finset.mem_insert, forall_eq_or_imp] at hf 
+  simp [ha, ih hf.2, logb_mul hf.1 (Finset.prod_ne_zero_iff.2 hf.2)]
 #align real.logb_prod Real.logb_prod
 -/
 
Diff
@@ -561,12 +561,7 @@ open scoped BigOperators
 
 #print Real.logb_prod /-
 theorem logb_prod {α : Type _} (s : Finset α) (f : α → ℝ) (hf : ∀ x ∈ s, f x ≠ 0) :
-    logb b (∏ i in s, f i) = ∑ i in s, logb b (f i) := by
-  classical
-  induction' s using Finset.induction_on with a s ha ih
-  · simp
-  simp only [Finset.mem_insert, forall_eq_or_imp] at hf 
-  simp [ha, ih hf.2, logb_mul hf.1 (Finset.prod_ne_zero_iff.2 hf.2)]
+    logb b (∏ i in s, f i) = ∑ i in s, logb b (f i) := by classical
 #align real.logb_prod Real.logb_prod
 -/
 
Diff
@@ -3,8 +3,8 @@ Copyright (c) 2022 Bolton Bailey. All rights reserved.
 Released under Apache 2.0 license as described in the file LICENSE.
 Authors: Bolton Bailey, Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne
 -/
-import Mathbin.Analysis.SpecialFunctions.Pow.Real
-import Mathbin.Data.Int.Log
+import Analysis.SpecialFunctions.Pow.Real
+import Data.Int.Log
 
 #align_import analysis.special_functions.log.base from "leanprover-community/mathlib"@"f23a09ce6d3f367220dc3cecad6b7eb69eb01690"
 
Diff
@@ -211,7 +211,7 @@ theorem range_logb : range (logb b) = univ :=
 theorem surjOn_logb' : SurjOn (logb b) (Iio 0) univ :=
   by
   intro x x_in_univ
-  use -b ^ x
+  use-b ^ x
   constructor
   · simp only [Right.neg_neg_iff, Set.mem_Iio]; apply rpow_pos_of_pos b_pos
   · rw [logb_neg_eq_logb, logb_rpow b_pos b_ne_one]
Diff
@@ -2,15 +2,12 @@
 Copyright (c) 2022 Bolton Bailey. All rights reserved.
 Released under Apache 2.0 license as described in the file LICENSE.
 Authors: Bolton Bailey, Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne
-
-! This file was ported from Lean 3 source module analysis.special_functions.log.base
-! leanprover-community/mathlib commit f23a09ce6d3f367220dc3cecad6b7eb69eb01690
-! Please do not edit these lines, except to modify the commit id
-! if you have ported upstream changes.
 -/
 import Mathbin.Analysis.SpecialFunctions.Pow.Real
 import Mathbin.Data.Int.Log
 
+#align_import analysis.special_functions.log.base from "leanprover-community/mathlib"@"f23a09ce6d3f367220dc3cecad6b7eb69eb01690"
+
 /-!
 # Real logarithm base `b`
 
Diff
@@ -69,7 +69,7 @@ theorem logb_one : logb b 1 = 0 := by simp [logb]
 
 #print Real.logb_abs /-
 @[simp]
-theorem logb_abs (x : ℝ) : logb b (|x|) = logb b x := by rw [logb, logb, log_abs]
+theorem logb_abs (x : ℝ) : logb b |x| = logb b x := by rw [logb, logb, log_abs]
 #align real.logb_abs Real.logb_abs
 -/
 
Diff
@@ -98,39 +98,53 @@ theorem logb_inv (x : ℝ) : logb b x⁻¹ = -logb b x := by simp [logb, neg_div
 #align real.logb_inv Real.logb_inv
 -/
 
+#print Real.inv_logb /-
 theorem inv_logb (a b : ℝ) : (logb a b)⁻¹ = logb b a := by simp_rw [logb, inv_div]
 #align real.inv_logb Real.inv_logb
+-/
 
+#print Real.inv_logb_mul_base /-
 theorem inv_logb_mul_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
     (logb (a * b) c)⁻¹ = (logb a c)⁻¹ + (logb b c)⁻¹ := by
   simp_rw [inv_logb] <;> exact logb_mul h₁ h₂
 #align real.inv_logb_mul_base Real.inv_logb_mul_base
+-/
 
+#print Real.inv_logb_div_base /-
 theorem inv_logb_div_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
     (logb (a / b) c)⁻¹ = (logb a c)⁻¹ - (logb b c)⁻¹ := by
   simp_rw [inv_logb] <;> exact logb_div h₁ h₂
 #align real.inv_logb_div_base Real.inv_logb_div_base
+-/
 
+#print Real.logb_mul_base /-
 theorem logb_mul_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
     logb (a * b) c = ((logb a c)⁻¹ + (logb b c)⁻¹)⁻¹ := by rw [← inv_logb_mul_base h₁ h₂ c, inv_inv]
 #align real.logb_mul_base Real.logb_mul_base
+-/
 
+#print Real.logb_div_base /-
 theorem logb_div_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
     logb (a / b) c = ((logb a c)⁻¹ - (logb b c)⁻¹)⁻¹ := by rw [← inv_logb_div_base h₁ h₂ c, inv_inv]
 #align real.logb_div_base Real.logb_div_base
+-/
 
+#print Real.mul_logb /-
 theorem mul_logb {a b c : ℝ} (h₁ : b ≠ 0) (h₂ : b ≠ 1) (h₃ : b ≠ -1) :
     logb a b * logb b c = logb a c := by
   unfold logb
   rw [mul_comm, div_mul_div_cancel _ (log_ne_zero.mpr ⟨h₁, h₂, h₃⟩)]
 #align real.mul_logb Real.mul_logb
+-/
 
+#print Real.div_logb /-
 theorem div_logb {a b c : ℝ} (h₁ : c ≠ 0) (h₂ : c ≠ 1) (h₃ : c ≠ -1) :
     logb a c / logb b c = logb a b := by
   unfold logb
   -- TODO: div_div_div_cancel_left is missing for `group_with_zero`,
   rw [div_div_div_eq, mul_comm, mul_div_mul_right _ _ (log_ne_zero.mpr ⟨h₁, h₂, h₃⟩)]
 #align real.div_logb Real.div_logb
+-/
 
 section BPosAndNeOne
 
Diff
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
 Authors: Bolton Bailey, Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne
 
 ! This file was ported from Lean 3 source module analysis.special_functions.log.base
-! leanprover-community/mathlib commit 1b0a28e1c93409dbf6d69526863cd9984ef652ce
+! leanprover-community/mathlib commit f23a09ce6d3f367220dc3cecad6b7eb69eb01690
 ! Please do not edit these lines, except to modify the commit id
 ! if you have ported upstream changes.
 -/
@@ -98,6 +98,40 @@ theorem logb_inv (x : ℝ) : logb b x⁻¹ = -logb b x := by simp [logb, neg_div
 #align real.logb_inv Real.logb_inv
 -/
 
+theorem inv_logb (a b : ℝ) : (logb a b)⁻¹ = logb b a := by simp_rw [logb, inv_div]
+#align real.inv_logb Real.inv_logb
+
+theorem inv_logb_mul_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
+    (logb (a * b) c)⁻¹ = (logb a c)⁻¹ + (logb b c)⁻¹ := by
+  simp_rw [inv_logb] <;> exact logb_mul h₁ h₂
+#align real.inv_logb_mul_base Real.inv_logb_mul_base
+
+theorem inv_logb_div_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
+    (logb (a / b) c)⁻¹ = (logb a c)⁻¹ - (logb b c)⁻¹ := by
+  simp_rw [inv_logb] <;> exact logb_div h₁ h₂
+#align real.inv_logb_div_base Real.inv_logb_div_base
+
+theorem logb_mul_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
+    logb (a * b) c = ((logb a c)⁻¹ + (logb b c)⁻¹)⁻¹ := by rw [← inv_logb_mul_base h₁ h₂ c, inv_inv]
+#align real.logb_mul_base Real.logb_mul_base
+
+theorem logb_div_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
+    logb (a / b) c = ((logb a c)⁻¹ - (logb b c)⁻¹)⁻¹ := by rw [← inv_logb_div_base h₁ h₂ c, inv_inv]
+#align real.logb_div_base Real.logb_div_base
+
+theorem mul_logb {a b c : ℝ} (h₁ : b ≠ 0) (h₂ : b ≠ 1) (h₃ : b ≠ -1) :
+    logb a b * logb b c = logb a c := by
+  unfold logb
+  rw [mul_comm, div_mul_div_cancel _ (log_ne_zero.mpr ⟨h₁, h₂, h₃⟩)]
+#align real.mul_logb Real.mul_logb
+
+theorem div_logb {a b c : ℝ} (h₁ : c ≠ 0) (h₂ : c ≠ 1) (h₃ : c ≠ -1) :
+    logb a c / logb b c = logb a b := by
+  unfold logb
+  -- TODO: div_div_div_cancel_left is missing for `group_with_zero`,
+  rw [div_div_div_eq, mul_comm, mul_div_mul_right _ _ (log_ne_zero.mpr ⟨h₁, h₂, h₃⟩)]
+#align real.div_logb Real.div_logb
+
 section BPosAndNeOne
 
 variable (b_pos : 0 < b)
Diff
@@ -49,38 +49,54 @@ noncomputable def logb (b x : ℝ) : ℝ :=
 #align real.logb Real.logb
 -/
 
+#print Real.log_div_log /-
 theorem log_div_log : log x / log b = logb b x :=
   rfl
 #align real.log_div_log Real.log_div_log
+-/
 
+#print Real.logb_zero /-
 @[simp]
 theorem logb_zero : logb b 0 = 0 := by simp [logb]
 #align real.logb_zero Real.logb_zero
+-/
 
+#print Real.logb_one /-
 @[simp]
 theorem logb_one : logb b 1 = 0 := by simp [logb]
 #align real.logb_one Real.logb_one
+-/
 
+#print Real.logb_abs /-
 @[simp]
 theorem logb_abs (x : ℝ) : logb b (|x|) = logb b x := by rw [logb, logb, log_abs]
 #align real.logb_abs Real.logb_abs
+-/
 
+#print Real.logb_neg_eq_logb /-
 @[simp]
 theorem logb_neg_eq_logb (x : ℝ) : logb b (-x) = logb b x := by
   rw [← logb_abs x, ← logb_abs (-x), abs_neg]
 #align real.logb_neg_eq_logb Real.logb_neg_eq_logb
+-/
 
+#print Real.logb_mul /-
 theorem logb_mul (hx : x ≠ 0) (hy : y ≠ 0) : logb b (x * y) = logb b x + logb b y := by
   simp_rw [logb, log_mul hx hy, add_div]
 #align real.logb_mul Real.logb_mul
+-/
 
+#print Real.logb_div /-
 theorem logb_div (hx : x ≠ 0) (hy : y ≠ 0) : logb b (x / y) = logb b x - logb b y := by
   simp_rw [logb, log_div hx hy, sub_div]
 #align real.logb_div Real.logb_div
+-/
 
+#print Real.logb_inv /-
 @[simp]
 theorem logb_inv (x : ℝ) : logb b x⁻¹ = -logb b x := by simp [logb, neg_div]
 #align real.logb_inv Real.logb_inv
+-/
 
 section BPosAndNeOne
 
@@ -88,21 +104,22 @@ variable (b_pos : 0 < b)
 
 variable (b_ne_one : b ≠ 1)
 
-include b_pos b_ne_one
-
 private theorem log_b_ne_zero : log b ≠ 0 :=
   by
   have b_ne_zero : b ≠ 0; linarith
   have b_ne_minus_one : b ≠ -1; linarith
   simp [b_ne_one, b_ne_zero, b_ne_minus_one]
 
+#print Real.logb_rpow /-
 @[simp]
 theorem logb_rpow : logb b (b ^ x) = x :=
   by
   rw [logb, div_eq_iff, log_rpow b_pos]
   exact log_b_ne_zero b_pos b_ne_one
 #align real.logb_rpow Real.logb_rpow
+-/
 
+#print Real.rpow_logb_eq_abs /-
 theorem rpow_logb_eq_abs (hx : x ≠ 0) : b ^ logb b x = |x| :=
   by
   apply log_inj_on_pos
@@ -112,28 +129,40 @@ theorem rpow_logb_eq_abs (hx : x ≠ 0) : b ^ logb b x = |x| :=
   rw [log_rpow b_pos, logb, log_abs]
   field_simp [log_b_ne_zero b_pos b_ne_one]
 #align real.rpow_logb_eq_abs Real.rpow_logb_eq_abs
+-/
 
+#print Real.rpow_logb /-
 @[simp]
 theorem rpow_logb (hx : 0 < x) : b ^ logb b x = x := by rw [rpow_logb_eq_abs b_pos b_ne_one hx.ne'];
   exact abs_of_pos hx
 #align real.rpow_logb Real.rpow_logb
+-/
 
+#print Real.rpow_logb_of_neg /-
 theorem rpow_logb_of_neg (hx : x < 0) : b ^ logb b x = -x := by
   rw [rpow_logb_eq_abs b_pos b_ne_one (ne_of_lt hx)]; exact abs_of_neg hx
 #align real.rpow_logb_of_neg Real.rpow_logb_of_neg
+-/
 
+#print Real.surjOn_logb /-
 theorem surjOn_logb : SurjOn (logb b) (Ioi 0) univ := fun x _ =>
   ⟨rpow b x, rpow_pos_of_pos b_pos x, logb_rpow b_pos b_ne_one⟩
 #align real.surj_on_logb Real.surjOn_logb
+-/
 
+#print Real.logb_surjective /-
 theorem logb_surjective : Surjective (logb b) := fun x => ⟨b ^ x, logb_rpow b_pos b_ne_one⟩
 #align real.logb_surjective Real.logb_surjective
+-/
 
+#print Real.range_logb /-
 @[simp]
 theorem range_logb : range (logb b) = univ :=
   (logb_surjective b_pos b_ne_one).range_eq
 #align real.range_logb Real.range_logb
+-/
 
+#print Real.surjOn_logb' /-
 theorem surjOn_logb' : SurjOn (logb b) (Iio 0) univ :=
   by
   intro x x_in_univ
@@ -142,6 +171,7 @@ theorem surjOn_logb' : SurjOn (logb b) (Iio 0) univ :=
   · simp only [Right.neg_neg_iff, Set.mem_Iio]; apply rpow_pos_of_pos b_pos
   · rw [logb_neg_eq_logb, logb_rpow b_pos b_ne_one]
 #align real.surj_on_logb' Real.surjOn_logb'
+-/
 
 end BPosAndNeOne
 
@@ -149,85 +179,118 @@ section OneLtB
 
 variable (hb : 1 < b)
 
-include hb
-
 private theorem b_pos : 0 < b := by linarith
 
 private theorem b_ne_one : b ≠ 1 := by linarith
 
+#print Real.logb_le_logb /-
 @[simp]
 theorem logb_le_logb (h : 0 < x) (h₁ : 0 < y) : logb b x ≤ logb b y ↔ x ≤ y := by
   rw [logb, logb, div_le_div_right (log_pos hb), log_le_log h h₁]
 #align real.logb_le_logb Real.logb_le_logb
+-/
 
+#print Real.logb_lt_logb /-
 theorem logb_lt_logb (hx : 0 < x) (hxy : x < y) : logb b x < logb b y := by
   rw [logb, logb, div_lt_div_right (log_pos hb)]; exact log_lt_log hx hxy
 #align real.logb_lt_logb Real.logb_lt_logb
+-/
 
+#print Real.logb_lt_logb_iff /-
 @[simp]
 theorem logb_lt_logb_iff (hx : 0 < x) (hy : 0 < y) : logb b x < logb b y ↔ x < y := by
   rw [logb, logb, div_lt_div_right (log_pos hb)]; exact log_lt_log_iff hx hy
 #align real.logb_lt_logb_iff Real.logb_lt_logb_iff
+-/
 
+#print Real.logb_le_iff_le_rpow /-
 theorem logb_le_iff_le_rpow (hx : 0 < x) : logb b x ≤ y ↔ x ≤ b ^ y := by
   rw [← rpow_le_rpow_left_iff hb, rpow_logb (b_pos hb) (b_ne_one hb) hx]
 #align real.logb_le_iff_le_rpow Real.logb_le_iff_le_rpow
+-/
 
+#print Real.logb_lt_iff_lt_rpow /-
 theorem logb_lt_iff_lt_rpow (hx : 0 < x) : logb b x < y ↔ x < b ^ y := by
   rw [← rpow_lt_rpow_left_iff hb, rpow_logb (b_pos hb) (b_ne_one hb) hx]
 #align real.logb_lt_iff_lt_rpow Real.logb_lt_iff_lt_rpow
+-/
 
+#print Real.le_logb_iff_rpow_le /-
 theorem le_logb_iff_rpow_le (hy : 0 < y) : x ≤ logb b y ↔ b ^ x ≤ y := by
   rw [← rpow_le_rpow_left_iff hb, rpow_logb (b_pos hb) (b_ne_one hb) hy]
 #align real.le_logb_iff_rpow_le Real.le_logb_iff_rpow_le
+-/
 
+#print Real.lt_logb_iff_rpow_lt /-
 theorem lt_logb_iff_rpow_lt (hy : 0 < y) : x < logb b y ↔ b ^ x < y := by
   rw [← rpow_lt_rpow_left_iff hb, rpow_logb (b_pos hb) (b_ne_one hb) hy]
 #align real.lt_logb_iff_rpow_lt Real.lt_logb_iff_rpow_lt
+-/
 
+#print Real.logb_pos_iff /-
 theorem logb_pos_iff (hx : 0 < x) : 0 < logb b x ↔ 1 < x := by rw [← @logb_one b];
   rw [logb_lt_logb_iff hb zero_lt_one hx]
 #align real.logb_pos_iff Real.logb_pos_iff
+-/
 
+#print Real.logb_pos /-
 theorem logb_pos (hx : 1 < x) : 0 < logb b x := by rw [logb_pos_iff hb (lt_trans zero_lt_one hx)];
   exact hx
 #align real.logb_pos Real.logb_pos
+-/
 
+#print Real.logb_neg_iff /-
 theorem logb_neg_iff (h : 0 < x) : logb b x < 0 ↔ x < 1 := by rw [← logb_one];
   exact logb_lt_logb_iff hb h zero_lt_one
 #align real.logb_neg_iff Real.logb_neg_iff
+-/
 
+#print Real.logb_neg /-
 theorem logb_neg (h0 : 0 < x) (h1 : x < 1) : logb b x < 0 :=
   (logb_neg_iff hb h0).2 h1
 #align real.logb_neg Real.logb_neg
+-/
 
+#print Real.logb_nonneg_iff /-
 theorem logb_nonneg_iff (hx : 0 < x) : 0 ≤ logb b x ↔ 1 ≤ x := by
   rw [← not_lt, logb_neg_iff hb hx, not_lt]
 #align real.logb_nonneg_iff Real.logb_nonneg_iff
+-/
 
+#print Real.logb_nonneg /-
 theorem logb_nonneg (hx : 1 ≤ x) : 0 ≤ logb b x :=
   (logb_nonneg_iff hb (zero_lt_one.trans_le hx)).2 hx
 #align real.logb_nonneg Real.logb_nonneg
+-/
 
+#print Real.logb_nonpos_iff /-
 theorem logb_nonpos_iff (hx : 0 < x) : logb b x ≤ 0 ↔ x ≤ 1 := by
   rw [← not_lt, logb_pos_iff hb hx, not_lt]
 #align real.logb_nonpos_iff Real.logb_nonpos_iff
+-/
 
+#print Real.logb_nonpos_iff' /-
 theorem logb_nonpos_iff' (hx : 0 ≤ x) : logb b x ≤ 0 ↔ x ≤ 1 :=
   by
   rcases hx.eq_or_lt with (rfl | hx)
   · simp [le_refl, zero_le_one]
   exact logb_nonpos_iff hb hx
 #align real.logb_nonpos_iff' Real.logb_nonpos_iff'
+-/
 
+#print Real.logb_nonpos /-
 theorem logb_nonpos (hx : 0 ≤ x) (h'x : x ≤ 1) : logb b x ≤ 0 :=
   (logb_nonpos_iff' hb hx).2 h'x
 #align real.logb_nonpos Real.logb_nonpos
+-/
 
+#print Real.strictMonoOn_logb /-
 theorem strictMonoOn_logb : StrictMonoOn (logb b) (Set.Ioi 0) := fun x hx y hy hxy =>
   logb_lt_logb hb hx hxy
 #align real.strict_mono_on_logb Real.strictMonoOn_logb
+-/
 
+#print Real.strictAntiOn_logb /-
 theorem strictAntiOn_logb : StrictAntiOn (logb b) (Set.Iio 0) :=
   by
   rintro x (hx : x < 0) y (hy : y < 0) hxy
@@ -235,22 +298,31 @@ theorem strictAntiOn_logb : StrictAntiOn (logb b) (Set.Iio 0) :=
   refine' logb_lt_logb hb (abs_pos.2 hy.ne) _
   rwa [abs_of_neg hy, abs_of_neg hx, neg_lt_neg_iff]
 #align real.strict_anti_on_logb Real.strictAntiOn_logb
+-/
 
+#print Real.logb_injOn_pos /-
 theorem logb_injOn_pos : Set.InjOn (logb b) (Set.Ioi 0) :=
   (strictMonoOn_logb hb).InjOn
 #align real.logb_inj_on_pos Real.logb_injOn_pos
+-/
 
+#print Real.eq_one_of_pos_of_logb_eq_zero /-
 theorem eq_one_of_pos_of_logb_eq_zero (h₁ : 0 < x) (h₂ : logb b x = 0) : x = 1 :=
   logb_injOn_pos hb (Set.mem_Ioi.2 h₁) (Set.mem_Ioi.2 zero_lt_one) (h₂.trans Real.logb_one.symm)
 #align real.eq_one_of_pos_of_logb_eq_zero Real.eq_one_of_pos_of_logb_eq_zero
+-/
 
+#print Real.logb_ne_zero_of_pos_of_ne_one /-
 theorem logb_ne_zero_of_pos_of_ne_one (hx_pos : 0 < x) (hx : x ≠ 1) : logb b x ≠ 0 :=
   mt (eq_one_of_pos_of_logb_eq_zero hb hx_pos) hx
 #align real.logb_ne_zero_of_pos_of_ne_one Real.logb_ne_zero_of_pos_of_ne_one
+-/
 
+#print Real.tendsto_logb_atTop /-
 theorem tendsto_logb_atTop : Tendsto (logb b) atTop atTop :=
   Tendsto.atTop_div_const (log_pos hb) tendsto_log_atTop
 #align real.tendsto_logb_at_top Real.tendsto_logb_atTop
+-/
 
 end OneLtB
 
@@ -260,74 +332,101 @@ variable (b_pos : 0 < b)
 
 variable (b_lt_one : b < 1)
 
-include b_lt_one
-
 private theorem b_ne_one : b ≠ 1 := by linarith
 
-include b_pos
-
+#print Real.logb_le_logb_of_base_lt_one /-
 @[simp]
 theorem logb_le_logb_of_base_lt_one (h : 0 < x) (h₁ : 0 < y) : logb b x ≤ logb b y ↔ y ≤ x := by
   rw [logb, logb, div_le_div_right_of_neg (log_neg b_pos b_lt_one), log_le_log h₁ h]
 #align real.logb_le_logb_of_base_lt_one Real.logb_le_logb_of_base_lt_one
+-/
 
+#print Real.logb_lt_logb_of_base_lt_one /-
 theorem logb_lt_logb_of_base_lt_one (hx : 0 < x) (hxy : x < y) : logb b y < logb b x := by
   rw [logb, logb, div_lt_div_right_of_neg (log_neg b_pos b_lt_one)]; exact log_lt_log hx hxy
 #align real.logb_lt_logb_of_base_lt_one Real.logb_lt_logb_of_base_lt_one
+-/
 
+#print Real.logb_lt_logb_iff_of_base_lt_one /-
 @[simp]
 theorem logb_lt_logb_iff_of_base_lt_one (hx : 0 < x) (hy : 0 < y) : logb b x < logb b y ↔ y < x :=
   by rw [logb, logb, div_lt_div_right_of_neg (log_neg b_pos b_lt_one)]; exact log_lt_log_iff hy hx
 #align real.logb_lt_logb_iff_of_base_lt_one Real.logb_lt_logb_iff_of_base_lt_one
+-/
 
+#print Real.logb_le_iff_le_rpow_of_base_lt_one /-
 theorem logb_le_iff_le_rpow_of_base_lt_one (hx : 0 < x) : logb b x ≤ y ↔ b ^ y ≤ x := by
   rw [← rpow_le_rpow_left_iff_of_base_lt_one b_pos b_lt_one, rpow_logb b_pos (b_ne_one b_lt_one) hx]
 #align real.logb_le_iff_le_rpow_of_base_lt_one Real.logb_le_iff_le_rpow_of_base_lt_one
+-/
 
+#print Real.logb_lt_iff_lt_rpow_of_base_lt_one /-
 theorem logb_lt_iff_lt_rpow_of_base_lt_one (hx : 0 < x) : logb b x < y ↔ b ^ y < x := by
   rw [← rpow_lt_rpow_left_iff_of_base_lt_one b_pos b_lt_one, rpow_logb b_pos (b_ne_one b_lt_one) hx]
 #align real.logb_lt_iff_lt_rpow_of_base_lt_one Real.logb_lt_iff_lt_rpow_of_base_lt_one
+-/
 
+#print Real.le_logb_iff_rpow_le_of_base_lt_one /-
 theorem le_logb_iff_rpow_le_of_base_lt_one (hy : 0 < y) : x ≤ logb b y ↔ y ≤ b ^ x := by
   rw [← rpow_le_rpow_left_iff_of_base_lt_one b_pos b_lt_one, rpow_logb b_pos (b_ne_one b_lt_one) hy]
 #align real.le_logb_iff_rpow_le_of_base_lt_one Real.le_logb_iff_rpow_le_of_base_lt_one
+-/
 
+#print Real.lt_logb_iff_rpow_lt_of_base_lt_one /-
 theorem lt_logb_iff_rpow_lt_of_base_lt_one (hy : 0 < y) : x < logb b y ↔ y < b ^ x := by
   rw [← rpow_lt_rpow_left_iff_of_base_lt_one b_pos b_lt_one, rpow_logb b_pos (b_ne_one b_lt_one) hy]
 #align real.lt_logb_iff_rpow_lt_of_base_lt_one Real.lt_logb_iff_rpow_lt_of_base_lt_one
+-/
 
+#print Real.logb_pos_iff_of_base_lt_one /-
 theorem logb_pos_iff_of_base_lt_one (hx : 0 < x) : 0 < logb b x ↔ x < 1 := by
   rw [← @logb_one b, logb_lt_logb_iff_of_base_lt_one b_pos b_lt_one zero_lt_one hx]
 #align real.logb_pos_iff_of_base_lt_one Real.logb_pos_iff_of_base_lt_one
+-/
 
+#print Real.logb_pos_of_base_lt_one /-
 theorem logb_pos_of_base_lt_one (hx : 0 < x) (hx' : x < 1) : 0 < logb b x := by
   rw [logb_pos_iff_of_base_lt_one b_pos b_lt_one hx]; exact hx'
 #align real.logb_pos_of_base_lt_one Real.logb_pos_of_base_lt_one
+-/
 
+#print Real.logb_neg_iff_of_base_lt_one /-
 theorem logb_neg_iff_of_base_lt_one (h : 0 < x) : logb b x < 0 ↔ 1 < x := by
   rw [← @logb_one b, logb_lt_logb_iff_of_base_lt_one b_pos b_lt_one h zero_lt_one]
 #align real.logb_neg_iff_of_base_lt_one Real.logb_neg_iff_of_base_lt_one
+-/
 
+#print Real.logb_neg_of_base_lt_one /-
 theorem logb_neg_of_base_lt_one (h1 : 1 < x) : logb b x < 0 :=
   (logb_neg_iff_of_base_lt_one b_pos b_lt_one (lt_trans zero_lt_one h1)).2 h1
 #align real.logb_neg_of_base_lt_one Real.logb_neg_of_base_lt_one
+-/
 
+#print Real.logb_nonneg_iff_of_base_lt_one /-
 theorem logb_nonneg_iff_of_base_lt_one (hx : 0 < x) : 0 ≤ logb b x ↔ x ≤ 1 := by
   rw [← not_lt, logb_neg_iff_of_base_lt_one b_pos b_lt_one hx, not_lt]
 #align real.logb_nonneg_iff_of_base_lt_one Real.logb_nonneg_iff_of_base_lt_one
+-/
 
+#print Real.logb_nonneg_of_base_lt_one /-
 theorem logb_nonneg_of_base_lt_one (hx : 0 < x) (hx' : x ≤ 1) : 0 ≤ logb b x := by
   rw [logb_nonneg_iff_of_base_lt_one b_pos b_lt_one hx]; exact hx'
 #align real.logb_nonneg_of_base_lt_one Real.logb_nonneg_of_base_lt_one
+-/
 
+#print Real.logb_nonpos_iff_of_base_lt_one /-
 theorem logb_nonpos_iff_of_base_lt_one (hx : 0 < x) : logb b x ≤ 0 ↔ 1 ≤ x := by
   rw [← not_lt, logb_pos_iff_of_base_lt_one b_pos b_lt_one hx, not_lt]
 #align real.logb_nonpos_iff_of_base_lt_one Real.logb_nonpos_iff_of_base_lt_one
+-/
 
+#print Real.strictAntiOn_logb_of_base_lt_one /-
 theorem strictAntiOn_logb_of_base_lt_one : StrictAntiOn (logb b) (Set.Ioi 0) := fun x hx y hy hxy =>
   logb_lt_logb_of_base_lt_one b_pos b_lt_one hx hxy
 #align real.strict_anti_on_logb_of_base_lt_one Real.strictAntiOn_logb_of_base_lt_one
+-/
 
+#print Real.strictMonoOn_logb_of_base_lt_one /-
 theorem strictMonoOn_logb_of_base_lt_one : StrictMonoOn (logb b) (Set.Iio 0) :=
   by
   rintro x (hx : x < 0) y (hy : y < 0) hxy
@@ -335,20 +434,28 @@ theorem strictMonoOn_logb_of_base_lt_one : StrictMonoOn (logb b) (Set.Iio 0) :=
   refine' logb_lt_logb_of_base_lt_one b_pos b_lt_one (abs_pos.2 hy.ne) _
   rwa [abs_of_neg hy, abs_of_neg hx, neg_lt_neg_iff]
 #align real.strict_mono_on_logb_of_base_lt_one Real.strictMonoOn_logb_of_base_lt_one
+-/
 
+#print Real.logb_injOn_pos_of_base_lt_one /-
 theorem logb_injOn_pos_of_base_lt_one : Set.InjOn (logb b) (Set.Ioi 0) :=
   (strictAntiOn_logb_of_base_lt_one b_pos b_lt_one).InjOn
 #align real.logb_inj_on_pos_of_base_lt_one Real.logb_injOn_pos_of_base_lt_one
+-/
 
+#print Real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_one /-
 theorem eq_one_of_pos_of_logb_eq_zero_of_base_lt_one (h₁ : 0 < x) (h₂ : logb b x = 0) : x = 1 :=
   logb_injOn_pos_of_base_lt_one b_pos b_lt_one (Set.mem_Ioi.2 h₁) (Set.mem_Ioi.2 zero_lt_one)
     (h₂.trans Real.logb_one.symm)
 #align real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_one Real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_one
+-/
 
+#print Real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_one /-
 theorem logb_ne_zero_of_pos_of_ne_one_of_base_lt_one (hx_pos : 0 < x) (hx : x ≠ 1) : logb b x ≠ 0 :=
   mt (eq_one_of_pos_of_logb_eq_zero_of_base_lt_one b_pos b_lt_one hx_pos) hx
 #align real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_one Real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_one
+-/
 
+#print Real.tendsto_logb_atTop_of_base_lt_one /-
 theorem tendsto_logb_atTop_of_base_lt_one : Tendsto (logb b) atTop atBot :=
   by
   rw [tendsto_at_top_at_bot]
@@ -361,9 +468,11 @@ theorem tendsto_logb_atTop_of_base_lt_one : Tendsto (logb b) atTop atBot :=
   tauto
   exact lt_of_lt_of_le zero_lt_one ha
 #align real.tendsto_logb_at_top_of_base_lt_one Real.tendsto_logb_atTop_of_base_lt_one
+-/
 
 end BPosAndBLtOne
 
+#print Real.floor_logb_nat_cast /-
 theorem floor_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : ⌊logb b r⌋ = Int.log b r :=
   by
   obtain rfl | hr := hr.eq_or_lt
@@ -376,7 +485,9 @@ theorem floor_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : 
   · rw [Int.le_floor, le_logb_iff_rpow_le hb1' hr, rpow_int_cast]
     exact Int.zpow_log_le_self hb hr
 #align real.floor_logb_nat_cast Real.floor_logb_nat_cast
+-/
 
+#print Real.ceil_logb_nat_cast /-
 theorem ceil_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : ⌈logb b r⌉ = Int.clog b r :=
   by
   obtain rfl | hr := hr.eq_or_lt
@@ -389,17 +500,21 @@ theorem ceil_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : ⌈
     refine' (rpow_logb (zero_lt_one.trans hb1') hb1'.ne' hr).symm.trans_le _
     exact rpow_le_rpow_of_exponent_le hb1'.le (Int.le_ceil _)
 #align real.ceil_logb_nat_cast Real.ceil_logb_nat_cast
+-/
 
+#print Real.logb_eq_zero /-
 @[simp]
 theorem logb_eq_zero : logb b x = 0 ↔ b = 0 ∨ b = 1 ∨ b = -1 ∨ x = 0 ∨ x = 1 ∨ x = -1 :=
   by
   simp_rw [logb, div_eq_zero_iff, log_eq_zero]
   tauto
 #align real.logb_eq_zero Real.logb_eq_zero
+-/
 
 -- TODO add other limits and continuous API lemmas analogous to those in log.lean
 open scoped BigOperators
 
+#print Real.logb_prod /-
 theorem logb_prod {α : Type _} (s : Finset α) (f : α → ℝ) (hf : ∀ x ∈ s, f x ≠ 0) :
     logb b (∏ i in s, f i) = ∑ i in s, logb b (f i) := by
   classical
@@ -408,6 +523,7 @@ theorem logb_prod {α : Type _} (s : Finset α) (f : α → ℝ) (hf : ∀ x ∈
   simp only [Finset.mem_insert, forall_eq_or_imp] at hf 
   simp [ha, ih hf.2, logb_mul hf.1 (Finset.prod_ne_zero_iff.2 hf.2)]
 #align real.logb_prod Real.logb_prod
+-/
 
 end Real
 
Diff
@@ -403,10 +403,10 @@ open scoped BigOperators
 theorem logb_prod {α : Type _} (s : Finset α) (f : α → ℝ) (hf : ∀ x ∈ s, f x ≠ 0) :
     logb b (∏ i in s, f i) = ∑ i in s, logb b (f i) := by
   classical
-    induction' s using Finset.induction_on with a s ha ih
-    · simp
-    simp only [Finset.mem_insert, forall_eq_or_imp] at hf 
-    simp [ha, ih hf.2, logb_mul hf.1 (Finset.prod_ne_zero_iff.2 hf.2)]
+  induction' s using Finset.induction_on with a s ha ih
+  · simp
+  simp only [Finset.mem_insert, forall_eq_or_imp] at hf 
+  simp [ha, ih hf.2, logb_mul hf.1 (Finset.prod_ne_zero_iff.2 hf.2)]
 #align real.logb_prod Real.logb_prod
 
 end Real
Diff
@@ -405,7 +405,7 @@ theorem logb_prod {α : Type _} (s : Finset α) (f : α → ℝ) (hf : ∀ x ∈
   classical
     induction' s using Finset.induction_on with a s ha ih
     · simp
-    simp only [Finset.mem_insert, forall_eq_or_imp] at hf
+    simp only [Finset.mem_insert, forall_eq_or_imp] at hf 
     simp [ha, ih hf.2, logb_mul hf.1 (Finset.prod_ne_zero_iff.2 hf.2)]
 #align real.logb_prod Real.logb_prod
 
Diff
@@ -32,7 +32,7 @@ logarithm, continuity
 
 open Set Filter Function
 
-open Topology
+open scoped Topology
 
 noncomputable section
 
@@ -398,7 +398,7 @@ theorem logb_eq_zero : logb b x = 0 ↔ b = 0 ∨ b = 1 ∨ b = -1 ∨ x = 0 ∨
 #align real.logb_eq_zero Real.logb_eq_zero
 
 -- TODO add other limits and continuous API lemmas analogous to those in log.lean
-open BigOperators
+open scoped BigOperators
 
 theorem logb_prod {α : Type _} (s : Finset α) (f : α → ℝ) (hf : ∀ x ∈ s, f x ≠ 0) :
     logb b (∏ i in s, f i) = ∑ i in s, logb b (f i) := by
Diff
@@ -49,83 +49,35 @@ noncomputable def logb (b x : ℝ) : ℝ :=
 #align real.logb Real.logb
 -/
 
-/- warning: real.log_div_log -> Real.log_div_log is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, Eq.{1} Real (HDiv.hDiv.{0, 0, 0} Real Real Real (instHDiv.{0} Real (DivInvMonoid.toHasDiv.{0} Real (DivisionRing.toDivInvMonoid.{0} Real Real.divisionRing))) (Real.log x) (Real.log b)) (Real.logb b x)
-but is expected to have type
-  forall {b : Real} {x : Real}, Eq.{1} Real (HDiv.hDiv.{0, 0, 0} Real Real Real (instHDiv.{0} Real (LinearOrderedField.toDiv.{0} Real Real.instLinearOrderedFieldReal)) (Real.log x) (Real.log b)) (Real.logb b x)
-Case conversion may be inaccurate. Consider using '#align real.log_div_log Real.log_div_logₓ'. -/
 theorem log_div_log : log x / log b = logb b x :=
   rfl
 #align real.log_div_log Real.log_div_log
 
-/- warning: real.logb_zero -> Real.logb_zero is a dubious translation:
-lean 3 declaration is
-  forall {b : Real}, Eq.{1} Real (Real.logb b (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))
-but is expected to have type
-  forall {b : Real}, Eq.{1} Real (Real.logb b (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))
-Case conversion may be inaccurate. Consider using '#align real.logb_zero Real.logb_zeroₓ'. -/
 @[simp]
 theorem logb_zero : logb b 0 = 0 := by simp [logb]
 #align real.logb_zero Real.logb_zero
 
-/- warning: real.logb_one -> Real.logb_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real}, Eq.{1} Real (Real.logb b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))
-but is expected to have type
-  forall {b : Real}, Eq.{1} Real (Real.logb b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))
-Case conversion may be inaccurate. Consider using '#align real.logb_one Real.logb_oneₓ'. -/
 @[simp]
 theorem logb_one : logb b 1 = 0 := by simp [logb]
 #align real.logb_one Real.logb_one
 
-/- warning: real.logb_abs -> Real.logb_abs is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} (x : Real), Eq.{1} Real (Real.logb b (Abs.abs.{0} Real (Neg.toHasAbs.{0} Real Real.hasNeg Real.hasSup) x)) (Real.logb b x)
-but is expected to have type
-  forall {b : Real} (x : Real), Eq.{1} Real (Real.logb b (Abs.abs.{0} Real (Neg.toHasAbs.{0} Real Real.instNegReal Real.instSupReal) x)) (Real.logb b x)
-Case conversion may be inaccurate. Consider using '#align real.logb_abs Real.logb_absₓ'. -/
 @[simp]
 theorem logb_abs (x : ℝ) : logb b (|x|) = logb b x := by rw [logb, logb, log_abs]
 #align real.logb_abs Real.logb_abs
 
-/- warning: real.logb_neg_eq_logb -> Real.logb_neg_eq_logb is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} (x : Real), Eq.{1} Real (Real.logb b (Neg.neg.{0} Real Real.hasNeg x)) (Real.logb b x)
-but is expected to have type
-  forall {b : Real} (x : Real), Eq.{1} Real (Real.logb b (Neg.neg.{0} Real Real.instNegReal x)) (Real.logb b x)
-Case conversion may be inaccurate. Consider using '#align real.logb_neg_eq_logb Real.logb_neg_eq_logbₓ'. -/
 @[simp]
 theorem logb_neg_eq_logb (x : ℝ) : logb b (-x) = logb b x := by
   rw [← logb_abs x, ← logb_abs (-x), abs_neg]
 #align real.logb_neg_eq_logb Real.logb_neg_eq_logb
 
-/- warning: real.logb_mul -> Real.logb_mul is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (Ne.{1} Real x (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Ne.{1} Real y (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Eq.{1} Real (Real.logb b (HMul.hMul.{0, 0, 0} Real Real Real (instHMul.{0} Real Real.hasMul) x y)) (HAdd.hAdd.{0, 0, 0} Real Real Real (instHAdd.{0} Real Real.hasAdd) (Real.logb b x) (Real.logb b y)))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (Ne.{1} Real x (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Ne.{1} Real y (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Eq.{1} Real (Real.logb b (HMul.hMul.{0, 0, 0} Real Real Real (instHMul.{0} Real Real.instMulReal) x y)) (HAdd.hAdd.{0, 0, 0} Real Real Real (instHAdd.{0} Real Real.instAddReal) (Real.logb b x) (Real.logb b y)))
-Case conversion may be inaccurate. Consider using '#align real.logb_mul Real.logb_mulₓ'. -/
 theorem logb_mul (hx : x ≠ 0) (hy : y ≠ 0) : logb b (x * y) = logb b x + logb b y := by
   simp_rw [logb, log_mul hx hy, add_div]
 #align real.logb_mul Real.logb_mul
 
-/- warning: real.logb_div -> Real.logb_div is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (Ne.{1} Real x (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Ne.{1} Real y (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Eq.{1} Real (Real.logb b (HDiv.hDiv.{0, 0, 0} Real Real Real (instHDiv.{0} Real (DivInvMonoid.toHasDiv.{0} Real (DivisionRing.toDivInvMonoid.{0} Real Real.divisionRing))) x y)) (HSub.hSub.{0, 0, 0} Real Real Real (instHSub.{0} Real Real.hasSub) (Real.logb b x) (Real.logb b y)))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (Ne.{1} Real x (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Ne.{1} Real y (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Eq.{1} Real (Real.logb b (HDiv.hDiv.{0, 0, 0} Real Real Real (instHDiv.{0} Real (LinearOrderedField.toDiv.{0} Real Real.instLinearOrderedFieldReal)) x y)) (HSub.hSub.{0, 0, 0} Real Real Real (instHSub.{0} Real Real.instSubReal) (Real.logb b x) (Real.logb b y)))
-Case conversion may be inaccurate. Consider using '#align real.logb_div Real.logb_divₓ'. -/
 theorem logb_div (hx : x ≠ 0) (hy : y ≠ 0) : logb b (x / y) = logb b x - logb b y := by
   simp_rw [logb, log_div hx hy, sub_div]
 #align real.logb_div Real.logb_div
 
-/- warning: real.logb_inv -> Real.logb_inv is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} (x : Real), Eq.{1} Real (Real.logb b (Inv.inv.{0} Real Real.hasInv x)) (Neg.neg.{0} Real Real.hasNeg (Real.logb b x))
-but is expected to have type
-  forall {b : Real} (x : Real), Eq.{1} Real (Real.logb b (Inv.inv.{0} Real Real.instInvReal x)) (Neg.neg.{0} Real Real.instNegReal (Real.logb b x))
-Case conversion may be inaccurate. Consider using '#align real.logb_inv Real.logb_invₓ'. -/
 @[simp]
 theorem logb_inv (x : ℝ) : logb b x⁻¹ = -logb b x := by simp [logb, neg_div]
 #align real.logb_inv Real.logb_inv
@@ -144,12 +96,6 @@ private theorem log_b_ne_zero : log b ≠ 0 :=
   have b_ne_minus_one : b ≠ -1; linarith
   simp [b_ne_one, b_ne_zero, b_ne_minus_one]
 
-/- warning: real.logb_rpow -> Real.logb_rpow is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Eq.{1} Real (Real.logb b (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b x)) x)
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Eq.{1} Real (Real.logb b (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b x)) x)
-Case conversion may be inaccurate. Consider using '#align real.logb_rpow Real.logb_rpowₓ'. -/
 @[simp]
 theorem logb_rpow : logb b (b ^ x) = x :=
   by
@@ -157,12 +103,6 @@ theorem logb_rpow : logb b (b ^ x) = x :=
   exact log_b_ne_zero b_pos b_ne_one
 #align real.logb_rpow Real.logb_rpow
 
-/- warning: real.rpow_logb_eq_abs -> Real.rpow_logb_eq_abs is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Ne.{1} Real x (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Eq.{1} Real (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b (Real.logb b x)) (Abs.abs.{0} Real (Neg.toHasAbs.{0} Real Real.hasNeg Real.hasSup) x))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Ne.{1} Real x (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Eq.{1} Real (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b (Real.logb b x)) (Abs.abs.{0} Real (Neg.toHasAbs.{0} Real Real.instNegReal Real.instSupReal) x))
-Case conversion may be inaccurate. Consider using '#align real.rpow_logb_eq_abs Real.rpow_logb_eq_absₓ'. -/
 theorem rpow_logb_eq_abs (hx : x ≠ 0) : b ^ logb b x = |x| :=
   by
   apply log_inj_on_pos
@@ -173,63 +113,27 @@ theorem rpow_logb_eq_abs (hx : x ≠ 0) : b ^ logb b x = |x| :=
   field_simp [log_b_ne_zero b_pos b_ne_one]
 #align real.rpow_logb_eq_abs Real.rpow_logb_eq_abs
 
-/- warning: real.rpow_logb -> Real.rpow_logb is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Eq.{1} Real (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b (Real.logb b x)) x)
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Eq.{1} Real (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b (Real.logb b x)) x)
-Case conversion may be inaccurate. Consider using '#align real.rpow_logb Real.rpow_logbₓ'. -/
 @[simp]
 theorem rpow_logb (hx : 0 < x) : b ^ logb b x = x := by rw [rpow_logb_eq_abs b_pos b_ne_one hx.ne'];
   exact abs_of_pos hx
 #align real.rpow_logb Real.rpow_logb
 
-/- warning: real.rpow_logb_of_neg -> Real.rpow_logb_of_neg is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt x (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Eq.{1} Real (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b (Real.logb b x)) (Neg.neg.{0} Real Real.hasNeg x))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal x (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Eq.{1} Real (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b (Real.logb b x)) (Neg.neg.{0} Real Real.instNegReal x))
-Case conversion may be inaccurate. Consider using '#align real.rpow_logb_of_neg Real.rpow_logb_of_negₓ'. -/
 theorem rpow_logb_of_neg (hx : x < 0) : b ^ logb b x = -x := by
   rw [rpow_logb_eq_abs b_pos b_ne_one (ne_of_lt hx)]; exact abs_of_neg hx
 #align real.rpow_logb_of_neg Real.rpow_logb_of_neg
 
-/- warning: real.surj_on_logb -> Real.surjOn_logb is a dubious translation:
-lean 3 declaration is
-  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Set.SurjOn.{0, 0} Real Real (Real.logb b) (Set.Ioi.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (Set.univ.{0} Real))
-but is expected to have type
-  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Set.SurjOn.{0, 0} Real Real (Real.logb b) (Set.Ioi.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (Set.univ.{0} Real))
-Case conversion may be inaccurate. Consider using '#align real.surj_on_logb Real.surjOn_logbₓ'. -/
 theorem surjOn_logb : SurjOn (logb b) (Ioi 0) univ := fun x _ =>
   ⟨rpow b x, rpow_pos_of_pos b_pos x, logb_rpow b_pos b_ne_one⟩
 #align real.surj_on_logb Real.surjOn_logb
 
-/- warning: real.logb_surjective -> Real.logb_surjective is a dubious translation:
-lean 3 declaration is
-  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Function.Surjective.{1, 1} Real Real (Real.logb b))
-but is expected to have type
-  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Function.Surjective.{1, 1} Real Real (Real.logb b))
-Case conversion may be inaccurate. Consider using '#align real.logb_surjective Real.logb_surjectiveₓ'. -/
 theorem logb_surjective : Surjective (logb b) := fun x => ⟨b ^ x, logb_rpow b_pos b_ne_one⟩
 #align real.logb_surjective Real.logb_surjective
 
-/- warning: real.range_logb -> Real.range_logb is a dubious translation:
-lean 3 declaration is
-  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Eq.{1} (Set.{0} Real) (Set.range.{0, 1} Real Real (Real.logb b)) (Set.univ.{0} Real))
-but is expected to have type
-  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Eq.{1} (Set.{0} Real) (Set.range.{0, 1} Real Real (Real.logb b)) (Set.univ.{0} Real))
-Case conversion may be inaccurate. Consider using '#align real.range_logb Real.range_logbₓ'. -/
 @[simp]
 theorem range_logb : range (logb b) = univ :=
   (logb_surjective b_pos b_ne_one).range_eq
 #align real.range_logb Real.range_logb
 
-/- warning: real.surj_on_logb' -> Real.surjOn_logb' is a dubious translation:
-lean 3 declaration is
-  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Set.SurjOn.{0, 0} Real Real (Real.logb b) (Set.Iio.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (Set.univ.{0} Real))
-but is expected to have type
-  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Set.SurjOn.{0, 0} Real Real (Real.logb b) (Set.Iio.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (Set.univ.{0} Real))
-Case conversion may be inaccurate. Consider using '#align real.surj_on_logb' Real.surjOn_logb'ₓ'. -/
 theorem surjOn_logb' : SurjOn (logb b) (Iio 0) univ :=
   by
   intro x x_in_univ
@@ -251,154 +155,64 @@ private theorem b_pos : 0 < b := by linarith
 
 private theorem b_ne_one : b ≠ 1 := by linarith
 
-/- warning: real.logb_le_logb -> Real.logb_le_logb is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LE.le.{0} Real Real.hasLe (Real.logb b x) (Real.logb b y)) (LE.le.{0} Real Real.hasLe x y))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LE.le.{0} Real Real.instLEReal (Real.logb b x) (Real.logb b y)) (LE.le.{0} Real Real.instLEReal x y))
-Case conversion may be inaccurate. Consider using '#align real.logb_le_logb Real.logb_le_logbₓ'. -/
 @[simp]
 theorem logb_le_logb (h : 0 < x) (h₁ : 0 < y) : logb b x ≤ logb b y ↔ x ≤ y := by
   rw [logb, logb, div_le_div_right (log_pos hb), log_le_log h h₁]
 #align real.logb_le_logb Real.logb_le_logb
 
-/- warning: real.logb_lt_logb -> Real.logb_lt_logb is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt x y) -> (LT.lt.{0} Real Real.hasLt (Real.logb b x) (Real.logb b y))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal x y) -> (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (Real.logb b y))
-Case conversion may be inaccurate. Consider using '#align real.logb_lt_logb Real.logb_lt_logbₓ'. -/
 theorem logb_lt_logb (hx : 0 < x) (hxy : x < y) : logb b x < logb b y := by
   rw [logb, logb, div_lt_div_right (log_pos hb)]; exact log_lt_log hx hxy
 #align real.logb_lt_logb Real.logb_lt_logb
 
-/- warning: real.logb_lt_logb_iff -> Real.logb_lt_logb_iff is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LT.lt.{0} Real Real.hasLt (Real.logb b x) (Real.logb b y)) (LT.lt.{0} Real Real.hasLt x y))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (Real.logb b y)) (LT.lt.{0} Real Real.instLTReal x y))
-Case conversion may be inaccurate. Consider using '#align real.logb_lt_logb_iff Real.logb_lt_logb_iffₓ'. -/
 @[simp]
 theorem logb_lt_logb_iff (hx : 0 < x) (hy : 0 < y) : logb b x < logb b y ↔ x < y := by
   rw [logb, logb, div_lt_div_right (log_pos hb)]; exact log_lt_log_iff hx hy
 #align real.logb_lt_logb_iff Real.logb_lt_logb_iff
 
-/- warning: real.logb_le_iff_le_rpow -> Real.logb_le_iff_le_rpow is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LE.le.{0} Real Real.hasLe (Real.logb b x) y) (LE.le.{0} Real Real.hasLe x (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b y)))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LE.le.{0} Real Real.instLEReal (Real.logb b x) y) (LE.le.{0} Real Real.instLEReal x (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b y)))
-Case conversion may be inaccurate. Consider using '#align real.logb_le_iff_le_rpow Real.logb_le_iff_le_rpowₓ'. -/
 theorem logb_le_iff_le_rpow (hx : 0 < x) : logb b x ≤ y ↔ x ≤ b ^ y := by
   rw [← rpow_le_rpow_left_iff hb, rpow_logb (b_pos hb) (b_ne_one hb) hx]
 #align real.logb_le_iff_le_rpow Real.logb_le_iff_le_rpow
 
-/- warning: real.logb_lt_iff_lt_rpow -> Real.logb_lt_iff_lt_rpow is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LT.lt.{0} Real Real.hasLt (Real.logb b x) y) (LT.lt.{0} Real Real.hasLt x (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b y)))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LT.lt.{0} Real Real.instLTReal (Real.logb b x) y) (LT.lt.{0} Real Real.instLTReal x (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b y)))
-Case conversion may be inaccurate. Consider using '#align real.logb_lt_iff_lt_rpow Real.logb_lt_iff_lt_rpowₓ'. -/
 theorem logb_lt_iff_lt_rpow (hx : 0 < x) : logb b x < y ↔ x < b ^ y := by
   rw [← rpow_lt_rpow_left_iff hb, rpow_logb (b_pos hb) (b_ne_one hb) hx]
 #align real.logb_lt_iff_lt_rpow Real.logb_lt_iff_lt_rpow
 
-/- warning: real.le_logb_iff_rpow_le -> Real.le_logb_iff_rpow_le is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LE.le.{0} Real Real.hasLe x (Real.logb b y)) (LE.le.{0} Real Real.hasLe (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b x) y))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LE.le.{0} Real Real.instLEReal x (Real.logb b y)) (LE.le.{0} Real Real.instLEReal (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b x) y))
-Case conversion may be inaccurate. Consider using '#align real.le_logb_iff_rpow_le Real.le_logb_iff_rpow_leₓ'. -/
 theorem le_logb_iff_rpow_le (hy : 0 < y) : x ≤ logb b y ↔ b ^ x ≤ y := by
   rw [← rpow_le_rpow_left_iff hb, rpow_logb (b_pos hb) (b_ne_one hb) hy]
 #align real.le_logb_iff_rpow_le Real.le_logb_iff_rpow_le
 
-/- warning: real.lt_logb_iff_rpow_lt -> Real.lt_logb_iff_rpow_lt is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LT.lt.{0} Real Real.hasLt x (Real.logb b y)) (LT.lt.{0} Real Real.hasLt (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b x) y))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LT.lt.{0} Real Real.instLTReal x (Real.logb b y)) (LT.lt.{0} Real Real.instLTReal (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b x) y))
-Case conversion may be inaccurate. Consider using '#align real.lt_logb_iff_rpow_lt Real.lt_logb_iff_rpow_ltₓ'. -/
 theorem lt_logb_iff_rpow_lt (hy : 0 < y) : x < logb b y ↔ b ^ x < y := by
   rw [← rpow_lt_rpow_left_iff hb, rpow_logb (b_pos hb) (b_ne_one hb) hy]
 #align real.lt_logb_iff_rpow_lt Real.lt_logb_iff_rpow_lt
 
-/- warning: real.logb_pos_iff -> Real.logb_pos_iff is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x)) (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) x))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x)) (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x))
-Case conversion may be inaccurate. Consider using '#align real.logb_pos_iff Real.logb_pos_iffₓ'. -/
 theorem logb_pos_iff (hx : 0 < x) : 0 < logb b x ↔ 1 < x := by rw [← @logb_one b];
   rw [logb_lt_logb_iff hb zero_lt_one hx]
 #align real.logb_pos_iff Real.logb_pos_iff
 
-/- warning: real.logb_pos -> Real.logb_pos is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) x) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x))
-Case conversion may be inaccurate. Consider using '#align real.logb_pos Real.logb_posₓ'. -/
 theorem logb_pos (hx : 1 < x) : 0 < logb b x := by rw [logb_pos_iff hb (lt_trans zero_lt_one hx)];
   exact hx
 #align real.logb_pos Real.logb_pos
 
-/- warning: real.logb_neg_iff -> Real.logb_neg_iff is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LT.lt.{0} Real Real.hasLt (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (LT.lt.{0} Real Real.hasLt x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (LT.lt.{0} Real Real.instLTReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))))
-Case conversion may be inaccurate. Consider using '#align real.logb_neg_iff Real.logb_neg_iffₓ'. -/
 theorem logb_neg_iff (h : 0 < x) : logb b x < 0 ↔ x < 1 := by rw [← logb_one];
   exact logb_lt_logb_iff hb h zero_lt_one
 #align real.logb_neg_iff Real.logb_neg_iff
 
-/- warning: real.logb_neg -> Real.logb_neg is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)))
-Case conversion may be inaccurate. Consider using '#align real.logb_neg Real.logb_negₓ'. -/
 theorem logb_neg (h0 : 0 < x) (h1 : x < 1) : logb b x < 0 :=
   (logb_neg_iff hb h0).2 h1
 #align real.logb_neg Real.logb_neg
 
-/- warning: real.logb_nonneg_iff -> Real.logb_nonneg_iff is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x)) (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) x))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x)) (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x))
-Case conversion may be inaccurate. Consider using '#align real.logb_nonneg_iff Real.logb_nonneg_iffₓ'. -/
 theorem logb_nonneg_iff (hx : 0 < x) : 0 ≤ logb b x ↔ 1 ≤ x := by
   rw [← not_lt, logb_neg_iff hb hx, not_lt]
 #align real.logb_nonneg_iff Real.logb_nonneg_iff
 
-/- warning: real.logb_nonneg -> Real.logb_nonneg is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) x) -> (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x))
-Case conversion may be inaccurate. Consider using '#align real.logb_nonneg Real.logb_nonnegₓ'. -/
 theorem logb_nonneg (hx : 1 ≤ x) : 0 ≤ logb b x :=
   (logb_nonneg_iff hb (zero_lt_one.trans_le hx)).2 hx
 #align real.logb_nonneg Real.logb_nonneg
 
-/- warning: real.logb_nonpos_iff -> Real.logb_nonpos_iff is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LE.le.{0} Real Real.hasLe (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (LE.le.{0} Real Real.hasLe x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LE.le.{0} Real Real.instLEReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (LE.le.{0} Real Real.instLEReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))))
-Case conversion may be inaccurate. Consider using '#align real.logb_nonpos_iff Real.logb_nonpos_iffₓ'. -/
 theorem logb_nonpos_iff (hx : 0 < x) : logb b x ≤ 0 ↔ x ≤ 1 := by
   rw [← not_lt, logb_pos_iff hb hx, not_lt]
 #align real.logb_nonpos_iff Real.logb_nonpos_iff
 
-/- warning: real.logb_nonpos_iff' -> Real.logb_nonpos_iff' is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LE.le.{0} Real Real.hasLe (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (LE.le.{0} Real Real.hasLe x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LE.le.{0} Real Real.instLEReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (LE.le.{0} Real Real.instLEReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))))
-Case conversion may be inaccurate. Consider using '#align real.logb_nonpos_iff' Real.logb_nonpos_iff'ₓ'. -/
 theorem logb_nonpos_iff' (hx : 0 ≤ x) : logb b x ≤ 0 ↔ x ≤ 1 :=
   by
   rcases hx.eq_or_lt with (rfl | hx)
@@ -406,32 +220,14 @@ theorem logb_nonpos_iff' (hx : 0 ≤ x) : logb b x ≤ 0 ↔ x ≤ 1 :=
   exact logb_nonpos_iff hb hx
 #align real.logb_nonpos_iff' Real.logb_nonpos_iff'
 
-/- warning: real.logb_nonpos -> Real.logb_nonpos is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LE.le.{0} Real Real.hasLe x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LE.le.{0} Real Real.hasLe (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LE.le.{0} Real Real.instLEReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LE.le.{0} Real Real.instLEReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)))
-Case conversion may be inaccurate. Consider using '#align real.logb_nonpos Real.logb_nonposₓ'. -/
 theorem logb_nonpos (hx : 0 ≤ x) (h'x : x ≤ 1) : logb b x ≤ 0 :=
   (logb_nonpos_iff' hb hx).2 h'x
 #align real.logb_nonpos Real.logb_nonpos
 
-/- warning: real.strict_mono_on_logb -> Real.strictMonoOn_logb is a dubious translation:
-lean 3 declaration is
-  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (StrictMonoOn.{0, 0} Real Real Real.preorder Real.preorder (Real.logb b) (Set.Ioi.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))))
-but is expected to have type
-  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (StrictMonoOn.{0, 0} Real Real Real.instPreorderReal Real.instPreorderReal (Real.logb b) (Set.Ioi.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))))
-Case conversion may be inaccurate. Consider using '#align real.strict_mono_on_logb Real.strictMonoOn_logbₓ'. -/
 theorem strictMonoOn_logb : StrictMonoOn (logb b) (Set.Ioi 0) := fun x hx y hy hxy =>
   logb_lt_logb hb hx hxy
 #align real.strict_mono_on_logb Real.strictMonoOn_logb
 
-/- warning: real.strict_anti_on_logb -> Real.strictAntiOn_logb is a dubious translation:
-lean 3 declaration is
-  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (StrictAntiOn.{0, 0} Real Real Real.preorder Real.preorder (Real.logb b) (Set.Iio.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))))
-but is expected to have type
-  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (StrictAntiOn.{0, 0} Real Real Real.instPreorderReal Real.instPreorderReal (Real.logb b) (Set.Iio.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))))
-Case conversion may be inaccurate. Consider using '#align real.strict_anti_on_logb Real.strictAntiOn_logbₓ'. -/
 theorem strictAntiOn_logb : StrictAntiOn (logb b) (Set.Iio 0) :=
   by
   rintro x (hx : x < 0) y (hy : y < 0) hxy
@@ -440,42 +236,18 @@ theorem strictAntiOn_logb : StrictAntiOn (logb b) (Set.Iio 0) :=
   rwa [abs_of_neg hy, abs_of_neg hx, neg_lt_neg_iff]
 #align real.strict_anti_on_logb Real.strictAntiOn_logb
 
-/- warning: real.logb_inj_on_pos -> Real.logb_injOn_pos is a dubious translation:
-lean 3 declaration is
-  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (Set.InjOn.{0, 0} Real Real (Real.logb b) (Set.Ioi.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))))
-but is expected to have type
-  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (Set.InjOn.{0, 0} Real Real (Real.logb b) (Set.Ioi.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))))
-Case conversion may be inaccurate. Consider using '#align real.logb_inj_on_pos Real.logb_injOn_posₓ'. -/
 theorem logb_injOn_pos : Set.InjOn (logb b) (Set.Ioi 0) :=
   (strictMonoOn_logb hb).InjOn
 #align real.logb_inj_on_pos Real.logb_injOn_pos
 
-/- warning: real.eq_one_of_pos_of_logb_eq_zero -> Real.eq_one_of_pos_of_logb_eq_zero is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Eq.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Eq.{1} Real x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Eq.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Eq.{1} Real x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)))
-Case conversion may be inaccurate. Consider using '#align real.eq_one_of_pos_of_logb_eq_zero Real.eq_one_of_pos_of_logb_eq_zeroₓ'. -/
 theorem eq_one_of_pos_of_logb_eq_zero (h₁ : 0 < x) (h₂ : logb b x = 0) : x = 1 :=
   logb_injOn_pos hb (Set.mem_Ioi.2 h₁) (Set.mem_Ioi.2 zero_lt_one) (h₂.trans Real.logb_one.symm)
 #align real.eq_one_of_pos_of_logb_eq_zero Real.eq_one_of_pos_of_logb_eq_zero
 
-/- warning: real.logb_ne_zero_of_pos_of_ne_one -> Real.logb_ne_zero_of_pos_of_ne_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Ne.{1} Real x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Ne.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Ne.{1} Real x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Ne.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)))
-Case conversion may be inaccurate. Consider using '#align real.logb_ne_zero_of_pos_of_ne_one Real.logb_ne_zero_of_pos_of_ne_oneₓ'. -/
 theorem logb_ne_zero_of_pos_of_ne_one (hx_pos : 0 < x) (hx : x ≠ 1) : logb b x ≠ 0 :=
   mt (eq_one_of_pos_of_logb_eq_zero hb hx_pos) hx
 #align real.logb_ne_zero_of_pos_of_ne_one Real.logb_ne_zero_of_pos_of_ne_one
 
-/- warning: real.tendsto_logb_at_top -> Real.tendsto_logb_atTop is a dubious translation:
-lean 3 declaration is
-  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (Filter.Tendsto.{0, 0} Real Real (Real.logb b) (Filter.atTop.{0} Real Real.preorder) (Filter.atTop.{0} Real Real.preorder))
-but is expected to have type
-  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (Filter.Tendsto.{0, 0} Real Real (Real.logb b) (Filter.atTop.{0} Real Real.instPreorderReal) (Filter.atTop.{0} Real Real.instPreorderReal))
-Case conversion may be inaccurate. Consider using '#align real.tendsto_logb_at_top Real.tendsto_logb_atTopₓ'. -/
 theorem tendsto_logb_atTop : Tendsto (logb b) atTop atTop :=
   Tendsto.atTop_div_const (log_pos hb) tendsto_log_atTop
 #align real.tendsto_logb_at_top Real.tendsto_logb_atTop
@@ -494,164 +266,68 @@ private theorem b_ne_one : b ≠ 1 := by linarith
 
 include b_pos
 
-/- warning: real.logb_le_logb_of_base_lt_one -> Real.logb_le_logb_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LE.le.{0} Real Real.hasLe (Real.logb b x) (Real.logb b y)) (LE.le.{0} Real Real.hasLe y x))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LE.le.{0} Real Real.instLEReal (Real.logb b x) (Real.logb b y)) (LE.le.{0} Real Real.instLEReal y x))
-Case conversion may be inaccurate. Consider using '#align real.logb_le_logb_of_base_lt_one Real.logb_le_logb_of_base_lt_oneₓ'. -/
 @[simp]
 theorem logb_le_logb_of_base_lt_one (h : 0 < x) (h₁ : 0 < y) : logb b x ≤ logb b y ↔ y ≤ x := by
   rw [logb, logb, div_le_div_right_of_neg (log_neg b_pos b_lt_one), log_le_log h₁ h]
 #align real.logb_le_logb_of_base_lt_one Real.logb_le_logb_of_base_lt_one
 
-/- warning: real.logb_lt_logb_of_base_lt_one -> Real.logb_lt_logb_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt x y) -> (LT.lt.{0} Real Real.hasLt (Real.logb b y) (Real.logb b x))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal x y) -> (LT.lt.{0} Real Real.instLTReal (Real.logb b y) (Real.logb b x))
-Case conversion may be inaccurate. Consider using '#align real.logb_lt_logb_of_base_lt_one Real.logb_lt_logb_of_base_lt_oneₓ'. -/
 theorem logb_lt_logb_of_base_lt_one (hx : 0 < x) (hxy : x < y) : logb b y < logb b x := by
   rw [logb, logb, div_lt_div_right_of_neg (log_neg b_pos b_lt_one)]; exact log_lt_log hx hxy
 #align real.logb_lt_logb_of_base_lt_one Real.logb_lt_logb_of_base_lt_one
 
-/- warning: real.logb_lt_logb_iff_of_base_lt_one -> Real.logb_lt_logb_iff_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LT.lt.{0} Real Real.hasLt (Real.logb b x) (Real.logb b y)) (LT.lt.{0} Real Real.hasLt y x))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (Real.logb b y)) (LT.lt.{0} Real Real.instLTReal y x))
-Case conversion may be inaccurate. Consider using '#align real.logb_lt_logb_iff_of_base_lt_one Real.logb_lt_logb_iff_of_base_lt_oneₓ'. -/
 @[simp]
 theorem logb_lt_logb_iff_of_base_lt_one (hx : 0 < x) (hy : 0 < y) : logb b x < logb b y ↔ y < x :=
   by rw [logb, logb, div_lt_div_right_of_neg (log_neg b_pos b_lt_one)]; exact log_lt_log_iff hy hx
 #align real.logb_lt_logb_iff_of_base_lt_one Real.logb_lt_logb_iff_of_base_lt_one
 
-/- warning: real.logb_le_iff_le_rpow_of_base_lt_one -> Real.logb_le_iff_le_rpow_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LE.le.{0} Real Real.hasLe (Real.logb b x) y) (LE.le.{0} Real Real.hasLe (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b y) x))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LE.le.{0} Real Real.instLEReal (Real.logb b x) y) (LE.le.{0} Real Real.instLEReal (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b y) x))
-Case conversion may be inaccurate. Consider using '#align real.logb_le_iff_le_rpow_of_base_lt_one Real.logb_le_iff_le_rpow_of_base_lt_oneₓ'. -/
 theorem logb_le_iff_le_rpow_of_base_lt_one (hx : 0 < x) : logb b x ≤ y ↔ b ^ y ≤ x := by
   rw [← rpow_le_rpow_left_iff_of_base_lt_one b_pos b_lt_one, rpow_logb b_pos (b_ne_one b_lt_one) hx]
 #align real.logb_le_iff_le_rpow_of_base_lt_one Real.logb_le_iff_le_rpow_of_base_lt_one
 
-/- warning: real.logb_lt_iff_lt_rpow_of_base_lt_one -> Real.logb_lt_iff_lt_rpow_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LT.lt.{0} Real Real.hasLt (Real.logb b x) y) (LT.lt.{0} Real Real.hasLt (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b y) x))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LT.lt.{0} Real Real.instLTReal (Real.logb b x) y) (LT.lt.{0} Real Real.instLTReal (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b y) x))
-Case conversion may be inaccurate. Consider using '#align real.logb_lt_iff_lt_rpow_of_base_lt_one Real.logb_lt_iff_lt_rpow_of_base_lt_oneₓ'. -/
 theorem logb_lt_iff_lt_rpow_of_base_lt_one (hx : 0 < x) : logb b x < y ↔ b ^ y < x := by
   rw [← rpow_lt_rpow_left_iff_of_base_lt_one b_pos b_lt_one, rpow_logb b_pos (b_ne_one b_lt_one) hx]
 #align real.logb_lt_iff_lt_rpow_of_base_lt_one Real.logb_lt_iff_lt_rpow_of_base_lt_one
 
-/- warning: real.le_logb_iff_rpow_le_of_base_lt_one -> Real.le_logb_iff_rpow_le_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LE.le.{0} Real Real.hasLe x (Real.logb b y)) (LE.le.{0} Real Real.hasLe y (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b x)))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LE.le.{0} Real Real.instLEReal x (Real.logb b y)) (LE.le.{0} Real Real.instLEReal y (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b x)))
-Case conversion may be inaccurate. Consider using '#align real.le_logb_iff_rpow_le_of_base_lt_one Real.le_logb_iff_rpow_le_of_base_lt_oneₓ'. -/
 theorem le_logb_iff_rpow_le_of_base_lt_one (hy : 0 < y) : x ≤ logb b y ↔ y ≤ b ^ x := by
   rw [← rpow_le_rpow_left_iff_of_base_lt_one b_pos b_lt_one, rpow_logb b_pos (b_ne_one b_lt_one) hy]
 #align real.le_logb_iff_rpow_le_of_base_lt_one Real.le_logb_iff_rpow_le_of_base_lt_one
 
-/- warning: real.lt_logb_iff_rpow_lt_of_base_lt_one -> Real.lt_logb_iff_rpow_lt_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LT.lt.{0} Real Real.hasLt x (Real.logb b y)) (LT.lt.{0} Real Real.hasLt y (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b x)))
-but is expected to have type
-  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LT.lt.{0} Real Real.instLTReal x (Real.logb b y)) (LT.lt.{0} Real Real.instLTReal y (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b x)))
-Case conversion may be inaccurate. Consider using '#align real.lt_logb_iff_rpow_lt_of_base_lt_one Real.lt_logb_iff_rpow_lt_of_base_lt_oneₓ'. -/
 theorem lt_logb_iff_rpow_lt_of_base_lt_one (hy : 0 < y) : x < logb b y ↔ y < b ^ x := by
   rw [← rpow_lt_rpow_left_iff_of_base_lt_one b_pos b_lt_one, rpow_logb b_pos (b_ne_one b_lt_one) hy]
 #align real.lt_logb_iff_rpow_lt_of_base_lt_one Real.lt_logb_iff_rpow_lt_of_base_lt_one
 
-/- warning: real.logb_pos_iff_of_base_lt_one -> Real.logb_pos_iff_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x)) (LT.lt.{0} Real Real.hasLt x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x)) (LT.lt.{0} Real Real.instLTReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))))
-Case conversion may be inaccurate. Consider using '#align real.logb_pos_iff_of_base_lt_one Real.logb_pos_iff_of_base_lt_oneₓ'. -/
 theorem logb_pos_iff_of_base_lt_one (hx : 0 < x) : 0 < logb b x ↔ x < 1 := by
   rw [← @logb_one b, logb_lt_logb_iff_of_base_lt_one b_pos b_lt_one zero_lt_one hx]
 #align real.logb_pos_iff_of_base_lt_one Real.logb_pos_iff_of_base_lt_one
 
-/- warning: real.logb_pos_of_base_lt_one -> Real.logb_pos_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x))
-Case conversion may be inaccurate. Consider using '#align real.logb_pos_of_base_lt_one Real.logb_pos_of_base_lt_oneₓ'. -/
 theorem logb_pos_of_base_lt_one (hx : 0 < x) (hx' : x < 1) : 0 < logb b x := by
   rw [logb_pos_iff_of_base_lt_one b_pos b_lt_one hx]; exact hx'
 #align real.logb_pos_of_base_lt_one Real.logb_pos_of_base_lt_one
 
-/- warning: real.logb_neg_iff_of_base_lt_one -> Real.logb_neg_iff_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LT.lt.{0} Real Real.hasLt (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) x))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x))
-Case conversion may be inaccurate. Consider using '#align real.logb_neg_iff_of_base_lt_one Real.logb_neg_iff_of_base_lt_oneₓ'. -/
 theorem logb_neg_iff_of_base_lt_one (h : 0 < x) : logb b x < 0 ↔ 1 < x := by
   rw [← @logb_one b, logb_lt_logb_iff_of_base_lt_one b_pos b_lt_one h zero_lt_one]
 #align real.logb_neg_iff_of_base_lt_one Real.logb_neg_iff_of_base_lt_one
 
-/- warning: real.logb_neg_of_base_lt_one -> Real.logb_neg_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) x) -> (LT.lt.{0} Real Real.hasLt (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x) -> (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)))
-Case conversion may be inaccurate. Consider using '#align real.logb_neg_of_base_lt_one Real.logb_neg_of_base_lt_oneₓ'. -/
 theorem logb_neg_of_base_lt_one (h1 : 1 < x) : logb b x < 0 :=
   (logb_neg_iff_of_base_lt_one b_pos b_lt_one (lt_trans zero_lt_one h1)).2 h1
 #align real.logb_neg_of_base_lt_one Real.logb_neg_of_base_lt_one
 
-/- warning: real.logb_nonneg_iff_of_base_lt_one -> Real.logb_nonneg_iff_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x)) (LE.le.{0} Real Real.hasLe x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x)) (LE.le.{0} Real Real.instLEReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))))
-Case conversion may be inaccurate. Consider using '#align real.logb_nonneg_iff_of_base_lt_one Real.logb_nonneg_iff_of_base_lt_oneₓ'. -/
 theorem logb_nonneg_iff_of_base_lt_one (hx : 0 < x) : 0 ≤ logb b x ↔ x ≤ 1 := by
   rw [← not_lt, logb_neg_iff_of_base_lt_one b_pos b_lt_one hx, not_lt]
 #align real.logb_nonneg_iff_of_base_lt_one Real.logb_nonneg_iff_of_base_lt_one
 
-/- warning: real.logb_nonneg_of_base_lt_one -> Real.logb_nonneg_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LE.le.{0} Real Real.hasLe x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LE.le.{0} Real Real.instLEReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x))
-Case conversion may be inaccurate. Consider using '#align real.logb_nonneg_of_base_lt_one Real.logb_nonneg_of_base_lt_oneₓ'. -/
 theorem logb_nonneg_of_base_lt_one (hx : 0 < x) (hx' : x ≤ 1) : 0 ≤ logb b x := by
   rw [logb_nonneg_iff_of_base_lt_one b_pos b_lt_one hx]; exact hx'
 #align real.logb_nonneg_of_base_lt_one Real.logb_nonneg_of_base_lt_one
 
-/- warning: real.logb_nonpos_iff_of_base_lt_one -> Real.logb_nonpos_iff_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LE.le.{0} Real Real.hasLe (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) x))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LE.le.{0} Real Real.instLEReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x))
-Case conversion may be inaccurate. Consider using '#align real.logb_nonpos_iff_of_base_lt_one Real.logb_nonpos_iff_of_base_lt_oneₓ'. -/
 theorem logb_nonpos_iff_of_base_lt_one (hx : 0 < x) : logb b x ≤ 0 ↔ 1 ≤ x := by
   rw [← not_lt, logb_pos_iff_of_base_lt_one b_pos b_lt_one hx, not_lt]
 #align real.logb_nonpos_iff_of_base_lt_one Real.logb_nonpos_iff_of_base_lt_one
 
-/- warning: real.strict_anti_on_logb_of_base_lt_one -> Real.strictAntiOn_logb_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (StrictAntiOn.{0, 0} Real Real Real.preorder Real.preorder (Real.logb b) (Set.Ioi.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))))
-but is expected to have type
-  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (StrictAntiOn.{0, 0} Real Real Real.instPreorderReal Real.instPreorderReal (Real.logb b) (Set.Ioi.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))))
-Case conversion may be inaccurate. Consider using '#align real.strict_anti_on_logb_of_base_lt_one Real.strictAntiOn_logb_of_base_lt_oneₓ'. -/
 theorem strictAntiOn_logb_of_base_lt_one : StrictAntiOn (logb b) (Set.Ioi 0) := fun x hx y hy hxy =>
   logb_lt_logb_of_base_lt_one b_pos b_lt_one hx hxy
 #align real.strict_anti_on_logb_of_base_lt_one Real.strictAntiOn_logb_of_base_lt_one
 
-/- warning: real.strict_mono_on_logb_of_base_lt_one -> Real.strictMonoOn_logb_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (StrictMonoOn.{0, 0} Real Real Real.preorder Real.preorder (Real.logb b) (Set.Iio.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))))
-but is expected to have type
-  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (StrictMonoOn.{0, 0} Real Real Real.instPreorderReal Real.instPreorderReal (Real.logb b) (Set.Iio.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))))
-Case conversion may be inaccurate. Consider using '#align real.strict_mono_on_logb_of_base_lt_one Real.strictMonoOn_logb_of_base_lt_oneₓ'. -/
 theorem strictMonoOn_logb_of_base_lt_one : StrictMonoOn (logb b) (Set.Iio 0) :=
   by
   rintro x (hx : x < 0) y (hy : y < 0) hxy
@@ -660,43 +336,19 @@ theorem strictMonoOn_logb_of_base_lt_one : StrictMonoOn (logb b) (Set.Iio 0) :=
   rwa [abs_of_neg hy, abs_of_neg hx, neg_lt_neg_iff]
 #align real.strict_mono_on_logb_of_base_lt_one Real.strictMonoOn_logb_of_base_lt_one
 
-/- warning: real.logb_inj_on_pos_of_base_lt_one -> Real.logb_injOn_pos_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Set.InjOn.{0, 0} Real Real (Real.logb b) (Set.Ioi.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))))
-but is expected to have type
-  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Set.InjOn.{0, 0} Real Real (Real.logb b) (Set.Ioi.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))))
-Case conversion may be inaccurate. Consider using '#align real.logb_inj_on_pos_of_base_lt_one Real.logb_injOn_pos_of_base_lt_oneₓ'. -/
 theorem logb_injOn_pos_of_base_lt_one : Set.InjOn (logb b) (Set.Ioi 0) :=
   (strictAntiOn_logb_of_base_lt_one b_pos b_lt_one).InjOn
 #align real.logb_inj_on_pos_of_base_lt_one Real.logb_injOn_pos_of_base_lt_one
 
-/- warning: real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_one -> Real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Eq.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Eq.{1} Real x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Eq.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Eq.{1} Real x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)))
-Case conversion may be inaccurate. Consider using '#align real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_one Real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_oneₓ'. -/
 theorem eq_one_of_pos_of_logb_eq_zero_of_base_lt_one (h₁ : 0 < x) (h₂ : logb b x = 0) : x = 1 :=
   logb_injOn_pos_of_base_lt_one b_pos b_lt_one (Set.mem_Ioi.2 h₁) (Set.mem_Ioi.2 zero_lt_one)
     (h₂.trans Real.logb_one.symm)
 #align real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_one Real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_one
 
-/- warning: real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_one -> Real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Ne.{1} Real x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Ne.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))))
-but is expected to have type
-  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Ne.{1} Real x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Ne.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)))
-Case conversion may be inaccurate. Consider using '#align real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_one Real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_oneₓ'. -/
 theorem logb_ne_zero_of_pos_of_ne_one_of_base_lt_one (hx_pos : 0 < x) (hx : x ≠ 1) : logb b x ≠ 0 :=
   mt (eq_one_of_pos_of_logb_eq_zero_of_base_lt_one b_pos b_lt_one hx_pos) hx
 #align real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_one Real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_one
 
-/- warning: real.tendsto_logb_at_top_of_base_lt_one -> Real.tendsto_logb_atTop_of_base_lt_one is a dubious translation:
-lean 3 declaration is
-  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Filter.Tendsto.{0, 0} Real Real (Real.logb b) (Filter.atTop.{0} Real Real.preorder) (Filter.atBot.{0} Real Real.preorder))
-but is expected to have type
-  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Filter.Tendsto.{0, 0} Real Real (Real.logb b) (Filter.atTop.{0} Real Real.instPreorderReal) (Filter.atBot.{0} Real Real.instPreorderReal))
-Case conversion may be inaccurate. Consider using '#align real.tendsto_logb_at_top_of_base_lt_one Real.tendsto_logb_atTop_of_base_lt_oneₓ'. -/
 theorem tendsto_logb_atTop_of_base_lt_one : Tendsto (logb b) atTop atBot :=
   by
   rw [tendsto_at_top_at_bot]
@@ -712,12 +364,6 @@ theorem tendsto_logb_atTop_of_base_lt_one : Tendsto (logb b) atTop atBot :=
 
 end BPosAndBLtOne
 
-/- warning: real.floor_logb_nat_cast -> Real.floor_logb_nat_cast is a dubious translation:
-lean 3 declaration is
-  forall {b : Nat} {r : Real}, (LT.lt.{0} Nat Nat.hasLt (OfNat.ofNat.{0} Nat 1 (OfNat.mk.{0} Nat 1 (One.one.{0} Nat Nat.hasOne))) b) -> (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) r) -> (Eq.{1} Int (Int.floor.{0} Real Real.linearOrderedRing Real.floorRing (Real.logb ((fun (a : Type) (b : Type) [self : HasLiftT.{1, 1} a b] => self.0) Nat Real (HasLiftT.mk.{1, 1} Nat Real (CoeTCₓ.coe.{1, 1} Nat Real (Nat.castCoe.{0} Real Real.hasNatCast))) b) r)) (Int.log.{0} Real (LinearOrderedField.toLinearOrderedSemifield.{0} Real Real.linearOrderedField) (FloorRing.toFloorSemiring.{0} Real Real.linearOrderedRing Real.floorRing) b r))
-but is expected to have type
-  forall {b : Nat} {r : Real}, (LT.lt.{0} Nat instLTNat (OfNat.ofNat.{0} Nat 1 (instOfNatNat 1)) b) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) r) -> (Eq.{1} Int (Int.floor.{0} Real Real.instLinearOrderedRingReal Real.instFloorRingRealInstLinearOrderedRingReal (Real.logb (Nat.cast.{0} Real Real.natCast b) r)) (Int.log.{0} Real (LinearOrderedField.toLinearOrderedSemifield.{0} Real Real.instLinearOrderedFieldReal) (FloorRing.toFloorSemiring.{0} Real (LinearOrderedCommRing.toLinearOrderedRing.{0} Real (LinearOrderedField.toLinearOrderedCommRing.{0} Real Real.instLinearOrderedFieldReal)) Real.instFloorRingRealInstLinearOrderedRingReal) b r))
-Case conversion may be inaccurate. Consider using '#align real.floor_logb_nat_cast Real.floor_logb_nat_castₓ'. -/
 theorem floor_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : ⌊logb b r⌋ = Int.log b r :=
   by
   obtain rfl | hr := hr.eq_or_lt
@@ -731,12 +377,6 @@ theorem floor_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : 
     exact Int.zpow_log_le_self hb hr
 #align real.floor_logb_nat_cast Real.floor_logb_nat_cast
 
-/- warning: real.ceil_logb_nat_cast -> Real.ceil_logb_nat_cast is a dubious translation:
-lean 3 declaration is
-  forall {b : Nat} {r : Real}, (LT.lt.{0} Nat Nat.hasLt (OfNat.ofNat.{0} Nat 1 (OfNat.mk.{0} Nat 1 (One.one.{0} Nat Nat.hasOne))) b) -> (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) r) -> (Eq.{1} Int (Int.ceil.{0} Real Real.linearOrderedRing Real.floorRing (Real.logb ((fun (a : Type) (b : Type) [self : HasLiftT.{1, 1} a b] => self.0) Nat Real (HasLiftT.mk.{1, 1} Nat Real (CoeTCₓ.coe.{1, 1} Nat Real (Nat.castCoe.{0} Real Real.hasNatCast))) b) r)) (Int.clog.{0} Real (LinearOrderedField.toLinearOrderedSemifield.{0} Real Real.linearOrderedField) (FloorRing.toFloorSemiring.{0} Real Real.linearOrderedRing Real.floorRing) b r))
-but is expected to have type
-  forall {b : Nat} {r : Real}, (LT.lt.{0} Nat instLTNat (OfNat.ofNat.{0} Nat 1 (instOfNatNat 1)) b) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) r) -> (Eq.{1} Int (Int.ceil.{0} Real Real.instLinearOrderedRingReal Real.instFloorRingRealInstLinearOrderedRingReal (Real.logb (Nat.cast.{0} Real Real.natCast b) r)) (Int.clog.{0} Real (LinearOrderedField.toLinearOrderedSemifield.{0} Real Real.instLinearOrderedFieldReal) (FloorRing.toFloorSemiring.{0} Real (LinearOrderedCommRing.toLinearOrderedRing.{0} Real (LinearOrderedField.toLinearOrderedCommRing.{0} Real Real.instLinearOrderedFieldReal)) Real.instFloorRingRealInstLinearOrderedRingReal) b r))
-Case conversion may be inaccurate. Consider using '#align real.ceil_logb_nat_cast Real.ceil_logb_nat_castₓ'. -/
 theorem ceil_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : ⌈logb b r⌉ = Int.clog b r :=
   by
   obtain rfl | hr := hr.eq_or_lt
@@ -750,12 +390,6 @@ theorem ceil_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : ⌈
     exact rpow_le_rpow_of_exponent_le hb1'.le (Int.le_ceil _)
 #align real.ceil_logb_nat_cast Real.ceil_logb_nat_cast
 
-/- warning: real.logb_eq_zero -> Real.logb_eq_zero is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {x : Real}, Iff (Eq.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (Or (Eq.{1} Real b (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (Or (Eq.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) (Or (Eq.{1} Real b (Neg.neg.{0} Real Real.hasNeg (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))))) (Or (Eq.{1} Real x (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (Or (Eq.{1} Real x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) (Eq.{1} Real x (Neg.neg.{0} Real Real.hasNeg (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))))))))))
-but is expected to have type
-  forall {b : Real} {x : Real}, Iff (Eq.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (Or (Eq.{1} Real b (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (Or (Eq.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) (Or (Eq.{1} Real b (Neg.neg.{0} Real Real.instNegReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)))) (Or (Eq.{1} Real x (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (Or (Eq.{1} Real x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) (Eq.{1} Real x (Neg.neg.{0} Real Real.instNegReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)))))))))
-Case conversion may be inaccurate. Consider using '#align real.logb_eq_zero Real.logb_eq_zeroₓ'. -/
 @[simp]
 theorem logb_eq_zero : logb b x = 0 ↔ b = 0 ∨ b = 1 ∨ b = -1 ∨ x = 0 ∨ x = 1 ∨ x = -1 :=
   by
@@ -766,12 +400,6 @@ theorem logb_eq_zero : logb b x = 0 ↔ b = 0 ∨ b = 1 ∨ b = -1 ∨ x = 0 ∨
 -- TODO add other limits and continuous API lemmas analogous to those in log.lean
 open BigOperators
 
-/- warning: real.logb_prod -> Real.logb_prod is a dubious translation:
-lean 3 declaration is
-  forall {b : Real} {α : Type.{u1}} (s : Finset.{u1} α) (f : α -> Real), (forall (x : α), (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) -> (Ne.{1} Real (f x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))))) -> (Eq.{1} Real (Real.logb b (Finset.prod.{0, u1} Real α Real.commMonoid s (fun (i : α) => f i))) (Finset.sum.{0, u1} Real α Real.addCommMonoid s (fun (i : α) => Real.logb b (f i))))
-but is expected to have type
-  forall {b : Real} {α : Type.{u1}} (s : Finset.{u1} α) (f : α -> Real), (forall (x : α), (Membership.mem.{u1, u1} α (Finset.{u1} α) (Finset.instMembershipFinset.{u1} α) x s) -> (Ne.{1} Real (f x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)))) -> (Eq.{1} Real (Real.logb b (Finset.prod.{0, u1} Real α Real.instCommMonoidReal s (fun (i : α) => f i))) (Finset.sum.{0, u1} Real α Real.instAddCommMonoidReal s (fun (i : α) => Real.logb b (f i))))
-Case conversion may be inaccurate. Consider using '#align real.logb_prod Real.logb_prodₓ'. -/
 theorem logb_prod {α : Type _} (s : Finset α) (f : α → ℝ) (hf : ∀ x ∈ s, f x ≠ 0) :
     logb b (∏ i in s, f i) = ∑ i in s, logb b (f i) := by
   classical
Diff
@@ -180,9 +180,7 @@ but is expected to have type
   forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Eq.{1} Real (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b (Real.logb b x)) x)
 Case conversion may be inaccurate. Consider using '#align real.rpow_logb Real.rpow_logbₓ'. -/
 @[simp]
-theorem rpow_logb (hx : 0 < x) : b ^ logb b x = x :=
-  by
-  rw [rpow_logb_eq_abs b_pos b_ne_one hx.ne']
+theorem rpow_logb (hx : 0 < x) : b ^ logb b x = x := by rw [rpow_logb_eq_abs b_pos b_ne_one hx.ne'];
   exact abs_of_pos hx
 #align real.rpow_logb Real.rpow_logb
 
@@ -192,10 +190,8 @@ lean 3 declaration is
 but is expected to have type
   forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal x (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Eq.{1} Real (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b (Real.logb b x)) (Neg.neg.{0} Real Real.instNegReal x))
 Case conversion may be inaccurate. Consider using '#align real.rpow_logb_of_neg Real.rpow_logb_of_negₓ'. -/
-theorem rpow_logb_of_neg (hx : x < 0) : b ^ logb b x = -x :=
-  by
-  rw [rpow_logb_eq_abs b_pos b_ne_one (ne_of_lt hx)]
-  exact abs_of_neg hx
+theorem rpow_logb_of_neg (hx : x < 0) : b ^ logb b x = -x := by
+  rw [rpow_logb_eq_abs b_pos b_ne_one (ne_of_lt hx)]; exact abs_of_neg hx
 #align real.rpow_logb_of_neg Real.rpow_logb_of_neg
 
 /- warning: real.surj_on_logb -> Real.surjOn_logb is a dubious translation:
@@ -239,8 +235,7 @@ theorem surjOn_logb' : SurjOn (logb b) (Iio 0) univ :=
   intro x x_in_univ
   use -b ^ x
   constructor
-  · simp only [Right.neg_neg_iff, Set.mem_Iio]
-    apply rpow_pos_of_pos b_pos
+  · simp only [Right.neg_neg_iff, Set.mem_Iio]; apply rpow_pos_of_pos b_pos
   · rw [logb_neg_eq_logb, logb_rpow b_pos b_ne_one]
 #align real.surj_on_logb' Real.surjOn_logb'
 
@@ -273,10 +268,8 @@ lean 3 declaration is
 but is expected to have type
   forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal x y) -> (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (Real.logb b y))
 Case conversion may be inaccurate. Consider using '#align real.logb_lt_logb Real.logb_lt_logbₓ'. -/
-theorem logb_lt_logb (hx : 0 < x) (hxy : x < y) : logb b x < logb b y :=
-  by
-  rw [logb, logb, div_lt_div_right (log_pos hb)]
-  exact log_lt_log hx hxy
+theorem logb_lt_logb (hx : 0 < x) (hxy : x < y) : logb b x < logb b y := by
+  rw [logb, logb, div_lt_div_right (log_pos hb)]; exact log_lt_log hx hxy
 #align real.logb_lt_logb Real.logb_lt_logb
 
 /- warning: real.logb_lt_logb_iff -> Real.logb_lt_logb_iff is a dubious translation:
@@ -286,10 +279,8 @@ but is expected to have type
   forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (Real.logb b y)) (LT.lt.{0} Real Real.instLTReal x y))
 Case conversion may be inaccurate. Consider using '#align real.logb_lt_logb_iff Real.logb_lt_logb_iffₓ'. -/
 @[simp]
-theorem logb_lt_logb_iff (hx : 0 < x) (hy : 0 < y) : logb b x < logb b y ↔ x < y :=
-  by
-  rw [logb, logb, div_lt_div_right (log_pos hb)]
-  exact log_lt_log_iff hx hy
+theorem logb_lt_logb_iff (hx : 0 < x) (hy : 0 < y) : logb b x < logb b y ↔ x < y := by
+  rw [logb, logb, div_lt_div_right (log_pos hb)]; exact log_lt_log_iff hx hy
 #align real.logb_lt_logb_iff Real.logb_lt_logb_iff
 
 /- warning: real.logb_le_iff_le_rpow -> Real.logb_le_iff_le_rpow is a dubious translation:
@@ -338,9 +329,7 @@ lean 3 declaration is
 but is expected to have type
   forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x)) (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x))
 Case conversion may be inaccurate. Consider using '#align real.logb_pos_iff Real.logb_pos_iffₓ'. -/
-theorem logb_pos_iff (hx : 0 < x) : 0 < logb b x ↔ 1 < x :=
-  by
-  rw [← @logb_one b]
+theorem logb_pos_iff (hx : 0 < x) : 0 < logb b x ↔ 1 < x := by rw [← @logb_one b];
   rw [logb_lt_logb_iff hb zero_lt_one hx]
 #align real.logb_pos_iff Real.logb_pos_iff
 
@@ -350,9 +339,7 @@ lean 3 declaration is
 but is expected to have type
   forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x))
 Case conversion may be inaccurate. Consider using '#align real.logb_pos Real.logb_posₓ'. -/
-theorem logb_pos (hx : 1 < x) : 0 < logb b x :=
-  by
-  rw [logb_pos_iff hb (lt_trans zero_lt_one hx)]
+theorem logb_pos (hx : 1 < x) : 0 < logb b x := by rw [logb_pos_iff hb (lt_trans zero_lt_one hx)];
   exact hx
 #align real.logb_pos Real.logb_pos
 
@@ -362,9 +349,7 @@ lean 3 declaration is
 but is expected to have type
   forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (LT.lt.{0} Real Real.instLTReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))))
 Case conversion may be inaccurate. Consider using '#align real.logb_neg_iff Real.logb_neg_iffₓ'. -/
-theorem logb_neg_iff (h : 0 < x) : logb b x < 0 ↔ x < 1 :=
-  by
-  rw [← logb_one]
+theorem logb_neg_iff (h : 0 < x) : logb b x < 0 ↔ x < 1 := by rw [← logb_one];
   exact logb_lt_logb_iff hb h zero_lt_one
 #align real.logb_neg_iff Real.logb_neg_iff
 
@@ -526,10 +511,8 @@ lean 3 declaration is
 but is expected to have type
   forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal x y) -> (LT.lt.{0} Real Real.instLTReal (Real.logb b y) (Real.logb b x))
 Case conversion may be inaccurate. Consider using '#align real.logb_lt_logb_of_base_lt_one Real.logb_lt_logb_of_base_lt_oneₓ'. -/
-theorem logb_lt_logb_of_base_lt_one (hx : 0 < x) (hxy : x < y) : logb b y < logb b x :=
-  by
-  rw [logb, logb, div_lt_div_right_of_neg (log_neg b_pos b_lt_one)]
-  exact log_lt_log hx hxy
+theorem logb_lt_logb_of_base_lt_one (hx : 0 < x) (hxy : x < y) : logb b y < logb b x := by
+  rw [logb, logb, div_lt_div_right_of_neg (log_neg b_pos b_lt_one)]; exact log_lt_log hx hxy
 #align real.logb_lt_logb_of_base_lt_one Real.logb_lt_logb_of_base_lt_one
 
 /- warning: real.logb_lt_logb_iff_of_base_lt_one -> Real.logb_lt_logb_iff_of_base_lt_one is a dubious translation:
@@ -540,9 +523,7 @@ but is expected to have type
 Case conversion may be inaccurate. Consider using '#align real.logb_lt_logb_iff_of_base_lt_one Real.logb_lt_logb_iff_of_base_lt_oneₓ'. -/
 @[simp]
 theorem logb_lt_logb_iff_of_base_lt_one (hx : 0 < x) (hy : 0 < y) : logb b x < logb b y ↔ y < x :=
-  by
-  rw [logb, logb, div_lt_div_right_of_neg (log_neg b_pos b_lt_one)]
-  exact log_lt_log_iff hy hx
+  by rw [logb, logb, div_lt_div_right_of_neg (log_neg b_pos b_lt_one)]; exact log_lt_log_iff hy hx
 #align real.logb_lt_logb_iff_of_base_lt_one Real.logb_lt_logb_iff_of_base_lt_one
 
 /- warning: real.logb_le_iff_le_rpow_of_base_lt_one -> Real.logb_le_iff_le_rpow_of_base_lt_one is a dubious translation:
@@ -601,10 +582,8 @@ lean 3 declaration is
 but is expected to have type
   forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x))
 Case conversion may be inaccurate. Consider using '#align real.logb_pos_of_base_lt_one Real.logb_pos_of_base_lt_oneₓ'. -/
-theorem logb_pos_of_base_lt_one (hx : 0 < x) (hx' : x < 1) : 0 < logb b x :=
-  by
-  rw [logb_pos_iff_of_base_lt_one b_pos b_lt_one hx]
-  exact hx'
+theorem logb_pos_of_base_lt_one (hx : 0 < x) (hx' : x < 1) : 0 < logb b x := by
+  rw [logb_pos_iff_of_base_lt_one b_pos b_lt_one hx]; exact hx'
 #align real.logb_pos_of_base_lt_one Real.logb_pos_of_base_lt_one
 
 /- warning: real.logb_neg_iff_of_base_lt_one -> Real.logb_neg_iff_of_base_lt_one is a dubious translation:
@@ -643,10 +622,8 @@ lean 3 declaration is
 but is expected to have type
   forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LE.le.{0} Real Real.instLEReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x))
 Case conversion may be inaccurate. Consider using '#align real.logb_nonneg_of_base_lt_one Real.logb_nonneg_of_base_lt_oneₓ'. -/
-theorem logb_nonneg_of_base_lt_one (hx : 0 < x) (hx' : x ≤ 1) : 0 ≤ logb b x :=
-  by
-  rw [logb_nonneg_iff_of_base_lt_one b_pos b_lt_one hx]
-  exact hx'
+theorem logb_nonneg_of_base_lt_one (hx : 0 < x) (hx' : x ≤ 1) : 0 ≤ logb b x := by
+  rw [logb_nonneg_iff_of_base_lt_one b_pos b_lt_one hx]; exact hx'
 #align real.logb_nonneg_of_base_lt_one Real.logb_nonneg_of_base_lt_one
 
 /- warning: real.logb_nonpos_iff_of_base_lt_one -> Real.logb_nonpos_iff_of_base_lt_one is a dubious translation:
Diff
@@ -143,7 +143,6 @@ private theorem log_b_ne_zero : log b ≠ 0 :=
   have b_ne_zero : b ≠ 0; linarith
   have b_ne_minus_one : b ≠ -1; linarith
   simp [b_ne_one, b_ne_zero, b_ne_minus_one]
-#align real.log_b_ne_zero real.log_b_ne_zero
 
 /- warning: real.logb_rpow -> Real.logb_rpow is a dubious translation:
 lean 3 declaration is
@@ -254,10 +253,8 @@ variable (hb : 1 < b)
 include hb
 
 private theorem b_pos : 0 < b := by linarith
-#align real.b_pos real.b_pos
 
 private theorem b_ne_one : b ≠ 1 := by linarith
-#align real.b_ne_one real.b_ne_one
 
 /- warning: real.logb_le_logb -> Real.logb_le_logb is a dubious translation:
 lean 3 declaration is
@@ -509,7 +506,6 @@ variable (b_lt_one : b < 1)
 include b_lt_one
 
 private theorem b_ne_one : b ≠ 1 := by linarith
-#align real.b_ne_one real.b_ne_one
 
 include b_pos
 
Diff
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
 Authors: Bolton Bailey, Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne
 
 ! This file was ported from Lean 3 source module analysis.special_functions.log.base
-! leanprover-community/mathlib commit 0b9eaaa7686280fad8cce467f5c3c57ee6ce77f8
+! leanprover-community/mathlib commit 1b0a28e1c93409dbf6d69526863cd9984ef652ce
 ! Please do not edit these lines, except to modify the commit id
 ! if you have ported upstream changes.
 -/
@@ -14,6 +14,9 @@ import Mathbin.Data.Int.Log
 /-!
 # Real logarithm base `b`
 
+> THIS FILE IS SYNCHRONIZED WITH MATHLIB4.
+> Any changes to this file require a corresponding PR to mathlib4.
+
 In this file we define `real.logb` to be the logarithm of a real number in a given base `b`. We
 define this as the division of the natural logarithms of the argument and the base, so that we have
 a globally defined function with `logb b 0 = 0`, `logb b (-x) = logb b x` `logb 0 x = 0` and
Diff
@@ -37,42 +37,92 @@ namespace Real
 
 variable {b x y : ℝ}
 
+#print Real.logb /-
 /-- The real logarithm in a given base. As with the natural logarithm, we define `logb b x` to
 be `logb b |x|` for `x < 0`, and `0` for `x = 0`.-/
 @[pp_nodot]
 noncomputable def logb (b x : ℝ) : ℝ :=
   log x / log b
 #align real.logb Real.logb
+-/
 
+/- warning: real.log_div_log -> Real.log_div_log is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, Eq.{1} Real (HDiv.hDiv.{0, 0, 0} Real Real Real (instHDiv.{0} Real (DivInvMonoid.toHasDiv.{0} Real (DivisionRing.toDivInvMonoid.{0} Real Real.divisionRing))) (Real.log x) (Real.log b)) (Real.logb b x)
+but is expected to have type
+  forall {b : Real} {x : Real}, Eq.{1} Real (HDiv.hDiv.{0, 0, 0} Real Real Real (instHDiv.{0} Real (LinearOrderedField.toDiv.{0} Real Real.instLinearOrderedFieldReal)) (Real.log x) (Real.log b)) (Real.logb b x)
+Case conversion may be inaccurate. Consider using '#align real.log_div_log Real.log_div_logₓ'. -/
 theorem log_div_log : log x / log b = logb b x :=
   rfl
 #align real.log_div_log Real.log_div_log
 
+/- warning: real.logb_zero -> Real.logb_zero is a dubious translation:
+lean 3 declaration is
+  forall {b : Real}, Eq.{1} Real (Real.logb b (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))
+but is expected to have type
+  forall {b : Real}, Eq.{1} Real (Real.logb b (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))
+Case conversion may be inaccurate. Consider using '#align real.logb_zero Real.logb_zeroₓ'. -/
 @[simp]
 theorem logb_zero : logb b 0 = 0 := by simp [logb]
 #align real.logb_zero Real.logb_zero
 
+/- warning: real.logb_one -> Real.logb_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real}, Eq.{1} Real (Real.logb b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))
+but is expected to have type
+  forall {b : Real}, Eq.{1} Real (Real.logb b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))
+Case conversion may be inaccurate. Consider using '#align real.logb_one Real.logb_oneₓ'. -/
 @[simp]
 theorem logb_one : logb b 1 = 0 := by simp [logb]
 #align real.logb_one Real.logb_one
 
+/- warning: real.logb_abs -> Real.logb_abs is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} (x : Real), Eq.{1} Real (Real.logb b (Abs.abs.{0} Real (Neg.toHasAbs.{0} Real Real.hasNeg Real.hasSup) x)) (Real.logb b x)
+but is expected to have type
+  forall {b : Real} (x : Real), Eq.{1} Real (Real.logb b (Abs.abs.{0} Real (Neg.toHasAbs.{0} Real Real.instNegReal Real.instSupReal) x)) (Real.logb b x)
+Case conversion may be inaccurate. Consider using '#align real.logb_abs Real.logb_absₓ'. -/
 @[simp]
 theorem logb_abs (x : ℝ) : logb b (|x|) = logb b x := by rw [logb, logb, log_abs]
 #align real.logb_abs Real.logb_abs
 
+/- warning: real.logb_neg_eq_logb -> Real.logb_neg_eq_logb is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} (x : Real), Eq.{1} Real (Real.logb b (Neg.neg.{0} Real Real.hasNeg x)) (Real.logb b x)
+but is expected to have type
+  forall {b : Real} (x : Real), Eq.{1} Real (Real.logb b (Neg.neg.{0} Real Real.instNegReal x)) (Real.logb b x)
+Case conversion may be inaccurate. Consider using '#align real.logb_neg_eq_logb Real.logb_neg_eq_logbₓ'. -/
 @[simp]
 theorem logb_neg_eq_logb (x : ℝ) : logb b (-x) = logb b x := by
   rw [← logb_abs x, ← logb_abs (-x), abs_neg]
 #align real.logb_neg_eq_logb Real.logb_neg_eq_logb
 
+/- warning: real.logb_mul -> Real.logb_mul is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (Ne.{1} Real x (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Ne.{1} Real y (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Eq.{1} Real (Real.logb b (HMul.hMul.{0, 0, 0} Real Real Real (instHMul.{0} Real Real.hasMul) x y)) (HAdd.hAdd.{0, 0, 0} Real Real Real (instHAdd.{0} Real Real.hasAdd) (Real.logb b x) (Real.logb b y)))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (Ne.{1} Real x (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Ne.{1} Real y (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Eq.{1} Real (Real.logb b (HMul.hMul.{0, 0, 0} Real Real Real (instHMul.{0} Real Real.instMulReal) x y)) (HAdd.hAdd.{0, 0, 0} Real Real Real (instHAdd.{0} Real Real.instAddReal) (Real.logb b x) (Real.logb b y)))
+Case conversion may be inaccurate. Consider using '#align real.logb_mul Real.logb_mulₓ'. -/
 theorem logb_mul (hx : x ≠ 0) (hy : y ≠ 0) : logb b (x * y) = logb b x + logb b y := by
   simp_rw [logb, log_mul hx hy, add_div]
 #align real.logb_mul Real.logb_mul
 
+/- warning: real.logb_div -> Real.logb_div is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (Ne.{1} Real x (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Ne.{1} Real y (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Eq.{1} Real (Real.logb b (HDiv.hDiv.{0, 0, 0} Real Real Real (instHDiv.{0} Real (DivInvMonoid.toHasDiv.{0} Real (DivisionRing.toDivInvMonoid.{0} Real Real.divisionRing))) x y)) (HSub.hSub.{0, 0, 0} Real Real Real (instHSub.{0} Real Real.hasSub) (Real.logb b x) (Real.logb b y)))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (Ne.{1} Real x (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Ne.{1} Real y (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Eq.{1} Real (Real.logb b (HDiv.hDiv.{0, 0, 0} Real Real Real (instHDiv.{0} Real (LinearOrderedField.toDiv.{0} Real Real.instLinearOrderedFieldReal)) x y)) (HSub.hSub.{0, 0, 0} Real Real Real (instHSub.{0} Real Real.instSubReal) (Real.logb b x) (Real.logb b y)))
+Case conversion may be inaccurate. Consider using '#align real.logb_div Real.logb_divₓ'. -/
 theorem logb_div (hx : x ≠ 0) (hy : y ≠ 0) : logb b (x / y) = logb b x - logb b y := by
   simp_rw [logb, log_div hx hy, sub_div]
 #align real.logb_div Real.logb_div
 
+/- warning: real.logb_inv -> Real.logb_inv is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} (x : Real), Eq.{1} Real (Real.logb b (Inv.inv.{0} Real Real.hasInv x)) (Neg.neg.{0} Real Real.hasNeg (Real.logb b x))
+but is expected to have type
+  forall {b : Real} (x : Real), Eq.{1} Real (Real.logb b (Inv.inv.{0} Real Real.instInvReal x)) (Neg.neg.{0} Real Real.instNegReal (Real.logb b x))
+Case conversion may be inaccurate. Consider using '#align real.logb_inv Real.logb_invₓ'. -/
 @[simp]
 theorem logb_inv (x : ℝ) : logb b x⁻¹ = -logb b x := by simp [logb, neg_div]
 #align real.logb_inv Real.logb_inv
@@ -92,6 +142,12 @@ private theorem log_b_ne_zero : log b ≠ 0 :=
   simp [b_ne_one, b_ne_zero, b_ne_minus_one]
 #align real.log_b_ne_zero real.log_b_ne_zero
 
+/- warning: real.logb_rpow -> Real.logb_rpow is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Eq.{1} Real (Real.logb b (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b x)) x)
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Eq.{1} Real (Real.logb b (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b x)) x)
+Case conversion may be inaccurate. Consider using '#align real.logb_rpow Real.logb_rpowₓ'. -/
 @[simp]
 theorem logb_rpow : logb b (b ^ x) = x :=
   by
@@ -99,6 +155,12 @@ theorem logb_rpow : logb b (b ^ x) = x :=
   exact log_b_ne_zero b_pos b_ne_one
 #align real.logb_rpow Real.logb_rpow
 
+/- warning: real.rpow_logb_eq_abs -> Real.rpow_logb_eq_abs is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Ne.{1} Real x (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Eq.{1} Real (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b (Real.logb b x)) (Abs.abs.{0} Real (Neg.toHasAbs.{0} Real Real.hasNeg Real.hasSup) x))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Ne.{1} Real x (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Eq.{1} Real (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b (Real.logb b x)) (Abs.abs.{0} Real (Neg.toHasAbs.{0} Real Real.instNegReal Real.instSupReal) x))
+Case conversion may be inaccurate. Consider using '#align real.rpow_logb_eq_abs Real.rpow_logb_eq_absₓ'. -/
 theorem rpow_logb_eq_abs (hx : x ≠ 0) : b ^ logb b x = |x| :=
   by
   apply log_inj_on_pos
@@ -109,6 +171,12 @@ theorem rpow_logb_eq_abs (hx : x ≠ 0) : b ^ logb b x = |x| :=
   field_simp [log_b_ne_zero b_pos b_ne_one]
 #align real.rpow_logb_eq_abs Real.rpow_logb_eq_abs
 
+/- warning: real.rpow_logb -> Real.rpow_logb is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Eq.{1} Real (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b (Real.logb b x)) x)
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Eq.{1} Real (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b (Real.logb b x)) x)
+Case conversion may be inaccurate. Consider using '#align real.rpow_logb Real.rpow_logbₓ'. -/
 @[simp]
 theorem rpow_logb (hx : 0 < x) : b ^ logb b x = x :=
   by
@@ -116,24 +184,54 @@ theorem rpow_logb (hx : 0 < x) : b ^ logb b x = x :=
   exact abs_of_pos hx
 #align real.rpow_logb Real.rpow_logb
 
+/- warning: real.rpow_logb_of_neg -> Real.rpow_logb_of_neg is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt x (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Eq.{1} Real (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b (Real.logb b x)) (Neg.neg.{0} Real Real.hasNeg x))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal x (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Eq.{1} Real (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b (Real.logb b x)) (Neg.neg.{0} Real Real.instNegReal x))
+Case conversion may be inaccurate. Consider using '#align real.rpow_logb_of_neg Real.rpow_logb_of_negₓ'. -/
 theorem rpow_logb_of_neg (hx : x < 0) : b ^ logb b x = -x :=
   by
   rw [rpow_logb_eq_abs b_pos b_ne_one (ne_of_lt hx)]
   exact abs_of_neg hx
 #align real.rpow_logb_of_neg Real.rpow_logb_of_neg
 
+/- warning: real.surj_on_logb -> Real.surjOn_logb is a dubious translation:
+lean 3 declaration is
+  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Set.SurjOn.{0, 0} Real Real (Real.logb b) (Set.Ioi.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (Set.univ.{0} Real))
+but is expected to have type
+  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Set.SurjOn.{0, 0} Real Real (Real.logb b) (Set.Ioi.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (Set.univ.{0} Real))
+Case conversion may be inaccurate. Consider using '#align real.surj_on_logb Real.surjOn_logbₓ'. -/
 theorem surjOn_logb : SurjOn (logb b) (Ioi 0) univ := fun x _ =>
   ⟨rpow b x, rpow_pos_of_pos b_pos x, logb_rpow b_pos b_ne_one⟩
 #align real.surj_on_logb Real.surjOn_logb
 
+/- warning: real.logb_surjective -> Real.logb_surjective is a dubious translation:
+lean 3 declaration is
+  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Function.Surjective.{1, 1} Real Real (Real.logb b))
+but is expected to have type
+  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Function.Surjective.{1, 1} Real Real (Real.logb b))
+Case conversion may be inaccurate. Consider using '#align real.logb_surjective Real.logb_surjectiveₓ'. -/
 theorem logb_surjective : Surjective (logb b) := fun x => ⟨b ^ x, logb_rpow b_pos b_ne_one⟩
 #align real.logb_surjective Real.logb_surjective
 
+/- warning: real.range_logb -> Real.range_logb is a dubious translation:
+lean 3 declaration is
+  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Eq.{1} (Set.{0} Real) (Set.range.{0, 1} Real Real (Real.logb b)) (Set.univ.{0} Real))
+but is expected to have type
+  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Eq.{1} (Set.{0} Real) (Set.range.{0, 1} Real Real (Real.logb b)) (Set.univ.{0} Real))
+Case conversion may be inaccurate. Consider using '#align real.range_logb Real.range_logbₓ'. -/
 @[simp]
 theorem range_logb : range (logb b) = univ :=
   (logb_surjective b_pos b_ne_one).range_eq
 #align real.range_logb Real.range_logb
 
+/- warning: real.surj_on_logb' -> Real.surjOn_logb' is a dubious translation:
+lean 3 declaration is
+  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Set.SurjOn.{0, 0} Real Real (Real.logb b) (Set.Iio.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (Set.univ.{0} Real))
+but is expected to have type
+  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (Ne.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Set.SurjOn.{0, 0} Real Real (Real.logb b) (Set.Iio.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (Set.univ.{0} Real))
+Case conversion may be inaccurate. Consider using '#align real.surj_on_logb' Real.surjOn_logb'ₓ'. -/
 theorem surjOn_logb' : SurjOn (logb b) (Iio 0) univ :=
   by
   intro x x_in_univ
@@ -158,17 +256,35 @@ private theorem b_pos : 0 < b := by linarith
 private theorem b_ne_one : b ≠ 1 := by linarith
 #align real.b_ne_one real.b_ne_one
 
+/- warning: real.logb_le_logb -> Real.logb_le_logb is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LE.le.{0} Real Real.hasLe (Real.logb b x) (Real.logb b y)) (LE.le.{0} Real Real.hasLe x y))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LE.le.{0} Real Real.instLEReal (Real.logb b x) (Real.logb b y)) (LE.le.{0} Real Real.instLEReal x y))
+Case conversion may be inaccurate. Consider using '#align real.logb_le_logb Real.logb_le_logbₓ'. -/
 @[simp]
 theorem logb_le_logb (h : 0 < x) (h₁ : 0 < y) : logb b x ≤ logb b y ↔ x ≤ y := by
   rw [logb, logb, div_le_div_right (log_pos hb), log_le_log h h₁]
 #align real.logb_le_logb Real.logb_le_logb
 
+/- warning: real.logb_lt_logb -> Real.logb_lt_logb is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt x y) -> (LT.lt.{0} Real Real.hasLt (Real.logb b x) (Real.logb b y))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal x y) -> (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (Real.logb b y))
+Case conversion may be inaccurate. Consider using '#align real.logb_lt_logb Real.logb_lt_logbₓ'. -/
 theorem logb_lt_logb (hx : 0 < x) (hxy : x < y) : logb b x < logb b y :=
   by
   rw [logb, logb, div_lt_div_right (log_pos hb)]
   exact log_lt_log hx hxy
 #align real.logb_lt_logb Real.logb_lt_logb
 
+/- warning: real.logb_lt_logb_iff -> Real.logb_lt_logb_iff is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LT.lt.{0} Real Real.hasLt (Real.logb b x) (Real.logb b y)) (LT.lt.{0} Real Real.hasLt x y))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (Real.logb b y)) (LT.lt.{0} Real Real.instLTReal x y))
+Case conversion may be inaccurate. Consider using '#align real.logb_lt_logb_iff Real.logb_lt_logb_iffₓ'. -/
 @[simp]
 theorem logb_lt_logb_iff (hx : 0 < x) (hy : 0 < y) : logb b x < logb b y ↔ x < y :=
   by
@@ -176,56 +292,128 @@ theorem logb_lt_logb_iff (hx : 0 < x) (hy : 0 < y) : logb b x < logb b y ↔ x <
   exact log_lt_log_iff hx hy
 #align real.logb_lt_logb_iff Real.logb_lt_logb_iff
 
+/- warning: real.logb_le_iff_le_rpow -> Real.logb_le_iff_le_rpow is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LE.le.{0} Real Real.hasLe (Real.logb b x) y) (LE.le.{0} Real Real.hasLe x (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b y)))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LE.le.{0} Real Real.instLEReal (Real.logb b x) y) (LE.le.{0} Real Real.instLEReal x (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b y)))
+Case conversion may be inaccurate. Consider using '#align real.logb_le_iff_le_rpow Real.logb_le_iff_le_rpowₓ'. -/
 theorem logb_le_iff_le_rpow (hx : 0 < x) : logb b x ≤ y ↔ x ≤ b ^ y := by
   rw [← rpow_le_rpow_left_iff hb, rpow_logb (b_pos hb) (b_ne_one hb) hx]
 #align real.logb_le_iff_le_rpow Real.logb_le_iff_le_rpow
 
+/- warning: real.logb_lt_iff_lt_rpow -> Real.logb_lt_iff_lt_rpow is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LT.lt.{0} Real Real.hasLt (Real.logb b x) y) (LT.lt.{0} Real Real.hasLt x (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b y)))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LT.lt.{0} Real Real.instLTReal (Real.logb b x) y) (LT.lt.{0} Real Real.instLTReal x (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b y)))
+Case conversion may be inaccurate. Consider using '#align real.logb_lt_iff_lt_rpow Real.logb_lt_iff_lt_rpowₓ'. -/
 theorem logb_lt_iff_lt_rpow (hx : 0 < x) : logb b x < y ↔ x < b ^ y := by
   rw [← rpow_lt_rpow_left_iff hb, rpow_logb (b_pos hb) (b_ne_one hb) hx]
 #align real.logb_lt_iff_lt_rpow Real.logb_lt_iff_lt_rpow
 
+/- warning: real.le_logb_iff_rpow_le -> Real.le_logb_iff_rpow_le is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LE.le.{0} Real Real.hasLe x (Real.logb b y)) (LE.le.{0} Real Real.hasLe (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b x) y))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LE.le.{0} Real Real.instLEReal x (Real.logb b y)) (LE.le.{0} Real Real.instLEReal (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b x) y))
+Case conversion may be inaccurate. Consider using '#align real.le_logb_iff_rpow_le Real.le_logb_iff_rpow_leₓ'. -/
 theorem le_logb_iff_rpow_le (hy : 0 < y) : x ≤ logb b y ↔ b ^ x ≤ y := by
   rw [← rpow_le_rpow_left_iff hb, rpow_logb (b_pos hb) (b_ne_one hb) hy]
 #align real.le_logb_iff_rpow_le Real.le_logb_iff_rpow_le
 
+/- warning: real.lt_logb_iff_rpow_lt -> Real.lt_logb_iff_rpow_lt is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LT.lt.{0} Real Real.hasLt x (Real.logb b y)) (LT.lt.{0} Real Real.hasLt (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b x) y))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LT.lt.{0} Real Real.instLTReal x (Real.logb b y)) (LT.lt.{0} Real Real.instLTReal (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b x) y))
+Case conversion may be inaccurate. Consider using '#align real.lt_logb_iff_rpow_lt Real.lt_logb_iff_rpow_ltₓ'. -/
 theorem lt_logb_iff_rpow_lt (hy : 0 < y) : x < logb b y ↔ b ^ x < y := by
   rw [← rpow_lt_rpow_left_iff hb, rpow_logb (b_pos hb) (b_ne_one hb) hy]
 #align real.lt_logb_iff_rpow_lt Real.lt_logb_iff_rpow_lt
 
+/- warning: real.logb_pos_iff -> Real.logb_pos_iff is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x)) (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) x))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x)) (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x))
+Case conversion may be inaccurate. Consider using '#align real.logb_pos_iff Real.logb_pos_iffₓ'. -/
 theorem logb_pos_iff (hx : 0 < x) : 0 < logb b x ↔ 1 < x :=
   by
   rw [← @logb_one b]
   rw [logb_lt_logb_iff hb zero_lt_one hx]
 #align real.logb_pos_iff Real.logb_pos_iff
 
+/- warning: real.logb_pos -> Real.logb_pos is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) x) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x))
+Case conversion may be inaccurate. Consider using '#align real.logb_pos Real.logb_posₓ'. -/
 theorem logb_pos (hx : 1 < x) : 0 < logb b x :=
   by
   rw [logb_pos_iff hb (lt_trans zero_lt_one hx)]
   exact hx
 #align real.logb_pos Real.logb_pos
 
+/- warning: real.logb_neg_iff -> Real.logb_neg_iff is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LT.lt.{0} Real Real.hasLt (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (LT.lt.{0} Real Real.hasLt x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (LT.lt.{0} Real Real.instLTReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))))
+Case conversion may be inaccurate. Consider using '#align real.logb_neg_iff Real.logb_neg_iffₓ'. -/
 theorem logb_neg_iff (h : 0 < x) : logb b x < 0 ↔ x < 1 :=
   by
   rw [← logb_one]
   exact logb_lt_logb_iff hb h zero_lt_one
 #align real.logb_neg_iff Real.logb_neg_iff
 
+/- warning: real.logb_neg -> Real.logb_neg is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)))
+Case conversion may be inaccurate. Consider using '#align real.logb_neg Real.logb_negₓ'. -/
 theorem logb_neg (h0 : 0 < x) (h1 : x < 1) : logb b x < 0 :=
   (logb_neg_iff hb h0).2 h1
 #align real.logb_neg Real.logb_neg
 
+/- warning: real.logb_nonneg_iff -> Real.logb_nonneg_iff is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x)) (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) x))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x)) (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x))
+Case conversion may be inaccurate. Consider using '#align real.logb_nonneg_iff Real.logb_nonneg_iffₓ'. -/
 theorem logb_nonneg_iff (hx : 0 < x) : 0 ≤ logb b x ↔ 1 ≤ x := by
   rw [← not_lt, logb_neg_iff hb hx, not_lt]
 #align real.logb_nonneg_iff Real.logb_nonneg_iff
 
+/- warning: real.logb_nonneg -> Real.logb_nonneg is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) x) -> (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x))
+Case conversion may be inaccurate. Consider using '#align real.logb_nonneg Real.logb_nonnegₓ'. -/
 theorem logb_nonneg (hx : 1 ≤ x) : 0 ≤ logb b x :=
   (logb_nonneg_iff hb (zero_lt_one.trans_le hx)).2 hx
 #align real.logb_nonneg Real.logb_nonneg
 
+/- warning: real.logb_nonpos_iff -> Real.logb_nonpos_iff is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LE.le.{0} Real Real.hasLe (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (LE.le.{0} Real Real.hasLe x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LE.le.{0} Real Real.instLEReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (LE.le.{0} Real Real.instLEReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))))
+Case conversion may be inaccurate. Consider using '#align real.logb_nonpos_iff Real.logb_nonpos_iffₓ'. -/
 theorem logb_nonpos_iff (hx : 0 < x) : logb b x ≤ 0 ↔ x ≤ 1 := by
   rw [← not_lt, logb_pos_iff hb hx, not_lt]
 #align real.logb_nonpos_iff Real.logb_nonpos_iff
 
+/- warning: real.logb_nonpos_iff' -> Real.logb_nonpos_iff' is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LE.le.{0} Real Real.hasLe (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (LE.le.{0} Real Real.hasLe x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LE.le.{0} Real Real.instLEReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (LE.le.{0} Real Real.instLEReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))))
+Case conversion may be inaccurate. Consider using '#align real.logb_nonpos_iff' Real.logb_nonpos_iff'ₓ'. -/
 theorem logb_nonpos_iff' (hx : 0 ≤ x) : logb b x ≤ 0 ↔ x ≤ 1 :=
   by
   rcases hx.eq_or_lt with (rfl | hx)
@@ -233,14 +421,32 @@ theorem logb_nonpos_iff' (hx : 0 ≤ x) : logb b x ≤ 0 ↔ x ≤ 1 :=
   exact logb_nonpos_iff hb hx
 #align real.logb_nonpos_iff' Real.logb_nonpos_iff'
 
+/- warning: real.logb_nonpos -> Real.logb_nonpos is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LE.le.{0} Real Real.hasLe x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LE.le.{0} Real Real.hasLe (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LE.le.{0} Real Real.instLEReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LE.le.{0} Real Real.instLEReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)))
+Case conversion may be inaccurate. Consider using '#align real.logb_nonpos Real.logb_nonposₓ'. -/
 theorem logb_nonpos (hx : 0 ≤ x) (h'x : x ≤ 1) : logb b x ≤ 0 :=
   (logb_nonpos_iff' hb hx).2 h'x
 #align real.logb_nonpos Real.logb_nonpos
 
+/- warning: real.strict_mono_on_logb -> Real.strictMonoOn_logb is a dubious translation:
+lean 3 declaration is
+  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (StrictMonoOn.{0, 0} Real Real Real.preorder Real.preorder (Real.logb b) (Set.Ioi.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))))
+but is expected to have type
+  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (StrictMonoOn.{0, 0} Real Real Real.instPreorderReal Real.instPreorderReal (Real.logb b) (Set.Ioi.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))))
+Case conversion may be inaccurate. Consider using '#align real.strict_mono_on_logb Real.strictMonoOn_logbₓ'. -/
 theorem strictMonoOn_logb : StrictMonoOn (logb b) (Set.Ioi 0) := fun x hx y hy hxy =>
   logb_lt_logb hb hx hxy
 #align real.strict_mono_on_logb Real.strictMonoOn_logb
 
+/- warning: real.strict_anti_on_logb -> Real.strictAntiOn_logb is a dubious translation:
+lean 3 declaration is
+  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (StrictAntiOn.{0, 0} Real Real Real.preorder Real.preorder (Real.logb b) (Set.Iio.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))))
+but is expected to have type
+  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (StrictAntiOn.{0, 0} Real Real Real.instPreorderReal Real.instPreorderReal (Real.logb b) (Set.Iio.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))))
+Case conversion may be inaccurate. Consider using '#align real.strict_anti_on_logb Real.strictAntiOn_logbₓ'. -/
 theorem strictAntiOn_logb : StrictAntiOn (logb b) (Set.Iio 0) :=
   by
   rintro x (hx : x < 0) y (hy : y < 0) hxy
@@ -249,18 +455,42 @@ theorem strictAntiOn_logb : StrictAntiOn (logb b) (Set.Iio 0) :=
   rwa [abs_of_neg hy, abs_of_neg hx, neg_lt_neg_iff]
 #align real.strict_anti_on_logb Real.strictAntiOn_logb
 
+/- warning: real.logb_inj_on_pos -> Real.logb_injOn_pos is a dubious translation:
+lean 3 declaration is
+  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (Set.InjOn.{0, 0} Real Real (Real.logb b) (Set.Ioi.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))))
+but is expected to have type
+  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (Set.InjOn.{0, 0} Real Real (Real.logb b) (Set.Ioi.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))))
+Case conversion may be inaccurate. Consider using '#align real.logb_inj_on_pos Real.logb_injOn_posₓ'. -/
 theorem logb_injOn_pos : Set.InjOn (logb b) (Set.Ioi 0) :=
   (strictMonoOn_logb hb).InjOn
 #align real.logb_inj_on_pos Real.logb_injOn_pos
 
+/- warning: real.eq_one_of_pos_of_logb_eq_zero -> Real.eq_one_of_pos_of_logb_eq_zero is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Eq.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Eq.{1} Real x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Eq.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Eq.{1} Real x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)))
+Case conversion may be inaccurate. Consider using '#align real.eq_one_of_pos_of_logb_eq_zero Real.eq_one_of_pos_of_logb_eq_zeroₓ'. -/
 theorem eq_one_of_pos_of_logb_eq_zero (h₁ : 0 < x) (h₂ : logb b x = 0) : x = 1 :=
   logb_injOn_pos hb (Set.mem_Ioi.2 h₁) (Set.mem_Ioi.2 zero_lt_one) (h₂.trans Real.logb_one.symm)
 #align real.eq_one_of_pos_of_logb_eq_zero Real.eq_one_of_pos_of_logb_eq_zero
 
+/- warning: real.logb_ne_zero_of_pos_of_ne_one -> Real.logb_ne_zero_of_pos_of_ne_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Ne.{1} Real x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Ne.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Ne.{1} Real x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Ne.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)))
+Case conversion may be inaccurate. Consider using '#align real.logb_ne_zero_of_pos_of_ne_one Real.logb_ne_zero_of_pos_of_ne_oneₓ'. -/
 theorem logb_ne_zero_of_pos_of_ne_one (hx_pos : 0 < x) (hx : x ≠ 1) : logb b x ≠ 0 :=
   mt (eq_one_of_pos_of_logb_eq_zero hb hx_pos) hx
 #align real.logb_ne_zero_of_pos_of_ne_one Real.logb_ne_zero_of_pos_of_ne_one
 
+/- warning: real.tendsto_logb_at_top -> Real.tendsto_logb_atTop is a dubious translation:
+lean 3 declaration is
+  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) b) -> (Filter.Tendsto.{0, 0} Real Real (Real.logb b) (Filter.atTop.{0} Real Real.preorder) (Filter.atTop.{0} Real Real.preorder))
+but is expected to have type
+  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) b) -> (Filter.Tendsto.{0, 0} Real Real (Real.logb b) (Filter.atTop.{0} Real Real.instPreorderReal) (Filter.atTop.{0} Real Real.instPreorderReal))
+Case conversion may be inaccurate. Consider using '#align real.tendsto_logb_at_top Real.tendsto_logb_atTopₓ'. -/
 theorem tendsto_logb_atTop : Tendsto (logb b) atTop atTop :=
   Tendsto.atTop_div_const (log_pos hb) tendsto_log_atTop
 #align real.tendsto_logb_at_top Real.tendsto_logb_atTop
@@ -280,17 +510,35 @@ private theorem b_ne_one : b ≠ 1 := by linarith
 
 include b_pos
 
+/- warning: real.logb_le_logb_of_base_lt_one -> Real.logb_le_logb_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LE.le.{0} Real Real.hasLe (Real.logb b x) (Real.logb b y)) (LE.le.{0} Real Real.hasLe y x))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LE.le.{0} Real Real.instLEReal (Real.logb b x) (Real.logb b y)) (LE.le.{0} Real Real.instLEReal y x))
+Case conversion may be inaccurate. Consider using '#align real.logb_le_logb_of_base_lt_one Real.logb_le_logb_of_base_lt_oneₓ'. -/
 @[simp]
 theorem logb_le_logb_of_base_lt_one (h : 0 < x) (h₁ : 0 < y) : logb b x ≤ logb b y ↔ y ≤ x := by
   rw [logb, logb, div_le_div_right_of_neg (log_neg b_pos b_lt_one), log_le_log h₁ h]
 #align real.logb_le_logb_of_base_lt_one Real.logb_le_logb_of_base_lt_one
 
+/- warning: real.logb_lt_logb_of_base_lt_one -> Real.logb_lt_logb_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt x y) -> (LT.lt.{0} Real Real.hasLt (Real.logb b y) (Real.logb b x))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal x y) -> (LT.lt.{0} Real Real.instLTReal (Real.logb b y) (Real.logb b x))
+Case conversion may be inaccurate. Consider using '#align real.logb_lt_logb_of_base_lt_one Real.logb_lt_logb_of_base_lt_oneₓ'. -/
 theorem logb_lt_logb_of_base_lt_one (hx : 0 < x) (hxy : x < y) : logb b y < logb b x :=
   by
   rw [logb, logb, div_lt_div_right_of_neg (log_neg b_pos b_lt_one)]
   exact log_lt_log hx hxy
 #align real.logb_lt_logb_of_base_lt_one Real.logb_lt_logb_of_base_lt_one
 
+/- warning: real.logb_lt_logb_iff_of_base_lt_one -> Real.logb_lt_logb_iff_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LT.lt.{0} Real Real.hasLt (Real.logb b x) (Real.logb b y)) (LT.lt.{0} Real Real.hasLt y x))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (Real.logb b y)) (LT.lt.{0} Real Real.instLTReal y x))
+Case conversion may be inaccurate. Consider using '#align real.logb_lt_logb_iff_of_base_lt_one Real.logb_lt_logb_iff_of_base_lt_oneₓ'. -/
 @[simp]
 theorem logb_lt_logb_iff_of_base_lt_one (hx : 0 < x) (hy : 0 < y) : logb b x < logb b y ↔ y < x :=
   by
@@ -298,58 +546,136 @@ theorem logb_lt_logb_iff_of_base_lt_one (hx : 0 < x) (hy : 0 < y) : logb b x < l
   exact log_lt_log_iff hy hx
 #align real.logb_lt_logb_iff_of_base_lt_one Real.logb_lt_logb_iff_of_base_lt_one
 
+/- warning: real.logb_le_iff_le_rpow_of_base_lt_one -> Real.logb_le_iff_le_rpow_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LE.le.{0} Real Real.hasLe (Real.logb b x) y) (LE.le.{0} Real Real.hasLe (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b y) x))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LE.le.{0} Real Real.instLEReal (Real.logb b x) y) (LE.le.{0} Real Real.instLEReal (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b y) x))
+Case conversion may be inaccurate. Consider using '#align real.logb_le_iff_le_rpow_of_base_lt_one Real.logb_le_iff_le_rpow_of_base_lt_oneₓ'. -/
 theorem logb_le_iff_le_rpow_of_base_lt_one (hx : 0 < x) : logb b x ≤ y ↔ b ^ y ≤ x := by
   rw [← rpow_le_rpow_left_iff_of_base_lt_one b_pos b_lt_one, rpow_logb b_pos (b_ne_one b_lt_one) hx]
 #align real.logb_le_iff_le_rpow_of_base_lt_one Real.logb_le_iff_le_rpow_of_base_lt_one
 
+/- warning: real.logb_lt_iff_lt_rpow_of_base_lt_one -> Real.logb_lt_iff_lt_rpow_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LT.lt.{0} Real Real.hasLt (Real.logb b x) y) (LT.lt.{0} Real Real.hasLt (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b y) x))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LT.lt.{0} Real Real.instLTReal (Real.logb b x) y) (LT.lt.{0} Real Real.instLTReal (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b y) x))
+Case conversion may be inaccurate. Consider using '#align real.logb_lt_iff_lt_rpow_of_base_lt_one Real.logb_lt_iff_lt_rpow_of_base_lt_oneₓ'. -/
 theorem logb_lt_iff_lt_rpow_of_base_lt_one (hx : 0 < x) : logb b x < y ↔ b ^ y < x := by
   rw [← rpow_lt_rpow_left_iff_of_base_lt_one b_pos b_lt_one, rpow_logb b_pos (b_ne_one b_lt_one) hx]
 #align real.logb_lt_iff_lt_rpow_of_base_lt_one Real.logb_lt_iff_lt_rpow_of_base_lt_one
 
+/- warning: real.le_logb_iff_rpow_le_of_base_lt_one -> Real.le_logb_iff_rpow_le_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LE.le.{0} Real Real.hasLe x (Real.logb b y)) (LE.le.{0} Real Real.hasLe y (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b x)))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LE.le.{0} Real Real.instLEReal x (Real.logb b y)) (LE.le.{0} Real Real.instLEReal y (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b x)))
+Case conversion may be inaccurate. Consider using '#align real.le_logb_iff_rpow_le_of_base_lt_one Real.le_logb_iff_rpow_le_of_base_lt_oneₓ'. -/
 theorem le_logb_iff_rpow_le_of_base_lt_one (hy : 0 < y) : x ≤ logb b y ↔ y ≤ b ^ x := by
   rw [← rpow_le_rpow_left_iff_of_base_lt_one b_pos b_lt_one, rpow_logb b_pos (b_ne_one b_lt_one) hy]
 #align real.le_logb_iff_rpow_le_of_base_lt_one Real.le_logb_iff_rpow_le_of_base_lt_one
 
+/- warning: real.lt_logb_iff_rpow_lt_of_base_lt_one -> Real.lt_logb_iff_rpow_lt_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) y) -> (Iff (LT.lt.{0} Real Real.hasLt x (Real.logb b y)) (LT.lt.{0} Real Real.hasLt y (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.hasPow) b x)))
+but is expected to have type
+  forall {b : Real} {x : Real} {y : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) y) -> (Iff (LT.lt.{0} Real Real.instLTReal x (Real.logb b y)) (LT.lt.{0} Real Real.instLTReal y (HPow.hPow.{0, 0, 0} Real Real Real (instHPow.{0, 0} Real Real Real.instPowReal) b x)))
+Case conversion may be inaccurate. Consider using '#align real.lt_logb_iff_rpow_lt_of_base_lt_one Real.lt_logb_iff_rpow_lt_of_base_lt_oneₓ'. -/
 theorem lt_logb_iff_rpow_lt_of_base_lt_one (hy : 0 < y) : x < logb b y ↔ y < b ^ x := by
   rw [← rpow_lt_rpow_left_iff_of_base_lt_one b_pos b_lt_one, rpow_logb b_pos (b_ne_one b_lt_one) hy]
 #align real.lt_logb_iff_rpow_lt_of_base_lt_one Real.lt_logb_iff_rpow_lt_of_base_lt_one
 
+/- warning: real.logb_pos_iff_of_base_lt_one -> Real.logb_pos_iff_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x)) (LT.lt.{0} Real Real.hasLt x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x)) (LT.lt.{0} Real Real.instLTReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))))
+Case conversion may be inaccurate. Consider using '#align real.logb_pos_iff_of_base_lt_one Real.logb_pos_iff_of_base_lt_oneₓ'. -/
 theorem logb_pos_iff_of_base_lt_one (hx : 0 < x) : 0 < logb b x ↔ x < 1 := by
   rw [← @logb_one b, logb_lt_logb_iff_of_base_lt_one b_pos b_lt_one zero_lt_one hx]
 #align real.logb_pos_iff_of_base_lt_one Real.logb_pos_iff_of_base_lt_one
 
+/- warning: real.logb_pos_of_base_lt_one -> Real.logb_pos_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LT.lt.{0} Real Real.hasLt x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LT.lt.{0} Real Real.instLTReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x))
+Case conversion may be inaccurate. Consider using '#align real.logb_pos_of_base_lt_one Real.logb_pos_of_base_lt_oneₓ'. -/
 theorem logb_pos_of_base_lt_one (hx : 0 < x) (hx' : x < 1) : 0 < logb b x :=
   by
   rw [logb_pos_iff_of_base_lt_one b_pos b_lt_one hx]
   exact hx'
 #align real.logb_pos_of_base_lt_one Real.logb_pos_of_base_lt_one
 
+/- warning: real.logb_neg_iff_of_base_lt_one -> Real.logb_neg_iff_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LT.lt.{0} Real Real.hasLt (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) x))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x))
+Case conversion may be inaccurate. Consider using '#align real.logb_neg_iff_of_base_lt_one Real.logb_neg_iff_of_base_lt_oneₓ'. -/
 theorem logb_neg_iff_of_base_lt_one (h : 0 < x) : logb b x < 0 ↔ 1 < x := by
   rw [← @logb_one b, logb_lt_logb_iff_of_base_lt_one b_pos b_lt_one h zero_lt_one]
 #align real.logb_neg_iff_of_base_lt_one Real.logb_neg_iff_of_base_lt_one
 
+/- warning: real.logb_neg_of_base_lt_one -> Real.logb_neg_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) x) -> (LT.lt.{0} Real Real.hasLt (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x) -> (LT.lt.{0} Real Real.instLTReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)))
+Case conversion may be inaccurate. Consider using '#align real.logb_neg_of_base_lt_one Real.logb_neg_of_base_lt_oneₓ'. -/
 theorem logb_neg_of_base_lt_one (h1 : 1 < x) : logb b x < 0 :=
   (logb_neg_iff_of_base_lt_one b_pos b_lt_one (lt_trans zero_lt_one h1)).2 h1
 #align real.logb_neg_of_base_lt_one Real.logb_neg_of_base_lt_one
 
+/- warning: real.logb_nonneg_iff_of_base_lt_one -> Real.logb_nonneg_iff_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x)) (LE.le.{0} Real Real.hasLe x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x)) (LE.le.{0} Real Real.instLEReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))))
+Case conversion may be inaccurate. Consider using '#align real.logb_nonneg_iff_of_base_lt_one Real.logb_nonneg_iff_of_base_lt_oneₓ'. -/
 theorem logb_nonneg_iff_of_base_lt_one (hx : 0 < x) : 0 ≤ logb b x ↔ x ≤ 1 := by
   rw [← not_lt, logb_neg_iff_of_base_lt_one b_pos b_lt_one hx, not_lt]
 #align real.logb_nonneg_iff_of_base_lt_one Real.logb_nonneg_iff_of_base_lt_one
 
+/- warning: real.logb_nonneg_of_base_lt_one -> Real.logb_nonneg_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (LE.le.{0} Real Real.hasLe x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) (Real.logb b x))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (LE.le.{0} Real Real.instLEReal x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) (Real.logb b x))
+Case conversion may be inaccurate. Consider using '#align real.logb_nonneg_of_base_lt_one Real.logb_nonneg_of_base_lt_oneₓ'. -/
 theorem logb_nonneg_of_base_lt_one (hx : 0 < x) (hx' : x ≤ 1) : 0 ≤ logb b x :=
   by
   rw [logb_nonneg_iff_of_base_lt_one b_pos b_lt_one hx]
   exact hx'
 #align real.logb_nonneg_of_base_lt_one Real.logb_nonneg_of_base_lt_one
 
+/- warning: real.logb_nonpos_iff_of_base_lt_one -> Real.logb_nonpos_iff_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Iff (LE.le.{0} Real Real.hasLe (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))) x))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Iff (LE.le.{0} Real Real.instLEReal (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)) x))
+Case conversion may be inaccurate. Consider using '#align real.logb_nonpos_iff_of_base_lt_one Real.logb_nonpos_iff_of_base_lt_oneₓ'. -/
 theorem logb_nonpos_iff_of_base_lt_one (hx : 0 < x) : logb b x ≤ 0 ↔ 1 ≤ x := by
   rw [← not_lt, logb_pos_iff_of_base_lt_one b_pos b_lt_one hx, not_lt]
 #align real.logb_nonpos_iff_of_base_lt_one Real.logb_nonpos_iff_of_base_lt_one
 
+/- warning: real.strict_anti_on_logb_of_base_lt_one -> Real.strictAntiOn_logb_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (StrictAntiOn.{0, 0} Real Real Real.preorder Real.preorder (Real.logb b) (Set.Ioi.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))))
+but is expected to have type
+  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (StrictAntiOn.{0, 0} Real Real Real.instPreorderReal Real.instPreorderReal (Real.logb b) (Set.Ioi.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))))
+Case conversion may be inaccurate. Consider using '#align real.strict_anti_on_logb_of_base_lt_one Real.strictAntiOn_logb_of_base_lt_oneₓ'. -/
 theorem strictAntiOn_logb_of_base_lt_one : StrictAntiOn (logb b) (Set.Ioi 0) := fun x hx y hy hxy =>
   logb_lt_logb_of_base_lt_one b_pos b_lt_one hx hxy
 #align real.strict_anti_on_logb_of_base_lt_one Real.strictAntiOn_logb_of_base_lt_one
 
+/- warning: real.strict_mono_on_logb_of_base_lt_one -> Real.strictMonoOn_logb_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (StrictMonoOn.{0, 0} Real Real Real.preorder Real.preorder (Real.logb b) (Set.Iio.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))))
+but is expected to have type
+  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (StrictMonoOn.{0, 0} Real Real Real.instPreorderReal Real.instPreorderReal (Real.logb b) (Set.Iio.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))))
+Case conversion may be inaccurate. Consider using '#align real.strict_mono_on_logb_of_base_lt_one Real.strictMonoOn_logb_of_base_lt_oneₓ'. -/
 theorem strictMonoOn_logb_of_base_lt_one : StrictMonoOn (logb b) (Set.Iio 0) :=
   by
   rintro x (hx : x < 0) y (hy : y < 0) hxy
@@ -358,19 +684,43 @@ theorem strictMonoOn_logb_of_base_lt_one : StrictMonoOn (logb b) (Set.Iio 0) :=
   rwa [abs_of_neg hy, abs_of_neg hx, neg_lt_neg_iff]
 #align real.strict_mono_on_logb_of_base_lt_one Real.strictMonoOn_logb_of_base_lt_one
 
+/- warning: real.logb_inj_on_pos_of_base_lt_one -> Real.logb_injOn_pos_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Set.InjOn.{0, 0} Real Real (Real.logb b) (Set.Ioi.{0} Real Real.preorder (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))))
+but is expected to have type
+  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Set.InjOn.{0, 0} Real Real (Real.logb b) (Set.Ioi.{0} Real Real.instPreorderReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))))
+Case conversion may be inaccurate. Consider using '#align real.logb_inj_on_pos_of_base_lt_one Real.logb_injOn_pos_of_base_lt_oneₓ'. -/
 theorem logb_injOn_pos_of_base_lt_one : Set.InjOn (logb b) (Set.Ioi 0) :=
   (strictAntiOn_logb_of_base_lt_one b_pos b_lt_one).InjOn
 #align real.logb_inj_on_pos_of_base_lt_one Real.logb_injOn_pos_of_base_lt_one
 
+/- warning: real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_one -> Real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Eq.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) -> (Eq.{1} Real x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Eq.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) -> (Eq.{1} Real x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)))
+Case conversion may be inaccurate. Consider using '#align real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_one Real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_oneₓ'. -/
 theorem eq_one_of_pos_of_logb_eq_zero_of_base_lt_one (h₁ : 0 < x) (h₂ : logb b x = 0) : x = 1 :=
   logb_injOn_pos_of_base_lt_one b_pos b_lt_one (Set.mem_Ioi.2 h₁) (Set.mem_Ioi.2 zero_lt_one)
     (h₂.trans Real.logb_one.symm)
 #align real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_one Real.eq_one_of_pos_of_logb_eq_zero_of_base_lt_one
 
+/- warning: real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_one -> Real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) x) -> (Ne.{1} Real x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Ne.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))))
+but is expected to have type
+  forall {b : Real} {x : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) x) -> (Ne.{1} Real x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Ne.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)))
+Case conversion may be inaccurate. Consider using '#align real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_one Real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_oneₓ'. -/
 theorem logb_ne_zero_of_pos_of_ne_one_of_base_lt_one (hx_pos : 0 < x) (hx : x ≠ 1) : logb b x ≠ 0 :=
   mt (eq_one_of_pos_of_logb_eq_zero_of_base_lt_one b_pos b_lt_one hx_pos) hx
 #align real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_one Real.logb_ne_zero_of_pos_of_ne_one_of_base_lt_one
 
+/- warning: real.tendsto_logb_at_top_of_base_lt_one -> Real.tendsto_logb_atTop_of_base_lt_one is a dubious translation:
+lean 3 declaration is
+  forall {b : Real}, (LT.lt.{0} Real Real.hasLt (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) b) -> (LT.lt.{0} Real Real.hasLt b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) -> (Filter.Tendsto.{0, 0} Real Real (Real.logb b) (Filter.atTop.{0} Real Real.preorder) (Filter.atBot.{0} Real Real.preorder))
+but is expected to have type
+  forall {b : Real}, (LT.lt.{0} Real Real.instLTReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) b) -> (LT.lt.{0} Real Real.instLTReal b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) -> (Filter.Tendsto.{0, 0} Real Real (Real.logb b) (Filter.atTop.{0} Real Real.instPreorderReal) (Filter.atBot.{0} Real Real.instPreorderReal))
+Case conversion may be inaccurate. Consider using '#align real.tendsto_logb_at_top_of_base_lt_one Real.tendsto_logb_atTop_of_base_lt_oneₓ'. -/
 theorem tendsto_logb_atTop_of_base_lt_one : Tendsto (logb b) atTop atBot :=
   by
   rw [tendsto_at_top_at_bot]
@@ -386,6 +736,12 @@ theorem tendsto_logb_atTop_of_base_lt_one : Tendsto (logb b) atTop atBot :=
 
 end BPosAndBLtOne
 
+/- warning: real.floor_logb_nat_cast -> Real.floor_logb_nat_cast is a dubious translation:
+lean 3 declaration is
+  forall {b : Nat} {r : Real}, (LT.lt.{0} Nat Nat.hasLt (OfNat.ofNat.{0} Nat 1 (OfNat.mk.{0} Nat 1 (One.one.{0} Nat Nat.hasOne))) b) -> (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) r) -> (Eq.{1} Int (Int.floor.{0} Real Real.linearOrderedRing Real.floorRing (Real.logb ((fun (a : Type) (b : Type) [self : HasLiftT.{1, 1} a b] => self.0) Nat Real (HasLiftT.mk.{1, 1} Nat Real (CoeTCₓ.coe.{1, 1} Nat Real (Nat.castCoe.{0} Real Real.hasNatCast))) b) r)) (Int.log.{0} Real (LinearOrderedField.toLinearOrderedSemifield.{0} Real Real.linearOrderedField) (FloorRing.toFloorSemiring.{0} Real Real.linearOrderedRing Real.floorRing) b r))
+but is expected to have type
+  forall {b : Nat} {r : Real}, (LT.lt.{0} Nat instLTNat (OfNat.ofNat.{0} Nat 1 (instOfNatNat 1)) b) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) r) -> (Eq.{1} Int (Int.floor.{0} Real Real.instLinearOrderedRingReal Real.instFloorRingRealInstLinearOrderedRingReal (Real.logb (Nat.cast.{0} Real Real.natCast b) r)) (Int.log.{0} Real (LinearOrderedField.toLinearOrderedSemifield.{0} Real Real.instLinearOrderedFieldReal) (FloorRing.toFloorSemiring.{0} Real (LinearOrderedCommRing.toLinearOrderedRing.{0} Real (LinearOrderedField.toLinearOrderedCommRing.{0} Real Real.instLinearOrderedFieldReal)) Real.instFloorRingRealInstLinearOrderedRingReal) b r))
+Case conversion may be inaccurate. Consider using '#align real.floor_logb_nat_cast Real.floor_logb_nat_castₓ'. -/
 theorem floor_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : ⌊logb b r⌋ = Int.log b r :=
   by
   obtain rfl | hr := hr.eq_or_lt
@@ -399,6 +755,12 @@ theorem floor_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : 
     exact Int.zpow_log_le_self hb hr
 #align real.floor_logb_nat_cast Real.floor_logb_nat_cast
 
+/- warning: real.ceil_logb_nat_cast -> Real.ceil_logb_nat_cast is a dubious translation:
+lean 3 declaration is
+  forall {b : Nat} {r : Real}, (LT.lt.{0} Nat Nat.hasLt (OfNat.ofNat.{0} Nat 1 (OfNat.mk.{0} Nat 1 (One.one.{0} Nat Nat.hasOne))) b) -> (LE.le.{0} Real Real.hasLe (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))) r) -> (Eq.{1} Int (Int.ceil.{0} Real Real.linearOrderedRing Real.floorRing (Real.logb ((fun (a : Type) (b : Type) [self : HasLiftT.{1, 1} a b] => self.0) Nat Real (HasLiftT.mk.{1, 1} Nat Real (CoeTCₓ.coe.{1, 1} Nat Real (Nat.castCoe.{0} Real Real.hasNatCast))) b) r)) (Int.clog.{0} Real (LinearOrderedField.toLinearOrderedSemifield.{0} Real Real.linearOrderedField) (FloorRing.toFloorSemiring.{0} Real Real.linearOrderedRing Real.floorRing) b r))
+but is expected to have type
+  forall {b : Nat} {r : Real}, (LT.lt.{0} Nat instLTNat (OfNat.ofNat.{0} Nat 1 (instOfNatNat 1)) b) -> (LE.le.{0} Real Real.instLEReal (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)) r) -> (Eq.{1} Int (Int.ceil.{0} Real Real.instLinearOrderedRingReal Real.instFloorRingRealInstLinearOrderedRingReal (Real.logb (Nat.cast.{0} Real Real.natCast b) r)) (Int.clog.{0} Real (LinearOrderedField.toLinearOrderedSemifield.{0} Real Real.instLinearOrderedFieldReal) (FloorRing.toFloorSemiring.{0} Real (LinearOrderedCommRing.toLinearOrderedRing.{0} Real (LinearOrderedField.toLinearOrderedCommRing.{0} Real Real.instLinearOrderedFieldReal)) Real.instFloorRingRealInstLinearOrderedRingReal) b r))
+Case conversion may be inaccurate. Consider using '#align real.ceil_logb_nat_cast Real.ceil_logb_nat_castₓ'. -/
 theorem ceil_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : ⌈logb b r⌉ = Int.clog b r :=
   by
   obtain rfl | hr := hr.eq_or_lt
@@ -412,6 +774,12 @@ theorem ceil_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) : ⌈
     exact rpow_le_rpow_of_exponent_le hb1'.le (Int.le_ceil _)
 #align real.ceil_logb_nat_cast Real.ceil_logb_nat_cast
 
+/- warning: real.logb_eq_zero -> Real.logb_eq_zero is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {x : Real}, Iff (Eq.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (Or (Eq.{1} Real b (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (Or (Eq.{1} Real b (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) (Or (Eq.{1} Real b (Neg.neg.{0} Real Real.hasNeg (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))))) (Or (Eq.{1} Real x (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero)))) (Or (Eq.{1} Real x (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne)))) (Eq.{1} Real x (Neg.neg.{0} Real Real.hasNeg (OfNat.ofNat.{0} Real 1 (OfNat.mk.{0} Real 1 (One.one.{0} Real Real.hasOne))))))))))
+but is expected to have type
+  forall {b : Real} {x : Real}, Iff (Eq.{1} Real (Real.logb b x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (Or (Eq.{1} Real b (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (Or (Eq.{1} Real b (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) (Or (Eq.{1} Real b (Neg.neg.{0} Real Real.instNegReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)))) (Or (Eq.{1} Real x (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal))) (Or (Eq.{1} Real x (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal))) (Eq.{1} Real x (Neg.neg.{0} Real Real.instNegReal (OfNat.ofNat.{0} Real 1 (One.toOfNat1.{0} Real Real.instOneReal)))))))))
+Case conversion may be inaccurate. Consider using '#align real.logb_eq_zero Real.logb_eq_zeroₓ'. -/
 @[simp]
 theorem logb_eq_zero : logb b x = 0 ↔ b = 0 ∨ b = 1 ∨ b = -1 ∨ x = 0 ∨ x = 1 ∨ x = -1 :=
   by
@@ -422,6 +790,12 @@ theorem logb_eq_zero : logb b x = 0 ↔ b = 0 ∨ b = 1 ∨ b = -1 ∨ x = 0 ∨
 -- TODO add other limits and continuous API lemmas analogous to those in log.lean
 open BigOperators
 
+/- warning: real.logb_prod -> Real.logb_prod is a dubious translation:
+lean 3 declaration is
+  forall {b : Real} {α : Type.{u1}} (s : Finset.{u1} α) (f : α -> Real), (forall (x : α), (Membership.Mem.{u1, u1} α (Finset.{u1} α) (Finset.hasMem.{u1} α) x s) -> (Ne.{1} Real (f x) (OfNat.ofNat.{0} Real 0 (OfNat.mk.{0} Real 0 (Zero.zero.{0} Real Real.hasZero))))) -> (Eq.{1} Real (Real.logb b (Finset.prod.{0, u1} Real α Real.commMonoid s (fun (i : α) => f i))) (Finset.sum.{0, u1} Real α Real.addCommMonoid s (fun (i : α) => Real.logb b (f i))))
+but is expected to have type
+  forall {b : Real} {α : Type.{u1}} (s : Finset.{u1} α) (f : α -> Real), (forall (x : α), (Membership.mem.{u1, u1} α (Finset.{u1} α) (Finset.instMembershipFinset.{u1} α) x s) -> (Ne.{1} Real (f x) (OfNat.ofNat.{0} Real 0 (Zero.toOfNat0.{0} Real Real.instZeroReal)))) -> (Eq.{1} Real (Real.logb b (Finset.prod.{0, u1} Real α Real.instCommMonoidReal s (fun (i : α) => f i))) (Finset.sum.{0, u1} Real α Real.instAddCommMonoidReal s (fun (i : α) => Real.logb b (f i))))
+Case conversion may be inaccurate. Consider using '#align real.logb_prod Real.logb_prodₓ'. -/
 theorem logb_prod {α : Type _} (s : Finset α) (f : α → ℝ) (hf : ∀ x ∈ s, f x ≠ 0) :
     logb b (∏ i in s, f i) = ∑ i in s, logb b (f i) := by
   classical
Diff
@@ -4,11 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE.
 Authors: Bolton Bailey, Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne
 
 ! This file was ported from Lean 3 source module analysis.special_functions.log.base
-! leanprover-community/mathlib commit 2c1d8ca2812b64f88992a5294ea3dba144755cd1
+! leanprover-community/mathlib commit 0b9eaaa7686280fad8cce467f5c3c57ee6ce77f8
 ! Please do not edit these lines, except to modify the commit id
 ! if you have ported upstream changes.
 -/
-import Mathbin.Analysis.SpecialFunctions.Pow
+import Mathbin.Analysis.SpecialFunctions.Pow.Real
 import Mathbin.Data.Int.Log
 
 /-!
Diff
@@ -4,11 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE.
 Authors: Bolton Bailey, Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne
 
 ! This file was ported from Lean 3 source module analysis.special_functions.log.base
-! leanprover-community/mathlib commit f2ce6086713c78a7f880485f7917ea547a215982
+! leanprover-community/mathlib commit 2c1d8ca2812b64f88992a5294ea3dba144755cd1
 ! Please do not edit these lines, except to modify the commit id
 ! if you have ported upstream changes.
 -/
-import Mathbin.Analysis.SpecialFunctions.Log.Basic
 import Mathbin.Analysis.SpecialFunctions.Pow
 import Mathbin.Data.Int.Log
 

Changes in mathlib4

mathlib3
mathlib4
chore: Rename nat_cast/int_cast/rat_cast to natCast/intCast/ratCast (#11486)

Now that I am defining NNRat.cast, I want a definitive answer to this naming issue. Plenty of lemmas in mathlib already use natCast/intCast/ratCast over nat_cast/int_cast/rat_cast, and this matches with the general expectation that underscore-separated name parts correspond to a single declaration.

Diff
@@ -117,7 +117,7 @@ theorem logb_rpow_eq_mul_logb_of_pos (hx : 0 < x) : logb b (x ^ y) = y * logb b
   rw [logb, log_rpow hx, logb, mul_div_assoc]
 
 theorem logb_pow {k : ℕ} (hx : 0 < x) : logb b (x ^ k) = k * logb b x := by
-  rw [← rpow_nat_cast, logb_rpow_eq_mul_logb_of_pos hx]
+  rw [← rpow_natCast, logb_rpow_eq_mul_logb_of_pos hx]
 
 section BPosAndNeOne
 
@@ -405,31 +405,31 @@ theorem tendsto_logb_atTop_of_base_lt_one : Tendsto (logb b) atTop atBot := by
 
 end BPosAndBLtOne
 
-theorem floor_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) :
+theorem floor_logb_natCast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) :
     ⌊logb b r⌋ = Int.log b r := by
   obtain rfl | hr := hr.eq_or_lt
   · rw [logb_zero, Int.log_zero_right, Int.floor_zero]
   have hb1' : 1 < (b : ℝ) := Nat.one_lt_cast.mpr hb
   apply le_antisymm
-  · rw [← Int.zpow_le_iff_le_log hb hr, ← rpow_int_cast b]
+  · rw [← Int.zpow_le_iff_le_log hb hr, ← rpow_intCast b]
     refine' le_of_le_of_eq _ (rpow_logb (zero_lt_one.trans hb1') hb1'.ne' hr)
     exact rpow_le_rpow_of_exponent_le hb1'.le (Int.floor_le _)
-  · rw [Int.le_floor, le_logb_iff_rpow_le hb1' hr, rpow_int_cast]
+  · rw [Int.le_floor, le_logb_iff_rpow_le hb1' hr, rpow_intCast]
     exact Int.zpow_log_le_self hb hr
-#align real.floor_logb_nat_cast Real.floor_logb_nat_cast
+#align real.floor_logb_nat_cast Real.floor_logb_natCast
 
-theorem ceil_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) :
+theorem ceil_logb_natCast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) :
     ⌈logb b r⌉ = Int.clog b r := by
   obtain rfl | hr := hr.eq_or_lt
   · rw [logb_zero, Int.clog_zero_right, Int.ceil_zero]
   have hb1' : 1 < (b : ℝ) := Nat.one_lt_cast.mpr hb
   apply le_antisymm
-  · rw [Int.ceil_le, logb_le_iff_le_rpow hb1' hr, rpow_int_cast]
+  · rw [Int.ceil_le, logb_le_iff_le_rpow hb1' hr, rpow_intCast]
     exact Int.self_le_zpow_clog hb r
-  · rw [← Int.le_zpow_iff_clog_le hb hr, ← rpow_int_cast b]
+  · rw [← Int.le_zpow_iff_clog_le hb hr, ← rpow_intCast b]
     refine' (rpow_logb (zero_lt_one.trans hb1') hb1'.ne' hr).symm.trans_le _
     exact rpow_le_rpow_of_exponent_le hb1'.le (Int.le_ceil _)
-#align real.ceil_logb_nat_cast Real.ceil_logb_nat_cast
+#align real.ceil_logb_nat_cast Real.ceil_logb_natCast
 
 @[simp]
 theorem logb_eq_zero : logb b x = 0 ↔ b = 0 ∨ b = 1 ∨ b = -1 ∨ x = 0 ∨ x = 1 ∨ x = -1 := by
@@ -465,7 +465,7 @@ lemma Real.induction_Ico_mul {P : ℝ → Prop} (x₀ r : ℝ) (hr : 1 < r) (hx
     intro x hx
     have hx' : 0 < x / x₀ := div_pos (hx₀.trans_le hx) hx₀
     refine this ⌊logb r (x / x₀)⌋₊ x ?_
-    rw [mem_Ico, ← div_lt_iff hx₀, ← rpow_nat_cast, ← logb_lt_iff_lt_rpow hr hx', Nat.cast_add,
+    rw [mem_Ico, ← div_lt_iff hx₀, ← rpow_natCast, ← logb_lt_iff_lt_rpow hr hx', Nat.cast_add,
       Nat.cast_one]
     exact ⟨hx, Nat.lt_floor_add_one _⟩
   intro n
Feat(Log/Base): Adding 3 theorems related to Real.logb (#12110)
Diff
@@ -113,6 +113,12 @@ theorem div_logb {a b c : ℝ} (h₁ : c ≠ 0) (h₂ : c ≠ 1) (h₃ : c ≠ -
   div_div_div_cancel_left' _ _ <| log_ne_zero.mpr ⟨h₁, h₂, h₃⟩
 #align real.div_logb Real.div_logb
 
+theorem logb_rpow_eq_mul_logb_of_pos (hx : 0 < x) : logb b (x ^ y) = y * logb b x := by
+  rw [logb, log_rpow hx, logb, mul_div_assoc]
+
+theorem logb_pow {k : ℕ} (hx : 0 < x) : logb b (x ^ k) = k * logb b x := by
+  rw [← rpow_nat_cast, logb_rpow_eq_mul_logb_of_pos hx]
+
 section BPosAndNeOne
 
 variable (b_pos : 0 < b) (b_ne_one : b ≠ 1)
@@ -148,6 +154,11 @@ theorem rpow_logb_of_neg (hx : x < 0) : b ^ logb b x = -x := by
   exact abs_of_neg hx
 #align real.rpow_logb_of_neg Real.rpow_logb_of_neg
 
+theorem logb_eq_iff_rpow_eq (hy : 0 < y) : logb b y = x ↔ b ^ x = y := by
+  constructor <;> rintro rfl
+  · exact rpow_logb b_pos b_ne_one hy
+  · exact logb_rpow b_pos b_ne_one
+
 theorem surjOn_logb : SurjOn (logb b) (Ioi 0) univ := fun x _ =>
   ⟨rpow b x, rpow_pos_of_pos b_pos x, logb_rpow b_pos b_ne_one⟩
 #align real.surj_on_logb Real.surjOn_logb
style: replace '.-/' by '. -/' (#11938)

Purely automatic replacement. If this is in any way controversial; I'm happy to just close this PR.

Diff
@@ -35,7 +35,7 @@ namespace Real
 variable {b x y : ℝ}
 
 /-- The real logarithm in a given base. As with the natural logarithm, we define `logb b x` to
-be `logb b |x|` for `x < 0`, and `0` for `x = 0`.-/
+be `logb b |x|` for `x < 0`, and `0` for `x = 0`. -/
 -- @[pp_nodot] -- Porting note: removed
 noncomputable def logb (b x : ℝ) : ℝ :=
   log x / log b
chore: avoid Ne.def (adaptation for nightly-2024-03-27) (#11801)
Diff
@@ -132,7 +132,7 @@ theorem rpow_logb_eq_abs (hx : x ≠ 0) : b ^ logb b x = |x| := by
   apply log_injOn_pos
   simp only [Set.mem_Ioi]
   apply rpow_pos_of_pos b_pos
-  simp only [abs_pos, mem_Ioi, Ne.def, hx, not_false_iff]
+  simp only [abs_pos, mem_Ioi, Ne, hx, not_false_iff]
   rw [log_rpow b_pos, logb, log_abs]
   field_simp [log_b_ne_zero b_pos b_ne_one]
 #align real.rpow_logb_eq_abs Real.rpow_logb_eq_abs
chore: remove terminal, terminal refines (#10762)

I replaced a few "terminal" refine/refine's with exact.

The strategy was very simple-minded: essentially any refine whose following line had smaller indentation got replaced by exact and then I cleaned up the mess.

This PR certainly leaves some further terminal refines, but maybe the current change is beneficial.

Diff
@@ -414,7 +414,7 @@ theorem ceil_logb_nat_cast {b : ℕ} {r : ℝ} (hb : 1 < b) (hr : 0 ≤ r) :
   have hb1' : 1 < (b : ℝ) := Nat.one_lt_cast.mpr hb
   apply le_antisymm
   · rw [Int.ceil_le, logb_le_iff_le_rpow hb1' hr, rpow_int_cast]
-    refine' Int.self_le_zpow_clog hb r
+    exact Int.self_le_zpow_clog hb r
   · rw [← Int.le_zpow_iff_clog_le hb hr, ← rpow_int_cast b]
     refine' (rpow_logb (zero_lt_one.trans hb1') hb1'.ne' hr).symm.trans_le _
     exact rpow_le_rpow_of_exponent_le hb1'.le (Int.le_ceil _)
chore: remove stream-of-consciousness uses of have, replace and suffices (#10640)

No changes to tactic file, it's just boring fixes throughout the library.

This follows on from #6964.

Co-authored-by: sgouezel <sebastien.gouezel@univ-rennes1.fr> Co-authored-by: Eric Wieser <wieser.eric@gmail.com>

Diff
@@ -118,8 +118,8 @@ section BPosAndNeOne
 variable (b_pos : 0 < b) (b_ne_one : b ≠ 1)
 
 private theorem log_b_ne_zero : log b ≠ 0 := by
-  have b_ne_zero : b ≠ 0; linarith
-  have b_ne_minus_one : b ≠ -1; linarith
+  have b_ne_zero : b ≠ 0 := by linarith
+  have b_ne_minus_one : b ≠ -1 := by linarith
   simp [b_ne_one, b_ne_zero, b_ne_minus_one]
 
 @[simp]
feat: More rpow lemmas (#9108)

A bunch of easy lemmas about Real.pow and the golf of existing lemmas with them.

Also rename log_le_log to log_le_log_iff and log_le_log' to log_le_log. Those misnames caused several proofs to bother with side conditions they didn't need.

From LeanAPAP

Diff
@@ -182,7 +182,7 @@ private theorem b_ne_one' : b ≠ 1 := by linarith
 
 @[simp]
 theorem logb_le_logb (h : 0 < x) (h₁ : 0 < y) : logb b x ≤ logb b y ↔ x ≤ y := by
-  rw [logb, logb, div_le_div_right (log_pos hb), log_le_log h h₁]
+  rw [logb, logb, div_le_div_right (log_pos hb), log_le_log_iff h h₁]
 #align real.logb_le_logb Real.logb_le_logb
 
 @[gcongr]
@@ -295,7 +295,7 @@ private theorem b_ne_one : b ≠ 1 := by linarith
 
 @[simp]
 theorem logb_le_logb_of_base_lt_one (h : 0 < x) (h₁ : 0 < y) : logb b x ≤ logb b y ↔ y ≤ x := by
-  rw [logb, logb, div_le_div_right_of_neg (log_neg b_pos b_lt_one), log_le_log h₁ h]
+  rw [logb, logb, div_le_div_right_of_neg (log_neg b_pos b_lt_one), log_le_log_iff h₁ h]
 #align real.logb_le_logb_of_base_lt_one Real.logb_le_logb_of_base_lt_one
 
 theorem logb_lt_logb_of_base_lt_one (hx : 0 < x) (hxy : x < y) : logb b y < logb b x := by
chore: tidy various files (#8818)
Diff
@@ -55,7 +55,7 @@ theorem logb_one : logb b 1 = 0 := by simp [logb]
 
 @[simp]
 lemma logb_self_eq_one (hb : 1 < b) : logb b b = 1 :=
-   div_self (log_pos hb).ne'
+  div_self (log_pos hb).ne'
 
 lemma logb_self_eq_one_iff : logb b b = 1 ↔ b ≠ 0 ∧ b ≠ 1 ∧ b ≠ -1 :=
   Iff.trans ⟨fun h h' => by simp [logb, h'] at h, div_self⟩ log_ne_zero
@@ -450,18 +450,17 @@ lemma Real.induction_Ico_mul {P : ℝ → Prop} (x₀ r : ℝ) (hr : 1 < r) (hx
     (step : ∀ n : ℕ, n ≥ 1 → (∀ z ∈ Set.Ico x₀ (r ^ n * x₀), P z) →
       (∀ z ∈ Set.Ico (r ^ n * x₀) (r ^ (n+1) * x₀), P z)) :
     ∀ x ≥ x₀, P x := by
-  suffices : ∀ n : ℕ, ∀ x ∈ Set.Ico x₀ (r ^ (n + 1) * x₀), P x
-  · intro x hx
+  suffices ∀ n : ℕ, ∀ x ∈ Set.Ico x₀ (r ^ (n + 1) * x₀), P x by
+    intro x hx
     have hx' : 0 < x / x₀ := div_pos (hx₀.trans_le hx) hx₀
     refine this ⌊logb r (x / x₀)⌋₊ x ?_
     rw [mem_Ico, ← div_lt_iff hx₀, ← rpow_nat_cast, ← logb_lt_iff_lt_rpow hr hx', Nat.cast_add,
       Nat.cast_one]
     exact ⟨hx, Nat.lt_floor_add_one _⟩
   intro n
-  induction n
-  case zero => simpa using base
-  case succ n ih =>
-    specialize step (n + 1) (by simp)
-    exact fun x hx => (Ico_subset_Ico_union_Ico hx).elim (ih x) (step ih _)
+  induction n with
+  | zero => simpa using base
+  | succ n ih =>
+    exact fun x hx => (Ico_subset_Ico_union_Ico hx).elim (ih x) (step (n + 1) (by simp) ih _)
 
 end Induction
chore: space after (#8178)

Co-authored-by: Moritz Firsching <firsching@google.com>

Diff
@@ -454,7 +454,7 @@ lemma Real.induction_Ico_mul {P : ℝ → Prop} (x₀ r : ℝ) (hr : 1 < r) (hx
   · intro x hx
     have hx' : 0 < x / x₀ := div_pos (hx₀.trans_le hx) hx₀
     refine this ⌊logb r (x / x₀)⌋₊ x ?_
-    rw [mem_Ico, ←div_lt_iff hx₀, ← rpow_nat_cast, ←logb_lt_iff_lt_rpow hr hx', Nat.cast_add,
+    rw [mem_Ico, ← div_lt_iff hx₀, ← rpow_nat_cast, ← logb_lt_iff_lt_rpow hr hx', Nat.cast_add,
       Nat.cast_one]
     exact ⟨hx, Nat.lt_floor_add_one _⟩
   intro n
chore: bump to v4.3.0-rc2 (#8366)

PR contents

This is the supremum of

along with some minor fixes from failures on nightly-testing as Mathlib master is merged into it.

Note that some PRs for changes that are already compatible with the current toolchain and will be necessary have already been split out: #8380.

I am hopeful that in future we will be able to progressively merge adaptation PRs into a bump/v4.X.0 branch, so we never end up with a "big merge" like this. However one of these adaptation PRs (#8056) predates my new scheme for combined CI, and it wasn't possible to keep that PR viable in the meantime.

Lean PRs involved in this bump

In particular this includes adjustments for the Lean PRs

leanprover/lean4#2778

We can get rid of all the

local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue [lean4#2220](https://github.com/leanprover/lean4/pull/2220)

macros across Mathlib (and in any projects that want to write natural number powers of reals).

leanprover/lean4#2722

Changes the default behaviour of simp to (config := {decide := false}). This makes simp (and consequentially norm_num) less powerful, but also more consistent, and less likely to blow up in long failures. This requires a variety of changes: changing some previously by simp or norm_num to decide or rfl, or adding (config := {decide := true}).

leanprover/lean4#2783

This changed the behaviour of simp so that simp [f] will only unfold "fully applied" occurrences of f. The old behaviour can be recovered with simp (config := { unfoldPartialApp := true }). We may in future add a syntax for this, e.g. simp [!f]; please provide feedback! In the meantime, we have made the following changes:

  • switching to using explicit lemmas that have the intended level of application
  • (config := { unfoldPartialApp := true }) in some places, to recover the old behaviour
  • Using @[eqns] to manually adjust the equation lemmas for a particular definition, recovering the old behaviour just for that definition. See #8371, where we do this for Function.comp and Function.flip.

This change in Lean may require further changes down the line (e.g. adding the !f syntax, and/or upstreaming the special treatment for Function.comp and Function.flip, and/or removing this special treatment). Please keep an open and skeptical mind about these changes!

Co-authored-by: leanprover-community-mathlib4-bot <leanprover-community-mathlib4-bot@users.noreply.github.com> Co-authored-by: Scott Morrison <scott.morrison@gmail.com> Co-authored-by: Eric Wieser <wieser.eric@gmail.com> Co-authored-by: Mauricio Collares <mauricio@collares.org>

Diff
@@ -454,14 +454,14 @@ lemma Real.induction_Ico_mul {P : ℝ → Prop} (x₀ r : ℝ) (hr : 1 < r) (hx
   · intro x hx
     have hx' : 0 < x / x₀ := div_pos (hx₀.trans_le hx) hx₀
     refine this ⌊logb r (x / x₀)⌋₊ x ?_
-    rw [mem_Ico, ←div_lt_iff hx₀, ←logb_lt_iff_lt_rpow hr hx']
+    rw [mem_Ico, ←div_lt_iff hx₀, ← rpow_nat_cast, ←logb_lt_iff_lt_rpow hr hx', Nat.cast_add,
+      Nat.cast_one]
     exact ⟨hx, Nat.lt_floor_add_one _⟩
   intro n
   induction n
   case zero => simpa using base
   case succ n ih =>
     specialize step (n + 1) (by simp)
-    simp only [Nat.cast_add_one] at step ⊢
     exact fun x hx => (Ico_subset_Ico_union_Ico hx).elim (ih x) (step ih _)
 
 end Induction
feat: add induction principle for intervals (#8350)

This PR adds an induction principle that says that if a proposition P is true on [x₀, r * x₀) and if P for [x₀, r^n * x₀) implies P for [r^n * x₀, r^(n+1) * x₀), then P is true for all x ≥ x₀.

It also adds a few missing lemmas about logb.

Co-authored-by: Frédéric Dupuis <31101893+dupuisf@users.noreply.github.com>

Diff
@@ -53,6 +53,13 @@ theorem logb_zero : logb b 0 = 0 := by simp [logb]
 theorem logb_one : logb b 1 = 0 := by simp [logb]
 #align real.logb_one Real.logb_one
 
+@[simp]
+lemma logb_self_eq_one (hb : 1 < b) : logb b b = 1 :=
+   div_self (log_pos hb).ne'
+
+lemma logb_self_eq_one_iff : logb b b = 1 ↔ b ≠ 0 ∧ b ≠ 1 ∧ b ≠ -1 :=
+  Iff.trans ⟨fun h h' => by simp [logb, h'] at h, div_self⟩ log_ne_zero
+
 @[simp]
 theorem logb_abs (x : ℝ) : logb b |x| = logb b x := by rw [logb, logb, log_abs]
 #align real.logb_abs Real.logb_abs
@@ -178,6 +185,11 @@ theorem logb_le_logb (h : 0 < x) (h₁ : 0 < y) : logb b x ≤ logb b y ↔ x 
   rw [logb, logb, div_le_div_right (log_pos hb), log_le_log h h₁]
 #align real.logb_le_logb Real.logb_le_logb
 
+@[gcongr]
+theorem logb_le_logb_of_le (h : 0 < x) (hxy : x ≤ y) : logb b x ≤ logb b y :=
+  (logb_le_logb hb h (by linarith)).mpr hxy
+
+@[gcongr]
 theorem logb_lt_logb (hx : 0 < x) (hxy : x < y) : logb b x < logb b y := by
   rw [logb, logb, div_lt_div_right (log_pos hb)]
   exact log_lt_log hx hxy
@@ -427,3 +439,29 @@ theorem logb_prod {α : Type*} (s : Finset α) (f : α → ℝ) (hf : ∀ x ∈
 #align real.logb_prod Real.logb_prod
 
 end Real
+
+section Induction
+
+/-- Induction principle for intervals of real numbers: if a proposition `P` is true
+on `[x₀, r * x₀)` and if `P` for `[x₀, r^n * x₀)` implies `P` for `[r^n * x₀, r^(n+1) * x₀)`,
+then `P` is true for all `x ≥ x₀`. -/
+lemma Real.induction_Ico_mul {P : ℝ → Prop} (x₀ r : ℝ) (hr : 1 < r) (hx₀ : 0 < x₀)
+    (base : ∀ x ∈ Set.Ico x₀ (r * x₀), P x)
+    (step : ∀ n : ℕ, n ≥ 1 → (∀ z ∈ Set.Ico x₀ (r ^ n * x₀), P z) →
+      (∀ z ∈ Set.Ico (r ^ n * x₀) (r ^ (n+1) * x₀), P z)) :
+    ∀ x ≥ x₀, P x := by
+  suffices : ∀ n : ℕ, ∀ x ∈ Set.Ico x₀ (r ^ (n + 1) * x₀), P x
+  · intro x hx
+    have hx' : 0 < x / x₀ := div_pos (hx₀.trans_le hx) hx₀
+    refine this ⌊logb r (x / x₀)⌋₊ x ?_
+    rw [mem_Ico, ←div_lt_iff hx₀, ←logb_lt_iff_lt_rpow hr hx']
+    exact ⟨hx, Nat.lt_floor_add_one _⟩
+  intro n
+  induction n
+  case zero => simpa using base
+  case succ n ih =>
+    specialize step (n + 1) (by simp)
+    simp only [Nat.cast_add_one] at step ⊢
+    exact fun x hx => (Ico_subset_Ico_union_Ico hx).elim (ih x) (step ih _)
+
+end Induction
feat(Algebra/GroupWithZero/Units/Lemmas): div_div_div_cancel_left' (#6606)

Co-authored-by: Moritz Firsching <firsching@google.com>

Diff
@@ -102,10 +102,8 @@ theorem mul_logb {a b c : ℝ} (h₁ : b ≠ 0) (h₂ : b ≠ 1) (h₃ : b ≠ -
 #align real.mul_logb Real.mul_logb
 
 theorem div_logb {a b c : ℝ} (h₁ : c ≠ 0) (h₂ : c ≠ 1) (h₃ : c ≠ -1) :
-    logb a c / logb b c = logb a b := by
-  unfold logb
-  -- TODO: div_div_div_cancel_left is missing for `group_with_zero`,
-  rw [div_div_div_eq, mul_comm, mul_div_mul_right _ _ (log_ne_zero.mpr ⟨h₁, h₂, h₃⟩)]
+    logb a c / logb b c = logb a b :=
+  div_div_div_cancel_left' _ _ <| log_ne_zero.mpr ⟨h₁, h₂, h₃⟩
 #align real.div_logb Real.div_logb
 
 section BPosAndNeOne
chore: banish Type _ and Sort _ (#6499)

We remove all possible occurences of Type _ and Sort _ in favor of Type* and Sort*.

This has nice performance benefits.

Diff
@@ -419,7 +419,7 @@ theorem logb_eq_zero : logb b x = 0 ↔ b = 0 ∨ b = 1 ∨ b = -1 ∨ x = 0 ∨
 -- TODO add other limits and continuous API lemmas analogous to those in Log.lean
 open BigOperators
 
-theorem logb_prod {α : Type _} (s : Finset α) (f : α → ℝ) (hf : ∀ x ∈ s, f x ≠ 0) :
+theorem logb_prod {α : Type*} (s : Finset α) (f : α → ℝ) (hf : ∀ x ∈ s, f x ≠ 0) :
     logb b (∏ i in s, f i) = ∑ i in s, logb b (f i) := by
   classical
     induction' s using Finset.induction_on with a s ha ih
chore: script to replace headers with #align_import statements (#5979)

Open in Gitpod

Co-authored-by: Eric Wieser <wieser.eric@gmail.com> Co-authored-by: Scott Morrison <scott.morrison@gmail.com>

Diff
@@ -2,15 +2,12 @@
 Copyright (c) 2022 Bolton Bailey. All rights reserved.
 Released under Apache 2.0 license as described in the file LICENSE.
 Authors: Bolton Bailey, Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne
-
-! This file was ported from Lean 3 source module analysis.special_functions.log.base
-! leanprover-community/mathlib commit f23a09ce6d3f367220dc3cecad6b7eb69eb01690
-! Please do not edit these lines, except to modify the commit id
-! if you have ported upstream changes.
 -/
 import Mathlib.Analysis.SpecialFunctions.Pow.Real
 import Mathlib.Data.Int.Log
 
+#align_import analysis.special_functions.log.base from "leanprover-community/mathlib"@"f23a09ce6d3f367220dc3cecad6b7eb69eb01690"
+
 /-!
 # Real logarithm base `b`
 
fix: precedence of , and abs (#5619)
Diff
@@ -57,7 +57,7 @@ theorem logb_one : logb b 1 = 0 := by simp [logb]
 #align real.logb_one Real.logb_one
 
 @[simp]
-theorem logb_abs (x : ℝ) : logb b (|x|) = logb b x := by rw [logb, logb, log_abs]
+theorem logb_abs (x : ℝ) : logb b |x| = logb b x := by rw [logb, logb, log_abs]
 #align real.logb_abs Real.logb_abs
 
 @[simp]
feat: add Real.logb_mul_base (#5609)

Forward-port leanprover-community/mathlib#18979

Diff
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
 Authors: Bolton Bailey, Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne
 
 ! This file was ported from Lean 3 source module analysis.special_functions.log.base
-! leanprover-community/mathlib commit 0b9eaaa7686280fad8cce467f5c3c57ee6ce77f8
+! leanprover-community/mathlib commit f23a09ce6d3f367220dc3cecad6b7eb69eb01690
 ! Please do not edit these lines, except to modify the commit id
 ! if you have ported upstream changes.
 -/
@@ -77,6 +77,40 @@ theorem logb_div (hx : x ≠ 0) (hy : y ≠ 0) : logb b (x / y) = logb b x - log
 theorem logb_inv (x : ℝ) : logb b x⁻¹ = -logb b x := by simp [logb, neg_div]
 #align real.logb_inv Real.logb_inv
 
+theorem inv_logb (a b : ℝ) : (logb a b)⁻¹ = logb b a := by simp_rw [logb, inv_div]
+#align real.inv_logb Real.inv_logb
+
+theorem inv_logb_mul_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
+    (logb (a * b) c)⁻¹ = (logb a c)⁻¹ + (logb b c)⁻¹ := by
+  simp_rw [inv_logb]; exact logb_mul h₁ h₂
+#align real.inv_logb_mul_base Real.inv_logb_mul_base
+
+theorem inv_logb_div_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
+    (logb (a / b) c)⁻¹ = (logb a c)⁻¹ - (logb b c)⁻¹ := by
+  simp_rw [inv_logb]; exact logb_div h₁ h₂
+#align real.inv_logb_div_base Real.inv_logb_div_base
+
+theorem logb_mul_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
+    logb (a * b) c = ((logb a c)⁻¹ + (logb b c)⁻¹)⁻¹ := by rw [← inv_logb_mul_base h₁ h₂ c, inv_inv]
+#align real.logb_mul_base Real.logb_mul_base
+
+theorem logb_div_base {a b : ℝ} (h₁ : a ≠ 0) (h₂ : b ≠ 0) (c : ℝ) :
+    logb (a / b) c = ((logb a c)⁻¹ - (logb b c)⁻¹)⁻¹ := by rw [← inv_logb_div_base h₁ h₂ c, inv_inv]
+#align real.logb_div_base Real.logb_div_base
+
+theorem mul_logb {a b c : ℝ} (h₁ : b ≠ 0) (h₂ : b ≠ 1) (h₃ : b ≠ -1) :
+    logb a b * logb b c = logb a c := by
+  unfold logb
+  rw [mul_comm, div_mul_div_cancel _ (log_ne_zero.mpr ⟨h₁, h₂, h₃⟩)]
+#align real.mul_logb Real.mul_logb
+
+theorem div_logb {a b c : ℝ} (h₁ : c ≠ 0) (h₂ : c ≠ 1) (h₃ : c ≠ -1) :
+    logb a c / logb b c = logb a b := by
+  unfold logb
+  -- TODO: div_div_div_cancel_left is missing for `group_with_zero`,
+  rw [div_div_div_eq, mul_comm, mul_div_mul_right _ _ (log_ne_zero.mpr ⟨h₁, h₂, h₃⟩)]
+#align real.div_logb Real.div_logb
+
 section BPosAndNeOne
 
 variable (b_pos : 0 < b) (b_ne_one : b ≠ 1)
feat: port Analysis.SpecialFunctions.Log.Base (#4140)

Dependencies 12 + 748

749 files ported (98.4%)
328214 lines ported (98.3%)
Show graph

The unported dependencies are

The following 1 dependencies have changed in mathlib3 since they were ported, which may complicate porting this file