quox/lib/Quox/Syntax/Dim.idr

150 lines
3.4 KiB
Idris

module Quox.Syntax.Dim
import Quox.Thin
import Quox.Syntax.Var
import Quox.Syntax.Subst
import Quox.Pretty
import Quox.Name
import Quox.Loc
import Quox.Context
import Decidable.Equality
import Control.Function
import Derive.Prelude
import Data.DPair
import Data.SnocVect
%default total
%language ElabReflection
public export
data DimConst = Zero | One
%name DimConst e
%runElab derive "DimConst" [Eq, Ord, Show]
||| `ends l r e` returns `l` if `e` is `Zero`, or `r` if it is `One`.
public export
ends : Lazy a -> Lazy a -> DimConst -> a
ends l r Zero = l
ends l r One = r
export Uninhabited (Zero = One) where uninhabited _ impossible
export Uninhabited (One = Zero) where uninhabited _ impossible
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
data Dim : Nat -> Type where
K : DimConst -> Loc -> Dim 0
B : Loc -> Dim 1
%name Dim.Dim p, q
%runElab deriveIndexed "Dim" [Eq, Ord, Show]
public export
DimT : Nat -> Type
DimT = Thinned Dim
public export %inline
KT : DimConst -> Loc -> DimT d
KT e loc = Th zero $ K e loc
||| `endsOr l r x p` returns `ends l r ε` if `p` is a constant ε, and
||| `x` otherwise.
public export
endsOr : Lazy a -> Lazy a -> Lazy a -> Dim n -> a
endsOr l r x (K e _) = ends l r e
endsOr l r x (B _) = x
export
Located (Dim d) where
(K _ loc).loc = loc
(B loc).loc = loc
export
Relocatable (Dim d) where
setLoc loc (K e _) = K e loc
setLoc loc (B _) = B loc
parameters {opts : LayoutOpts}
export
prettyDimConst : DimConst -> Eff Pretty (Doc opts)
prettyDimConst = hl Dim . text . ends "0" "1"
export
prettyDim : {d : Nat} -> BContext d -> DimT d -> Eff Pretty (Doc opts)
prettyDim names (Th _ (K e _)) = prettyDimConst e
prettyDim names (Th i (B _)) = prettyDBind $ names !!! i.fin
public export %inline
toConst : Dim 0 -> DimConst
toConst (K e _) = e
public export
DSubst : Nat -> Nat -> Type
DSubst = Subst Dim
-- public export FromVar Dim where fromVarLoc = B
-- export
-- CanShift Dim where
-- K e loc // _ = K e loc
-- B i loc // by = B (i // by) loc
export %inline FromVar Dim where var = B
export %inline
CanSubstSelf Dim where
Th _ (K e loc) // _ = KT e loc
Th i (B loc) // th = get th i.fin
export Uninhabited (B loc1 = K e loc2) where uninhabited _ impossible
export Uninhabited (K e loc1 = B loc2) where uninhabited _ impossible
-- public export
-- data Eqv : Dim d1 -> Dim d2 -> Type where
-- EK : K e _ `Eqv` K e _
-- EB : i `Eqv` j -> B i _ `Eqv` B j _
-- export Uninhabited (K e l1 `Eqv` B i l2) where uninhabited _ impossible
-- export Uninhabited (B i l1 `Eqv` K e l2) where uninhabited _ impossible
-- export
-- injectiveB : B i loc1 `Eqv` B j loc2 -> i `Eqv` j
-- injectiveB (EB e) = e
-- export
-- injectiveK : K e loc1 `Eqv` K f loc2 -> e = f
-- injectiveK EK = Refl
-- public export
-- decEqv : Dec2 Dim.Eqv
-- decEqv (K e _) (K f _) = case decEq e f of
-- Yes Refl => Yes EK
-- No n => No $ n . injectiveK
-- decEqv (B i _) (B j _) = case decEqv i j of
-- Yes y => Yes $ EB y
-- No n => No $ \(EB y) => n y
-- decEqv (B _ _) (K _ _) = No absurd
-- decEqv (K _ _) (B _ _) = No absurd
||| abbreviation for a bound variable like `BV 4` instead of
||| `B (VS (VS (VS (VS VZ))))`
public export %inline
BV : (i : Fin d) -> (loc : Loc) -> DimT d
BV i loc = Th (one' i) $ B loc