Skip to content

Commit

Permalink
Lawful Ord Unit proofs
Browse files Browse the repository at this point in the history
  • Loading branch information
odderwiser committed Dec 27, 2024
1 parent 09853c7 commit 02e27dd
Show file tree
Hide file tree
Showing 3 changed files with 24 additions and 2 deletions.
1 change: 1 addition & 0 deletions lib/Haskell/Law/Ord.agda
Original file line number Diff line number Diff line change
Expand Up @@ -7,3 +7,4 @@ open import Haskell.Law.Ord.Integer public
open import Haskell.Law.Ord.Maybe public
open import Haskell.Law.Ord.Nat public
open import Haskell.Law.Ord.Ordering public
open import Haskell.Law.Ord.Unit public
2 changes: 0 additions & 2 deletions lib/Haskell/Law/Ord/Def.agda
Original file line number Diff line number Diff line change
Expand Up @@ -178,8 +178,6 @@ postulate instance

iLawfulOrdDouble : IsLawfulOrd Double

iLawfulOrdUnit : IsLawfulOrd ⊤

iLawfulOrdTuple₂ : ⦃ iOrdA : Ord a ⦄ ⦃ iOrdB : Ord b ⦄
⦃ IsLawfulOrd a ⦄ ⦃ IsLawfulOrd b ⦄
IsLawfulOrd (a × b)
Expand Down
23 changes: 23 additions & 0 deletions lib/Haskell/Law/Ord/Unit.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,23 @@
module Haskell.Law.Ord.Unit where

open import Haskell.Prim
open import Haskell.Prim.Ord

open import Haskell.Law.Ord.Def
open import Haskell.Law.Eq.Instances

instance
iLawfulOrdUnit : IsLawfulOrd ⊤

iLawfulOrdUnit .comparability _ _ = refl
iLawfulOrdUnit .transitivity _ _ _ _ = refl
iLawfulOrdUnit .reflexivity _ = refl
iLawfulOrdUnit .antisymmetry _ _ _ = refl
iLawfulOrdUnit .lte2gte _ _ = refl
iLawfulOrdUnit .lt2LteNeq _ _ = refl
iLawfulOrdUnit .lt2gt _ _ = refl
iLawfulOrdUnit .compareLt _ _ = refl
iLawfulOrdUnit .compareGt _ _ = refl
iLawfulOrdUnit .compareEq _ _ = refl
iLawfulOrdUnit .min2if _ _ = refl
iLawfulOrdUnit .max2if _ _ = refl

0 comments on commit 02e27dd

Please sign in to comment.