Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit b494f8f

Browse files
committed
split(algebra/order/nonneg): Separate ring and field instances (#17348)
Split `algebra.order.nonneg` into: * `algebra.order.nonneg.ring` for ring instances * `algebra.order.nonneg.field` for field and archimedean instances
1 parent 8108363 commit b494f8f

File tree

4 files changed

+107
-81
lines changed

4 files changed

+107
-81
lines changed
Lines changed: 95 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,95 @@
1+
/-
2+
Copyright (c) 2021 Floris van Doorn. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Floris van Doorn
5+
-/
6+
import algebra.order.archimedean
7+
import algebra.order.nonneg.ring
8+
9+
/-!
10+
# Semifield structure on the type of nonnegative elements
11+
12+
This file defines instances and prove some properties about the nonnegative elements
13+
`{x : α // 0 ≤ x}` of an arbitrary type `α`.
14+
15+
This is used to derive algebraic structures on `ℝ≥0` and `ℚ≥0` automatically.
16+
17+
## Main declarations
18+
19+
* `{x : α // 0 ≤ x}` is a `canonically_linear_ordered_semifield` if `α` is a `linear_ordered_field`.
20+
-/
21+
22+
open set
23+
24+
variables {α : Type*}
25+
26+
namespace nonneg
27+
28+
section linear_ordered_semifield
29+
variables [linear_ordered_semifield α] {x y : α}
30+
31+
instance has_inv : has_inv {x : α // 0 ≤ x} := ⟨λ x, ⟨x⁻¹, inv_nonneg.2 x.2⟩⟩
32+
33+
@[simp, norm_cast]
34+
protected lemma coe_inv (a : {x : α // 0 ≤ x}) : ((a⁻¹ : {x : α // 0 ≤ x}) : α) = a⁻¹ := rfl
35+
36+
@[simp] lemma inv_mk (hx : 0 ≤ x) : (⟨x, hx⟩ : {x : α // 0 ≤ x})⁻¹ = ⟨x⁻¹, inv_nonneg.2 hx⟩ := rfl
37+
38+
instance has_div : has_div {x : α // 0 ≤ x} := ⟨λ x y, ⟨x / y, div_nonneg x.2 y.2⟩⟩
39+
40+
@[simp, norm_cast] protected lemma coe_div (a b : {x : α // 0 ≤ x}) :
41+
((a / b : {x : α // 0 ≤ x}) : α) = a / b := rfl
42+
43+
@[simp] lemma mk_div_mk (hx : 0 ≤ x) (hy : 0 ≤ y) :
44+
(⟨x, hx⟩ : {x : α // 0 ≤ x}) / ⟨y, hy⟩ = ⟨x / y, div_nonneg hx hy⟩ := rfl
45+
46+
instance has_zpow : has_pow {x : α // 0 ≤ x} ℤ := ⟨λ a n, ⟨a ^ n, zpow_nonneg a.2 _⟩⟩
47+
48+
@[simp, norm_cast] protected lemma coe_zpow (a : {x : α // 0 ≤ x}) (n : ℤ) :
49+
((a ^ n : {x : α // 0 ≤ x}) : α) = a ^ n := rfl
50+
51+
@[simp] lemma mk_zpow (hx : 0 ≤ x) (n : ℤ) :
52+
(⟨x, hx⟩ : {x : α // 0 ≤ x}) ^ n = ⟨x ^ n, zpow_nonneg hx n⟩ := rfl
53+
54+
instance linear_ordered_semifield : linear_ordered_semifield {x : α // 0 ≤ x} :=
55+
subtype.coe_injective.linear_ordered_semifield _ nonneg.coe_zero nonneg.coe_one nonneg.coe_add
56+
nonneg.coe_mul nonneg.coe_inv nonneg.coe_div (λ _ _, rfl) nonneg.coe_pow nonneg.coe_zpow
57+
nonneg.coe_nat_cast (λ _ _, rfl) (λ _ _, rfl)
58+
59+
end linear_ordered_semifield
60+
61+
instance canonically_linear_ordered_semifield [linear_ordered_field α] :
62+
canonically_linear_ordered_semifield {x : α // 0 ≤ x} :=
63+
{ ..nonneg.linear_ordered_semifield, ..nonneg.canonically_ordered_comm_semiring }
64+
65+
instance linear_ordered_comm_group_with_zero [linear_ordered_field α] :
66+
linear_ordered_comm_group_with_zero {x : α // 0 ≤ x} :=
67+
infer_instance
68+
69+
/-! ### Floor -/
70+
71+
instance archimedean [ordered_add_comm_monoid α] [archimedean α] : archimedean {x : α // 0 ≤ x} :=
72+
⟨λ x y hy,
73+
let ⟨n, hr⟩ := archimedean.arch (x : α) (hy : (0 : α) < y) in
74+
⟨n, show (x : α) ≤ (n • y : {x : α // 0 ≤ x}), by simp [*, -nsmul_eq_mul, nsmul_coe]⟩⟩
75+
76+
instance floor_semiring [ordered_semiring α] [floor_semiring α] : floor_semiring {r : α // 0 ≤ r} :=
77+
{ floor := λ a, ⌊(a : α)⌋₊,
78+
ceil := λ a, ⌈(a : α)⌉₊,
79+
floor_of_neg := λ a ha, floor_semiring.floor_of_neg ha,
80+
gc_floor := λ a n ha, begin
81+
refine (floor_semiring.gc_floor (show 0 ≤ (a : α), from ha)).trans _,
82+
rw [←subtype.coe_le_coe, nonneg.coe_nat_cast]
83+
end,
84+
gc_ceil := λ a n, begin
85+
refine (floor_semiring.gc_ceil (a : α) n).trans _,
86+
rw [←subtype.coe_le_coe, nonneg.coe_nat_cast]
87+
end}
88+
89+
@[norm_cast] lemma nat_floor_coe [ordered_semiring α] [floor_semiring α] (a : {r : α // 0 ≤ r}) :
90+
⌊(a : α)⌋₊ = ⌊a⌋₊ := rfl
91+
92+
@[norm_cast] lemma nat_ceil_coe [ordered_semiring α] [floor_semiring α] (a : {r : α // 0 ≤ r}) :
93+
⌈(a : α)⌉₊ = ⌈a⌉₊ := rfl
94+
95+
end nonneg

src/algebra/order/nonneg.lean renamed to src/algebra/order/nonneg/ring.lean

Lines changed: 10 additions & 79 deletions
Original file line numberDiff line numberDiff line change
@@ -3,9 +3,9 @@ Copyright (c) 2021 Floris van Doorn. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Floris van Doorn
55
-/
6-
import algebra.order.archimedean
7-
import order.lattice_intervals
6+
import algebra.order.ring
87
import order.complete_lattice_intervals
8+
import order.lattice_intervals
99

1010
/-!
1111
# The type of nonnegative elements
@@ -20,7 +20,6 @@ When `α` is `ℝ`, this will give us some properties about `ℝ≥0`.
2020
## Main declarations
2121
2222
* `{x : α // 0 ≤ x}` is a `canonically_linear_ordered_add_monoid` if `α` is a `linear_ordered_ring`.
23-
* `{x : α // 0 ≤ x}` is a `linear_ordered_comm_group_with_zero` if `α` is a `linear_ordered_field`.
2423
2524
## Implementation Notes
2625
@@ -149,11 +148,6 @@ lemma nsmul_coe [ordered_add_comm_monoid α] (n : ℕ) (r : {x : α // 0 ≤ x})
149148
↑(n • r) = n • (r : α) :=
150149
nonneg.coe_add_monoid_hom.map_nsmul _ _
151150

152-
instance archimedean [ordered_add_comm_monoid α] [archimedean α] : archimedean {x : α // 0 ≤ x} :=
153-
⟨ assume x y pos_y,
154-
let ⟨n, hr⟩ := archimedean.arch (x : α) (pos_y : (0 : α) < y) in
155-
⟨n, show (x : α) ≤ (n • y : {x : α // 0 ≤ x}), by simp [*, -nsmul_eq_mul, nsmul_coe]⟩ ⟩
156-
157151
instance has_one [ordered_semiring α] : has_one {x : α // 0 ≤ x} :=
158152
{ one := ⟨1, zero_le_one⟩ }
159153

@@ -181,16 +175,21 @@ instance add_monoid_with_one [ordered_semiring α] : add_monoid_with_one {x : α
181175
nat_cast_succ := λ _, by simp [nat.cast]; refl,
182176
.. nonneg.has_one, .. nonneg.ordered_add_comm_monoid }
183177

178+
@[simp, norm_cast]
179+
protected lemma coe_nat_cast [ordered_semiring α] (n : ℕ) : ((↑n : {x : α // 0 ≤ x}) : α) = n := rfl
180+
181+
@[simp] lemma mk_nat_cast [ordered_semiring α] (n : ℕ) :
182+
(⟨n, n.cast_nonneg⟩ : {x : α // 0 ≤ x}) = n := rfl
183+
184184
instance has_pow [ordered_semiring α] : has_pow {x : α // 0 ≤ x} ℕ :=
185185
{ pow := λ x n, ⟨x ^ n, pow_nonneg x.2 n⟩ }
186186

187187
@[simp, norm_cast]
188188
protected lemma coe_pow [ordered_semiring α] (a : {x : α // 0 ≤ x}) (n : ℕ) :
189-
((a ^ n: {x : α // 0 ≤ x}) : α) = a ^ n := rfl
189+
((a ^ n) : α) = a ^ n := rfl
190190

191191
@[simp] lemma mk_pow [ordered_semiring α] {x : α} (hx : 0 ≤ x) (n : ℕ) :
192-
(⟨x, hx⟩ : {x : α // 0 ≤ x}) ^ n = ⟨x ^ n, pow_nonneg hx n⟩ :=
193-
rfl
192+
(⟨x, hx⟩ : {x : α // 0 ≤ x}) ^ n = ⟨x ^ n, pow_nonneg hx n⟩ := rfl
194193

195194
instance ordered_semiring [ordered_semiring α] : ordered_semiring {x : α // 0 ≤ x} :=
196195
subtype.coe_injective.ordered_semiring _
@@ -239,10 +238,6 @@ instance linear_ordered_comm_monoid_with_zero [linear_ordered_comm_ring α] :
239238
def coe_ring_hom [ordered_semiring α] : {x : α // 0 ≤ x} →+* α :=
240239
⟨coe, nonneg.coe_one, nonneg.coe_mul, nonneg.coe_zero, nonneg.coe_add⟩
241240

242-
@[simp, norm_cast]
243-
protected lemma coe_nat_cast [ordered_semiring α] (n : ℕ) : ((↑n : {x : α // 0 ≤ x}) : α) = n :=
244-
map_nat_cast (coe_ring_hom : {x : α // 0 ≤ x} →+* α) n
245-
246241
instance canonically_ordered_add_monoid [ordered_ring α] :
247242
canonically_ordered_add_monoid {x : α // 0 ≤ x} :=
248243
{ le_self_add := λ a b, le_add_of_nonneg_right b.2,
@@ -261,70 +256,6 @@ instance canonically_linear_ordered_add_monoid [linear_ordered_ring α] :
261256
canonically_linear_ordered_add_monoid {x : α // 0 ≤ x} :=
262257
{ ..subtype.linear_order _, ..nonneg.canonically_ordered_add_monoid }
263258

264-
section linear_ordered_semifield
265-
variables [linear_ordered_semifield α] {x y : α}
266-
267-
instance has_inv : has_inv {x : α // 0 ≤ x} := ⟨λ x, ⟨x⁻¹, inv_nonneg.mpr x.2⟩⟩
268-
269-
@[simp, norm_cast]
270-
protected lemma coe_inv (a : {x : α // 0 ≤ x}) : ((a⁻¹ : {x : α // 0 ≤ x}) : α) = a⁻¹ := rfl
271-
272-
@[simp] lemma inv_mk (hx : 0 ≤ x) : (⟨x, hx⟩ : {x : α // 0 ≤ x})⁻¹ = ⟨x⁻¹, inv_nonneg.mpr hx⟩ := rfl
273-
274-
instance has_div : has_div {x : α // 0 ≤ x} := ⟨λ x y, ⟨x / y, div_nonneg x.2 y.2⟩⟩
275-
276-
@[simp, norm_cast] protected lemma coe_div (a b : {x : α // 0 ≤ x}) :
277-
((a / b : {x : α // 0 ≤ x}) : α) = a / b := rfl
278-
279-
@[simp] lemma mk_div_mk (hx : 0 ≤ x) (hy : 0 ≤ y) :
280-
(⟨x, hx⟩ : {x : α // 0 ≤ x}) / ⟨y, hy⟩ = ⟨x / y, div_nonneg hx hy⟩ := rfl
281-
282-
instance has_zpow : has_pow {x : α // 0 ≤ x} ℤ := ⟨λ a n, ⟨a ^ n, zpow_nonneg a.2 _⟩⟩
283-
284-
@[simp, norm_cast] protected lemma coe_zpow (a : {x : α // 0 ≤ x}) (n : ℤ) :
285-
((a ^ n : {x : α // 0 ≤ x}) : α) = a ^ n := rfl
286-
287-
@[simp] lemma mk_zpow (hx : 0 ≤ x) (n : ℤ) :
288-
(⟨x, hx⟩ : {x : α // 0 ≤ x}) ^ n = ⟨x ^ n, zpow_nonneg hx n⟩ := rfl
289-
290-
instance linear_ordered_semifield : linear_ordered_semifield {x : α // 0 ≤ x} :=
291-
subtype.coe_injective.linear_ordered_semifield _ nonneg.coe_zero nonneg.coe_one nonneg.coe_add
292-
nonneg.coe_mul nonneg.coe_inv nonneg.coe_div (λ _ _, rfl) nonneg.coe_pow nonneg.coe_zpow
293-
nonneg.coe_nat_cast (λ _ _, rfl) (λ _ _, rfl)
294-
295-
end linear_ordered_semifield
296-
297-
instance linear_ordered_comm_group_with_zero [linear_ordered_field α] :
298-
linear_ordered_comm_group_with_zero {x : α // 0 ≤ x} :=
299-
{ inv_zero := by { ext, exact inv_zero },
300-
mul_inv_cancel := by { intros a ha, ext, refine mul_inv_cancel (mt (λ h, _) ha), ext, exact h },
301-
..nonneg.nontrivial,
302-
..nonneg.has_inv,
303-
..nonneg.linear_ordered_comm_monoid_with_zero }
304-
305-
instance canonically_linear_ordered_semifield [linear_ordered_field α] :
306-
canonically_linear_ordered_semifield {x : α // 0 ≤ x} :=
307-
{ ..nonneg.linear_ordered_semifield, ..nonneg.canonically_ordered_comm_semiring }
308-
309-
instance floor_semiring [ordered_semiring α] [floor_semiring α] : floor_semiring {r : α // 0 ≤ r} :=
310-
{ floor := λ a, ⌊(a : α)⌋₊,
311-
ceil := λ a, ⌈(a : α)⌉₊,
312-
floor_of_neg := λ a ha, floor_semiring.floor_of_neg ha,
313-
gc_floor := λ a n ha, begin
314-
refine (floor_semiring.gc_floor (show 0 ≤ (a : α), from ha)).trans _,
315-
rw [←subtype.coe_le_coe, nonneg.coe_nat_cast]
316-
end,
317-
gc_ceil := λ a n, begin
318-
refine (floor_semiring.gc_ceil (a : α) n).trans _,
319-
rw [←subtype.coe_le_coe, nonneg.coe_nat_cast]
320-
end}
321-
322-
@[norm_cast] lemma nat_floor_coe [ordered_semiring α] [floor_semiring α] (a : {r : α // 0 ≤ r}) :
323-
⌊(a : α)⌋₊ = ⌊a⌋₊ := rfl
324-
325-
@[norm_cast] lemma nat_ceil_coe [ordered_semiring α] [floor_semiring α] (a : {r : α // 0 ≤ r}) :
326-
⌈(a : α)⌉₊ = ⌈a⌉₊ := rfl
327-
328259
section linear_order
329260

330261
variables [has_zero α] [linear_order α]

src/data/rat/nnrat.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Yaël Dillies, Bhavik Mehta
55
-/
66
import algebra.algebra.basic
7-
import algebra.order.nonneg
7+
import algebra.order.nonneg.field
88

99
/-!
1010
# Nonnegative rationals

src/data/real/nnreal.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johan Commelin
55
-/
66
import algebra.algebra.basic
7-
import algebra.order.nonneg
7+
import algebra.order.nonneg.field
88
import data.real.pointwise
99
import tactic.positivity
1010

0 commit comments

Comments
 (0)