module Data.Rational.Order where
open Explicit ℤ-comm
infix 7 _<_ _≤ᶠ_
Ordering on rationals🔗
The field of rational numbers inherits a partial order from the ring of integers defining the relation to be As usual, we define this first at the level of fractions, then prove that it extends to the quotient
_≤ᶠ_ : Fraction → Fraction → Type
(x / s [ _ ]) ≤ᶠ (y / t [ _ ]) = (x *ℤ t) ℤ.≤ (y *ℤ s)
The easiest way to show this is transitivity at the level of fractions, since it will follow directly from the definition that implies We can then prove that is invariant under on either side, by appealing to transitivity. The proof is not complicated, just annoying:
: ∀ {x y} → x ≈ y → x ≤ᶠ y
≤ᶠ-refl' {x@record{}} {y@record{}} p = ℤ.≤-refl' (from-same-rational p)
≤ᶠ-refl'
: ∀ x y z → x ≤ᶠ y → y ≤ᶠ z → x ≤ᶠ z
≤ᶠ-trans (x / s [ pos s' ]) (y / t [ pos t' ]) (z / u [ pos u' ]) α β =
≤ᶠ-trans (ℤ.*ℤ-cancel-≤r {t} $
3 (λ x u t → x :* u :* t ≔ x :* t :* u) x u t refl ⟩
x *ℤ u *ℤ t =⟨ solve .*ℤ-preserves-≤r u α ⟩
x *ℤ t *ℤ u ≤⟨ ℤ3 (λ y s u → y :* s :* u ≔ y :* u :* s) y s u refl ⟩
y *ℤ s *ℤ u =⟨ solve .*ℤ-preserves-≤r s β ⟩
y *ℤ u *ℤ s ≤⟨ ℤ3 (λ z t s → z :* t :* s ≔ z :* s :* t) z t s refl ⟩
z *ℤ t *ℤ s =⟨ solve ) z *ℤ s *ℤ t ≤∎
We can then lift this to a family of propositions over
The strategy for showing it respects the quotient is as outlined above,
so we’ll leave this in a <details>
block.
private
: ℚ → ℚ → Prop lzero
leℚ (inc x) (inc y) =
leℚ (hlevel 2) work
Coeq-rec₂ (λ a (x , y , r) → r2 x y a r)
(λ a (x , y , r) → r1 a x y r) x y
where
: Fraction → Fraction → Prop lzero
work = x ≤ᶠ y
∣ work x y ∣ record{} record{} .is-tr = hlevel 1
work
: ∀ u x y → x ≈ y → work u x ≡ work u y
r1 @record{} x@record{} y@record{} p' = n-ua (prop-ext!
r1 u(λ α → ≤ᶠ-trans u x y α (≤ᶠ-refl' p'))
(λ α → ≤ᶠ-trans u y x α (≤ᶠ-refl' (≈.symᶜ p'))))
: ∀ x y u → x ≈ y → work x u ≡ work y u
r2 @record{} x@record{} y@record{} p' = n-ua (prop-ext!
r2 u(λ α → ≤ᶠ-trans x u y (≤ᶠ-refl' (≈.symᶜ p')) α)
(λ α → ≤ᶠ-trans u x y (≤ᶠ-refl' p') α))
We define the type family _≤_
as a record type, wrapping orderℚ
,
to help out type inference: orderℚ
is a pretty nasty definition, whereas an application of a record name is
always a normal form.
record _≤_ (x y : ℚ) : Type where
constructor inc
field
: ⌞ leℚ x y ⌟ lower
open _≤_
unquoteDecl H-Level-≤ = declare-record-hlevel 1 H-Level-≤ (quote _≤_)
By lifting our proofs about _≤ᶠ_
through this record type, we can prove that the ordering on the rational
numbers is decidable, reflexive, transitive, and antisymmetric.
instance
: ∀ {x y} → Dec (x ≤ y)
Dec-≤ {inc x} {inc y} = elim! {P = λ x → ∀ y → Dec (ℚ.inc x ≤ inc y)} go x y where
Dec-≤ : ∀ x y → Dec (toℚ x ≤ toℚ y)
go @record{} y@record{} with holds? ((↑ x *ℤ ↓ y) ℤ.≤ (↑ y *ℤ ↓ x))
go x... | yes p = yes (inc p)
... | no ¬p = no (¬p ∘ lower)
: ∀ {x y} ⦃ p : (↑ x *ℤ ↓ y) ℤ.≤ (↑ y *ℤ ↓ x) ⦄ → toℚ x ≤ toℚ y
≤-from {record{}} {record{}} ⦃ p ⦄ = inc p
≤-from
abstract
: ∀ {x y} → (↑ x *ℤ ↓ y) ℤ.≤ (↑ y *ℤ ↓ x) → toℚ x ≤ toℚ y
toℚ≤ {record{}} {record{}} p = inc p
toℚ≤
: ∀ {x} → x ≤ x
≤-refl : ∀ {x y z} → x ≤ y → y ≤ z → x ≤ z
≤-trans : ∀ {x y} → x ≤ y → y ≤ x → x ≡ y
≤-antisym
unquoteDef ≤-refl ≤-trans ≤-antisym = do
λ d x → inc ℤ.≤-refl
by-elim-ℚ ≤-refl
λ d x y z (inc p) (inc q) → inc
by-elim-ℚ ≤-trans (≤ᶠ-trans (x / d [ auto ]) (y / d [ auto ]) (z / d [ auto ]) p q)
λ d x y (inc α) (inc β) →
by-elim-ℚ ≤-antisym (to-same-rational (ℤ.≤-antisym α β))
quotℚ
: ∀ {x y} (p : x ≡ y) → x ≤ y
≤-refl' {x} p = transport (λ i → x ≤ p i) ≤-refl ≤-refl'
Algebraic properties🔗
We can also show that the ordering on behaves well with respect to its algebraic structure. For example, the addition function is monotone in both arguments, and division by a positive integer preserves the order.
: ∀ {x y d} (p : ℤ.Positive d) → x ℤ.≤ y → toℚ (x / d [ p ]) ≤ toℚ (y / d [ p ])
common-denominator-≤ {d = d} (pos _) p = inc (ℤ.*ℤ-preserves-≤r d p)
common-denominator-≤
: ∀ {x y a b} → x ≤ y → a ≤ b → (x +ℚ a) ≤ (y +ℚ b)
+ℚ-preserves-≤ : ∀ {x y} → 0 ≤ x → 0 ≤ y → 0 ≤ (x *ℚ y)
*ℚ-nonnegative : ∀ {x} ⦃ p : Nonzero x ⦄ → 0 ≤ x → 0 ≤ (invℚ x ⦃ p ⦄)
invℚ-nonnegative
unquoteDef +ℚ-preserves-≤ *ℚ-nonnegative invℚ-nonnegative = do
λ d x y a b (inc p) (inc q) →
by-elim-ℚ +ℚ-preserves-≤ (ℤ.*ℤ-positive auto auto) $
common-denominator-≤ .+ℤ-preserves-≤r (a *ℤ d) _ _ p ⟩
x *ℤ d +ℤ a *ℤ d ≤⟨ ℤ.+ℤ-preserves-≤l (y *ℤ d) _ _ q ⟩
y *ℤ d +ℤ a *ℤ d ≤⟨ ℤ
y *ℤ d +ℤ b *ℤ d ≤∎
λ d x y (inc p) (inc q) → inc
by-elim-ℚ *ℚ-nonnegative (ℤ.≤-trans
(ℤ.*ℤ-nonnegative
(ℤ.≤-trans p (ℤ.≤-refl' (ℤ.*ℤ-oner x)))
(ℤ.≤-trans q (ℤ.≤-refl' (ℤ.*ℤ-oner y))))
(ℤ.≤-refl' (sym (ℤ.*ℤ-oner (x *ℤ y)))))
λ where
by-elim-ℚ invℚ-nonnegative (possuc x) posz ⦃ inc nz ⦄ z → absurd (nz (quotℚ (to-same-rational refl)))
(possuc x) (possuc y) z → inc (ℤ.pos≤pos 0≤x)
: ∀ {x y} ⦃ p : Nonzero y ⦄ → 0 ≤ x → 0 ≤ y → 0 ≤ (x /ℚ y)
/ℚ-nonnegative {inc x} {inc y} a b = *ℚ-nonnegative a (invℚ-nonnegative {inc y} b) /ℚ-nonnegative
Positivity🔗
We can also lift the notion of positivity from the integers to the rational numbers. A fraction is positive if its numerator is positive.
private
: ℚ → Prop lzero
positiveℚ (inc x) = Coeq-rec (λ f → el! (ℤ.Positive (↑ f))) (λ (x , y , p) → r x y p) x where
positiveℚ : ∀ x y → x ≈ y → el! (ℤ.Positive (↑ x)) ≡ el! (ℤ.Positive (↑ y))
r (x / s [ ps ]) (y / t [ pt ]) r with r ← from-same-rational r = n-ua (prop-ext!
r (λ px → ℤ.*ℤ-positive-cancel ps (subst ℤ.Positive (r ∙ ℤ.*ℤ-commutative y s) (ℤ.*ℤ-positive px pt)))
λ py → ℤ.*ℤ-positive-cancel pt (subst ℤ.Positive (sym r ∙ ℤ.*ℤ-commutative x t) (ℤ.*ℤ-positive py ps)))
record Positive (q : ℚ) : Type where
constructor inc
field
: ⌞ positiveℚ q ⌟ lower
open Positive
unquoteDecl H-Level-Positive = declare-record-hlevel 1 H-Level-Positive (quote Positive)
instance
: ∀ {x} → Dec (Positive x)
Dec-Positive {x} with (r@(n / d [ p ]) , q) ← splitℚ x | holds? (ℤ.Positive n)
Dec-Positive ... | yes p = yes (subst Positive q (inc (recover p)))
... | no ¬p = no λ x → absurd (case subst Positive (sym q) x of λ (inc p) → ¬p p)
: ∀ {x s p} → Positive (toℚ (possuc x / s [ p ]))
Positive-pos = inc (pos _) Positive-pos
This has the expected interaction with the ordering and the algebraic operations.
: ∀ {x} → 0 ≤ x → x ≠ 0 → Positive x
nonnegative-nonzero→positive : ∀ {x} → Positive x → x ≠ 0
positive→nonzero : ∀ {x} → Positive x → 0 ≤ x
positive→nonnegative
unquoteDef nonnegative-nonzero→positive positive→nonzero positive→nonnegative = do
λ where
by-elim-ℚ nonnegative-nonzero→positive (inc q) r → absurd (r (ext refl))
d posz (possuc x) (inc q) r → inc (pos x)
d
λ where
by-elim-ℚ positive→nonzero (inc α) β → ℤ.positive→nonzero α (sym (ℤ.*ℤ-oner a) ∙ from-same-rational (unquotℚ β))
d a
λ where
by-elim-ℚ positive→nonnegative (inc α) → inc (ℤ.≤-trans (ℤ.positive→nonnegative α) (ℤ.≤-refl' (sym (ℤ.*ℤ-oner a)))) d a
The strict order on rationals🔗
With a bit more effort, we can also lift the strict ordering on the integers to the rationals. The definitions end up being pretty much the same as for lifting the partial order! We start by defining it on fractions, then showing that it respects the quotient on both sides — we can’t use transitivity and reflexivity like we did above, unfortunately.
_<ᶠ_ : Fraction → Fraction → Type
(x / s [ _ ]) <ᶠ (y / t [ _ ]) = (x *ℤ t) ℤ.< (y *ℤ s)
: ∀ {x y f} → x ≈ y → f <ᶠ y → f <ᶠ x
<ᶠ-respr {x / s [ p ]} {y / t [ q ]} {z / u [ r ]} xt~ys α = ℤ.*ℤ-cancel-<r {t} ⦃ q ⦄ $
<ᶠ-respr .≤-<-trans
ℤ(ℤ.≤-refl' (solve 3 (λ z s t → z :* s :* t ≔ (z :* t) :* s) z s t refl))
(ℤ.<-≤-trans (ℤ.*ℤ-preserves-<r s ⦃ p ⦄ α) (ℤ.≤-refl'
( solve 3 (λ y u s → y :* u :* s ≔ y :* s :* u) y u s refl
(_*ℤ u) (sym (from-same-rational xt~ys))
∙ ap 3 (λ x t u → x :* t :* u ≔ x :* u :* t) x t u refl)))
∙ solve
: ∀ {x y f} → x ≈ y → x <ᶠ f → y <ᶠ f
<ᶠ-respl {x / s [ p ]} {y / t [ q ]} {z / u [ r ]} xt~ys α = ℤ.*ℤ-cancel-<l {s} ⦃ p ⦄ $
<ᶠ-respl .≤-<-trans
ℤ(ℤ.≤-refl'
( solve 3 (λ s y u → s :* (y :* u) ≔ (y :* s) :* u) s y u refl
(_*ℤ u) (sym (from-same-rational xt~ys))
∙ ap 3 (λ x t u → x :* t :* u ≔ x :* u :* t) x t u refl))
∙ solve (ℤ.<-≤-trans (ℤ.*ℤ-preserves-<r t ⦃ q ⦄ α)
(ℤ.≤-refl' (solve 3 (λ z s t → (z :* s) :* t ≔ s :* (z :* t)) z s t refl)))
Having gotten the bulk of the construction out of the way, we can lift it to a type family over pairs of rationals, and do the same dance as to make this into something definitionally injective.
private
: ℚ → ℚ → Prop lzero
ltℚ (inc x) (inc y) = Coeq-rec₂ (hlevel 2) work (λ u (x , y , r) → r2 x y u r) (λ u (x , y , r) → r1 u x y r) x y where
ltℚ : Fraction → Fraction → Prop lzero
work = x <ᶠ y
∣ work x y ∣ record{} record{} .is-tr = hlevel 1
work
: ∀ u x y → x ≈ y → work u x ≡ work u y
r1 @record{} x@record{} y@record{} p' = n-ua (prop-ext!
r1 u(λ α → <ᶠ-respr {y} {x} {u} (≈.symᶜ p') α)
(λ α → <ᶠ-respr {x} {y} {u} p' α))
: ∀ x y u → x ≈ y → work x u ≡ work y u
r2 @record{} y@record{} u@record{} p' = n-ua (prop-ext!
r2 x(λ α → <ᶠ-respl {x} {y} {u} p' α)
(λ α → <ᶠ-respl {y} {x} {u} (≈.symᶜ p') α))
record _<_ (x y : ℚ) : Type where
constructor inc
field
: ⌞ ltℚ x y ⌟
lower
open _<_
unquoteDecl H-Level-< = declare-record-hlevel 1 H-Level-< (quote _<_)
As before, everything we want to show about strict inequality on the rationals follows by lifting the analogous facts from the integers, where we’re free to assume any set of related rationals has the same denominator.
instance
: ∀ {x y} → Dec (x < y)
Dec-< {inc x} {inc y} = elim! {P = λ x → ∀ y → Dec (inc x < inc y)} go x y where
Dec-< : ∀ x y → Dec (toℚ x < toℚ y)
go @record{} y@record{} with holds? ((↑ x *ℤ ↓ y) ℤ.< (↑ y *ℤ ↓ x))
go x... | yes p = yes (inc p)
... | no ¬p = no (¬p ∘ lower)
abstract
: ∀ {x y} → (↑ x *ℤ ↓ y) ℤ.< (↑ y *ℤ ↓ x) → toℚ x < toℚ y
toℚ< {record{}} {record{}} p = inc p
toℚ<
: ∀ {x y d} ⦃ _ : ℤ.Positive d ⦄ → x ℤ.< y → (x / d) < (y / d)
common-denom-< {x} {y} {d} p = inc (ℤ.*ℤ-preserves-<r {x} {y} d p)
common-denom-<
: ∀ {x y d} ⦃ _ : ℤ.Positive d ⦄ → (x / d) < (y / d) → x ℤ.< y
<-common-denom {x} {y} {d} (inc p) = ℤ.*ℤ-cancel-<r {d} {x} {y} p
<-common-denom
: ∀ {x y} → x ≡ y → ¬ x < y
<-irrefl : ∀ {x y z} → x < y → y < z → x < z
<-trans : ∀ {x y} → x < y → x ≤ y
<-weaken : ∀ {x y} → x ≤ y → (x ≡ y) ⊎ (x < y)
≤-strengthen
private instance
: ∀ {x y} → H-Level ((x ≡ y) ⊎ (x < y)) 1
hl-str = prop-instance (disjoint-⊎-is-prop (hlevel 1) (hlevel 1) λ (x , y) → <-irrefl x y)
hl-str {-# OVERLAPPING hl-str #-}
unquoteDef <-irrefl <-trans <-weaken ≤-strengthen = do
λ d x y p a →
by-elim-ℚ <-irrefl .<-irrefl (from-same-denom p) (<-common-denom a)
ℤ
λ d x y z p q →
by-elim-ℚ <-trans (ℤ.<-trans (<-common-denom p) (<-common-denom q))
common-denom-<
λ d x y (inc p) → inc (ℤ.<-weaken p)
by-elim-ℚ <-weaken
λ d x y (inc p) → case ℤ.≤-strengthen p of λ where
by-elim-ℚ ≤-strengthen (inl dx=dy) → inl (quotℚ (to-same-rational dx=dy))
(inr dx<dy) → inr (inc dx<dy)
: ∀ {x y z} → x < y → y ≤ z → x < z
<-≤-trans with ≤-strengthen q
<-≤-trans p q ... | inl y=z = subst₂ _<_ refl y=z p
... | inr y<z = <-trans p y<z
: ∀ {x y z} → x ≤ y → y < z → x < z
≤-<-trans with ≤-strengthen p
≤-<-trans p q ... | inl x=y = subst₂ _<_ (sym x=y) refl q
... | inr x<z = <-trans x<z q
: ∀ {x y} → x ≤ y → x ≠ y → x < y
<-from-≤ with ≤-strengthen x≤y
<-from-≤ x≤y x≠y ... | inl x=y = absurd (x≠y x=y)
... | inr x<y = x<y
: ℚ → ℚ
absℚ (inc x) = ℚ.inc $ ℚ-rec absᶠ absᶠ-resp (inc x) where
absℚ : Fraction → _
absᶠ (x / s [ p ]) = inc (pos (abs x) / s [ p ])
absᶠ
: ∀ x y → x ≈ y → absᶠ x ≡ absᶠ y
absᶠ-resp (pos x / possuc s [ p ]) (pos y / possuc t [ q ]) α = quot α
absᶠ-resp (pos x / possuc s [ p ]) (negsuc y / possuc t [ q ]) α = absurd (pos≠negsuc (sym (ℤ.assign-pos (x * suc t)) ∙ from-same-rational α))
absᶠ-resp (negsuc x / possuc s [ p ]) (pos y / possuc t [ q ]) α = absurd (negsuc≠pos (from-same-rational α ∙ ℤ.assign-pos (y * suc s)))
absᶠ-resp (negsuc x / possuc s [ p ]) (negsuc y / possuc t [ q ]) α = quot (to-same-rational (ap possuc (negsuc-injective (from-same-rational α)))) absᶠ-resp