From 47b51515e69f59bca5cf34ef456e6000fe205a69 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 8 Nov 2022 20:55:33 +0000 Subject: [PATCH] refactor(*): supremum of refactoring PRs #17405 #17418 #17419 #17420 #17421 #17422 #17423 #17427 #17430 (#17424) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yury G. Kudryashov Co-authored-by: Scott Morrison Co-authored-by: Yaël Dillies --- ...linear_order_with_pos_mul_pos_eq_zero.lean | 2 +- src/algebra/associated.lean | 2 +- src/algebra/big_operators/basic.lean | 2 +- .../Group/equivalence_Group_AddGroup.lean | 1 + src/algebra/category/Semigroup/basic.lean | 2 +- src/algebra/divisibility.lean | 325 ------------ src/algebra/divisibility/basic.lean | 150 ++++++ src/algebra/divisibility/units.lean | 114 +++++ .../basic.lean} | 226 +-------- src/algebra/euclidean_domain/defs.lean | 221 ++++++++ src/algebra/euclidean_domain/instances.lean | 56 +++ src/algebra/group/opposite.lean | 2 +- src/algebra/group/prod.lean | 1 + src/algebra/group/ulift.lean | 3 +- src/algebra/group/with_one.lean | 6 +- src/algebra/group_power/basic.lean | 3 +- src/algebra/group_with_zero/divisibility.lean | 87 ++++ src/algebra/hom/commute.lean | 28 ++ .../hom/{equiv.lean => equiv/basic.lean} | 242 +-------- src/algebra/hom/equiv/type_tags.lean | 71 +++ src/algebra/hom/equiv/units.lean | 192 +++++++ src/algebra/hom/group.lean | 37 +- src/algebra/hom/ring.lean | 2 +- src/algebra/hom/units.lean | 25 +- .../order/euclidean_absolute_value.lean | 2 +- src/algebra/order/field/basic.lean | 1 + src/algebra/order/group/abs.lean | 4 +- src/algebra/order/group/bounds.lean | 2 +- .../order/group/{basic.lean => defs.lean} | 38 +- src/algebra/order/group/densely_ordered.lean | 4 +- src/algebra/order/group/inj_surj.lean | 47 ++ src/algebra/order/group/prod.lean | 2 +- src/algebra/order/group/type_tags.lean | 2 +- src/algebra/order/group/units.lean | 2 +- src/algebra/order/group/with_top.lean | 2 +- src/algebra/order/hom/ring.lean | 2 +- src/algebra/order/invertible.lean | 2 +- src/algebra/order/monoid/basic.lean | 108 +--- src/algebra/order/monoid/cancel/basic.lean | 54 ++ .../monoid/{cancel.lean => cancel/defs.lean} | 33 +- .../{canonical.lean => canonical/defs.lean} | 4 +- src/algebra/order/monoid/defs.lean | 118 +++++ src/algebra/order/monoid/lemmas.lean | 1 - src/algebra/order/monoid/min_max.lean | 2 +- src/algebra/order/monoid/order_dual.lean | 2 +- src/algebra/order/monoid/prod.lean | 4 +- src/algebra/order/monoid/type_tags.lean | 4 +- src/algebra/order/monoid/units.lean | 4 +- src/algebra/order/monoid/with_zero.lean | 2 +- src/algebra/order/nonneg/ring.lean | 3 +- src/algebra/order/pi.lean | 2 +- src/algebra/order/positive/ring.lean | 2 +- src/algebra/order/ring/abs.lean | 2 +- src/algebra/order/ring/canonical.lean | 2 +- src/algebra/order/ring/cone.lean | 77 +++ .../order/ring/{basic.lean => defs.lean} | 247 +-------- src/algebra/order/ring/inj_surj.lean | 194 +++++++ src/algebra/order/ring/nontrivial.lean | 2 +- src/algebra/order/sub/basic.lean | 338 +------------ src/algebra/order/sub/canonical.lean | 4 +- src/algebra/order/sub/defs.lean | 363 +++++++++++++ src/algebra/order/sub/with_top.lean | 2 +- src/algebra/order/upper_lower.lean | 2 +- src/algebra/order/with_zero.lean | 1 + src/algebra/quandle.lean | 2 +- src/algebra/ring/basic.lean | 475 +----------------- src/algebra/ring/defs.lean | 57 +++ src/algebra/ring/divisibility.lean | 2 +- src/algebra/ring/equiv.lean | 2 - src/algebra/ring/inj_surj.lean | 455 +++++++++++++++++ src/category_theory/conj.lean | 1 + src/category_theory/endomorphism.lean | 2 +- src/category_theory/groupoid/subgroupoid.lean | 2 +- .../groupoid/vertex_group.lean | 2 +- src/data/finset/noncomm_prod.lean | 2 +- src/data/int/basic.lean | 1 + src/data/int/modeq.lean | 1 + src/data/int/order.lean | 1 + src/data/nat/order.lean | 1 + src/data/nat/part_enat.lean | 2 +- src/data/ordmap/ordnode.lean | 1 + src/data/ordmap/ordset.lean | 2 +- src/data/pi/lex.lean | 2 +- src/data/rat/basic.lean | 3 +- src/data/rat/floor.lean | 1 + src/deprecated/group.lean | 2 +- src/group_theory/congruence.lean | 2 +- .../free_abelian_group_finsupp.lean | 1 + src/group_theory/group_action/defs.lean | 1 + src/group_theory/group_action/opposite.lean | 1 + src/group_theory/subgroup/basic.lean | 1 + src/group_theory/submonoid/operations.lean | 3 +- src/order/iterate.lean | 1 + src/order/semiconj_Sup.lean | 2 +- .../dedekind_domain/selmer_group.lean | 2 +- src/ring_theory/principal_ideal_domain.lean | 1 + src/tactic/linarith/lemmas.lean | 2 +- src/tactic/monotonicity/lemmas.lean | 2 +- test/assert_exists/test2.lean | 2 +- test/finish4.lean | 2 +- test/library_search/ordered_ring.lean | 2 +- test/monotonicity.lean | 2 +- test/nontriviality.lean | 2 +- 103 files changed, 2423 insertions(+), 2112 deletions(-) delete mode 100644 src/algebra/divisibility.lean create mode 100644 src/algebra/divisibility/basic.lean create mode 100644 src/algebra/divisibility/units.lean rename src/algebra/{euclidean_domain.lean => euclidean_domain/basic.lean} (51%) create mode 100644 src/algebra/euclidean_domain/defs.lean create mode 100644 src/algebra/euclidean_domain/instances.lean create mode 100644 src/algebra/group_with_zero/divisibility.lean create mode 100644 src/algebra/hom/commute.lean rename src/algebra/hom/{equiv.lean => equiv/basic.lean} (70%) create mode 100644 src/algebra/hom/equiv/type_tags.lean create mode 100644 src/algebra/hom/equiv/units.lean rename src/algebra/order/group/{basic.lean => defs.lean} (94%) create mode 100644 src/algebra/order/group/inj_surj.lean create mode 100644 src/algebra/order/monoid/cancel/basic.lean rename src/algebra/order/monoid/{cancel.lean => cancel/defs.lean} (69%) rename src/algebra/order/monoid/{canonical.lean => canonical/defs.lean} (99%) create mode 100644 src/algebra/order/monoid/defs.lean create mode 100644 src/algebra/order/ring/cone.lean rename src/algebra/order/ring/{basic.lean => defs.lean} (75%) create mode 100644 src/algebra/order/ring/inj_surj.lean create mode 100644 src/algebra/order/sub/defs.lean create mode 100644 src/algebra/ring/inj_surj.lean diff --git a/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean b/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean index 25b019fcc0266..bbb15d9c98c05 100644 --- a/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean +++ b/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean @@ -4,7 +4,7 @@ All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Damiano Testa, Kevin Buzzard -/ -import algebra.order.monoid.basic +import algebra.order.monoid.defs /-! An example of a `linear_ordered_comm_monoid_with_zero` in which the product of two positive diff --git a/src/algebra/associated.lean b/src/algebra/associated.lean index ebb3b8e913410..dcd4752d5302e 100644 --- a/src/algebra/associated.lean +++ b/src/algebra/associated.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Jens Wagemaker -/ -import algebra.divisibility +import algebra.divisibility.basic import algebra.group_power.lemmas import algebra.invertible import order.atoms diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index f1ed09081352f..2679a31759d1e 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -5,7 +5,7 @@ Authors: Johannes Hölzl -/ import algebra.group.pi -import algebra.hom.equiv +import algebra.hom.equiv.basic import algebra.ring.opposite import data.finset.fold import data.fintype.basic diff --git a/src/algebra/category/Group/equivalence_Group_AddGroup.lean b/src/algebra/category/Group/equivalence_Group_AddGroup.lean index a57db0bacd0e5..00d5f4ec48617 100644 --- a/src/algebra/category/Group/equivalence_Group_AddGroup.lean +++ b/src/algebra/category/Group/equivalence_Group_AddGroup.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ import algebra.category.Group.basic +import algebra.hom.equiv.type_tags /-! # Equivalence between `Group` and `AddGroup` diff --git a/src/algebra/category/Semigroup/basic.lean b/src/algebra/category/Semigroup/basic.lean index ad3ebe18aa566..65bb2df05eb96 100644 --- a/src/algebra/category/Semigroup/basic.lean +++ b/src/algebra/category/Semigroup/basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Julian Kuelshammer -/ import algebra.pempty_instances -import algebra.hom.equiv +import algebra.hom.equiv.basic import category_theory.concrete_category.bundled_hom import category_theory.functor.reflects_isomorphisms import category_theory.elementwise diff --git a/src/algebra/divisibility.lean b/src/algebra/divisibility.lean deleted file mode 100644 index 63bcda5bb13cc..0000000000000 --- a/src/algebra/divisibility.lean +++ /dev/null @@ -1,325 +0,0 @@ -/- -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, Floris van Doorn, Amelia Livingston, Yury Kudryashov, -Neil Strickland, Aaron Anderson --/ - -import algebra.hom.group -import algebra.group_with_zero.basic - -/-! -# Divisibility - -This file defines the basics of the divisibility relation in the context of `(comm_)` `monoid`s -`(_with_zero)`. - -## Main definitions - - * `monoid.has_dvd` - -## Implementation notes - -The divisibility relation is defined for all monoids, and as such, depends on the order of - multiplication if the monoid is not commutative. There are two possible conventions for - divisibility in the noncommutative context, and this relation follows the convention for ordinals, - so `a | b` is defined as `∃ c, b = a * c`. - -## Tags - -divisibility, divides --/ - -variables {α : Type*} - -section semigroup - -variables [semigroup α] {a b c : α} - -/-- There are two possible conventions for divisibility, which coincide in a `comm_monoid`. - This matches the convention for ordinals. -/ -@[priority 100] -instance semigroup_has_dvd : has_dvd α := -has_dvd.mk (λ a b, ∃ c, b = a * c) - --- TODO: this used to not have `c` explicit, but that seems to be important --- for use with tactics, similar to `exists.intro` -theorem dvd.intro (c : α) (h : a * c = b) : a ∣ b := -exists.intro c h^.symm - -alias dvd.intro ← dvd_of_mul_right_eq - -theorem exists_eq_mul_right_of_dvd (h : a ∣ b) : ∃ c, b = a * c := h - -theorem dvd.elim {P : Prop} {a b : α} (H₁ : a ∣ b) (H₂ : ∀ c, b = a * c → P) : P := -exists.elim H₁ H₂ - -local attribute [simp] mul_assoc mul_comm mul_left_comm - -@[trans] theorem dvd_trans : a ∣ b → b ∣ c → a ∣ c -| ⟨d, h₁⟩ ⟨e, h₂⟩ := ⟨d * e, h₁ ▸ h₂.trans $ mul_assoc a d e⟩ - -alias dvd_trans ← has_dvd.dvd.trans - -instance : is_trans α (∣) := ⟨λ a b c, dvd_trans⟩ - -@[simp] theorem dvd_mul_right (a b : α) : a ∣ a * b := dvd.intro b rfl - -theorem dvd_mul_of_dvd_left (h : a ∣ b) (c : α) : a ∣ b * c := -h.trans (dvd_mul_right b c) - -alias dvd_mul_of_dvd_left ← has_dvd.dvd.mul_right - -theorem dvd_of_mul_right_dvd (h : a * b ∣ c) : a ∣ c := -(dvd_mul_right a b).trans h - -section map_dvd - -variables {M N : Type*} [monoid M] [monoid N] - -lemma map_dvd {F : Type*} [mul_hom_class F M N] (f : F) {a b} : a ∣ b → f a ∣ f b -| ⟨c, h⟩ := ⟨f c, h.symm ▸ map_mul f a c⟩ - -lemma mul_hom.map_dvd (f : M →ₙ* N) {a b} : a ∣ b → f a ∣ f b := map_dvd f - -lemma monoid_hom.map_dvd (f : M →* N) {a b} : a ∣ b → f a ∣ f b := map_dvd f - -end map_dvd - -end semigroup - -section monoid - -variables [monoid α] - -@[refl, simp] theorem dvd_refl (a : α) : a ∣ a := dvd.intro 1 (mul_one a) -theorem dvd_rfl : ∀ {a : α}, a ∣ a := dvd_refl -instance : is_refl α (∣) := ⟨dvd_refl⟩ - -theorem one_dvd (a : α) : 1 ∣ a := dvd.intro a (one_mul a) - -end monoid - -section comm_semigroup - -variables [comm_semigroup α] {a b c : α} - -theorem dvd.intro_left (c : α) (h : c * a = b) : a ∣ b := -dvd.intro _ (begin rewrite mul_comm at h, apply h end) - -alias dvd.intro_left ← dvd_of_mul_left_eq - -theorem exists_eq_mul_left_of_dvd (h : a ∣ b) : ∃ c, b = c * a := -dvd.elim h (assume c, assume H1 : b = a * c, exists.intro c (eq.trans H1 (mul_comm a c))) - -lemma dvd_iff_exists_eq_mul_left : a ∣ b ↔ ∃ c, b = c * a := -⟨exists_eq_mul_left_of_dvd, by { rintro ⟨c, rfl⟩, exact ⟨c, mul_comm _ _⟩, }⟩ - -theorem dvd.elim_left {P : Prop} (h₁ : a ∣ b) (h₂ : ∀ c, b = c * a → P) : P := -exists.elim (exists_eq_mul_left_of_dvd h₁) (assume c, assume h₃ : b = c * a, h₂ c h₃) - -@[simp] theorem dvd_mul_left (a b : α) : a ∣ b * a := dvd.intro b (mul_comm a b) - -theorem dvd_mul_of_dvd_right (h : a ∣ b) (c : α) : a ∣ c * b := -begin rw mul_comm, exact h.mul_right _ end - -alias dvd_mul_of_dvd_right ← has_dvd.dvd.mul_left - -local attribute [simp] mul_assoc mul_comm mul_left_comm - -theorem mul_dvd_mul : ∀ {a b c d : α}, a ∣ b → c ∣ d → a * c ∣ b * d -| a ._ c ._ ⟨e, rfl⟩ ⟨f, rfl⟩ := ⟨e * f, by simp⟩ - -theorem dvd_of_mul_left_dvd (h : a * b ∣ c) : b ∣ c := -dvd.elim h (λ d ceq, dvd.intro (a * d) (by simp [ceq])) - -end comm_semigroup - -section comm_monoid - -variables [comm_monoid α] {a b : α} - -theorem mul_dvd_mul_left (a : α) {b c : α} (h : b ∣ c) : a * b ∣ a * c := -mul_dvd_mul (dvd_refl a) h - -theorem mul_dvd_mul_right (h : a ∣ b) (c : α) : a * c ∣ b * c := -mul_dvd_mul h (dvd_refl c) - -theorem pow_dvd_pow_of_dvd {a b : α} (h : a ∣ b) : ∀ n : ℕ, a ^ n ∣ b ^ n -| 0 := by rw [pow_zero, pow_zero] -| (n+1) := by { rw [pow_succ, pow_succ], exact mul_dvd_mul h (pow_dvd_pow_of_dvd n) } - -end comm_monoid - -section semigroup_with_zero - -variables [semigroup_with_zero α] {a : α} - -theorem eq_zero_of_zero_dvd (h : 0 ∣ a) : a = 0 := -dvd.elim h (λ c H', H'.trans (zero_mul c)) - -/-- Given an element `a` of a commutative semigroup with zero, there exists another element whose - product with zero equals `a` iff `a` equals zero. -/ -@[simp] lemma zero_dvd_iff : 0 ∣ a ↔ a = 0 := -⟨eq_zero_of_zero_dvd, λ h, by { rw h, use 0, simp }⟩ - -@[simp] theorem dvd_zero (a : α) : a ∣ 0 := dvd.intro 0 (by simp) - -end semigroup_with_zero - -/-- Given two elements `b`, `c` of a `cancel_monoid_with_zero` and a nonzero element `a`, - `a*b` divides `a*c` iff `b` divides `c`. -/ -theorem mul_dvd_mul_iff_left [cancel_monoid_with_zero α] {a b c : α} - (ha : a ≠ 0) : a * b ∣ a * c ↔ b ∣ c := -exists_congr $ λ d, by rw [mul_assoc, mul_right_inj' ha] - -/-- Given two elements `a`, `b` of a commutative `cancel_monoid_with_zero` and a nonzero - element `c`, `a*c` divides `b*c` iff `a` divides `b`. -/ -theorem mul_dvd_mul_iff_right [cancel_comm_monoid_with_zero α] {a b c : α} (hc : c ≠ 0) : - a * c ∣ b * c ↔ a ∣ b := -exists_congr $ λ d, by rw [mul_right_comm, mul_left_inj' hc] - -/-! -### Units in various monoids --/ - -namespace units - -section monoid -variables [monoid α] {a b : α} {u : αˣ} - -/-- Elements of the unit group of a monoid represented as elements of the monoid - divide any element of the monoid. -/ -lemma coe_dvd : ↑u ∣ a := ⟨↑u⁻¹ * a, by simp⟩ - -/-- In a monoid, an element `a` divides an element `b` iff `a` divides all - associates of `b`. -/ -lemma dvd_mul_right : a ∣ b * u ↔ a ∣ b := -iff.intro - (assume ⟨c, eq⟩, ⟨c * ↑u⁻¹, by rw [← mul_assoc, ← eq, units.mul_inv_cancel_right]⟩) - (assume ⟨c, eq⟩, eq.symm ▸ (dvd_mul_right _ _).mul_right _) - -/-- In a monoid, an element `a` divides an element `b` iff all associates of `a` divide `b`. -/ -lemma mul_right_dvd : a * u ∣ b ↔ a ∣ b := -iff.intro - (λ ⟨c, eq⟩, ⟨↑u * c, eq.trans (mul_assoc _ _ _)⟩) - (λ h, dvd_trans (dvd.intro ↑u⁻¹ (by rw [mul_assoc, u.mul_inv, mul_one])) h) - -end monoid - -section comm_monoid -variables [comm_monoid α] {a b : α} {u : αˣ} - -/-- In a commutative monoid, an element `a` divides an element `b` iff `a` divides all left - associates of `b`. -/ -lemma dvd_mul_left : a ∣ u * b ↔ a ∣ b := by { rw mul_comm, apply dvd_mul_right } - -/-- In a commutative monoid, an element `a` divides an element `b` iff all - left associates of `a` divide `b`.-/ -lemma mul_left_dvd : ↑u * a ∣ b ↔ a ∣ b := -by { rw mul_comm, apply mul_right_dvd } - -end comm_monoid - -end units - -namespace is_unit - -section monoid - -variables [monoid α] {a b u : α} (hu : is_unit u) -include hu - -/-- Units of a monoid divide any element of the monoid. -/ -@[simp] lemma dvd : u ∣ a := by { rcases hu with ⟨u, rfl⟩, apply units.coe_dvd, } - -@[simp] lemma dvd_mul_right : a ∣ b * u ↔ a ∣ b := -by { rcases hu with ⟨u, rfl⟩, apply units.dvd_mul_right, } - -/-- In a monoid, an element a divides an element b iff all associates of `a` divide `b`.-/ -@[simp] lemma mul_right_dvd : a * u ∣ b ↔ a ∣ b := -by { rcases hu with ⟨u, rfl⟩, apply units.mul_right_dvd, } - -end monoid - -section comm_monoid -variables [comm_monoid α] (a b u : α) (hu : is_unit u) -include hu - -/-- In a commutative monoid, an element `a` divides an element `b` iff `a` divides all left - associates of `b`. -/ -@[simp] lemma dvd_mul_left : a ∣ u * b ↔ a ∣ b := -by { rcases hu with ⟨u, rfl⟩, apply units.dvd_mul_left, } - -/-- In a commutative monoid, an element `a` divides an element `b` iff all - left associates of `a` divide `b`.-/ -@[simp] lemma mul_left_dvd : u * a ∣ b ↔ a ∣ b := -by { rcases hu with ⟨u, rfl⟩, apply units.mul_left_dvd, } - -end comm_monoid - -end is_unit - -section comm_monoid -variables [comm_monoid α] - -theorem is_unit_iff_dvd_one {x : α} : is_unit x ↔ x ∣ 1 := -⟨is_unit.dvd, λ ⟨y, h⟩, ⟨⟨x, y, h.symm, by rw [h, mul_comm]⟩, rfl⟩⟩ - -theorem is_unit_iff_forall_dvd {x : α} : - is_unit x ↔ ∀ y, x ∣ y := -is_unit_iff_dvd_one.trans ⟨λ h y, h.trans (one_dvd _), λ h, h _⟩ - -theorem is_unit_of_dvd_unit {x y : α} - (xy : x ∣ y) (hu : is_unit y) : is_unit x := -is_unit_iff_dvd_one.2 $ xy.trans $ is_unit_iff_dvd_one.1 hu - -lemma is_unit_of_dvd_one : ∀a ∣ 1, is_unit (a:α) -| a ⟨b, eq⟩ := ⟨units.mk_of_mul_eq_one a b eq.symm, rfl⟩ - -lemma not_is_unit_of_not_is_unit_dvd {a b : α} (ha : ¬is_unit a) (hb : a ∣ b) : - ¬ is_unit b := -mt (is_unit_of_dvd_unit hb) ha - -end comm_monoid - -section comm_monoid_with_zero - -variable [comm_monoid_with_zero α] - -/-- `dvd_not_unit a b` expresses that `a` divides `b` "strictly", i.e. that `b` divided by `a` -is not a unit. -/ -def dvd_not_unit (a b : α) : Prop := a ≠ 0 ∧ ∃ x, ¬is_unit x ∧ b = a * x - -lemma dvd_not_unit_of_dvd_of_not_dvd {a b : α} (hd : a ∣ b) (hnd : ¬ b ∣ a) : - dvd_not_unit a b := -begin - split, - { rintro rfl, exact hnd (dvd_zero _) }, - { rcases hd with ⟨c, rfl⟩, - refine ⟨c, _, rfl⟩, - rintro ⟨u, rfl⟩, - simpa using hnd } -end - -end comm_monoid_with_zero - -lemma dvd_and_not_dvd_iff [cancel_comm_monoid_with_zero α] {x y : α} : - x ∣ y ∧ ¬y ∣ x ↔ dvd_not_unit x y := -⟨λ ⟨⟨d, hd⟩, hyx⟩, ⟨λ hx0, by simpa [hx0] using hyx, ⟨d, - mt is_unit_iff_dvd_one.1 (λ ⟨e, he⟩, hyx ⟨e, by rw [hd, mul_assoc, ← he, mul_one]⟩), hd⟩⟩, - λ ⟨hx0, d, hdu, hdx⟩, ⟨⟨d, hdx⟩, λ ⟨e, he⟩, hdu (is_unit_of_dvd_one _ - ⟨e, mul_left_cancel₀ hx0 $ by conv {to_lhs, rw [he, hdx]};simp [mul_assoc]⟩)⟩⟩ - -section monoid_with_zero - -variable [monoid_with_zero α] - -theorem ne_zero_of_dvd_ne_zero {p q : α} (h₁ : q ≠ 0) - (h₂ : p ∣ q) : p ≠ 0 := -begin - rcases h₂ with ⟨u, rfl⟩, - exact left_ne_zero_of_mul h₁, -end - -end monoid_with_zero diff --git a/src/algebra/divisibility/basic.lean b/src/algebra/divisibility/basic.lean new file mode 100644 index 0000000000000..79400687391e8 --- /dev/null +++ b/src/algebra/divisibility/basic.lean @@ -0,0 +1,150 @@ +/- +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, Floris van Doorn, Amelia Livingston, Yury Kudryashov, +Neil Strickland, Aaron Anderson +-/ + +import algebra.hom.group + +/-! +# Divisibility + +This file defines the basics of the divisibility relation in the context of `(comm_)` `monoid`s. + +## Main definitions + + * `monoid.has_dvd` + +## Implementation notes + +The divisibility relation is defined for all monoids, and as such, depends on the order of + multiplication if the monoid is not commutative. There are two possible conventions for + divisibility in the noncommutative context, and this relation follows the convention for ordinals, + so `a | b` is defined as `∃ c, b = a * c`. + +## Tags + +divisibility, divides +-/ + +variables {α : Type*} + +section semigroup + +variables [semigroup α] {a b c : α} + +/-- There are two possible conventions for divisibility, which coincide in a `comm_monoid`. + This matches the convention for ordinals. -/ +@[priority 100] +instance semigroup_has_dvd : has_dvd α := +has_dvd.mk (λ a b, ∃ c, b = a * c) + +-- TODO: this used to not have `c` explicit, but that seems to be important +-- for use with tactics, similar to `exists.intro` +theorem dvd.intro (c : α) (h : a * c = b) : a ∣ b := +exists.intro c h^.symm + +alias dvd.intro ← dvd_of_mul_right_eq + +theorem exists_eq_mul_right_of_dvd (h : a ∣ b) : ∃ c, b = a * c := h + +theorem dvd.elim {P : Prop} {a b : α} (H₁ : a ∣ b) (H₂ : ∀ c, b = a * c → P) : P := +exists.elim H₁ H₂ + +local attribute [simp] mul_assoc mul_comm mul_left_comm + +@[trans] theorem dvd_trans : a ∣ b → b ∣ c → a ∣ c +| ⟨d, h₁⟩ ⟨e, h₂⟩ := ⟨d * e, h₁ ▸ h₂.trans $ mul_assoc a d e⟩ + +alias dvd_trans ← has_dvd.dvd.trans + +instance : is_trans α (∣) := ⟨λ a b c, dvd_trans⟩ + +@[simp] theorem dvd_mul_right (a b : α) : a ∣ a * b := dvd.intro b rfl + +theorem dvd_mul_of_dvd_left (h : a ∣ b) (c : α) : a ∣ b * c := +h.trans (dvd_mul_right b c) + +alias dvd_mul_of_dvd_left ← has_dvd.dvd.mul_right + +theorem dvd_of_mul_right_dvd (h : a * b ∣ c) : a ∣ c := +(dvd_mul_right a b).trans h + +section map_dvd + +variables {M N : Type*} [monoid M] [monoid N] + +lemma map_dvd {F : Type*} [mul_hom_class F M N] (f : F) {a b} : a ∣ b → f a ∣ f b +| ⟨c, h⟩ := ⟨f c, h.symm ▸ map_mul f a c⟩ + +lemma mul_hom.map_dvd (f : M →ₙ* N) {a b} : a ∣ b → f a ∣ f b := map_dvd f + +lemma monoid_hom.map_dvd (f : M →* N) {a b} : a ∣ b → f a ∣ f b := map_dvd f + +end map_dvd + +end semigroup + +section monoid + +variables [monoid α] + +@[refl, simp] theorem dvd_refl (a : α) : a ∣ a := dvd.intro 1 (mul_one a) +theorem dvd_rfl : ∀ {a : α}, a ∣ a := dvd_refl +instance : is_refl α (∣) := ⟨dvd_refl⟩ + +theorem one_dvd (a : α) : 1 ∣ a := dvd.intro a (one_mul a) + +end monoid + +section comm_semigroup + +variables [comm_semigroup α] {a b c : α} + +theorem dvd.intro_left (c : α) (h : c * a = b) : a ∣ b := +dvd.intro _ (begin rewrite mul_comm at h, apply h end) + +alias dvd.intro_left ← dvd_of_mul_left_eq + +theorem exists_eq_mul_left_of_dvd (h : a ∣ b) : ∃ c, b = c * a := +dvd.elim h (assume c, assume H1 : b = a * c, exists.intro c (eq.trans H1 (mul_comm a c))) + +lemma dvd_iff_exists_eq_mul_left : a ∣ b ↔ ∃ c, b = c * a := +⟨exists_eq_mul_left_of_dvd, by { rintro ⟨c, rfl⟩, exact ⟨c, mul_comm _ _⟩, }⟩ + +theorem dvd.elim_left {P : Prop} (h₁ : a ∣ b) (h₂ : ∀ c, b = c * a → P) : P := +exists.elim (exists_eq_mul_left_of_dvd h₁) (assume c, assume h₃ : b = c * a, h₂ c h₃) + +@[simp] theorem dvd_mul_left (a b : α) : a ∣ b * a := dvd.intro b (mul_comm a b) + +theorem dvd_mul_of_dvd_right (h : a ∣ b) (c : α) : a ∣ c * b := +begin rw mul_comm, exact h.mul_right _ end + +alias dvd_mul_of_dvd_right ← has_dvd.dvd.mul_left + +local attribute [simp] mul_assoc mul_comm mul_left_comm + +theorem mul_dvd_mul : ∀ {a b c d : α}, a ∣ b → c ∣ d → a * c ∣ b * d +| a ._ c ._ ⟨e, rfl⟩ ⟨f, rfl⟩ := ⟨e * f, by simp⟩ + +theorem dvd_of_mul_left_dvd (h : a * b ∣ c) : b ∣ c := +dvd.elim h (λ d ceq, dvd.intro (a * d) (by simp [ceq])) + +end comm_semigroup + +section comm_monoid + +variables [comm_monoid α] {a b : α} + +theorem mul_dvd_mul_left (a : α) {b c : α} (h : b ∣ c) : a * b ∣ a * c := +mul_dvd_mul (dvd_refl a) h + +theorem mul_dvd_mul_right (h : a ∣ b) (c : α) : a * c ∣ b * c := +mul_dvd_mul h (dvd_refl c) + +theorem pow_dvd_pow_of_dvd {a b : α} (h : a ∣ b) : ∀ n : ℕ, a ^ n ∣ b ^ n +| 0 := by rw [pow_zero, pow_zero] +| (n+1) := by { rw [pow_succ, pow_succ], exact mul_dvd_mul h (pow_dvd_pow_of_dvd n) } + +end comm_monoid diff --git a/src/algebra/divisibility/units.lean b/src/algebra/divisibility/units.lean new file mode 100644 index 0000000000000..b5f0a67edf160 --- /dev/null +++ b/src/algebra/divisibility/units.lean @@ -0,0 +1,114 @@ +/- +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, Floris van Doorn, Amelia Livingston, Yury Kudryashov, +Neil Strickland, Aaron Anderson +-/ +import algebra.divisibility.basic +import algebra.hom.units + +/-! +# Lemmas about divisibility and units +-/ + +variables {α : Type*} + +namespace units + +section monoid +variables [monoid α] {a b : α} {u : αˣ} + +/-- Elements of the unit group of a monoid represented as elements of the monoid + divide any element of the monoid. -/ +lemma coe_dvd : ↑u ∣ a := ⟨↑u⁻¹ * a, by simp⟩ + +/-- In a monoid, an element `a` divides an element `b` iff `a` divides all + associates of `b`. -/ +lemma dvd_mul_right : a ∣ b * u ↔ a ∣ b := +iff.intro + (assume ⟨c, eq⟩, ⟨c * ↑u⁻¹, by rw [← mul_assoc, ← eq, units.mul_inv_cancel_right]⟩) + (assume ⟨c, eq⟩, eq.symm ▸ (dvd_mul_right _ _).mul_right _) + +/-- In a monoid, an element `a` divides an element `b` iff all associates of `a` divide `b`. -/ +lemma mul_right_dvd : a * u ∣ b ↔ a ∣ b := +iff.intro + (λ ⟨c, eq⟩, ⟨↑u * c, eq.trans (mul_assoc _ _ _)⟩) + (λ h, dvd_trans (dvd.intro ↑u⁻¹ (by rw [mul_assoc, u.mul_inv, mul_one])) h) + +end monoid + +section comm_monoid +variables [comm_monoid α] {a b : α} {u : αˣ} + +/-- In a commutative monoid, an element `a` divides an element `b` iff `a` divides all left + associates of `b`. -/ +lemma dvd_mul_left : a ∣ u * b ↔ a ∣ b := by { rw mul_comm, apply dvd_mul_right } + +/-- In a commutative monoid, an element `a` divides an element `b` iff all + left associates of `a` divide `b`.-/ +lemma mul_left_dvd : ↑u * a ∣ b ↔ a ∣ b := +by { rw mul_comm, apply mul_right_dvd } + +end comm_monoid + +end units + +namespace is_unit + +section monoid + +variables [monoid α] {a b u : α} (hu : is_unit u) +include hu + +/-- Units of a monoid divide any element of the monoid. -/ +@[simp] lemma dvd : u ∣ a := by { rcases hu with ⟨u, rfl⟩, apply units.coe_dvd, } + +@[simp] lemma dvd_mul_right : a ∣ b * u ↔ a ∣ b := +by { rcases hu with ⟨u, rfl⟩, apply units.dvd_mul_right, } + +/-- In a monoid, an element a divides an element b iff all associates of `a` divide `b`.-/ +@[simp] lemma mul_right_dvd : a * u ∣ b ↔ a ∣ b := +by { rcases hu with ⟨u, rfl⟩, apply units.mul_right_dvd, } + +end monoid + +section comm_monoid +variables [comm_monoid α] (a b u : α) (hu : is_unit u) +include hu + +/-- In a commutative monoid, an element `a` divides an element `b` iff `a` divides all left + associates of `b`. -/ +@[simp] lemma dvd_mul_left : a ∣ u * b ↔ a ∣ b := +by { rcases hu with ⟨u, rfl⟩, apply units.dvd_mul_left, } + +/-- In a commutative monoid, an element `a` divides an element `b` iff all + left associates of `a` divide `b`.-/ +@[simp] lemma mul_left_dvd : u * a ∣ b ↔ a ∣ b := +by { rcases hu with ⟨u, rfl⟩, apply units.mul_left_dvd, } + +end comm_monoid + +end is_unit + +section comm_monoid +variables [comm_monoid α] + +theorem is_unit_iff_dvd_one {x : α} : is_unit x ↔ x ∣ 1 := +⟨is_unit.dvd, λ ⟨y, h⟩, ⟨⟨x, y, h.symm, by rw [h, mul_comm]⟩, rfl⟩⟩ + +theorem is_unit_iff_forall_dvd {x : α} : + is_unit x ↔ ∀ y, x ∣ y := +is_unit_iff_dvd_one.trans ⟨λ h y, h.trans (one_dvd _), λ h, h _⟩ + +theorem is_unit_of_dvd_unit {x y : α} + (xy : x ∣ y) (hu : is_unit y) : is_unit x := +is_unit_iff_dvd_one.2 $ xy.trans $ is_unit_iff_dvd_one.1 hu + +lemma is_unit_of_dvd_one : ∀a ∣ 1, is_unit (a:α) +| a ⟨b, eq⟩ := ⟨units.mk_of_mul_eq_one a b eq.symm, rfl⟩ + +lemma not_is_unit_of_not_is_unit_dvd {a b : α} (ha : ¬is_unit a) (hb : a ∣ b) : + ¬ is_unit b := +mt (is_unit_of_dvd_unit hb) ha + +end comm_monoid diff --git a/src/algebra/euclidean_domain.lean b/src/algebra/euclidean_domain/basic.lean similarity index 51% rename from src/algebra/euclidean_domain.lean rename to src/algebra/euclidean_domain/basic.lean index f749980d318ee..c311017586490 100644 --- a/src/algebra/euclidean_domain.lean +++ b/src/algebra/euclidean_domain/basic.lean @@ -3,112 +3,29 @@ Copyright (c) 2018 Louis Carlin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Louis Carlin, Mario Carneiro -/ -import algebra.field.basic +import algebra.euclidean_domain.defs +import algebra.ring.divisibility +import algebra.ring.regular +import algebra.field.defs +import algebra.group_with_zero.divisibility /-! -# Euclidean domains - -This file introduces Euclidean domains and provides the extended Euclidean algorithm. To be precise, -a slightly more general version is provided which is sometimes called a transfinite Euclidean domain -and differs in the fact that the degree function need not take values in `ℕ` but can take values in -any well-ordered set. Transfinite Euclidean domains were introduced by Motzkin and examples which -don't satisfy the classical notion were provided independently by Hiblot and Nagata. - -## Main definitions - -* `euclidean_domain`: Defines Euclidean domain with functions `quotient` and `remainder`. Instances - of `has_div` and `has_mod` are provided, so that one can write `a = b * (a / b) + a % b`. -* `gcd`: defines the greatest common divisors of two elements of a Euclidean domain. -* `xgcd`: given two elements `a b : R`, `xgcd a b` defines the pair `(x, y)` such that - `x * a + y * b = gcd a b`. -* `lcm`: defines the lowest common multiple of two elements `a` and `b` of a Euclidean domain as - `a * b / (gcd a b)` +# Lemmas about Euclidean domains ## Main statements * `gcd_eq_gcd_ab`: states Bézout's lemma for Euclidean domains. -* `int.euclidean_domain`: shows that `ℤ` is a Euclidean domain. -* `field.to_euclidean_domain`: shows that any field is a Euclidean domain. - -## Notation - -`≺` denotes the well founded relation on the Euclidean domain, e.g. in the example of the polynomial -ring over a field, `p ≺ q` for polynomials `p` and `q` if and only if the degree of `p` is less than -the degree of `q`. - -## Implementation details -Instead of working with a valuation, `euclidean_domain` is implemented with the existence of a well -founded relation `r` on the integral domain `R`, which in the example of `ℤ` would correspond to -setting `i ≺ j` for integers `i` and `j` if the absolute value of `i` is smaller than the absolute -value of `j`. - -## References - -* [Th. Motzkin, *The Euclidean algorithm*][MR32592] -* [J.-J. Hiblot, *Des anneaux euclidiens dont le plus petit algorithme n'est pas à valeurs finies*] - [MR399081] -* [M. Nagata, *On Euclid algorithm*][MR541021] - - -## Tags - -Euclidean domain, transfinite Euclidean domain, Bézout's lemma -/ universe u -/-- A `euclidean_domain` is an non-trivial commutative ring with a division and a remainder, - satisfying `b * (a / b) + a % b = a`. - The definition of a euclidean domain usually includes a valuation function `R → ℕ`. - This definition is slightly generalised to include a well founded relation - `r` with the property that `r (a % b) b`, instead of a valuation. -/ -@[protect_proj without mul_left_not_lt r_well_founded] -class euclidean_domain (R : Type u) extends comm_ring R, nontrivial R := -(quotient : R → R → R) -(quotient_zero : ∀ a, quotient a 0 = 0) -(remainder : R → R → R) -(quotient_mul_add_remainder_eq : ∀ a b, b * quotient a b + remainder a b = a) -(r : R → R → Prop) -(r_well_founded : well_founded r) -(remainder_lt : ∀ a {b}, b ≠ 0 → r (remainder a b) b) -(mul_left_not_lt : ∀ a {b}, b ≠ 0 → ¬r (a * b) a) - namespace euclidean_domain variable {R : Type u} variables [euclidean_domain R] local infix ` ≺ `:50 := euclidean_domain.r -@[priority 70] -- see Note [lower instance priority] -instance : has_div R := ⟨euclidean_domain.quotient⟩ - -@[priority 70] -- see Note [lower instance priority] -instance : has_mod R := ⟨euclidean_domain.remainder⟩ - -theorem div_add_mod (a b : R) : b * (a / b) + a % b = a := -euclidean_domain.quotient_mul_add_remainder_eq _ _ - -lemma mod_add_div (a b : R) : a % b + b * (a / b) = a := -(add_comm _ _).trans (div_add_mod _ _) - -lemma mod_add_div' (m k : R) : m % k + (m / k) * k = m := -by { rw mul_comm, exact mod_add_div _ _ } - -lemma div_add_mod' (m k : R) : (m / k) * k + m % k = m := -by { rw mul_comm, exact div_add_mod _ _ } - -lemma mod_eq_sub_mul_div {R : Type*} [euclidean_domain R] (a b : R) : - a % b = a - b * (a / b) := -calc a % b = b * (a / b) + a % b - b * (a / b) : (add_sub_cancel' _ _).symm -... = a - b * (a / b) : by rw div_add_mod - -theorem mod_lt : ∀ a {b : R}, b ≠ 0 → (a % b) ≺ b := -euclidean_domain.remainder_lt - -theorem mul_right_not_lt {a : R} (b) (h : a ≠ 0) : ¬(a * b) ≺ b := -by { rw mul_comm, exact mul_left_not_lt b h } - lemma mul_div_cancel_left {a : R} (b) (a0 : a ≠ 0) : a * b / a = b := eq.symm $ eq_of_sub_eq_zero $ classical.by_contradiction $ λ h, begin @@ -120,9 +37,6 @@ end lemma mul_div_cancel (a) {b : R} (b0 : b ≠ 0) : a * b / b = a := by { rw mul_comm, exact mul_div_cancel_left a b0 } -@[simp] lemma mod_zero (a : R) : a % 0 = a := -by simpa only [zero_mul, zero_add] using div_add_mod a 0 - @[simp] lemma mod_eq_zero {a b : R} : a % b = 0 ↔ b ∣ a := ⟨λ h, by { rw [← div_add_mod a b, h, add_zero], exact dvd_mul_right _ _ }, λ ⟨c, e⟩, begin @@ -139,22 +53,12 @@ mod_eq_zero.2 dvd_rfl lemma dvd_mod_iff {a b c : R} (h : c ∣ b) : c ∣ a % b ↔ c ∣ a := by rw [dvd_add_iff_right (h.mul_right _), div_add_mod] -lemma lt_one (a : R) : a ≺ (1:R) → a = 0 := -by { haveI := classical.dec, exact - not_imp_not.1 (λ h, by simpa only [one_mul] using mul_left_not_lt 1 h) } - -lemma val_dvd_le : ∀ a b : R, b ∣ a → a ≠ 0 → ¬a ≺ b -| _ b ⟨d, rfl⟩ ha := mul_left_not_lt b (mt (by { rintro rfl, exact mul_zero _ }) ha) - @[simp] lemma mod_one (a : R) : a % 1 = 0 := mod_eq_zero.2 (one_dvd _) @[simp] lemma zero_mod (b : R) : 0 % b = 0 := mod_eq_zero.2 (dvd_zero _) -@[simp, priority 900] lemma div_zero (a : R) : a / 0 = 0 := -euclidean_domain.quotient_zero a - @[simp, priority 900] lemma zero_div {a : R} : 0 / a = 0 := classical.by_cases (λ a0 : a = 0, a0.symm ▸ div_zero 0) @@ -202,37 +106,9 @@ begin rw [mul_assoc, mul_div_cancel_left _ ha] end -section -open_locale classical - -@[elab_as_eliminator] -theorem gcd.induction {P : R → R → Prop} : ∀ a b : R, - (∀ x, P 0 x) → - (∀ a b, a ≠ 0 → P (b % a) a → P a b) → - P a b -| a := λ b H0 H1, if a0 : a = 0 then a0.symm ▸ H0 _ else - have h:_ := mod_lt b a0, - H1 _ _ a0 (gcd.induction (b%a) a H0 H1) -using_well_founded {dec_tac := tactic.assumption, - rel_tac := λ _ _, `[exact ⟨_, r_well_founded⟩]} - -end - section gcd variable [decidable_eq R] -/-- `gcd a b` is a (non-unique) element such that `gcd a b ∣ a` `gcd a b ∣ b`, and for - any element `c` such that `c ∣ a` and `c ∣ b`, then `c ∣ gcd a b` -/ -def gcd : R → R → R -| a := λ b, if a0 : a = 0 then b else - have h:_ := mod_lt b a0, - gcd (b%a) a -using_well_founded {dec_tac := tactic.assumption, - rel_tac := λ _ _, `[exact ⟨_, r_well_founded⟩]} - -@[simp] theorem gcd_zero_left (a : R) : gcd 0 a = a := -by { rw gcd, exact if_pos rfl } - @[simp] theorem gcd_zero_right (a : R) : gcd a 0 = a := by { rw gcd, split_ifs; simp only [h, zero_mod, gcd_zero_left] } @@ -270,48 +146,6 @@ gcd_eq_left.2 (one_dvd _) @[simp] theorem gcd_self (a : R) : gcd a a = a := gcd_eq_left.2 dvd_rfl -/-- -An implementation of the extended GCD algorithm. -At each step we are computing a triple `(r, s, t)`, where `r` is the next value of the GCD -algorithm, to compute the greatest common divisor of the input (say `x` and `y`), and `s` and `t` -are the coefficients in front of `x` and `y` to obtain `r` (i.e. `r = s * x + t * y`). -The function `xgcd_aux` takes in two triples, and from these recursively computes the next triple: -``` -xgcd_aux (r, s, t) (r', s', t') = xgcd_aux (r' % r, s' - (r' / r) * s, t' - (r' / r) * t) (r, s, t) -``` --/ -def xgcd_aux : R → R → R → R → R → R → R × R × R -| r := λ s t r' s' t', -if hr : r = 0 then (r', s', t') - else - have r' % r ≺ r, from mod_lt _ hr, - let q := r' / r in xgcd_aux (r' % r) (s' - q * s) (t' - q * t) r s t -using_well_founded {dec_tac := tactic.assumption, - rel_tac := λ _ _, `[exact ⟨_, r_well_founded⟩]} - -@[simp] theorem xgcd_zero_left {s t r' s' t' : R} : xgcd_aux 0 s t r' s' t' = (r', s', t') := -by { unfold xgcd_aux, exact if_pos rfl } - -theorem xgcd_aux_rec {r s t r' s' t' : R} (h : r ≠ 0) : - xgcd_aux r s t r' s' t' = xgcd_aux (r' % r) (s' - (r' / r) * s) (t' - (r' / r) * t) r s t := -by { conv {to_lhs, rw [xgcd_aux]}, exact if_neg h} - -/-- Use the extended GCD algorithm to generate the `a` and `b` values - satisfying `gcd x y = x * a + y * b`. -/ -def xgcd (x y : R) : R × R := (xgcd_aux x 1 0 y 0 1).2 - -/-- The extended GCD `a` value in the equation `gcd x y = x * a + y * b`. -/ -def gcd_a (x y : R) : R := (xgcd x y).1 - -/-- The extended GCD `b` value in the equation `gcd x y = x * a + y * b`. -/ -def gcd_b (x y : R) : R := (xgcd x y).2 - -@[simp] theorem gcd_a_zero_left {s : R} : gcd_a 0 s = 0 := -by { unfold gcd_a, rw [xgcd, xgcd_zero_left] } - -@[simp] theorem gcd_b_zero_left {s : R} : gcd_b 0 s = 1 := -by { unfold gcd_b, rw [xgcd, xgcd_zero_left] } - @[simp] theorem xgcd_aux_fst (x y : R) : ∀ s t s' t', (xgcd_aux x s t y s' t').1 = gcd x y := gcd.induction x y (by { intros, rw [xgcd_zero_left, gcd_zero_left] }) @@ -320,9 +154,6 @@ gcd.induction x y (by { intros, rw [xgcd_zero_left, gcd_zero_left] }) theorem xgcd_aux_val (x y : R) : xgcd_aux x 1 0 y 0 1 = (gcd x y, xgcd x y) := by rw [xgcd, ← xgcd_aux_fst x y 1 0 0 1, prod.mk.eta] -theorem xgcd_val (x y : R) : xgcd x y = (gcd_a x y, gcd_b x y) := -prod.mk.eta.symm - private def P (a b : R) : R × R × R → Prop | (r, s, t) := (r : R) = a * s + b * t theorem xgcd_aux_P (a b : R) {r r' : R} : ∀ {s t s' t'}, P a b (r, s, t) → @@ -353,11 +184,6 @@ end gcd section lcm variables [decidable_eq R] -/-- `lcm a b` is a (non-unique) element such that `a ∣ lcm a b` `b ∣ lcm a b`, and for - any element `c` such that `a ∣ c` and `b ∣ c`, then `lcm a b ∣ c` -/ -def lcm (x y : R) : R := -x * y / gcd x y - theorem dvd_lcm_left (x y : R) : x ∣ lcm x y := classical.by_cases (assume hxy : gcd x y = 0, by { rw [lcm, hxy, div_zero], exact dvd_zero _ }) @@ -446,43 +272,3 @@ end end div end euclidean_domain - -instance int.euclidean_domain : euclidean_domain ℤ := -{ add := (+), - mul := (*), - one := 1, - zero := 0, - neg := has_neg.neg, - quotient := (/), - quotient_zero := int.div_zero, - remainder := (%), - quotient_mul_add_remainder_eq := λ a b, int.div_add_mod _ _, - r := λ a b, a.nat_abs < b.nat_abs, - r_well_founded := measure_wf (λ a, int.nat_abs a), - remainder_lt := λ a b b0, int.coe_nat_lt.1 $ - by { rw [int.nat_abs_of_nonneg (int.mod_nonneg _ b0), ← int.abs_eq_nat_abs], - exact int.mod_lt _ b0 }, - mul_left_not_lt := λ a b b0, not_lt_of_ge $ - by {rw [← mul_one a.nat_abs, int.nat_abs_mul], - exact mul_le_mul_of_nonneg_left (int.nat_abs_pos_of_ne_zero b0) (nat.zero_le _) }, - .. int.comm_ring, - .. int.nontrivial } - -@[priority 100] -- see Note [lower instance priority] -instance field.to_euclidean_domain {K : Type u} [field K] : euclidean_domain K := -{ add := (+), - mul := (*), - one := 1, - zero := 0, - neg := has_neg.neg, - quotient := (/), - remainder := λ a b, a - a * b / b, - quotient_zero := div_zero, - quotient_mul_add_remainder_eq := λ a b, - by { classical, by_cases b = 0; simp [h, mul_div_cancel'] }, - r := λ a b, a = 0 ∧ b ≠ 0, - r_well_founded := well_founded.intro $ λ a, acc.intro _ $ λ b ⟨hb, hna⟩, - acc.intro _ $ λ c ⟨hc, hnb⟩, false.elim $ hnb hb, - remainder_lt := λ a b hnb, by simp [hnb], - mul_left_not_lt := λ a b hnb ⟨hab, hna⟩, or.cases_on (mul_eq_zero.1 hab) hna hnb, - .. ‹field K› } diff --git a/src/algebra/euclidean_domain/defs.lean b/src/algebra/euclidean_domain/defs.lean new file mode 100644 index 0000000000000..c6bfb4ab3ec08 --- /dev/null +++ b/src/algebra/euclidean_domain/defs.lean @@ -0,0 +1,221 @@ +/- +Copyright (c) 2018 Louis Carlin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Louis Carlin, Mario Carneiro +-/ +import logic.nontrivial +import algebra.opposites +import algebra.field.defs +import algebra.hom.group +import algebra.divisibility.basic +import algebra.group.basic +import algebra.order.monoid.lemmas + +/-! +# Euclidean domains + +This file introduces Euclidean domains and provides the extended Euclidean algorithm. To be precise, +a slightly more general version is provided which is sometimes called a transfinite Euclidean domain +and differs in the fact that the degree function need not take values in `ℕ` but can take values in +any well-ordered set. Transfinite Euclidean domains were introduced by Motzkin and examples which +don't satisfy the classical notion were provided independently by Hiblot and Nagata. + +## Main definitions + +* `euclidean_domain`: Defines Euclidean domain with functions `quotient` and `remainder`. Instances + of `has_div` and `has_mod` are provided, so that one can write `a = b * (a / b) + a % b`. +* `gcd`: defines the greatest common divisors of two elements of a Euclidean domain. +* `xgcd`: given two elements `a b : R`, `xgcd a b` defines the pair `(x, y)` such that + `x * a + y * b = gcd a b`. +* `lcm`: defines the lowest common multiple of two elements `a` and `b` of a Euclidean domain as + `a * b / (gcd a b)` + +## Main statements + +See `algebra.euclidean_domain.basic` for most of the theorems about Eucliean domains, +including Bézout's lemma. + +See `algebra.euclidean_domain.instances` for that facts that `ℤ` is a Euclidean domain, +as is any field. + +## Notation + +`≺` denotes the well founded relation on the Euclidean domain, e.g. in the example of the polynomial +ring over a field, `p ≺ q` for polynomials `p` and `q` if and only if the degree of `p` is less than +the degree of `q`. + +## Implementation details + +Instead of working with a valuation, `euclidean_domain` is implemented with the existence of a well +founded relation `r` on the integral domain `R`, which in the example of `ℤ` would correspond to +setting `i ≺ j` for integers `i` and `j` if the absolute value of `i` is smaller than the absolute +value of `j`. + +## References + +* [Th. Motzkin, *The Euclidean algorithm*][MR32592] +* [J.-J. Hiblot, *Des anneaux euclidiens dont le plus petit algorithme n'est pas à valeurs finies*] + [MR399081] +* [M. Nagata, *On Euclid algorithm*][MR541021] + + +## Tags + +Euclidean domain, transfinite Euclidean domain, Bézout's lemma +-/ + +universe u + +/-- A `euclidean_domain` is an non-trivial commutative ring with a division and a remainder, + satisfying `b * (a / b) + a % b = a`. + The definition of a euclidean domain usually includes a valuation function `R → ℕ`. + This definition is slightly generalised to include a well founded relation + `r` with the property that `r (a % b) b`, instead of a valuation. -/ +@[protect_proj without mul_left_not_lt r_well_founded] +class euclidean_domain (R : Type u) extends comm_ring R, nontrivial R := +(quotient : R → R → R) +(quotient_zero : ∀ a, quotient a 0 = 0) +(remainder : R → R → R) +(quotient_mul_add_remainder_eq : ∀ a b, b * quotient a b + remainder a b = a) +(r : R → R → Prop) +(r_well_founded : well_founded r) +(remainder_lt : ∀ a {b}, b ≠ 0 → r (remainder a b) b) +(mul_left_not_lt : ∀ a {b}, b ≠ 0 → ¬r (a * b) a) + +namespace euclidean_domain +variable {R : Type u} +variables [euclidean_domain R] + +local infix ` ≺ `:50 := euclidean_domain.r + +@[priority 70] -- see Note [lower instance priority] +instance : has_div R := ⟨euclidean_domain.quotient⟩ + +@[priority 70] -- see Note [lower instance priority] +instance : has_mod R := ⟨euclidean_domain.remainder⟩ + +theorem div_add_mod (a b : R) : b * (a / b) + a % b = a := +euclidean_domain.quotient_mul_add_remainder_eq _ _ + +lemma mod_add_div (a b : R) : a % b + b * (a / b) = a := +(add_comm _ _).trans (div_add_mod _ _) + +lemma mod_add_div' (m k : R) : m % k + (m / k) * k = m := +by { rw mul_comm, exact mod_add_div _ _ } + +lemma div_add_mod' (m k : R) : (m / k) * k + m % k = m := +by { rw mul_comm, exact div_add_mod _ _ } + +lemma mod_eq_sub_mul_div {R : Type*} [euclidean_domain R] (a b : R) : + a % b = a - b * (a / b) := +calc a % b = b * (a / b) + a % b - b * (a / b) : (add_sub_cancel' _ _).symm +... = a - b * (a / b) : by rw div_add_mod + +theorem mod_lt : ∀ a {b : R}, b ≠ 0 → (a % b) ≺ b := +euclidean_domain.remainder_lt + +theorem mul_right_not_lt {a : R} (b) (h : a ≠ 0) : ¬(a * b) ≺ b := +by { rw mul_comm, exact mul_left_not_lt b h } + +@[simp] lemma mod_zero (a : R) : a % 0 = a := +by simpa only [zero_mul, zero_add] using div_add_mod a 0 + +lemma lt_one (a : R) : a ≺ (1:R) → a = 0 := +by { haveI := classical.dec, exact + not_imp_not.1 (λ h, by simpa only [one_mul] using mul_left_not_lt 1 h) } + +lemma val_dvd_le : ∀ a b : R, b ∣ a → a ≠ 0 → ¬a ≺ b +| _ b ⟨d, rfl⟩ ha := mul_left_not_lt b (mt (by { rintro rfl, exact mul_zero _ }) ha) + +@[simp, priority 900] lemma div_zero (a : R) : a / 0 = 0 := +euclidean_domain.quotient_zero a + +section +open_locale classical + +@[elab_as_eliminator] +theorem gcd.induction {P : R → R → Prop} : ∀ a b : R, + (∀ x, P 0 x) → + (∀ a b, a ≠ 0 → P (b % a) a → P a b) → + P a b +| a := λ b H0 H1, if a0 : a = 0 then a0.symm ▸ H0 _ else + have h:_ := mod_lt b a0, + H1 _ _ a0 (gcd.induction (b%a) a H0 H1) +using_well_founded {dec_tac := tactic.assumption, + rel_tac := λ _ _, `[exact ⟨_, r_well_founded⟩]} + +end + +section gcd +variable [decidable_eq R] + +/-- `gcd a b` is a (non-unique) element such that `gcd a b ∣ a` `gcd a b ∣ b`, and for + any element `c` such that `c ∣ a` and `c ∣ b`, then `c ∣ gcd a b` -/ +def gcd : R → R → R +| a := λ b, if a0 : a = 0 then b else + have h:_ := mod_lt b a0, + gcd (b%a) a +using_well_founded {dec_tac := tactic.assumption, + rel_tac := λ _ _, `[exact ⟨_, r_well_founded⟩]} + +@[simp] theorem gcd_zero_left (a : R) : gcd 0 a = a := +by { rw gcd, exact if_pos rfl } + +/-- +An implementation of the extended GCD algorithm. +At each step we are computing a triple `(r, s, t)`, where `r` is the next value of the GCD +algorithm, to compute the greatest common divisor of the input (say `x` and `y`), and `s` and `t` +are the coefficients in front of `x` and `y` to obtain `r` (i.e. `r = s * x + t * y`). +The function `xgcd_aux` takes in two triples, and from these recursively computes the next triple: +``` +xgcd_aux (r, s, t) (r', s', t') = xgcd_aux (r' % r, s' - (r' / r) * s, t' - (r' / r) * t) (r, s, t) +``` +-/ +def xgcd_aux : R → R → R → R → R → R → R × R × R +| r := λ s t r' s' t', +if hr : r = 0 then (r', s', t') + else + have r' % r ≺ r, from mod_lt _ hr, + let q := r' / r in xgcd_aux (r' % r) (s' - q * s) (t' - q * t) r s t +using_well_founded {dec_tac := tactic.assumption, + rel_tac := λ _ _, `[exact ⟨_, r_well_founded⟩]} + +@[simp] theorem xgcd_zero_left {s t r' s' t' : R} : xgcd_aux 0 s t r' s' t' = (r', s', t') := +by { unfold xgcd_aux, exact if_pos rfl } + +theorem xgcd_aux_rec {r s t r' s' t' : R} (h : r ≠ 0) : + xgcd_aux r s t r' s' t' = xgcd_aux (r' % r) (s' - (r' / r) * s) (t' - (r' / r) * t) r s t := +by { conv {to_lhs, rw [xgcd_aux]}, exact if_neg h} + +/-- Use the extended GCD algorithm to generate the `a` and `b` values + satisfying `gcd x y = x * a + y * b`. -/ +def xgcd (x y : R) : R × R := (xgcd_aux x 1 0 y 0 1).2 + +/-- The extended GCD `a` value in the equation `gcd x y = x * a + y * b`. -/ +def gcd_a (x y : R) : R := (xgcd x y).1 + +/-- The extended GCD `b` value in the equation `gcd x y = x * a + y * b`. -/ +def gcd_b (x y : R) : R := (xgcd x y).2 + +@[simp] theorem gcd_a_zero_left {s : R} : gcd_a 0 s = 0 := +by { unfold gcd_a, rw [xgcd, xgcd_zero_left] } + +@[simp] theorem gcd_b_zero_left {s : R} : gcd_b 0 s = 1 := +by { unfold gcd_b, rw [xgcd, xgcd_zero_left] } + +theorem xgcd_val (x y : R) : xgcd x y = (gcd_a x y, gcd_b x y) := +prod.mk.eta.symm + +end gcd + +section lcm +variables [decidable_eq R] + +/-- `lcm a b` is a (non-unique) element such that `a ∣ lcm a b` `b ∣ lcm a b`, and for + any element `c` such that `a ∣ c` and `b ∣ c`, then `lcm a b ∣ c` -/ +def lcm (x y : R) : R := +x * y / gcd x y + +end lcm + +end euclidean_domain diff --git a/src/algebra/euclidean_domain/instances.lean b/src/algebra/euclidean_domain/instances.lean new file mode 100644 index 0000000000000..7a1e38757ee86 --- /dev/null +++ b/src/algebra/euclidean_domain/instances.lean @@ -0,0 +1,56 @@ +/- +Copyright (c) 2018 Louis Carlin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Louis Carlin, Mario Carneiro +-/ +import algebra.euclidean_domain.defs +import data.nat.order +import data.int.order + +/-! +# Instances for Euclidean domains + +* `int.euclidean_domain`: shows that `ℤ` is a Euclidean domain. +* `field.to_euclidean_domain`: shows that any field is a Euclidean domain. + +-/ + +instance int.euclidean_domain : euclidean_domain ℤ := +{ add := (+), + mul := (*), + one := 1, + zero := 0, + neg := has_neg.neg, + quotient := (/), + quotient_zero := int.div_zero, + remainder := (%), + quotient_mul_add_remainder_eq := λ a b, int.div_add_mod _ _, + r := λ a b, a.nat_abs < b.nat_abs, + r_well_founded := measure_wf (λ a, int.nat_abs a), + remainder_lt := λ a b b0, int.coe_nat_lt.1 $ + by { rw [int.nat_abs_of_nonneg (int.mod_nonneg _ b0), ← int.abs_eq_nat_abs], + exact int.mod_lt _ b0 }, + mul_left_not_lt := λ a b b0, not_lt_of_ge $ + by {rw [← mul_one a.nat_abs, int.nat_abs_mul], + exact mul_le_mul_of_nonneg_left (int.nat_abs_pos_of_ne_zero b0) (nat.zero_le _) }, + .. int.comm_ring, + .. int.nontrivial } + +@[priority 100] -- see Note [lower instance priority] +instance field.to_euclidean_domain {K : Type*} [field K] : euclidean_domain K := +{ add := (+), + mul := (*), + one := 1, + zero := 0, + neg := has_neg.neg, + quotient := (/), + remainder := λ a b, a - a * b / b, + quotient_zero := div_zero, + quotient_mul_add_remainder_eq := λ a b, + by { classical, by_cases b = 0; simp [h, mul_div_cancel'] }, + r := λ a b, a = 0 ∧ b ≠ 0, + r_well_founded := well_founded.intro $ λ a, acc.intro _ $ λ b ⟨hb, hna⟩, + acc.intro _ $ λ c ⟨hc, hnb⟩, false.elim $ hnb hb, + remainder_lt := λ a b hnb, by simp [hnb], + mul_left_not_lt := λ a b hnb ⟨hab, hna⟩, or.cases_on (mul_eq_zero.1 hab) hna hnb, + .. ‹field K› } diff --git a/src/algebra/group/opposite.lean b/src/algebra/group/opposite.lean index 5a34c0f781951..ee00fdab71b3d 100644 --- a/src/algebra/group/opposite.lean +++ b/src/algebra/group/opposite.lean @@ -5,7 +5,7 @@ Authors: Kenny Lau -/ import algebra.group.inj_surj import algebra.group.commute -import algebra.hom.equiv +import algebra.hom.equiv.basic import algebra.opposites import data.int.cast.defs diff --git a/src/algebra/group/prod.lean b/src/algebra/group/prod.lean index e4c95bf518499..95dad1f39ae3e 100644 --- a/src/algebra/group/prod.lean +++ b/src/algebra/group/prod.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot, Yury Kudryashov -/ +import algebra.hom.equiv.units import algebra.group.opposite /-! diff --git a/src/algebra/group/ulift.lean b/src/algebra/group/ulift.lean index 6f74e32584af4..bf989a83612f3 100644 --- a/src/algebra/group/ulift.lean +++ b/src/algebra/group/ulift.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import data.int.cast.defs -import algebra.hom.equiv +import algebra.hom.equiv.basic +import algebra.group_with_zero.basic /-! # `ulift` instances for groups and monoids diff --git a/src/algebra/group/with_one.lean b/src/algebra/group/with_one.lean index d066f3b36b293..3129bb96ac504 100644 --- a/src/algebra/group/with_one.lean +++ b/src/algebra/group/with_one.lean @@ -3,10 +3,10 @@ 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 algebra.hom.equiv +import order.bounded_order +import algebra.hom.equiv.basic +import algebra.group_with_zero.units import algebra.ring.defs -import logic.equiv.defs -import logic.equiv.option /-! # Adjoining a zero/one to semigroups and related algebraic structures diff --git a/src/algebra/group_power/basic.lean b/src/algebra/group_power/basic.lean index 774b0049ab5d7..9910c50bcb1b9 100644 --- a/src/algebra/group_power/basic.lean +++ b/src/algebra/group_power/basic.lean @@ -3,7 +3,8 @@ Copyright (c) 2015 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Robert Y. Lewis -/ -import algebra.divisibility +import algebra.divisibility.basic +import algebra.group.commute import algebra.group.type_tags /-! diff --git a/src/algebra/group_with_zero/divisibility.lean b/src/algebra/group_with_zero/divisibility.lean new file mode 100644 index 0000000000000..91c4a44532674 --- /dev/null +++ b/src/algebra/group_with_zero/divisibility.lean @@ -0,0 +1,87 @@ +/- +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, Floris van Doorn, Amelia Livingston, Yury Kudryashov, +Neil Strickland, Aaron Anderson +-/ +import algebra.group_with_zero.basic +import algebra.divisibility.units + +/-! +# Divisibility in groups with zero. + +Lemmas about divisibility in groups and monoids with zero. + +-/ + +variables {α : Type*} + + +section semigroup_with_zero + +variables [semigroup_with_zero α] {a : α} + +theorem eq_zero_of_zero_dvd (h : 0 ∣ a) : a = 0 := +dvd.elim h (λ c H', H'.trans (zero_mul c)) + +/-- Given an element `a` of a commutative semigroup with zero, there exists another element whose + product with zero equals `a` iff `a` equals zero. -/ +@[simp] lemma zero_dvd_iff : 0 ∣ a ↔ a = 0 := +⟨eq_zero_of_zero_dvd, λ h, by { rw h, use 0, simp }⟩ + +@[simp] theorem dvd_zero (a : α) : a ∣ 0 := dvd.intro 0 (by simp) + +end semigroup_with_zero + +/-- Given two elements `b`, `c` of a `cancel_monoid_with_zero` and a nonzero element `a`, + `a*b` divides `a*c` iff `b` divides `c`. -/ +theorem mul_dvd_mul_iff_left [cancel_monoid_with_zero α] {a b c : α} + (ha : a ≠ 0) : a * b ∣ a * c ↔ b ∣ c := +exists_congr $ λ d, by rw [mul_assoc, mul_right_inj' ha] + +/-- Given two elements `a`, `b` of a commutative `cancel_monoid_with_zero` and a nonzero + element `c`, `a*c` divides `b*c` iff `a` divides `b`. -/ +theorem mul_dvd_mul_iff_right [cancel_comm_monoid_with_zero α] {a b c : α} (hc : c ≠ 0) : + a * c ∣ b * c ↔ a ∣ b := +exists_congr $ λ d, by rw [mul_right_comm, mul_left_inj' hc] + +section comm_monoid_with_zero + +variable [comm_monoid_with_zero α] + +/-- `dvd_not_unit a b` expresses that `a` divides `b` "strictly", i.e. that `b` divided by `a` +is not a unit. -/ +def dvd_not_unit (a b : α) : Prop := a ≠ 0 ∧ ∃ x, ¬is_unit x ∧ b = a * x + +lemma dvd_not_unit_of_dvd_of_not_dvd {a b : α} (hd : a ∣ b) (hnd : ¬ b ∣ a) : + dvd_not_unit a b := +begin + split, + { rintro rfl, exact hnd (dvd_zero _) }, + { rcases hd with ⟨c, rfl⟩, + refine ⟨c, _, rfl⟩, + rintro ⟨u, rfl⟩, + simpa using hnd } +end + +end comm_monoid_with_zero + +lemma dvd_and_not_dvd_iff [cancel_comm_monoid_with_zero α] {x y : α} : + x ∣ y ∧ ¬y ∣ x ↔ dvd_not_unit x y := +⟨λ ⟨⟨d, hd⟩, hyx⟩, ⟨λ hx0, by simpa [hx0] using hyx, ⟨d, + mt is_unit_iff_dvd_one.1 (λ ⟨e, he⟩, hyx ⟨e, by rw [hd, mul_assoc, ← he, mul_one]⟩), hd⟩⟩, + λ ⟨hx0, d, hdu, hdx⟩, ⟨⟨d, hdx⟩, λ ⟨e, he⟩, hdu (is_unit_of_dvd_one _ + ⟨e, mul_left_cancel₀ hx0 $ by conv {to_lhs, rw [he, hdx]};simp [mul_assoc]⟩)⟩⟩ + +section monoid_with_zero + +variable [monoid_with_zero α] + +theorem ne_zero_of_dvd_ne_zero {p q : α} (h₁ : q ≠ 0) + (h₂ : p ∣ q) : p ≠ 0 := +begin + rcases h₂ with ⟨u, rfl⟩, + exact left_ne_zero_of_mul h₁, +end + +end monoid_with_zero diff --git a/src/algebra/hom/commute.lean b/src/algebra/hom/commute.lean new file mode 100644 index 0000000000000..5f4ef0a44fc8d --- /dev/null +++ b/src/algebra/hom/commute.lean @@ -0,0 +1,28 @@ +/- +Copyright (c) 2018 Patrick Massot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Massot, Kevin Buzzard, Scott Morrison, Johan Commelin, Chris Hughes, + Johannes Hölzl, Yury Kudryashov +-/ +import algebra.hom.group +import algebra.group.commute + +/-! +# Multiplicative homomorphisms respect semiconjugation and commutation. +-/ + +section commute + +variables {F M N : Type*} [has_mul M] [has_mul N] {a x y : M} + +@[simp, to_additive] +protected lemma semiconj_by.map [mul_hom_class F M N] (h : semiconj_by a x y) (f : F) : + semiconj_by (f a) (f x) (f y) := +by simpa only [semiconj_by, map_mul] using congr_arg f h + +@[simp, to_additive] +protected lemma commute.map [mul_hom_class F M N] (h : commute x y) (f : F) : + commute (f x) (f y) := +h.map f + +end commute diff --git a/src/algebra/hom/equiv.lean b/src/algebra/hom/equiv/basic.lean similarity index 70% rename from src/algebra/hom/equiv.lean rename to src/algebra/hom/equiv/basic.lean index 8d3d4bcd95c3b..cedc374820889 100644 --- a/src/algebra/hom/equiv.lean +++ b/src/algebra/hom/equiv/basic.lean @@ -3,8 +3,9 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ -import algebra.group.type_tags -import algebra.group_with_zero.units +import algebra.hom.group +import data.fun_like.equiv +import logic.equiv.basic import data.pi.algebra /-! @@ -524,74 +525,6 @@ def monoid_hom.to_mul_equiv [mul_one_class M] [mul_one_class N] (f : M →* N) ( right_inv := monoid_hom.congr_fun h₂, map_mul' := f.map_mul } -/-- A group is isomorphic to its group of units. -/ -@[to_additive "An additive group is isomorphic to its group of additive units"] -def to_units [group G] : G ≃* Gˣ := -{ to_fun := λ x, ⟨x, x⁻¹, mul_inv_self _, inv_mul_self _⟩, - inv_fun := coe, - left_inv := λ x, rfl, - right_inv := λ u, units.ext rfl, - map_mul' := λ x y, units.ext rfl } - -@[simp, to_additive] lemma coe_to_units [group G] (g : G) : - (to_units g : G) = g := rfl - -namespace units - -variables [monoid M] [monoid N] [monoid P] - -/-- A multiplicative equivalence of monoids defines a multiplicative equivalence -of their groups of units. -/ -def map_equiv (h : M ≃* N) : Mˣ ≃* Nˣ := -{ inv_fun := map h.symm.to_monoid_hom, - left_inv := λ u, ext $ h.left_inv u, - right_inv := λ u, ext $ h.right_inv u, - .. map h.to_monoid_hom } - -@[simp] -lemma map_equiv_symm (h : M ≃* N) : (map_equiv h).symm = map_equiv h.symm := -rfl - -@[simp] -lemma coe_map_equiv (h : M ≃* N) (x : Mˣ) : (map_equiv h x : N) = h x := -rfl - -/-- Left multiplication by a unit of a monoid is a permutation of the underlying type. -/ -@[to_additive "Left addition of an additive unit is a permutation of the underlying type.", - simps apply {fully_applied := ff}] -def mul_left (u : Mˣ) : equiv.perm M := -{ to_fun := λx, u * x, - inv_fun := λx, ↑u⁻¹ * x, - left_inv := u.inv_mul_cancel_left, - right_inv := u.mul_inv_cancel_left } - -@[simp, to_additive] -lemma mul_left_symm (u : Mˣ) : u.mul_left.symm = u⁻¹.mul_left := -equiv.ext $ λ x, rfl - -@[to_additive] -lemma mul_left_bijective (a : Mˣ) : function.bijective ((*) a : M → M) := -(mul_left a).bijective - -/-- Right multiplication by a unit of a monoid is a permutation of the underlying type. -/ -@[to_additive "Right addition of an additive unit is a permutation of the underlying type.", - simps apply {fully_applied := ff}] -def mul_right (u : Mˣ) : equiv.perm M := -{ to_fun := λx, x * u, - inv_fun := λx, x * ↑u⁻¹, - left_inv := λ x, mul_inv_cancel_right x u, - right_inv := λ x, inv_mul_cancel_right x u } - -@[simp, to_additive] -lemma mul_right_symm (u : Mˣ) : u.mul_right.symm = u⁻¹.mul_right := -equiv.ext $ λ x, rfl - -@[to_additive] -lemma mul_right_bijective (a : Mˣ) : function.bijective ((* a) : M → M) := -(mul_right a).bijective - -end units - namespace equiv section has_involutive_neg @@ -610,173 +543,4 @@ lemma inv_symm : (equiv.inv G).symm = equiv.inv G := rfl end has_involutive_neg -section group -variables [group G] - -/-- Left multiplication in a `group` is a permutation of the underlying type. -/ -@[to_additive "Left addition in an `add_group` is a permutation of the underlying type."] -protected def mul_left (a : G) : perm G := (to_units a).mul_left - -@[simp, to_additive] -lemma coe_mul_left (a : G) : ⇑(equiv.mul_left a) = (*) a := rfl - -/-- Extra simp lemma that `dsimp` can use. `simp` will never use this. -/ -@[simp, nolint simp_nf, - to_additive "Extra simp lemma that `dsimp` can use. `simp` will never use this."] -lemma mul_left_symm_apply (a : G) : ((equiv.mul_left a).symm : G → G) = (*) a⁻¹ := rfl - -@[simp, to_additive] -lemma mul_left_symm (a : G) : (equiv.mul_left a).symm = equiv.mul_left a⁻¹ := -ext $ λ x, rfl - -@[to_additive] -lemma _root_.group.mul_left_bijective (a : G) : function.bijective ((*) a) := -(equiv.mul_left a).bijective - -/-- Right multiplication in a `group` is a permutation of the underlying type. -/ -@[to_additive "Right addition in an `add_group` is a permutation of the underlying type."] -protected def mul_right (a : G) : perm G := (to_units a).mul_right - -@[simp, to_additive] -lemma coe_mul_right (a : G) : ⇑(equiv.mul_right a) = λ x, x * a := rfl - -@[simp, to_additive] -lemma mul_right_symm (a : G) : (equiv.mul_right a).symm = equiv.mul_right a⁻¹ := -ext $ λ x, rfl - -/-- Extra simp lemma that `dsimp` can use. `simp` will never use this. -/ -@[simp, nolint simp_nf, - to_additive "Extra simp lemma that `dsimp` can use. `simp` will never use this."] -lemma mul_right_symm_apply (a : G) : ((equiv.mul_right a).symm : G → G) = λ x, x * a⁻¹ := rfl - -@[to_additive] -lemma _root_.group.mul_right_bijective (a : G) : function.bijective (* a) := -(equiv.mul_right a).bijective - -/-- A version of `equiv.mul_left a b⁻¹` that is defeq to `a / b`. -/ -@[to_additive /-" A version of `equiv.add_left a (-b)` that is defeq to `a - b`. "-/, simps] -protected def div_left (a : G) : G ≃ G := -{ to_fun := λ b, a / b, - inv_fun := λ b, b⁻¹ * a, - left_inv := λ b, by simp [div_eq_mul_inv], - right_inv := λ b, by simp [div_eq_mul_inv] } - -@[to_additive] -lemma div_left_eq_inv_trans_mul_left (a : G) : - equiv.div_left a = (equiv.inv G).trans (equiv.mul_left a) := -ext $ λ _, div_eq_mul_inv _ _ - -/-- A version of `equiv.mul_right a⁻¹ b` that is defeq to `b / a`. -/ -@[to_additive /-" A version of `equiv.add_right (-a) b` that is defeq to `b - a`. "-/, simps] -protected def div_right (a : G) : G ≃ G := -{ to_fun := λ b, b / a, - inv_fun := λ b, b * a, - left_inv := λ b, by simp [div_eq_mul_inv], - right_inv := λ b, by simp [div_eq_mul_inv] } - -@[to_additive] -lemma div_right_eq_mul_right_inv (a : G) : equiv.div_right a = equiv.mul_right a⁻¹ := -ext $ λ _, div_eq_mul_inv _ _ - -end group - -section group_with_zero -variables [group_with_zero G] - -/-- Left multiplication by a nonzero element in a `group_with_zero` is a permutation of the -underlying type. -/ -@[simps {fully_applied := ff}] -protected def mul_left₀ (a : G) (ha : a ≠ 0) : perm G := -(units.mk0 a ha).mul_left - -lemma _root_.mul_left_bijective₀ (a : G) (ha : a ≠ 0) : - function.bijective ((*) a : G → G) := -(equiv.mul_left₀ a ha).bijective - -/-- Right multiplication by a nonzero element in a `group_with_zero` is a permutation of the -underlying type. -/ -@[simps {fully_applied := ff}] -protected def mul_right₀ (a : G) (ha : a ≠ 0) : perm G := -(units.mk0 a ha).mul_right - -lemma _root_.mul_right_bijective₀ (a : G) (ha : a ≠ 0) : - function.bijective ((* a) : G → G) := -(equiv.mul_right₀ a ha).bijective - -end group_with_zero - end equiv - -/-- In a `division_comm_monoid`, `equiv.inv` is a `mul_equiv`. There is a variant of this -`mul_equiv.inv' G : G ≃* Gᵐᵒᵖ` for the non-commutative case. -/ -@[to_additive "When the `add_group` is commutative, `equiv.neg` is an `add_equiv`.", simps apply] -def mul_equiv.inv (G : Type*) [division_comm_monoid G] : G ≃* G := -{ to_fun := has_inv.inv, - inv_fun := has_inv.inv, - map_mul' := mul_inv, - ..equiv.inv G } - -@[simp] lemma mul_equiv.inv_symm (G : Type*) [division_comm_monoid G] : - (mul_equiv.inv G).symm = mul_equiv.inv G := rfl - -section type_tags - -/-- Reinterpret `G ≃+ H` as `multiplicative G ≃* multiplicative H`. -/ -def add_equiv.to_multiplicative [add_zero_class G] [add_zero_class H] : - (G ≃+ H) ≃ (multiplicative G ≃* multiplicative H) := -{ to_fun := λ f, ⟨f.to_add_monoid_hom.to_multiplicative, - f.symm.to_add_monoid_hom.to_multiplicative, f.3, f.4, f.5⟩, - inv_fun := λ f, ⟨f.to_monoid_hom, f.symm.to_monoid_hom, f.3, f.4, f.5⟩, - left_inv := λ x, by { ext, refl, }, - right_inv := λ x, by { ext, refl, }, } - -/-- Reinterpret `G ≃* H` as `additive G ≃+ additive H`. -/ -def mul_equiv.to_additive [mul_one_class G] [mul_one_class H] : - (G ≃* H) ≃ (additive G ≃+ additive H) := -{ to_fun := λ f, ⟨f.to_monoid_hom.to_additive, f.symm.to_monoid_hom.to_additive, f.3, f.4, f.5⟩, - inv_fun := λ f, ⟨f.to_add_monoid_hom, f.symm.to_add_monoid_hom, f.3, f.4, f.5⟩, - left_inv := λ x, by { ext, refl, }, - right_inv := λ x, by { ext, refl, }, } - -/-- Reinterpret `additive G ≃+ H` as `G ≃* multiplicative H`. -/ -def add_equiv.to_multiplicative' [mul_one_class G] [add_zero_class H] : - (additive G ≃+ H) ≃ (G ≃* multiplicative H) := -{ to_fun := λ f, ⟨f.to_add_monoid_hom.to_multiplicative', - f.symm.to_add_monoid_hom.to_multiplicative'', f.3, f.4, f.5⟩, - inv_fun := λ f, ⟨f.to_monoid_hom, f.symm.to_monoid_hom, f.3, f.4, f.5⟩, - left_inv := λ x, by { ext, refl, }, - right_inv := λ x, by { ext, refl, }, } - -/-- Reinterpret `G ≃* multiplicative H` as `additive G ≃+ H` as. -/ -def mul_equiv.to_additive' [mul_one_class G] [add_zero_class H] : - (G ≃* multiplicative H) ≃ (additive G ≃+ H) := -add_equiv.to_multiplicative'.symm - -/-- Reinterpret `G ≃+ additive H` as `multiplicative G ≃* H`. -/ -def add_equiv.to_multiplicative'' [add_zero_class G] [mul_one_class H] : - (G ≃+ additive H) ≃ (multiplicative G ≃* H) := -{ to_fun := λ f, ⟨f.to_add_monoid_hom.to_multiplicative'', - f.symm.to_add_monoid_hom.to_multiplicative', f.3, f.4, f.5⟩, - inv_fun := λ f, ⟨f.to_monoid_hom, f.symm.to_monoid_hom, f.3, f.4, f.5⟩, - left_inv := λ x, by { ext, refl, }, - right_inv := λ x, by { ext, refl, }, } - -/-- Reinterpret `multiplicative G ≃* H` as `G ≃+ additive H` as. -/ -def mul_equiv.to_additive'' [add_zero_class G] [mul_one_class H] : - (multiplicative G ≃* H) ≃ (G ≃+ additive H) := -add_equiv.to_multiplicative''.symm - -end type_tags - -section -variables (G) (H) - -/-- `additive (multiplicative G)` is just `G`. -/ -def add_equiv.additive_multiplicative [add_zero_class G] : additive (multiplicative G) ≃+ G := -mul_equiv.to_additive'' (mul_equiv.refl (multiplicative G)) - -/-- `multiplicative (additive H)` is just `H`. -/ -def mul_equiv.multiplicative_additive [mul_one_class H] : multiplicative (additive H) ≃* H := -add_equiv.to_multiplicative'' (add_equiv.refl (additive H)) - -end diff --git a/src/algebra/hom/equiv/type_tags.lean b/src/algebra/hom/equiv/type_tags.lean new file mode 100644 index 0000000000000..25edf2ef357eb --- /dev/null +++ b/src/algebra/hom/equiv/type_tags.lean @@ -0,0 +1,71 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov +-/ +import algebra.hom.equiv.basic +import algebra.group_with_zero.units + +/-! +# Additive and multiplicative equivalences associated to `multiplicative` and `additive`. +-/ + +variables {G H : Type*} + +/-- Reinterpret `G ≃+ H` as `multiplicative G ≃* multiplicative H`. -/ +def add_equiv.to_multiplicative [add_zero_class G] [add_zero_class H] : + (G ≃+ H) ≃ (multiplicative G ≃* multiplicative H) := +{ to_fun := λ f, ⟨f.to_add_monoid_hom.to_multiplicative, + f.symm.to_add_monoid_hom.to_multiplicative, f.3, f.4, f.5⟩, + inv_fun := λ f, ⟨f.to_monoid_hom, f.symm.to_monoid_hom, f.3, f.4, f.5⟩, + left_inv := λ x, by { ext, refl, }, + right_inv := λ x, by { ext, refl, }, } + +/-- Reinterpret `G ≃* H` as `additive G ≃+ additive H`. -/ +def mul_equiv.to_additive [mul_one_class G] [mul_one_class H] : + (G ≃* H) ≃ (additive G ≃+ additive H) := +{ to_fun := λ f, ⟨f.to_monoid_hom.to_additive, f.symm.to_monoid_hom.to_additive, f.3, f.4, f.5⟩, + inv_fun := λ f, ⟨f.to_add_monoid_hom, f.symm.to_add_monoid_hom, f.3, f.4, f.5⟩, + left_inv := λ x, by { ext, refl, }, + right_inv := λ x, by { ext, refl, }, } + +/-- Reinterpret `additive G ≃+ H` as `G ≃* multiplicative H`. -/ +def add_equiv.to_multiplicative' [mul_one_class G] [add_zero_class H] : + (additive G ≃+ H) ≃ (G ≃* multiplicative H) := +{ to_fun := λ f, ⟨f.to_add_monoid_hom.to_multiplicative', + f.symm.to_add_monoid_hom.to_multiplicative'', f.3, f.4, f.5⟩, + inv_fun := λ f, ⟨f.to_monoid_hom, f.symm.to_monoid_hom, f.3, f.4, f.5⟩, + left_inv := λ x, by { ext, refl, }, + right_inv := λ x, by { ext, refl, }, } + +/-- Reinterpret `G ≃* multiplicative H` as `additive G ≃+ H` as. -/ +def mul_equiv.to_additive' [mul_one_class G] [add_zero_class H] : + (G ≃* multiplicative H) ≃ (additive G ≃+ H) := +add_equiv.to_multiplicative'.symm + +/-- Reinterpret `G ≃+ additive H` as `multiplicative G ≃* H`. -/ +def add_equiv.to_multiplicative'' [add_zero_class G] [mul_one_class H] : + (G ≃+ additive H) ≃ (multiplicative G ≃* H) := +{ to_fun := λ f, ⟨f.to_add_monoid_hom.to_multiplicative'', + f.symm.to_add_monoid_hom.to_multiplicative', f.3, f.4, f.5⟩, + inv_fun := λ f, ⟨f.to_monoid_hom, f.symm.to_monoid_hom, f.3, f.4, f.5⟩, + left_inv := λ x, by { ext, refl, }, + right_inv := λ x, by { ext, refl, }, } + +/-- Reinterpret `multiplicative G ≃* H` as `G ≃+ additive H` as. -/ +def mul_equiv.to_additive'' [add_zero_class G] [mul_one_class H] : + (multiplicative G ≃* H) ≃ (G ≃+ additive H) := +add_equiv.to_multiplicative''.symm + +section +variables (G) (H) + +/-- `additive (multiplicative G)` is just `G`. -/ +def add_equiv.additive_multiplicative [add_zero_class G] : additive (multiplicative G) ≃+ G := +mul_equiv.to_additive'' (mul_equiv.refl (multiplicative G)) + +/-- `multiplicative (additive H)` is just `H`. -/ +def mul_equiv.multiplicative_additive [mul_one_class H] : multiplicative (additive H) ≃* H := +add_equiv.to_multiplicative'' (add_equiv.refl (additive H)) + +end diff --git a/src/algebra/hom/equiv/units.lean b/src/algebra/hom/equiv/units.lean new file mode 100644 index 0000000000000..855e8eef1325d --- /dev/null +++ b/src/algebra/hom/equiv/units.lean @@ -0,0 +1,192 @@ +/- +Copyright (c) 2018 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov +-/ +import algebra.hom.equiv.basic +import algebra.group_with_zero.units + +/-! +# Multiplicative and additive equivalence acting on units. +-/ + +variables {F α β A B M N P Q G H : Type*} + +/-- A group is isomorphic to its group of units. -/ +@[to_additive "An additive group is isomorphic to its group of additive units"] +def to_units [group G] : G ≃* Gˣ := +{ to_fun := λ x, ⟨x, x⁻¹, mul_inv_self _, inv_mul_self _⟩, + inv_fun := coe, + left_inv := λ x, rfl, + right_inv := λ u, units.ext rfl, + map_mul' := λ x y, units.ext rfl } + +@[simp, to_additive] lemma coe_to_units [group G] (g : G) : + (to_units g : G) = g := rfl + +namespace units + +variables [monoid M] [monoid N] [monoid P] + +/-- A multiplicative equivalence of monoids defines a multiplicative equivalence +of their groups of units. -/ +def map_equiv (h : M ≃* N) : Mˣ ≃* Nˣ := +{ inv_fun := map h.symm.to_monoid_hom, + left_inv := λ u, ext $ h.left_inv u, + right_inv := λ u, ext $ h.right_inv u, + .. map h.to_monoid_hom } + +@[simp] +lemma map_equiv_symm (h : M ≃* N) : (map_equiv h).symm = map_equiv h.symm := +rfl + +@[simp] +lemma coe_map_equiv (h : M ≃* N) (x : Mˣ) : (map_equiv h x : N) = h x := +rfl + +/-- Left multiplication by a unit of a monoid is a permutation of the underlying type. -/ +@[to_additive "Left addition of an additive unit is a permutation of the underlying type.", + simps apply {fully_applied := ff}] +def mul_left (u : Mˣ) : equiv.perm M := +{ to_fun := λx, u * x, + inv_fun := λx, ↑u⁻¹ * x, + left_inv := u.inv_mul_cancel_left, + right_inv := u.mul_inv_cancel_left } + +@[simp, to_additive] +lemma mul_left_symm (u : Mˣ) : u.mul_left.symm = u⁻¹.mul_left := +equiv.ext $ λ x, rfl + +@[to_additive] +lemma mul_left_bijective (a : Mˣ) : function.bijective ((*) a : M → M) := +(mul_left a).bijective + +/-- Right multiplication by a unit of a monoid is a permutation of the underlying type. -/ +@[to_additive "Right addition of an additive unit is a permutation of the underlying type.", + simps apply {fully_applied := ff}] +def mul_right (u : Mˣ) : equiv.perm M := +{ to_fun := λx, x * u, + inv_fun := λx, x * ↑u⁻¹, + left_inv := λ x, mul_inv_cancel_right x u, + right_inv := λ x, inv_mul_cancel_right x u } + +@[simp, to_additive] +lemma mul_right_symm (u : Mˣ) : u.mul_right.symm = u⁻¹.mul_right := +equiv.ext $ λ x, rfl + +@[to_additive] +lemma mul_right_bijective (a : Mˣ) : function.bijective ((* a) : M → M) := +(mul_right a).bijective + +end units + +namespace equiv + +section group +variables [group G] + +/-- Left multiplication in a `group` is a permutation of the underlying type. -/ +@[to_additive "Left addition in an `add_group` is a permutation of the underlying type."] +protected def mul_left (a : G) : perm G := (to_units a).mul_left + +@[simp, to_additive] +lemma coe_mul_left (a : G) : ⇑(equiv.mul_left a) = (*) a := rfl + +/-- Extra simp lemma that `dsimp` can use. `simp` will never use this. -/ +@[simp, nolint simp_nf, + to_additive "Extra simp lemma that `dsimp` can use. `simp` will never use this."] +lemma mul_left_symm_apply (a : G) : ((equiv.mul_left a).symm : G → G) = (*) a⁻¹ := rfl + +@[simp, to_additive] +lemma mul_left_symm (a : G) : (equiv.mul_left a).symm = equiv.mul_left a⁻¹ := +ext $ λ x, rfl + +@[to_additive] +lemma _root_.group.mul_left_bijective (a : G) : function.bijective ((*) a) := +(equiv.mul_left a).bijective + +/-- Right multiplication in a `group` is a permutation of the underlying type. -/ +@[to_additive "Right addition in an `add_group` is a permutation of the underlying type."] +protected def mul_right (a : G) : perm G := (to_units a).mul_right + +@[simp, to_additive] +lemma coe_mul_right (a : G) : ⇑(equiv.mul_right a) = λ x, x * a := rfl + +@[simp, to_additive] +lemma mul_right_symm (a : G) : (equiv.mul_right a).symm = equiv.mul_right a⁻¹ := +ext $ λ x, rfl + +/-- Extra simp lemma that `dsimp` can use. `simp` will never use this. -/ +@[simp, nolint simp_nf, + to_additive "Extra simp lemma that `dsimp` can use. `simp` will never use this."] +lemma mul_right_symm_apply (a : G) : ((equiv.mul_right a).symm : G → G) = λ x, x * a⁻¹ := rfl + +@[to_additive] +lemma _root_.group.mul_right_bijective (a : G) : function.bijective (* a) := +(equiv.mul_right a).bijective + +/-- A version of `equiv.mul_left a b⁻¹` that is defeq to `a / b`. -/ +@[to_additive /-" A version of `equiv.add_left a (-b)` that is defeq to `a - b`. "-/, simps] +protected def div_left (a : G) : G ≃ G := +{ to_fun := λ b, a / b, + inv_fun := λ b, b⁻¹ * a, + left_inv := λ b, by simp [div_eq_mul_inv], + right_inv := λ b, by simp [div_eq_mul_inv] } + +@[to_additive] +lemma div_left_eq_inv_trans_mul_left (a : G) : + equiv.div_left a = (equiv.inv G).trans (equiv.mul_left a) := +ext $ λ _, div_eq_mul_inv _ _ + +/-- A version of `equiv.mul_right a⁻¹ b` that is defeq to `b / a`. -/ +@[to_additive /-" A version of `equiv.add_right (-a) b` that is defeq to `b - a`. "-/, simps] +protected def div_right (a : G) : G ≃ G := +{ to_fun := λ b, b / a, + inv_fun := λ b, b * a, + left_inv := λ b, by simp [div_eq_mul_inv], + right_inv := λ b, by simp [div_eq_mul_inv] } + +@[to_additive] +lemma div_right_eq_mul_right_inv (a : G) : equiv.div_right a = equiv.mul_right a⁻¹ := +ext $ λ _, div_eq_mul_inv _ _ + +end group + +section group_with_zero +variables [group_with_zero G] + +/-- Left multiplication by a nonzero element in a `group_with_zero` is a permutation of the +underlying type. -/ +@[simps {fully_applied := ff}] +protected def mul_left₀ (a : G) (ha : a ≠ 0) : perm G := +(units.mk0 a ha).mul_left + +lemma _root_.mul_left_bijective₀ (a : G) (ha : a ≠ 0) : + function.bijective ((*) a : G → G) := +(equiv.mul_left₀ a ha).bijective + +/-- Right multiplication by a nonzero element in a `group_with_zero` is a permutation of the +underlying type. -/ +@[simps {fully_applied := ff}] +protected def mul_right₀ (a : G) (ha : a ≠ 0) : perm G := +(units.mk0 a ha).mul_right + +lemma _root_.mul_right_bijective₀ (a : G) (ha : a ≠ 0) : + function.bijective ((* a) : G → G) := +(equiv.mul_right₀ a ha).bijective + +end group_with_zero + +end equiv + +/-- In a `division_comm_monoid`, `equiv.inv` is a `mul_equiv`. There is a variant of this +`mul_equiv.inv' G : G ≃* Gᵐᵒᵖ` for the non-commutative case. -/ +@[to_additive "When the `add_group` is commutative, `equiv.neg` is an `add_equiv`.", simps apply] +def mul_equiv.inv (G : Type*) [division_comm_monoid G] : G ≃* G := +{ to_fun := has_inv.inv, + inv_fun := has_inv.inv, + map_mul' := mul_inv, + ..equiv.inv G } + +@[simp] lemma mul_equiv.inv_symm (G : Type*) [division_comm_monoid G] : + (mul_equiv.inv G).symm = mul_equiv.inv G := rfl diff --git a/src/algebra/hom/group.lean b/src/algebra/hom/group.lean index d202b69d240a9..6eb978dccd5d9 100644 --- a/src/algebra/hom/group.lean +++ b/src/algebra/hom/group.lean @@ -5,7 +5,7 @@ Authors: Patrick Massot, Kevin Buzzard, Scott Morrison, Johan Commelin, Chris Hu Johannes Hölzl, Yury Kudryashov -/ import algebra.ne_zero -import algebra.group.commute +import algebra.group.basic import algebra.group_with_zero.defs import data.fun_like.basic @@ -377,9 +377,6 @@ attribute [to_additive_reorder 8, to_additive] map_zpow end mul_one -@[to_additive] lemma group.is_unit [group G] (g : G) : is_unit g := -⟨⟨g, g⁻¹, mul_inv_self g, inv_mul_self g⟩, rfl⟩ - section mul_zero_one variables [mul_zero_one_class M] [mul_zero_one_class N] @@ -1089,22 +1086,6 @@ by { ext, simp only [map_one, coe_comp, function.comp_app, one_apply] } g.comp (f₁ * f₂) = g.comp f₁ * g.comp f₂ := by { ext, simp only [mul_apply, function.comp_app, map_mul, coe_comp] } -/-- If two homomorphisms from a division monoid to a monoid are equal at a unit `x`, then they are -equal at `x⁻¹`. -/ -@[to_additive "If two homomorphisms from a subtraction monoid to an additive monoid are equal at an -additive unit `x`, then they are equal at `-x`."] -lemma _root_.is_unit.eq_on_inv {G N} [division_monoid G] [monoid N] [monoid_hom_class F G N] {x : G} - (hx : is_unit x) (f g : F) (h : f x = g x) : f x⁻¹ = g x⁻¹ := -left_inv_eq_right_inv (map_mul_eq_one f hx.inv_mul_cancel) $ - h.symm ▸ map_mul_eq_one g $ hx.mul_inv_cancel - -/-- If two homomorphism from a group to a monoid are equal at `x`, then they are equal at `x⁻¹`. -/ -@[to_additive "If two homomorphism from an additive group to an additive monoid are equal at `x`, -then they are equal at `-x`." ] -lemma _root_.eq_on_inv {G} [group G] [monoid M] [monoid_hom_class F G M] (f g : F) {x : G} - (h : f x = g x) : f x⁻¹ = g x⁻¹ := -(group.is_unit x).eq_on_inv f g h - /-- Group homomorphisms preserve inverse. -/ @[to_additive "Additive group homomorphisms preserve negation."] protected lemma map_inv [group α] [division_monoid β] (f : α →* β) (a : α) : f a⁻¹ = (f a)⁻¹ := @@ -1233,19 +1214,3 @@ instance {M N} {hM : mul_zero_one_class M} [comm_monoid_with_zero N] : has_mul ( { to_fun := λ a, f a * g a, map_zero' := by rw [map_zero, zero_mul], ..(f * g : M →* N) }⟩ - -section commute - -variables [has_mul M] [has_mul N] {a x y : M} - -@[simp, to_additive] -protected lemma semiconj_by.map [mul_hom_class F M N] (h : semiconj_by a x y) (f : F) : - semiconj_by (f a) (f x) (f y) := -by simpa only [semiconj_by, map_mul] using congr_arg f h - -@[simp, to_additive] -protected lemma commute.map [mul_hom_class F M N] (h : commute x y) (f : F) : - commute (f x) (f y) := -h.map f - -end commute diff --git a/src/algebra/hom/ring.lean b/src/algebra/hom/ring.lean index ec0f5ea918f90..3398e23b3dbb6 100644 --- a/src/algebra/hom/ring.lean +++ b/src/algebra/hom/ring.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston, Jireh Loreaux -/ import algebra.ring.basic -import algebra.divisibility +import algebra.divisibility.basic import data.pi.algebra import algebra.hom.units import data.set.basic diff --git a/src/algebra/hom/units.lean b/src/algebra/hom/units.lean index 40746ceff609e..73b4cff34507a 100644 --- a/src/algebra/hom/units.lean +++ b/src/algebra/hom/units.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Chris Hughes, Kevin Buzzard -/ import algebra.hom.group - +import algebra.group.commute /-! # Monoid homomorphisms and units @@ -27,6 +27,29 @@ open function universes u v w +@[to_additive] lemma group.is_unit {G} [group G] (g : G) : is_unit g := +⟨⟨g, g⁻¹, mul_inv_self g, inv_mul_self g⟩, rfl⟩ + +section monoid_hom_class + +/-- If two homomorphisms from a division monoid to a monoid are equal at a unit `x`, then they are +equal at `x⁻¹`. -/ +@[to_additive "If two homomorphisms from a subtraction monoid to an additive monoid are equal at an +additive unit `x`, then they are equal at `-x`."] +lemma is_unit.eq_on_inv {F G N} [division_monoid G] [monoid N] [monoid_hom_class F G N] {x : G} + (hx : is_unit x) (f g : F) (h : f x = g x) : f x⁻¹ = g x⁻¹ := +left_inv_eq_right_inv (map_mul_eq_one f hx.inv_mul_cancel) $ + h.symm ▸ map_mul_eq_one g $ hx.mul_inv_cancel + +/-- If two homomorphism from a group to a monoid are equal at `x`, then they are equal at `x⁻¹`. -/ +@[to_additive "If two homomorphism from an additive group to an additive monoid are equal at `x`, +then they are equal at `-x`." ] +lemma eq_on_inv {F G M} [group G] [monoid M] [monoid_hom_class F G M] (f g : F) {x : G} + (h : f x = g x) : f x⁻¹ = g x⁻¹ := +(group.is_unit x).eq_on_inv f g h + +end monoid_hom_class + namespace units variables {α : Type*} {M : Type u} {N : Type v} {P : Type w} [monoid M] [monoid N] [monoid P] diff --git a/src/algebra/order/euclidean_absolute_value.lean b/src/algebra/order/euclidean_absolute_value.lean index ddfbe8b958802..913d36d4b278b 100644 --- a/src/algebra/order/euclidean_absolute_value.lean +++ b/src/algebra/order/euclidean_absolute_value.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ import algebra.order.absolute_value -import algebra.euclidean_domain +import algebra.euclidean_domain.instances /-! # Euclidean absolute values diff --git a/src/algebra/order/field/basic.lean b/src/algebra/order/field/basic.lean index 0d66775d93bca..de5b9638d3602 100644 --- a/src/algebra/order/field/basic.lean +++ b/src/algebra/order/field/basic.lean @@ -6,6 +6,7 @@ Authors: Robert Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn import order.bounds.order_iso import algebra.field.basic import algebra.order.field.defs +import algebra.order.ring.inj_surj /-! # Linear ordered (semi)fields diff --git a/src/algebra/order/group/abs.lean b/src/algebra/order/group/abs.lean index 05d4c10b3a9f5..96bf83787a035 100644 --- a/src/algebra/order/group/abs.lean +++ b/src/algebra/order/group/abs.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ import algebra.abs -import algebra.order.group.basic -import algebra.order.monoid.canonical +import algebra.order.group.defs +import algebra.order.monoid.canonical.defs /-! # Absolute values in ordered groups. diff --git a/src/algebra/order/group/bounds.lean b/src/algebra/order/group/bounds.lean index 3f42c06822f28..85767cd6b059e 100644 --- a/src/algebra/order/group/bounds.lean +++ b/src/algebra/order/group/bounds.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov -/ import order.bounds.basic -import algebra.order.group.basic +import algebra.order.group.defs /-! # Least upper bound and the greatest lower bound in linear ordered additive commutative groups diff --git a/src/algebra/order/group/basic.lean b/src/algebra/order/group/defs.lean similarity index 94% rename from src/algebra/order/group/basic.lean rename to src/algebra/order/group/defs.lean index cb1a67e54c780..37276e85cfd96 100644 --- a/src/algebra/order/group/basic.lean +++ b/src/algebra/order/group/defs.lean @@ -3,7 +3,9 @@ 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 algebra.order.sub.basic +import order.hom.basic +import algebra.hom.equiv.units -- I would prefer not to be importing this here. (@semorrison) +import algebra.order.sub.defs import algebra.order.monoid.order_dual /-! @@ -483,23 +485,6 @@ alias lt_of_mul_lt_mul_left' ← ordered_comm_group.lt_of_mul_lt_mul_left attribute [to_additive ordered_add_comm_group.lt_of_add_lt_add_left] ordered_comm_group.lt_of_mul_lt_mul_left -/-- Pullback an `ordered_comm_group` under an injective map. -See note [reducible non-instances]. -/ -@[reducible, to_additive function.injective.ordered_add_comm_group -"Pullback an `ordered_add_comm_group` under an injective map."] -def function.injective.ordered_comm_group [ordered_comm_group α] {β : Type*} - [has_one β] [has_mul β] [has_inv β] [has_div β] [has_pow β ℕ] [has_pow β ℤ] - (f : β → α) (hf : function.injective f) (one : f 1 = 1) - (mul : ∀ x y, f (x * y) = f x * f y) - (inv : ∀ x, f (x⁻¹) = (f x)⁻¹) - (div : ∀ x y, f (x / y) = f x / f y) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (zpow : ∀ x (n : ℤ), f (x ^ n) = f x ^ n) : - ordered_comm_group β := -{ ..partial_order.lift f hf, - ..hf.ordered_comm_monoid f one mul npow, - ..hf.comm_group f one mul inv div npow zpow } - /- Most of the lemmas that are primed in this section appear in ordered_field. -/ /- I (DT) did not try to minimise the assumptions. -/ section group @@ -832,23 +817,6 @@ instance linear_ordered_comm_group.to_linear_ordered_cancel_comm_monoid : { le_of_mul_le_mul_left := λ x y z, le_of_mul_le_mul_left', ..‹linear_ordered_comm_group α› } -/-- Pullback a `linear_ordered_comm_group` under an injective map. -See note [reducible non-instances]. -/ -@[reducible, to_additive function.injective.linear_ordered_add_comm_group -"Pullback a `linear_ordered_add_comm_group` under an injective map."] -def function.injective.linear_ordered_comm_group {β : Type*} - [has_one β] [has_mul β] [has_inv β] [has_div β] [has_pow β ℕ] [has_pow β ℤ] - [has_sup β] [has_inf β] (f : β → α) (hf : function.injective f) (one : f 1 = 1) - (mul : ∀ x y, f (x * y) = f x * f y) - (inv : ∀ x, f (x⁻¹) = (f x)⁻¹) - (div : ∀ x y, f (x / y) = f x / f y) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (zpow : ∀ x (n : ℤ), f (x ^ n) = f x ^ n) - (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : - linear_ordered_comm_group β := -{ ..linear_order.lift f hf hsup hinf, - ..hf.ordered_comm_group f one mul inv div npow zpow } - @[to_additive linear_ordered_add_comm_group.add_lt_add_left] lemma linear_ordered_comm_group.mul_lt_mul_left' (a b : α) (h : a < b) (c : α) : c * a < c * b := diff --git a/src/algebra/order/group/densely_ordered.lean b/src/algebra/order/group/densely_ordered.lean index cedc235ee20c2..a477157fd766a 100644 --- a/src/algebra/order/group/densely_ordered.lean +++ b/src/algebra/order/group/densely_ordered.lean @@ -3,8 +3,8 @@ 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 algebra.order.group.basic -import algebra.order.monoid.canonical +import algebra.order.group.defs +import algebra.order.monoid.canonical.defs /-! # Lemmas about densely linearly ordered groups. diff --git a/src/algebra/order/group/inj_surj.lean b/src/algebra/order/group/inj_surj.lean new file mode 100644 index 0000000000000..b44b707501cf9 --- /dev/null +++ b/src/algebra/order/group/inj_surj.lean @@ -0,0 +1,47 @@ +/- +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 algebra.order.group.defs +import algebra.order.monoid.basic + +/-! +# Pull back ordered groups along injective maps. +-/ + +variables {α β : Type*} + +/-- Pullback an `ordered_comm_group` under an injective map. +See note [reducible non-instances]. -/ +@[reducible, to_additive function.injective.ordered_add_comm_group +"Pullback an `ordered_add_comm_group` under an injective map."] +def function.injective.ordered_comm_group [ordered_comm_group α] {β : Type*} + [has_one β] [has_mul β] [has_inv β] [has_div β] [has_pow β ℕ] [has_pow β ℤ] + (f : β → α) (hf : function.injective f) (one : f 1 = 1) + (mul : ∀ x y, f (x * y) = f x * f y) + (inv : ∀ x, f (x⁻¹) = (f x)⁻¹) + (div : ∀ x y, f (x / y) = f x / f y) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (zpow : ∀ x (n : ℤ), f (x ^ n) = f x ^ n) : + ordered_comm_group β := +{ ..partial_order.lift f hf, + ..hf.ordered_comm_monoid f one mul npow, + ..hf.comm_group f one mul inv div npow zpow } + +/-- Pullback a `linear_ordered_comm_group` under an injective map. +See note [reducible non-instances]. -/ +@[reducible, to_additive function.injective.linear_ordered_add_comm_group +"Pullback a `linear_ordered_add_comm_group` under an injective map."] +def function.injective.linear_ordered_comm_group [linear_ordered_comm_group α] {β : Type*} + [has_one β] [has_mul β] [has_inv β] [has_div β] [has_pow β ℕ] [has_pow β ℤ] + [has_sup β] [has_inf β] (f : β → α) (hf : function.injective f) (one : f 1 = 1) + (mul : ∀ x y, f (x * y) = f x * f y) + (inv : ∀ x, f (x⁻¹) = (f x)⁻¹) + (div : ∀ x y, f (x / y) = f x / f y) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (zpow : ∀ x (n : ℤ), f (x ^ n) = f x ^ n) + (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : + linear_ordered_comm_group β := +{ ..linear_order.lift f hf hsup hinf, + ..hf.ordered_comm_group f one mul inv div npow zpow } diff --git a/src/algebra/order/group/prod.lean b/src/algebra/order/group/prod.lean index 2912a9e4474d1..d9605fbd0cb52 100644 --- a/src/algebra/order/group/prod.lean +++ b/src/algebra/order/group/prod.lean @@ -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 algebra.order.group.basic +import algebra.order.group.defs import algebra.order.monoid.prod /-! diff --git a/src/algebra/order/group/type_tags.lean b/src/algebra/order/group/type_tags.lean index 88310593826b7..396e7392adeae 100644 --- a/src/algebra/order/group/type_tags.lean +++ b/src/algebra/order/group/type_tags.lean @@ -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 algebra.order.group.basic +import algebra.order.group.defs import algebra.order.monoid.type_tags /-! # Ordered group structures on `multiplicative α` and `additive α`. -/ diff --git a/src/algebra/order/group/units.lean b/src/algebra/order/group/units.lean index 7a83791187aaf..4a4874d23aedb 100644 --- a/src/algebra/order/group/units.lean +++ b/src/algebra/order/group/units.lean @@ -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 algebra.order.group.basic +import algebra.order.group.defs import algebra.order.monoid.units /-! diff --git a/src/algebra/order/group/with_top.lean b/src/algebra/order/group/with_top.lean index 012955c4aa09c..d1fe799b1b39d 100644 --- a/src/algebra/order/group/with_top.lean +++ b/src/algebra/order/group/with_top.lean @@ -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 algebra.order.group.basic +import algebra.order.group.defs import algebra.order.monoid.with_top /-! diff --git a/src/algebra/order/hom/ring.lean b/src/algebra/order/hom/ring.lean index 12b15382e7496..2f07b52607c51 100644 --- a/src/algebra/order/hom/ring.lean +++ b/src/algebra/order/hom/ring.lean @@ -5,7 +5,7 @@ Authors: Alex J. Best, Yaël Dillies -/ import algebra.order.archimedean import algebra.order.hom.monoid -import algebra.order.ring.basic +import algebra.order.ring.defs import algebra.ring.equiv /-! diff --git a/src/algebra/order/invertible.lean b/src/algebra/order/invertible.lean index e2677da2e6068..d862f4ae21ef8 100644 --- a/src/algebra/order/invertible.lean +++ b/src/algebra/order/invertible.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ -import algebra.order.ring.basic +import algebra.order.ring.defs import algebra.invertible /-! diff --git a/src/algebra/order/monoid/basic.lean b/src/algebra/order/monoid/basic.lean index 7bd4c83de9d34..d93995fa753b2 100644 --- a/src/algebra/order/monoid/basic.lean +++ b/src/algebra/order/monoid/basic.lean @@ -3,22 +3,15 @@ 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 algebra.order.monoid.defs import algebra.group.units -import algebra.group_with_zero.defs import algebra.group.inj_surj -import algebra.order.zero_le_one import order.hom.basic /-! # Ordered monoids -This file develops the basics of ordered monoids. - -## Implementation details - -Unfortunately, the number of `'` appended to lemmas in this file -may differ between the multiplicative and the additive version of a lemma. -The reason is that we did not want to change existing names in the library. +This file develops some additional material on ordered monoids. -/ set_option old_structure_cmd true @@ -27,103 +20,6 @@ open function universe u variables {α : Type u} {β : Type*} -/-- An ordered commutative monoid is a commutative monoid -with a partial order such that `a ≤ b → c * a ≤ c * b` (multiplication is monotone) --/ -@[protect_proj, ancestor comm_monoid partial_order] -class ordered_comm_monoid (α : Type*) extends comm_monoid α, partial_order α := -(mul_le_mul_left : ∀ a b : α, a ≤ b → ∀ c : α, c * a ≤ c * b) - -/-- An ordered (additive) commutative monoid is a commutative monoid - with a partial order such that `a ≤ b → c + a ≤ c + b` (addition is monotone) --/ -@[protect_proj, ancestor add_comm_monoid partial_order] -class ordered_add_comm_monoid (α : Type*) extends add_comm_monoid α, partial_order α := -(add_le_add_left : ∀ a b : α, a ≤ b → ∀ c : α, c + a ≤ c + b) - -attribute [to_additive] ordered_comm_monoid - -section ordered_instances - -@[to_additive] -instance ordered_comm_monoid.to_covariant_class_left (M : Type*) [ordered_comm_monoid M] : - covariant_class M M (*) (≤) := -{ elim := λ a b c bc, ordered_comm_monoid.mul_le_mul_left _ _ bc a } - -/- This instance can be proven with `by apply_instance`. However, `with_bot ℕ` does not -pick up a `covariant_class M M (function.swap (*)) (≤)` instance without it (see PR #7940). -/ -@[to_additive] -instance ordered_comm_monoid.to_covariant_class_right (M : Type*) [ordered_comm_monoid M] : - covariant_class M M (swap (*)) (≤) := -covariant_swap_mul_le_of_covariant_mul_le M - -/- This is not an instance, to avoid creating a loop in the type-class system: in a -`left_cancel_semigroup` with a `partial_order`, assuming `covariant_class M M (*) (≤)` implies -`covariant_class M M (*) (<)`, see `left_cancel_semigroup.covariant_mul_lt_of_covariant_mul_le`. -/ -@[to_additive] lemma has_mul.to_covariant_class_left - (M : Type*) [has_mul M] [partial_order M] [covariant_class M M (*) (<)] : - covariant_class M M (*) (≤) := -⟨covariant_le_of_covariant_lt _ _ _ covariant_class.elim⟩ - -/- This is not an instance, to avoid creating a loop in the type-class system: in a -`right_cancel_semigroup` with a `partial_order`, assuming `covariant_class M M (swap (*)) (<)` -implies `covariant_class M M (swap (*)) (≤)`, see -`right_cancel_semigroup.covariant_swap_mul_lt_of_covariant_swap_mul_le`. -/ -@[to_additive] lemma has_mul.to_covariant_class_right - (M : Type*) [has_mul M] [partial_order M] [covariant_class M M (swap (*)) (<)] : - covariant_class M M (swap (*)) (≤) := -⟨covariant_le_of_covariant_lt _ _ _ covariant_class.elim⟩ - -end ordered_instances - -lemma bit0_pos [ordered_add_comm_monoid α] {a : α} (h : 0 < a) : 0 < bit0 a := -add_pos' h h - -/-- A linearly ordered additive commutative monoid. -/ -@[protect_proj, ancestor linear_order ordered_add_comm_monoid] -class linear_ordered_add_comm_monoid (α : Type*) - extends linear_order α, ordered_add_comm_monoid α. - -/-- A linearly ordered commutative monoid. -/ -@[protect_proj, ancestor linear_order ordered_comm_monoid, to_additive] -class linear_ordered_comm_monoid (α : Type*) - extends linear_order α, ordered_comm_monoid α. - -/-- A linearly ordered commutative monoid with a zero element. -/ -class linear_ordered_comm_monoid_with_zero (α : Type*) - extends linear_ordered_comm_monoid α, comm_monoid_with_zero α := -(zero_le_one : (0 : α) ≤ 1) - -@[priority 100] -instance linear_ordered_comm_monoid_with_zero.zero_le_one_class - [h : linear_ordered_comm_monoid_with_zero α] : zero_le_one_class α := -{ ..h } - -/-- A linearly ordered commutative monoid with an additively absorbing `⊤` element. - Instances should include number systems with an infinite element adjoined.` -/ -@[protect_proj, ancestor linear_ordered_add_comm_monoid has_top] -class linear_ordered_add_comm_monoid_with_top (α : Type*) - extends linear_ordered_add_comm_monoid α, has_top α := -(le_top : ∀ x : α, x ≤ ⊤) -(top_add' : ∀ x : α, ⊤ + x = ⊤) - -@[priority 100] -- see Note [lower instance priority] -instance linear_ordered_add_comm_monoid_with_top.to_order_top (α : Type u) - [h : linear_ordered_add_comm_monoid_with_top α] : order_top α := -{ ..h } - -section linear_ordered_add_comm_monoid_with_top -variables [linear_ordered_add_comm_monoid_with_top α] {a b : α} - -@[simp] -lemma top_add (a : α) : ⊤ + a = ⊤ := linear_ordered_add_comm_monoid_with_top.top_add' a - -@[simp] -lemma add_top (a : α) : a + ⊤ = ⊤ := -trans (add_comm _ _) (top_add _) - -end linear_ordered_add_comm_monoid_with_top - /-- Pullback an `ordered_comm_monoid` under an injective map. See note [reducible non-instances]. -/ @[reducible, to_additive function.injective.ordered_add_comm_monoid diff --git a/src/algebra/order/monoid/cancel/basic.lean b/src/algebra/order/monoid/cancel/basic.lean new file mode 100644 index 0000000000000..e0c897c779871 --- /dev/null +++ b/src/algebra/order/monoid/cancel/basic.lean @@ -0,0 +1,54 @@ +/- +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 algebra.order.monoid.basic +import algebra.order.monoid.cancel.defs + +/-! +# Basic results on ordered cancellative monoids. + +We pull back ordered cancellative monoids along injective maps. +-/ + +universe u +variables {α : Type u} + +open function + +section ordered_cancel_comm_monoid +variables [ordered_cancel_comm_monoid α] + +/-- Pullback an `ordered_cancel_comm_monoid` under an injective map. +See note [reducible non-instances]. -/ +@[reducible, to_additive function.injective.ordered_cancel_add_comm_monoid +"Pullback an `ordered_cancel_add_comm_monoid` under an injective map."] +def function.injective.ordered_cancel_comm_monoid {β : Type*} + [has_one β] [has_mul β] [has_pow β ℕ] + (f : β → α) (hf : function.injective f) (one : f 1 = 1) + (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) : + ordered_cancel_comm_monoid β := +{ le_of_mul_le_mul_left := λ a b c (bc : f (a * b) ≤ f (a * c)), + (mul_le_mul_iff_left (f a)).mp (by rwa [← mul, ← mul]), + ..hf.ordered_comm_monoid f one mul npow } + +end ordered_cancel_comm_monoid + +section linear_ordered_cancel_comm_monoid +variables [linear_ordered_cancel_comm_monoid α] + +/-- Pullback a `linear_ordered_cancel_comm_monoid` under an injective map. +See note [reducible non-instances]. -/ +@[reducible, to_additive function.injective.linear_ordered_cancel_add_comm_monoid +"Pullback a `linear_ordered_cancel_add_comm_monoid` under an injective map."] +def function.injective.linear_ordered_cancel_comm_monoid {β : Type*} + [has_one β] [has_mul β] [has_pow β ℕ] [has_sup β] [has_inf β] + (f : β → α) (hf : function.injective f) (one : f 1 = 1) + (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : + linear_ordered_cancel_comm_monoid β := +{ ..hf.linear_ordered_comm_monoid f one mul npow hsup hinf, + ..hf.ordered_cancel_comm_monoid f one mul npow } + +end linear_ordered_cancel_comm_monoid diff --git a/src/algebra/order/monoid/cancel.lean b/src/algebra/order/monoid/cancel/defs.lean similarity index 69% rename from src/algebra/order/monoid/cancel.lean rename to src/algebra/order/monoid/cancel/defs.lean index 02bc3b4bca0dd..7ba8f9e489a4d 100644 --- a/src/algebra/order/monoid/cancel.lean +++ b/src/algebra/order/monoid/cancel/defs.lean @@ -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 algebra.order.monoid.basic +import algebra.order.monoid.defs /-! # Ordered cancellative monoids @@ -72,19 +72,6 @@ instance ordered_cancel_comm_monoid.to_cancel_comm_monoid : cancel_comm_monoid (le_of_mul_le_mul_left' h.le).antisymm $ le_of_mul_le_mul_left' h.ge, ..‹ordered_cancel_comm_monoid α› } -/-- Pullback an `ordered_cancel_comm_monoid` under an injective map. -See note [reducible non-instances]. -/ -@[reducible, to_additive function.injective.ordered_cancel_add_comm_monoid -"Pullback an `ordered_cancel_add_comm_monoid` under an injective map."] -def function.injective.ordered_cancel_comm_monoid {β : Type*} - [has_one β] [has_mul β] [has_pow β ℕ] - (f : β → α) (hf : function.injective f) (one : f 1 = 1) - (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) : - ordered_cancel_comm_monoid β := -{ le_of_mul_le_mul_left := λ a b c (bc : f (a * b) ≤ f (a * c)), - (mul_le_mul_iff_left (f a)).mp (by rwa [← mul, ← mul]), - ..hf.ordered_comm_monoid f one mul npow } - end ordered_cancel_comm_monoid /-- A linearly ordered cancellative additive commutative monoid @@ -100,21 +87,3 @@ in which multiplication is cancellative and monotone. -/ @[protect_proj, ancestor ordered_cancel_comm_monoid linear_ordered_comm_monoid, to_additive] class linear_ordered_cancel_comm_monoid (α : Type u) extends ordered_cancel_comm_monoid α, linear_ordered_comm_monoid α - -section linear_ordered_cancel_comm_monoid -variables [linear_ordered_cancel_comm_monoid α] - -/-- Pullback a `linear_ordered_cancel_comm_monoid` under an injective map. -See note [reducible non-instances]. -/ -@[reducible, to_additive function.injective.linear_ordered_cancel_add_comm_monoid -"Pullback a `linear_ordered_cancel_add_comm_monoid` under an injective map."] -def function.injective.linear_ordered_cancel_comm_monoid {β : Type*} - [has_one β] [has_mul β] [has_pow β ℕ] [has_sup β] [has_inf β] - (f : β → α) (hf : function.injective f) (one : f 1 = 1) - (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : - linear_ordered_cancel_comm_monoid β := -{ ..hf.linear_ordered_comm_monoid f one mul npow hsup hinf, - ..hf.ordered_cancel_comm_monoid f one mul npow } - -end linear_ordered_cancel_comm_monoid diff --git a/src/algebra/order/monoid/canonical.lean b/src/algebra/order/monoid/canonical/defs.lean similarity index 99% rename from src/algebra/order/monoid/canonical.lean rename to src/algebra/order/monoid/canonical/defs.lean index f84882b55b98f..1962af430782d 100644 --- a/src/algebra/order/monoid/canonical.lean +++ b/src/algebra/order/monoid/canonical/defs.lean @@ -3,8 +3,10 @@ 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 algebra.order.monoid.basic +import order.bounded_order import order.min_max +import algebra.ne_zero +import algebra.order.monoid.defs /-! # Canonically ordered monoids diff --git a/src/algebra/order/monoid/defs.lean b/src/algebra/order/monoid/defs.lean new file mode 100644 index 0000000000000..3ade646971e3b --- /dev/null +++ b/src/algebra/order/monoid/defs.lean @@ -0,0 +1,118 @@ +/- +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 order.bounded_order +import algebra.group_with_zero.defs +import algebra.order.zero_le_one + +/-! +# Ordered monoids + +This file provides the definitions of ordered monoids. + +-/ + +set_option old_structure_cmd true +open function + +universe u +variables {α : Type u} {β : Type*} + +/-- An ordered commutative monoid is a commutative monoid +with a partial order such that `a ≤ b → c * a ≤ c * b` (multiplication is monotone) +-/ +@[protect_proj, ancestor comm_monoid partial_order] +class ordered_comm_monoid (α : Type*) extends comm_monoid α, partial_order α := +(mul_le_mul_left : ∀ a b : α, a ≤ b → ∀ c : α, c * a ≤ c * b) + +/-- An ordered (additive) commutative monoid is a commutative monoid + with a partial order such that `a ≤ b → c + a ≤ c + b` (addition is monotone) +-/ +@[protect_proj, ancestor add_comm_monoid partial_order] +class ordered_add_comm_monoid (α : Type*) extends add_comm_monoid α, partial_order α := +(add_le_add_left : ∀ a b : α, a ≤ b → ∀ c : α, c + a ≤ c + b) + +attribute [to_additive] ordered_comm_monoid + +section ordered_instances + +@[to_additive] +instance ordered_comm_monoid.to_covariant_class_left (M : Type*) [ordered_comm_monoid M] : + covariant_class M M (*) (≤) := +{ elim := λ a b c bc, ordered_comm_monoid.mul_le_mul_left _ _ bc a } + +/- This instance can be proven with `by apply_instance`. However, `with_bot ℕ` does not +pick up a `covariant_class M M (function.swap (*)) (≤)` instance without it (see PR #7940). -/ +@[to_additive] +instance ordered_comm_monoid.to_covariant_class_right (M : Type*) [ordered_comm_monoid M] : + covariant_class M M (swap (*)) (≤) := +covariant_swap_mul_le_of_covariant_mul_le M + +/- This is not an instance, to avoid creating a loop in the type-class system: in a +`left_cancel_semigroup` with a `partial_order`, assuming `covariant_class M M (*) (≤)` implies +`covariant_class M M (*) (<)`, see `left_cancel_semigroup.covariant_mul_lt_of_covariant_mul_le`. -/ +@[to_additive] lemma has_mul.to_covariant_class_left + (M : Type*) [has_mul M] [partial_order M] [covariant_class M M (*) (<)] : + covariant_class M M (*) (≤) := +⟨covariant_le_of_covariant_lt _ _ _ covariant_class.elim⟩ + +/- This is not an instance, to avoid creating a loop in the type-class system: in a +`right_cancel_semigroup` with a `partial_order`, assuming `covariant_class M M (swap (*)) (<)` +implies `covariant_class M M (swap (*)) (≤)`, see +`right_cancel_semigroup.covariant_swap_mul_lt_of_covariant_swap_mul_le`. -/ +@[to_additive] lemma has_mul.to_covariant_class_right + (M : Type*) [has_mul M] [partial_order M] [covariant_class M M (swap (*)) (<)] : + covariant_class M M (swap (*)) (≤) := +⟨covariant_le_of_covariant_lt _ _ _ covariant_class.elim⟩ + +end ordered_instances + +lemma bit0_pos [ordered_add_comm_monoid α] {a : α} (h : 0 < a) : 0 < bit0 a := +add_pos' h h + +/-- A linearly ordered additive commutative monoid. -/ +@[protect_proj, ancestor linear_order ordered_add_comm_monoid] +class linear_ordered_add_comm_monoid (α : Type*) + extends linear_order α, ordered_add_comm_monoid α. + +/-- A linearly ordered commutative monoid. -/ +@[protect_proj, ancestor linear_order ordered_comm_monoid, to_additive] +class linear_ordered_comm_monoid (α : Type*) + extends linear_order α, ordered_comm_monoid α. + +/-- A linearly ordered commutative monoid with a zero element. -/ +class linear_ordered_comm_monoid_with_zero (α : Type*) + extends linear_ordered_comm_monoid α, comm_monoid_with_zero α := +(zero_le_one : (0 : α) ≤ 1) + +@[priority 100] +instance linear_ordered_comm_monoid_with_zero.zero_le_one_class + [h : linear_ordered_comm_monoid_with_zero α] : zero_le_one_class α := +{ ..h } + +/-- A linearly ordered commutative monoid with an additively absorbing `⊤` element. + Instances should include number systems with an infinite element adjoined.` -/ +@[protect_proj, ancestor linear_ordered_add_comm_monoid has_top] +class linear_ordered_add_comm_monoid_with_top (α : Type*) + extends linear_ordered_add_comm_monoid α, has_top α := +(le_top : ∀ x : α, x ≤ ⊤) +(top_add' : ∀ x : α, ⊤ + x = ⊤) + +@[priority 100] -- see Note [lower instance priority] +instance linear_ordered_add_comm_monoid_with_top.to_order_top (α : Type u) + [h : linear_ordered_add_comm_monoid_with_top α] : order_top α := +{ ..h } + +section linear_ordered_add_comm_monoid_with_top +variables [linear_ordered_add_comm_monoid_with_top α] {a b : α} + +@[simp] +lemma top_add (a : α) : ⊤ + a = ⊤ := linear_ordered_add_comm_monoid_with_top.top_add' a + +@[simp] +lemma add_top (a : α) : a + ⊤ = ⊤ := +trans (add_comm _ _) (top_add _) + +end linear_ordered_add_comm_monoid_with_top diff --git a/src/algebra/order/monoid/lemmas.lean b/src/algebra/order/monoid/lemmas.lean index 37beb06d52e1c..1e2e127da7454 100644 --- a/src/algebra/order/monoid/lemmas.lean +++ b/src/algebra/order/monoid/lemmas.lean @@ -5,7 +5,6 @@ Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl, Dami Yuyang Zhao -/ import algebra.covariant_and_contravariant -import order.monotone /-! # Ordered monoids diff --git a/src/algebra/order/monoid/min_max.lean b/src/algebra/order/monoid/min_max.lean index a94dda7230c1d..0bba54e2649b2 100644 --- a/src/algebra/order/monoid/min_max.lean +++ b/src/algebra/order/monoid/min_max.lean @@ -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 algebra.order.monoid.basic +import algebra.order.monoid.defs import order.min_max /-! diff --git a/src/algebra/order/monoid/order_dual.lean b/src/algebra/order/monoid/order_dual.lean index 6802456dfa4ff..eff995e0cfc17 100644 --- a/src/algebra/order/monoid/order_dual.lean +++ b/src/algebra/order/monoid/order_dual.lean @@ -5,7 +5,7 @@ Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ import algebra.group.order_synonym -import algebra.order.monoid.cancel +import algebra.order.monoid.cancel.defs /-! # Ordered monoid structures on the order dual. -/ diff --git a/src/algebra/order/monoid/prod.lean b/src/algebra/order/monoid/prod.lean index af0535d8324d8..e7f0d95f7729e 100644 --- a/src/algebra/order/monoid/prod.lean +++ b/src/algebra/order/monoid/prod.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ import algebra.group.prod -import algebra.order.monoid.cancel -import algebra.order.monoid.canonical +import algebra.order.monoid.cancel.defs +import algebra.order.monoid.canonical.defs /-! # Products of ordered monoids -/ diff --git a/src/algebra/order/monoid/type_tags.lean b/src/algebra/order/monoid/type_tags.lean index da4904966ce75..f0544665e062f 100644 --- a/src/algebra/order/monoid/type_tags.lean +++ b/src/algebra/order/monoid/type_tags.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ import algebra.group.type_tags -import algebra.order.monoid.cancel -import algebra.order.monoid.canonical +import algebra.order.monoid.cancel.defs +import algebra.order.monoid.canonical.defs /-! # Ordered monoid structures on `multiplicative α` and `additive α`. -/ diff --git a/src/algebra/order/monoid/units.lean b/src/algebra/order/monoid/units.lean index d632ac1820813..03f50bd00760e 100644 --- a/src/algebra/order/monoid/units.lean +++ b/src/algebra/order/monoid/units.lean @@ -3,8 +3,10 @@ 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 algebra.order.monoid.basic +import order.hom.basic import order.min_max +import algebra.group.units +import algebra.order.monoid.defs /-! # Units in ordered monoids diff --git a/src/algebra/order/monoid/with_zero.lean b/src/algebra/order/monoid/with_zero.lean index 38922f65d2e3b..e9435c98b9188 100644 --- a/src/algebra/order/monoid/with_zero.lean +++ b/src/algebra/order/monoid/with_zero.lean @@ -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 algebra.order.monoid.canonical +import algebra.order.monoid.canonical.defs import algebra.group.with_one /-! diff --git a/src/algebra/order/nonneg/ring.lean b/src/algebra/order/nonneg/ring.lean index ff60acf5db16d..6913bd3c861c8 100644 --- a/src/algebra/order/nonneg/ring.lean +++ b/src/algebra/order/nonneg/ring.lean @@ -3,7 +3,8 @@ Copyright (c) 2021 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import algebra.order.ring.basic +import algebra.order.ring.defs +import algebra.order.ring.inj_surj import order.complete_lattice_intervals import order.lattice_intervals diff --git a/src/algebra/order/pi.lean b/src/algebra/order/pi.lean index 0086c3a5d5e38..bc62e3627bcc2 100644 --- a/src/algebra/order/pi.lean +++ b/src/algebra/order/pi.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot -/ -import algebra.order.ring.basic +import algebra.order.ring.defs import algebra.ring.pi import tactic.positivity diff --git a/src/algebra/order/positive/ring.lean b/src/algebra/order/positive/ring.lean index 944a0ae69aafe..785d291cae35c 100644 --- a/src/algebra/order/positive/ring.lean +++ b/src/algebra/order/positive/ring.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import algebra.order.ring.basic +import algebra.order.ring.inj_surj /-! # Algebraic structures on the set of positive numbers diff --git a/src/algebra/order/ring/abs.lean b/src/algebra/order/ring/abs.lean index 246d43deaed4f..05e6c1a7efc11 100644 --- a/src/algebra/order/ring/abs.lean +++ b/src/algebra/order/ring/abs.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro -/ -import algebra.order.ring.basic +import algebra.order.ring.defs import algebra.ring.divisibility import algebra.order.group.abs diff --git a/src/algebra/order/ring/canonical.lean b/src/algebra/order/ring/canonical.lean index 3701a2c241283..4c2e7fbca594a 100644 --- a/src/algebra/order/ring/canonical.lean +++ b/src/algebra/order/ring/canonical.lean @@ -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 -/ -import algebra.order.ring.basic +import algebra.order.ring.defs import algebra.order.sub.canonical /-! diff --git a/src/algebra/order/ring/cone.lean b/src/algebra/order/ring/cone.lean new file mode 100644 index 0000000000000..dc7f31fb6a246 --- /dev/null +++ b/src/algebra/order/ring/cone.lean @@ -0,0 +1,77 @@ +/- +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 +-/ +import algebra.order.ring.defs + +/-! +# Constructing an ordered ring from a ring with a specified positive cone. + +-/ + +/-! ### Positive cones -/ + +set_option old_structure_cmd true + +namespace ring + +/-- A positive cone in a ring consists of a positive cone in underlying `add_comm_group`, +which contains `1` and such that the positive elements are closed under multiplication. -/ +@[nolint has_nonempty_instance] +structure positive_cone (α : Type*) [ring α] extends add_comm_group.positive_cone α := +(one_nonneg : nonneg 1) +(mul_pos : ∀ (a b), pos a → pos b → pos (a * b)) + +/-- Forget that a positive cone in a ring respects the multiplicative structure. -/ +add_decl_doc positive_cone.to_positive_cone + +/-- A positive cone in a ring induces a linear order if `1` is a positive element. -/ +@[nolint has_nonempty_instance] +structure total_positive_cone (α : Type*) [ring α] + extends positive_cone α, add_comm_group.total_positive_cone α := +(one_pos : pos 1) + +/-- Forget that a `total_positive_cone` in a ring is total. -/ +add_decl_doc total_positive_cone.to_positive_cone + +/-- Forget that a `total_positive_cone` in a ring respects the multiplicative structure. -/ +add_decl_doc total_positive_cone.to_total_positive_cone + +end ring + +namespace strict_ordered_ring + +open ring + +/-- Construct a `strict_ordered_ring` by designating a positive cone in an existing `ring`. -/ +def mk_of_positive_cone {α : Type*} [ring α] (C : positive_cone α) : strict_ordered_ring α := +{ zero_le_one := by { change C.nonneg (1 - 0), convert C.one_nonneg, simp, }, + mul_pos := λ x y xp yp, begin + change C.pos (x*y - 0), + convert C.mul_pos x y (by { convert xp, simp, }) (by { convert yp, simp, }), + simp, + end, + ..‹ring α›, + ..ordered_add_comm_group.mk_of_positive_cone C.to_positive_cone } + +end strict_ordered_ring + +namespace linear_ordered_ring + +open ring + +/-- Construct a `linear_ordered_ring` by +designating a positive cone in an existing `ring`. -/ +def mk_of_positive_cone {α : Type*} [ring α] (C : total_positive_cone α) : + linear_ordered_ring α := +{ exists_pair_ne := ⟨0, 1, begin + intro h, + have one_pos := C.one_pos, + rw [←h, C.pos_iff] at one_pos, + simpa using one_pos, + end⟩, + ..strict_ordered_ring.mk_of_positive_cone C.to_positive_cone, + ..linear_ordered_add_comm_group.mk_of_positive_cone C.to_total_positive_cone, } + +end linear_ordered_ring diff --git a/src/algebra/order/ring/basic.lean b/src/algebra/order/ring/defs.lean similarity index 75% rename from src/algebra/order/ring/basic.lean rename to src/algebra/order/ring/defs.lean index 69cc1d4372506..4a5905e0574dc 100644 --- a/src/algebra/order/ring/basic.lean +++ b/src/algebra/order/ring/defs.lean @@ -4,7 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro -/ import order.min_max -import algebra.order.group.basic +import algebra.order.monoid.cancel.defs +import algebra.order.sub.defs +import algebra.order.group.defs import algebra.order.ring.lemmas /-! @@ -44,10 +46,6 @@ For short, * `canonically_ordered_comm_semiring`: Commutative semiring with a partial order such that `+` respects `≤`, `*` respects `<`, and `a ≤ b ↔ ∃ c, b = a + c`. -and some typeclasses to define ordered rings by specifying their nonnegative elements: -* `nonneg_ring`: To define `ordered_ring`s. -* `linear_nonneg_ring`: To define `linear_ordered_ring`s. - ## Hierarchy The hardest part of proving order lemmas might be to figure out the correct generality and its @@ -968,242 +966,3 @@ max_le (by simpa [mul_comm, max_comm] using cd) end linear_ordered_comm_ring - -namespace function.injective - -/-- Pullback an `ordered_semiring` under an injective map. -/ -@[reducible] -- See note [reducible non-instances] -protected def ordered_semiring [ordered_semiring α] [has_zero β] [has_one β] [has_add β] [has_mul β] - [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) - (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) : - ordered_semiring β := -{ zero_le_one := show f 0 ≤ f 1, by simp only [zero, one, zero_le_one], - mul_le_mul_of_nonneg_left := λ a b c h hc, show f (c * a) ≤ f (c * b), - by { rw [mul, mul], refine mul_le_mul_of_nonneg_left h _, rwa ←zero }, - mul_le_mul_of_nonneg_right := λ a b c h hc, show f (a * c) ≤ f (b * c), - by { rw [mul, mul], refine mul_le_mul_of_nonneg_right h _, rwa ←zero }, - ..hf.ordered_add_comm_monoid f zero add nsmul, - ..hf.semiring f zero one add mul nsmul npow nat_cast } - -/-- Pullback an `ordered_comm_semiring` under an injective map. -/ -@[reducible] -- See note [reducible non-instances] -protected def ordered_comm_semiring [ordered_comm_semiring α] [has_zero β] [has_one β] [has_add β] - [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] (f : β → α) (hf : injective f) - (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) - (mul : ∀ x y, f (x * y) = f x * f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) (nat_cast : ∀ n : ℕ, f n = n) : - ordered_comm_semiring β := -{ ..hf.comm_semiring f zero one add mul nsmul npow nat_cast, - ..hf.ordered_semiring f zero one add mul nsmul npow nat_cast } - -/-- Pullback an `ordered_ring` under an injective map. -/ -@[reducible] -- See note [reducible non-instances] -protected def ordered_ring [ordered_ring α] [has_zero β] [has_one β] [has_add β] [has_mul β] - [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] - [has_int_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : - ordered_ring β := -{ mul_nonneg := λ a b ha hb, show f 0 ≤ f (a * b), - by { rw [zero, mul], apply mul_nonneg; rwa ← zero }, - ..hf.ordered_semiring f zero one add mul nsmul npow nat_cast, - ..hf.ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } - -/-- Pullback an `ordered_comm_ring` under an injective map. -/ -@[reducible] -- See note [reducible non-instances] -protected def ordered_comm_ring [ordered_comm_ring α] [has_zero β] [has_one β] [has_add β] - [has_mul β] [has_neg β] [has_sub β] [has_pow β ℕ] [has_smul ℕ β] [has_smul ℤ β] [has_nat_cast β] - [has_int_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : - ordered_comm_ring β := -{ ..hf.ordered_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast, - ..hf.comm_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } - -/-- Pullback a `strict_ordered_semiring` under an injective map. -/ -@[reducible] -- See note [reducible non-instances] -protected def strict_ordered_semiring [strict_ordered_semiring α] [has_zero β] [has_one β] - [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] (f : β → α) - (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) - (mul : ∀ x y, f (x * y) = f x * f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) (nat_cast : ∀ n : ℕ, f n = n) : - strict_ordered_semiring β := -{ mul_lt_mul_of_pos_left := λ a b c h hc, show f (c * a) < f (c * b), - by simpa only [mul, zero] using mul_lt_mul_of_pos_left ‹f a < f b› (by rwa ←zero), - mul_lt_mul_of_pos_right := λ a b c h hc, show f (a * c) < f (b * c), - by simpa only [mul, zero] using mul_lt_mul_of_pos_right ‹f a < f b› (by rwa ←zero), - ..hf.ordered_cancel_add_comm_monoid f zero add nsmul, - ..hf.ordered_semiring f zero one add mul nsmul npow nat_cast } - -/-- Pullback a `strict_ordered_comm_semiring` under an injective map. -/ -@[reducible] -- See note [reducible non-instances] -protected def strict_ordered_comm_semiring [strict_ordered_comm_semiring α] [has_zero β] [has_one β] - [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] (f : β → α) - (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) - (mul : ∀ x y, f (x * y) = f x * f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) (nat_cast : ∀ n : ℕ, f n = n) : - strict_ordered_comm_semiring β := -{ ..hf.comm_semiring f zero one add mul nsmul npow nat_cast, - ..hf.strict_ordered_semiring f zero one add mul nsmul npow nat_cast } - -/-- Pullback a `strict_ordered_ring` under an injective map. -/ -@[reducible] -- See note [reducible non-instances] -protected def strict_ordered_ring [strict_ordered_ring α] [has_zero β] [has_one β] [has_add β] - [has_mul β] [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] - [has_int_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : - strict_ordered_ring β := -{ mul_pos := λ a b a0 b0, show f 0 < f (a * b), by { rw [zero, mul], apply mul_pos; rwa ← zero }, - ..hf.strict_ordered_semiring f zero one add mul nsmul npow nat_cast, - ..hf.ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } - -/-- Pullback a `strict_ordered_comm_ring` under an injective map. -/ -@[reducible] -- See note [reducible non-instances] -protected def strict_ordered_comm_ring [strict_ordered_comm_ring α] [has_zero β] - [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_pow β ℕ] [has_smul ℕ β] - [has_smul ℤ β] [has_nat_cast β] [has_int_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) - (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : - strict_ordered_comm_ring β := -{ ..hf.strict_ordered_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast, - ..hf.comm_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } - -/-- Pullback a `linear_ordered_semiring` under an injective map. -/ -@[reducible] -- See note [reducible non-instances] -protected def linear_ordered_semiring [linear_ordered_semiring α] [has_zero β] [has_one β] - [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] [has_sup β] [has_inf β] - (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) - (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : - linear_ordered_semiring β := -{ .. linear_order.lift f hf hsup hinf, - .. pullback_nonzero f zero one, - .. hf.strict_ordered_semiring f zero one add mul nsmul npow nat_cast } - -/-- Pullback a `linear_ordered_semiring` under an injective map. -/ -@[reducible] -- See note [reducible non-instances] -protected def linear_ordered_comm_semiring [linear_ordered_comm_semiring α] - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] - [has_sup β] [has_inf β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) - (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : - linear_ordered_comm_semiring β := -{ ..hf.linear_ordered_semiring f zero one add mul nsmul npow nat_cast hsup hinf, - ..hf.strict_ordered_comm_semiring f zero one add mul nsmul npow nat_cast } - -/-- Pullback a `linear_ordered_ring` under an injective map. -/ -@[reducible] -- See note [reducible non-instances] -def linear_ordered_ring [linear_ordered_ring α] [has_zero β] [has_one β] [has_add β] [has_mul β] - [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] - [has_int_cast β] [has_sup β] [has_inf β] (f : β → α) (hf : injective f) (zero : f 0 = 0) - (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) - (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : - linear_ordered_ring β := -{ .. linear_order.lift f hf hsup hinf, - .. pullback_nonzero f zero one, - .. hf.strict_ordered_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } - -/-- Pullback a `linear_ordered_comm_ring` under an injective map. -/ -@[reducible] -- See note [reducible non-instances] -protected def linear_ordered_comm_ring [linear_ordered_comm_ring α] [has_zero β] - [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_pow β ℕ] [has_smul ℕ β] - [has_smul ℤ β] [has_nat_cast β] [has_int_cast β] [has_sup β] [has_inf β] (f : β → α) - (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) - (mul : ∀ x y, f (x * y) = f x * f y) (neg : ∀ x, f (-x) = -f x) - (sub : ∀ x y, f (x - y) = f x - f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) - (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) - (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : - linear_ordered_comm_ring β := -{ .. linear_order.lift f hf hsup hinf, - .. pullback_nonzero f zero one, - .. hf.strict_ordered_comm_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } - -end function.injective - -/-! ### Positive cones -/ - -namespace ring - -/-- A positive cone in a ring consists of a positive cone in underlying `add_comm_group`, -which contains `1` and such that the positive elements are closed under multiplication. -/ -@[nolint has_nonempty_instance] -structure positive_cone (α : Type*) [ring α] extends add_comm_group.positive_cone α := -(one_nonneg : nonneg 1) -(mul_pos : ∀ (a b), pos a → pos b → pos (a * b)) - -/-- Forget that a positive cone in a ring respects the multiplicative structure. -/ -add_decl_doc positive_cone.to_positive_cone - -/-- A positive cone in a ring induces a linear order if `1` is a positive element. -/ -@[nolint has_nonempty_instance] -structure total_positive_cone (α : Type*) [ring α] - extends positive_cone α, add_comm_group.total_positive_cone α := -(one_pos : pos 1) - -/-- Forget that a `total_positive_cone` in a ring is total. -/ -add_decl_doc total_positive_cone.to_positive_cone - -/-- Forget that a `total_positive_cone` in a ring respects the multiplicative structure. -/ -add_decl_doc total_positive_cone.to_total_positive_cone - -end ring - -namespace strict_ordered_ring - -open ring - -/-- Construct a `strict_ordered_ring` by designating a positive cone in an existing `ring`. -/ -def mk_of_positive_cone {α : Type*} [ring α] (C : positive_cone α) : strict_ordered_ring α := -{ zero_le_one := by { change C.nonneg (1 - 0), convert C.one_nonneg, simp, }, - mul_pos := λ x y xp yp, begin - change C.pos (x*y - 0), - convert C.mul_pos x y (by { convert xp, simp, }) (by { convert yp, simp, }), - simp, - end, - ..‹ring α›, - ..ordered_add_comm_group.mk_of_positive_cone C.to_positive_cone } - -end strict_ordered_ring - -namespace linear_ordered_ring - -open ring - -/-- Construct a `linear_ordered_ring` by -designating a positive cone in an existing `ring`. -/ -def mk_of_positive_cone {α : Type*} [ring α] (C : total_positive_cone α) : - linear_ordered_ring α := -{ exists_pair_ne := ⟨0, 1, begin - intro h, - have one_pos := C.one_pos, - rw [←h, C.pos_iff] at one_pos, - simpa using one_pos, - end⟩, - ..strict_ordered_ring.mk_of_positive_cone C.to_positive_cone, - ..linear_ordered_add_comm_group.mk_of_positive_cone C.to_total_positive_cone, } - -end linear_ordered_ring diff --git a/src/algebra/order/ring/inj_surj.lean b/src/algebra/order/ring/inj_surj.lean new file mode 100644 index 0000000000000..1cefd24e6fd6b --- /dev/null +++ b/src/algebra/order/ring/inj_surj.lean @@ -0,0 +1,194 @@ +/- +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 +-/ +import algebra.order.ring.defs +import algebra.order.group.inj_surj +import algebra.order.monoid.cancel.basic +import algebra.ring.inj_surj + +/-! +# Pulling back ordered rings along injective maps. + +-/ + +open function + +universe u +variables {α : Type u} {β : Type*} + +namespace function.injective + +/-- Pullback an `ordered_semiring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def ordered_semiring [ordered_semiring α] [has_zero β] [has_one β] [has_add β] [has_mul β] + [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) + (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) : + ordered_semiring β := +{ zero_le_one := show f 0 ≤ f 1, by simp only [zero, one, zero_le_one], + mul_le_mul_of_nonneg_left := λ a b c h hc, show f (c * a) ≤ f (c * b), + by { rw [mul, mul], refine mul_le_mul_of_nonneg_left h _, rwa ←zero }, + mul_le_mul_of_nonneg_right := λ a b c h hc, show f (a * c) ≤ f (b * c), + by { rw [mul, mul], refine mul_le_mul_of_nonneg_right h _, rwa ←zero }, + ..hf.ordered_add_comm_monoid f zero add nsmul, + ..hf.semiring f zero one add mul nsmul npow nat_cast } + +/-- Pullback an `ordered_comm_semiring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def ordered_comm_semiring [ordered_comm_semiring α] [has_zero β] [has_one β] [has_add β] + [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] (f : β → α) (hf : injective f) + (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) + (mul : ∀ x y, f (x * y) = f x * f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) (nat_cast : ∀ n : ℕ, f n = n) : + ordered_comm_semiring β := +{ ..hf.comm_semiring f zero one add mul nsmul npow nat_cast, + ..hf.ordered_semiring f zero one add mul nsmul npow nat_cast } + +/-- Pullback an `ordered_ring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def ordered_ring [ordered_ring α] [has_zero β] [has_one β] [has_add β] [has_mul β] + [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] + [has_int_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + ordered_ring β := +{ mul_nonneg := λ a b ha hb, show f 0 ≤ f (a * b), + by { rw [zero, mul], apply mul_nonneg; rwa ← zero }, + ..hf.ordered_semiring f zero one add mul nsmul npow nat_cast, + ..hf.ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } + +/-- Pullback an `ordered_comm_ring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def ordered_comm_ring [ordered_comm_ring α] [has_zero β] [has_one β] [has_add β] + [has_mul β] [has_neg β] [has_sub β] [has_pow β ℕ] [has_smul ℕ β] [has_smul ℤ β] [has_nat_cast β] + [has_int_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + ordered_comm_ring β := +{ ..hf.ordered_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast, + ..hf.comm_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } + +/-- Pullback a `strict_ordered_semiring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def strict_ordered_semiring [strict_ordered_semiring α] [has_zero β] [has_one β] + [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] (f : β → α) + (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) + (mul : ∀ x y, f (x * y) = f x * f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) (nat_cast : ∀ n : ℕ, f n = n) : + strict_ordered_semiring β := +{ mul_lt_mul_of_pos_left := λ a b c h hc, show f (c * a) < f (c * b), + by simpa only [mul, zero] using mul_lt_mul_of_pos_left ‹f a < f b› (by rwa ←zero), + mul_lt_mul_of_pos_right := λ a b c h hc, show f (a * c) < f (b * c), + by simpa only [mul, zero] using mul_lt_mul_of_pos_right ‹f a < f b› (by rwa ←zero), + ..hf.ordered_cancel_add_comm_monoid f zero add nsmul, + ..hf.ordered_semiring f zero one add mul nsmul npow nat_cast } + +/-- Pullback a `strict_ordered_comm_semiring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def strict_ordered_comm_semiring [strict_ordered_comm_semiring α] [has_zero β] [has_one β] + [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] (f : β → α) + (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) + (mul : ∀ x y, f (x * y) = f x * f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) (nat_cast : ∀ n : ℕ, f n = n) : + strict_ordered_comm_semiring β := +{ ..hf.comm_semiring f zero one add mul nsmul npow nat_cast, + ..hf.strict_ordered_semiring f zero one add mul nsmul npow nat_cast } + +/-- Pullback a `strict_ordered_ring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def strict_ordered_ring [strict_ordered_ring α] [has_zero β] [has_one β] [has_add β] + [has_mul β] [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] + [has_int_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + strict_ordered_ring β := +{ mul_pos := λ a b a0 b0, show f 0 < f (a * b), by { rw [zero, mul], apply mul_pos; rwa ← zero }, + ..hf.strict_ordered_semiring f zero one add mul nsmul npow nat_cast, + ..hf.ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } + +/-- Pullback a `strict_ordered_comm_ring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def strict_ordered_comm_ring [strict_ordered_comm_ring α] [has_zero β] + [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_pow β ℕ] [has_smul ℕ β] + [has_smul ℤ β] [has_nat_cast β] [has_int_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) + (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + strict_ordered_comm_ring β := +{ ..hf.strict_ordered_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast, + ..hf.comm_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } + +/-- Pullback a `linear_ordered_semiring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def linear_ordered_semiring [linear_ordered_semiring α] [has_zero β] [has_one β] + [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] [has_sup β] [has_inf β] + (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) + (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : + linear_ordered_semiring β := +{ .. linear_order.lift f hf hsup hinf, + .. pullback_nonzero f zero one, + .. hf.strict_ordered_semiring f zero one add mul nsmul npow nat_cast } + +/-- Pullback a `linear_ordered_semiring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def linear_ordered_comm_semiring [linear_ordered_comm_semiring α] + [has_zero β] [has_one β] [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] + [has_sup β] [has_inf β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) + (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : + linear_ordered_comm_semiring β := +{ ..hf.linear_ordered_semiring f zero one add mul nsmul npow nat_cast hsup hinf, + ..hf.strict_ordered_comm_semiring f zero one add mul nsmul npow nat_cast } + +/-- Pullback a `linear_ordered_ring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +def linear_ordered_ring [linear_ordered_ring α] [has_zero β] [has_one β] [has_add β] [has_mul β] + [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] + [has_int_cast β] [has_sup β] [has_inf β] (f : β → α) (hf : injective f) (zero : f 0 = 0) + (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) + (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : + linear_ordered_ring β := +{ .. linear_order.lift f hf hsup hinf, + .. pullback_nonzero f zero one, + .. hf.strict_ordered_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } + +/-- Pullback a `linear_ordered_comm_ring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def linear_ordered_comm_ring [linear_ordered_comm_ring α] [has_zero β] + [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_pow β ℕ] [has_smul ℕ β] + [has_smul ℤ β] [has_nat_cast β] [has_int_cast β] [has_sup β] [has_inf β] (f : β → α) + (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) + (mul : ∀ x y, f (x * y) = f x * f y) (neg : ∀ x, f (-x) = -f x) + (sub : ∀ x y, f (x - y) = f x - f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) + (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) + (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : + linear_ordered_comm_ring β := +{ .. linear_order.lift f hf hsup hinf, + .. pullback_nonzero f zero one, + .. hf.strict_ordered_comm_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } + +end function.injective diff --git a/src/algebra/order/ring/nontrivial.lean b/src/algebra/order/ring/nontrivial.lean index cb5d70b0cd9bf..d09ce734c7d83 100644 --- a/src/algebra/order/ring/nontrivial.lean +++ b/src/algebra/order/ring/nontrivial.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro -/ import algebra.char_zero.defs -import algebra.order.ring.basic +import algebra.order.ring.defs /-! # Nontrivial strict ordered semirings (and hence linear ordered semirings) are characteristic zero. diff --git a/src/algebra/order/sub/basic.lean b/src/algebra/order/sub/basic.lean index 306d4427a8276..35a78f620700f 100644 --- a/src/algebra/order/sub/basic.lean +++ b/src/algebra/order/sub/basic.lean @@ -4,67 +4,21 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ import order.hom.basic -import algebra.hom.equiv +import algebra.hom.equiv.basic import algebra.ring.basic +import algebra.order.sub.defs /-! -# Ordered Subtraction +# Additional results about ordered Subtraction -This file proves lemmas relating (truncated) subtraction with an order. We provide a class -`has_ordered_sub` stating that `a - b ≤ c ↔ a ≤ c + b`. - -The subtraction discussed here could both be normal subtraction in an additive group or truncated -subtraction on a canonically ordered monoid (`ℕ`, `multiset`, `part_enat`, `ennreal`, ...) - -## Implementation details - -`has_ordered_sub` is a mixin type-class, so that we can use the results in this file even in cases -where we don't have a `canonically_ordered_add_monoid` instance -(even though that is our main focus). Conversely, this means we can use -`canonically_ordered_add_monoid` without necessarily having to define a subtraction. - -The results in this file are ordered by the type-class assumption needed to prove it. -This means that similar results might not be close to each other. Furthermore, we don't prove -implications if a bi-implication can be proven under the same assumptions. - -Lemmas using this class are named using `tsub` instead of `sub` (short for "truncated subtraction"). -This is to avoid naming conflicts with similar lemmas about ordered groups. - -We provide a second version of most results that require `[contravariant_class α α (+) (≤)]`. In the -second version we replace this type-class assumption by explicit `add_le_cancellable` assumptions. - -TODO: maybe we should make a multiplicative version of this, so that we can replace some identical -lemmas about subtraction/division in `ordered_[add_]comm_group` with these. - -TODO: generalize `nat.le_of_le_of_sub_le_sub_right`, `nat.sub_le_sub_right_iff`, - `nat.mul_self_sub_mul_self_eq` -/ variables {α β : Type*} -/-- `has_ordered_sub α` means that `α` has a subtraction characterized by `a - b ≤ c ↔ a ≤ c + b`. -In other words, `a - b` is the least `c` such that `a ≤ b + c`. - -This is satisfied both by the subtraction in additive ordered groups and by truncated subtraction -in canonically ordered monoids on many specific types. --/ -class has_ordered_sub (α : Type*) [has_le α] [has_add α] [has_sub α] := -(tsub_le_iff_right : ∀ a b c : α, a - b ≤ c ↔ a ≤ c + b) - section has_add variables [preorder α] [has_add α] [has_sub α] [has_ordered_sub α] {a b c d : α} -@[simp] lemma tsub_le_iff_right : a - b ≤ c ↔ a ≤ c + b := -has_ordered_sub.tsub_le_iff_right a b c - -/-- See `add_tsub_cancel_right` for the equality if `contravariant_class α α (+) (≤)`. -/ -lemma add_tsub_le_right : a + b - b ≤ a := -tsub_le_iff_right.mpr le_rfl - -lemma le_tsub_add : b ≤ (b - a) + a := -tsub_le_iff_right.mp le_rfl - lemma add_hom.le_map_tsub [preorder β] [has_add β] [has_sub β] [has_ordered_sub β] (f : add_hom α β) (hf : monotone f) (a b : α) : f a - f b ≤ f (a - b) := @@ -97,300 +51,14 @@ end /-! ### Preorder -/ -section ordered_add_comm_semigroup - section preorder variables [preorder α] -section add_comm_semigroup -variables [add_comm_semigroup α] [has_sub α] [has_ordered_sub α] {a b c d : α} - -lemma tsub_le_iff_left : a - b ≤ c ↔ a ≤ b + c := -by rw [tsub_le_iff_right, add_comm] - -lemma le_add_tsub : a ≤ b + (a - b) := -tsub_le_iff_left.mp le_rfl - -/-- See `add_tsub_cancel_left` for the equality if `contravariant_class α α (+) (≤)`. -/ -lemma add_tsub_le_left : a + b - a ≤ b := -tsub_le_iff_left.mpr le_rfl - -lemma tsub_le_tsub_right (h : a ≤ b) (c : α) : a - c ≤ b - c := -tsub_le_iff_left.mpr $ h.trans le_add_tsub - -lemma tsub_le_iff_tsub_le : a - b ≤ c ↔ a - c ≤ b := -by rw [tsub_le_iff_left, tsub_le_iff_right] - -/-- See `tsub_tsub_cancel_of_le` for the equality. -/ -lemma tsub_tsub_le : b - (b - a) ≤ a := -tsub_le_iff_right.mpr le_add_tsub - -section cov -variable [covariant_class α α (+) (≤)] - -lemma tsub_le_tsub_left (h : a ≤ b) (c : α) : c - b ≤ c - a := -tsub_le_iff_left.mpr $ le_add_tsub.trans $ add_le_add_right h _ - -lemma tsub_le_tsub (hab : a ≤ b) (hcd : c ≤ d) : a - d ≤ b - c := -(tsub_le_tsub_right hab _).trans $ tsub_le_tsub_left hcd _ - -lemma antitone_const_tsub : antitone (λ x, c - x) := -λ x y hxy, tsub_le_tsub rfl.le hxy - -/-- See `add_tsub_assoc_of_le` for the equality. -/ -lemma add_tsub_le_assoc : a + b - c ≤ a + (b - c) := -by { rw [tsub_le_iff_left, add_left_comm], exact add_le_add_left le_add_tsub a } - -/-- See `tsub_add_eq_add_tsub` for the equality. -/ -lemma add_tsub_le_tsub_add : a + b - c ≤ a - c + b := -by { rw [add_comm, add_comm _ b], exact add_tsub_le_assoc } - -lemma add_le_add_add_tsub : a + b ≤ (a + c) + (b - c) := -by { rw [add_assoc], exact add_le_add_left le_add_tsub a } - -lemma le_tsub_add_add : a + b ≤ (a - c) + (b + c) := -by { rw [add_comm a, add_comm (a - c)], exact add_le_add_add_tsub } - -lemma tsub_le_tsub_add_tsub : a - c ≤ (a - b) + (b - c) := -begin - rw [tsub_le_iff_left, ← add_assoc, add_right_comm], - exact le_add_tsub.trans (add_le_add_right le_add_tsub _), -end - -lemma tsub_tsub_tsub_le_tsub : (c - a) - (c - b) ≤ b - a := -begin - rw [tsub_le_iff_left, tsub_le_iff_left, add_left_comm], - exact le_tsub_add.trans (add_le_add_left le_add_tsub _), -end - -lemma tsub_tsub_le_tsub_add {a b c : α} : a - (b - c) ≤ a - b + c := -tsub_le_iff_right.2 $ calc - a ≤ a - b + b : le_tsub_add - ... ≤ a - b + (c + (b - c)) : add_le_add_left le_add_tsub _ - ... = a - b + c + (b - c) : (add_assoc _ _ _).symm - -/-- See `tsub_add_tsub_comm` for the equality. -/ -lemma add_tsub_add_le_tsub_add_tsub : a + b - (c + d) ≤ a - c + (b - d) := -begin - rw [add_comm c, tsub_le_iff_left, add_assoc, ←tsub_le_iff_left, ←tsub_le_iff_left], - refine (tsub_le_tsub_right add_tsub_le_assoc c).trans _, - rw [add_comm a, add_comm (a - c)], - exact add_tsub_le_assoc, -end - -/-- See `add_tsub_add_eq_tsub_left` for the equality. -/ -lemma add_tsub_add_le_tsub_left : a + b - (a + c) ≤ b - c := -by { rw [tsub_le_iff_left, add_assoc], exact add_le_add_left le_add_tsub _ } - -/-- See `add_tsub_add_eq_tsub_right` for the equality. -/ -lemma add_tsub_add_le_tsub_right : a + c - (b + c) ≤ a - b := -by { rw [tsub_le_iff_left, add_right_comm], exact add_le_add_right le_add_tsub c } - -end cov - -/-! #### Lemmas that assume that an element is `add_le_cancellable` -/ - -namespace add_le_cancellable - -protected lemma le_add_tsub_swap (hb : add_le_cancellable b) : a ≤ b + a - b := hb le_add_tsub - -protected lemma le_add_tsub (hb : add_le_cancellable b) : a ≤ a + b - b := -by { rw add_comm, exact hb.le_add_tsub_swap } - -protected lemma le_tsub_of_add_le_left (ha : add_le_cancellable a) (h : a + b ≤ c) : b ≤ c - a := -ha $ h.trans le_add_tsub - -protected lemma le_tsub_of_add_le_right (hb : add_le_cancellable b) (h : a + b ≤ c) : a ≤ c - b := -hb.le_tsub_of_add_le_left $ by rwa add_comm - -end add_le_cancellable - -/-! ### Lemmas where addition is order-reflecting -/ - -section contra -variable [contravariant_class α α (+) (≤)] - -lemma le_add_tsub_swap : a ≤ b + a - b := contravariant.add_le_cancellable.le_add_tsub_swap - -lemma le_add_tsub' : a ≤ a + b - b := contravariant.add_le_cancellable.le_add_tsub - -lemma le_tsub_of_add_le_left (h : a + b ≤ c) : b ≤ c - a := -contravariant.add_le_cancellable.le_tsub_of_add_le_left h - -lemma le_tsub_of_add_le_right (h : a + b ≤ c) : a ≤ c - b := -contravariant.add_le_cancellable.le_tsub_of_add_le_right h - -end contra - -end add_comm_semigroup - variables [add_comm_monoid α] [has_sub α] [has_ordered_sub α] {a b c d : α} -lemma tsub_nonpos : a - b ≤ 0 ↔ a ≤ b := by rw [tsub_le_iff_left, add_zero] - -alias tsub_nonpos ↔ _ tsub_nonpos_of_le - lemma add_monoid_hom.le_map_tsub [preorder β] [add_comm_monoid β] [has_sub β] [has_ordered_sub β] (f : α →+ β) (hf : monotone f) (a b : α) : f a - f b ≤ f (a - b) := f.to_add_hom.le_map_tsub hf a b end preorder - -/-! ### Partial order -/ - -variables [partial_order α] [add_comm_semigroup α] [has_sub α] [has_ordered_sub α] {a b c d : α} - -lemma tsub_tsub (b a c : α) : b - a - c = b - (a + c) := -begin - apply le_antisymm, - { rw [tsub_le_iff_left, tsub_le_iff_left, ← add_assoc, ← tsub_le_iff_left] }, - { rw [tsub_le_iff_left, add_assoc, ← tsub_le_iff_left, ← tsub_le_iff_left] } -end - -lemma tsub_add_eq_tsub_tsub (a b c : α) : a - (b + c) = a - b - c := (tsub_tsub _ _ _).symm - -lemma tsub_add_eq_tsub_tsub_swap (a b c : α) : a - (b + c) = a - c - b := -by { rw [add_comm], apply tsub_add_eq_tsub_tsub } - -lemma tsub_right_comm : a - b - c = a - c - b := -by simp_rw [← tsub_add_eq_tsub_tsub, add_comm] - -/-! ### Lemmas that assume that an element is `add_le_cancellable`. -/ - -namespace add_le_cancellable - -protected lemma tsub_eq_of_eq_add (hb : add_le_cancellable b) (h : a = c + b) : a - b = c := -le_antisymm (tsub_le_iff_right.mpr h.le) $ - by { rw h, exact hb.le_add_tsub } - -protected lemma eq_tsub_of_add_eq (hc : add_le_cancellable c) (h : a + c = b) : a = b - c := -(hc.tsub_eq_of_eq_add h.symm).symm - -protected theorem tsub_eq_of_eq_add_rev (hb : add_le_cancellable b) (h : a = b + c) : a - b = c := -hb.tsub_eq_of_eq_add $ by rw [add_comm, h] - -@[simp] -protected lemma add_tsub_cancel_right (hb : add_le_cancellable b) : a + b - b = a := -hb.tsub_eq_of_eq_add $ by rw [add_comm] - -@[simp] -protected lemma add_tsub_cancel_left (ha : add_le_cancellable a) : a + b - a = b := -ha.tsub_eq_of_eq_add $ add_comm a b - -protected lemma lt_add_of_tsub_lt_left (hb : add_le_cancellable b) (h : a - b < c) : a < b + c := -begin - rw [lt_iff_le_and_ne, ← tsub_le_iff_left], - refine ⟨h.le, _⟩, - rintro rfl, - simpa [hb] using h, -end - -protected lemma lt_add_of_tsub_lt_right (hc : add_le_cancellable c) (h : a - c < b) : a < b + c := -begin - rw [lt_iff_le_and_ne, ← tsub_le_iff_right], - refine ⟨h.le, _⟩, - rintro rfl, - simpa [hc] using h, -end - -protected lemma lt_tsub_of_add_lt_right (hc : add_le_cancellable c) (h : a + c < b) : a < b - c := -(hc.le_tsub_of_add_le_right h.le).lt_of_ne $ by { rintro rfl, exact h.not_le le_tsub_add } - -protected lemma lt_tsub_of_add_lt_left (ha : add_le_cancellable a) (h : a + c < b) : c < b - a := -ha.lt_tsub_of_add_lt_right $ by rwa add_comm - -end add_le_cancellable - -/-! #### Lemmas where addition is order-reflecting. -/ - -section contra -variable [contravariant_class α α (+) (≤)] - -lemma tsub_eq_of_eq_add (h : a = c + b) : a - b = c := -contravariant.add_le_cancellable.tsub_eq_of_eq_add h - -lemma eq_tsub_of_add_eq (h : a + c = b) : a = b - c := -contravariant.add_le_cancellable.eq_tsub_of_add_eq h - -lemma tsub_eq_of_eq_add_rev (h : a = b + c) : a - b = c := -contravariant.add_le_cancellable.tsub_eq_of_eq_add_rev h - -@[simp] -lemma add_tsub_cancel_right (a b : α) : a + b - b = a := -contravariant.add_le_cancellable.add_tsub_cancel_right - -@[simp] -lemma add_tsub_cancel_left (a b : α) : a + b - a = b := -contravariant.add_le_cancellable.add_tsub_cancel_left - -lemma lt_add_of_tsub_lt_left (h : a - b < c) : a < b + c := -contravariant.add_le_cancellable.lt_add_of_tsub_lt_left h - -lemma lt_add_of_tsub_lt_right (h : a - c < b) : a < b + c := -contravariant.add_le_cancellable.lt_add_of_tsub_lt_right h - -/-- This lemma (and some of its corollaries) also holds for `ennreal`, but this proof doesn't work -for it. Maybe we should add this lemma as field to `has_ordered_sub`? -/ -lemma lt_tsub_of_add_lt_left : a + c < b → c < b - a := -contravariant.add_le_cancellable.lt_tsub_of_add_lt_left - -lemma lt_tsub_of_add_lt_right : a + c < b → a < b - c := -contravariant.add_le_cancellable.lt_tsub_of_add_lt_right - -end contra - -section both -variables [covariant_class α α (+) (≤)] [contravariant_class α α (+) (≤)] - -lemma add_tsub_add_eq_tsub_right (a c b : α) : (a + c) - (b + c) = a - b := -begin - refine add_tsub_add_le_tsub_right.antisymm (tsub_le_iff_right.2 $ le_of_add_le_add_right _), swap, - rw add_assoc, - exact le_tsub_add, -end - -lemma add_tsub_add_eq_tsub_left (a b c : α) : (a + b) - (a + c) = b - c := -by rw [add_comm a b, add_comm a c, add_tsub_add_eq_tsub_right] - -end both - -end ordered_add_comm_semigroup - -/-! ### Lemmas in a linearly ordered monoid. -/ -section linear_order -variables {a b c d : α} [linear_order α] [add_comm_semigroup α] [has_sub α] [has_ordered_sub α] - -/-- See `lt_of_tsub_lt_tsub_right_of_le` for a weaker statement in a partial order. -/ -lemma lt_of_tsub_lt_tsub_right (h : a - c < b - c) : a < b := -lt_imp_lt_of_le_imp_le (λ h, tsub_le_tsub_right h c) h - -/-- See `lt_tsub_iff_right_of_le` for a weaker statement in a partial order. -/ -lemma lt_tsub_iff_right : a < b - c ↔ a + c < b := -lt_iff_lt_of_le_iff_le tsub_le_iff_right - -/-- See `lt_tsub_iff_left_of_le` for a weaker statement in a partial order. -/ -lemma lt_tsub_iff_left : a < b - c ↔ c + a < b := -lt_iff_lt_of_le_iff_le tsub_le_iff_left - -lemma lt_tsub_comm : a < b - c ↔ c < b - a := -lt_tsub_iff_left.trans lt_tsub_iff_right.symm - -section cov -variable [covariant_class α α (+) (≤)] - -/-- See `lt_of_tsub_lt_tsub_left_of_le` for a weaker statement in a partial order. -/ -lemma lt_of_tsub_lt_tsub_left (h : a - b < a - c) : c < b := -lt_imp_lt_of_le_imp_le (λ h, tsub_le_tsub_left h a) h - -end cov - -end linear_order - -section ordered_add_comm_monoid -variables [partial_order α] [add_comm_monoid α] [has_sub α] [has_ordered_sub α] - -@[simp] lemma tsub_zero (a : α) : a - 0 = a := -add_le_cancellable.tsub_eq_of_eq_add add_le_cancellable_zero (add_zero _).symm - -end ordered_add_comm_monoid diff --git a/src/algebra/order/sub/canonical.lean b/src/algebra/order/sub/canonical.lean index 5d3f58c8983b7..b938d38dd2132 100644 --- a/src/algebra/order/sub/canonical.lean +++ b/src/algebra/order/sub/canonical.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import algebra.order.monoid.canonical -import algebra.order.sub.basic +import algebra.order.monoid.canonical.defs +import algebra.order.sub.defs /-! # Lemmas about subtraction in canonically ordered monoids diff --git a/src/algebra/order/sub/defs.lean b/src/algebra/order/sub/defs.lean new file mode 100644 index 0000000000000..785cf2443e986 --- /dev/null +++ b/src/algebra/order/sub/defs.lean @@ -0,0 +1,363 @@ +/- +Copyright (c) 2021 Floris van Doorn. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn +-/ +import algebra.covariant_and_contravariant +import algebra.order.monoid.defs +import algebra.ring.defs + +/-! +# Ordered Subtraction + +This file proves lemmas relating (truncated) subtraction with an order. We provide a class +`has_ordered_sub` stating that `a - b ≤ c ↔ a ≤ c + b`. + +The subtraction discussed here could both be normal subtraction in an additive group or truncated +subtraction on a canonically ordered monoid (`ℕ`, `multiset`, `part_enat`, `ennreal`, ...) + +## Implementation details + +`has_ordered_sub` is a mixin type-class, so that we can use the results in this file even in cases +where we don't have a `canonically_ordered_add_monoid` instance +(even though that is our main focus). Conversely, this means we can use +`canonically_ordered_add_monoid` without necessarily having to define a subtraction. + +The results in this file are ordered by the type-class assumption needed to prove it. +This means that similar results might not be close to each other. Furthermore, we don't prove +implications if a bi-implication can be proven under the same assumptions. + +Lemmas using this class are named using `tsub` instead of `sub` (short for "truncated subtraction"). +This is to avoid naming conflicts with similar lemmas about ordered groups. + +We provide a second version of most results that require `[contravariant_class α α (+) (≤)]`. In the +second version we replace this type-class assumption by explicit `add_le_cancellable` assumptions. + +TODO: maybe we should make a multiplicative version of this, so that we can replace some identical +lemmas about subtraction/division in `ordered_[add_]comm_group` with these. + +TODO: generalize `nat.le_of_le_of_sub_le_sub_right`, `nat.sub_le_sub_right_iff`, + `nat.mul_self_sub_mul_self_eq` +-/ + +variables {α β : Type*} + +/-- `has_ordered_sub α` means that `α` has a subtraction characterized by `a - b ≤ c ↔ a ≤ c + b`. +In other words, `a - b` is the least `c` such that `a ≤ b + c`. + +This is satisfied both by the subtraction in additive ordered groups and by truncated subtraction +in canonically ordered monoids on many specific types. +-/ +class has_ordered_sub (α : Type*) [has_le α] [has_add α] [has_sub α] := +(tsub_le_iff_right : ∀ a b c : α, a - b ≤ c ↔ a ≤ c + b) + +section has_add + +variables [preorder α] [has_add α] [has_sub α] [has_ordered_sub α] {a b c d : α} + +@[simp] lemma tsub_le_iff_right : a - b ≤ c ↔ a ≤ c + b := +has_ordered_sub.tsub_le_iff_right a b c + +/-- See `add_tsub_cancel_right` for the equality if `contravariant_class α α (+) (≤)`. -/ +lemma add_tsub_le_right : a + b - b ≤ a := +tsub_le_iff_right.mpr le_rfl + +lemma le_tsub_add : b ≤ (b - a) + a := +tsub_le_iff_right.mp le_rfl + +end has_add + +/-! ### Preorder -/ + +section ordered_add_comm_semigroup + +section preorder +variables [preorder α] + +section add_comm_semigroup +variables [add_comm_semigroup α] [has_sub α] [has_ordered_sub α] {a b c d : α} + +lemma tsub_le_iff_left : a - b ≤ c ↔ a ≤ b + c := +by rw [tsub_le_iff_right, add_comm] + +lemma le_add_tsub : a ≤ b + (a - b) := +tsub_le_iff_left.mp le_rfl + +/-- See `add_tsub_cancel_left` for the equality if `contravariant_class α α (+) (≤)`. -/ +lemma add_tsub_le_left : a + b - a ≤ b := +tsub_le_iff_left.mpr le_rfl + +lemma tsub_le_tsub_right (h : a ≤ b) (c : α) : a - c ≤ b - c := +tsub_le_iff_left.mpr $ h.trans le_add_tsub + +lemma tsub_le_iff_tsub_le : a - b ≤ c ↔ a - c ≤ b := +by rw [tsub_le_iff_left, tsub_le_iff_right] + +/-- See `tsub_tsub_cancel_of_le` for the equality. -/ +lemma tsub_tsub_le : b - (b - a) ≤ a := +tsub_le_iff_right.mpr le_add_tsub + +section cov +variable [covariant_class α α (+) (≤)] + +lemma tsub_le_tsub_left (h : a ≤ b) (c : α) : c - b ≤ c - a := +tsub_le_iff_left.mpr $ le_add_tsub.trans $ add_le_add_right h _ + +lemma tsub_le_tsub (hab : a ≤ b) (hcd : c ≤ d) : a - d ≤ b - c := +(tsub_le_tsub_right hab _).trans $ tsub_le_tsub_left hcd _ + +lemma antitone_const_tsub : antitone (λ x, c - x) := +λ x y hxy, tsub_le_tsub rfl.le hxy + +/-- See `add_tsub_assoc_of_le` for the equality. -/ +lemma add_tsub_le_assoc : a + b - c ≤ a + (b - c) := +by { rw [tsub_le_iff_left, add_left_comm], exact add_le_add_left le_add_tsub a } + +/-- See `tsub_add_eq_add_tsub` for the equality. -/ +lemma add_tsub_le_tsub_add : a + b - c ≤ a - c + b := +by { rw [add_comm, add_comm _ b], exact add_tsub_le_assoc } + +lemma add_le_add_add_tsub : a + b ≤ (a + c) + (b - c) := +by { rw [add_assoc], exact add_le_add_left le_add_tsub a } + +lemma le_tsub_add_add : a + b ≤ (a - c) + (b + c) := +by { rw [add_comm a, add_comm (a - c)], exact add_le_add_add_tsub } + +lemma tsub_le_tsub_add_tsub : a - c ≤ (a - b) + (b - c) := +begin + rw [tsub_le_iff_left, ← add_assoc, add_right_comm], + exact le_add_tsub.trans (add_le_add_right le_add_tsub _), +end + +lemma tsub_tsub_tsub_le_tsub : (c - a) - (c - b) ≤ b - a := +begin + rw [tsub_le_iff_left, tsub_le_iff_left, add_left_comm], + exact le_tsub_add.trans (add_le_add_left le_add_tsub _), +end + +lemma tsub_tsub_le_tsub_add {a b c : α} : a - (b - c) ≤ a - b + c := +tsub_le_iff_right.2 $ calc + a ≤ a - b + b : le_tsub_add + ... ≤ a - b + (c + (b - c)) : add_le_add_left le_add_tsub _ + ... = a - b + c + (b - c) : (add_assoc _ _ _).symm + +/-- See `tsub_add_tsub_comm` for the equality. -/ +lemma add_tsub_add_le_tsub_add_tsub : a + b - (c + d) ≤ a - c + (b - d) := +begin + rw [add_comm c, tsub_le_iff_left, add_assoc, ←tsub_le_iff_left, ←tsub_le_iff_left], + refine (tsub_le_tsub_right add_tsub_le_assoc c).trans _, + rw [add_comm a, add_comm (a - c)], + exact add_tsub_le_assoc, +end + +/-- See `add_tsub_add_eq_tsub_left` for the equality. -/ +lemma add_tsub_add_le_tsub_left : a + b - (a + c) ≤ b - c := +by { rw [tsub_le_iff_left, add_assoc], exact add_le_add_left le_add_tsub _ } + +/-- See `add_tsub_add_eq_tsub_right` for the equality. -/ +lemma add_tsub_add_le_tsub_right : a + c - (b + c) ≤ a - b := +by { rw [tsub_le_iff_left, add_right_comm], exact add_le_add_right le_add_tsub c } + +end cov + +/-! #### Lemmas that assume that an element is `add_le_cancellable` -/ + +namespace add_le_cancellable + +protected lemma le_add_tsub_swap (hb : add_le_cancellable b) : a ≤ b + a - b := hb le_add_tsub + +protected lemma le_add_tsub (hb : add_le_cancellable b) : a ≤ a + b - b := +by { rw add_comm, exact hb.le_add_tsub_swap } + +protected lemma le_tsub_of_add_le_left (ha : add_le_cancellable a) (h : a + b ≤ c) : b ≤ c - a := +ha $ h.trans le_add_tsub + +protected lemma le_tsub_of_add_le_right (hb : add_le_cancellable b) (h : a + b ≤ c) : a ≤ c - b := +hb.le_tsub_of_add_le_left $ by rwa add_comm + +end add_le_cancellable + +/-! ### Lemmas where addition is order-reflecting -/ + +section contra +variable [contravariant_class α α (+) (≤)] + +lemma le_add_tsub_swap : a ≤ b + a - b := contravariant.add_le_cancellable.le_add_tsub_swap + +lemma le_add_tsub' : a ≤ a + b - b := contravariant.add_le_cancellable.le_add_tsub + +lemma le_tsub_of_add_le_left (h : a + b ≤ c) : b ≤ c - a := +contravariant.add_le_cancellable.le_tsub_of_add_le_left h + +lemma le_tsub_of_add_le_right (h : a + b ≤ c) : a ≤ c - b := +contravariant.add_le_cancellable.le_tsub_of_add_le_right h + +end contra + +end add_comm_semigroup + +variables [add_comm_monoid α] [has_sub α] [has_ordered_sub α] {a b c d : α} + +lemma tsub_nonpos : a - b ≤ 0 ↔ a ≤ b := by rw [tsub_le_iff_left, add_zero] + +alias tsub_nonpos ↔ _ tsub_nonpos_of_le + +end preorder + +/-! ### Partial order -/ + +variables [partial_order α] [add_comm_semigroup α] [has_sub α] [has_ordered_sub α] {a b c d : α} + +lemma tsub_tsub (b a c : α) : b - a - c = b - (a + c) := +begin + apply le_antisymm, + { rw [tsub_le_iff_left, tsub_le_iff_left, ← add_assoc, ← tsub_le_iff_left] }, + { rw [tsub_le_iff_left, add_assoc, ← tsub_le_iff_left, ← tsub_le_iff_left] } +end + +lemma tsub_add_eq_tsub_tsub (a b c : α) : a - (b + c) = a - b - c := (tsub_tsub _ _ _).symm + +lemma tsub_add_eq_tsub_tsub_swap (a b c : α) : a - (b + c) = a - c - b := +by { rw [add_comm], apply tsub_add_eq_tsub_tsub } + +lemma tsub_right_comm : a - b - c = a - c - b := +by simp_rw [← tsub_add_eq_tsub_tsub, add_comm] + +/-! ### Lemmas that assume that an element is `add_le_cancellable`. -/ + +namespace add_le_cancellable + +protected lemma tsub_eq_of_eq_add (hb : add_le_cancellable b) (h : a = c + b) : a - b = c := +le_antisymm (tsub_le_iff_right.mpr h.le) $ + by { rw h, exact hb.le_add_tsub } + +protected lemma eq_tsub_of_add_eq (hc : add_le_cancellable c) (h : a + c = b) : a = b - c := +(hc.tsub_eq_of_eq_add h.symm).symm + +protected theorem tsub_eq_of_eq_add_rev (hb : add_le_cancellable b) (h : a = b + c) : a - b = c := +hb.tsub_eq_of_eq_add $ by rw [add_comm, h] + +@[simp] +protected lemma add_tsub_cancel_right (hb : add_le_cancellable b) : a + b - b = a := +hb.tsub_eq_of_eq_add $ by rw [add_comm] + +@[simp] +protected lemma add_tsub_cancel_left (ha : add_le_cancellable a) : a + b - a = b := +ha.tsub_eq_of_eq_add $ add_comm a b + +protected lemma lt_add_of_tsub_lt_left (hb : add_le_cancellable b) (h : a - b < c) : a < b + c := +begin + rw [lt_iff_le_and_ne, ← tsub_le_iff_left], + refine ⟨h.le, _⟩, + rintro rfl, + simpa [hb] using h, +end + +protected lemma lt_add_of_tsub_lt_right (hc : add_le_cancellable c) (h : a - c < b) : a < b + c := +begin + rw [lt_iff_le_and_ne, ← tsub_le_iff_right], + refine ⟨h.le, _⟩, + rintro rfl, + simpa [hc] using h, +end + +protected lemma lt_tsub_of_add_lt_right (hc : add_le_cancellable c) (h : a + c < b) : a < b - c := +(hc.le_tsub_of_add_le_right h.le).lt_of_ne $ by { rintro rfl, exact h.not_le le_tsub_add } + +protected lemma lt_tsub_of_add_lt_left (ha : add_le_cancellable a) (h : a + c < b) : c < b - a := +ha.lt_tsub_of_add_lt_right $ by rwa add_comm + +end add_le_cancellable + +/-! #### Lemmas where addition is order-reflecting. -/ + +section contra +variable [contravariant_class α α (+) (≤)] + +lemma tsub_eq_of_eq_add (h : a = c + b) : a - b = c := +contravariant.add_le_cancellable.tsub_eq_of_eq_add h + +lemma eq_tsub_of_add_eq (h : a + c = b) : a = b - c := +contravariant.add_le_cancellable.eq_tsub_of_add_eq h + +lemma tsub_eq_of_eq_add_rev (h : a = b + c) : a - b = c := +contravariant.add_le_cancellable.tsub_eq_of_eq_add_rev h + +@[simp] +lemma add_tsub_cancel_right (a b : α) : a + b - b = a := +contravariant.add_le_cancellable.add_tsub_cancel_right + +@[simp] +lemma add_tsub_cancel_left (a b : α) : a + b - a = b := +contravariant.add_le_cancellable.add_tsub_cancel_left + +lemma lt_add_of_tsub_lt_left (h : a - b < c) : a < b + c := +contravariant.add_le_cancellable.lt_add_of_tsub_lt_left h + +lemma lt_add_of_tsub_lt_right (h : a - c < b) : a < b + c := +contravariant.add_le_cancellable.lt_add_of_tsub_lt_right h + +/-- This lemma (and some of its corollaries) also holds for `ennreal`, but this proof doesn't work +for it. Maybe we should add this lemma as field to `has_ordered_sub`? -/ +lemma lt_tsub_of_add_lt_left : a + c < b → c < b - a := +contravariant.add_le_cancellable.lt_tsub_of_add_lt_left + +lemma lt_tsub_of_add_lt_right : a + c < b → a < b - c := +contravariant.add_le_cancellable.lt_tsub_of_add_lt_right + +end contra + +section both +variables [covariant_class α α (+) (≤)] [contravariant_class α α (+) (≤)] + +lemma add_tsub_add_eq_tsub_right (a c b : α) : (a + c) - (b + c) = a - b := +begin + refine add_tsub_add_le_tsub_right.antisymm (tsub_le_iff_right.2 $ le_of_add_le_add_right _), swap, + rw add_assoc, + exact le_tsub_add, +end + +lemma add_tsub_add_eq_tsub_left (a b c : α) : (a + b) - (a + c) = b - c := +by rw [add_comm a b, add_comm a c, add_tsub_add_eq_tsub_right] + +end both + +end ordered_add_comm_semigroup + +/-! ### Lemmas in a linearly ordered monoid. -/ +section linear_order +variables {a b c d : α} [linear_order α] [add_comm_semigroup α] [has_sub α] [has_ordered_sub α] + +/-- See `lt_of_tsub_lt_tsub_right_of_le` for a weaker statement in a partial order. -/ +lemma lt_of_tsub_lt_tsub_right (h : a - c < b - c) : a < b := +lt_imp_lt_of_le_imp_le (λ h, tsub_le_tsub_right h c) h + +/-- See `lt_tsub_iff_right_of_le` for a weaker statement in a partial order. -/ +lemma lt_tsub_iff_right : a < b - c ↔ a + c < b := +lt_iff_lt_of_le_iff_le tsub_le_iff_right + +/-- See `lt_tsub_iff_left_of_le` for a weaker statement in a partial order. -/ +lemma lt_tsub_iff_left : a < b - c ↔ c + a < b := +lt_iff_lt_of_le_iff_le tsub_le_iff_left + +lemma lt_tsub_comm : a < b - c ↔ c < b - a := +lt_tsub_iff_left.trans lt_tsub_iff_right.symm + +section cov +variable [covariant_class α α (+) (≤)] + +/-- See `lt_of_tsub_lt_tsub_left_of_le` for a weaker statement in a partial order. -/ +lemma lt_of_tsub_lt_tsub_left (h : a - b < a - c) : c < b := +lt_imp_lt_of_le_imp_le (λ h, tsub_le_tsub_left h a) h + +end cov + +end linear_order + +section ordered_add_comm_monoid +variables [partial_order α] [add_comm_monoid α] [has_sub α] [has_ordered_sub α] + +@[simp] lemma tsub_zero (a : α) : a - 0 = a := +add_le_cancellable.tsub_eq_of_eq_add add_le_cancellable_zero (add_zero _).symm + +end ordered_add_comm_monoid diff --git a/src/algebra/order/sub/with_top.lean b/src/algebra/order/sub/with_top.lean index 2c38902326f17..9068932b60cd7 100644 --- a/src/algebra/order/sub/with_top.lean +++ b/src/algebra/order/sub/with_top.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import algebra.order.sub.basic +import algebra.order.sub.defs import algebra.order.monoid.with_top /-! diff --git a/src/algebra/order/upper_lower.lean b/src/algebra/order/upper_lower.lean index a4db7dcc24c4e..e6845bdde35db 100644 --- a/src/algebra/order/upper_lower.lean +++ b/src/algebra/order/upper_lower.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import algebra.order.group.basic +import algebra.order.group.defs import data.set.pointwise.basic import order.upper_lower diff --git a/src/algebra/order/with_zero.lean b/src/algebra/order/with_zero.lean index 724f603462641..9b4491eab1a32 100644 --- a/src/algebra/order/with_zero.lean +++ b/src/algebra/order/with_zero.lean @@ -5,6 +5,7 @@ Authors: Kenny Lau, Johan Commelin, Patrick Massot -/ import algebra.order.group.type_tags import algebra.order.group.units +import algebra.order.monoid.basic import algebra.order.monoid.with_zero /-! diff --git a/src/algebra/quandle.lean b/src/algebra/quandle.lean index 92aaa5c0fb566..7db721f07c95b 100644 --- a/src/algebra/quandle.lean +++ b/src/algebra/quandle.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ -import algebra.hom.equiv +import algebra.hom.equiv.basic import data.zmod.basic import tactic.group diff --git a/src/algebra/ring/basic.lean b/src/algebra/ring/basic.lean index d68aa0dcff8d8..29497f57bd38e 100644 --- a/src/algebra/ring/basic.lean +++ b/src/algebra/ring/basic.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Yury Kudryashov, Neil Strickland -/ import algebra.ring.defs -import algebra.regular.basic +import algebra.group.commute import algebra.hom.group import algebra.opposites +import algebra.ring.inj_surj /-! # Semirings and rings @@ -16,7 +17,7 @@ This is analogous to `algebra.group.basic`, the difference being that the former is about `+` and `*` separately, while the present file is about their interaction. -For the definitions of semirings and rings see `algebra.ring.degs`. +For the definitions of semirings and rings see `algebra.ring.defs`. -/ universes u v w x @@ -24,151 +25,6 @@ variables {α : Type u} {β : Type v} {γ : Type w} {R : Type x} open function -/-! -### `distrib` class --/ - -/-- Pullback a `distrib` instance along an injective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.injective.distrib {S} [has_mul R] [has_add R] [distrib S] - (f : R → S) (hf : injective f) (add : ∀ x y, f (x + y) = f x + f y) - (mul : ∀ x y, f (x * y) = f x * f y) : - distrib R := -{ mul := (*), - add := (+), - left_distrib := λ x y z, hf $ by simp only [*, left_distrib], - right_distrib := λ x y z, hf $ by simp only [*, right_distrib] } - -/-- Pushforward a `distrib` instance along a surjective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.surjective.distrib {S} [distrib R] [has_add S] [has_mul S] - (f : R → S) (hf : surjective f) (add : ∀ x y, f (x + y) = f x + f y) - (mul : ∀ x y, f (x * y) = f x * f y) : - distrib S := -{ mul := (*), - add := (+), - left_distrib := hf.forall₃.2 $ λ x y z, by simp only [← add, ← mul, left_distrib], - right_distrib := hf.forall₃.2 $ λ x y z, by simp only [← add, ← mul, right_distrib] } - -/-! -### Semirings --/ - -section injective_surjective_maps - -variables [has_zero β] [has_add β] [has_mul β] [has_smul ℕ β] - -/-- Pullback a `non_unital_non_assoc_semiring` instance along an injective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.injective.non_unital_non_assoc_semiring - {α : Type u} [non_unital_non_assoc_semiring α] - (f : β → α) (hf : injective f) (zero : f 0 = 0) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) : - non_unital_non_assoc_semiring β := -{ .. hf.mul_zero_class f zero mul, .. hf.add_comm_monoid f zero add nsmul, .. hf.distrib f add mul } - -/-- Pullback a `non_unital_semiring` instance along an injective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.injective.non_unital_semiring - {α : Type u} [non_unital_semiring α] - (f : β → α) (hf : injective f) (zero : f 0 = 0) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) : - non_unital_semiring β := -{ .. hf.non_unital_non_assoc_semiring f zero add mul nsmul, .. hf.semigroup_with_zero f zero mul } - -/-- Pullback a `non_assoc_semiring` instance along an injective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.injective.non_assoc_semiring - {α : Type u} [non_assoc_semiring α] - {β : Type v} [has_zero β] [has_one β] [has_mul β] [has_add β] - [has_smul ℕ β] [has_nat_cast β] - (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) - (nat_cast : ∀ n : ℕ, f n = n) : - non_assoc_semiring β := -{ .. hf.add_monoid_with_one f zero one add nsmul nat_cast, - .. hf.non_unital_non_assoc_semiring f zero add mul nsmul, - .. hf.mul_one_class f one mul } - -/-- Pullback a `semiring` instance along an injective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.injective.semiring - {α : Type u} [semiring α] - {β : Type v} [has_zero β] [has_one β] [has_add β] [has_mul β] [has_pow β ℕ] - [has_smul ℕ β] [has_nat_cast β] - (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) : - semiring β := -{ .. hf.non_assoc_semiring f zero one add mul nsmul nat_cast, - .. hf.monoid_with_zero f zero one mul npow, - .. hf.distrib f add mul } - -/-- Pushforward a `non_unital_non_assoc_semiring` instance along a surjective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.surjective.non_unital_non_assoc_semiring - {α : Type u} [non_unital_non_assoc_semiring α] - (f : α → β) (hf : surjective f) (zero : f 0 = 0) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) : - non_unital_non_assoc_semiring β := -{ .. hf.mul_zero_class f zero mul, .. hf.add_comm_monoid f zero add nsmul, .. hf.distrib f add mul } - -/-- Pushforward a `non_unital_semiring` instance along a surjective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.surjective.non_unital_semiring - {α : Type u} [non_unital_semiring α] - (f : α → β) (hf : surjective f) (zero : f 0 = 0) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) : - non_unital_semiring β := -{ .. hf.non_unital_non_assoc_semiring f zero add mul nsmul, .. hf.semigroup_with_zero f zero mul } - -/-- Pushforward a `non_assoc_semiring` instance along a surjective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.surjective.non_assoc_semiring - {α : Type u} [non_assoc_semiring α] - {β : Type v} [has_zero β] [has_one β] [has_add β] [has_mul β] - [has_smul ℕ β] [has_nat_cast β] - (f : α → β) (hf : surjective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) - (nat_cast : ∀ n : ℕ, f n = n) : - non_assoc_semiring β := -{ .. hf.add_monoid_with_one f zero one add nsmul nat_cast, - .. hf.non_unital_non_assoc_semiring f zero add mul nsmul, .. hf.mul_one_class f one mul } - -/-- Pushforward a `semiring` instance along a surjective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.surjective.semiring - {α : Type u} [semiring α] - {β : Type v} [has_zero β] [has_one β] [has_add β] [has_mul β] [has_pow β ℕ] - [has_smul ℕ β] [has_nat_cast β] - (f : α → β) (hf : surjective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) : - semiring β := -{ .. hf.non_assoc_semiring f zero one add mul nsmul nat_cast, - .. hf.monoid_with_zero f zero one mul npow, .. hf.add_comm_monoid f zero add nsmul, - .. hf.distrib f add mul } - -end injective_surjective_maps - namespace add_hom /-- Left multiplication by an element of a type with distributive multiplication is an `add_hom`. -/ @@ -216,85 +72,11 @@ lemma mul_right_apply {R : Type*} [non_unital_non_assoc_semiring R] (a r : R) : end add_monoid_hom -section non_unital_comm_semiring -variables [non_unital_comm_semiring α] [non_unital_comm_semiring β] {a b c : α} - -/-- Pullback a `non_unital_semiring` instance along an injective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.injective.non_unital_comm_semiring [has_zero γ] [has_add γ] [has_mul γ] - [has_smul ℕ γ] (f : γ → α) (hf : injective f) (zero : f 0 = 0) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) : - non_unital_comm_semiring γ := -{ .. hf.non_unital_semiring f zero add mul nsmul, .. hf.comm_semigroup f mul } - -/-- Pushforward a `non_unital_semiring` instance along a surjective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.surjective.non_unital_comm_semiring [has_zero γ] [has_add γ] [has_mul γ] - [has_smul ℕ γ] (f : α → γ) (hf : surjective f) (zero : f 0 = 0) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) : - non_unital_comm_semiring γ := -{ .. hf.non_unital_semiring f zero add mul nsmul, .. hf.comm_semigroup f mul } - -end non_unital_comm_semiring - -section comm_semiring -variables [comm_semiring α] [comm_semiring β] {a b c : α} - -/-- Pullback a `semiring` instance along an injective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.injective.comm_semiring - [has_zero γ] [has_one γ] [has_add γ] [has_mul γ] [has_smul ℕ γ] [has_nat_cast γ] - [has_pow γ ℕ] (f : γ → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) : - comm_semiring γ := -{ .. hf.semiring f zero one add mul nsmul npow nat_cast, .. hf.comm_semigroup f mul } - -/-- Pushforward a `semiring` instance along a surjective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.surjective.comm_semiring - [has_zero γ] [has_one γ] [has_add γ] [has_mul γ] [has_smul ℕ γ] [has_nat_cast γ] - [has_pow γ ℕ] (f : α → γ) (hf : surjective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) : - comm_semiring γ := -{ .. hf.semiring f zero one add mul nsmul npow nat_cast, .. hf.comm_semigroup f mul } - -end comm_semiring - section has_distrib_neg section has_mul variables [has_mul α] [has_distrib_neg α] -/-- A type endowed with `-` and `*` has distributive negation, if it admits an injective map that -preserves `-` and `*` to a type which has distributive negation. -/ -@[reducible] -- See note [reducible non-instances] -protected def function.injective.has_distrib_neg [has_neg β] [has_mul β] (f : β → α) - (hf : injective f) (neg : ∀ a, f (-a) = -f a) (mul : ∀ a b, f (a * b) = f a * f b) : - has_distrib_neg β := -{ neg_mul := λ x y, hf $ by erw [neg, mul, neg, neg_mul, mul], - mul_neg := λ x y, hf $ by erw [neg, mul, neg, mul_neg, mul], - ..hf.has_involutive_neg _ neg, ..‹has_mul β› } - -/-- A type endowed with `-` and `*` has distributive negation, if it admits a surjective map that -preserves `-` and `*` from a type which has distributive negation. -/ -@[reducible] -- See note [reducible non-instances] -protected def function.surjective.has_distrib_neg [has_neg β] [has_mul β] (f : α → β) - (hf : surjective f) (neg : ∀ a, f (-a) = -f a) (mul : ∀ a b, f (a * b) = f a * f b) : - has_distrib_neg β := -{ neg_mul := hf.forall₂.2 $ λ x y, by { erw [←neg, ← mul, neg_mul, neg, mul], refl }, - mul_neg := hf.forall₂.2 $ λ x y, by { erw [←neg, ← mul, mul_neg, neg, mul], refl }, - ..hf.has_involutive_neg _ neg, ..‹has_mul β› } - namespace add_opposite instance : has_distrib_neg αᵃᵒᵖ := unop_injective.has_distrib_neg _ unop_neg unop_mul @@ -320,194 +102,6 @@ end group end has_distrib_neg -/-! -### Rings --/ - -section non_unital_non_assoc_ring -variables [non_unital_non_assoc_ring α] - -/-- Pullback a `non_unital_non_assoc_ring` instance along an injective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.injective.non_unital_non_assoc_ring - [has_zero β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] - (f : β → α) (hf : injective f) (zero : f 0 = 0) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) : - non_unital_non_assoc_ring β := -{ .. hf.add_comm_group f zero add neg sub nsmul zsmul, ..hf.mul_zero_class f zero mul, - .. hf.distrib f add mul } - -/-- Pushforward a `non_unital_non_assoc_ring` instance along a surjective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.surjective.non_unital_non_assoc_ring - [has_zero β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] - (f : α → β) (hf : surjective f) (zero : f 0 = 0) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) : - non_unital_non_assoc_ring β := -{ .. hf.add_comm_group f zero add neg sub nsmul zsmul, .. hf.mul_zero_class f zero mul, - .. hf.distrib f add mul } - -@[priority 100] -instance non_unital_non_assoc_ring.to_has_distrib_neg : has_distrib_neg α := -{ neg := has_neg.neg, - neg_neg := neg_neg, - neg_mul := λ a b, eq_neg_of_add_eq_zero_left $ by rw [←right_distrib, add_left_neg, zero_mul], - mul_neg := λ a b, eq_neg_of_add_eq_zero_left $ by rw [←left_distrib, add_left_neg, mul_zero] } - -lemma mul_sub_left_distrib (a b c : α) : a * (b - c) = a * b - a * c := -by simpa only [sub_eq_add_neg, neg_mul_eq_mul_neg] using mul_add a b (-c) - -alias mul_sub_left_distrib ← mul_sub - -lemma mul_sub_right_distrib (a b c : α) : (a - b) * c = a * c - b * c := -by simpa only [sub_eq_add_neg, neg_mul_eq_neg_mul] using add_mul a (-b) c - -alias mul_sub_right_distrib ← sub_mul - -variables {a b c d e : α} - -/-- An iff statement following from right distributivity in rings and the definition - of subtraction. -/ -theorem mul_add_eq_mul_add_iff_sub_mul_add_eq : a * e + c = b * e + d ↔ (a - b) * e + c = d := -calc - a * e + c = b * e + d ↔ a * e + c = d + b * e : by simp [add_comm] - ... ↔ a * e + c - b * e = d : iff.intro (λ h, begin rw h, simp end) (λ h, - begin rw ← h, simp end) - ... ↔ (a - b) * e + c = d : begin simp [sub_mul, sub_add_eq_add_sub] end - -/-- A simplification of one side of an equation exploiting right distributivity in rings - and the definition of subtraction. -/ -theorem sub_mul_add_eq_of_mul_add_eq_mul_add : a * e + c = b * e + d → (a - b) * e + c = d := -assume h, -calc - (a - b) * e + c = (a * e + c) - b * e : begin simp [sub_mul, sub_add_eq_add_sub] end - ... = d : begin rw h, simp [@add_sub_cancel α] end - -end non_unital_non_assoc_ring - -section non_unital_ring -variables [non_unital_ring α] - -/-- Pullback a `non_unital_ring` instance along an injective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.injective.non_unital_ring - [has_zero β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] - (f : β → α) (hf : injective f) (zero : f 0 = 0) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (gsmul : ∀ x (n : ℤ), f (n • x) = n • f x) : - non_unital_ring β := -{ .. hf.add_comm_group f zero add neg sub nsmul gsmul, ..hf.mul_zero_class f zero mul, - .. hf.distrib f add mul, .. hf.semigroup f mul } - -/-- Pushforward a `non_unital_ring` instance along a surjective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.surjective.non_unital_ring - [has_zero β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] - (f : α → β) (hf : surjective f) (zero : f 0 = 0) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (gsmul : ∀ x (n : ℤ), f (n • x) = n • f x) : - non_unital_ring β := -{ .. hf.add_comm_group f zero add neg sub nsmul gsmul, .. hf.mul_zero_class f zero mul, - .. hf.distrib f add mul, .. hf.semigroup f mul } - -end non_unital_ring - -section non_assoc_ring -variables [non_assoc_ring α] - -/-- Pullback a `non_assoc_ring` instance along an injective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.injective.non_assoc_ring - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] - [has_smul ℕ β] [has_smul ℤ β] [has_nat_cast β] [has_int_cast β] - (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (gsmul : ∀ x (n : ℤ), f (n • x) = n • f x) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : - non_assoc_ring β := -{ .. hf.add_comm_group f zero add neg sub nsmul gsmul, - .. hf.add_group_with_one f zero one add neg sub nsmul gsmul nat_cast int_cast, - .. hf.mul_zero_class f zero mul, .. hf.distrib f add mul, - .. hf.mul_one_class f one mul } - -/-- Pushforward a `non_unital_ring` instance along a surjective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.surjective.non_assoc_ring - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] - [has_smul ℕ β] [has_smul ℤ β] [has_nat_cast β] [has_int_cast β] - (f : α → β) (hf : surjective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (gsmul : ∀ x (n : ℤ), f (n • x) = n • f x) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : - non_assoc_ring β := -{ .. hf.add_comm_group f zero add neg sub nsmul gsmul, .. hf.mul_zero_class f zero mul, - .. hf.add_group_with_one f zero one add neg sub nsmul gsmul nat_cast int_cast, - .. hf.distrib f add mul, .. hf.mul_one_class f one mul } - -lemma sub_one_mul (a b : α) : (a - 1) * b = a * b - b := -by rw [sub_mul, one_mul] -lemma mul_sub_one (a b : α) : a * (b - 1) = a * b - a := -by rw [mul_sub, mul_one] -lemma one_sub_mul (a b : α) : (1 - a) * b = b - a * b := -by rw [sub_mul, one_mul] -lemma mul_one_sub (a b : α) : a * (1 - b) = a - a * b := -by rw [mul_sub, mul_one] - -end non_assoc_ring - -section ring -variables [ring α] {a b c d e : α} - -/-- Pullback a `ring` instance along an injective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.injective.ring - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] - [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] [has_int_cast β] - (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : - ring β := -{ .. hf.add_group_with_one f zero one add neg sub nsmul zsmul nat_cast int_cast, - .. hf.add_comm_group f zero add neg sub nsmul zsmul, - .. hf.monoid f one mul npow, .. hf.distrib f add mul } - -/-- Pushforward a `ring` instance along a surjective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.surjective.ring - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] - [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] [has_int_cast β] - (f : α → β) (hf : surjective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : - ring β := -{ .. hf.add_group_with_one f zero one add neg sub nsmul zsmul nat_cast int_cast, - .. hf.add_comm_group f zero add neg sub nsmul zsmul, - .. hf.monoid f one mul npow, .. hf.distrib f add mul } - -end ring - namespace units section has_distrib_neg @@ -574,32 +168,6 @@ lemma is_unit.sub_iff [ring α] {x y : α} : section non_unital_comm_ring variables [non_unital_comm_ring α] {a b c : α} -/-- Pullback a `comm_ring` instance along an injective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.injective.non_unital_comm_ring - [has_zero β] [has_add β] [has_mul β] [has_neg β] [has_sub β] - [has_smul ℕ β] [has_smul ℤ β] - (f : β → α) (hf : injective f) (zero : f 0 = 0) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) : - non_unital_comm_ring β := -{ .. hf.non_unital_ring f zero add mul neg sub nsmul zsmul, .. hf.comm_semigroup f mul } - -/-- Pushforward a `non_unital_comm_ring` instance along a surjective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.surjective.non_unital_comm_ring - [has_zero β] [has_add β] [has_mul β] [has_neg β] [has_sub β] - [has_smul ℕ β] [has_smul ℤ β] - (f : α → β) (hf : surjective f) (zero : f 0 = 0) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) : - non_unital_comm_ring β := -{ .. hf.non_unital_ring f zero add mul neg sub nsmul zsmul, .. hf.comm_semigroup f mul } - local attribute [simp] add_assoc add_comm add_left_comm mul_comm /-- Vieta's formula for a quadratic equation, relating the coefficients of the polynomial with @@ -616,43 +184,6 @@ end end non_unital_comm_ring -section comm_ring -variables [comm_ring α] {a b c : α} - -/-- Pullback a `comm_ring` instance along an injective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.injective.comm_ring - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] - [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] [has_int_cast β] - (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : - comm_ring β := -{ .. hf.ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast, - .. hf.comm_semigroup f mul } - -/-- Pushforward a `comm_ring` instance along a surjective function. -See note [reducible non-instances]. -/ -@[reducible] -protected def function.surjective.comm_ring - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] - [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] [has_int_cast β] - (f : α → β) (hf : surjective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : - comm_ring β := -{ .. hf.ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast, - .. hf.comm_semigroup f mul } - -end comm_ring - lemma succ_ne_self [non_assoc_ring α] [nontrivial α] (a : α) : a + 1 ≠ a := λ h, one_ne_zero ((add_right_inj a).mp (by simp [h])) diff --git a/src/algebra/ring/defs.lean b/src/algebra/ring/defs.lean index 0f65066588e07..28f46e4bd80f8 100644 --- a/src/algebra/ring/defs.lean +++ b/src/algebra/ring/defs.lean @@ -3,6 +3,7 @@ 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, Floris van Doorn, Yury Kudryashov, Neil Strickland -/ +import algebra.group.basic import algebra.group_with_zero.defs import data.int.cast.defs @@ -298,6 +299,62 @@ multiplicative monoid (`monoid`), and distributive laws (`distrib`). Equivalent @[protect_proj, ancestor add_comm_group monoid distrib] class ring (α : Type u) extends add_comm_group_with_one α, monoid α, distrib α +section non_unital_non_assoc_ring +variables [non_unital_non_assoc_ring α] + +@[priority 100] +instance non_unital_non_assoc_ring.to_has_distrib_neg : has_distrib_neg α := +{ neg := has_neg.neg, + neg_neg := neg_neg, + neg_mul := λ a b, eq_neg_of_add_eq_zero_left $ by rw [←right_distrib, add_left_neg, zero_mul], + mul_neg := λ a b, eq_neg_of_add_eq_zero_left $ by rw [←left_distrib, add_left_neg, mul_zero] } + +lemma mul_sub_left_distrib (a b c : α) : a * (b - c) = a * b - a * c := +by simpa only [sub_eq_add_neg, neg_mul_eq_mul_neg] using mul_add a b (-c) + +alias mul_sub_left_distrib ← mul_sub + +lemma mul_sub_right_distrib (a b c : α) : (a - b) * c = a * c - b * c := +by simpa only [sub_eq_add_neg, neg_mul_eq_neg_mul] using add_mul a (-b) c + +alias mul_sub_right_distrib ← sub_mul + +variables {a b c d e : α} + +/-- An iff statement following from right distributivity in rings and the definition + of subtraction. -/ +theorem mul_add_eq_mul_add_iff_sub_mul_add_eq : a * e + c = b * e + d ↔ (a - b) * e + c = d := +calc + a * e + c = b * e + d ↔ a * e + c = d + b * e : by simp [add_comm] + ... ↔ a * e + c - b * e = d : iff.intro (λ h, begin rw h, simp end) (λ h, + begin rw ← h, simp end) + ... ↔ (a - b) * e + c = d : begin simp [sub_mul, sub_add_eq_add_sub] end + +/-- A simplification of one side of an equation exploiting right distributivity in rings + and the definition of subtraction. -/ +theorem sub_mul_add_eq_of_mul_add_eq_mul_add : a * e + c = b * e + d → (a - b) * e + c = d := +assume h, +calc + (a - b) * e + c = (a * e + c) - b * e : begin simp [sub_mul, sub_add_eq_add_sub] end + ... = d : begin rw h, simp [@add_sub_cancel α] end + +end non_unital_non_assoc_ring + + +section non_assoc_ring +variables [non_assoc_ring α] + +lemma sub_one_mul (a b : α) : (a - 1) * b = a * b - b := +by rw [sub_mul, one_mul] +lemma mul_sub_one (a b : α) : a * (b - 1) = a * b - a := +by rw [mul_sub, mul_one] +lemma one_sub_mul (a b : α) : (1 - a) * b = b - a * b := +by rw [sub_mul, one_mul] +lemma mul_one_sub (a b : α) : a * (1 - b) = a - a * b := +by rw [mul_sub, mul_one] + +end non_assoc_ring + section ring variables [ring α] {a b c d e : α} diff --git a/src/algebra/ring/divisibility.lean b/src/algebra/ring/divisibility.lean index 329a97b6611e6..a4c9f7bbeee7e 100644 --- a/src/algebra/ring/divisibility.lean +++ b/src/algebra/ring/divisibility.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Yury Kudryashov, Neil Strickland -/ import algebra.ring.basic -import algebra.divisibility +import algebra.divisibility.basic /-! # Lemmas about divisibility in rings diff --git a/src/algebra/ring/equiv.lean b/src/algebra/ring/equiv.lean index 25ba6f38a69db..0df7b75533e6f 100644 --- a/src/algebra/ring/equiv.lean +++ b/src/algebra/ring/equiv.lean @@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ import algebra.big_operators.basic -import algebra.hom.equiv -import algebra.ring.opposite /-! # (Semi)ring equivs diff --git a/src/algebra/ring/inj_surj.lean b/src/algebra/ring/inj_surj.lean new file mode 100644 index 0000000000000..d4eb386c4a49a --- /dev/null +++ b/src/algebra/ring/inj_surj.lean @@ -0,0 +1,455 @@ +/- +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, Floris van Doorn, Yury Kudryashov, Neil Strickland +-/ +import algebra.ring.defs +import algebra.group_with_zero.basic + +/-! +# Pulling back rings along injective maps, and pushing them forward along surjective maps. + +-/ +universes u v w x +variables {α : Type u} {β : Type v} {γ : Type w} {R : Type x} + +open function + +/-! +### `distrib` class +-/ + +/-- Pullback a `distrib` instance along an injective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.distrib {S} [has_mul R] [has_add R] [distrib S] + (f : R → S) (hf : injective f) (add : ∀ x y, f (x + y) = f x + f y) + (mul : ∀ x y, f (x * y) = f x * f y) : + distrib R := +{ mul := (*), + add := (+), + left_distrib := λ x y z, hf $ by simp only [*, left_distrib], + right_distrib := λ x y z, hf $ by simp only [*, right_distrib] } + +/-- Pushforward a `distrib` instance along a surjective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.surjective.distrib {S} [distrib R] [has_add S] [has_mul S] + (f : R → S) (hf : surjective f) (add : ∀ x y, f (x + y) = f x + f y) + (mul : ∀ x y, f (x * y) = f x * f y) : + distrib S := +{ mul := (*), + add := (+), + left_distrib := hf.forall₃.2 $ λ x y z, by simp only [← add, ← mul, left_distrib], + right_distrib := hf.forall₃.2 $ λ x y z, by simp only [← add, ← mul, right_distrib] } + +section injective_surjective_maps + +/-! +### Semirings +-/ + +variables [has_zero β] [has_add β] [has_mul β] [has_smul ℕ β] + +/-- Pullback a `non_unital_non_assoc_semiring` instance along an injective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.non_unital_non_assoc_semiring + {α : Type u} [non_unital_non_assoc_semiring α] + (f : β → α) (hf : injective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) : + non_unital_non_assoc_semiring β := +{ .. hf.mul_zero_class f zero mul, .. hf.add_comm_monoid f zero add nsmul, .. hf.distrib f add mul } + +/-- Pullback a `non_unital_semiring` instance along an injective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.non_unital_semiring + {α : Type u} [non_unital_semiring α] + (f : β → α) (hf : injective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) : + non_unital_semiring β := +{ .. hf.non_unital_non_assoc_semiring f zero add mul nsmul, .. hf.semigroup_with_zero f zero mul } + +/-- Pullback a `non_assoc_semiring` instance along an injective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.non_assoc_semiring + {α : Type u} [non_assoc_semiring α] + {β : Type v} [has_zero β] [has_one β] [has_mul β] [has_add β] + [has_smul ℕ β] [has_nat_cast β] + (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) + (nat_cast : ∀ n : ℕ, f n = n) : + non_assoc_semiring β := +{ .. hf.add_monoid_with_one f zero one add nsmul nat_cast, + .. hf.non_unital_non_assoc_semiring f zero add mul nsmul, + .. hf.mul_one_class f one mul } + +/-- Pullback a `semiring` instance along an injective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.semiring + {α : Type u} [semiring α] + {β : Type v} [has_zero β] [has_one β] [has_add β] [has_mul β] [has_pow β ℕ] + [has_smul ℕ β] [has_nat_cast β] + (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) : + semiring β := +{ .. hf.non_assoc_semiring f zero one add mul nsmul nat_cast, + .. hf.monoid_with_zero f zero one mul npow, + .. hf.distrib f add mul } + +/-- Pushforward a `non_unital_non_assoc_semiring` instance along a surjective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.surjective.non_unital_non_assoc_semiring + {α : Type u} [non_unital_non_assoc_semiring α] + (f : α → β) (hf : surjective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) : + non_unital_non_assoc_semiring β := +{ .. hf.mul_zero_class f zero mul, .. hf.add_comm_monoid f zero add nsmul, .. hf.distrib f add mul } + +/-- Pushforward a `non_unital_semiring` instance along a surjective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.surjective.non_unital_semiring + {α : Type u} [non_unital_semiring α] + (f : α → β) (hf : surjective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) : + non_unital_semiring β := +{ .. hf.non_unital_non_assoc_semiring f zero add mul nsmul, .. hf.semigroup_with_zero f zero mul } + +/-- Pushforward a `non_assoc_semiring` instance along a surjective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.surjective.non_assoc_semiring + {α : Type u} [non_assoc_semiring α] + {β : Type v} [has_zero β] [has_one β] [has_add β] [has_mul β] + [has_smul ℕ β] [has_nat_cast β] + (f : α → β) (hf : surjective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) + (nat_cast : ∀ n : ℕ, f n = n) : + non_assoc_semiring β := +{ .. hf.add_monoid_with_one f zero one add nsmul nat_cast, + .. hf.non_unital_non_assoc_semiring f zero add mul nsmul, .. hf.mul_one_class f one mul } + +/-- Pushforward a `semiring` instance along a surjective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.surjective.semiring + {α : Type u} [semiring α] + {β : Type v} [has_zero β] [has_one β] [has_add β] [has_mul β] [has_pow β ℕ] + [has_smul ℕ β] [has_nat_cast β] + (f : α → β) (hf : surjective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) : + semiring β := +{ .. hf.non_assoc_semiring f zero one add mul nsmul nat_cast, + .. hf.monoid_with_zero f zero one mul npow, .. hf.add_comm_monoid f zero add nsmul, + .. hf.distrib f add mul } + +end injective_surjective_maps + +section non_unital_comm_semiring +variables [non_unital_comm_semiring α] [non_unital_comm_semiring β] {a b c : α} + +/-- Pullback a `non_unital_semiring` instance along an injective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.non_unital_comm_semiring [has_zero γ] [has_add γ] [has_mul γ] + [has_smul ℕ γ] (f : γ → α) (hf : injective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) : + non_unital_comm_semiring γ := +{ .. hf.non_unital_semiring f zero add mul nsmul, .. hf.comm_semigroup f mul } + +/-- Pushforward a `non_unital_semiring` instance along a surjective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.surjective.non_unital_comm_semiring [has_zero γ] [has_add γ] [has_mul γ] + [has_smul ℕ γ] (f : α → γ) (hf : surjective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) : + non_unital_comm_semiring γ := +{ .. hf.non_unital_semiring f zero add mul nsmul, .. hf.comm_semigroup f mul } + +end non_unital_comm_semiring + +section comm_semiring +variables [comm_semiring α] [comm_semiring β] {a b c : α} + +/-- Pullback a `semiring` instance along an injective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.comm_semiring + [has_zero γ] [has_one γ] [has_add γ] [has_mul γ] [has_smul ℕ γ] [has_nat_cast γ] + [has_pow γ ℕ] (f : γ → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) : + comm_semiring γ := +{ .. hf.semiring f zero one add mul nsmul npow nat_cast, .. hf.comm_semigroup f mul } + +/-- Pushforward a `semiring` instance along a surjective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.surjective.comm_semiring + [has_zero γ] [has_one γ] [has_add γ] [has_mul γ] [has_smul ℕ γ] [has_nat_cast γ] + [has_pow γ ℕ] (f : α → γ) (hf : surjective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) : + comm_semiring γ := +{ .. hf.semiring f zero one add mul nsmul npow nat_cast, .. hf.comm_semigroup f mul } + +end comm_semiring + +section has_distrib_neg + +section has_mul +variables [has_mul α] [has_distrib_neg α] + +/-- A type endowed with `-` and `*` has distributive negation, if it admits an injective map that +preserves `-` and `*` to a type which has distributive negation. -/ +@[reducible] -- See note [reducible non-instances] +protected def function.injective.has_distrib_neg [has_neg β] [has_mul β] (f : β → α) + (hf : injective f) (neg : ∀ a, f (-a) = -f a) (mul : ∀ a b, f (a * b) = f a * f b) : + has_distrib_neg β := +{ neg_mul := λ x y, hf $ by erw [neg, mul, neg, neg_mul, mul], + mul_neg := λ x y, hf $ by erw [neg, mul, neg, mul_neg, mul], + ..hf.has_involutive_neg _ neg, ..‹has_mul β› } + +/-- A type endowed with `-` and `*` has distributive negation, if it admits a surjective map that +preserves `-` and `*` from a type which has distributive negation. -/ +@[reducible] -- See note [reducible non-instances] +protected def function.surjective.has_distrib_neg [has_neg β] [has_mul β] (f : α → β) + (hf : surjective f) (neg : ∀ a, f (-a) = -f a) (mul : ∀ a b, f (a * b) = f a * f b) : + has_distrib_neg β := +{ neg_mul := hf.forall₂.2 $ λ x y, by { erw [←neg, ← mul, neg_mul, neg, mul], refl }, + mul_neg := hf.forall₂.2 $ λ x y, by { erw [←neg, ← mul, mul_neg, neg, mul], refl }, + ..hf.has_involutive_neg _ neg, ..‹has_mul β› } + +end has_mul + +end has_distrib_neg + +/-! +### Rings +-/ + +section non_unital_non_assoc_ring +variables [non_unital_non_assoc_ring α] + +/-- Pullback a `non_unital_non_assoc_ring` instance along an injective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.non_unital_non_assoc_ring + [has_zero β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] + (f : β → α) (hf : injective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) : + non_unital_non_assoc_ring β := +{ .. hf.add_comm_group f zero add neg sub nsmul zsmul, ..hf.mul_zero_class f zero mul, + .. hf.distrib f add mul } + +/-- Pushforward a `non_unital_non_assoc_ring` instance along a surjective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.surjective.non_unital_non_assoc_ring + [has_zero β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] + (f : α → β) (hf : surjective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) : + non_unital_non_assoc_ring β := +{ .. hf.add_comm_group f zero add neg sub nsmul zsmul, .. hf.mul_zero_class f zero mul, + .. hf.distrib f add mul } + +end non_unital_non_assoc_ring + +section non_unital_ring +variables [non_unital_ring α] + +/-- Pullback a `non_unital_ring` instance along an injective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.non_unital_ring + [has_zero β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] + (f : β → α) (hf : injective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (gsmul : ∀ x (n : ℤ), f (n • x) = n • f x) : + non_unital_ring β := +{ .. hf.add_comm_group f zero add neg sub nsmul gsmul, ..hf.mul_zero_class f zero mul, + .. hf.distrib f add mul, .. hf.semigroup f mul } + +/-- Pushforward a `non_unital_ring` instance along a surjective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.surjective.non_unital_ring + [has_zero β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] + (f : α → β) (hf : surjective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (gsmul : ∀ x (n : ℤ), f (n • x) = n • f x) : + non_unital_ring β := +{ .. hf.add_comm_group f zero add neg sub nsmul gsmul, .. hf.mul_zero_class f zero mul, + .. hf.distrib f add mul, .. hf.semigroup f mul } + +end non_unital_ring + +section non_assoc_ring +variables [non_assoc_ring α] + +/-- Pullback a `non_assoc_ring` instance along an injective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.non_assoc_ring + [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] + [has_smul ℕ β] [has_smul ℤ β] [has_nat_cast β] [has_int_cast β] + (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (gsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + non_assoc_ring β := +{ .. hf.add_comm_group f zero add neg sub nsmul gsmul, + .. hf.add_group_with_one f zero one add neg sub nsmul gsmul nat_cast int_cast, + .. hf.mul_zero_class f zero mul, .. hf.distrib f add mul, + .. hf.mul_one_class f one mul } + +/-- Pushforward a `non_unital_ring` instance along a surjective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.surjective.non_assoc_ring + [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] + [has_smul ℕ β] [has_smul ℤ β] [has_nat_cast β] [has_int_cast β] + (f : α → β) (hf : surjective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (gsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + non_assoc_ring β := +{ .. hf.add_comm_group f zero add neg sub nsmul gsmul, .. hf.mul_zero_class f zero mul, + .. hf.add_group_with_one f zero one add neg sub nsmul gsmul nat_cast int_cast, + .. hf.distrib f add mul, .. hf.mul_one_class f one mul } + +end non_assoc_ring + +section ring +variables [ring α] {a b c d e : α} + +/-- Pullback a `ring` instance along an injective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.ring + [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] + [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] [has_int_cast β] + (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + ring β := +{ .. hf.add_group_with_one f zero one add neg sub nsmul zsmul nat_cast int_cast, + .. hf.add_comm_group f zero add neg sub nsmul zsmul, + .. hf.monoid f one mul npow, .. hf.distrib f add mul } + +/-- Pushforward a `ring` instance along a surjective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.surjective.ring + [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] + [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] [has_int_cast β] + (f : α → β) (hf : surjective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + ring β := +{ .. hf.add_group_with_one f zero one add neg sub nsmul zsmul nat_cast int_cast, + .. hf.add_comm_group f zero add neg sub nsmul zsmul, + .. hf.monoid f one mul npow, .. hf.distrib f add mul } + +end ring + +section non_unital_comm_ring +variables [non_unital_comm_ring α] {a b c : α} + +/-- Pullback a `comm_ring` instance along an injective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.non_unital_comm_ring + [has_zero β] [has_add β] [has_mul β] [has_neg β] [has_sub β] + [has_smul ℕ β] [has_smul ℤ β] + (f : β → α) (hf : injective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) : + non_unital_comm_ring β := +{ .. hf.non_unital_ring f zero add mul neg sub nsmul zsmul, .. hf.comm_semigroup f mul } + +/-- Pushforward a `non_unital_comm_ring` instance along a surjective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.surjective.non_unital_comm_ring + [has_zero β] [has_add β] [has_mul β] [has_neg β] [has_sub β] + [has_smul ℕ β] [has_smul ℤ β] + (f : α → β) (hf : surjective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) : + non_unital_comm_ring β := +{ .. hf.non_unital_ring f zero add mul neg sub nsmul zsmul, .. hf.comm_semigroup f mul } + +end non_unital_comm_ring + +section comm_ring +variables [comm_ring α] {a b c : α} + +/-- Pullback a `comm_ring` instance along an injective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.comm_ring + [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] + [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] [has_int_cast β] + (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + comm_ring β := +{ .. hf.ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast, + .. hf.comm_semigroup f mul } + +/-- Pushforward a `comm_ring` instance along a surjective function. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.surjective.comm_ring + [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] + [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] [has_int_cast β] + (f : α → β) (hf : surjective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + comm_ring β := +{ .. hf.ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast, + .. hf.comm_semigroup f mul } + +end comm_ring diff --git a/src/category_theory/conj.lean b/src/category_theory/conj.lean index ffd088fa0e901..375f959fb0159 100644 --- a/src/category_theory/conj.lean +++ b/src/category_theory/conj.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import algebra.hom.equiv.units import category_theory.endomorphism /-! diff --git a/src/category_theory/endomorphism.lean b/src/category_theory/endomorphism.lean index c12f58e70f7f0..d65377635225f 100644 --- a/src/category_theory/endomorphism.lean +++ b/src/category_theory/endomorphism.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Scott Morrison, Simon Hudon -/ -import algebra.hom.equiv +import algebra.hom.equiv.basic import category_theory.groupoid import category_theory.opposites import group_theory.group_action.defs diff --git a/src/category_theory/groupoid/subgroupoid.lean b/src/category_theory/groupoid/subgroupoid.lean index 2d988374df0d9..cd719f2a760c8 100644 --- a/src/category_theory/groupoid/subgroupoid.lean +++ b/src/category_theory/groupoid/subgroupoid.lean @@ -8,7 +8,7 @@ import category_theory.groupoid.basic import category_theory.groupoid import algebra.group.defs import algebra.hom.group -import algebra.hom.equiv +import algebra.hom.equiv.basic import data.set.lattice import combinatorics.quiver.connected_component import group_theory.subgroup.basic diff --git a/src/category_theory/groupoid/vertex_group.lean b/src/category_theory/groupoid/vertex_group.lean index a65b09cc7ea28..58c8c6e968a1b 100644 --- a/src/category_theory/groupoid/vertex_group.lean +++ b/src/category_theory/groupoid/vertex_group.lean @@ -7,7 +7,7 @@ import category_theory.groupoid import category_theory.path_category import algebra.group.defs import algebra.hom.group -import algebra.hom.equiv +import algebra.hom.equiv.basic import combinatorics.quiver.path /-! diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index 0b9423790895f..0aa4c141fb0ee 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yakov Pechersky. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yakov Pechersky -/ - +import algebra.hom.commute import algebra.big_operators.basic /-! diff --git a/src/data/int/basic.lean b/src/data/int/basic.lean index 75bac92ff5fa6..9fe9075ee9b3b 100644 --- a/src/data/int/basic.lean +++ b/src/data/int/basic.lean @@ -3,6 +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 -/ +import order.monotone import data.nat.basic /-! diff --git a/src/data/int/modeq.lean b/src/data/int/modeq.lean index c50554a490059..d03364f443f1d 100644 --- a/src/data/int/modeq.lean +++ b/src/data/int/modeq.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ import data.nat.modeq +import algebra.euclidean_domain.basic import tactic.ring /-! diff --git a/src/data/int/order.lean b/src/data/int/order.lean index 845d233b30994..0890809f5441a 100644 --- a/src/data/int/order.lean +++ b/src/data/int/order.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ import data.int.basic +import algebra.group_with_zero.divisibility import algebra.order.ring.abs import algebra.order.ring.nontrivial import algebra.char_zero.defs diff --git a/src/data/nat/order.lean b/src/data/nat/order.lean index 9b3dc61b2b4ae..e175210f43d7a 100644 --- a/src/data/nat/order.lean +++ b/src/data/nat/order.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import algebra.ring.divisibility +import algebra.group_with_zero.divisibility import algebra.order.ring.canonical import algebra.order.with_zero import data.nat.basic diff --git a/src/data/nat/part_enat.lean b/src/data/nat/part_enat.lean index 305b0e17db444..772fa87fa06d1 100644 --- a/src/data/nat/part_enat.lean +++ b/src/data/nat/part_enat.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ import order.well_founded -import algebra.hom.equiv +import algebra.hom.equiv.basic import data.part import data.enat.basic import tactic.norm_num diff --git a/src/data/ordmap/ordnode.lean b/src/data/ordmap/ordnode.lean index 899fe54c74d88..3f03aeca20423 100644 --- a/src/data/ordmap/ordnode.lean +++ b/src/data/ordmap/ordnode.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import order.compare import data.list.defs import data.nat.psub diff --git a/src/data/ordmap/ordset.lean b/src/data/ordmap/ordset.lean index 6b23c1b79c2be..bf6dbbdba64a5 100644 --- a/src/data/ordmap/ordset.lean +++ b/src/data/ordmap/ordset.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import data.ordmap.ordnode -import algebra.order.ring.basic +import algebra.order.ring.defs import data.nat.dist import tactic.linarith diff --git a/src/data/pi/lex.lean b/src/data/pi/lex.lean index 692a93c700abc..18603396c2809 100644 --- a/src/data/pi/lex.lean +++ b/src/data/pi/lex.lean @@ -5,7 +5,7 @@ Authors: Chris Hughes -/ import order.well_founded import algebra.group.pi -import algebra.order.group.basic +import algebra.order.group.defs import order.min_max diff --git a/src/data/rat/basic.lean b/src/data/rat/basic.lean index 089d227ffda4a..1ad0a5bdba2ee 100644 --- a/src/data/rat/basic.lean +++ b/src/data/rat/basic.lean @@ -3,7 +3,8 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import algebra.euclidean_domain +import algebra.field.defs +import data.rat.defs /-! # Field Structure on the Rational Numbers diff --git a/src/data/rat/floor.lean b/src/data/rat/floor.lean index 4b76ae26228f2..b5096eae42c92 100644 --- a/src/data/rat/floor.lean +++ b/src/data/rat/floor.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Kevin Kappelmann -/ import algebra.order.floor +import algebra.euclidean_domain.instances import tactic.field_simp /-! diff --git a/src/deprecated/group.lean b/src/deprecated/group.lean index d5e6f30d512be..c85156a5faca6 100644 --- a/src/deprecated/group.lean +++ b/src/deprecated/group.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import algebra.group.type_tags -import algebra.hom.equiv +import algebra.hom.equiv.basic import algebra.hom.ring import algebra.hom.units diff --git a/src/group_theory/congruence.lean b/src/group_theory/congruence.lean index 2ad849299b97a..ecf6cd1c003ec 100644 --- a/src/group_theory/congruence.lean +++ b/src/group_theory/congruence.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston -/ import algebra.group.prod -import algebra.hom.equiv +import algebra.hom.equiv.basic import data.setoid.basic import group_theory.submonoid.operations diff --git a/src/group_theory/free_abelian_group_finsupp.lean b/src/group_theory/free_abelian_group_finsupp.lean index d648ab34ff5a6..4eaefccabea1f 100644 --- a/src/group_theory/free_abelian_group_finsupp.lean +++ b/src/group_theory/free_abelian_group_finsupp.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ +import algebra.hom.equiv.type_tags import algebra.module.equiv import data.finsupp.defs import group_theory.free_abelian_group diff --git a/src/group_theory/group_action/defs.lean b/src/group_theory/group_action/defs.lean index 517f23ece2d32..a21958228a9ff 100644 --- a/src/group_theory/group_action/defs.lean +++ b/src/group_theory/group_action/defs.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Yury Kudryashov -/ import algebra.group.type_tags +import algebra.group.commute import algebra.hom.group import algebra.opposites import logic.embedding diff --git a/src/group_theory/group_action/opposite.lean b/src/group_theory/group_action/opposite.lean index cd5674bf9bd1e..1d5257e0532c9 100644 --- a/src/group_theory/group_action/opposite.lean +++ b/src/group_theory/group_action/opposite.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import algebra.group.opposite +import algebra.group_with_zero.basic import group_theory.group_action.defs /-! diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index a06f4fcb0e867..b303326d9dcba 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -8,6 +8,7 @@ import group_theory.submonoid.membership import group_theory.submonoid.centralizer import algebra.group.conj import algebra.module.basic +import algebra.order.group.inj_surj import order.atoms import order.sup_indep diff --git a/src/group_theory/submonoid/operations.lean b/src/group_theory/submonoid/operations.lean index 401f552008396..adaf03fb24014 100644 --- a/src/group_theory/submonoid/operations.lean +++ b/src/group_theory/submonoid/operations.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Kenny Lau, Johan Commelin, Mario Carneiro, Kevin Buzzard, Amelia Livingston, Yury Kudryashov -/ -import algebra.order.group.basic +import algebra.order.monoid.cancel.basic +import algebra.order.group.defs import group_theory.group_action.defs import group_theory.submonoid.basic import group_theory.subsemigroup.operations diff --git a/src/order/iterate.lean b/src/order/iterate.lean index 2b0b78c82ce61..89c2f4111e018 100644 --- a/src/order/iterate.lean +++ b/src/order/iterate.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ import logic.function.iterate +import order.monotone import data.nat.basic /-! diff --git a/src/order/semiconj_Sup.lean b/src/order/semiconj_Sup.lean index 7d3148c8c3842..810bebcc0a8a6 100644 --- a/src/order/semiconj_Sup.lean +++ b/src/order/semiconj_Sup.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ -import algebra.hom.equiv +import algebra.hom.equiv.basic import logic.function.conjugate import order.bounds.order_iso import order.conditionally_complete_lattice diff --git a/src/ring_theory/dedekind_domain/selmer_group.lean b/src/ring_theory/dedekind_domain/selmer_group.lean index dce7f6f6cf459..433814c96a50e 100644 --- a/src/ring_theory/dedekind_domain/selmer_group.lean +++ b/src/ring_theory/dedekind_domain/selmer_group.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 David Kurniadi Angdinata. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata -/ - +import algebra.hom.equiv.type_tags import data.zmod.quotient import ring_theory.dedekind_domain.adic_valuation diff --git a/src/ring_theory/principal_ideal_domain.lean b/src/ring_theory/principal_ideal_domain.lean index c0603f90d6d8f..4906cdcf429a9 100644 --- a/src/ring_theory/principal_ideal_domain.lean +++ b/src/ring_theory/principal_ideal_domain.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Morenikeji Neri -/ +import algebra.euclidean_domain.instances import ring_theory.unique_factorization_domain /-! diff --git a/src/tactic/linarith/lemmas.lean b/src/tactic/linarith/lemmas.lean index 391600f131e19..7fef5adb7fbd3 100644 --- a/src/tactic/linarith/lemmas.lean +++ b/src/tactic/linarith/lemmas.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert Y. Lewis -/ -import algebra.order.ring.basic +import algebra.order.ring.defs /-! # Lemmas for `linarith` diff --git a/src/tactic/monotonicity/lemmas.lean b/src/tactic/monotonicity/lemmas.lean index 470c17e7070dc..4912d311ab8c8 100644 --- a/src/tactic/monotonicity/lemmas.lean +++ b/src/tactic/monotonicity/lemmas.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon -/ import algebra.order.group.abs -import algebra.order.ring.basic +import algebra.order.ring.defs import algebra.order.sub.canonical import data.set.lattice import tactic.monotonicity.basic diff --git a/test/assert_exists/test2.lean b/test/assert_exists/test2.lean index 5fb5e1fc0cb39..1e3453a11c3d9 100644 --- a/test/assert_exists/test2.lean +++ b/test/assert_exists/test2.lean @@ -1,5 +1,5 @@ import tactic.assert_exists -import algebra.order.ring.basic +import algebra.order.ring.defs import data.int.basic assert_exists int diff --git a/test/finish4.lean b/test/finish4.lean index ed900d192fae7..8b75d9f0a874b 100644 --- a/test/finish4.lean +++ b/test/finish4.lean @@ -7,7 +7,7 @@ Tests for `finish using [...]` -/ import tactic.finish -import algebra.order.ring.basic +import algebra.order.ring.defs section list_rev open list diff --git a/test/library_search/ordered_ring.lean b/test/library_search/ordered_ring.lean index fb33b229dabd6..1a111ab6a2c61 100644 --- a/test/library_search/ordered_ring.lean +++ b/test/library_search/ordered_ring.lean @@ -5,7 +5,7 @@ Authors: Scott Morrison -/ import tactic.basic import data.nat.order -import algebra.order.ring.basic +import algebra.order.ring.defs /- Turn off trace messages so they don't pollute the test build: -/ set_option trace.silence_library_search true diff --git a/test/monotonicity.lean b/test/monotonicity.lean index 1fc83c445a805..30ee14d29d5ca 100644 --- a/test/monotonicity.lean +++ b/test/monotonicity.lean @@ -5,7 +5,7 @@ Authors: Simon Hudon -/ import tactic.monotonicity import tactic.norm_num -import algebra.order.ring.basic +import algebra.order.ring.defs import measure_theory.measure.lebesgue import measure_theory.function.locally_integrable import data.list.defs diff --git a/test/nontriviality.lean b/test/nontriviality.lean index 7174ec3059737..befc053c44ee4 100644 --- a/test/nontriviality.lean +++ b/test/nontriviality.lean @@ -1,5 +1,5 @@ import logic.nontrivial -import algebra.order.ring.basic +import algebra.order.ring.defs import data.nat.basic /-! ### Test `nontriviality` with inequality hypotheses -/