-
Notifications
You must be signed in to change notification settings - Fork 246
/
Copy pathCore.agda
81 lines (61 loc) · 2.5 KB
/
Core.agda
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
------------------------------------------------------------------------
-- The Agda standard library
--
-- Core definition of divisibility
------------------------------------------------------------------------
-- The definition of divisibility is split out from
-- `Data.Nat.Divisibility` to avoid a dependency cycle with
-- `Data.Nat.DivMod`.
{-# OPTIONS --cubical-compatible --safe #-}
module Data.Nat.Divisibility.Core where
open import Data.Nat.Base using (ℕ; _*_; _<_; NonTrivial)
open import Data.Nat.Properties using ([m*n]*[o*p]≡[m*o]*[n*p])
open import Relation.Nullary.Negation using (¬_)
open import Relation.Binary.Core using (Rel)
open import Relation.Binary.PropositionalEquality.Core
using (_≡_; refl)
private
variable m n o p : ℕ
------------------------------------------------------------------------
-- Main definition
--
-- m ∣ n is inhabited iff m divides n. Some sources, like Hardy and
-- Wright's "An Introduction to the Theory of Numbers", require m to
-- be non-zero. However, some things become a bit nicer if m is
-- allowed to be zero. For instance, _∣_ becomes a partial order, and
-- the gcd of 0 and 0 becomes defined.
infix 4 _∣_ _∤_
record _∣_ (m n : ℕ) : Set where
constructor divides
field quotient : ℕ
equality : n ≡ quotient * m
_∤_ : Rel ℕ _
m ∤ n = ¬ (m ∣ n)
-- Smart constructor
pattern divides-refl q = divides q refl
open _∣_ using (quotient) public
------------------------------------------------------------------------
-- Restricted divisor relation
-- Relation for having a non-trivial divisor below a given bound.
-- Useful when reasoning about primality.
infix 10 _HasNonTrivialDivisorLessThan_
record _HasNonTrivialDivisorLessThan_ (m n : ℕ) : Set where
constructor hasNonTrivialDivisor
field
{divisor} : ℕ
.{{nontrivial}} : NonTrivial divisor
divisor-< : divisor < n
divisor-∣ : divisor ∣ m
------------------------------------------------------------------------
-- DEPRECATED NAMES
------------------------------------------------------------------------
-- Please use the new names as continuing support for the old names is
-- not guaranteed.
-- Version 2.1
*-pres-∣ : o ∣ m → p ∣ n → o * p ∣ m * n
*-pres-∣ {o} {m@.(q * o)} {p} {n@.(r * p)} (divides-refl q) (divides-refl r) =
divides (q * r) ([m*n]*[o*p]≡[m*o]*[n*p] q o r p)
{-# WARNING_ON_USAGE *-pres-∣
"Warning: *-pres-∣ was deprecated in v2.1.
Please use Data.Nat.Divisibility.*-pres-∣ instead."
#-}