Skip to content

Commit d025702

Browse files
committed
chore: Move pow_lt_pow_succ to Algebra.Order.WithZero (#11507)
These lemmas can be defined earlier, ridding us of an import in `Algebra.GroupPower.Order`
1 parent 119c442 commit d025702

File tree

6 files changed

+38
-51
lines changed

6 files changed

+38
-51
lines changed

Mathlib/Algebra/GroupPower/Order.lean

Lines changed: 0 additions & 27 deletions
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,6 @@ Authors: Jeremy Avigad, Robert Y. Lewis
55
-/
66
import Mathlib.Algebra.GroupPower.CovariantClass
77
import Mathlib.Algebra.GroupPower.Ring
8-
import Mathlib.Algebra.Order.WithZero
98

109
#align_import algebra.group_power.order from "leanprover-community/mathlib"@"00f91228655eecdcd3ac97a7fd8dbcb139fe990a"
1110

@@ -508,31 +507,6 @@ lemma pow_four_le_pow_two_of_pow_two_le (h : a ^ 2 ≤ b) : a ^ 4 ≤ b ^ 2 :=
508507

509508
end LinearOrderedRing
510509

511-
section LinearOrderedCommMonoidWithZero
512-
513-
variable [LinearOrderedCommMonoidWithZero M] [NoZeroDivisors M] {a : M} {n : ℕ}
514-
515-
theorem pow_pos_iff (hn : n ≠ 0) : 0 < a ^ n ↔ 0 < a := by simp_rw [zero_lt_iff, pow_ne_zero_iff hn]
516-
#align pow_pos_iff pow_pos_iff
517-
518-
end LinearOrderedCommMonoidWithZero
519-
520-
section LinearOrderedCommGroupWithZero
521-
522-
variable [LinearOrderedCommGroupWithZero M] {a : M} {m n : ℕ}
523-
524-
theorem pow_lt_pow_succ (ha : 1 < a) : a ^ n < a ^ n.succ := by
525-
rw [← one_mul (a ^ n), pow_succ]
526-
exact mul_lt_right₀ _ ha (pow_ne_zero _ (zero_lt_one.trans ha).ne')
527-
#align pow_lt_pow_succ pow_lt_pow_succ
528-
529-
theorem pow_lt_pow_right₀ (ha : 1 < a) (hmn : m < n) : a ^ m < a ^ n := by
530-
induction' hmn with n _ ih
531-
exacts [pow_lt_pow_succ ha, lt_trans ih (pow_lt_pow_succ ha)]
532-
#align pow_lt_pow₀ pow_lt_pow_right₀
533-
534-
end LinearOrderedCommGroupWithZero
535-
536510
namespace MonoidHom
537511

538512
variable [Ring R] [Monoid M] [LinearOrder M] [CovariantClass M M (· * ·) (· ≤ ·)] (f : R →* M)
@@ -571,5 +545,4 @@ Those lemmas have been deprecated on 2023-12-23.
571545
@[deprecated] alias pow_lt_pow_of_lt_one := pow_lt_pow_right_of_lt_one
572546
@[deprecated] alias lt_of_pow_lt_pow := lt_of_pow_lt_pow_left
573547
@[deprecated] alias le_of_pow_le_pow := le_of_pow_le_pow_left
574-
@[deprecated] alias pow_lt_pow₀ := pow_lt_pow_right₀
575548
@[deprecated] alias self_le_pow := le_self_pow

Mathlib/Algebra/Order/Field/Basic.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Robert Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn
55
-/
66
import Mathlib.Algebra.Field.Basic
7+
import Mathlib.Algebra.GroupWithZero.Units.Equiv
78
import Mathlib.Algebra.Order.Field.Defs
89
import Mathlib.Algebra.Order.Ring.Abs
910
import Mathlib.Order.Bounds.OrderIso

Mathlib/Algebra/Order/Floor.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -334,7 +334,7 @@ theorem ceil_one : ⌈(1 : α)⌉₊ = 1 := by rw [← Nat.cast_one, ceil_natCas
334334
#align nat.ceil_one Nat.ceil_one
335335

336336
@[simp]
337-
theorem ceil_eq_zero : ⌈a⌉₊ = 0 ↔ a ≤ 0 := by rw [← le_zero_iff, ceil_le, Nat.cast_zero]
337+
theorem ceil_eq_zero : ⌈a⌉₊ = 0 ↔ a ≤ 0 := by rw [← Nat.le_zero, ceil_le, Nat.cast_zero]
338338
#align nat.ceil_eq_zero Nat.ceil_eq_zero
339339

340340
@[simp]

Mathlib/Algebra/Order/WithZero.lean

Lines changed: 22 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -36,8 +36,7 @@ class LinearOrderedCommGroupWithZero (α : Type*) extends LinearOrderedCommMonoi
3636
CommGroupWithZero α
3737
#align linear_ordered_comm_group_with_zero LinearOrderedCommGroupWithZero
3838

39-
variable {α : Type*}
40-
variable {a b c d x y z : α}
39+
variable {α : Type*} {a b c d x y z : α}
4140

4241
instance instLinearOrderedCommMonoidWithZeroMultiplicativeOrderDual
4342
[LinearOrderedAddCommMonoidWithTop α] :
@@ -67,9 +66,8 @@ instance instLinearOrderedCommMonoidWithZeroWithZero [LinearOrderedCommMonoid α
6766
instance [LinearOrderedCommGroup α] : LinearOrderedCommGroupWithZero (WithZero α) :=
6867
{ instLinearOrderedCommMonoidWithZeroWithZero, WithZero.commGroupWithZero with }
6968

70-
section LinearOrderedCommMonoid
71-
72-
variable [LinearOrderedCommMonoidWithZero α]
69+
section LinearOrderedCommMonoidWithZero
70+
variable [LinearOrderedCommMonoidWithZero α] {n : ℕ}
7371

7472
/-
7573
The following facts are true more generally in a (linearly) ordered commutative monoid.
@@ -117,9 +115,14 @@ instance instLinearOrderedAddCommMonoidWithTopAdditiveOrderDual :
117115
le_top := fun _ ↦ zero_le' }
118116
#align additive.linear_ordered_add_comm_monoid_with_top instLinearOrderedAddCommMonoidWithTopAdditiveOrderDual
119117

120-
end LinearOrderedCommMonoid
118+
variable [NoZeroDivisors α]
119+
120+
lemma pow_pos_iff (hn : n ≠ 0) : 0 < a ^ n ↔ 0 < a := by simp_rw [zero_lt_iff, pow_ne_zero_iff hn]
121+
#align pow_pos_iff pow_pos_iff
122+
123+
end LinearOrderedCommMonoidWithZero
121124

122-
variable [LinearOrderedCommGroupWithZero α]
125+
variable [LinearOrderedCommGroupWithZero α] {m n : ℕ}
123126

124127
-- TODO: Do we really need the following two?
125128
/-- Alias of `mul_le_one'` for unification. -/
@@ -285,3 +288,15 @@ instance : LinearOrderedAddCommGroupWithTop (Additive αᵒᵈ) :=
285288
Additive.instNontrivial with
286289
neg_top := @inv_zero _ (_)
287290
add_neg_cancel := fun a ha ↦ mul_inv_cancel (G₀ := α) (id ha : Additive.toMul a ≠ 0) }
291+
292+
lemma pow_lt_pow_succ (ha : 1 < a) : a ^ n < a ^ n.succ := by
293+
rw [← one_mul (a ^ n), pow_succ];
294+
exact mul_lt_right₀ _ ha (pow_ne_zero _ (zero_lt_one.trans ha).ne')
295+
#align pow_lt_pow_succ pow_lt_pow_succ
296+
297+
lemma pow_lt_pow_right₀ (ha : 1 < a) (hmn : m < n) : a ^ m < a ^ n := by
298+
induction' hmn with n _ ih; exacts [pow_lt_pow_succ ha, lt_trans ih (pow_lt_pow_succ ha)]
299+
#align pow_lt_pow₀ pow_lt_pow_right₀
300+
301+
-- 2023-12-23
302+
@[deprecated] alias pow_lt_pow₀ := pow_lt_pow_right₀

Mathlib/Data/Ordmap/Ordset.lean

Lines changed: 11 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -3,12 +3,12 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Mario Carneiro
55
-/
6-
import Mathlib.Data.Ordmap.Ordnode
76
import Mathlib.Algebra.Order.Ring.Defs
8-
import Mathlib.Data.Nat.Dist
97
import Mathlib.Data.Int.Basic
10-
import Mathlib.Tactic.Linarith
8+
import Mathlib.Data.Nat.Dist
9+
import Mathlib.Data.Ordmap.Ordnode
1110
import Mathlib.Tactic.Abel
11+
import Mathlib.Tactic.Linarith
1212

1313
#align_import data.ordmap.ordset from "leanprover-community/mathlib"@"47b51515e69f59bca5cf34ef456e6000fe205a69"
1414

@@ -680,15 +680,15 @@ theorem balance_eq_balance' {l x r} (hl : Balanced l) (hr : Balanced r) (sl : Si
680680
· rfl
681681
· have : size rrl = 0 ∧ size rrr = 0 := by
682682
have := balancedSz_zero.1 hr.1.symm
683-
rwa [size, sr.2.2.1, Nat.succ_le_succ_iff, le_zero_iff, add_eq_zero_iff] at this
683+
rwa [size, sr.2.2.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero_iff] at this
684684
cases sr.2.2.2.1.size_eq_zero.1 this.1
685685
cases sr.2.2.2.2.size_eq_zero.1 this.2
686686
obtain rfl : rrs = 1 := sr.2.2.1
687687
rw [if_neg, if_pos, rotateL, if_pos]; · rfl
688688
all_goals dsimp only [size]; decide
689689
· have : size rll = 0 ∧ size rlr = 0 := by
690690
have := balancedSz_zero.1 hr.1
691-
rwa [size, sr.2.1.1, Nat.succ_le_succ_iff, le_zero_iff, add_eq_zero_iff] at this
691+
rwa [size, sr.2.1.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero_iff] at this
692692
cases sr.2.1.2.1.size_eq_zero.1 this.1
693693
cases sr.2.1.2.2.size_eq_zero.1 this.2
694694
obtain rfl : rls = 1 := sr.2.1.1
@@ -707,15 +707,15 @@ theorem balance_eq_balance' {l x r} (hl : Balanced l) (hr : Balanced r) (sl : Si
707707
· rfl
708708
· have : size lrl = 0 ∧ size lrr = 0 := by
709709
have := balancedSz_zero.1 hl.1.symm
710-
rwa [size, sl.2.2.1, Nat.succ_le_succ_iff, le_zero_iff, add_eq_zero_iff] at this
710+
rwa [size, sl.2.2.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero_iff] at this
711711
cases sl.2.2.2.1.size_eq_zero.1 this.1
712712
cases sl.2.2.2.2.size_eq_zero.1 this.2
713713
obtain rfl : lrs = 1 := sl.2.2.1
714714
rw [if_neg, if_neg, if_pos, rotateR, if_neg]; · rfl
715715
all_goals dsimp only [size]; decide
716716
· have : size lll = 0 ∧ size llr = 0 := by
717717
have := balancedSz_zero.1 hl.1
718-
rwa [size, sl.2.1.1, Nat.succ_le_succ_iff, le_zero_iff, add_eq_zero_iff] at this
718+
rwa [size, sl.2.1.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero_iff] at this
719719
cases sl.2.1.2.1.size_eq_zero.1 this.1
720720
cases sl.2.1.2.2.size_eq_zero.1 this.2
721721
obtain rfl : lls = 1 := sl.2.1.1
@@ -765,7 +765,7 @@ theorem balanceL_eq_balance {l x r} (sl : Sized l) (sr : Sized r) (H1 : size l =
765765
· cases' l with ls ll lx lr
766766
· have : size rl = 0 ∧ size rr = 0 := by
767767
have := H1 rfl
768-
rwa [size, sr.1, Nat.succ_le_succ_iff, le_zero_iff, add_eq_zero_iff] at this
768+
rwa [size, sr.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero_iff] at this
769769
cases sr.2.1.size_eq_zero.1 this.1
770770
cases sr.2.2.size_eq_zero.1 this.2
771771
rw [sr.eq_node']; rfl
@@ -1195,7 +1195,7 @@ theorem Valid'.node4L {l} {x : α} {m} {y : α} {r o₁ o₂} (hl : Valid' o₁
11951195
· rw [Nat.succ_add] at mm; rcases mm with (_ | ⟨⟨⟩⟩)
11961196
rcases hm.3.1.resolve_left mm with ⟨mm₁, mm₂⟩
11971197
rcases Nat.eq_zero_or_pos (size ml) with ml0 | ml0
1198-
· rw [ml0, mul_zero, le_zero_iff] at mm₂
1198+
· rw [ml0, mul_zero, Nat.le_zero] at mm₂
11991199
rw [ml0, mm₂] at mm; cases mm (by decide)
12001200
have : 2 * size l ≤ size ml + size mr + 1 := by
12011201
have := Nat.mul_le_mul_left ratio lr₁
@@ -1274,9 +1274,9 @@ theorem Valid'.rotateL {l} {x : α} {r o₁ o₂} (hl : Valid' o₁ l x) (hr : V
12741274
le_trans hb₂
12751275
(Nat.mul_le_mul_left _ <| le_trans (Nat.le_add_left _ _) (Nat.le_add_right _ _))
12761276
· rcases Nat.eq_zero_or_pos (size rl) with rl0 | rl0
1277-
· rw [rl0, not_lt, le_zero_iff, Nat.mul_eq_zero] at h
1277+
· rw [rl0, not_lt, Nat.le_zero, Nat.mul_eq_zero] at h
12781278
replace h := h.resolve_left (by decide)
1279-
erw [rl0, h, le_zero_iff, Nat.mul_eq_zero] at H2
1279+
erw [rl0, h, Nat.le_zero, Nat.mul_eq_zero] at H2
12801280
rw [hr.2.size_eq, rl0, h, H2.resolve_left (by decide)] at H1
12811281
cases H1 (by decide)
12821282
refine' hl.node4L hr.left hr.right rl0 _

Mathlib/Order/RelSeries.lean

Lines changed: 3 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -452,11 +452,9 @@ lemma smash_succ_castAdd {p q : RelSeries r} (h : p.last = q.head)
452452

453453
lemma smash_natAdd {p q : RelSeries r} (h : p.last = q.head) (i : Fin q.length) :
454454
smash p q h (Fin.castSucc <| i.natAdd p.length) = q (Fin.castSucc i) := by
455-
rw [smash_toFun]
456-
split_ifs with H
457-
· simp only [Fin.coe_castSucc, Fin.coe_natAdd, add_lt_iff_neg_left, not_lt_zero'] at H
458-
· congr
459-
exact Nat.add_sub_self_left _ _
455+
rw [smash_toFun, dif_neg (by simp)]
456+
congr
457+
exact Nat.add_sub_self_left _ _
460458

461459
lemma smash_succ_natAdd {p q : RelSeries r} (h : p.last = q.head) (i : Fin q.length) :
462460
smash p q h (i.natAdd p.length).succ = q i.succ := by

0 commit comments

Comments
 (0)