Idris2/libs/contrib/Data/Nat/Order/Strict.idr
Nick Drozd 61b9a3e4e5
Define and implement Relation interfaces (#1472)
Co-authored-by: Guillaume ALLAIS <guillaume.allais@ens-lyon.org>
2021-07-09 09:06:27 +01:00

38 lines
900 B
Idris

||| Implementing `Decidable.Order.Strict` for `Data.Nat.LT`
module Data.Nat.Order.Strict
import Data.Nat
import Decidable.Order.Strict
import Decidable.Equality
import Data.Nat.Order
%default total
public export
Irreflexive Nat LT where
irreflexive {x = 0} _ impossible
irreflexive {x = S _} (LTESucc prf) =
irreflexive {rel = Nat.LT} prf
public export
Transitive Nat LT where
transitive {x} {y} xy yz =
transitive {rel = LTE} (lteSuccRight xy) yz
public export
StrictPreorder Nat LT where
public export
decLT : (a, b : Nat) -> DecOrdering {lt = LT} a b
decLT 0 0 = DecEQ Refl
decLT 0 (S b) = DecLT (LTESucc LTEZero)
decLT (S a) 0 = DecGT (LTESucc LTEZero)
decLT (S a) (S b) = case decLT a b of
DecLT a_lt_b => DecLT (LTESucc a_lt_b)
DecEQ Refl => DecEQ Refl
DecGT b_lt_a => DecGT (LTESucc b_lt_a)
public export
StrictOrdered Nat LT where
order = decLT