Skip to content

Commit

Permalink
chore: refactor of Algebra/Group/Defs to reduce imports (#9606)
Browse files Browse the repository at this point in the history
We are not that far from the point that `Algebra/Group/Defs` depends on nothing significant besides `simps` and `to_additive`.

This removes from `Mathlib.Algebra.Group.Defs` the dependencies on
* `Mathlib.Tactic.Basic` (which is a grab-bag of random stuff)
* `Mathlib.Init.Algebra.Classes` (which is ancient and half-baked)
* `Mathlib.Logic.Function.Basic` (not particularly important, but it is barely used in this file)

The goal is to avoid all unnecessary imports to set up the *definitions* of basic algebraic structures. 

We also separate out `Mathlib.Tactic.TypeStar` and `Mathlib.Tactic.Lemma` as prerequisites to `Mathlib.Tactic.Basic`, but which can be imported separately when the rest of `Mathlib.Tactic.Basic` is not needed.



Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
Co-authored-by: Yaël Dillies <yael.dillies@gmail.com>
  • Loading branch information
3 people committed Jan 17, 2024
1 parent b075cdd commit 45df823
Show file tree
Hide file tree
Showing 25 changed files with 186 additions and 108 deletions.
2 changes: 2 additions & 0 deletions Mathlib.lean
Expand Up @@ -3364,6 +3364,7 @@ import Mathlib.Tactic.InferParam
import Mathlib.Tactic.Inhabit
import Mathlib.Tactic.IntervalCases
import Mathlib.Tactic.IrreducibleDef
import Mathlib.Tactic.Lemma
import Mathlib.Tactic.LibrarySearch
import Mathlib.Tactic.Lift
import Mathlib.Tactic.LiftLets
Expand Down Expand Up @@ -3465,6 +3466,7 @@ import Mathlib.Tactic.ToLevel
import Mathlib.Tactic.Trace
import Mathlib.Tactic.TryThis
import Mathlib.Tactic.TypeCheck
import Mathlib.Tactic.TypeStar
import Mathlib.Tactic.UnsetOption
import Mathlib.Tactic.Use
import Mathlib.Tactic.Variable
Expand Down
1 change: 1 addition & 0 deletions Mathlib/Algebra/CharZero/Defs.lean
Expand Up @@ -7,6 +7,7 @@ import Mathlib.Init.Data.Nat.Lemmas
import Mathlib.Data.Int.Cast.Defs
import Mathlib.Tactic.Cases
import Mathlib.Algebra.NeZero
import Mathlib.Logic.Function.Basic

#align_import algebra.char_zero.defs from "leanprover-community/mathlib"@"d6aae1bcbd04b8de2022b9b83a5b5b10e10c777d"

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/CovariantAndContravariant.lean
Expand Up @@ -3,7 +3,7 @@ Copyright (c) 2021 Damiano Testa. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Damiano Testa
-/
import Mathlib.Algebra.Group.Defs
import Mathlib.Algebra.Group.Basic
import Mathlib.Order.Basic
import Mathlib.Order.Monotone.Basic

Expand Down
62 changes: 60 additions & 2 deletions Mathlib/Algebra/Group/Basic.lean
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro
-/
import Mathlib.Algebra.Group.Defs
import Mathlib.Logic.Function.Basic
import Mathlib.Tactic.Cases
import Mathlib.Tactic.SimpRw
import Mathlib.Tactic.SplitIfs
Expand All @@ -25,14 +26,66 @@ universe u

variable {α β G : Type*}

section IsLeftCancelMul

variable [Mul G] [IsLeftCancelMul G]

@[to_additive]
theorem mul_right_injective (a : G) : Injective (a * ·) := fun _ _ ↦ mul_left_cancel
#align mul_right_injective mul_right_injective
#align add_right_injective add_right_injective

@[to_additive (attr := simp)]
theorem mul_right_inj (a : G) {b c : G} : a * b = a * c ↔ b = c :=
(mul_right_injective a).eq_iff
#align mul_right_inj mul_right_inj
#align add_right_inj add_right_inj

@[to_additive]
theorem mul_ne_mul_right (a : G) {b c : G} : a * b ≠ a * c ↔ b ≠ c :=
(mul_right_injective a).ne_iff
#align mul_ne_mul_right mul_ne_mul_right
#align add_ne_add_right add_ne_add_right

end IsLeftCancelMul

section IsRightCancelMul

variable [Mul G] [IsRightCancelMul G]

@[to_additive]
theorem mul_left_injective (a : G) : Function.Injective (· * a) := fun _ _ ↦ mul_right_cancel
#align mul_left_injective mul_left_injective
#align add_left_injective add_left_injective

@[to_additive (attr := simp)]
theorem mul_left_inj (a : G) {b c : G} : b * a = c * a ↔ b = c :=
(mul_left_injective a).eq_iff
#align mul_left_inj mul_left_inj
#align add_left_inj add_left_inj

@[to_additive]
theorem mul_ne_mul_left (a : G) {b c : G} : b * a ≠ c * a ↔ b ≠ c :=
(mul_left_injective a).ne_iff
#align mul_ne_mul_left mul_ne_mul_left
#align add_ne_add_left add_ne_add_left

end IsRightCancelMul

section Semigroup
variable [Semigroup α]

@[to_additive]
instance Semigroup.to_isAssociative : IsAssociative α (· * ·) := ⟨mul_assoc⟩
#align semigroup.to_is_associative Semigroup.to_isAssociative
#align add_semigroup.to_is_associative AddSemigroup.to_isAssociative

/-- Composing two multiplications on the left by `y` then `x`
is equal to a multiplication on the left by `x * y`.
-/
@[to_additive (attr := simp) "Composing two additions on the left by `y` then `x`
is equal to an addition on the left by `x + y`."]
theorem comp_mul_left [Semigroup α] (x y : α) : (x * ·) ∘ (y * ·) = (x * y * ·) := by
theorem comp_mul_left (x y : α) : (x * ·) ∘ (y * ·) = (x * y * ·) := by
ext z
simp [mul_assoc]
#align comp_mul_left comp_mul_left
Expand All @@ -43,14 +96,19 @@ is equal to a multiplication on the right by `y * x`.
-/
@[to_additive (attr := simp) "Composing two additions on the right by `y` and `x`
is equal to an addition on the right by `y + x`."]
theorem comp_mul_right [Semigroup α] (x y : α) : (· * x) ∘ (· * y) = (· * (y * x)) := by
theorem comp_mul_right (x y : α) : (· * x) ∘ (· * y) = (· * (y * x)) := by
ext z
simp [mul_assoc]
#align comp_mul_right comp_mul_right
#align comp_add_right comp_add_right

end Semigroup

@[to_additive]
instance CommMagma.to_isCommutative [CommMagma G] : IsCommutative G (· * ·) := ⟨mul_comm⟩
#align comm_semigroup.to_is_commutative CommMagma.to_isCommutative
#align add_comm_semigroup.to_is_commutative AddCommMagma.to_isCommutative

section MulOneClass

variable {M : Type u} [MulOneClass M]
Expand Down
1 change: 1 addition & 0 deletions Mathlib/Algebra/Group/Commute/Defs.lean
Expand Up @@ -5,6 +5,7 @@ Authors: Neil Strickland, Yury Kudryashov
-/
import Mathlib.Algebra.Group.Semiconj.Defs
import Mathlib.Data.Nat.Defs
import Mathlib.Init.Algebra.Classes

#align_import algebra.group.commute from "leanprover-community/mathlib"@"05101c3df9d9cfe9430edc205860c79b6d660102"

Expand Down
74 changes: 11 additions & 63 deletions Mathlib/Algebra/Group/Defs.lean
Expand Up @@ -3,11 +3,13 @@ Copyright (c) 2014 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro
-/
import Mathlib.Init.ZeroOne
import Mathlib.Init.Data.Int.Basic
import Mathlib.Logic.Function.Basic
import Mathlib.Tactic.Basic
import Mathlib.Init.ZeroOne
import Mathlib.Tactic.Lemma
import Mathlib.Tactic.TypeStar
import Mathlib.Tactic.Simps.Basic
import Mathlib.Tactic.ToAdditive
import Mathlib.Util.AssertExists

#align_import algebra.group.defs from "leanprover-community/mathlib"@"48fb5b5280e7c81672afc9524185ae994553ebf4"

Expand Down Expand Up @@ -221,27 +223,10 @@ theorem mul_left_cancel : a * b = a * c → b = c :=

@[to_additive]
theorem mul_left_cancel_iff : a * b = a * c ↔ b = c :=
⟨mul_left_cancel, congr_arg _⟩
⟨mul_left_cancel, congrArg _⟩
#align mul_left_cancel_iff mul_left_cancel_iff
#align add_left_cancel_iff add_left_cancel_iff

@[to_additive]
theorem mul_right_injective (a : G) : Injective (a * ·) := fun _ _ ↦ mul_left_cancel
#align mul_right_injective mul_right_injective
#align add_right_injective add_right_injective

@[to_additive (attr := simp)]
theorem mul_right_inj (a : G) {b c : G} : a * b = a * c ↔ b = c :=
(mul_right_injective a).eq_iff
#align mul_right_inj mul_right_inj
#align add_right_inj add_right_inj

@[to_additive]
theorem mul_ne_mul_right (a : G) {b c : G} : a * b ≠ a * c ↔ b ≠ c :=
(mul_right_injective a).ne_iff
#align mul_ne_mul_right mul_ne_mul_right
#align add_ne_add_right add_ne_add_right

end IsLeftCancelMul

section IsRightCancelMul
Expand All @@ -256,27 +241,10 @@ theorem mul_right_cancel : a * b = c * b → a = c :=

@[to_additive]
theorem mul_right_cancel_iff : b * a = c * a ↔ b = c :=
⟨mul_right_cancel, congr_arg (· * a)⟩
⟨mul_right_cancel, congrArg (· * a)⟩
#align mul_right_cancel_iff mul_right_cancel_iff
#align add_right_cancel_iff add_right_cancel_iff

@[to_additive]
theorem mul_left_injective (a : G) : Function.Injective (· * a) := fun _ _ ↦ mul_right_cancel
#align mul_left_injective mul_left_injective
#align add_left_injective add_left_injective

@[to_additive (attr := simp)]
theorem mul_left_inj (a : G) {b c : G} : b * a = c * a ↔ b = c :=
(mul_left_injective a).eq_iff
#align mul_left_inj mul_left_inj
#align add_left_inj add_left_inj

@[to_additive]
theorem mul_ne_mul_left (a : G) {b c : G} : b * a ≠ c * a ↔ b ≠ c :=
(mul_left_injective a).ne_iff
#align mul_ne_mul_left mul_ne_mul_left
#align add_ne_add_left add_ne_add_left

end IsRightCancelMul

end Mul
Expand Down Expand Up @@ -311,12 +279,6 @@ theorem mul_assoc : ∀ a b c : G, a * b * c = a * (b * c) :=
#align mul_assoc mul_assoc
#align add_assoc add_assoc

@[to_additive]
instance Semigroup.to_isAssociative : IsAssociative G (· * ·) :=
⟨mul_assoc⟩
#align semigroup.to_is_associative Semigroup.to_isAssociative
#align add_semigroup.to_is_associative AddSemigroup.to_isAssociative

end Semigroup

/-- A commutative addition is a type with an addition which commutes-/
Expand Down Expand Up @@ -359,11 +321,6 @@ theorem mul_comm : ∀ a b : G, a * b = b * a := CommMagma.mul_comm
#align mul_comm mul_comm
#align add_comm add_comm

@[to_additive]
instance CommMagma.to_isCommutative : IsCommutative G (· * ·) := ⟨mul_comm⟩
#align comm_semigroup.to_is_commutative CommMagma.to_isCommutative
#align add_comm_semigroup.to_is_commutative AddCommMagma.to_isCommutative

/-- Any `CommMagma G` that satisfies `IsRightCancelMul G` also satisfies `IsLeftCancelMul G`. -/
@[to_additive AddCommMagma.IsRightCancelAdd.toIsLeftCancelAdd "Any `AddCommMagma G` that satisfies
`IsRightCancelAdd G` also satisfies `IsLeftCancelAdd G`."]
Expand Down Expand Up @@ -993,7 +950,7 @@ theorem zpow_ofNat (a : G) : ∀ n : ℕ, a ^ (n : ℤ) = a ^ n
| 0 => (zpow_zero _).trans (pow_zero _).symm
| n + 1 => calc
a ^ (↑(n + 1) : ℤ) = a * a ^ (n : ℤ) := DivInvMonoid.zpow_succ' _ _
_ = a * a ^ n := congr_arg (a * ·) (zpow_ofNat a n)
_ = a * a ^ n := congrArg (a * ·) (zpow_ofNat a n)
_ = a ^ (n + 1) := (pow_succ _ _).symm
#align zpow_coe_nat zpow_ofNat
#align zpow_of_nat zpow_ofNat
Expand Down Expand Up @@ -1231,12 +1188,6 @@ instance (priority := 100) Group.toCancelMonoid : CancelMonoid G :=

end Group

@[to_additive]
theorem Group.toDivInvMonoid_injective {G : Type*} :
Function.Injective (@Group.toDivInvMonoid G) := by rintro ⟨⟩ ⟨⟩ ⟨⟩; rfl
#align group.to_div_inv_monoid_injective Group.toDivInvMonoid_injective
#align add_group.to_sub_neg_add_monoid_injective AddGroup.toSubNegAddMonoid_injective

/-- An additive commutative group is an additive group with commutative `(+)`. -/
class AddCommGroup (G : Type u) extends AddGroup G, AddCommMonoid G
#align add_comm_group AddCommGroup
Expand All @@ -1248,12 +1199,6 @@ class CommGroup (G : Type u) extends Group G, CommMonoid G

attribute [to_additive existing] CommGroup.toCommMonoid

@[to_additive]
theorem CommGroup.toGroup_injective {G : Type u} : Function.Injective (@CommGroup.toGroup G) := by
rintro ⟨⟩ ⟨⟩ ⟨⟩; rfl
#align comm_group.to_group_injective CommGroup.toGroup_injective
#align add_comm_group.to_add_group_injective AddCommGroup.toAddGroup_injective

section CommGroup

variable [CommGroup G]
Expand Down Expand Up @@ -1318,3 +1263,6 @@ initialize_simps_projections Group
initialize_simps_projections AddGroup
initialize_simps_projections CommGroup
initialize_simps_projections AddCommGroup

assert_not_exists Function.Injective
assert_not_exists IsCommutative
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Group/Embedding.lean
Expand Up @@ -3,7 +3,7 @@ Copyright (c) 2021 Damiano Testa. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Damiano Testa
-/
import Mathlib.Algebra.Group.Defs
import Mathlib.Algebra.Group.Basic
import Mathlib.Logic.Embedding.Basic

#align_import algebra.hom.embedding from "leanprover-community/mathlib"@"70d50ecfd4900dd6d328da39ab7ebd516abe4025"
Expand Down
14 changes: 14 additions & 0 deletions Mathlib/Algebra/Group/Ext.lean
Expand Up @@ -27,6 +27,8 @@ former uses `HMul.hMul` which is the canonical spelling.
monoid, group, extensionality
-/

open Function

universe u

@[to_additive (attr := ext)]
Expand Down Expand Up @@ -152,6 +154,12 @@ theorem DivInvMonoid.ext {M : Type*} ⦃m₁ m₂ : DivInvMonoid M⦄
#align div_inv_monoid.ext DivInvMonoid.ext
#align sub_neg_monoid.ext SubNegMonoid.ext

@[to_additive]
lemma Group.toDivInvMonoid_injective {G : Type*} : Injective (@Group.toDivInvMonoid G) := by
rintro ⟨⟩ ⟨⟩ ⟨⟩; rfl
#align group.to_div_inv_monoid_injective Group.toDivInvMonoid_injective
#align add_group.to_sub_neg_add_monoid_injective AddGroup.toSubNegAddMonoid_injective

@[to_additive (attr := ext)]
theorem Group.ext {G : Type*} ⦃g₁ g₂ : Group G⦄ (h_mul : g₁.mul = g₂.mul) : g₁ = g₂ := by
have h₁ : g₁.one = g₂.one := congr_arg (·.one) (Monoid.ext h_mul)
Expand All @@ -165,6 +173,12 @@ theorem Group.ext {G : Type*} ⦃g₁ g₂ : Group G⦄ (h_mul : g₁.mul = g₂
#align group.ext Group.ext
#align add_group.ext AddGroup.ext

@[to_additive]
lemma CommGroup.toGroup_injective {G : Type*} : Injective (@CommGroup.toGroup G) := by
rintro ⟨⟩ ⟨⟩ ⟨⟩; rfl
#align comm_group.to_group_injective CommGroup.toGroup_injective
#align add_comm_group.to_add_group_injective AddCommGroup.toAddGroup_injective

@[to_additive (attr := ext)]
theorem CommGroup.ext {G : Type*} ⦃g₁ g₂ : CommGroup G⦄ (h_mul : g₁.mul = g₂.mul) : g₁ = g₂ :=
CommGroup.toGroup_injective <| Group.ext h_mul
Expand Down
1 change: 1 addition & 0 deletions Mathlib/Algebra/Group/Semiconj/Defs.lean
Expand Up @@ -6,6 +6,7 @@ Authors: Yury Kudryashov
Some proofs and docs came from `algebra/commute` (c) Neil Strickland
-/
import Mathlib.Algebra.Group.Defs
import Mathlib.Init.Logic
import Mathlib.Tactic.Cases

#align_import algebra.group.semiconj from "leanprover-community/mathlib"@"a148d797a1094ab554ad4183a4ad6f130358ef64"
Expand Down
36 changes: 25 additions & 11 deletions Mathlib/Algebra/Group/WithOne/Defs.lean
Expand Up @@ -3,8 +3,8 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Mario Carneiro, Johan Commelin
-/
import Mathlib.Order.WithBot
import Mathlib.Algebra.Ring.Defs
import Mathlib.Order.WithBot

#align_import algebra.group.with_one.defs from "leanprover-community/mathlib"@"995b47e555f1b6297c7cf16855f1023e355219fb"

Expand Down Expand Up @@ -187,20 +187,28 @@ instance mulOneClass [Mul α] : MulOneClass (WithOne α) where
one_mul := (Option.liftOrGet_isLeftId _).1
mul_one := (Option.liftOrGet_isRightId _).1

@[to_additive]
instance monoid [Semigroup α] : Monoid (WithOne α) :=
{ WithOne.mulOneClass with mul_assoc := (Option.liftOrGet_isAssociative _).1 }

@[to_additive]
instance commMonoid [CommSemigroup α] : CommMonoid (WithOne α) :=
{ WithOne.monoid with mul_comm := (Option.liftOrGet_isCommutative _).1 }

@[to_additive (attr := simp, norm_cast)]
theorem coe_mul [Mul α] (a b : α) : ((a * b : α) : WithOne α) = a * b :=
rfl
lemma coe_mul [Mul α] (a b : α) : (↑(a * b) : WithOne α) = a * b := rfl
#align with_one.coe_mul WithOne.coe_mul
#align with_zero.coe_add WithZero.coe_add

@[to_additive]
instance monoid [Semigroup α] : Monoid (WithOne α) where
__ := mulOneClass
mul_assoc a b c := match a, b, c with
| 1, b, c => by simp
| (a : α), 1, c => by simp
| (a : α), (b : α), 1 => by simp
| (a : α), (b : α), (c : α) => by simp_rw [← coe_mul, mul_assoc]

@[to_additive]
instance commMonoid [CommSemigroup α] : CommMonoid (WithOne α) where
mul_comm := fun a b => match a, b with
| (a : α), (b : α) => congr_arg some (mul_comm a b)
| (_ : α), 1 => rfl
| 1, (_ : α) => rfl
| 1, 1 => rfl

@[to_additive (attr := simp, norm_cast)]
theorem coe_inv [Inv α] (a : α) : ((a⁻¹ : α) : WithOne α) = (a : WithOne α)⁻¹ :=
rfl
Expand Down Expand Up @@ -392,3 +400,9 @@ instance semiring [Semiring α] : Semiring (WithZero α) :=
WithZero.monoidWithZero, WithZero.instDistrib with }

end WithZero

-- Check that we haven't needed to import all the basic lemmas about groups,
-- by asserting a random sample don't exist here:
assert_not_exists inv_involutive
assert_not_exists div_right_inj
assert_not_exists pow_ite
1 change: 1 addition & 0 deletions Mathlib/Algebra/GroupWithZero/Defs.lean
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johan Commelin
-/
import Mathlib.Algebra.Group.Defs
import Mathlib.Logic.Function.Basic
import Mathlib.Logic.Nontrivial.Defs

#align_import algebra.group_with_zero.defs from "leanprover-community/mathlib"@"2f3994e1b117b1e1da49bcfb67334f33460c3ce4"
Expand Down

0 comments on commit 45df823

Please sign in to comment.