IdrisDoc: Prelude.Nat

Prelude.Nat

toIntegerNat : Nat -> Integer

Convert a Nat to an Integer

toIntNat : Nat -> Int

Tail recursive cast Nat to Int
Note that this can overflow

succNotLTEzero : Not (LTE (S m) 0)

A successor is never less than or equal zero

succInjective : (left : Nat) -> (right : Nat) -> (p : S left = S right) -> left = right

S is injective

sucMinR : (l : Nat) -> minimum l (S l) = l
sucMinL : (l : Nat) -> minimum (S l) l = l
sucMaxR : (l : Nat) -> maximum l (S l) = S l
sucMaxL : (l : Nat) -> maximum (S l) l = S l
predSucc : (n : Nat) -> pred (S n) = n
pred : Nat -> Nat

The predecessor of a natural number. pred Z is Z.

powerZeroOne : (base : Nat) -> power base (fromInteger 0) = 1
powerSuccSuccMult : (base : Nat) -> power base (fromInteger 2) = mult base base
powerSuccPowerLeft : (base : Nat) -> (exp : Nat) -> power base (S exp) = base * power base exp
powerPowerMultPower : (base : Nat) -> (exp : Nat) -> (exp' : Nat) -> power (power base exp) exp' = power base (exp * exp')
powerOneSuccOne : (exp : Nat) -> power (fromInteger 1) exp = 1
powerOneNeutral : (base : Nat) -> power base (fromInteger 1) = base
power : Nat -> Nat -> Nat

Exponentiation of natural numbers

plusZeroRightNeutral : (left : Nat) -> left + fromInteger 0 = left
plusZeroLeftNeutral : (right : Nat) -> fromInteger 0 + right = right
plusSuccRightSucc : (left : Nat) -> (right : Nat) -> S (left + right) = left + S right
plusRightCancel : (left : Nat) -> (left' : Nat) -> (right : Nat) -> (p : left + right = left' + right) -> left = left'
plusOneSucc : (right : Nat) -> fromInteger 1 + right = S right
plusMinusLeftCancel : (left : Nat) -> (right : Nat) -> (right' : Nat) -> left + right - (left + right') = right - right'
plusLeftLeftRightZero : (left : Nat) -> (right : Nat) -> (p : left + right = left) -> right = 0
plusLeftCancel : (left : Nat) -> (right : Nat) -> (right' : Nat) -> (p : left + right = left + right') -> right = right'
plusConstantRight : (left : Nat) -> (right : Nat) -> (c : Nat) -> (p : left = right) -> left + c = right + c
plusConstantLeft : (left : Nat) -> (right : Nat) -> (c : Nat) -> (p : left = right) -> c + left = c + right
plusCommutative : (left : Nat) -> (right : Nat) -> left + right = right + left
plusAssociative : (left : Nat) -> (centre : Nat) -> (right : Nat) -> left + (centre + right) = left + centre + right
plus : (n : Nat) -> (m : Nat) -> Nat

Add two natural numbers.

n

the number to case-split on

m

the other number

multZeroRightZero : (left : Nat) -> left * 0 = 0
multZeroLeftZero : (right : Nat) -> 0 * right = 0
multRightSuccPlus : (left : Nat) -> (right : Nat) -> left * S right = left + left * right
multPowerPowerPlus : (base : Nat) -> (exp : Nat) -> (exp' : Nat) -> power base exp * power base exp' = power base (exp + exp')
multOneRightNeutral : (left : Nat) -> left * fromInteger 1 = left
multOneLeftNeutral : (right : Nat) -> fromInteger 1 * right = right
multLeftSuccPlus : (left : Nat) -> (right : Nat) -> S left * right = right + left * right
multDistributesOverPlusRight : (left : Nat) -> (centre : Nat) -> (right : Nat) -> left * (centre + right) = left * centre + left * right
multDistributesOverPlusLeft : (left : Nat) -> (centre : Nat) -> (right : Nat) -> (left + centre) * right = left * right + centre * right
multDistributesOverMinusRight : (left : Nat) -> (centre : Nat) -> (right : Nat) -> left * (centre - right) = left * centre - left * right
multDistributesOverMinusLeft : (left : Nat) -> (centre : Nat) -> (right : Nat) -> (left - centre) * right = left * right - centre * right
multCommutative : (left : Nat) -> (right : Nat) -> left * right = right * left
multAssociative : (left : Nat) -> (centre : Nat) -> (right : Nat) -> left * (centre * right) = left * centre * right
mult : Nat -> Nat -> Nat

Multiply natural numbers

modNatNZ : Nat -> (y : Nat) -> Not (y = 0) -> Nat
modNat : Nat -> Nat -> Nat
minusZeroRight : (left : Nat) -> left - fromInteger 0 = left
minusZeroN : (n : Nat) -> 0 = n - n
minusZeroLeft : (right : Nat) -> fromInteger 0 - right = 0
minusSuccSucc : (left : Nat) -> (right : Nat) -> S left - S right = left - right
minusSuccPred : (left : Nat) -> (right : Nat) -> left - S right = pred (left - right)
minusSuccOne : (n : Nat) -> S n - fromInteger 1 = n
minusPlusZero : (n : Nat) -> (m : Nat) -> n - (n + m) = 0
minusOneSuccN : (n : Nat) -> 1 = S n - n
minusMinusMinusPlus : (left : Nat) -> (centre : Nat) -> (right : Nat) -> left - centre - right = left - (centre + right)
minus : Nat -> Nat -> Nat

Subtract natural numbers. If the second number is larger than the first, return 0.

minimumZeroZeroRight : (right : Nat) -> minimum (fromInteger 0) right = 0
minimumZeroZeroLeft : (left : Nat) -> minimum left (fromInteger 0) = 0
minimumSuccSucc : (left : Nat) -> (right : Nat) -> minimum (S left) (S right) = S (minimum left right)
minimumIdempotent : (n : Nat) -> minimum n n = n
minimumCommutative : (l : Nat) -> (r : Nat) -> minimum l r = minimum r l
minimumAssociative : (l : Nat) -> (c : Nat) -> (r : Nat) -> minimum l (minimum c r) = minimum (minimum l c) r
minimum : Nat -> Nat -> Nat

Find the least of two natural numbers

maximumZeroNRight : (right : Nat) -> maximum 0 right = right
maximumZeroNLeft : (left : Nat) -> maximum left 0 = left
maximumSuccSucc : (left : Nat) -> (right : Nat) -> S (maximum left right) = maximum (S left) (S right)
maximumIdempotent : (n : Nat) -> maximum n n = n
maximumCommutative : (l : Nat) -> (r : Nat) -> maximum l r = maximum r l
maximumAssociative : (l : Nat) -> (c : Nat) -> (r : Nat) -> maximum l (maximum c r) = maximum (maximum l c) r
maximum : Nat -> Nat -> Nat

Find the greatest of two natural numbers

lteSuccRight : LTE n m -> LTE n (S m)

n < m implies n < m + 1

lteRefl : LTE n n

LTE is reflexive.

lteNTrue : (n : Nat) -> lte n n = True
lte : Nat -> Nat -> Bool

Boolean test than one Nat is less than or equal to another

lt : Nat -> Nat -> Bool

Boolean test than one Nat is strictly less than another

log2NZ : (x : Nat) -> Not (x = 0) -> Nat
log2 : Nat -> Nat
lcm : Nat -> Nat -> Nat
isZero : Nat -> Bool
isSucc : Nat -> Bool
isLTE : (m : Nat) -> (n : Nat) -> Dec (LTE m n)

A decision procedure for LTE

ifThenElseSuccSucc : (cond : Bool) -> (t : Nat) -> (f : Nat) -> S (ifThenElse cond (Delay t) (Delay f)) = ifThenElse cond (Delay (S t)) (Delay (S f))
ifThenElsePlusPlusRight : (cond : Bool) -> (right : Nat) -> (t : Nat) -> (f : Nat) -> ifThenElse cond (Delay t) (Delay f) + right = ifThenElse cond (Delay (t + right)) (Delay (f + right))
ifThenElsePlusPlusLeft : (cond : Bool) -> (left : Nat) -> (t : Nat) -> (f : Nat) -> left + ifThenElse cond (Delay t) (Delay f) = ifThenElse cond (Delay (left + t)) (Delay (left + f))
ifThenElseMultMultRight : (cond : Bool) -> (right : Nat) -> (t : Nat) -> (f : Nat) -> ifThenElse cond (Delay t) (Delay f) * right = ifThenElse cond (Delay (t * right)) (Delay (f * right))
ifThenElseMultMultLeft : (cond : Bool) -> (left : Nat) -> (t : Nat) -> (f : Nat) -> left * ifThenElse cond (Delay t) (Delay f) = ifThenElse cond (Delay (left * t)) (Delay (left * f))
hyper : Nat -> Nat -> Nat -> Nat
gte : Nat -> Nat -> Bool

Boolean test than one Nat is greater than or equal to another

gt : Nat -> Nat -> Bool

Boolean test than one Nat is strictly greater than another

gcd : Nat -> Nat -> Nat
fromLteSucc : LTE (S m) (S n) -> LTE m n

If two numbers are ordered, their predecessors are too

fromIntegerNat : Integer -> Nat

Convert an Integer to a Nat, mapping negative numbers to 0

fib : Nat -> Nat

Fibonacci numbers

fact : Nat -> Nat

Factorial function

eqSucc : (left : Nat) -> (right : Nat) -> (p : left = right) -> S left = S right

S preserves equality

divNatNZ : Nat -> (y : Nat) -> Not (y = 0) -> Nat
divNat : Nat -> Nat -> Nat
divCeilNZ : Nat -> (y : Nat) -> Not (y = fromInteger 0) -> Nat
divCeil : Nat -> Nat -> Nat
cmp : (x : Nat) -> (y : Nat) -> CmpNat x y
SIsNotZ : (S x = 0) -> Void
data Nat : Type

Natural numbers: unbounded, unsigned integers which can be pattern
matched.

Z : Nat

Zero

S : Nat -> Nat

Successor

data Multiplicative : Type

A wrapper for Nat that specifies the semigroup and monad instances that use (*)

getMultiplicative : Nat -> Multiplicative
LTESuccZeroFalse : (n : Nat) -> lte (S n) 0 = False
data LTE : (n : Nat) -> (m : Nat) -> Type

Proofs that n is less than or equal to m

n

the smaller number

m

the larger number

LTEZero : LTE 0 right

Zero is the smallest Nat

LTESucc : LTE left right -> LTE (S left) (S right)

If n <= m, then n + 1 <= m + 1

LT : Nat -> Nat -> Type

Strict less than

GTE : Nat -> Nat -> Type

Greater than or equal to

GT : Nat -> Nat -> Type

Strict greater than

data CmpNat : Nat -> Nat -> Type
CmpLT : (y : Nat) -> CmpNat x (plus x (S y))
CmpEQ : CmpNat x x
CmpGT : (x : Nat) -> CmpNat (plus y (S x)) y
data Additive : Type

A wrapper for Nat that specifies the semigroup and monad instances that use (+)

getAdditive : Nat -> Additive