Idris2/libs/contrib/Syntax/PreorderReasoning/Generic.idr
Ohad Kammar 0c1a124704
Division theorem (#695)
Division Theorem. For every natural number `x` and positive natural
number `n`, there is a unique decomposition:
`x = q*n + r`
with `q`,`r` natural and `r` < `n`.

`q` is the quotient when dividing `x` by `n`
`r` is the remainder when dividing `x` by `n`.

This commit adds a proof for this fact, in case
we want to reason about modular arithmetic (for example, when dealing
with binary representations). A future, more systematic, development could
perhaps follow: @clayrat 's (idris1) port of Coq's binary arithmetics:

    https://github.com/sbp/idris-bi/blob/master/src/Data/Bin/DivMod.idr
    https://github.com/sbp/idris-bi/blob/master/src/Data/Biz/DivMod.idr
    https://github.com/sbp/idris-bi/blob/master/src/Data/BizMod2/DivMod.idr

In the process, it bulks up the stdlib with:
+ a generic PreorderReasoning module for arbitrary preorders,
analogous for the equational reasoning module
+ some missing facts about Nat operations.
+ Refactor some Nat order properties using a 'reflect' function

Co-authored-by: Ohad Kammar <ohad.kammar@ed.ac.uk>
Co-authored-by: G. Allais <guillaume.allais@ens-lyon.org>
2020-10-06 13:09:02 +01:00

30 lines
978 B
Idris

module Syntax.PreorderReasoning.Generic
import Decidable.Order
infixl 0 ~~
infixl 0 <~
prefix 1 |~
infix 1 ...
public export
data Step : (leq : a -> a -> Type) -> a -> a -> Type where
(...) : {leq : a -> a -> Type} -> (y : a) -> x `leq` y -> Step leq x y
public export
data FastDerivation : {leq : a -> a -> Type} -> (x : a) -> (y : a) -> Type where
(|~) : (x : a) -> FastDerivation x x
(<~) : {leq : a -> a -> Type} -> {x,y : a}
-> FastDerivation {leq = leq} x y -> {z : a} -> (step : Step leq y z)
-> FastDerivation {leq = leq} x z
public export
CalcWith : Preorder dom leq => {x,y : dom} -> FastDerivation {leq = leq} x y -> x `leq` y
CalcWith (|~ x) = reflexive x
CalcWith ((<~) der (z ... step)) = transitive {po = leq} _ _ _ (CalcWith der) step
public export
(~~) : {x,y : dom}
-> FastDerivation {leq = leq} x y -> {z : dom} -> (step : Step Equal y z)
-> FastDerivation {leq = leq} x z
(~~) der (z ... Refl) = der