Q.S.T.Reduce ⇒ Q.Reduce and make it use Definition directly

This commit is contained in:
rhiannon morris 2023-02-19 18:22:53 +01:00
parent ae43c324c0
commit 7895fa37e5
8 changed files with 43 additions and 97 deletions

View file

@ -88,43 +88,3 @@ public export %inline
lookupElim : forall isGlobal. lookupElim : forall isGlobal.
Name -> Definitions' q isGlobal -> Maybe (Elim q d n) Name -> Definitions' q isGlobal -> Maybe (Elim q d n)
lookupElim x defs = toElim !(lookup x defs) lookupElim x defs = toElim !(lookup x defs)
parameters {0 isGlobal : _} (defs : Definitions' q isGlobal)
namespace Term
public export %inline
isRedex : Term q d n -> Bool
isRedex = isRedex $ \x => lookupElim x defs
public export
0 IsRedex, NotRedex : Pred $ Term q d n
IsRedex = So . isRedex
NotRedex = No . isRedex
namespace Elim
public export %inline
isRedex : Elim q d n -> Bool
isRedex = isRedex $ \x => lookupElim x defs
public export
0 IsRedex, NotRedex : Pred $ Elim q d n
IsRedex = So . isRedex
NotRedex = No . isRedex
public export
0 NonRedexElim, NonRedexTerm :
(q : Type) -> (d, n : Nat) -> {isGlobal : Pred q} ->
Definitions' q isGlobal -> Type
NonRedexElim q d n defs = Subset (Elim q d n) (NotRedex defs)
NonRedexTerm q d n defs = Subset (Term q d n) (NotRedex defs)
parameters {0 isGlobal : _} (defs : Definitions' q isGlobal)
namespace Term
export covering %inline
whnfD : Term q d n -> NonRedexTerm q d n defs
whnfD = whnf $ \x => lookupElim x defs
namespace Elim
export covering %inline
whnfD : Elim q d n -> NonRedexElim q d n defs
whnfD = whnf $ \x => lookupElim x defs

View file

@ -1,7 +1,5 @@
module Quox.Equal module Quox.Equal
import public Quox.Syntax
import public Quox.Definition
import public Quox.Typing import public Quox.Typing
import public Control.Monad.Either import public Control.Monad.Either
import public Control.Monad.Reader import public Control.Monad.Reader
@ -83,7 +81,7 @@ parameters (defs : Definitions' q g)
public export public export
isSubSing : Term q 0 n -> Bool isSubSing : Term q 0 n -> Bool
isSubSing ty = isSubSing ty =
let Element ty nc = whnfD defs ty in let Element ty nc = whnf defs ty in
case ty of case ty of
TYPE _ => False TYPE _ => False
Pi {res, _} => isSubSing res.term Pi {res, _} => isSubSing res.term
@ -118,9 +116,9 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
compare0 : TContext q 0 n -> (ty, s, t : Term q 0 n) -> m () compare0 : TContext q 0 n -> (ty, s, t : Term q 0 n) -> m ()
compare0 ctx ty s t = compare0 ctx ty s t =
wrapErr (WhileComparingT (MkTyContext new ctx) !mode ty s t) $ do wrapErr (WhileComparingT (MkTyContext new ctx) !mode ty s t) $ do
let Element ty nty = whnfD defs ty let Element ty nty = whnf defs ty
Element s ns = whnfD defs s Element s ns = whnf defs s
Element t nt = whnfD defs t Element t nt = whnf defs t
tty <- ensureTyCon ty tty <- ensureTyCon ty
compare0' ctx ty s t compare0' ctx ty s t
@ -192,8 +190,8 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
export covering %inline export covering %inline
compareType : TContext q 0 n -> (s, t : Term q 0 n) -> m () compareType : TContext q 0 n -> (s, t : Term q 0 n) -> m ()
compareType ctx s t = do compareType ctx s t = do
let Element s ns = whnfD defs s let Element s ns = whnf defs s
Element t nt = whnfD defs t Element t nt = whnf defs t
ts <- ensureTyCon s ts <- ensureTyCon s
tt <- ensureTyCon t tt <- ensureTyCon t
st <- either pure (const $ clashT (TYPE UAny) s t) $ st <- either pure (const $ clashT (TYPE UAny) s t) $
@ -245,7 +243,7 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
compareType' ctx (E e) (E f) = do compareType' ctx (E e) (E f) = do
-- no fanciness needed here cos anything other than a neutral -- no fanciness needed here cos anything other than a neutral
-- has been inlined by whnfD -- has been inlined by whnf
Elim.compare0 ctx e f Elim.compare0 ctx e f
||| performs the minimum work required to recompute the type of an elim. ||| performs the minimum work required to recompute the type of an elim.
@ -291,8 +289,8 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
compare0 : TContext q 0 n -> (e, f : Elim q 0 n) -> m () compare0 : TContext q 0 n -> (e, f : Elim q 0 n) -> m ()
compare0 ctx e f = compare0 ctx e f =
wrapErr (WhileComparingE (MkTyContext new ctx) !mode e f) $ do wrapErr (WhileComparingE (MkTyContext new ctx) !mode e f) $ do
let Element e ne = whnfD defs e let Element e ne = whnf defs e
Element f nf = whnfD defs f Element f nf = whnf defs f
-- [fixme] there is a better way to do this "isSubSing" stuff for sure -- [fixme] there is a better way to do this "isSubSing" stuff for sure
unless (isSubSing defs !(computeElimType ctx e ne)) $ unless (isSubSing defs !(computeElimType ctx e ne)) $
compare0' ctx e f ne nf compare0' ctx e f ne nf

View file

@ -1,8 +1,8 @@
module Quox.Syntax.Term.Reduce module Quox.Reduce
import Quox.No import Quox.No
import Quox.Syntax.Term.Base import Quox.Syntax
import Quox.Syntax.Term.Subst import Quox.Definition
import Data.Vect import Data.Vect
import Data.Maybe import Data.Maybe
@ -125,10 +125,6 @@ parameters (th : DSubst dfrom dto) (ph : TSubst q dto from to)
pushSubstsWith' e = (pushSubstsWith th ph e).fst pushSubstsWith' e = (pushSubstsWith th ph e).fst
public export 0
Lookup : TermLike
Lookup q d n = Name -> Maybe $ Elim q d n
public export %inline public export %inline
isLamHead : Elim {} -> Bool isLamHead : Elim {} -> Bool
isLamHead (Lam {} :# Pi {}) = True isLamHead (Lam {} :# Pi {}) = True
@ -154,12 +150,12 @@ isAnn : Elim {} -> Bool
isAnn (_ :# _) = True isAnn (_ :# _) = True
isAnn _ = False isAnn _ = False
parameters (g : Lookup q d n) parameters (defs : Definitions' q g)
mutual mutual
namespace Elim namespace Elim
public export public export
isRedex : Elim q d n -> Bool isRedex : Elim q d n -> Bool
isRedex (F x) = isJust $ g x isRedex (F x) {d, n} = isJust $ lookupElim x defs {d, n}
isRedex (B _) = False isRedex (B _) = False
isRedex (f :@ _) = isRedex f || isLamHead f isRedex (f :@ _) = isRedex f || isLamHead f
isRedex (CasePair {pair, _}) = isRedex pair || isPairHead pair isRedex (CasePair {pair, _}) = isRedex pair || isPairHead pair
@ -188,18 +184,19 @@ parameters (g : Lookup q d n)
IsRedex = So . isRedex IsRedex = So . isRedex
NotRedex = No . isRedex NotRedex = No . isRedex
public export parameters (q : Type) (d, n : Nat) {g : q -> Type} (defs : Definitions' q g)
0 NonRedexElim, NonRedexTerm : (q, d, n : _) -> Lookup q d n -> Type public export
NonRedexElim q d n g = Subset (Elim q d n) (NotRedex g) 0 NonRedexElim, NonRedexTerm : Type
NonRedexTerm q d n g = Subset (Term q d n) (NotRedex g) NonRedexElim = Subset (Elim q d n) (NotRedex defs)
NonRedexTerm = Subset (Term q d n) (NotRedex defs)
parameters (g : Lookup q d n) parameters (defs : Definitions' q g)
mutual mutual
namespace Elim namespace Elim
export covering export covering
whnf : Elim q d n -> NonRedexElim q d n g whnf : Elim q d n -> NonRedexElim q d n defs
whnf (F x) with (g x) proof eq whnf (F x) with (lookupElim x defs) proof eq
_ | Just y = whnf y _ | Just y = whnf y
_ | Nothing = Element (F x) $ rewrite eq in Ah _ | Nothing = Element (F x) $ rewrite eq in Ah
@ -251,7 +248,7 @@ parameters (g : Lookup q d n)
namespace Term namespace Term
export covering export covering
whnf : Term q d n -> NonRedexTerm q d n g whnf : Term q d n -> NonRedexTerm q d n defs
whnf t@(TYPE {}) = Element t Ah whnf t@(TYPE {}) = Element t Ah
whnf t@(Pi {}) = Element t Ah whnf t@(Pi {}) = Element t Ah
whnf t@(Lam {}) = Element t Ah whnf t@(Lam {}) = Element t Ah

View file

@ -3,5 +3,4 @@ module Quox.Syntax.Term
import public Quox.Syntax.Term.Base import public Quox.Syntax.Term.Base
import public Quox.Syntax.Term.Split import public Quox.Syntax.Term.Split
import public Quox.Syntax.Term.Subst import public Quox.Syntax.Term.Subst
import public Quox.Syntax.Term.Reduce
import public Quox.Syntax.Term.Pretty import public Quox.Syntax.Term.Pretty

View file

@ -1,14 +1,11 @@
module Quox.Typing module Quox.Typing
import public Quox.Syntax import public Quox.Syntax
import public Quox.Context
import public Quox.Definition import public Quox.Definition
import public Quox.Reduce
import Data.Nat
import public Data.SortedMap import public Data.SortedMap
import Control.Monad.Either import public Control.Monad.Either
import Control.Monad.Reader
import Control.Monad.State
import Generics.Derive import Generics.Derive
%hide TT.Name %hide TT.Name
@ -206,27 +203,27 @@ parameters {auto _ : HasErr q m} (defs : Definitions' q _)
export covering %inline export covering %inline
expectTYPE : Term q d n -> m Universe expectTYPE : Term q d n -> m Universe
expectTYPE s = expectTYPE s =
case fst $ whnfD defs s of case fst $ whnf defs s of
TYPE l => pure l TYPE l => pure l
_ => throwError $ ExpectedTYPE s _ => throwError $ ExpectedTYPE s
export covering %inline export covering %inline
expectPi : Term q d n -> m (q, Term q d n, ScopeTerm q d n) expectPi : Term q d n -> m (q, Term q d n, ScopeTerm q d n)
expectPi s = expectPi s =
case fst $ whnfD defs s of case fst $ whnf defs s of
Pi {qty, arg, res, _} => pure (qty, arg, res) Pi {qty, arg, res, _} => pure (qty, arg, res)
_ => throwError $ ExpectedPi s _ => throwError $ ExpectedPi s
export covering %inline export covering %inline
expectSig : Term q d n -> m (Term q d n, ScopeTerm q d n) expectSig : Term q d n -> m (Term q d n, ScopeTerm q d n)
expectSig s = expectSig s =
case fst $ whnfD defs s of case fst $ whnf defs s of
Sig {fst, snd, _} => pure (fst, snd) Sig {fst, snd, _} => pure (fst, snd)
_ => throwError $ ExpectedSig s _ => throwError $ ExpectedSig s
export covering %inline export covering %inline
expectEq : Term q d n -> m (DScopeTerm q d n, Term q d n, Term q d n) expectEq : Term q d n -> m (DScopeTerm q d n, Term q d n, Term q d n)
expectEq s = expectEq s =
case fst $ whnfD defs s of case fst $ whnf defs s of
Eq {ty, l, r, _} => pure (ty, l, r) Eq {ty, l, r, _} => pure (ty, l, r)
_ => throwError $ ExpectedEq s _ => throwError $ ExpectedEq s

View file

@ -22,12 +22,12 @@ modules =
Quox.Syntax.Term, Quox.Syntax.Term,
Quox.Syntax.Term.Base, Quox.Syntax.Term.Base,
Quox.Syntax.Term.Pretty, Quox.Syntax.Term.Pretty,
Quox.Syntax.Term.Reduce,
Quox.Syntax.Term.Split, Quox.Syntax.Term.Split,
Quox.Syntax.Term.Subst, Quox.Syntax.Term.Subst,
Quox.Syntax.Universe, Quox.Syntax.Universe,
Quox.Syntax.Var, Quox.Syntax.Var,
Quox.Definition, Quox.Definition,
Quox.Reduce,
Quox.Context, Quox.Context,
Quox.Equal, Quox.Equal,
Quox.Name, Quox.Name,

View file

@ -8,35 +8,30 @@ import TAP
testWhnf : Eq b => Show b => (a -> (Subset b _)) -> String -> a -> b -> Test testWhnf : Eq b => Show b => (a -> (Subset b _)) -> String -> a -> b -> Test
testWhnf whnf label from to = test "\{label} (whnf)" $ testWhnf whnf label from to = test "\{label} (whnf)" $ do
let result = fst (whnf from) in let result = fst (whnf from)
if result == to unless (result == to) $
then Right () Left [("exp", to), ("got", result)] {a = List (String, b)}
else with Prelude.(::)
Left [("expected", to), ("received", result)]
testNoStep : Eq a => Show a => (a -> (Subset a _)) -> String -> a -> Test testNoStep : Eq a => Show a => (a -> (Subset a _)) -> String -> a -> Test
testNoStep whnf label e = test "\{label} (no step)" $ testNoStep whnf label e = test "\{label} (no step)" $ do
let result = fst (whnf e) in let result = fst (whnf e)
if result == e unless (result == e) $ Left [("reduced", result)] {a = List (String, a)}
then Right ()
else with Prelude.(::)
Left [("reduced", result)]
parameters {default empty defs : Definitions Three} {default 0 d, n : Nat} parameters {default empty defs : Definitions Three} {default 0 d, n : Nat}
testWhnfT : String -> Term Three d n -> Term Three d n -> Test testWhnfT : String -> Term Three d n -> Term Three d n -> Test
testWhnfT = testWhnf (whnfD defs) testWhnfT = testWhnf (whnf defs)
testWhnfE : String -> Elim Three d n -> Elim Three d n -> Test testWhnfE : String -> Elim Three d n -> Elim Three d n -> Test
testWhnfE = testWhnf (whnfD defs) testWhnfE = testWhnf (whnf defs)
testNoStepE : String -> Elim Three d n -> Test testNoStepE : String -> Elim Three d n -> Test
testNoStepE = testNoStep (whnfD defs) testNoStepE = testNoStep (whnf defs)
testNoStepT : String -> Term Three d n -> Test testNoStepT : String -> Term Three d n -> Test
testNoStepT = testNoStep (whnfD defs) testNoStepT = testNoStep (whnf defs)
tests = "whnf" :- [ tests = "whnf" :- [

View file

@ -51,7 +51,7 @@ PrettyHL q => ToInfo (Error q) where
("right", prettyStr True q)] ("right", prettyStr True q)]
toInfo (NotType ty) = toInfo (NotType ty) =
[("type", "NotType"), [("type", "NotType"),
("actual", prettyStr True ty)] ("got", prettyStr True ty)]
toInfo (WrongType ty s t) = toInfo (WrongType ty s t) =
[("type", "WrongType"), [("type", "WrongType"),
("ty", prettyStr True ty), ("ty", prettyStr True ty),