Skip to content
New issue

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

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

Already on GitHub? Sign in to your account

[Merged by Bors] - chore(algebra): displace zero_ne_one_class with nonzero and make no_zero_divisors a Prop #2847

Closed
wants to merge 53 commits into from
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
Show all changes
53 commits
Select commit Hold shift + click to select a range
d1df178
chore(algebra): displace zero_ne_one_class with nonzero
kckennylau May 28, 2020
889420a
algebra.ring
kckennylau May 28, 2020
e142e51
Update src/algebra/ring.lean
kckennylau May 28, 2020
8fc64cc
update algebra.ring
kckennylau May 28, 2020
ba5497a
algebra.group_with_zero
kckennylau May 28, 2020
4f7929c
algebra.group.with_one
kckennylau May 28, 2020
3a0f538
ring_theory.prod
kckennylau May 28, 2020
217edd8
algebra.field
kckennylau May 28, 2020
5fcca92
algebra.opposites
kckennylau May 28, 2020
46b0420
init_.data.int.basic
kckennylau May 28, 2020
42f4d11
algebra.ordered_ring
kckennylau May 28, 2020
24c18fb
data.equiv.transfer_instances
kckennylau May 28, 2020
ce14467
init_.data.nat.lemmas
kckennylau May 28, 2020
d34f6f2
algebra.euclidean_domain
kckennylau May 28, 2020
778c52d
data.equiv.ring
kckennylau May 28, 2020
797c796
data.rat.basic
kckennylau May 28, 2020
16880d2
algebra.associated
kckennylau May 28, 2020
953c707
algebra.gcd_domain
kckennylau May 28, 2020
7c51a9c
data.zsqrtd.basic
kckennylau May 28, 2020
1001e52
set_theory.cardinal
kckennylau May 28, 2020
1bfa31b
field_theory.subfield
kckennylau May 28, 2020
a0d1d13
ring_theory.fintype
kckennylau May 28, 2020
d3146b3
algebra.char_p
kckennylau May 28, 2020
09263be
data.real.basic
kckennylau May 28, 2020
e09e2ad
data.zmod.basic
kckennylau May 28, 2020
2f93b3e
data.real.nnreal
kckennylau May 28, 2020
95897ff
order.filter.filter_product
kckennylau May 28, 2020
821eb6c
data.matrix.pequiv
kckennylau May 28, 2020
19f41df
ring_theory.ideals
kckennylau May 28, 2020
acf8982
data.complex.exponential
kckennylau May 28, 2020
d4fe0eb
set_theory.ordinal
kckennylau May 28, 2020
27cd8a3
linear_algebra.basis
kckennylau May 28, 2020
f51a240
data.polynomial
kckennylau May 28, 2020
4bacb83
ring_theory.ideal_operations
kckennylau May 28, 2020
0b3ca61
ring_theory.noetherian
kckennylau May 28, 2020
b0a7e7e
data.padics.padic_integers
kckennylau May 28, 2020
b1382b2
algebra.euclidean_domain
kckennylau May 29, 2020
6110deb
algebra.euclidean_domain: make more things protected
kckennylau May 29, 2020
396edbf
algebra.classical_lie_algebras
kckennylau May 29, 2020
b315a3e
data.polynomial
kckennylau May 29, 2020
77b7f76
data.mv_polynomial
kckennylau May 29, 2020
eb3fbe4
data.zsqrtd.gaussian_int
kckennylau May 29, 2020
da3bf8e
ring_theory.polynomial
kckennylau May 29, 2020
3d96bc9
ring_theory.free_comm_ring
kckennylau May 29, 2020
94a474f
ring_theory.algebraic
kckennylau May 29, 2020
a0a6e6f
algebra.direct_limit
kckennylau May 29, 2020
44db175
ring_theory.power_series
kckennylau May 29, 2020
aa83f7f
analysis.special_functions.pow
kckennylau May 29, 2020
5f693c0
ring_theory.fractional_ideal
kckennylau May 29, 2020
30ffafe
Merge remote-tracking branch 'origin/master' into nonzero
kckennylau May 29, 2020
f7835fa
fix errors
kckennylau May 29, 2020
ae3abc2
it has gotten faster
kckennylau May 29, 2020
6fb14e9
Merge branch 'master' into nonzero
kckennylau May 29, 2020
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Jump to
Jump to file
Failed to load files.
Diff view
Diff view
2 changes: 1 addition & 1 deletion src/algebra/associated.lean
Original file line number Diff line number Diff line change
Expand Up @@ -18,7 +18,7 @@ variables {α : Type*} {β : Type*} {γ : Type*} {δ : Type*}
refine ⟨⟨0, 0, _, _⟩, rfl⟩; apply subsingleton.elim
end⟩

@[simp] theorem not_is_unit_zero [nonzero_semiring α] : ¬ is_unit (0 : α) :=
@[simp] theorem not_is_unit_zero [semiring α] [nonzero α] : ¬ is_unit (0 : α) :=
mt is_unit_zero_iff.1 zero_ne_one

lemma is_unit_pow [monoid α] {a : α} (n : ℕ) : is_unit a → is_unit (a ^ n) :=
Expand Down
4 changes: 2 additions & 2 deletions src/algebra/char_p.lean
Original file line number Diff line number Diff line change
Expand Up @@ -258,10 +258,10 @@ calc r = 1 * r : by rw one_mul

end prio

lemma false_of_nonzero_of_char_one [nonzero_comm_ring R] [char_p R 1] : false :=
lemma false_of_nonzero_of_char_one [semiring R] [nonzero R] [char_p R 1] : false :=
zero_ne_one $ show (0:R) = 1, from subsingleton.elim 0 1

lemma ring_char_ne_one [nonzero_semiring R] : ring_char R ≠ 1 :=
lemma ring_char_ne_one [semiring R] [nonzero R] : ring_char R ≠ 1 :=
by { intros h, apply @zero_ne_one R, symmetry, rw [←nat.cast_one, ring_char.spec, h], }

end char_one
Expand Down
4 changes: 2 additions & 2 deletions src/algebra/classical_lie_algebras.lean
Original file line number Diff line number Diff line change
Expand Up @@ -28,7 +28,7 @@ namespace lie_algebra
open_locale matrix

variables (n : Type u₁) (R : Type u₂)
variables [fintype n] [decidable_eq n] [nonzero_comm_ring R]
variables [fintype n] [decidable_eq n] [comm_ring R]

local attribute [instance] matrix.lie_ring
local attribute [instance] matrix.lie_algebra
Expand Down Expand Up @@ -77,7 +77,7 @@ def Eb (h : j ≠ i) : sl n R :=

end elementary_basis

lemma sl_non_abelian (h : 1 < fintype.card n) : ¬lie_algebra.is_abelian ↥(sl n R) :=
lemma sl_non_abelian [nonzero R] (h : 1 < fintype.card n) : ¬lie_algebra.is_abelian ↥(sl n R) :=
begin
rcases fintype.exists_pair_of_one_lt_card h with ⟨i, j, hij⟩,
let A := Eb R i j hij,
Expand Down
8 changes: 4 additions & 4 deletions src/algebra/direct_limit.lean
Original file line number Diff line number Diff line change
Expand Up @@ -497,15 +497,14 @@ variables [directed_system G f]

namespace direct_limit

instance nonzero_comm_ring : nonzero_comm_ring (ring.direct_limit G f) :=
instance nonzero : nonzero (ring.direct_limit G f) :=
{ zero_ne_one := nonempty.elim (by apply_instance) $ assume i : ι, begin
change (0 : ring.direct_limit G f) ≠ 1,
rw ← ring.direct_limit.of_one,
intros H, rcases ring.direct_limit.of.zero_exact H.symm with ⟨j, hij, hf⟩,
rw is_ring_hom.map_one (f i j hij) at hf,
exact one_ne_zero hf
end,
.. ring.direct_limit.comm_ring G f }
end }

theorem exists_inv {p : ring.direct_limit G f} : p ≠ 0 → ∃ y, p * y = 1 :=
ring.direct_limit.induction_on p $ λ i x H,
Expand All @@ -529,7 +528,8 @@ protected noncomputable def field : field (ring.direct_limit G f) :=
{ inv := inv G f,
mul_inv_cancel := λ p, direct_limit.mul_inv_cancel G f,
inv_zero := dif_pos rfl,
.. direct_limit.nonzero_comm_ring G f }
.. ring.direct_limit.comm_ring G f,
.. direct_limit.nonzero G f }

end

Expand Down
17 changes: 11 additions & 6 deletions src/algebra/euclidean_domain.lean
Original file line number Diff line number Diff line change
Expand Up @@ -12,9 +12,9 @@ universe u

section prio
set_option default_priority 100 -- see Note [default priority]

set_option old_structure_cmd true
@[protect_proj without mul_left_not_lt r_well_founded]
class euclidean_domain (α : Type u) extends nonzero_comm_ring α :=
class euclidean_domain (α : Type u) extends comm_ring α :=
(quotient : α → α → α)
(quotient_zero : ∀ a, quotient a 0 = 0)
(remainder : α → α → α)
Expand All @@ -32,6 +32,7 @@ class euclidean_domain (α : Type u) extends nonzero_comm_ring α :=
function from weak to strong. I've currently divided the lemmas into
strong and weak depending on whether they require `val_le_mul_left` or not. -/
(mul_left_not_lt : ∀ a {b}, b ≠ 0 → ¬r (a * b) a)
(zero_ne_one : (0 : α) ≠ 1)
end prio

namespace euclidean_domain
Expand All @@ -40,6 +41,9 @@ variables [euclidean_domain α]

local infix ` ≺ `:50 := euclidean_domain.r

@[priority 70] -- see Note [lower instance priority]
instance : nonzero α := ⟨euclidean_domain.zero_ne_one⟩

@[priority 70] -- see Note [lower instance priority]
instance : has_div α := ⟨euclidean_domain.quotient⟩

Expand Down Expand Up @@ -328,8 +332,6 @@ end lcm

end euclidean_domain

open euclidean_domain

instance int.euclidean_domain : euclidean_domain ℤ :=
{ quotient := (/),
quotient_zero := int.div_zero,
Expand All @@ -342,7 +344,9 @@ instance int.euclidean_domain : euclidean_domain ℤ :=
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 _) }
exact mul_le_mul_of_nonneg_left (int.nat_abs_pos_of_ne_zero b0) (nat.zero_le _),
.. int.comm_ring,
.. int.nonzero }

@[priority 100] -- see Note [lower instance priority]
instance field.to_euclidean_domain {K : Type u} [field K] : euclidean_domain K :=
Expand All @@ -355,4 +359,5 @@ instance field.to_euclidean_domain {K : Type u} [field K] : euclidean_domain K :
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 }
mul_left_not_lt := λ a b hnb ⟨hab, hna⟩, or.cases_on (mul_eq_zero.1 hab) hna hnb,
.. ‹field K› }
11 changes: 8 additions & 3 deletions src/algebra/field.lean
Original file line number Diff line number Diff line change
Expand Up @@ -13,15 +13,19 @@ set_option old_structure_cmd true
universe u
variables {α : Type u}

@[protect_proj, ancestor ring has_inv zero_ne_one_class]
class division_ring (α : Type u) extends ring α, has_inv α, zero_ne_one_class α :=
@[protect_proj, ancestor ring has_inv]
class division_ring (α : Type u) extends ring α, has_inv α :=
(mul_inv_cancel : ∀ {a : α}, a ≠ 0 → a * a⁻¹ = 1)
(inv_mul_cancel : ∀ {a : α}, a ≠ 0 → a⁻¹ * a = 1)
(inv_zero : (0 : α)⁻¹ = 0)
(zero_ne_one : (0 : α) ≠ 1)

section division_ring
variables [division_ring α] {a b : α}

instance division_ring.to_nonzero : nonzero α :=
⟨division_ring.zero_ne_one⟩

protected definition algebra.div (a b : α) : α :=
a * b⁻¹

Expand Down Expand Up @@ -284,9 +288,10 @@ instance division_ring.to_domain : domain α :=
end division_ring

@[protect_proj, ancestor division_ring comm_ring]
class field (α : Type u) extends comm_ring α, has_inv α, zero_ne_one_class α :=
class field (α : Type u) extends comm_ring α, has_inv α :=
(mul_inv_cancel : ∀ {a : α}, a ≠ 0 → a * a⁻¹ = 1)
(inv_zero : (0 : α)⁻¹ = 0)
(zero_ne_one : (0 : α) ≠ 1)

section field

Expand Down
6 changes: 3 additions & 3 deletions src/algebra/gcd_domain.lean
Original file line number Diff line number Diff line change
Expand Up @@ -68,7 +68,7 @@ lemma normalize_eq_one {x : α} : normalize x = 1 ↔ is_unit x :=

theorem norm_unit_mul_norm_unit (a : α) : norm_unit (a * norm_unit a) = 1 :=
classical.by_cases (assume : a = 0, by simp only [this, norm_unit_zero, zero_mul]) $
assume h, by rw [norm_unit_mul h (units.ne_zero _), norm_unit_coe_units, mul_inv_eq_one]
assume h, by rw [norm_unit_mul h (units.coe_ne_zero _), norm_unit_coe_units, mul_inv_eq_one]

theorem normalize_idem (x : α) : normalize (normalize x) = normalize x :=
by rw [normalize, normalize, norm_unit_mul_norm_unit, units.coe_one, mul_one]
Expand All @@ -79,7 +79,7 @@ begin
rcases associated_of_dvd_dvd hab hba with ⟨u, rfl⟩,
refine classical.by_cases (by rintro rfl; simp only [zero_mul]) (assume ha : a ≠ 0, _),
suffices : a * ↑(norm_unit a) = a * ↑u * ↑(norm_unit a) * ↑u⁻¹,
by simpa only [normalize, mul_assoc, norm_unit_mul ha u.ne_zero, norm_unit_coe_units],
by simpa only [normalize, mul_assoc, norm_unit_mul ha u.coe_ne_zero, norm_unit_coe_units],
calc a * ↑(norm_unit a) = a * ↑(norm_unit a) * ↑u * ↑u⁻¹:
(units.mul_inv_cancel_right _ _).symm
... = a * ↑u * ↑(norm_unit a) * ↑u⁻¹ : by rw mul_right_comm a
Expand Down Expand Up @@ -282,7 +282,7 @@ iff.intro
(assume h : lcm a b = 0,
have normalize (a * b) = 0,
by rw [← gcd_mul_lcm _ _, h, mul_zero],
by simpa only [normalize_eq_zero, mul_eq_zero, units.ne_zero, or_false])
by simpa only [normalize_eq_zero, mul_eq_zero, units.coe_ne_zero, or_false])
(by rintro (rfl | rfl); [apply lcm_zero_left, apply lcm_zero_right])

@[simp] lemma normalize_lcm (a b : α) : normalize (lcm a b) = lcm a b :=
Expand Down
9 changes: 4 additions & 5 deletions src/algebra/group/with_one.lean
Original file line number Diff line number Diff line change
Expand Up @@ -117,10 +117,8 @@ namespace with_zero
instance [one : has_one α] : has_one (with_zero α) :=
{ ..one }

instance [has_one α] : zero_ne_one_class (with_zero α) :=
{ zero_ne_one := λ h, option.no_confusion h,
..with_zero.has_zero,
..with_zero.has_one }
instance [has_one α] : nonzero (with_zero α) :=
{ zero_ne_one := λ h, option.no_confusion h }

lemma coe_one [has_one α] : ((1 : α) : with_zero α) = 1 := rfl

Expand Down Expand Up @@ -159,7 +157,8 @@ instance [monoid α] : monoid (with_zero α) :=
| none := rfl
| some a := congr_arg some $ mul_one _
end,
..with_zero.zero_ne_one_class,
..with_zero.has_one,
..with_zero.nonzero,
..with_zero.semigroup }

instance [comm_monoid α] : comm_monoid (with_zero α) :=
Expand Down
7 changes: 5 additions & 2 deletions src/algebra/group_with_zero.lean
Original file line number Diff line number Diff line change
Expand Up @@ -50,10 +50,13 @@ The type is required to come with an “inverse” function, and the inverse of

Examples include division rings and the ordered monoids that are the
target of valuations in general valuation theory.-/
class group_with_zero (G₀ : Type*)
extends monoid_with_zero G₀, has_inv G₀, zero_ne_one_class G₀ :=
class group_with_zero (G₀ : Type*) extends monoid_with_zero G₀, has_inv G₀ :=
(inv_zero : (0 : G₀)⁻¹ = 0)
(mul_inv_cancel : ∀ a:G₀, a ≠ 0 → a * a⁻¹ = 1)
(zero_ne_one : (0 : G₀) ≠ 1)

instance group_with_zero.to_nonzero (G₀ : Type*) [group_with_zero G₀] : nonzero G₀ :=
⟨group_with_zero.zero_ne_one⟩

/-- A type `G₀` is a commutative “group with zero”
if it is a commutative monoid with zero element (distinct from `1`)
Expand Down
9 changes: 4 additions & 5 deletions src/algebra/opposites.lean
Original file line number Diff line number Diff line change
Expand Up @@ -109,20 +109,19 @@ instance [ring α] : ring (opposite α) :=
instance [comm_ring α] : comm_ring (opposite α) :=
{ .. opposite.ring α, .. opposite.comm_semigroup α }

instance [zero_ne_one_class α] : zero_ne_one_class (opposite α) :=
{ zero_ne_one := λ h, zero_ne_one $ op_inj h,
.. opposite.has_zero α, .. opposite.has_one α }
instance [has_zero α] [has_one α] [nonzero α] : nonzero (opposite α) :=
{ zero_ne_one := λ h : op (0 : α) = op 1, zero_ne_one (op_inj h) }

instance [integral_domain α] : integral_domain (opposite α) :=
{ eq_zero_or_eq_zero_of_mul_eq_zero := λ x y (H : op (_ * _) = op (0:α)),
or.cases_on (eq_zero_or_eq_zero_of_mul_eq_zero $ op_inj H)
(λ hy, or.inr $ unop_inj $ hy) (λ hx, or.inl $ unop_inj $ hx),
.. opposite.comm_ring α, .. opposite.zero_ne_one_class α }
.. opposite.comm_ring α, .. opposite.nonzero α }

instance [field α] : field (opposite α) :=
{ mul_inv_cancel := λ x hx, unop_inj $ inv_mul_cancel $ λ hx', hx $ unop_inj hx',
inv_zero := unop_inj inv_zero,
.. opposite.comm_ring α, .. opposite.zero_ne_one_class α, .. opposite.has_inv α }
.. opposite.comm_ring α, .. opposite.nonzero α, .. opposite.has_inv α }

@[simp] lemma op_zero [has_zero α] : op (0 : α) = 0 := rfl
@[simp] lemma unop_zero [has_zero α] : unop (0 : αᵒᵖ) = 0 := rfl
Expand Down
20 changes: 14 additions & 6 deletions src/algebra/ordered_ring.lean
Original file line number Diff line number Diff line change
Expand Up @@ -432,12 +432,16 @@ end decidable_linear_ordered_semiring
/-- An `ordered_ring α` is a ring `α` with a partial order such that
multiplication with a positive number and addition are monotone. -/
@[protect_proj]
class ordered_ring (α : Type u) extends ring α, ordered_add_comm_group α, zero_ne_one_class α :=
(mul_pos : ∀ a b : α, 0 < a → 0 < b → 0 < a * b)
class ordered_ring (α : Type u) extends ring α, ordered_add_comm_group α :=
(mul_pos : ∀ a b : α, 0 < a → 0 < b → 0 < a * b)
(zero_ne_one : (0 : α) ≠ 1)

section ordered_ring
variables [ordered_ring α] {a b c : α}

instance ordered_ring.to_nonzero : nonzero α :=
⟨ordered_ring.zero_ne_one⟩

lemma ordered_ring.mul_nonneg (a b : α) (h₁ : 0 ≤ a) (h₂ : 0 ≤ b) : 0 ≤ a * b :=
begin
cases classical.em (a ≤ 0), { simp [le_antisymm h h₁] },
Expand Down Expand Up @@ -813,10 +817,10 @@ end decidable_linear_ordered_comm_ring

/-- Extend `nonneg_add_comm_group` to support ordered rings
specified by their nonnegative elements -/
class nonneg_ring (α : Type*)
extends ring α, zero_ne_one_class α, nonneg_add_comm_group α :=
class nonneg_ring (α : Type*) extends ring α, nonneg_add_comm_group α :=
(mul_nonneg : ∀ {a b}, nonneg a → nonneg b → nonneg (a * b))
(mul_pos : ∀ {a b}, pos a → pos b → pos (a * b))
(zero_ne_one : (0 : α) ≠ 1)

/-- Extend `nonneg_add_comm_group` to support linearly ordered rings
specified by their nonnegative elements -/
Expand Down Expand Up @@ -925,14 +929,18 @@ end linear_nonneg_ring
in which `a ≤ b` iff there exists `c` with `b = a + c`. This is satisfied by the
natural numbers, for example, but not the integers or other ordered groups. -/
class canonically_ordered_comm_semiring (α : Type*) extends
canonically_ordered_add_monoid α, comm_semiring α, zero_ne_one_class α :=
canonically_ordered_add_monoid α, comm_semiring α :=
(mul_eq_zero_iff (a b : α) : a * b = 0 ↔ a = 0 ∨ b = 0)
(zero_ne_one : (0 : α) ≠ 1)

namespace canonically_ordered_semiring
variables [canonically_ordered_comm_semiring α] {a b : α}

open canonically_ordered_add_monoid (le_iff_exists_add)

instance canonically_ordered_comm_semiring.to_nonzero : nonzero α :=
⟨canonically_ordered_comm_semiring.zero_ne_one⟩

lemma mul_le_mul {a b c d : α} (hab : a ≤ b) (hcd : c ≤ d) :
a * c ≤ b * d :=
begin
Expand Down Expand Up @@ -1057,7 +1065,7 @@ instance : canonically_ordered_comm_semiring (with_top α) :=
mul_eq_zero_iff := mul_eq_zero,
one_mul := one_mul',
mul_one := assume a, by rw [comm, one_mul'],
zero_ne_one := assume h, @zero_ne_one α _ $ option.some.inj h,
zero_ne_one := assume h : ((0 : α) : with_top α) = 1, zero_ne_one $ option.some.inj h,
.. with_top.add_comm_monoid, .. with_top.mul_zero_class, .. with_top.canonically_ordered_add_monoid }

end with_top