112 lines
2.2 KiB
Idris
112 lines
2.2 KiB
Idris
|
module Quox.Syntax.Qty.Three
|
|||
|
|
|||
|
import Quox.Pretty
|
|||
|
import public Quox.Syntax.Qty
|
|||
|
import Generics.Derive
|
|||
|
|
|||
|
%default total
|
|||
|
%language ElabReflection
|
|||
|
|
|||
|
|
|||
|
public export
|
|||
|
data Three = Zero | One | Any
|
|||
|
%name Three pi, rh
|
|||
|
|
|||
|
%runElab derive "Three" [Generic, Meta, Eq, Ord, DecEq, Show]
|
|||
|
|
|||
|
|
|||
|
export
|
|||
|
PrettyHL Three where
|
|||
|
prettyM pi = hl Qty <$>
|
|||
|
case pi of
|
|||
|
Zero => ifUnicode "𝟬" "0"
|
|||
|
One => ifUnicode "𝟭" "1"
|
|||
|
Any => ifUnicode "𝛚" "*"
|
|||
|
|
|||
|
|
|||
|
public export
|
|||
|
plus : Three -> Three -> Three
|
|||
|
plus Zero rh = rh
|
|||
|
plus pi Zero = pi
|
|||
|
plus _ _ = Any
|
|||
|
|
|||
|
public export
|
|||
|
times : Three -> Three -> Three
|
|||
|
times Zero _ = Zero
|
|||
|
times _ Zero = Zero
|
|||
|
times One rh = rh
|
|||
|
times pi One = pi
|
|||
|
times Any Any = Any
|
|||
|
|
|||
|
public export
|
|||
|
data Compat3 : Three -> Three -> Type where
|
|||
|
CmpRefl : Compat3 pi pi
|
|||
|
CmpAny : Compat3 pi Any
|
|||
|
|
|||
|
public export
|
|||
|
compat3 : (pi, rh : Three) -> Dec (pi `Compat3` rh)
|
|||
|
compat3 pi rh with (decEq pi rh)
|
|||
|
compat3 pi pi | Yes Refl = Yes CmpRefl
|
|||
|
compat3 pi rh | No ne with (decEq rh Any)
|
|||
|
compat3 pi Any | No ne | Yes Refl = Yes CmpAny
|
|||
|
compat3 pi rh | No ne | No na = No $ \case
|
|||
|
CmpRefl => ne Refl
|
|||
|
CmpAny => na Refl
|
|||
|
|
|||
|
|
|||
|
public export
|
|||
|
data IsSubj3 : Three -> Type where
|
|||
|
SZero : IsSubj3 Zero
|
|||
|
SOne : IsSubj3 One
|
|||
|
|
|||
|
public export
|
|||
|
isSubj3 : (pi : Three) -> Dec (IsSubj3 pi)
|
|||
|
isSubj3 Zero = Yes SZero
|
|||
|
isSubj3 One = Yes SOne
|
|||
|
isSubj3 Any = No $ \case _ impossible
|
|||
|
|
|||
|
|
|||
|
public export
|
|||
|
data IsGlobal3 : Three -> Type where
|
|||
|
GZero : IsGlobal3 Zero
|
|||
|
GAny : IsGlobal3 Any
|
|||
|
|
|||
|
isGlobal3 : (pi : Three) -> Dec (IsGlobal3 pi)
|
|||
|
isGlobal3 Zero = Yes GZero
|
|||
|
isGlobal3 One = No $ \case _ impossible
|
|||
|
isGlobal3 Any = Yes GAny
|
|||
|
|
|||
|
|
|||
|
public export
|
|||
|
IsQty Three where
|
|||
|
zero = Zero
|
|||
|
one = One
|
|||
|
|
|||
|
(+) = plus
|
|||
|
(*) = times
|
|||
|
|
|||
|
IsZero = Equal Zero
|
|||
|
isZero = decEq Zero
|
|||
|
zeroIsZero = Refl
|
|||
|
|
|||
|
IsOne = Equal One
|
|||
|
isOne = decEq One
|
|||
|
oneIsOne = Refl
|
|||
|
|
|||
|
Compat = Compat3
|
|||
|
compat = compat3
|
|||
|
|
|||
|
IsSubj = IsSubj3
|
|||
|
isSubj = isSubj3
|
|||
|
zeroIsSubj = SZero
|
|||
|
oneIsSubj = SOne
|
|||
|
|
|||
|
IsGlobal = IsGlobal3
|
|||
|
isGlobal = isGlobal3
|
|||
|
zeroIsGlobal = GZero
|
|||
|
|
|||
|
|
|||
|
export Uninhabited (IsGlobal3 One) where uninhabited _ impossible
|
|||
|
|
|||
|
export Uninhabited (IsSubj3 Any) where uninhabited _ impossible
|