Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 53b216b

Browse files
committed
refactor(algebra/order/ring): Make strict_ordered_semirings nontrivial (#17394)
`strict_ordered_semiring` + `nontrivial` implies `char_zero`, but this can't be instance because `char_zero` implies `nontrivial` Instead of waiting for Lean 4, where such looping instances are possible, we make all `strict_ordered_semiring`s nontrivial, because we don't care about types with one element being`strict_ordered_semiring`s.
1 parent 36f5ed0 commit 53b216b

File tree

22 files changed

+122
-152
lines changed

22 files changed

+122
-152
lines changed

counterexamples/canonically_ordered_comm_semiring_two_mul.lean

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -140,7 +140,8 @@ instance socsN2 : strict_ordered_comm_semiring (ℕ × zmod 2) :=
140140
mul_lt_mul_of_pos_right := mul_lt_mul_of_pos_right,
141141
..Nxzmod_2.csrN2_1,
142142
..(infer_instance : partial_order (ℕ × zmod 2)),
143-
..(infer_instance : comm_semiring (ℕ × zmod 2)) }
143+
..(infer_instance : comm_semiring (ℕ × zmod 2)),
144+
..pullback_nonzero prod.fst prod.fst_zero prod.fst_one }
144145

145146
end Nxzmod_2
146147

src/algebra/geom_sum.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -409,7 +409,7 @@ section order
409409

410410
variables {n : ℕ} {x : α}
411411

412-
lemma geom_sum_pos [strict_ordered_semiring α] [nontrivial α] (hx : 0 ≤ x) (hn : n ≠ 0) :
412+
lemma geom_sum_pos [strict_ordered_semiring α] (hx : 0 ≤ x) (hn : n ≠ 0) :
413413
0 < ∑ i in range n, x ^ i :=
414414
sum_pos' (λ k hk, pow_nonneg hx _) ⟨0, mem_range.2 hn.bot_lt, by simp⟩
415415

src/algebra/order/archimedean.lean

Lines changed: 6 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -83,8 +83,7 @@ lemma exists_unique_add_zsmul_mem_Ioc {a : α} (ha : 0 < a) (b c : α) :
8383

8484
end linear_ordered_add_comm_group
8585

86-
theorem exists_nat_gt [strict_ordered_semiring α] [nontrivial α] [archimedean α]
87-
(x : α) : ∃ n : ℕ, x < n :=
86+
theorem exists_nat_gt [strict_ordered_semiring α] [archimedean α] (x : α) : ∃ n : ℕ, x < n :=
8887
let ⟨n, h⟩ := archimedean.arch x zero_lt_one in
8988
⟨n+1, lt_of_le_of_lt (by rwa ← nsmul_one)
9089
(nat.cast_lt.2 (nat.lt_succ_self _))⟩
@@ -96,8 +95,8 @@ begin
9695
exact (exists_nat_gt x).imp (λ n, le_of_lt)
9796
end
9897

99-
lemma add_one_pow_unbounded_of_pos [strict_ordered_semiring α] [nontrivial α] [archimedean α]
100-
(x : α) {y : α} (hy : 0 < y) :
98+
lemma add_one_pow_unbounded_of_pos [strict_ordered_semiring α] [archimedean α] (x : α) {y : α}
99+
(hy : 0 < y) :
101100
∃ n : ℕ, x < (y + 1) ^ n :=
102101
have 01 + y, from add_nonneg zero_le_one hy.le,
103102
let ⟨n, h⟩ := archimedean.arch x hy in
@@ -109,7 +108,7 @@ let ⟨n, h⟩ := archimedean.arch x hy in
109108
... = (y + 1) ^ n : by rw [add_comm]⟩
110109

111110
section strict_ordered_ring
112-
variables [strict_ordered_ring α] [nontrivial α] [archimedean α]
111+
variables [strict_ordered_ring α] [archimedean α]
113112

114113
lemma pow_unbounded_of_one_lt (x : α) {y : α} (hy1 : 1 < y) :
115114
∃ n : ℕ, x < y ^ n :=
@@ -270,7 +269,7 @@ section linear_ordered_field
270269
variables [linear_ordered_field α]
271270

272271
lemma archimedean_iff_nat_lt : archimedean α ↔ ∀ x : α, ∃ n : ℕ, x < n :=
273-
⟨@exists_nat_gt α _ _, λ H, ⟨λ x y y0,
272+
⟨@exists_nat_gt α _, λ H, ⟨λ x y y0,
274273
(H (x / y)).imp $ λ n h, le_of_lt $
275274
by rwa [div_lt_iff y0, ← nsmul_eq_mul] at h⟩⟩
276275

@@ -281,7 +280,7 @@ archimedean_iff_nat_lt.trans
281280
lt_of_le_of_lt h (nat.cast_lt.2 (lt_add_one _))⟩⟩
282281

283282
lemma archimedean_iff_int_lt : archimedean α ↔ ∀ x : α, ∃ n : ℤ, x < n :=
284-
⟨@exists_int_gt α _ _,
283+
⟨@exists_int_gt α _,
285284
begin
286285
rw archimedean_iff_nat_lt,
287286
intros h x,

src/algebra/order/ring/cone.lean

Lines changed: 12 additions & 24 deletions
Original file line numberDiff line numberDiff line change
@@ -14,6 +14,8 @@ import algebra.order.ring.defs
1414

1515
set_option old_structure_cmd true
1616

17+
variables {α : Type*} [ring α] [nontrivial α]
18+
1719
namespace ring
1820

1921
/-- A positive cone in a ring consists of a positive cone in underlying `add_comm_group`,
@@ -26,27 +28,28 @@ structure positive_cone (α : Type*) [ring α] extends add_comm_group.positive_c
2628
/-- Forget that a positive cone in a ring respects the multiplicative structure. -/
2729
add_decl_doc positive_cone.to_positive_cone
2830

29-
/-- A positive cone in a ring induces a linear order if `1` is a positive element. -/
31+
/-- A total positive cone in a nontrivial ring induces a linear order. -/
3032
@[nolint has_nonempty_instance]
3133
structure total_positive_cone (α : Type*) [ring α]
32-
extends positive_cone α, add_comm_group.total_positive_cone α :=
33-
(one_pos : pos 1)
34+
extends positive_cone α, add_comm_group.total_positive_cone α
3435

3536
/-- Forget that a `total_positive_cone` in a ring is total. -/
3637
add_decl_doc total_positive_cone.to_positive_cone
3738

3839
/-- Forget that a `total_positive_cone` in a ring respects the multiplicative structure. -/
3940
add_decl_doc total_positive_cone.to_total_positive_cone
4041

41-
end ring
42+
lemma positive_cone.one_pos (C : positive_cone α) : C.pos 1 :=
43+
(C.pos_iff _).2 ⟨C.one_nonneg, λ h, one_ne_zero $ C.nonneg_antisymm C.one_nonneg h⟩
4244

43-
namespace strict_ordered_ring
45+
end ring
4446

4547
open ring
4648

4749
/-- Construct a `strict_ordered_ring` by designating a positive cone in an existing `ring`. -/
48-
def mk_of_positive_cone {α : Type*} [ring α] (C : positive_cone α) : strict_ordered_ring α :=
49-
{ zero_le_one := by { change C.nonneg (1 - 0), convert C.one_nonneg, simp, },
50+
def strict_ordered_ring.mk_of_positive_cone (C : positive_cone α) : strict_ordered_ring α :=
51+
{ exists_pair_ne := ⟨0, 1, λ h, by simpa [←h, C.pos_iff] using C.one_pos⟩,
52+
zero_le_one := by { change C.nonneg (1 - 0), convert C.one_nonneg, simp, },
5053
mul_pos := λ x y xp yp, begin
5154
change C.pos (x*y - 0),
5255
convert C.mul_pos x y (by { convert xp, simp, }) (by { convert yp, simp, }),
@@ -55,23 +58,8 @@ def mk_of_positive_cone {α : Type*} [ring α] (C : positive_cone α) : strict_o
5558
..‹ring α›,
5659
..ordered_add_comm_group.mk_of_positive_cone C.to_positive_cone }
5760

58-
end strict_ordered_ring
59-
60-
namespace linear_ordered_ring
61-
62-
open ring
63-
6461
/-- Construct a `linear_ordered_ring` by
6562
designating a positive cone in an existing `ring`. -/
66-
def mk_of_positive_cone {α : Type*} [ring α] (C : total_positive_cone α) :
67-
linear_ordered_ring α :=
68-
{ exists_pair_ne := ⟨0, 1, begin
69-
intro h,
70-
have one_pos := C.one_pos,
71-
rw [←h, C.pos_iff] at one_pos,
72-
simpa using one_pos,
73-
end⟩,
74-
..strict_ordered_ring.mk_of_positive_cone C.to_positive_cone,
63+
def linear_ordered_ring.mk_of_positive_cone (C : total_positive_cone α) : linear_ordered_ring α :=
64+
{ ..strict_ordered_ring.mk_of_positive_cone C.to_positive_cone,
7565
..linear_ordered_add_comm_group.mk_of_positive_cone C.to_total_positive_cone, }
76-
77-
end linear_ordered_ring

src/algebra/order/ring/defs.lean

Lines changed: 36 additions & 38 deletions
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,7 @@
11
/-
22
Copyright (c) 2016 Jeremy Avigad. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
4-
Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro
4+
Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Yaël Dillies
55
-/
66
import order.min_max
77
import algebra.order.monoid.cancel.defs
@@ -29,21 +29,23 @@ For short,
2929
## Typeclasses
3030
3131
* `ordered_semiring`: Semiring with a partial order such that `+` and `*` respect `≤`.
32-
* `strict_ordered_semiring`: Semiring with a partial order such that `+` and `*` respects `<`.
32+
* `strict_ordered_semiring`: Nontrivial semiring with a partial order such that `+` and `*` respects
33+
`<`.
3334
* `ordered_comm_semiring`: Commutative semiring with a partial order such that `+` and `*` respect
3435
`≤`.
35-
* `strict_ordered_comm_semiring`: Commutative semiring with a partial order such that `+` and `*`
36-
respect `<`.
36+
* `strict_ordered_comm_semiring`: Nontrivial commutative semiring with a partial order such that `+`
37+
and `*` respect `<`.
3738
* `ordered_ring`: Ring with a partial order such that `+` respects `≤` and `*` respects `<`.
3839
* `ordered_comm_ring`: Commutative ring with a partial order such that `+` respects `≤` and
3940
`*` respects `<`.
40-
* `linear_ordered_semiring`: Semiring with a linear order such that `+` respects `≤` and
41+
* `linear_ordered_semiring`: Nontrivial semiring with a linear order such that `+` respects `≤` and
4142
`*` respects `<`.
42-
* `linear_ordered_comm_semiring`: Commutative semiring with a linear order such that `+` respects
43+
* `linear_ordered_comm_semiring`: Nontrivial commutative semiring with a linear order such that `+`
44+
respects `≤` and `*` respects `<`.
45+
* `linear_ordered_ring`: Nontrivial ring with a linear order such that `+` respects `≤` and `*`
46+
respects `<`.
47+
* `linear_ordered_comm_ring`: Nontrivial commutative ring with a linear order such that `+` respects
4348
`≤` and `*` respects `<`.
44-
* `linear_ordered_ring`: Ring with a linear order such that `+` respects `≤` and `*` respects `<`.
45-
* `linear_ordered_comm_ring`: Commutative ring with a linear order such that `+` respects `≤` and
46-
`*` respects `<`.
4749
* `canonically_ordered_comm_semiring`: Commutative semiring with a partial order such that `+`
4850
respects `≤`, `*` respects `<`, and `a ≤ b ↔ ∃ c, b = a + c`.
4951
@@ -57,35 +59,42 @@ immediate predecessors and what conditions are added to each of them.
5759
- `ordered_add_comm_monoid` & multiplication & `*` respects `≤`
5860
- `semiring` & partial order structure & `+` respects `≤` & `*` respects `≤`
5961
* `strict_ordered_semiring`
60-
- `ordered_cancel_add_comm_monoid` & multiplication & `*` respects `<`
61-
- `ordered_semiring` & `+` respects `<` & `*` respects `<`
62+
- `ordered_cancel_add_comm_monoid` & multiplication & `*` respects `<` & nontriviality
63+
- `ordered_semiring` & `+` respects `<` & `*` respects `<` & nontriviality
6264
* `ordered_comm_semiring`
6365
- `ordered_semiring` & commutativity of multiplication
6466
- `comm_semiring` & partial order structure & `+` respects `≤` & `*` respects `<`
6567
* `strict_ordered_comm_semiring`
6668
- `strict_ordered_semiring` & commutativity of multiplication
67-
- `ordered_comm_semiring` & `+` respects `<` & `*` respects `<`
69+
- `ordered_comm_semiring` & `+` respects `<` & `*` respects `<` & nontriviality
6870
* `ordered_ring`
69-
- `strict_ordered_semiring` & additive inverses
71+
- `ordered_semiring` & additive inverses
7072
- `ordered_add_comm_group` & multiplication & `*` respects `<`
7173
- `ring` & partial order structure & `+` respects `≤` & `*` respects `<`
74+
* `strict_ordered_ring`
75+
- `strict_ordered_semiring` & additive inverses
76+
- `ordered_semiring` & `+` respects `<` & `*` respects `<` & nontriviality
7277
* `ordered_comm_ring`
7378
- `ordered_ring` & commutativity of multiplication
7479
- `ordered_comm_semiring` & additive inverses
7580
- `comm_ring` & partial order structure & `+` respects `≤` & `*` respects `<`
81+
* `strict_ordered_comm_ring`
82+
- `strict_ordered_comm_semiring` & additive inverses
83+
- `strict_ordered_ring` & commutativity of multiplication
84+
- `ordered_comm_ring` & `+` respects `<` & `*` respects `<` & nontriviality
7685
* `linear_ordered_semiring`
77-
- `strict_ordered_semiring` & totality of the order & nontriviality
86+
- `strict_ordered_semiring` & totality of the order
7887
- `linear_ordered_add_comm_monoid` & multiplication & nontriviality & `*` respects `<`
7988
* `linear_ordered_comm_semiring`
80-
- `strict_ordered_comm_semiring` & totality of the order & nontriviality
89+
- `strict_ordered_comm_semiring` & totality of the order
8190
- `linear_ordered_semiring` & commutativity of multiplication
8291
* `linear_ordered_ring`
83-
- `ordered_ring` & totality of the order & nontriviality
92+
- `strict_ordered_ring` & totality of the order
8493
- `linear_ordered_semiring` & additive inverses
8594
- `linear_ordered_add_comm_group` & multiplication & `*` respects `<`
8695
- `domain` & linear order structure
8796
* `linear_ordered_comm_ring`
88-
- `ordered_comm_ring` & totality of the order & nontriviality
97+
- `strict_ordered_comm_ring` & totality of the order
8998
- `linear_ordered_ring` & commutativity of multiplication
9099
- `linear_ordered_comm_semiring` & additive inverses
91100
- `is_domain` & linear order structure
@@ -133,10 +142,11 @@ and multiplication by a nonnegative number is monotone. -/
133142
@[protect_proj, ancestor ordered_ring comm_ring]
134143
class ordered_comm_ring (α : Type u) extends ordered_ring α, comm_ring α
135144

136-
/-- A `strict_ordered_semiring` is a semiring with a partial order such that addition is strictly
137-
monotone and multiplication by a positive number is strictly monotone. -/
138-
@[protect_proj, ancestor semiring ordered_cancel_add_comm_monoid]
139-
class strict_ordered_semiring (α : Type u) extends semiring α, ordered_cancel_add_comm_monoid α :=
145+
/-- A `strict_ordered_semiring` is a nontrivial semiring with a partial order such that addition is
146+
strictly monotone and multiplication by a positive number is strictly monotone. -/
147+
@[protect_proj, ancestor semiring ordered_cancel_add_comm_monoid nontrivial]
148+
class strict_ordered_semiring (α : Type u)
149+
extends semiring α, ordered_cancel_add_comm_monoid α, nontrivial α :=
140150
(zero_le_one : (0 : α) ≤ 1)
141151
(mul_lt_mul_of_pos_left : ∀ a b c : α, a < b → 0 < c → c * a < c * b)
142152
(mul_lt_mul_of_pos_right : ∀ a b c : α, a < b → 0 < c → a * c < b * c)
@@ -148,8 +158,8 @@ class strict_ordered_comm_semiring (α : Type u) extends strict_ordered_semiring
148158

149159
/-- A `strict_ordered_ring` is a ring with a partial order such that addition is strictly monotone
150160
and multiplication by a positive number is strictly monotone. -/
151-
@[protect_proj, ancestor ring ordered_add_comm_group]
152-
class strict_ordered_ring (α : Type u) extends ring α, ordered_add_comm_group α :=
161+
@[protect_proj, ancestor ring ordered_add_comm_group nontrivial]
162+
class strict_ordered_ring (α : Type u) extends ring α, ordered_add_comm_group α, nontrivial α :=
153163
(zero_le_one : 0 ≤ (1 : α))
154164
(mul_pos : ∀ a b : α, 0 < a → 0 < b → 0 < a * b)
155165

@@ -165,7 +175,7 @@ explore changing this, but be warned that the instances involving `domain` may c
165175
search loops. -/
166176
@[protect_proj, ancestor strict_ordered_semiring linear_ordered_add_comm_monoid nontrivial]
167177
class linear_ordered_semiring (α : Type u)
168-
extends strict_ordered_semiring α, linear_ordered_add_comm_monoid α, nontrivial α
178+
extends strict_ordered_semiring α, linear_ordered_add_comm_monoid α
169179

170180
/-- A `linear_ordered_comm_semiring` is a nontrivial commutative semiring with a linear order such
171181
that addition is monotone and multiplication by a positive number is strictly monotone. -/
@@ -175,8 +185,8 @@ class linear_ordered_comm_semiring (α : Type*)
175185

176186
/-- A `linear_ordered_ring` is a ring with a linear order such that addition is monotone and
177187
multiplication by a positive number is strictly monotone. -/
178-
@[protect_proj, ancestor strict_ordered_ring linear_order nontrivial]
179-
class linear_ordered_ring (α : Type u) extends strict_ordered_ring α, linear_order α, nontrivial α
188+
@[protect_proj, ancestor strict_ordered_ring linear_order]
189+
class linear_ordered_ring (α : Type u) extends strict_ordered_ring α, linear_order α
180190

181191
/-- A `linear_ordered_comm_ring` is a commutative ring with a linear order such that addition is
182192
monotone and multiplication by a positive number is strictly monotone. -/
@@ -518,29 +528,17 @@ lemma strict_mono.mul (hf : strict_mono f) (hg : strict_mono g) (hf₀ : ∀ x,
518528

519529
end monotone
520530

521-
section nontrivial
522-
variables [nontrivial α]
523-
524531
lemma lt_one_add (a : α) : a < 1 + a := lt_add_of_pos_left _ zero_lt_one
525532
lemma lt_add_one (a : α) : a < a + 1 := lt_add_of_pos_right _ zero_lt_one
526533

527534
lemma one_lt_two : (1 : α) < 2 := lt_add_one _
528535

529536
lemma lt_two_mul_self (ha : 0 < a) : a < 2 * a := lt_mul_of_one_lt_left ha one_lt_two
530537

531-
lemma nat.strict_mono_cast : strict_mono (coe : ℕ → α) :=
532-
strict_mono_nat_of_lt_succ $ λ n, by { rw [nat.cast_succ], apply lt_add_one }
533-
534-
/-- `coe : ℕ → α` as an `order_embedding` -/
535-
@[simps { fully_applied := ff }] def nat.cast_order_embedding : ℕ ↪o α :=
536-
order_embedding.of_strict_mono coe nat.strict_mono_cast
537-
538538
@[priority 100] -- see Note [lower instance priority]
539539
instance strict_ordered_semiring.to_no_max_order : no_max_order α :=
540540
⟨λ a, ⟨a + 1, lt_add_of_pos_right _ one_pos⟩⟩
541541

542-
end nontrivial
543-
544542
end strict_ordered_semiring
545543

546544
section strict_ordered_comm_semiring

src/algebra/order/ring/inj_surj.lean

Lines changed: 2 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -90,7 +90,8 @@ protected def strict_ordered_semiring [strict_ordered_semiring α] [has_zero β]
9090
mul_lt_mul_of_pos_right := λ a b c h hc, show f (a * c) < f (b * c),
9191
by simpa only [mul, zero] using mul_lt_mul_of_pos_right ‹f a < f b› (by rwa ←zero),
9292
..hf.ordered_cancel_add_comm_monoid f zero add nsmul,
93-
..hf.ordered_semiring f zero one add mul nsmul npow nat_cast }
93+
..hf.ordered_semiring f zero one add mul nsmul npow nat_cast,
94+
..pullback_nonzero f zero one }
9495

9596
/-- Pullback a `strict_ordered_comm_semiring` under an injective map. -/
9697
@[reducible] -- See note [reducible non-instances]
@@ -143,7 +144,6 @@ protected def linear_ordered_semiring [linear_ordered_semiring α] [has_zero β]
143144
(hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) :
144145
linear_ordered_semiring β :=
145146
{ .. linear_order.lift f hf hsup hinf,
146-
.. pullback_nonzero f zero one,
147147
.. hf.strict_ordered_semiring f zero one add mul nsmul npow nat_cast }
148148

149149
/-- Pullback a `linear_ordered_semiring` under an injective map. -/
@@ -172,7 +172,6 @@ def linear_ordered_ring [linear_ordered_ring α] [has_zero β] [has_one β] [has
172172
(hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) :
173173
linear_ordered_ring β :=
174174
{ .. linear_order.lift f hf hsup hinf,
175-
.. pullback_nonzero f zero one,
176175
.. hf.strict_ordered_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast }
177176

178177
/-- Pullback a `linear_ordered_comm_ring` under an injective map. -/
@@ -188,7 +187,6 @@ protected def linear_ordered_comm_ring [linear_ordered_comm_ring α] [has_zero
188187
(hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) :
189188
linear_ordered_comm_ring β :=
190189
{ .. linear_order.lift f hf hsup hinf,
191-
.. pullback_nonzero f zero one,
192190
.. hf.strict_ordered_comm_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast }
193191

194192
end function.injective

src/algebra/order/ring/nontrivial.lean

Lines changed: 4 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -7,25 +7,12 @@ import algebra.char_zero.defs
77
import algebra.order.ring.defs
88

99
/-!
10-
# Nontrivial strict ordered semirings (and hence linear ordered semirings) are characteristic zero.
10+
# Strict ordered semiring have characteristic zero
1111
-/
1212

1313
variables {α : Type*}
1414

15-
section strict_ordered_semiring
16-
variables [strict_ordered_semiring α] [nontrivial α]
17-
18-
/-- Note this is not an instance as `char_zero` implies `nontrivial`, and this would risk forming a
19-
loop. -/
20-
lemma strict_ordered_semiring.to_char_zero : char_zero α := ⟨nat.strict_mono_cast.injective⟩
21-
22-
end strict_ordered_semiring
23-
24-
section linear_ordered_semiring
25-
variables [linear_ordered_semiring α]
26-
2715
@[priority 100] -- see Note [lower instance priority]
28-
instance linear_ordered_semiring.to_char_zero : char_zero α :=
29-
strict_ordered_semiring.to_char_zero
30-
31-
end linear_ordered_semiring
16+
instance strict_ordered_semiring.to_char_zero [strict_ordered_semiring α] : char_zero α :=
17+
⟨strict_mono.injective $ strict_mono_nat_of_lt_succ $ λ n,
18+
by { rw [nat.cast_succ], apply lt_add_one }⟩

src/analysis/box_integral/box/subbox_induction.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -147,7 +147,7 @@ begin
147147
{ suffices : tendsto (λ m, (J m).upper - (J m).lower) at_top (𝓝 0), by simpa using hJlz.add this,
148148
refine tendsto_pi_nhds.2 (λ i, _),
149149
simpa [hJsub] using tendsto_const_nhds.div_at_top
150-
(tendsto_pow_at_top_at_top_of_one_lt (@one_lt_two ℝ _ _)) },
150+
(tendsto_pow_at_top_at_top_of_one_lt $ @one_lt_two ℝ _) },
151151
replace hJlz : tendsto (λ m, (J m).lower) at_top (𝓝[Icc I.lower I.upper] z),
152152
from tendsto_nhds_within_of_tendsto_nhds_of_eventually_within _ hJlz
153153
(eventually_of_forall hJl_mem),

0 commit comments

Comments
 (0)