Skip to content

Commit

Permalink
lawful order word
Browse files Browse the repository at this point in the history
  • Loading branch information
odderwiser committed Dec 25, 2024
1 parent 02103a7 commit 866d212
Show file tree
Hide file tree
Showing 3 changed files with 73 additions and 1 deletion.
2 changes: 1 addition & 1 deletion lib/Haskell/Law/Ord/Def.agda
Original file line number Diff line number Diff line change
Expand Up @@ -178,7 +178,7 @@ reverseLte a b c d

postulate instance

iLawfulOrdWord : IsLawfulOrd Word
-- iLawfulOrdWord : IsLawfulOrd Word

iLawfulOrdDouble : IsLawfulOrd Double

Expand Down
1 change: 1 addition & 0 deletions lib/Haskell/Law/Ord/Int.agda
Original file line number Diff line number Diff line change
Expand Up @@ -11,6 +11,7 @@ open import Haskell.Law.Bool
open import Haskell.Law.Equality
open import Haskell.Law.Eq
open import Haskell.Law.Ord.Def
open import Haskell.Law.Ord.Word
open import Haskell.Law.Int


Expand Down
71 changes: 71 additions & 0 deletions lib/Haskell/Law/Ord/Word.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,71 @@
module Haskell.Law.Ord.Word where

open import Haskell.Prim
open import Haskell.Prim.Bool
open import Haskell.Prim.Eq
open import Haskell.Prim.Ord
open import Haskell.Prim.Word using ( Word )

open import Haskell.Law.Bool
open import Haskell.Law.Equality
open import Haskell.Law.Eq
open import Haskell.Law.Eq.Instances
open import Haskell.Law.Ord.Def
open import Haskell.Law.Ord.Nat

instance
iLawfulOrdWord : IsLawfulOrd Word

iLawfulOrdWord .comparability a b
with (w2n a) | (w2n b)
... | x | y = comparability x y

iLawfulOrdWord .transitivity a b c h
with (w2n a) | (w2n b) | (w2n c)
... | x | y | z = transitivity x y z h

iLawfulOrdWord .reflexivity a
with (w2n a)
... | x = reflexivity x

iLawfulOrdWord .antisymmetry a b h
with (w2n a) | (w2n b)
... | x | y = antisymmetry x y h

iLawfulOrdWord .lte2gte a b
with (w2n a) | (w2n b)
... | x | y = lte2gte x y

iLawfulOrdWord .lt2LteNeq a b
with (w2n a) | (w2n b)
... | x | y = lt2LteNeq x y

iLawfulOrdWord .lt2gt a b
with (w2n a) | (w2n b)
... | x | y = lt2gt x y

iLawfulOrdWord .compareLt a b
with (w2n a) | (w2n b)
... | x | y = compareLt x y

iLawfulOrdWord .compareGt a b
with (w2n a) | (w2n b)
... | x | y = compareGt x y

iLawfulOrdWord .compareEq a b
with (w2n a) | (w2n b)
... | x | y = compareEq x y

iLawfulOrdWord .min2if a b
with (w2n a) | (w2n b)
... | x | y
rewrite lte2ngt x y
| ifFlip (y < x) a b
= eqReflexivity (if (y < x) then b else a)

iLawfulOrdWord .max2if a b
with (w2n a) | (w2n b)
... | x | y
rewrite gte2nlt x y
| ifFlip (x < y) a b
= eqReflexivity (if (x < y) then b else a)

0 comments on commit 866d212

Please sign in to comment.