quox/lib/Quox/Syntax/Dim.idr

131 lines
2.9 KiB
Idris
Raw Normal View History

2021-07-20 16:05:19 -04:00
module Quox.Syntax.Dim
import Quox.Syntax.Var
import Quox.Syntax.Subst
import Quox.Pretty
2022-02-26 20:17:42 -05:00
import Decidable.Equality
import Control.Function
2022-05-02 16:40:28 -04:00
import Generics.Derive
2022-02-26 20:17:42 -05:00
2021-07-20 16:05:19 -04:00
%default total
2022-05-02 16:40:28 -04:00
%language ElabReflection
%hide SOP.from; %hide SOP.to
2021-07-20 16:05:19 -04:00
public export
data DimConst = Zero | One
%name DimConst e
2022-05-02 16:40:28 -04:00
%runElab derive "DimConst" [Generic, Meta, Eq, Ord, DecEq, Show]
2023-01-26 13:54:46 -05:00
||| `ends l r e` returns `l` if `e` is `Zero`, or `r` if it is `One`.
public export
2023-01-26 13:54:46 -05:00
ends : a -> a -> DimConst -> a
ends l r Zero = l
ends l r One = r
2021-07-20 16:05:19 -04:00
public export
data Dim : Nat -> Type where
K : DimConst -> Dim d
B : Var d -> Dim d
2021-07-20 16:05:19 -04:00
%name Dim.Dim p, q
2022-05-02 16:40:28 -04:00
private %inline
drepr : Dim n -> Either DimConst (Var n)
drepr (K k) = Left k
drepr (B x) = Right x
2021-07-20 16:05:19 -04:00
export Eq (Dim n) where (==) = (==) `on` drepr
export Ord (Dim n) where compare = compare `on` drepr
2022-05-06 18:57:23 -04:00
export
Show (Dim n) where
show (K k) = showCon App "K" $ show k
show (B i) = showCon App "B" $ show i
export
PrettyHL DimConst where
prettyM Zero = hl Dim <$> ifUnicode "𝟬" "0"
prettyM One = hl Dim <$> ifUnicode "𝟭" "1"
2021-07-20 16:05:19 -04:00
export
PrettyHL (Dim n) where
prettyM (K e) = prettyM e
prettyM (B i) = prettyVar DVar DVarErr (!ask).dnames i
2021-07-20 16:05:19 -04:00
2023-01-26 13:54:46 -05:00
||| `endsOr l r x e` returns:
||| - `l` if `p` is `K Zero`;
||| - `r` if `p` is `K One`;
||| - `x` otherwise.
public export
endsOr : a -> a -> Lazy a -> Dim n -> a
endsOr l r x (K e) = ends l r e
endsOr l r x (B _) = x
2022-04-06 14:32:56 -04:00
public export %inline
toConst : Dim 0 -> DimConst
toConst (K e) = e
2021-07-20 16:05:19 -04:00
public export
2021-09-03 10:34:57 -04:00
DSubst : Nat -> Nat -> Type
2021-07-20 16:05:19 -04:00
DSubst = Subst Dim
export %inline
prettyDSubst : Pretty.HasEnv m => DSubst from to -> m (Doc HL)
2021-07-20 16:05:19 -04:00
prettyDSubst th =
2021-09-03 10:31:53 -04:00
prettySubstM prettyM (!ask).dnames DVar
2021-07-20 16:05:19 -04:00
!(ifUnicode "" "<") !(ifUnicode "" ">") th
2023-01-20 20:34:28 -05:00
public export FromVar Dim where fromVar = B
2021-07-20 16:05:19 -04:00
2021-12-23 13:01:39 -05:00
export
CanShift Dim where
K e // _ = K e
B i // by = B (i // by)
2021-07-20 16:05:19 -04:00
export
CanSubst Dim Dim where
K e // _ = K e
B i // th = th !! i
2022-02-26 20:17:42 -05:00
export Uninhabited (Zero = One) where uninhabited _ impossible
export Uninhabited (One = Zero) where uninhabited _ impossible
export Uninhabited (B i = K e) where uninhabited _ impossible
export Uninhabited (K e = B i) where uninhabited _ impossible
2022-05-02 16:40:28 -04:00
public export %inline Injective Dim.B where injective Refl = Refl
public export %inline Injective Dim.K where injective Refl = Refl
2022-02-26 20:17:42 -05:00
public export
DecEq DimConst where
decEq Zero Zero = Yes Refl
decEq Zero One = No absurd
decEq One Zero = No absurd
decEq One One = Yes Refl
public export
DecEq (Dim d) where
decEq (K e) (K f) with (decEq e f)
_ | Yes prf = Yes $ cong K prf
_ | No contra = No $ contra . injective
decEq (K e) (B j) = No absurd
decEq (B i) (K f) = No absurd
decEq (B i) (B j) with (decEq i j)
_ | Yes prf = Yes $ cong B prf
_ | No contra = No $ contra . injective
2023-01-20 20:34:28 -05:00
||| abbreviation for a bound variable like `BV 4` instead of
||| `B (VS (VS (VS (VS VZ))))`
public export %inline
BV : (i : Nat) -> (0 _ : LT i d) => Dim d
BV i = B $ V i