unfold definitions in equality checking, plus cleanup
This commit is contained in:
parent
44778825c2
commit
8a55cc9581
4 changed files with 137 additions and 81 deletions
|
@ -2,52 +2,71 @@ module Quox.Equal
|
|||
|
||||
import public Quox.Syntax
|
||||
import public Quox.Definition
|
||||
import Control.Monad.Either
|
||||
import Generics.Derive
|
||||
import public Quox.Typing
|
||||
import public Control.Monad.Either
|
||||
import public Control.Monad.Reader
|
||||
import Data.Maybe
|
||||
|
||||
%default total
|
||||
%language ElabReflection
|
||||
|
||||
|
||||
public export
|
||||
data Mode = Equal | Sub
|
||||
%runElab derive "Mode" [Generic, Meta, Eq, Ord, DecEq, Show]
|
||||
|
||||
public export
|
||||
data Error
|
||||
= ClashT Mode (Term d n) (Term d n)
|
||||
| ClashU Mode Universe Universe
|
||||
| ClashQ Qty Qty
|
||||
|
||||
private %inline
|
||||
ClashE : Mode -> Elim d n -> Elim d n -> Error
|
||||
ClashE : EqMode -> Elim d n -> Elim d n -> Error
|
||||
ClashE mode = ClashT mode `on` E
|
||||
|
||||
parameters {auto _ : MonadError Error m}
|
||||
public export
|
||||
CanEq : (Type -> Type) -> Type
|
||||
CanEq m = (MonadError Error m, MonadReader Definitions m)
|
||||
|
||||
|
||||
parameters {auto _ : CanEq m}
|
||||
private %inline
|
||||
clashT : Mode -> Term d n -> Term d n -> m a
|
||||
clashT : EqMode -> Term d n -> Term d n -> m a
|
||||
clashT mode = throwError .: ClashT mode
|
||||
|
||||
private %inline
|
||||
clashE : Mode -> Elim d n -> Elim d n -> m a
|
||||
clashE : EqMode -> Elim d n -> Elim d n -> m a
|
||||
clashE mode = throwError .: ClashE mode
|
||||
|
||||
private %inline
|
||||
defE : Name -> m (Maybe (Elim d n))
|
||||
defE x = asks $ \defs => do
|
||||
g <- lookup x defs
|
||||
pure $ (!g.term).def :# g.type.def
|
||||
|
||||
private %inline
|
||||
defT : Name -> m (Maybe (Term d n))
|
||||
defT x = map E <$> defE x
|
||||
|
||||
export %inline
|
||||
compareU' : EqMode -> Universe -> Universe -> Bool
|
||||
compareU' = \case Equal => (==); Sub => (<=)
|
||||
|
||||
export %inline
|
||||
compareU : EqMode -> Universe -> Universe -> m ()
|
||||
compareU mode k l = unless (compareU' mode k l) $
|
||||
throwError $ ClashU mode k l
|
||||
|
||||
mutual
|
||||
private covering
|
||||
compareTN' : Mode ->
|
||||
compareTN' : EqMode ->
|
||||
(s, t : Term 0 n) ->
|
||||
(0 _ : NotRedexT s) -> (0 _ : NotRedexT t) -> m ()
|
||||
|
||||
compareTN' mode (TYPE k) (TYPE l) _ _ =
|
||||
case mode of
|
||||
Equal => unless (k == l) $ throwError $ ClashU Equal k l
|
||||
Sub => unless (k <= l) $ throwError $ ClashU Sub k l
|
||||
compareTN' mode (E e) (E f) ps pt = compareE0 mode e f
|
||||
-- if either term is a def, try to unfold it
|
||||
compareTN' mode s@(E (F x)) t ps pt = do
|
||||
Just s' <- defT x | Nothing => clashT mode s t
|
||||
compareT0 mode s' t
|
||||
compareTN' mode s t@(E (F y)) ps pt = do
|
||||
Just t' <- defT y | Nothing => clashT mode s t
|
||||
compareT0 mode s t'
|
||||
compareTN' mode s@(E _) t _ _ = clashT mode s t
|
||||
|
||||
compareTN' mode (TYPE k) (TYPE l) _ _ = compareU mode k l
|
||||
compareTN' mode s@(TYPE _) t _ _ = clashT mode s t
|
||||
|
||||
compareTN' mode (Pi qty1 _ arg1 res1) (Pi qty2 _ arg2 res2) _ _ = do
|
||||
-- [todo] this should probably always be ==, right..?
|
||||
unless (qty1 == qty2) $ throwError $ ClashQ qty1 qty2
|
||||
compareT0 mode arg2 arg1 -- reversed for contravariant Sub
|
||||
compareT0 mode arg2 arg1 -- reversed for contravariant domain
|
||||
compareTS0 mode res1 res2
|
||||
compareTN' mode s@(Pi {}) t _ _ = clashT mode s t
|
||||
|
||||
|
@ -56,20 +75,25 @@ parameters {auto _ : MonadError Error m}
|
|||
compareTS0 Equal body1 body2
|
||||
compareTN' mode s@(Lam {}) t _ _ = clashT mode s t
|
||||
|
||||
compareTN' mode (E e) (E f) ps pt = compareE0 mode e f
|
||||
compareTN' mode s@(E _) t _ _ = clashT mode s t
|
||||
|
||||
compareTN' _ (CloT {}) _ ps _ = void $ ps IsCloT
|
||||
compareTN' _ (DCloT {}) _ ps _ = void $ ps IsDCloT
|
||||
|
||||
private covering
|
||||
compareEN' : Mode ->
|
||||
compareEN' : EqMode ->
|
||||
(e, f : Elim 0 n) ->
|
||||
(0 _ : NotRedexE e) -> (0 _ : NotRedexE f) -> m ()
|
||||
|
||||
compareEN' mode e@(F x) f@(F y) _ _ =
|
||||
unless (x == y) $ clashE mode e f
|
||||
compareEN' mode e@(F _) f _ _ = clashE mode e f
|
||||
if x == y then pure () else
|
||||
case (!(defE x), !(defE y)) of
|
||||
(Nothing, Nothing) => clashE mode e f
|
||||
(s', t') => compareE0 mode (fromMaybe e s') (fromMaybe f t')
|
||||
compareEN' mode e@(F x) f _ _ = do
|
||||
Just e' <- defE x | Nothing => clashE mode e f
|
||||
compareE0 mode e' f
|
||||
compareEN' mode e f@(F y) _ _ = do
|
||||
Just f' <- defE y | Nothing => clashE mode e f
|
||||
compareE0 mode e f'
|
||||
|
||||
compareEN' mode e@(B i) f@(B j) _ _ =
|
||||
unless (i == j) $ clashE mode e f
|
||||
|
@ -93,35 +117,35 @@ parameters {auto _ : MonadError Error m}
|
|||
|
||||
|
||||
private covering %inline
|
||||
compareTN : Mode -> NonRedexTerm 0 n -> NonRedexTerm 0 n -> m ()
|
||||
compareTN : EqMode -> NonRedexTerm 0 n -> NonRedexTerm 0 n -> m ()
|
||||
compareTN mode s t = compareTN' mode s.fst t.fst s.snd t.snd
|
||||
|
||||
private covering %inline
|
||||
compareEN : Mode -> NonRedexElim 0 n -> NonRedexElim 0 n -> m ()
|
||||
compareEN : EqMode -> NonRedexElim 0 n -> NonRedexElim 0 n -> m ()
|
||||
compareEN mode e f = compareEN' mode e.fst f.fst e.snd f.snd
|
||||
|
||||
|
||||
export covering %inline
|
||||
compareT : Mode -> DimEq d -> Term d n -> Term d n -> m ()
|
||||
compareT : EqMode -> DimEq d -> Term d n -> Term d n -> m ()
|
||||
compareT mode eqs s t =
|
||||
for_ (splits eqs) $ \th => compareT0 mode (s /// th) (t /// th)
|
||||
|
||||
export covering %inline
|
||||
compareE : Mode -> DimEq d -> Elim d n -> Elim d n -> m ()
|
||||
compareE : EqMode -> DimEq d -> Elim d n -> Elim d n -> m ()
|
||||
compareE mode eqs e f =
|
||||
for_ (splits eqs) $ \th => compareE0 mode (e /// th) (f /// th)
|
||||
|
||||
|
||||
export covering %inline
|
||||
compareT0 : Mode -> Term 0 n -> Term 0 n -> m ()
|
||||
compareT0 : EqMode -> Term 0 n -> Term 0 n -> m ()
|
||||
compareT0 mode s t = compareTN mode (whnfT s) (whnfT t)
|
||||
|
||||
export covering %inline
|
||||
compareE0 : Mode -> Elim 0 n -> Elim 0 n -> m ()
|
||||
compareE0 : EqMode -> Elim 0 n -> Elim 0 n -> m ()
|
||||
compareE0 mode e f = compareEN mode (whnfE e) (whnfE f)
|
||||
|
||||
export covering %inline
|
||||
compareTS0 : Mode -> ScopeTerm 0 n -> ScopeTerm 0 n -> m ()
|
||||
compareTS0 : EqMode -> ScopeTerm 0 n -> ScopeTerm 0 n -> m ()
|
||||
compareTS0 mode (TUnused body1) (TUnused body2) =
|
||||
compareT0 mode body1 body2
|
||||
compareTS0 mode body1 body2 =
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue