Skip to content
Merged
Show file tree
Hide file tree
Changes from 11 commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
9 changes: 9 additions & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -16,6 +16,10 @@ Bug-fixes
Non-backwards compatible changes
--------------------------------

* Replaced O(n) implementation of `Data.Nat.Binary.fromℕ` with O(log n). The old
implementation is maintained under `Data.Nat.Binary.fromℕ'` and proven to be
equivalent.

Deprecated modules
------------------

Expand Down Expand Up @@ -61,3 +65,8 @@ Other minor additions
respʳ-flip : _≈_ Respectsʳ (flip _≈_)
respˡ-flip : _≈_ Respectsˡ (flip _≈_)
```

* Added new proof to `Data.Nat.DivMod`:
```agda
m/n≤m : ∀ m n {≢0} → (m / n) {≢0} ≤ m
```
20 changes: 17 additions & 3 deletions src/Data/Nat/Binary/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -13,7 +13,9 @@
module Data.Nat.Binary.Base where

open import Algebra.Core using (Op₂)
open import Data.Bool.Base using (if_then_else_)
open import Data.Nat.Base as ℕ using (ℕ)
open import Data.Nat.DivMod using (_%_ ; _/_)
open import Data.Sum.Base using (_⊎_)
open import Function.Base using (_on_)
open import Level using (0ℓ)
Expand Down Expand Up @@ -112,10 +114,22 @@ toℕ zero = 0
toℕ 2[1+ x ] = 2 ℕ.* (ℕ.suc (toℕ x))
toℕ 1+[2 x ] = ℕ.suc (2 ℕ.* (toℕ x))

-- Costs O(n), could be improved using `_/_` and `_%_`
fromℕ : ℕ → ℕᵇ
fromℕ 0 = zero
fromℕ (ℕ.suc n) = suc (fromℕ n)
fromℕ n = helper n n
module fromℕ where
helper : ℕ → ℕ → ℕᵇ
helper 0 _ = zero
helper (ℕ.suc n) (ℕ.suc w) =
if (n % 2 ℕ.≡ᵇ 0)
then 1+[2 helper (n / 2) w ]
else 2[1+ helper (n / 2) w ]
-- Impossible case
helper _ 0 = zero

-- An alternative slower definition
fromℕ' : ℕ → ℕᵇ
fromℕ' 0 = zero
fromℕ' (ℕ.suc n) = suc (fromℕ' n)

-- An alternative ordering lifted from ℕ

Expand Down
99 changes: 84 additions & 15 deletions src/Data/Nat/Binary/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -12,11 +12,16 @@ open import Algebra.Bundles
open import Algebra.Morphism.Structures
import Algebra.Morphism.MonoidMonomorphism as MonoidMonomorphism
open import Algebra.Consequences.Propositional
open import Data.Bool.Base using (if_then_else_; Bool; true; false)
open import Data.Maybe.Base using (Maybe; just; nothing)
open import Data.Nat.Binary.Base
open import Data.Nat as ℕ using (ℕ; z≤n; s≤s)
import Data.Nat.DivMod as ℕ÷
open import Data.Nat.DivMod.Core using (divₕ-extractAcc)
import Data.Nat.Properties as ℕₚ
open import Data.Nat.Solver
open import Data.Product using (_,_; proj₁; proj₂; ∃)
open import Data.Product using (_×_; _,_; proj₁; proj₂; ∃)
import Data.Product.Properties as Productₚ
open import Data.Sum.Base using (_⊎_; inj₁; inj₂)
open import Function.Base using (_∘_; _$_; id)
open import Function.Definitions using (Injective)
Expand Down Expand Up @@ -98,15 +103,75 @@ toℕ-pred zero = refl
toℕ-pred 2[1+ x ] = cong ℕ.pred $ sym $ ℕₚ.*-distribˡ-+ 2 1 (toℕ x)
toℕ-pred 1+[2 x ] = toℕ-double x

toℕ-fromℕ : toℕ ∘ fromℕ ≗ id
toℕ-fromℕ 0 = refl
toℕ-fromℕ (ℕ.suc n) = begin
toℕ (fromℕ (ℕ.suc n)) ≡⟨⟩
toℕ (suc (fromℕ n)) ≡⟨ toℕ-suc (fromℕ n) ⟩
ℕ.suc (toℕ (fromℕ n)) ≡⟨ cong ℕ.suc (toℕ-fromℕ n) ⟩
toℕ-fromℕ' : toℕ ∘ fromℕ' ≗ id
toℕ-fromℕ' 0 = refl
toℕ-fromℕ' (ℕ.suc n) = begin
toℕ (fromℕ' (ℕ.suc n)) ≡⟨⟩
toℕ (suc (fromℕ' n)) ≡⟨ toℕ-suc (fromℕ' n) ⟩
ℕ.suc (toℕ (fromℕ' n)) ≡⟨ cong ℕ.suc (toℕ-fromℕ' n) ⟩
ℕ.suc n ∎
where open ≡-Reasoning

fromℕ≡fromℕ' : fromℕ ≗ fromℕ'
fromℕ≡fromℕ' n = fromℕ-helper≡fromℕ' n n ℕₚ.≤-refl
where
split : ℕᵇ → Maybe Bool × ℕᵇ
split zero = nothing , zero
split 2[1+ n ] = just false , n
split 1+[2 n ] = just true , n

head = proj₁ ∘ split
tail = proj₂ ∘ split

split-cong : ∀ m n → split m ≡ split n → m ≡ n
split-cong zero zero refl = refl
split-cong 2[1+ m ] 2[1+ .m ] refl = refl
split-cong 1+[2 m ] 1+[2 .m ] refl = refl

split-if : ∀ x xs → split (if x then 1+[2 xs ] else 2[1+ xs ]) ≡ (just x , xs)
split-if false xs = refl
split-if true xs = refl

head-suc : ∀ n → head (suc (suc (suc n))) ≡ head (suc n)
head-suc zero = refl
head-suc 2[1+ n ] = refl
head-suc 1+[2 n ] = refl

tail-suc : ∀ n → suc (tail (suc n)) ≡ tail (suc (suc (suc n)))
tail-suc zero = refl
tail-suc 2[1+ n ] = refl
tail-suc 1+[2 n ] = refl

head-homo : ∀ n → head (suc (fromℕ' n)) ≡ just (n ℕ÷.% 2 ℕ.≡ᵇ 0)
head-homo ℕ.zero = refl
head-homo (ℕ.suc ℕ.zero) = refl
head-homo (ℕ.suc (ℕ.suc n)) = trans (head-suc (fromℕ' n)) (head-homo n)

open ≡-Reasoning

tail-homo : ∀ n → tail (suc (fromℕ' n)) ≡ fromℕ' (n ℕ÷./ 2)
tail-homo ℕ.zero = refl
tail-homo (ℕ.suc ℕ.zero) = refl
tail-homo (ℕ.suc (ℕ.suc n)) = begin
tail (suc (fromℕ' (ℕ.suc (ℕ.suc n)))) ≡˘⟨ tail-suc (fromℕ' n) ⟩
suc (tail (suc (fromℕ' n))) ≡⟨ cong suc (tail-homo n) ⟩
fromℕ' (ℕ.suc (n ℕ÷./ 2)) ≡⟨ cong fromℕ' (sym (divₕ-extractAcc 1 1 n 1)) ⟩
fromℕ' (ℕ.suc (ℕ.suc n) ℕ÷./ 2) ∎

fromℕ-helper≡fromℕ' : ∀ n w → n ℕ.≤ w → fromℕ.helper n n w ≡ fromℕ' n
fromℕ-helper≡fromℕ' ℕ.zero w p = refl
fromℕ-helper≡fromℕ' (ℕ.suc n) (ℕ.suc w) (s≤s n≤w) =
split-cong _ _ (begin
split (fromℕ.helper n (ℕ.suc n) (ℕ.suc w)) ≡⟨ split-if _ _ ⟩
just (n ℕ÷.% 2 ℕ.≡ᵇ 0) , fromℕ.helper n (n ℕ÷./ 2) w ≡⟨ cong (_ ,_) rec-n/2 ⟩
just (n ℕ÷.% 2 ℕ.≡ᵇ 0) , fromℕ' (n ℕ÷./ 2) ≡˘⟨ cong₂ _,_ (head-homo n) (tail-homo n) ⟩
head (fromℕ' (ℕ.suc n)) , tail (fromℕ' (ℕ.suc n)) ≡⟨⟩
split (fromℕ' (ℕ.suc n)) ∎)
where rec-n/2 = fromℕ-helper≡fromℕ' (n ℕ÷./ 2) w (ℕₚ.≤-trans (ℕ÷.m/n≤m n 2) n≤w)

toℕ-fromℕ : toℕ ∘ fromℕ ≗ id
toℕ-fromℕ n rewrite fromℕ≡fromℕ' n = toℕ-fromℕ' n

toℕ-injective : Injective _≡_ _≡_ toℕ
toℕ-injective {zero} {zero} _ = refl
toℕ-injective {2[1+ x ]} {2[1+ y ]} 2[1+xN]≡2[1+yN] = cong 2[1+_] x≡y
Expand Down Expand Up @@ -144,7 +209,7 @@ fromℕ-injective {x} {y} f[x]≡f[y] = begin
where open ≡-Reasoning

fromℕ-toℕ : fromℕ ∘ toℕ ≗ id
fromℕ-toℕ = toℕ-injective ∘ toℕ-fromℕ ∘ toℕ
fromℕ-toℕ = toℕ-injective ∘ toℕ-fromℕ ∘ toℕ

fromℕ-pred : ∀ n → fromℕ (ℕ.pred n) ≡ pred (fromℕ n)
fromℕ-pred n = begin
Expand Down Expand Up @@ -654,15 +719,19 @@ suc-+ 1+[2 x ] 1+[2 y ] = refl
1+≗suc : (1ᵇ +_) ≗ suc
1+≗suc = suc-+ zero

fromℕ-homo-+ : ∀ m n → fromℕ (m ℕ.+ n) ≡ fromℕ m + fromℕ n
fromℕ-homo-+ 0 _ = refl
fromℕ-homo-+ (ℕ.suc m) n = begin
fromℕ ((ℕ.suc m) ℕ.+ n) ≡⟨⟩
suc (fromℕ (m ℕ.+ n)) ≡⟨ cong suc (fromℕ-homo-+ m n) ⟩
fromℕ'-homo-+ : ∀ m n → fromℕ' (m ℕ.+ n) ≡ fromℕ' m + fromℕ' n
fromℕ'-homo-+ 0 _ = refl
fromℕ'-homo-+ (ℕ.suc m) n = begin
fromℕ' ((ℕ.suc m) ℕ.+ n) ≡⟨⟩
suc (fromℕ' (m ℕ.+ n)) ≡⟨ cong suc (fromℕ'-homo-+ m n) ⟩
suc (a + b) ≡⟨ sym (suc-+ a b) ⟩
(suc a) + b ≡⟨⟩
(fromℕ (ℕ.suc m)) + (fromℕ n) ∎
where open ≡-Reasoning; a = fromℕ m; b = fromℕ n
(fromℕ' (ℕ.suc m)) + (fromℕ' n) ∎
where open ≡-Reasoning; a = fromℕ' m; b = fromℕ' n

fromℕ-homo-+ : ∀ m n → fromℕ (m ℕ.+ n) ≡ fromℕ m + fromℕ n
fromℕ-homo-+ m n rewrite fromℕ≡fromℕ' (m ℕ.+ n) | fromℕ≡fromℕ' m | fromℕ≡fromℕ' n =
fromℕ'-homo-+ m n

------------------------------------------------------------------------
-- Algebraic properties of _+_
Expand Down
6 changes: 6 additions & 0 deletions src/Data/Nat/DivMod.agda
Original file line number Diff line number Diff line change
Expand Up @@ -185,6 +185,12 @@ m/n*n≤m m n@(suc n-1) = begin
m % n + (m / n) * n ≡⟨ sym (m≡m%n+[m/n]*n m n-1) ⟩
m ∎

m/n≤m : ∀ m n {≢0} → (m / n) {≢0} ≤ m
m/n≤m m n@(suc n-1) = *-cancelʳ-≤ (m / n) m n-1 (begin
(m / n) * n ≤⟨ m/n*n≤m m n ⟩
m ≤⟨ m≤m*n m (s≤s z≤n) ⟩
m * n ∎)

m/n<m : ∀ m n {≢0} → m ≥ 1 → n ≥ 2 → (m / n) {≢0} < m
m/n<m m n@(suc n-1) m≥1 n≥2 = *-cancelʳ-< {n} (m / n) m (begin-strict
(m / n) * n ≤⟨ m/n*n≤m m n ⟩
Expand Down