mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-18 00:31:57 +03:00
0c1a124704
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>
18 lines
589 B
Idris
18 lines
589 B
Idris
||| Additional properties/lemmata of Nats
|
|
module Data.Nat.Properties
|
|
|
|
import Data.Nat
|
|
|
|
%default total
|
|
|
|
export
|
|
multRightCancel : (a,b,r : Nat) -> Not (r = 0) -> a*r = b*r -> a = b
|
|
multRightCancel a b 0 r_nz ar_eq_br = void $ r_nz Refl
|
|
multRightCancel 0 0 r@(S predr) r_nz ar_eq_br = Refl
|
|
multRightCancel 0 (S b) r@(S predr) r_nz ar_eq_br impossible
|
|
multRightCancel (S a) 0 r@(S predr) r_nz ar_eq_br impossible
|
|
multRightCancel (S a) (S b) r@(S predr) r_nz ar_eq_br =
|
|
cong S $ multRightCancel a b r r_nz
|
|
$ plusLeftCancel r (a*r) (b*r) ar_eq_br
|
|
|