Skip to content

Commit afdc679

Browse files
definition of Irreducible and Rough; refactoring of Prime and Composite cf. #2180 (#2181)
* definition of `Irreducible`; refactoring of `Prime` and `Composite` * tidying up old cruft * knock-on consequences: `Coprimality` * considerable refactoring of `Primality` * knock-on consequences: `Coprimality` * refactoring: no appeal to `Data.Nat.Induction` * refactoring: renamed `SmoothAt` and its constructor; added pattern synonym; proved `Decidable Irreducible`; misc. tweaks * knock-on consequences: `Coprimality` * refactoring: removed `NonZero` analysis; misc. tweaks * knock-on consequences: `Coprimality`; tightened `import`s * knock-on consequences: `Coprimality`; tightened `import`s * refactoring: every number is `1-rough` * knock-on consequences: `Coprimality`; use of smart constructor * refactoring: every number is `0-rough`; streamlining; inverse to `prime`; documentation * attempt to optimise for a nearly-common case pseudo-constructor * fiddling now * refactoring: better use of `primality` API * `Rough` is bounded * removed unnecessary implicits * comment * refactoring: comprehensive shift over to `NonTrivial` instances * refactoring: oops! * tidying up: removed infixes * tidying up: restored `rough⇒≤` * further refinements; added `NonTrivial` proofs * tidying up * moving definitions to `Data.Nat.Base` * propagated changes; many more explicit s required? * `NonTrivial` is `Recomputable` * all instances of `NonTrivial` now irrelevant; weird to need `NonTrivial 2` explicitly * tidying up `Coprimality`, eg with `variable`s * `NonTrivial` is `Decidable` * systematise proofs of `Decidable` properties via the `UpTo` predicates * explicit quantifier now in `composite≢` * Nathan's simplification * interaction of `NonZero` and `NonTrivial` instances * divisor now a record field; final lemmas: closure under divisibility; plus partial `CHANGELOG` entries * '(non-)instances' become '(counter-)examples' * stylistics * removed `k` as a variable/parameter * renamed fields and smart constructors * moved smart constructors; made a local definition * tidying up * refactoring per review comments: equivalence of `UpTo` predicates; making more things `private` * tidying up: names congruent to ordering relation * removed variable `k`; removed old proof in favour of new one with `NonZero` instance * removed `recomputable` in favour of a private lemma * regularised use of instance brackets * made instances more explicit * made instances more explicit * blank line * made `nonTrivial⇒nonZero` take an explicit argument in lieu of being able to make it an `instance` * regularised use of instance brackets * regularised use of instance brackets * trimming * tidied up `Coprimality` entries * Make HasBoundedNonTrivialDivisor infix * Make NonTrivial into a record to fix instance resolution bug * Move HasNonTrivialDivisor to Divisibility.Core and hide UpTo lemmas * Rearrange file to follow standard ordering of lemmas in the rest of the library * Move UpTo predicates into decidability proofs * No-op refactoring to curb excessively long lines * Make a couple of names consistent with style-guide * new definition of `Prime` * renamed fundamental definition * one last reference in `CHANGELOG` * more better words; one fewer definition * tidied up `Definitions` section; rejigged order of proofs of properties to reflect definitional order * refactored proof of `prime⇒irreducible` * finishing touches * missing lemma from irrelevant instance list * regularised final proof to use `contradiction` * added fixity `infix 10` * added fixity `infix 10`; made `composite` a pattern synonym; knock-on improvements * comprehensive `CHNAGELOG` entry; whitespace fixes * Rename 1<2+ to sz<ss * Rename hasNonTrivialDivisor relation * Updated CHANGELOG --------- Co-authored-by: MatthewDaggitt <[email protected]>
1 parent 61335e5 commit afdc679

File tree

8 files changed

+550
-151
lines changed

8 files changed

+550
-151
lines changed

CHANGELOG.md

+74-30
Original file line numberDiff line numberDiff line change
@@ -645,10 +645,14 @@ Non-backwards compatible changes
645645
* To make it easier to use, reason about and read, the definition has been
646646
changed to:
647647
```agda
648-
Prime 0 = ⊥
649-
Prime 1 = ⊥
650-
Prime n = ∀ {d} → 2 ≤ d → d < n → d ∤ n
648+
record Prime (p : ℕ) : Set where
649+
constructor prime
650+
field
651+
.{{nontrivial}} : NonTrivial p
652+
notComposite : ¬ Composite p
651653
```
654+
where `Composite` is now defined as the diagonal of the new relation
655+
`_HasNonTrivialDivisorLessThan_` in `Data.Nat.Divisibility.Core`.
652656

653657
### Changes to operation reduction behaviour in `Data.Rational(.Unnormalised)`
654658

@@ -2754,6 +2758,18 @@ Additions to existing modules
27542758
s≤″s⁻¹ : suc m ≤″ suc n → m ≤″ n
27552759
s<″s⁻¹ : suc m <″ suc n → m <″ n
27562760
2761+
pattern 2+ n = suc (suc n)
2762+
pattern 1<2+n {n} = s<s (z<s {n})
2763+
2764+
NonTrivial : Pred ℕ 0ℓ
2765+
instance nonTrivial : NonTrivial (2+ n)
2766+
n>1⇒nonTrivial : 1 < n → NonTrivial n
2767+
nonZero⇒≢1⇒nonTrivial : .{{NonZero n}} → n ≢ 1 → NonTrivial n
2768+
recompute-nonTrivial : .{{NonTrivial n}} → NonTrivial n
2769+
nonTrivial⇒nonZero : .{{NonTrivial n}} → NonZero n
2770+
nonTrivial⇒n>1 : .{{NonTrivial n}} → 1 < n
2771+
nonTrivial⇒≢1 : .{{NonTrivial n}} → n ≢ 1
2772+
27572773
_⊔′_ : ℕ → ℕ → ℕ
27582774
_⊓′_ : ℕ → ℕ → ℕ
27592775
∣_-_∣′ : ℕ → ℕ → ℕ
@@ -2779,20 +2795,42 @@ Additions to existing modules
27792795
<-asym : Asymmetric _<_
27802796
```
27812797

2782-
* Added a new pattern synonym to `Data.Nat.Divisibility.Core`:
2798+
* Added a new pattern synonym and a new definition to `Data.Nat.Divisibility.Core`:
27832799
```agda
27842800
pattern divides-refl q = divides q refl
2801+
record _HasNonTrivialDivisorLessThan_ (m n : ℕ) : Set where
2802+
```
2803+
2804+
* Added new proofs to `Data.Nat.Divisibility`:
2805+
```agda
2806+
hasNonTrivialDivisor-≢ : .{{NonTrivial d}} → .{{NonZero n}} → d ≢ n → d ∣ n → n HasNonTrivialDivisorLessThan n
2807+
hasNonTrivialDivisor-∣ : m HasNonTrivialDivisorLessThan n → m ∣ o → o HasNonTrivialDivisorLessThan n
2808+
hasNonTrivialDivisor-≤ : m HasNonTrivialDivisorLessThan n → n ≤ o → m HasNonTrivialDivisorLessThan o
27852809
```
27862810

2787-
* Added new definitions and proofs to `Data.Nat.Primality`:
2811+
* Added new definitions, smart constructors and proofs to `Data.Nat.Primality`:
27882812
```agda
2789-
Composite : ℕ → Set
2790-
composite? : Decidable Composite
2791-
composite⇒¬prime : Composite n → ¬ Prime n
2792-
¬composite⇒prime : 2 ≤ n → ¬ Composite n → Prime n
2793-
prime⇒¬composite : Prime n → ¬ Composite n
2794-
¬prime⇒composite : 2 ≤ n → ¬ Prime n → Composite n
2795-
euclidsLemma : Prime p → p ∣ m * n → p ∣ m ⊎ p ∣ n
2813+
infix 10 _Rough_
2814+
_Rough_ : ℕ → Pred ℕ _
2815+
0-rough : 0 Rough n
2816+
1-rough : 1 Rough n
2817+
2-rough : 2 Rough n
2818+
rough⇒≤ : .{{NonTrivial n}} → m Rough n → m ≤ n
2819+
∤⇒rough-suc : m ∤ n → m Rough n → (suc m) Rough n
2820+
rough∧∣⇒rough : m Rough o → n ∣ o → m Rough n
2821+
Composite : ℕ → Set
2822+
composite-≢ : .{{NonTrivial d}} → .{{NonZero n}} → d ≢ n → d ∣ n → Composite n
2823+
composite-∣ : .{{NonZero n}} → Composite m → m ∣ n → Composite n
2824+
composite? : Decidable Composite
2825+
Irreducible : ℕ → Set
2826+
irreducible? : Decidable Irreducible
2827+
composite⇒¬prime : Composite n → ¬ Prime n
2828+
¬composite⇒prime : .{{NonTrivial n} → ¬ Composite n → Prime n
2829+
prime⇒¬composite : Prime n → ¬ Composite n
2830+
¬prime⇒composite : .{{NonTrivial n} → ¬ Prime n → Composite n
2831+
prime⇒irreducible : Prime p → Irreducible p
2832+
irreducible⇒prime : .{{NonTrivial p}} → Irreducible p → Prime p
2833+
euclidsLemma : Prime p → p ∣ m * n → p ∣ m ⊎ p ∣ n
27962834
```
27972835

27982836
* Added new proofs in `Data.Nat.Properties`:
@@ -2802,16 +2840,20 @@ Additions to existing modules
28022840
n+1+m≢m : n + suc m ≢ m
28032841
m*n≡0⇒m≡0 : .{{_ : NonZero n}} → m * n ≡ 0 → m ≡ 0
28042842
n>0⇒n≢0 : n > 0 → n ≢ 0
2805-
m^n≢0 : .{{_ : NonZero m}} → NonZero (m ^ n)
28062843
m*n≢0 : .{{_ : NonZero m}} .{{_ : NonZero n}} → NonZero (m * n)
2844+
m*n≢0⇒m≢0 : .{{NonZero (m * n)}} → NonZero m
2845+
m*n≢0⇒n≢0 : .{{NonZero (m * n)}} → NonZero n
2846+
m≢0∧n>1⇒m*n>1 : .{{_ : NonZero m}} .{{_ : NonTrivial n}} → NonTrivial (m * n)
2847+
n≢0∧m>1⇒m*n>1 : .{{_ : NonZero n}} .{{_ : NonTrivial m}} → NonTrivial (m * n)
2848+
m^n≢0 : .{{_ : NonZero m}} → NonZero (m ^ n)
28072849
m≤n⇒n∸m≤n : m ≤ n → n ∸ m ≤ n
28082850
28092851
s<s-injective : s<s p ≡ s<s q → p ≡ q
28102852
<-step : m < n → m < 1 + n
28112853
m<1+n⇒m<n∨m≡n : m < suc n → m < n ⊎ m ≡ n
28122854
28132855
pred-mono-≤ : m ≤ n → pred m ≤ pred n
2814-
pred-mono-< : ._ : NonZero m → m < n → pred m < pred n
2856+
pred-mono-< : .{{_ : NonZero m}} → m < n → pred m < pred n
28152857
28162858
z<′s : zero <′ suc n
28172859
s<′s : m <′ n → suc m <′ suc n
@@ -2852,7 +2894,7 @@ Additions to existing modules
28522894
⊓≡⊓′ : m ⊓ n ≡ m ⊓′ n
28532895
∣-∣≡∣-∣′ : ∣ m - n ∣ ≡ ∣ m - n ∣′
28542896
2855-
nonZero? : Decidable NonZero
2897+
nonTrivial? : Decidable NonTrivial
28562898
eq? : A ↣ ℕ → DecidableEquality A
28572899
≤-<-connex : Connex _≤_ _<_
28582900
≥->-connex : Connex _≥_ _>_
@@ -2878,21 +2920,21 @@ Additions to existing modules
28782920
m%n≤n : .{{_ : NonZero n}} → m % n ≤ n
28792921
m*n/m!≡n/[m∸1]! : .{{_ : NonZero m}} → m * n / m ! ≡ n / (pred m) !
28802922
2881-
%-congˡ : ._ : NonZero o → m ≡ n → m % o ≡ n % o
2882-
%-congʳ : ._ : NonZero m ⦄ .⦃ _ : NonZero n → m ≡ n → o % m ≡ o % n
2883-
m≤n⇒[n∸m]%m≡n%m : ._ : NonZero m → m ≤ n → (n ∸ m) % m ≡ n % m
2884-
m*n≤o⇒[o∸m*n]%n≡o%n : ._ : NonZero n → m * n ≤ o → (o ∸ m * n) % n ≡ o % n
2885-
m∣n⇒o%n%m≡o%m : ._ : NonZero m ⦄ .⦃ _ : NonZero n → m ∣ n → o % n % m ≡ o % m
2886-
m<n⇒m%n≡m : ._ : NonZero n → m < n → m % n ≡ m
2887-
m*n/o*n≡m/o : ._ : NonZero o ⦄ ⦃ _ : NonZero (o * n) → m * n / (o * n) ≡ m / o
2888-
m<n*o⇒m/o<n : ._ : NonZero o → m < n * o → m / o < n
2889-
[m∸n]/n≡m/n∸1 : ._ : NonZero n → (m ∸ n) / n ≡ pred (m / n)
2890-
[m∸n*o]/o≡m/o∸n : ._ : NonZero o → (m ∸ n * o) / o ≡ m / o ∸ n
2891-
m/n/o≡m/[n*o] : ._ : NonZero n ⦄ .⦃ _ : NonZero o ⦄ .⦃ _ : NonZero (n * o) → m / n / o ≡ m / (n * o)
2892-
m%[n*o]/o≡m/o%n : ._ : NonZero n ⦄ .⦃ _ : NonZero o ⦄ ⦃ _ : NonZero (n * o) → m % (n * o) / o ≡ m / o % n
2893-
m%n*o≡m*o%[n*o] : ._ : NonZero n ⦄ ⦃ _ : NonZero (n * o) → m % n * o ≡ m * o % (n * o)
2923+
%-congˡ : .{{_ : NonZero o}} → m ≡ n → m % o ≡ n % o
2924+
%-congʳ : .{{_ : NonZero m}} .{{_ : NonZero n}} → m ≡ n → o % m ≡ o % n
2925+
m≤n⇒[n∸m]%m≡n%m : .{{_ : NonZero m}} → m ≤ n → (n ∸ m) % m ≡ n % m
2926+
m*n≤o⇒[o∸m*n]%n≡o%n : .{{_ : NonZero n}} → m * n ≤ o → (o ∸ m * n) % n ≡ o % n
2927+
m∣n⇒o%n%m≡o%m : .{{_ : NonZero m}} .{{_ : NonZero n}} → m ∣ n → o % n % m ≡ o % m
2928+
m<n⇒m%n≡m : .{{_ : NonZero n}} → m < n → m % n ≡ m
2929+
m*n/o*n≡m/o : .{{_ : NonZero o}} {{_ : NonZero (o * n)}} → m * n / (o * n) ≡ m / o
2930+
m<n*o⇒m/o<n : .{{_ : NonZero o}} → m < n * o → m / o < n
2931+
[m∸n]/n≡m/n∸1 : .{{_ : NonZero n}} → (m ∸ n) / n ≡ pred (m / n)
2932+
[m∸n*o]/o≡m/o∸n : .{{_ : NonZero o}} → (m ∸ n * o) / o ≡ m / o ∸ n
2933+
m/n/o≡m/[n*o] : .{{_ : NonZero n}} .{{_ : NonZero o}} .{{_ : NonZero (n * o)}} → m / n / o ≡ m / (n * o)
2934+
m%[n*o]/o≡m/o%n : .{{_ : NonZero n}} .{{_ : NonZero o}} {{_ : NonZero (n * o)}} → m % (n * o) / o ≡ m / o % n
2935+
m%n*o≡m*o%[n*o] : .{{_ : NonZero n}} {{_ : NonZero (n * o)}} → m % n * o ≡ m * o % (n * o)
28942936
2895-
[m*n+o]%[p*n]≡[m*n]%[p*n]+o : _ : NonZero (p * n) → o < n → (m * n + o) % (p * n) ≡ (m * n) % (p * n) + o
2937+
[m*n+o]%[p*n]≡[m*n]%[p*n]+o : {{_ : NonZero (p * n)}} → o < n → (m * n + o) % (p * n) ≡ (m * n) % (p * n) + o
28962938
```
28972939

28982940
* Added new proofs in `Data.Nat.Divisibility`:
@@ -3629,7 +3671,9 @@ This is a full list of proofs that have changed form to use irrelevant instance
36293671

36303672
* In `Data.Nat.Coprimality`:
36313673
```
3632-
Bézout-coprime : ∀ {i j d} → Bézout.Identity (suc d) (i * suc d) (j * suc d) → Coprime i j
3674+
¬0-coprimeTo-2+ : ∀ {n} → ¬ Coprime 0 (2 + n)
3675+
Bézout-coprime : ∀ {i j d} → Bézout.Identity (suc d) (i * suc d) (j * suc d) → Coprime i j
3676+
prime⇒coprime : ∀ m → Prime m → ∀ n → 0 < n → n < m → Coprime m n
36333677
```
36343678

36353679
* In `Data.Nat.Divisibility`

src/Data/Nat/Base.agda

+47-8
Original file line numberDiff line numberDiff line change
@@ -21,13 +21,17 @@ open import Relation.Binary.Core using (Rel)
2121
open import Relation.Binary.PropositionalEquality.Core
2222
using (_≡_; _≢_; refl)
2323
open import Relation.Nullary.Negation.Core using (¬_; contradiction)
24+
open import Relation.Unary using (Pred)
2425

2526
------------------------------------------------------------------------
2627
-- Types
2728

2829
open import Agda.Builtin.Nat public
2930
using (zero; suc) renaming (Nat to ℕ)
3031

32+
--smart constructor
33+
pattern 2+ n = suc (suc n)
34+
3135
------------------------------------------------------------------------
3236
-- Boolean equality relation
3337

@@ -59,8 +63,9 @@ m < n = suc m ≤ n
5963

6064
-- Smart constructors of _<_
6165

62-
pattern z<s {n} = s≤s (z≤n {n})
66+
pattern z<s {n} = s≤s (z≤n {n})
6367
pattern s<s {m} {n} m<n = s≤s {m} {n} m<n
68+
pattern sz<ss {n} = s<s (z<s {n})
6469

6570
-- Smart destructors of _≤_, _<_
6671

@@ -72,7 +77,7 @@ s<s⁻¹ (s<s m<n) = m<n
7277

7378

7479
------------------------------------------------------------------------
75-
-- other ordering relations
80+
-- Other derived ordering relations
7681

7782
_≥_ : Rel ℕ 0ℓ
7883
m ≥ n = n ≤ m
@@ -95,13 +100,19 @@ a ≯ b = ¬ a > b
95100
------------------------------------------------------------------------
96101
-- Simple predicates
97102

98-
-- Defining `NonZero` in terms of `T` and therefore ultimately `⊤` and
99-
-- `⊥` allows Agda to automatically infer nonZero-ness for any natural
100-
-- of the form `suc n`. Consequently in many circumstances this
101-
-- eliminates the need to explicitly pass a proof when the NonZero
102-
-- argument is either an implicit or an instance argument.
103+
-- Defining these predicates in terms of `T` and therefore ultimately
104+
-- `⊤` and `⊥` allows Agda to automatically infer them for any natural
105+
-- of the correct form. Consequently in many circumstances this
106+
-- eliminates the need to explicitly pass a proof when the predicate
107+
-- argument is either an implicit or an instance argument. See `_/_`
108+
-- and `_%_` further down this file for examples.
103109
--
104-
-- See `Data.Nat.DivMod` for an example.
110+
-- Furthermore, defining these predicates as single-field records
111+
-- (rather defining them directly as the type of their field) is
112+
-- necessary as the current version of Agda is far better at
113+
-- reconstructing meta-variable values for the record parameters.
114+
115+
-- A predicate saying that a number is not equal to 0.
105116

106117
record NonZero (n : ℕ) : Set where
107118
field
@@ -130,6 +141,34 @@ instance
130141
>-nonZero⁻¹ : n .{{NonZero n}} n > 0
131142
>-nonZero⁻¹ (suc n) = z<s
132143

144+
-- The property of being a non-zero, non-unit
145+
146+
record NonTrivial (n : ℕ) : Set where
147+
field
148+
nonTrivial : T (1 <ᵇ n)
149+
150+
-- Instances
151+
152+
instance
153+
nonTrivial : {n} NonTrivial (2+ n)
154+
nonTrivial = _
155+
156+
-- Constructors
157+
158+
n>1⇒nonTrivial : {n} n > 1 NonTrivial n
159+
n>1⇒nonTrivial sz<ss = _
160+
161+
-- Destructors
162+
163+
nonTrivial⇒nonZero : n .{{NonTrivial n}} NonZero n
164+
nonTrivial⇒nonZero (2+ _) = _
165+
166+
nonTrivial⇒n>1 : n .{{NonTrivial n}} n > 1
167+
nonTrivial⇒n>1 (2+ _) = sz<ss
168+
169+
nonTrivial⇒≢1 : {n} .{{NonTrivial n}} n ≢ 1
170+
nonTrivial⇒≢1 {{()}} refl
171+
133172
------------------------------------------------------------------------
134173
-- Raw bundles
135174

0 commit comments

Comments
 (0)