Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - chore: Reduce scope of LinearOrderedCommGroupWithZero #11716

Closed
wants to merge 17 commits into from
Closed
5 changes: 2 additions & 3 deletions Archive/Imo/Imo1988Q6.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,7 +3,6 @@ Copyright (c) 2019 Johan Commelin. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johan Commelin
-/
import Mathlib.Algebra.Order.WithZero
import Mathlib.Data.Nat.Prime
import Mathlib.Data.Rat.Defs
import Mathlib.Order.WellFounded
Expand Down Expand Up @@ -220,7 +219,7 @@ theorem imo1988_q6 {a b : ℕ} (h : a * b + 1 ∣ a ^ 2 + b ^ 2) :
· -- Show that the claim is true if a = b.
intro x hx
suffices k ≤ 1 by
rw [Nat.le_add_one_iff, le_zero_iff] at this
rw [Nat.le_add_one_iff, Nat.le_zero] at this
rcases this with (rfl | rfl)
· use 0; simp
· use 1; simp
Expand Down Expand Up @@ -298,7 +297,7 @@ example {a b : ℕ} (h : a * b ∣ a ^ 2 + b ^ 2 + 1) : 3 * a * b = a ^ 2 + b ^
assumption_mod_cast
· -- Show the base case.
intro x y h h_base
obtain rfl | rfl : x = 0 ∨ x = 1 := by rwa [Nat.le_add_one_iff, le_zero_iff] at h_base
obtain rfl | rfl : x = 0 ∨ x = 1 := by rwa [Nat.le_add_one_iff, Nat.le_zero] at h_base
· simp at h
· rw [mul_one, one_mul, add_right_comm] at h
have y_dvd : y ∣ y * k := dvd_mul_right y k
Expand Down
2 changes: 1 addition & 1 deletion Archive/Imo/Imo2013Q5.lean
Original file line number Diff line number Diff line change
Expand Up @@ -163,7 +163,7 @@ theorem fixed_point_of_gt_1 {f : ℚ → ℝ} {x : ℚ} (hx : 1 < x)
(x : ℝ) + (a ^ N - x : ℚ) ≤ f x + (a ^ N - x : ℚ) := by gcongr; exact H5 x hx
_ ≤ f x + f (a ^ N - x) := by gcongr; exact H5 _ h_big_enough
have hxp : 0 < x := by positivity
have hNp : 0 < N := by by_contra! H; rw [le_zero_iff.mp H] at hN; linarith
have hNp : 0 < N := by by_contra! H; rw [Nat.le_zero.mp H] at hN; linarith
have h2 :=
calc
f x + f (a ^ N - x) ≤ f (x + (a ^ N - x)) := H2 x (a ^ N - x) hxp (by positivity)
Expand Down
2 changes: 1 addition & 1 deletion Counterexamples/LinearOrderWithPosMulPosEqZero.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,7 +5,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johan Commelin, Damiano Testa, Kevin Buzzard
-/
import Mathlib.Algebra.Order.Monoid.Defs
import Mathlib.Algebra.Order.Monoid.WithZero.Defs
import Mathlib.Algebra.Order.Monoid.WithZero

#align_import linear_order_with_pos_mul_pos_eq_zero from "leanprover-community/mathlib"@"328375597f2c0dd00522d9c2e5a33b6a6128feeb"

Expand Down
2 changes: 1 addition & 1 deletion Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean
Original file line number Diff line number Diff line change
Expand Up @@ -92,7 +92,7 @@ theorem zero_divisors_of_torsion {R A} [Nontrivial R] [Ring R] [AddMonoid A] (a
· intro b hb b0
rw [single_pow, one_pow, single_eq_of_ne]
exact nsmul_ne_zero_of_lt_addOrderOf' b0 (Finset.mem_range.mp hb)
· simp only [(zero_lt_two.trans_le o2).ne', Finset.mem_range, not_lt, le_zero_iff,
· simp only [(zero_lt_two.trans_le o2).ne', Finset.mem_range, not_lt, Nat.le_zero,
false_imp_iff]
· rw [single_pow, one_pow, zero_smul, single_eq_same]
· apply_fun fun x : R[A] => x 0
Expand Down
4 changes: 1 addition & 3 deletions Mathlib.lean
Original file line number Diff line number Diff line change
Expand Up @@ -434,8 +434,7 @@ import Mathlib.Algebra.Order.Monoid.ToMulBot
import Mathlib.Algebra.Order.Monoid.TypeTags
import Mathlib.Algebra.Order.Monoid.Units
import Mathlib.Algebra.Order.Monoid.WithTop
import Mathlib.Algebra.Order.Monoid.WithZero.Basic
import Mathlib.Algebra.Order.Monoid.WithZero.Defs
import Mathlib.Algebra.Order.Monoid.WithZero
import Mathlib.Algebra.Order.Monovary
import Mathlib.Algebra.Order.Nonneg.Field
import Mathlib.Algebra.Order.Nonneg.Floor
Expand Down Expand Up @@ -464,7 +463,6 @@ import Mathlib.Algebra.Order.Sub.WithTop
import Mathlib.Algebra.Order.Support
import Mathlib.Algebra.Order.ToIntervalMod
import Mathlib.Algebra.Order.UpperLower
import Mathlib.Algebra.Order.WithZero
import Mathlib.Algebra.Order.ZeroLEOne
import Mathlib.Algebra.PEmptyInstances
import Mathlib.Algebra.PUnitInstances
Expand Down
4 changes: 4 additions & 0 deletions Mathlib/Algebra/Group/WithOne/Defs.lean
Original file line number Diff line number Diff line change
Expand Up @@ -29,6 +29,10 @@ In Lean 3, we use `id` here and there to get correct types of proofs. This is re
`WithOne` and `WithZero` are marked as `Irreducible` at the end of
`Mathlib.Algebra.Group.WithOne.Defs`, so proofs that use `Option α` instead of `WithOne α` no
longer typecheck. In Lean 4, both types are plain `def`s, so we don't need these `id`s.

## TODO

`WithOne.coe_mul` and `WithZero.coe_mul` have inconsistent use of implicit parameters
-/


Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Lie/Nilpotent.lean
Original file line number Diff line number Diff line change
Expand Up @@ -134,7 +134,7 @@ variable (R L M)
theorem antitone_lowerCentralSeries : Antitone <| lowerCentralSeries R L M := by
intro l k
induction' k with k ih generalizing l <;> intro h
· exact (le_zero_iff.mp h).symm ▸ le_rfl
· exact (Nat.le_zero.mp h).symm ▸ le_rfl
· rcases Nat.of_le_succ h with (hk | hk)
· rw [lowerCentralSeries_succ]
exact (LieSubmodule.mono_lie_right _ _ ⊤ (ih hk)).trans (LieSubmodule.lie_le_right _ _)
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Lie/Solvable.lean
Original file line number Diff line number Diff line change
Expand Up @@ -89,7 +89,7 @@ theorem derivedSeriesOfIdeal_add (k l : ℕ) : D (k + l) I = D k (D l I) := by
theorem derivedSeriesOfIdeal_le {I J : LieIdeal R L} {k l : ℕ} (h₁ : I ≤ J) (h₂ : l ≤ k) :
D k I ≤ D l J := by
revert l; induction' k with k ih <;> intro l h₂
· rw [Nat.zero_eq, le_zero_iff] at h₂; rw [h₂, derivedSeriesOfIdeal_zero]; exact h₁
· rw [Nat.zero_eq, Nat.le_zero] at h₂; rw [h₂, derivedSeriesOfIdeal_zero]; exact h₁
· have h : l = k.succ ∨ l ≤ k := by rwa [le_iff_eq_or_lt, Nat.lt_succ_iff] at h₂
cases' h with h h
· rw [h, derivedSeriesOfIdeal_succ, derivedSeriesOfIdeal_succ]
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Order/Field/Canonical/Defs.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,7 +5,7 @@ Authors: Robert Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn
-/
import Mathlib.Algebra.Order.Field.Defs
import Mathlib.Algebra.Order.Ring.Canonical
import Mathlib.Algebra.Order.WithZero
import Mathlib.Algebra.Order.Monoid.WithZero

#align_import algebra.order.field.canonical.defs from "leanprover-community/mathlib"@"fc2ed6f838ce7c9b7c7171e58d78eaf7b438fb0e"

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Order/Hom/Monoid.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,7 +5,7 @@ Authors: Yaël Dillies
-/
import Mathlib.Algebra.GroupWithZero.Hom
import Mathlib.Algebra.Order.Group.Instances
import Mathlib.Algebra.Order.Monoid.WithZero.Defs
import Mathlib.Algebra.Order.Monoid.WithZero
import Mathlib.Order.Hom.Basic

#align_import algebra.order.hom.monoid from "leanprover-community/mathlib"@"3342d1b2178381196f818146ff79bc0e7ccd9e2d"
Expand Down
4 changes: 1 addition & 3 deletions Mathlib/Algebra/Order/Kleene.lean
Original file line number Diff line number Diff line change
Expand Up @@ -249,9 +249,7 @@ theorem kstar_eq_one : a∗ = 1 ↔ a ≤ 1 :=
fun h ↦ one_le_kstar.antisymm' <| kstar_le_of_mul_le_left le_rfl <| by rwa [one_mul]⟩
#align kstar_eq_one kstar_eq_one

@[simp]
theorem kstar_zero : (0 : α)∗ = 1 :=
kstar_eq_one.2 zero_le_one
@[simp] lemma kstar_zero : (0 : α)∗ = 1 := kstar_eq_one.2 (zero_le _)
#align kstar_zero kstar_zero

@[simp]
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Order/Monoid/ToMulBot.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,7 +3,7 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl
-/
import Mathlib.Algebra.Order.Monoid.WithZero.Defs
import Mathlib.Algebra.Order.Monoid.WithZero
import Mathlib.Algebra.Order.Monoid.TypeTags
import Mathlib.Algebra.Order.Monoid.WithTop
import Mathlib.Algebra.Group.Equiv.Basic
Expand Down
Loading
Loading