module Data.Int.Order where
Ordering integers🔗
The usual partial order on the integers relies on the observation that the number line looks like two copies of the natural numbers, smashed end-to-end at the number zero. This is precisely the definition of the order we use:
data _≤_ : Int → Int → Type where
: ∀ {x y} → y Nat.≤ x → negsuc x ≤ negsuc y
neg≤neg : ∀ {x y} → x Nat.≤ y → pos x ≤ pos y
pos≤pos : ∀ {x y} → negsuc x ≤ pos y neg≤pos
Note the key properties: the ordering between negative numbers is reversed, and every negative number is smaller than every positive number. This means that decomposes, as an order, into an ordinal sum .
Basic properties🔗
Proving that this is actually a partial order is a straightforward combination of case-bashing and appealing to the analogous properties for the ordering on natural numbers.
: ∀ {x y} → pos x ≤ negsuc y → ⊥
¬pos≤neg ()
¬pos≤neg
: ∀ {x y} → is-prop (x ≤ y)
≤-is-prop (neg≤neg p) (neg≤neg q) = ap neg≤neg (Nat.≤-is-prop p q)
≤-is-prop (pos≤pos p) (pos≤pos q) = ap pos≤pos (Nat.≤-is-prop p q)
≤-is-prop = refl
≤-is-prop neg≤pos neg≤pos
: ∀ {x} → x ≤ x
≤-refl {x = pos x} = pos≤pos Nat.≤-refl
≤-refl {x = negsuc x} = neg≤neg Nat.≤-refl
≤-refl
: ∀ {x y z} → x ≤ y → y ≤ z → x ≤ z
≤-trans (neg≤neg p) (neg≤neg q) = neg≤neg (Nat.≤-trans q p)
≤-trans (neg≤neg p) neg≤pos = neg≤pos
≤-trans (pos≤pos p) (pos≤pos q) = pos≤pos (Nat.≤-trans p q)
≤-trans (pos≤pos x) = neg≤pos
≤-trans neg≤pos
: ∀ {x y} → x ≤ y → y ≤ x → x ≡ y
≤-antisym (neg≤neg p) (neg≤neg q) = ap negsuc (Nat.≤-antisym q p)
≤-antisym (pos≤pos p) (pos≤pos q) = ap pos (Nat.≤-antisym p q) ≤-antisym
Totality🔗
The ordering on the integers is decidable, and moreover it is a total order. We show weak totality: if , then .
: ∀ x y → ¬ (x ≤ y) → y ≤ x
≤-is-weakly-total (pos x) (pos y) p = pos≤pos $
≤-is-weakly-total .≤-is-weakly-total x y (p ∘ pos≤pos)
Nat(pos x) (negsuc y) p = neg≤pos
≤-is-weakly-total (negsuc x) (pos y) p = absurd (p neg≤pos)
≤-is-weakly-total (negsuc x) (negsuc y) p = neg≤neg $
≤-is-weakly-total .≤-is-weakly-total y x (p ∘ neg≤neg)
Nat
instance
: ∀ {x y} → Dec (x ≤ y)
Dec-≤ {pos x} {pos y} with holds? (x Nat.≤ y)
Dec-≤ ... | yes p = yes (pos≤pos p)
... | no ¬p = no λ { (pos≤pos p) → ¬p p }
{negsuc x} {negsuc y} with holds? (y Nat.≤ x)
Dec-≤ ... | yes p = yes (neg≤neg p)
... | no ¬p = no λ { (neg≤neg p) → ¬p p }
{pos x} {negsuc y} = no ¬pos≤neg
Dec-≤ {negsuc x} {pos y} = yes neg≤pos Dec-≤
Universal properties of maximum and minimum🔗
This case bash shows that maxℤ
and minℤ
satisfy their universal properties.
: (x y : Int) → x ≤ maxℤ x y
maxℤ-≤l (pos x) (pos y) = pos≤pos (Nat.max-≤l x y)
maxℤ-≤l (pos _) (negsuc _) = ≤-refl
maxℤ-≤l (negsuc _) (pos _) = neg≤pos
maxℤ-≤l (negsuc x) (negsuc y) = neg≤neg (Nat.min-≤l x y)
maxℤ-≤l
: (x y : Int) → y ≤ maxℤ x y
maxℤ-≤r (pos x) (pos y) = pos≤pos (Nat.max-≤r x y)
maxℤ-≤r (pos _) (negsuc _) = neg≤pos
maxℤ-≤r (negsuc _) (pos _) = ≤-refl
maxℤ-≤r (negsuc x) (negsuc y) = neg≤neg (Nat.min-≤r x y)
maxℤ-≤r
: (x y z : Int) → x ≤ z → y ≤ z → maxℤ x y ≤ z
maxℤ-univ _ _ _ (pos≤pos x≤z) (pos≤pos y≤z) = pos≤pos (Nat.max-univ _ _ _ x≤z y≤z)
maxℤ-univ _ _ _ (pos≤pos x≤z) neg≤pos = pos≤pos x≤z
maxℤ-univ _ _ _ neg≤pos (pos≤pos y≤z) = pos≤pos y≤z
maxℤ-univ _ _ _ neg≤pos neg≤pos = neg≤pos
maxℤ-univ _ _ _ (neg≤neg x≥z) (neg≤neg y≥z) = neg≤neg (Nat.min-univ _ _ _ x≥z y≥z)
maxℤ-univ
: (x y : Int) → minℤ x y ≤ x
minℤ-≤l (pos x) (pos y) = pos≤pos (Nat.min-≤l x y)
minℤ-≤l (pos _) (negsuc _) = neg≤pos
minℤ-≤l (negsuc _) (pos _) = ≤-refl
minℤ-≤l (negsuc x) (negsuc y) = neg≤neg (Nat.max-≤l x y)
minℤ-≤l
: (x y : Int) → minℤ x y ≤ y
minℤ-≤r (pos x) (pos y) = pos≤pos (Nat.min-≤r x y)
minℤ-≤r (pos _) (negsuc _) = ≤-refl
minℤ-≤r (negsuc _) (pos _) = neg≤pos
minℤ-≤r (negsuc x) (negsuc y) = neg≤neg (Nat.max-≤r x y)
minℤ-≤r
: (x y z : Int) → z ≤ x → z ≤ y → z ≤ minℤ x y
minℤ-univ _ _ _ (pos≤pos x≥z) (pos≤pos y≥z) = pos≤pos (Nat.min-univ _ _ _ x≥z y≥z)
minℤ-univ _ _ _ neg≤pos neg≤pos = neg≤pos
minℤ-univ _ _ _ neg≤pos (neg≤neg y≤z) = neg≤neg y≤z
minℤ-univ _ _ _ (neg≤neg x≤z) neg≤pos = neg≤neg x≤z
minℤ-univ _ _ _ (neg≤neg x≤z) (neg≤neg y≤z) = neg≤neg (Nat.max-univ _ _ _ x≤z y≤z) minℤ-univ
Compatibility with the structure🔗
The last case bash in this module will establish that the ordering on integers is compatible with the successor, predecessor, and negation. Since addition is equivalent to iterated application of the successor and predecessor, we get as a corollary that addition also respects the order.
: ∀ x y → x ≤ y → sucℤ x ≤ sucℤ y
suc-≤ (pos x) (pos y) (pos≤pos p) = pos≤pos (Nat.s≤s p)
suc-≤ (negsuc zero) (pos y) p = pos≤pos Nat.0≤x
suc-≤ (negsuc zero) (negsuc zero) p = ≤-refl
suc-≤ (negsuc zero) (negsuc (suc y)) (neg≤neg ())
suc-≤ (negsuc (suc x)) (pos y) p = neg≤pos
suc-≤ (negsuc (suc x)) (negsuc zero) p = neg≤pos
suc-≤ (negsuc (suc x)) (negsuc (suc y)) (neg≤neg (Nat.s≤s p)) = neg≤neg p
suc-≤
: ∀ x y → x ≤ y → predℤ x ≤ predℤ y
pred-≤ = ≤-refl
pred-≤ posz posz p (possuc y) p = neg≤pos
pred-≤ posz (possuc x) posz (pos≤pos ())
pred-≤ (possuc x) (possuc y) (pos≤pos (Nat.s≤s p)) = pos≤pos p
pred-≤ (negsuc x) posz p = neg≤neg Nat.0≤x
pred-≤ (negsuc x) (possuc y) p = neg≤pos
pred-≤ (negsuc x) (negsuc y) (neg≤neg p) = neg≤neg (Nat.s≤s p)
pred-≤
: ∀ k x y → x ≤ y → rotℤ k x ≤ rotℤ k y
rotℤ≤l = p
rotℤ≤l posz x y p (possuc k) x y p = suc-≤ _ _ (rotℤ≤l (pos k) x y p)
rotℤ≤l (negsuc zero) x y p = pred-≤ _ _ p
rotℤ≤l (negsuc (suc k)) x y p = pred-≤ _ _ (rotℤ≤l (negsuc k) x y p)
rotℤ≤l
abstract
: ∀ k x y → x ≤ y → (k +ℤ x) ≤ (k +ℤ y)
+ℤ-mono-l = transport
+ℤ-mono-l k x y p (sym (ap₂ _≤_ (rot-is-add k x) (rot-is-add k y)))
(rotℤ≤l k x y p)
: ∀ k x y → x ≤ y → (x +ℤ k) ≤ (y +ℤ k)
+ℤ-mono-r = transport
+ℤ-mono-r k x y p (ap₂ _≤_ (+ℤ-commutative k x) (+ℤ-commutative k y))
(+ℤ-mono-l k x y p)
: ∀ x y → x ≤ y → negℤ y ≤ negℤ x
negℤ-anti = x≤y
negℤ-anti posz posz x≤y (possuc y) _ = neg≤pos
negℤ-anti posz (possuc x) (possuc y) (pos≤pos (Nat.s≤s x≤y)) = neg≤neg x≤y
negℤ-anti (negsuc _) posz _ = pos≤pos Nat.0≤x
negℤ-anti (negsuc _) (possuc y) _ = neg≤pos
negℤ-anti (negsuc x) (negsuc y) (neg≤neg x≤y) = pos≤pos (Nat.s≤s x≤y)
negℤ-anti
: ∀ x y → negℤ y ≤ negℤ x → x ≤ y
negℤ-anti-full (pos y) _ = pos≤pos Nat.0≤x
negℤ-anti-full posz (negsuc y) (pos≤pos ())
negℤ-anti-full posz (possuc x) (possuc y) (neg≤neg x≤y) = pos≤pos (Nat.s≤s x≤y)
negℤ-anti-full (negsuc x) (pos y) _ = neg≤pos
negℤ-anti-full (negsuc x) (negsuc y) (pos≤pos (Nat.s≤s y≤x)) = neg≤neg y≤x negℤ-anti-full