Skip to content

Commit 82c1b9d

Browse files
authored
Add and use pattern synonyms for strict order over naturals
1 parent 1798dfc commit 82c1b9d

File tree

9 files changed

+227
-130
lines changed

9 files changed

+227
-130
lines changed

CHANGELOG.md

+27
Original file line numberDiff line numberDiff line change
@@ -518,6 +518,10 @@ Deprecated modules
518518
Relation.Binary.Properties.MeetSemilattice.agda ↦ Relation.Binary.Lattice.Properties.MeetSemilattice.agda
519519
```
520520

521+
### Deprecation of `Data.Nat.Properties.Core`
522+
523+
* The module `Data.Nat.Properties.Core` has been deprecated, and its one entry moved to `Data.Nat.Properties`
524+
521525
Deprecated names
522526
----------------
523527

@@ -1003,6 +1007,8 @@ Other minor changes
10031007
combine-surjective : ∀ x → ∃₂ λ y z → combine y z ≡ x
10041008
10051009
lower₁-injective : lower₁ i n≢i ≡ lower₁ j n≢j → i ≡ j
1010+
1011+
i<1+i : i < suc i
10061012
```
10071013

10081014
* Added new functions in `Data.Integer.Base`:
@@ -1065,6 +1071,16 @@ Other minor changes
10651071
∷ʳ-++ : xs ∷ʳ a ++ ys ≡ xs ++ a ∷ ys
10661072
```
10671073

1074+
* Added new patterns and definitions to `Data.Nat.Base`:
1075+
```agda
1076+
pattern z<s {n} = s≤s (z≤n {n})
1077+
pattern s<s {m} {n} m<n = s≤s {m} {n} m<n
1078+
1079+
pattern <′-base = ≤′-refl
1080+
pattern <′-step {n} m<′n = ≤′-step {n} m<′n
1081+
```
1082+
1083+
10681084
* Added new definitions and proofs to `Data.Nat.Primality`:
10691085
```agda
10701086
Composite : ℕ → Set
@@ -1087,6 +1103,17 @@ Other minor changes
10871103
m*n≢0 : .{{_ : NonZero m}} .{{_ : NonZero n}} → NonZero (m * n)
10881104
m≤n⇒n∸m≤n : m ≤ n → n ∸ m ≤ n
10891105
1106+
≤-pred : suc m ≤ suc n → m ≤ n
1107+
s<s-injective : ∀ {p q : m < n} → s<s p ≡ s<s q → p ≡ q
1108+
<-pred : suc m < suc n → m < n
1109+
<-step : m < n → m < 1 + n
1110+
m<1+n⇒m<n∨m≡n : m < suc n → m < n ⊎ m ≡ n
1111+
1112+
z<′s : zero <′ suc n
1113+
s<′s : m <′ n → suc m <′ suc n
1114+
<⇒<′ : m < n → m <′ n
1115+
<′⇒< : m <′ n → m < n
1116+
10901117
1≤n! : 1 ≤ n !
10911118
_!≢0 : NonZero (n !)
10921119
_!*_!≢0 : NonZero (m ! * n !)

src/Data/Fin/Base.agda

+5-6
Original file line numberDiff line numberDiff line change
@@ -13,8 +13,7 @@
1313
module Data.Fin.Base where
1414

1515
open import Data.Empty using (⊥-elim)
16-
open import Data.Nat.Base as ℕ using (ℕ; zero; suc; z≤n; s≤s; _^_)
17-
open import Data.Nat.Properties.Core using (≤-pred)
16+
open import Data.Nat.Base as ℕ using (ℕ; zero; suc; z≤n; s≤s; z<s; s<s; _^_)
1817
open import Data.Product as Product using (_×_; _,_; proj₁; proj₂)
1918
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_]′)
2019
open import Function.Base using (id; _∘_; _on_; flip)
@@ -67,8 +66,8 @@ fromℕ (suc n) = suc (fromℕ n)
6766
-- fromℕ< {m} _ = "m".
6867

6968
fromℕ< : {m n} m ℕ.< n Fin n
70-
fromℕ< {zero} {suc n} m≤n = zero
71-
fromℕ< {suc m} {suc n} m≤n = suc (fromℕ< (≤-pred m≤n))
69+
fromℕ< {zero} {suc n} z<s = zero
70+
fromℕ< {suc m} {suc n} (s<s m<n) = suc (fromℕ< m<n)
7271

7372
-- fromℕ<″ m _ = "m".
7473

@@ -112,8 +111,8 @@ inject₁ zero = zero
112111
inject₁ (suc i) = suc (inject₁ i)
113112

114113
inject≤ : {m n} Fin m m ℕ.≤ n Fin n
115-
inject≤ {_} {suc n} zero le = zero
116-
inject≤ {_} {suc n} (suc i) le = suc (inject≤ i (≤-pred le))
114+
inject≤ {_} {suc n} zero _ = zero
115+
inject≤ {_} {suc n} (suc i) (s≤s m≤n) = suc (inject≤ i m≤n)
117116

118117
-- lower₁ "i" _ = "i".
119118

src/Data/Fin/Induction.agda

+5-5
Original file line numberDiff line numberDiff line change
@@ -8,7 +8,7 @@
88

99
open import Data.Fin.Base
1010
open import Data.Fin.Properties
11-
open import Data.Nat.Base as ℕ using (ℕ; zero; suc; z≤n; s≤s; _∸_)
11+
open import Data.Nat.Base as ℕ using (ℕ; zero; suc; _∸_)
1212
import Data.Nat.Induction as ℕ
1313
import Data.Nat.Properties as ℕ
1414
open import Induction
@@ -46,15 +46,15 @@ open WF public using (Acc; acc)
4646
induct : {i} Acc _<_ i P i
4747
induct {zero} _ = P₀
4848
induct {suc i} (acc rec) = Pᵢ⇒Pᵢ₊₁ i (induct (rec (inject₁ i) i<i+1))
49-
where i<i+1 = s≤s (ℕ.≤-reflexive (toℕ-inject₁ i))
49+
where i<i+1 = ℕ<⇒inject₁< (i<1+i i)
5050

5151
------------------------------------------------------------------------
5252
-- Induction over _>_
5353

5454
private
5555
acc-map : {x : Fin n} Acc ℕ._<_ (n ∸ toℕ x) Acc _>_ x
56-
acc-map {n} (acc rs) = acc (λ y y>x
57-
acc-map (rs (n ∸ toℕ y) (ℕ.∸-monoʳ-< y>x (toℕ≤n y))))
56+
acc-map {n} (acc rs) = acc λ y y>x
57+
acc-map (rs (n ∸ toℕ y) (ℕ.∸-monoʳ-< y>x (toℕ≤n y)))
5858

5959
>-wellFounded : WellFounded {A = Fin n} _>_
6060
>-wellFounded {n} x = acc-map (ℕ.<-wellFounded (n ∸ toℕ x))
@@ -69,7 +69,7 @@ private
6969
induct {i} (acc rec) with n ℕ.≟ toℕ i
7070
... | yes n≡i = subst P (toℕ-injective (trans (toℕ-fromℕ n) n≡i)) Pₙ
7171
... | no n≢i = subst P (inject₁-lower₁ i n≢i) (Pᵢ₊₁⇒Pᵢ _ Pᵢ₊₁)
72-
where Pᵢ₊₁ = induct (rec _ (s≤s (ℕ.≤-reflexive (sym (toℕ-lower₁ i n≢i)))))
72+
where Pᵢ₊₁ = induct (rec _ (ℕ.≤-reflexive (cong suc (sym (toℕ-lower₁ i n≢i)))))
7373

7474
------------------------------------------------------------------------
7575
-- Induction over _≺_

0 commit comments

Comments
 (0)