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

Commit 6de241f

Browse files
committed
feat(data/rat/nnrat): Nonnegative rationals (#16283)
Define `nnrat`, the subtype of nonnegative rational numbers, and show it forms a `canonically_linear_ordered_semifield`.
1 parent de62604 commit 6de241f

File tree

1 file changed

+301
-0
lines changed

1 file changed

+301
-0
lines changed

src/data/rat/nnrat.lean

Lines changed: 301 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,301 @@
1+
/-
2+
Copyright (c) 2022 Yaël Dillies, Bhavik Mehta. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Yaël Dillies, Bhavik Mehta
5+
-/
6+
import algebra.algebra.basic
7+
import algebra.order.nonneg
8+
9+
/-!
10+
# Nonnegative rationals
11+
12+
This file defines the nonnegative rationals as a subtype of `rat` and provides its algebraic order
13+
structure.
14+
15+
We also define an instance `can_lift ℚ ℚ≥0`. This instance can be used by the `lift` tactic to
16+
replace `x : ℚ` and `hx : 0 ≤ x` in the proof context with `x : ℚ≥0` while replacing all occurences
17+
of `x` with `↑x`. This tactic also works for a function `f : α → ℚ` with a hypothesis
18+
`hf : ∀ x, 0 ≤ f x`.
19+
20+
## Notation
21+
22+
`ℚ≥0` is notation for `nnrat` in locale `nnrat`.
23+
-/
24+
25+
open function
26+
open_locale big_operators
27+
28+
/-- Nonnegative rational numbers. -/
29+
@[derive [canonically_ordered_comm_semiring, canonically_linear_ordered_semifield,
30+
linear_ordered_comm_group_with_zero, has_sub, has_ordered_sub,
31+
densely_ordered, archimedean, inhabited]]
32+
def nnrat := {q : ℚ // 0 ≤ q}
33+
34+
localized "notation ` ℚ≥0 ` := nnrat" in nnrat
35+
36+
namespace nnrat
37+
variables {α : Type*} {p q : ℚ≥0}
38+
39+
instance : has_coe ℚ≥0 ℚ := ⟨subtype.val⟩
40+
41+
/- Simp lemma to put back `n.val` into the normal form given by the coercion. -/
42+
@[simp] lemma val_eq_coe (q : ℚ≥0) : q.val = q := rfl
43+
44+
instance : can_lift ℚ ℚ≥0 :=
45+
{ coe := coe,
46+
cond := λ q, 0 ≤ q,
47+
prf := λ q hq, ⟨⟨q, hq⟩, rfl⟩ }
48+
49+
@[ext] lemma ext : (p : ℚ) = (q : ℚ) → p = q := subtype.ext
50+
51+
protected lemma coe_injective : injective (coe : ℚ≥0 → ℚ) := subtype.coe_injective
52+
53+
@[simp, norm_cast] lemma coe_inj : (p : ℚ) = q ↔ p = q := subtype.coe_inj
54+
55+
lemma ext_iff : p = q ↔ (p : ℚ) = q := subtype.ext_iff
56+
57+
lemma ne_iff {x y : ℚ≥0} : (x : ℚ) ≠ (y : ℚ) ↔ x ≠ y := nnrat.coe_inj.not
58+
59+
@[norm_cast] lemma coe_mk (q : ℚ) (hq) : ((⟨q, hq⟩ : ℚ≥0) : ℚ) = q := rfl
60+
61+
/-- Reinterpret a rational number `q` as a non-negative rational number. Returns `0` if `q ≤ 0`. -/
62+
def _root_.rat.to_nnrat (q : ℚ) : ℚ≥0 := ⟨max q 0, le_max_right _ _⟩
63+
64+
lemma _root_.rat.coe_to_nnrat (q : ℚ) (hq : 0 ≤ q) : (q.to_nnrat : ℚ) = q := max_eq_left hq
65+
66+
lemma _root_.rat.le_coe_to_nnrat (q : ℚ) : q ≤ q.to_nnrat := le_max_left _ _
67+
68+
open _root_.rat (to_nnrat)
69+
70+
@[simp] lemma coe_nonneg (q : ℚ≥0) : (0 : ℚ) ≤ q := q.2
71+
72+
@[simp, norm_cast] lemma coe_zero : ((0 : ℚ≥0) : ℚ) = 0 := rfl
73+
@[simp, norm_cast] lemma coe_one : ((1 : ℚ≥0) : ℚ) = 1 := rfl
74+
@[simp, norm_cast] lemma coe_add (p q : ℚ≥0) : ((p + q : ℚ≥0) : ℚ) = p + q := rfl
75+
@[simp, norm_cast] lemma coe_mul (p q : ℚ≥0) : ((p * q : ℚ≥0) : ℚ) = p * q := rfl
76+
@[simp, norm_cast] lemma coe_inv (q : ℚ≥0) : ((q⁻¹ : ℚ≥0) : ℚ) = q⁻¹ := rfl
77+
@[simp, norm_cast] lemma coe_div (p q : ℚ≥0) : ((p / q : ℚ≥0) : ℚ) = p / q := rfl
78+
@[simp, norm_cast] lemma coe_bit0 (q : ℚ≥0) : ((bit0 q : ℚ≥0) : ℚ) = bit0 q := rfl
79+
@[simp, norm_cast] lemma coe_bit1 (q : ℚ≥0) : ((bit1 q : ℚ≥0) : ℚ) = bit1 q := rfl
80+
@[simp, norm_cast] lemma coe_sub (h : q ≤ p) : ((p - q : ℚ≥0) : ℚ) = p - q :=
81+
max_eq_left $ le_sub.2 $ by simp [show (q : ℚ) ≤ p, from h]
82+
83+
@[simp] lemma coe_eq_zero : (q : ℚ) = 0 ↔ q = 0 := by norm_cast
84+
lemma coe_ne_zero : (q : ℚ) ≠ 0 ↔ q ≠ 0 := coe_eq_zero.not
85+
86+
@[simp, norm_cast] lemma coe_le_coe : (p : ℚ) ≤ q ↔ p ≤ q := iff.rfl
87+
@[simp, norm_cast] lemma coe_lt_coe : (p : ℚ) < q ↔ p < q := iff.rfl
88+
@[simp, norm_cast] lemma coe_pos : (0 : ℚ) < q ↔ 0 < q := iff.rfl
89+
90+
lemma coe_mono : monotone (coe : ℚ≥0 → ℚ) := λ _ _, coe_le_coe.2
91+
92+
lemma to_nnrat_mono : monotone to_nnrat := λ x y h, max_le_max h le_rfl
93+
94+
@[simp] lemma to_nnrat_coe (q : ℚ≥0) : to_nnrat q = q := ext $ max_eq_left q.2
95+
96+
@[simp] lemma to_nnrat_coe_nat (n : ℕ) : to_nnrat n = n :=
97+
ext $ by simp [rat.coe_to_nnrat]
98+
99+
/-- `to_nnrat` and `coe : ℚ≥0 → ℚ` form a Galois insertion. -/
100+
protected def gi : galois_insertion to_nnrat coe :=
101+
galois_insertion.monotone_intro coe_mono to_nnrat_mono rat.le_coe_to_nnrat to_nnrat_coe
102+
103+
/-- Coercion `ℚ≥0 → ℚ` as a `ring_hom`. -/
104+
def coe_hom : ℚ≥0 →+* ℚ := ⟨coe, coe_one, coe_mul, coe_zero, coe_add⟩
105+
106+
@[simp, norm_cast] lemma coe_nat_cast (n : ℕ) : (↑(↑n : ℚ≥0) : ℚ) = n := map_nat_cast coe_hom n
107+
108+
@[simp] lemma mk_coe_nat (n : ℕ) : @eq ℚ≥0 (⟨(n : ℚ), n.cast_nonneg⟩ : ℚ≥0) n :=
109+
ext (coe_nat_cast n).symm
110+
111+
/-- The rational numbers are an algebra over the non-negative rationals. -/
112+
instance : algebra ℚ≥0 ℚ := coe_hom.to_algebra
113+
114+
/-- A `mul_action` over `ℚ` restricts to a `mul_action` over `ℚ≥0`. -/
115+
instance [mul_action ℚ α] : mul_action ℚ≥0 α := mul_action.comp_hom α coe_hom.to_monoid_hom
116+
117+
/-- A `distrib_mul_action` over `ℚ` restricts to a `distrib_mul_action` over `ℚ≥0`. -/
118+
instance [add_comm_monoid α] [distrib_mul_action ℚ α] : distrib_mul_action ℚ≥0 α :=
119+
distrib_mul_action.comp_hom α coe_hom.to_monoid_hom
120+
121+
/-- A `module` over `ℚ` restricts to a `module` over `ℚ≥0`. -/
122+
instance [add_comm_monoid α] [module ℚ α] : module ℚ≥0 α := module.comp_hom α coe_hom
123+
124+
@[simp] lemma coe_coe_hom : ⇑coe_hom = coe := rfl
125+
126+
@[simp, norm_cast] lemma coe_indicator (s : set α) (f : α → ℚ≥0) (a : α) :
127+
((s.indicator f a : ℚ≥0) : ℚ) = s.indicator (λ x, f x) a :=
128+
(coe_hom : ℚ≥0 →+ ℚ).map_indicator _ _ _
129+
130+
@[simp, norm_cast] lemma coe_pow (q : ℚ≥0) (n : ℕ) : (↑(q ^ n) : ℚ) = q ^ n := coe_hom.map_pow _ _
131+
132+
@[norm_cast] lemma coe_list_sum (l : list ℚ≥0) : (l.sum : ℚ) = (l.map coe).sum :=
133+
coe_hom.map_list_sum _
134+
135+
@[norm_cast] lemma coe_list_prod (l : list ℚ≥0) : (l.prod : ℚ) = (l.map coe).prod :=
136+
coe_hom.map_list_prod _
137+
138+
@[norm_cast] lemma coe_multiset_sum (s : multiset ℚ≥0) : (s.sum : ℚ) = (s.map coe).sum :=
139+
coe_hom.map_multiset_sum _
140+
141+
@[norm_cast] lemma coe_multiset_prod (s : multiset ℚ≥0) : (s.prod : ℚ) = (s.map coe).prod :=
142+
coe_hom.map_multiset_prod _
143+
144+
@[norm_cast] lemma coe_sum {s : finset α} {f : α → ℚ≥0} : ↑(∑ a in s, f a) = ∑ a in s, (f a : ℚ) :=
145+
coe_hom.map_sum _ _
146+
147+
lemma to_nnrat_sum_of_nonneg {s : finset α} {f : α → ℚ} (hf : ∀ a, a ∈ s → 0 ≤ f a) :
148+
(∑ a in s, f a).to_nnrat = ∑ a in s, (f a).to_nnrat :=
149+
begin
150+
rw [←coe_inj, coe_sum, rat.coe_to_nnrat _ (finset.sum_nonneg hf)],
151+
exact finset.sum_congr rfl (λ x hxs, by rw rat.coe_to_nnrat _ (hf x hxs)),
152+
end
153+
154+
@[norm_cast] lemma coe_prod {s : finset α} {f : α → ℚ≥0} : ↑(∏ a in s, f a) = ∏ a in s, (f a : ℚ) :=
155+
coe_hom.map_prod _ _
156+
157+
lemma to_nnrat_prod_of_nonneg {s : finset α} {f : α → ℚ} (hf : ∀ a ∈ s, 0 ≤ f a) :
158+
(∏ a in s, f a).to_nnrat = ∏ a in s, (f a).to_nnrat :=
159+
begin
160+
rw [←coe_inj, coe_prod, rat.coe_to_nnrat _ (finset.prod_nonneg hf)],
161+
exact finset.prod_congr rfl (λ x hxs, by rw rat.coe_to_nnrat _ (hf x hxs)),
162+
end
163+
164+
@[norm_cast] lemma nsmul_coe (q : ℚ≥0) (n : ℕ) : ↑(n • q) = n • (q : ℚ) :=
165+
coe_hom.to_add_monoid_hom.map_nsmul _ _
166+
167+
lemma bdd_above_coe {s : set ℚ≥0} : bdd_above (coe '' s : set ℚ) ↔ bdd_above s :=
168+
⟨λ ⟨b, hb⟩, ⟨to_nnrat b, λ ⟨y, hy⟩ hys, show y ≤ max b 0, from
169+
(hb $ set.mem_image_of_mem _ hys).trans $ le_max_left _ _⟩,
170+
λ ⟨b, hb⟩, ⟨b, λ y ⟨x, hx, eq⟩, eq ▸ hb hx⟩⟩
171+
172+
lemma bdd_below_coe (s : set ℚ≥0) : bdd_below ((coe : ℚ≥0 → ℚ) '' s) := ⟨0, λ r ⟨q, _, h⟩, h ▸ q.2
173+
174+
@[simp, norm_cast] lemma coe_max (x y : ℚ≥0) : ((max x y : ℚ≥0) : ℚ) = max (x : ℚ) (y : ℚ) :=
175+
coe_mono.map_max
176+
177+
@[simp, norm_cast] lemma coe_min (x y : ℚ≥0) : ((min x y : ℚ≥0) : ℚ) = min (x : ℚ) (y : ℚ) :=
178+
coe_mono.map_min
179+
180+
lemma sub_def (p q : ℚ≥0) : p - q = to_nnrat (p - q) := rfl
181+
182+
@[simp] lemma abs_coe (q : ℚ≥0) : |(q : ℚ)| = q := abs_of_nonneg q.2
183+
184+
end nnrat
185+
186+
open nnrat
187+
188+
namespace rat
189+
variables {p q : ℚ}
190+
191+
@[simp] lemma to_nnrat_zero : to_nnrat 0 = 0 := by simp [to_nnrat]; refl
192+
@[simp] lemma to_nnrat_one : to_nnrat 1 = 1 := by simp [to_nnrat, max_eq_left zero_le_one]
193+
194+
@[simp] lemma to_nnrat_pos : 0 < to_nnrat q ↔ 0 < q := by simp [to_nnrat, ←coe_lt_coe]
195+
196+
@[simp] lemma to_nnrat_eq_zero : to_nnrat q = 0 ↔ q ≤ 0 :=
197+
by simpa [-to_nnrat_pos] using (@to_nnrat_pos q).not
198+
199+
alias to_nnrat_eq_zero ↔ _ to_nnrat_of_nonpos
200+
201+
@[simp] lemma to_nnrat_le_to_nnrat_iff (hp : 0 ≤ p) : to_nnrat q ≤ to_nnrat p ↔ q ≤ p :=
202+
by simp [←coe_le_coe, to_nnrat, hp]
203+
204+
@[simp] lemma to_nnrat_lt_to_nnrat_iff' : to_nnrat q < to_nnrat p ↔ q < p ∧ 0 < p :=
205+
by { simp [←coe_lt_coe, to_nnrat, lt_irrefl], exact lt_trans' }
206+
207+
lemma to_nnrat_lt_to_nnrat_iff (h : 0 < p) : to_nnrat q < to_nnrat p ↔ q < p :=
208+
to_nnrat_lt_to_nnrat_iff'.trans (and_iff_left h)
209+
210+
lemma to_nnrat_lt_to_nnrat_iff_of_nonneg (hq : 0 ≤ q) : to_nnrat q < to_nnrat p ↔ q < p :=
211+
to_nnrat_lt_to_nnrat_iff'.trans ⟨and.left, λ h, ⟨h, hq.trans_lt h⟩⟩
212+
213+
@[simp] lemma to_nnrat_add (hq : 0 ≤ q) (hp : 0 ≤ p) : to_nnrat (q + p) = to_nnrat q + to_nnrat p :=
214+
nnrat.ext $ by simp [to_nnrat, hq, hp, add_nonneg]
215+
216+
lemma to_nnrat_add_le : to_nnrat (q + p) ≤ to_nnrat q + to_nnrat p :=
217+
coe_le_coe.1 $ max_le (add_le_add (le_max_left _ _) (le_max_left _ _)) $ coe_nonneg _
218+
219+
lemma to_nnrat_le_iff_le_coe {p : ℚ≥0} : to_nnrat q ≤ p ↔ q ≤ ↑p := nnrat.gi.gc q p
220+
221+
lemma le_to_nnrat_iff_coe_le {q : ℚ≥0} (hp : 0 ≤ p) : q ≤ to_nnrat p ↔ ↑q ≤ p :=
222+
by rw [←coe_le_coe, rat.coe_to_nnrat p hp]
223+
224+
lemma le_to_nnrat_iff_coe_le' {q : ℚ≥0} (hq : 0 < q) : q ≤ to_nnrat p ↔ ↑q ≤ p :=
225+
(le_or_lt 0 p).elim le_to_nnrat_iff_coe_le $ λ hp,
226+
by simp only [(hp.trans_le q.coe_nonneg).not_le, to_nnrat_eq_zero.2 hp.le, hq.not_le]
227+
228+
lemma to_nnrat_lt_iff_lt_coe {p : ℚ≥0} (hq : 0 ≤ q) : to_nnrat q < p ↔ q < ↑p :=
229+
by rw [←coe_lt_coe, rat.coe_to_nnrat q hq]
230+
231+
lemma lt_to_nnrat_iff_coe_lt {q : ℚ≥0} : q < to_nnrat p ↔ ↑q < p := nnrat.gi.gc.lt_iff_lt
232+
233+
@[simp] lemma to_nnrat_bit0 (hq : 0 ≤ q) : to_nnrat (bit0 q) = bit0 (to_nnrat q) :=
234+
to_nnrat_add hq hq
235+
236+
@[simp] lemma to_nnrat_bit1 (hq : 0 ≤ q) : to_nnrat (bit1 q) = bit1 (to_nnrat q) :=
237+
(to_nnrat_add (by simp [hq]) zero_le_one).trans $ by simp [to_nnrat_one, bit1, hq]
238+
239+
lemma to_nnrat_mul (hp : 0 ≤ p) : to_nnrat (p * q) = to_nnrat p * to_nnrat q :=
240+
begin
241+
cases le_total 0 q with hq hq,
242+
{ ext; simp [to_nnrat, hp, hq, max_eq_left, mul_nonneg] },
243+
{ have hpq := mul_nonpos_of_nonneg_of_nonpos hp hq,
244+
rw [to_nnrat_eq_zero.2 hq, to_nnrat_eq_zero.2 hpq, mul_zero] }
245+
end
246+
247+
lemma to_nnrat_inv (q : ℚ) : to_nnrat q⁻¹ = (to_nnrat q)⁻¹ :=
248+
begin
249+
obtain hq | hq := le_total q 0,
250+
{ rw [to_nnrat_eq_zero.mpr hq, inv_zero, to_nnrat_eq_zero.mpr (inv_nonpos.mpr hq)] },
251+
{ nth_rewrite 0 ←rat.coe_to_nnrat q hq,
252+
rw [←coe_inv, to_nnrat_coe] }
253+
end
254+
255+
lemma to_nnrat_div (hp : 0 ≤ p) : to_nnrat (p / q) = to_nnrat p / to_nnrat q :=
256+
by rw [div_eq_mul_inv, div_eq_mul_inv, ←to_nnrat_inv, ←to_nnrat_mul hp]
257+
258+
lemma to_nnrat_div' (hq : 0 ≤ q) : to_nnrat (p / q) = to_nnrat p / to_nnrat q :=
259+
by rw [div_eq_inv_mul, div_eq_inv_mul, to_nnrat_mul (inv_nonneg.2 hq), to_nnrat_inv]
260+
261+
end rat
262+
263+
/-- The absolute value on `ℚ` as a map to `ℚ≥0`. -/
264+
@[pp_nodot] def rat.nnabs (x : ℚ) : ℚ≥0 := ⟨abs x, abs_nonneg x⟩
265+
266+
@[norm_cast, simp] lemma rat.coe_nnabs (x : ℚ) : (rat.nnabs x : ℚ) = abs x := by simp [rat.nnabs]
267+
268+
/-! ### Numerator and denominator -/
269+
270+
namespace nnrat
271+
variables {p q : ℚ≥0}
272+
273+
/-- The numerator of a nonnegative rational. -/
274+
def num (q : ℚ≥0) : ℕ := (q : ℚ).num.nat_abs
275+
276+
/-- The denominator of a nonnegative rational. -/
277+
def denom (q : ℚ≥0) : ℕ := (q : ℚ).denom
278+
279+
@[simp] lemma nat_abs_num_coe : (q : ℚ).num.nat_abs = q.num := rfl
280+
@[simp] lemma denom_coe : (q : ℚ).denom = q.denom := rfl
281+
282+
lemma ext_num_denom (hn : p.num = q.num) (hd : p.denom = q.denom) : p = q :=
283+
ext $ rat.ext ((int.nat_abs_inj_of_nonneg_of_nonneg
284+
(rat.num_nonneg_iff_zero_le.2 p.2) $ rat.num_nonneg_iff_zero_le.2 q.2).1 hn) hd
285+
286+
lemma ext_num_denom_iff : p = q ↔ p.num = q.num ∧ p.denom = q.denom :=
287+
by { rintro rfl, exact ⟨rfl, rfl⟩ }, λ h, ext_num_denom h.1 h.2
288+
289+
@[simp] lemma num_div_denom (q : ℚ≥0) : (q.num : ℚ≥0) / q.denom = q :=
290+
begin
291+
ext1,
292+
rw [coe_div, coe_nat_cast, coe_nat_cast, num, ←int.cast_coe_nat,
293+
int.nat_abs_of_nonneg (rat.num_nonneg_iff_zero_le.2 q.prop)],
294+
exact rat.num_div_denom q,
295+
end
296+
297+
/-- A recursor for nonnegative rationals in terms of numerators and denominators. -/
298+
protected def rec {α : ℚ≥0Sort*} (h : Π m n : ℕ, α (m / n)) (q : ℚ≥0) : α q :=
299+
(num_div_denom _).rec (h _ _)
300+
301+
end nnrat

0 commit comments

Comments
 (0)