diff --git a/exe/Main.idr b/exe/Main.idr index 45a377d..177f438 100644 --- a/exe/Main.idr +++ b/exe/Main.idr @@ -164,7 +164,7 @@ private liftFromParser : Eff FromParserIO a -> Eff CompileStop a liftFromParser act = runEff act $ with Union.(::) - [handleExcept (\err => throw $ FromParserError err), + [handleExcept $ \err => throw $ FromParserError err, handleStateIORef !(asksAt STATE defs), handleStateIORef !(asksAt STATE ns), handleStateIORef !(asksAt STATE suf), @@ -174,8 +174,7 @@ private liftErase : Q.Definitions -> Eff Erase a -> Eff CompileStop a liftErase defs act = runEff act - [\case Err e => throw $ EraseError e, - \case Ask => pure defs, + [handleExcept $ \err => throw $ EraseError err, handleStateIORef !(asksAt STATE suf)] private @@ -207,7 +206,7 @@ processFile file = withEarlyStop $ do traverse (uncurry Q.prettyDef) defList let defs = SortedMap.fromList defList erased <- liftErase defs $ - traverse (\(x, d) => (x,) <$> eraseDef x d) defList + traverse (\(x, d) => (x,) <$> eraseDef defs x d) defList outputDocStopIf Erase $ traverse (uncurry U.prettyDef) erased (scheme, mains) <- liftScheme $ map catMaybes $ diff --git a/lib/Quox/Equal.idr b/lib/Quox/Equal.idr index 7055181..9a69ffb 100644 --- a/lib/Quox/Equal.idr +++ b/lib/Quox/Equal.idr @@ -170,12 +170,19 @@ compareType : Definitions -> EqContext n -> (s, t : Term 0 n) -> Eff EqualInner () +private +0 NotRedexEq : {isRedex : RedexTest tm} -> CanWhnf tm isRedex => + Definitions -> EqContext n -> SQty -> Pred (tm 0 n) +NotRedexEq defs ctx sg t = NotRedex defs (toWhnfContext ctx) sg t + namespace Term private covering - compare0' : (defs : Definitions) -> EqContext n -> (sg : SQty) -> + compare0' : (defs : Definitions) -> (ctx : EqContext n) -> (sg : SQty) -> (ty, s, t : Term 0 n) -> - (0 _ : NotRedex defs SZero ty) => (0 _ : So (isTyConE ty)) => - (0 _ : NotRedex defs sg s) => (0 _ : NotRedex defs sg t) => + (0 _ : NotRedexEq defs ctx SZero ty) => + (0 _ : So (isTyConE ty)) => + (0 _ : NotRedexEq defs ctx sg s) => + (0 _ : NotRedexEq defs ctx sg t) => Eff EqualInner () compare0' defs ctx sg (TYPE {}) s t = compareType defs ctx s t @@ -350,9 +357,10 @@ namespace Term private covering -compareType' : (defs : Definitions) -> EqContext n -> (s, t : Term 0 n) -> - (0 _ : NotRedex defs SZero s) => (0 _ : So (isTyConE s)) => - (0 _ : NotRedex defs SZero t) => (0 _ : So (isTyConE t)) => +compareType' : (defs : Definitions) -> (ctx : EqContext n) -> + (s, t : Term 0 n) -> + (0 _ : NotRedexEq defs ctx SZero s) => (0 _ : So (isTyConE s)) => + (0 _ : NotRedexEq defs ctx SZero t) => (0 _ : So (isTyConE t)) => (0 _ : So (sameTyCon s t)) => Eff EqualInner () -- equality is the same as subtyping, except with the @@ -477,8 +485,9 @@ namespace Elim EqualElim = InnerErrEff :: EqualInner private covering - computeElimTypeE : (defs : Definitions) -> EqContext n -> (sg : SQty) -> - (e : Elim 0 n) -> (0 ne : NotRedex defs sg e) => + computeElimTypeE : (defs : Definitions) -> (ctx : EqContext n) -> + (sg : SQty) -> + (e : Elim 0 n) -> (0 ne : NotRedexEq defs ctx sg e) => Eff EqualElim (Term 0 n) computeElimTypeE defs ectx sg e = lift $ computeElimType defs (toWhnfContext ectx) sg e @@ -492,8 +501,8 @@ namespace Elim try act = lift $ catch putError $ lift act {fs' = EqualElim} private covering %inline - clashE : (defs : Definitions) -> EqContext n -> (sg : SQty) -> - (e, f : Elim 0 n) -> (0 nf : NotRedex defs sg f) => + clashE : (defs : Definitions) -> (ctx : EqContext n) -> (sg : SQty) -> + (e, f : Elim 0 n) -> (0 nf : NotRedexEq defs ctx sg f) => Eff EqualElim (Term 0 n) clashE defs ctx sg e f = do putError $ ClashE e.loc ctx !mode e f @@ -522,9 +531,10 @@ namespace Elim (e, f : Elim 0 n) -> Eff EqualElim (Term 0 n) private covering - compare0Inner' : (defs : Definitions) -> EqContext n -> (sg : SQty) -> + compare0Inner' : (defs : Definitions) -> (ctx : EqContext n) -> (sg : SQty) -> (e, f : Elim 0 n) -> - (0 ne : NotRedex defs sg e) -> (0 nf : NotRedex defs sg f) -> + (0 ne : NotRedexEq defs ctx sg e) -> + (0 nf : NotRedexEq defs ctx sg f) -> Eff EqualElim (Term 0 n) compare0Inner' defs ctx sg e@(F {}) f _ _ = do diff --git a/lib/Quox/Typechecker.idr b/lib/Quox/Typechecker.idr index 13052b1..53677a5 100644 --- a/lib/Quox/Typechecker.idr +++ b/lib/Quox/Typechecker.idr @@ -215,8 +215,8 @@ mutual check' ctx sg (Let qty rhs body loc) ty = do eres <- inferC ctx (subjMult sg qty) rhs - qout <- checkC (extendTy (qty * sg.qty) body.name eres.type ctx) sg - body.term (weakT 1 ty) + qout <- checkC (extendTyLet (qty * sg.qty) body.name eres.type (E rhs) ctx) + sg body.term (weakT 1 ty) >>= popQ loc qty pure $ qty * eres.qout + qout @@ -338,8 +338,8 @@ mutual pure $ lookupBound sg.qty i ctx.tctx where lookupBound : forall n. Qty -> Var n -> TContext d n -> InferResult' d n - lookupBound pi VZ (ctx :< type) = - InfRes {type = weakT 1 type, qout = zeroFor ctx :< pi} + lookupBound pi VZ (ctx :< var) = + InfRes {type = weakT 1 var.type, qout = zeroFor ctx :< pi} lookupBound pi (VS i) (ctx :< _) = let InfRes {type, qout} = lookupBound pi i ctx in InfRes {type = weakT 1 type, qout = qout :< Zero} diff --git a/lib/Quox/Typing.idr b/lib/Quox/Typing.idr index 31c059f..5d62e8a 100644 --- a/lib/Quox/Typing.idr +++ b/lib/Quox/Typing.idr @@ -70,7 +70,7 @@ parameters (defs : Definitions) {auto _ : (Has ErrorEff fs, Has NameGen fs)} parameters (ctx : TyContext d n) (sg : SQty) (loc : Loc) export covering whnf : {0 isRedex : RedexTest tm} -> CanWhnf tm isRedex => - tm d n -> Eff fs (NonRedex tm d n defs sg) + tm d n -> Eff fs (NonRedex tm d n defs ? sg) whnf tm = do let Val n = ctx.termLen; Val d = ctx.dimLen res <- lift $ runExcept $ whnf defs (toWhnfContext ctx) sg tm @@ -120,7 +120,7 @@ parameters (defs : Definitions) {auto _ : (Has ErrorEff fs, Has NameGen fs)} parameters (ctx : EqContext n) (sg : SQty) (loc : Loc) export covering whnf : {0 isRedex : RedexTest tm} -> CanWhnf tm isRedex => - tm 0 n -> Eff fs (NonRedex tm 0 n defs sg) + tm 0 n -> Eff fs (NonRedex tm 0 n defs ? sg) whnf tm = do res <- lift $ runExcept $ whnf defs (toWhnfContext ctx) sg tm rethrow res diff --git a/lib/Quox/Typing/Context.idr b/lib/Quox/Typing/Context.idr index a1dd49a..f660a7c 100644 --- a/lib/Quox/Typing/Context.idr +++ b/lib/Quox/Typing/Context.idr @@ -14,9 +14,27 @@ public export QContext : Nat -> Type QContext = Context' Qty +public export +record LocalVar d n where + constructor MkLocal + type : Term d n + term : Maybe (Term d n) -- if from a `let` +%runElab deriveIndexed "LocalVar" [Show] + +export +CanShift (LocalVar d) where + l // by = {type $= (// by), term $= map (// by)} l + +namespace LocalVar + subD : DSubst d1 d2 -> LocalVar d1 n -> LocalVar d2 n + subD th = {type $= (// th), term $= map (// th)} + + weakD : LocalVar d n -> LocalVar (S d) n + weakD = subD $ shift 1 + public export TContext : TermLike -TContext d = Context (Term d) +TContext d = Context (\n => LocalVar d n) public export QOutput : Nat -> Type @@ -67,10 +85,6 @@ record WhnfContext d n where %runElab deriveIndexed "WhnfContext" [Show] namespace TContext - export %inline - pushD : TContext d n -> TContext (S d) n - pushD tel = map (// shift 1) tel - export %inline zeroFor : Context tm n -> QOutput n zeroFor ctx = Zero <$ ctx @@ -89,6 +103,14 @@ public export CtxExtension0 : Nat -> Nat -> Nat -> Type CtxExtension0 d = Telescope ((BindName,) . Term d) +public export +CtxExtensionLet : Nat -> Nat -> Nat -> Type +CtxExtensionLet d = Telescope ((Qty, BindName,) . LocalVar d) + +public export +CtxExtensionLet0 : Nat -> Nat -> Nat -> Type +CtxExtensionLet0 d = Telescope ((BindName,) . LocalVar d) + namespace TyContext public export %inline empty : TyContext 0 0 @@ -100,21 +122,34 @@ namespace TyContext null ctx = null ctx.dnames && null ctx.tnames export %inline - extendTyN : CtxExtension d n1 n2 -> TyContext d n1 -> TyContext d n2 - extendTyN xss (MkTyContext {termLen, dctx, dnames, tctx, tnames, qtys}) = - let (qs, xs, ss) = unzip3 xss in + extendTyLetN : CtxExtensionLet d n1 n2 -> TyContext d n1 -> TyContext d n2 + extendTyLetN xss (MkTyContext {termLen, dctx, dnames, tctx, tnames, qtys}) = + let (qs, xs, ls) = unzip3 xss in MkTyContext { dctx, dnames, termLen = extendLen xss termLen, - tctx = tctx . ss, + tctx = tctx . ls, tnames = tnames . xs, qtys = qtys . qs } + export %inline + extendTyN : CtxExtension d n1 n2 -> TyContext d n1 -> TyContext d n2 + extendTyN = extendTyLetN . map (\(q, x, s) => (q, x, MkLocal s Nothing)) + + export %inline + extendTyLetN0 : CtxExtensionLet0 d n1 n2 -> TyContext d n1 -> TyContext d n2 + extendTyLetN0 xss = extendTyLetN (map (Zero,) xss) + export %inline extendTyN0 : CtxExtension0 d n1 n2 -> TyContext d n1 -> TyContext d n2 extendTyN0 xss = extendTyN (map (Zero,) xss) + export %inline + extendTyLet : Qty -> BindName -> Term d n -> Term d n -> + TyContext d n -> TyContext d (S n) + extendTyLet q x s e = extendTyLetN [< (q, x, MkLocal s (Just e))] + export %inline extendTy : Qty -> BindName -> Term d n -> TyContext d n -> TyContext d (S n) extendTy q x s = extendTyN [< (q, x, s)] @@ -130,7 +165,7 @@ namespace TyContext dctx = dctx : EqContext n1 -> EqContext n2 - extendTyN xss (MkEqContext {termLen, dassign, dnames, tctx, tnames, qtys}) = - let (qs, xs, ss) = unzip3 xss in + extendTyLetN : CtxExtensionLet 0 n1 n2 -> EqContext n1 -> EqContext n2 + extendTyLetN xss (MkEqContext {termLen, dassign, dnames, tctx, tnames, qtys}) = + let (qs, xs, ls) = unzip3 xss in MkEqContext { termLen = extendLen xss termLen, - tctx = tctx . ss, + tctx = tctx . ls, tnames = tnames . xs, qtys = qtys . qs, dassign, dnames } + export %inline + extendTyN : CtxExtension 0 n1 n2 -> EqContext n1 -> EqContext n2 + extendTyN = extendTyLetN . map (\(q, x, s) => (q, x, MkLocal s Nothing)) + + export %inline + extendTyLetN0 : CtxExtensionLet0 0 n1 n2 -> EqContext n1 -> EqContext n2 + extendTyLetN0 xss = extendTyLetN (map (Zero,) xss) + export %inline extendTyN0 : CtxExtension0 0 n1 n2 -> EqContext n1 -> EqContext n2 extendTyN0 xss = extendTyN (map (Zero,) xss) + export %inline + extendTyLet : Qty -> BindName -> Term 0 n -> Term 0 n -> + EqContext n -> EqContext (S n) + extendTyLet q x s e = extendTyLetN [< (q, x, MkLocal s (Just e))] + export %inline extendTy : Qty -> BindName -> Term 0 n -> EqContext n -> EqContext (S n) extendTy q x s = extendTyN [< (q, x, s)] @@ -225,7 +273,7 @@ namespace EqContext toTyContext (MkEqContext {dimLen, dassign, dnames, tctx, tnames, qtys}) = MkTyContext { dctx = fromGround dassign, - tctx = map (// shift0 dimLen) tctx, + tctx = map (subD $ shift0 dimLen) tctx, dnames, tnames, qtys } @@ -252,7 +300,7 @@ namespace WhnfContext MkWhnfContext { dimLen = [|Val s + dimLen|], dnames = dnames ++ toSnocVect' ns, - tctx = dweakT s <$> tctx, + tctx = map (subD $ shift s) tctx, tnames } @@ -264,10 +312,20 @@ namespace WhnfContext private prettyTContextElt : {opts : _} -> BContext d -> BContext n -> - Qty -> BindName -> Term d n -> Eff Pretty (Doc opts) -prettyTContextElt dnames tnames q x s = - pure $ hsep [hcat [!(prettyQty q), !dotD, !(prettyTBind x)], !colonD, - !(withPrec Outer $ prettyTerm dnames tnames s)] + Qty -> BindName -> LocalVar d n -> Eff Pretty (Doc opts) +prettyTContextElt dnames tnames q x s = do + q <- prettyQty q; dot <- dotD + x <- prettyTBind x; colon <- colonD + ty <- withPrec Outer $ prettyTerm dnames tnames s.type; eq <- cstD + tm <- traverse (withPrec Outer . prettyTerm dnames tnames) s.term + d <- askAt INDENT + let qx = hcat [q, dot, x] + pure $ case tm of + Nothing => + ifMultiline (hsep [qx, colon, ty]) (vsep [qx, indent d $ colon <++> ty]) + Just tm => + ifMultiline (hsep [qx, colon, ty, eq, tm]) + (vsep [qx, indent d $ colon <++> ty, indent d $ eq <++> tm]) private prettyTContext' : {opts : _} -> diff --git a/lib/Quox/Untyped/Erase.idr b/lib/Quox/Untyped/Erase.idr index 701eaad..c51c5f5 100644 --- a/lib/Quox/Untyped/Erase.idr +++ b/lib/Quox/Untyped/Erase.idr @@ -88,16 +88,16 @@ parameters {opts : LayoutOpts} (showContext : Bool) public export Erase : List (Type -> Type) -Erase = [Except Error, DefsReader, NameGen] +Erase = [Except Error, NameGen] export liftWhnf : Eff Whnf a -> Eff Erase a liftWhnf act = lift $ wrapErr WrapTypeError act export covering -computeElimType : ErasureContext d n -> SQty -> Elim d n -> Eff Erase (Term d n) -computeElimType ctx sg e = do - defs <- askAt DEFS +computeElimType : Q.Definitions -> ErasureContext d n -> SQty -> + Elim d n -> Eff Erase (Term d n) +computeElimType defs ctx sg e = do let ctx = toWhnfContext ctx liftWhnf $ do Element e _ <- whnf defs ctx sg e @@ -106,10 +106,11 @@ computeElimType ctx sg e = do private %macro wrapExpect : TTImp -> - Elab (TyContext d n -> Loc -> Term d n -> Eff Erase a) + Elab (Q.Definitions -> TyContext d n -> Loc -> + Term d n -> Eff Erase a) wrapExpect f_ = do f <- check `(\x => ~(f_) x) - pure $ \ctx, loc, s => liftWhnf $ f !(askAt DEFS) ctx SZero loc s + pure $ \defs, ctx, loc, s => liftWhnf $ f defs ctx SZero loc s public export @@ -119,27 +120,36 @@ record EraseElimResult d n where term : U.Term n +export covering +eraseTerm' : (defs : Q.Definitions) -> (ctx : ErasureContext d n) -> + (ty, tm : Q.Term d n) -> + (0 _ : NotRedex defs (toWhnfContext ctx) SZero ty) => + Eff Erase (U.Term n) + -- "Ψ | Γ | Σ ⊢ s ⤋ s' ⇐ A" for `s' <- eraseTerm (Ψ,Γ,Σ) A s` -- -- in the below comments, Ψ, Γ, Σ are implicit and -- only their extensions are written export covering -eraseTerm : ErasureContext d n -> +eraseTerm : Q.Definitions -> ErasureContext d n -> (ty, tm : Q.Term d n) -> Eff Erase (U.Term n) +eraseTerm defs ctx ty tm = do + Element ty _ <- liftWhnf $ Interface.whnf defs (toWhnfContext ctx) SZero ty + eraseTerm' defs ctx ty tm -- "Ψ | Γ | Σ ⊢ e ⤋ e' ⇒ A" for `EraRes A e' <- eraseElim (Ψ,Γ,Σ) e` export covering -eraseElim : ErasureContext d n -> (tm : Q.Elim d n) -> +eraseElim : Q.Definitions -> ErasureContext d n -> (tm : Q.Elim d n) -> Eff Erase (EraseElimResult d n) -eraseTerm ctx _ s@(TYPE {}) = +eraseTerm' defs ctx _ s@(TYPE {}) = throw $ CompileTimeOnly ctx s -eraseTerm ctx _ s@(IOState {}) = +eraseTerm' defs ctx _ s@(IOState {}) = throw $ CompileTimeOnly ctx s -eraseTerm ctx _ s@(Pi {}) = +eraseTerm' defs ctx _ s@(Pi {}) = throw $ CompileTimeOnly ctx s -- x : A | 0.x ⊢ s ⤋ s' ⇐ B @@ -149,66 +159,66 @@ eraseTerm ctx _ s@(Pi {}) = -- x : A | π.x ⊢ s ⤋ s' ⇐ B π ≠ 0 -- ---------------------------------------- -- (λ x ⇒ s) ⤋ (λ x ⇒ s') ⇐ π.(x : A) → B -eraseTerm ctx ty (Lam body loc) = do +eraseTerm' defs ctx ty (Lam body loc) = do let x = body.name - (qty, arg, res) <- wrapExpect `(expectPi) ctx loc ty - body <- eraseTerm (extendTy qty x arg ctx) res.term body.term + (qty, arg, res) <- wrapExpect `(expectPi) defs ctx loc ty + body <- eraseTerm defs (extendTy qty x arg ctx) res.term body.term pure $ case isErased qty of Kept => U.Lam x body loc Erased => sub1 (Erased loc) body -eraseTerm ctx _ s@(Sig {}) = +eraseTerm' defs ctx _ s@(Sig {}) = throw $ CompileTimeOnly ctx s -- s ⤋ s' ⇐ A t ⤋ t' ⇐ B[s/x] -- --------------------------------- -- (s, t) ⤋ (s', t') ⇐ (x : A) × B -eraseTerm ctx ty (Pair fst snd loc) = do - (a, b) <- wrapExpect `(expectSig) ctx loc ty +eraseTerm' defs ctx ty (Pair fst snd loc) = do + (a, b) <- wrapExpect `(expectSig) defs ctx loc ty let b = sub1 b (Ann fst a a.loc) - fst <- eraseTerm ctx a fst - snd <- eraseTerm ctx b snd + fst <- eraseTerm defs ctx a fst + snd <- eraseTerm defs ctx b snd pure $ Pair fst snd loc -eraseTerm ctx _ s@(Enum {}) = +eraseTerm' defs ctx _ s@(Enum {}) = throw $ CompileTimeOnly ctx s -- '𝐚 ⤋ '𝐚 ⇐ {⋯} -eraseTerm ctx _ (Tag tag loc) = +eraseTerm' defs ctx _ (Tag tag loc) = pure $ Tag tag loc -eraseTerm ctx ty s@(Eq {}) = +eraseTerm' defs ctx ty s@(Eq {}) = throw $ CompileTimeOnly ctx s -- 𝑖 ⊢ s ⤋ s' ⇐ A -- --------------------------------- -- (δ 𝑖 ⇒ s) ⤋ s' ⇐ Eq (𝑖 ⇒ A) l r -eraseTerm ctx ty (DLam body loc) = do - a <- fst <$> wrapExpect `(expectEq) ctx loc ty - eraseTerm (extendDim body.name ctx) a.term body.term +eraseTerm' defs ctx ty (DLam body loc) = do + a <- fst <$> wrapExpect `(expectEq) defs ctx loc ty + eraseTerm defs (extendDim body.name ctx) a.term body.term -eraseTerm ctx _ s@(NAT {}) = +eraseTerm' defs ctx _ s@(NAT {}) = throw $ CompileTimeOnly ctx s -- n ⤋ n ⇐ ℕ -eraseTerm _ _ (Nat n loc) = +eraseTerm' _ _ _ (Nat n loc) = pure $ Nat n loc -- s ⤋ s' ⇐ ℕ -- ----------------------- -- succ s ⤋ succ s' ⇐ ℕ -eraseTerm ctx ty (Succ p loc) = do - p <- eraseTerm ctx ty p +eraseTerm' defs ctx ty (Succ p loc) = do + p <- eraseTerm defs ctx ty p pure $ Succ p loc -eraseTerm ctx ty s@(STRING {}) = +eraseTerm' defs ctx ty s@(STRING {}) = throw $ CompileTimeOnly ctx s -- s ⤋ s ⇐ String -eraseTerm _ _ (Str s loc) = +eraseTerm' _ _ _ (Str s loc) = pure $ Str s loc -eraseTerm ctx ty s@(BOX {}) = +eraseTerm' defs ctx ty s@(BOX {}) = throw $ CompileTimeOnly ctx s -- [s] ⤋ ⌷ ⇐ [0.A] @@ -216,48 +226,49 @@ eraseTerm ctx ty s@(BOX {}) = -- π ≠ 0 s ⤋ s' ⇐ A -- -------------------- -- [s] ⤋ s' ⇐ [π.A] -eraseTerm ctx ty (Box val loc) = do - (qty, a) <- wrapExpect `(expectBOX) ctx loc ty +eraseTerm' defs ctx ty (Box val loc) = do + (qty, a) <- wrapExpect `(expectBOX) defs ctx loc ty case isErased qty of Erased => pure $ Erased loc - Kept => eraseTerm ctx a val + Kept => eraseTerm defs ctx a val -- s ⤋ s' ⇐ A -- --------------------------------- -- let0 x = e in s ⤋ s'[⌷/x] ⇐ A -- --- e ⤋ e' s ⤋ s' ⇐ A π ≠ 0 +-- e ⤋ e' ⇒ E π ≠ 0 +-- x : E ≔ e ⊢ s ⤋ s' ⇐ A -- ------------------------------------- -- letπ x = e in s ⤋ let x = e' in s' -eraseTerm ctx ty (Let pi e s loc) = do +eraseTerm' defs ctx ty (Let pi e s loc) = do let x = s.name case isErased pi of Erased => do - ety <- computeElimType ctx SZero e - s' <- eraseTerm (extendTy pi x ety ctx) (weakT 1 ty) s.term + ety <- computeElimType defs ctx SZero e + s' <- eraseTerm defs (extendTyLet pi x ety (E e) ctx) (weakT 1 ty) s.term pure $ sub1 (Erased e.loc) s' Kept => do - EraRes ety e' <- eraseElim ctx e - s' <- eraseTerm (extendTy pi x ety ctx) (weakT 1 ty) s.term + EraRes ety e' <- eraseElim defs ctx e + s' <- eraseTerm defs (extendTyLet pi x ety (E e) ctx) (weakT 1 ty) s.term pure $ Let x e' s' loc -- e ⤋ e' ⇒ B -- ------------ -- e ⤋ e' ⇐ A -eraseTerm ctx ty (E e) = - term <$> eraseElim ctx e +eraseTerm' defs ctx ty (E e) = + term <$> eraseElim defs ctx e -eraseTerm ctx ty (CloT (Sub term th)) = - eraseTerm ctx ty $ pushSubstsWith' id th term +eraseTerm' defs ctx ty (CloT (Sub term th)) = + eraseTerm defs ctx ty $ pushSubstsWith' id th term -eraseTerm ctx ty (DCloT (Sub term th)) = - eraseTerm ctx ty $ pushSubstsWith' th id term +eraseTerm' defs ctx ty (DCloT (Sub term th)) = + eraseTerm defs ctx ty $ pushSubstsWith' th id term -- defω x : A = s -- ---------------- -- x ⤋ x ⇒ A -eraseElim ctx e@(F x u loc) = do - Just def <- asksAt DEFS $ lookup x +eraseElim defs ctx e@(F x u loc) = do + let Just def = lookup x defs | Nothing => throw $ notInScope loc x case isErased def.qty.qty of Erased => throw $ CompileTimeOnly ctx $ E e @@ -266,10 +277,10 @@ eraseElim ctx e@(F x u loc) = do -- π.x ∈ Σ π ≠ 0 -- ----------------- -- x ⤋ x ⇒ A -eraseElim ctx e@(B i loc) = do +eraseElim defs ctx e@(B i loc) = do case isErased $ ctx.qtys !!! i of Erased => throw $ CompileTimeOnly ctx $ E e - Kept => pure $ EraRes (ctx.tctx !! i) $ B i loc + Kept => pure $ EraRes (ctx.tctx !! i).type $ B i loc -- f ⤋ f' ⇒ π.(x : A) → B s ⤋ s' ⇒ A π ≠ 0 -- --------------------------------------------- @@ -278,13 +289,13 @@ eraseElim ctx e@(B i loc) = do -- f ⤋ f' ⇒ 0.(x : A) → B -- ------------------------- -- f s ⤋ f' ⇒ B[s/x] -eraseElim ctx (App fun arg loc) = do - efun <- eraseElim ctx fun - (qty, targ, tres) <- wrapExpect `(expectPi) ctx loc efun.type +eraseElim defs ctx (App fun arg loc) = do + efun <- eraseElim defs ctx fun + (qty, targ, tres) <- wrapExpect `(expectPi) defs ctx loc efun.type let ty = sub1 tres (Ann arg targ arg.loc) case isErased qty of Erased => pure $ EraRes ty efun.term - Kept => do arg <- eraseTerm ctx targ arg + Kept => do arg <- eraseTerm defs ctx targ arg pure $ EraRes ty $ App efun.term arg loc -- e ⇒ (x : A) × B @@ -298,16 +309,16 @@ eraseElim ctx (App fun arg loc) = do -- (caseρ e return z ⇒ R of {(x, y) ⇒ s}) ⤋ -- ⤋ -- let xy = e' in let x = fst xy in let y = snd xy in s' ⇒ R[e/z] -eraseElim ctx (CasePair qty pair ret body loc) = do +eraseElim defs ctx (CasePair qty pair ret body loc) = do let [< x, y] = body.names case isErased qty of Kept => do - EraRes ety eterm <- eraseElim ctx pair + EraRes ety eterm <- eraseElim defs ctx pair let ty = sub1 (ret // shift 2) $ Ann (Pair (BVT 0 loc) (BVT 1 loc) loc) (weakT 2 ety) loc - (tfst, tsnd) <- wrapExpect `(expectSig) ctx loc ety + (tfst, tsnd) <- wrapExpect `(expectSig) defs ctx loc ety let ctx' = extendTyN [< (qty, x, tfst), (qty, y, tsnd.term)] ctx - body' <- eraseTerm ctx' ty body.term + body' <- eraseTerm defs ctx' ty body.term p <- mnb "p" loc pure $ EraRes (sub1 ret pair) $ Let p eterm @@ -316,28 +327,28 @@ eraseElim ctx (CasePair qty pair ret body loc) = do (body' // (B VZ loc ::: B (VS VZ) loc ::: shift 3)) loc) loc) loc Erased => do - ety <- computeElimType ctx SOne pair + ety <- computeElimType defs ctx SOne pair let ty = sub1 (ret // shift 2) $ Ann (Pair (BVT 0 loc) (BVT 1 loc) loc) (weakT 2 ety) loc - (tfst, tsnd) <- wrapExpect `(expectSig) ctx loc ety + (tfst, tsnd) <- wrapExpect `(expectSig) defs ctx loc ety let ctx' = extendTyN0 [< (x, tfst), (y, tsnd.term)] ctx - body' <- eraseTerm ctx' ty body.term + body' <- eraseTerm defs ctx' ty body.term pure $ EraRes (sub1 ret pair) $ subN [< Erased loc, Erased loc] body' -- e ⤋ e' ⇒ (x : A) × B -- ---------------------- -- fst e ⤋ fst e' ⇒ A -eraseElim ctx (Fst pair loc) = do - epair <- eraseElim ctx pair - a <- fst <$> wrapExpect `(expectSig) ctx loc epair.type +eraseElim defs ctx (Fst pair loc) = do + epair <- eraseElim defs ctx pair + a <- fst <$> wrapExpect `(expectSig) defs ctx loc epair.type pure $ EraRes a $ Fst epair.term loc -- e ⤋ e' ⇒ (x : A) × B -- ----------------------------- -- snd e ⤋ snd e' ⇒ B[fst e/x] -eraseElim ctx (Snd pair loc) = do - epair <- eraseElim ctx pair - b <- snd <$> wrapExpect `(expectSig) ctx loc epair.type +eraseElim defs ctx (Snd pair loc) = do + epair <- eraseElim defs ctx pair + b <- snd <$> wrapExpect `(expectSig) defs ctx loc epair.type pure $ EraRes (sub1 b (Fst pair loc)) $ Snd epair.term loc -- caseρ e return z ⇒ R of {} ⤋ absurd ⇒ R[e/z] @@ -349,23 +360,23 @@ eraseElim ctx (Snd pair loc) = do -- e ⤋ e' ⇒ A sᵢ ⤋ s'ᵢ ⇐ R[𝐚ᵢ/z] ρ ≠ 0 i ≠ 0 -- ------------------------------------------------------------------- -- caseρ e return z ⇒ R of {𝐚ᵢ ⇒ sᵢ} ⤋ case e of {𝐚ᵢ ⇒ s'ᵢ} ⇒ R[e/z] -eraseElim ctx e@(CaseEnum qty tag ret arms loc) = do +eraseElim defs ctx e@(CaseEnum qty tag ret arms loc) = do let ty = sub1 ret tag case isErased qty of Erased => case SortedMap.toList arms of [] => pure $ EraRes ty $ Absurd loc [(t, rhs)] => do let ty' = sub1 ret (Ann (Tag t loc) (enum [t] loc) loc) - rhs' <- eraseTerm ctx ty' rhs + rhs' <- eraseTerm defs ctx ty' rhs pure $ EraRes ty rhs' _ => throw $ CompileTimeOnly ctx $ E e Kept => case List1.fromList $ SortedMap.toList arms of Nothing => pure $ EraRes ty $ Absurd loc Just arms => do - etag <- eraseElim ctx tag + etag <- eraseElim defs ctx tag arms <- for arms $ \(t, rhs) => do let ty' = sub1 ret (Ann (Tag t loc) etag.type loc) - rhs' <- eraseTerm ctx ty' rhs + rhs' <- eraseTerm defs ctx ty' rhs pure (t, rhs') pure $ EraRes ty $ CaseEnum etag.term arms loc @@ -382,12 +393,12 @@ eraseElim ctx e@(CaseEnum qty tag ret arms loc) = do -- caseρ n return z ⇒ R of {0 ⇒ z; succ m, 0.ih ⇒ s} -- ⤋ -- case n' of {0 ⇒ z'; succ m ⇒ s'[⌷/ih]} ⇒ R[n/z] -eraseElim ctx (CaseNat qty qtyIH nat ret zero succ loc) = do +eraseElim defs ctx (CaseNat qty qtyIH nat ret zero succ loc) = do let ty = sub1 ret nat - enat <- eraseElim ctx nat - zero <- eraseTerm ctx (sub1 ret (Ann (Zero loc) (NAT loc) loc)) zero + enat <- eraseElim defs ctx nat + zero <- eraseTerm defs ctx (sub1 ret (Ann (Zero loc) (NAT loc) loc)) zero let [< p, ih] = succ.names - succ' <- eraseTerm + succ' <- eraseTerm defs (extendTyN [< (qty, p, NAT loc), (qtyIH, ih, sub1 (ret // shift 1) (BV 0 loc))] ctx) (sub1 (ret // shift 2) (Ann (Succ (BVT 1 loc) loc) (NAT loc) loc)) @@ -404,56 +415,56 @@ eraseElim ctx (CaseNat qty qtyIH nat ret zero succ loc) = do -- b ⇒ [π.A] x : A | 0.x ⊢ s ⤋ s' ⇐ R[[x]∷[0.A]/z] πρ = 0 -- ------------------------------------------------------------- -- caseρ b return z ⇒ R of {[x] ⇒ s} ⤋ s'[⌷/x] ⇒ R[b/z] -eraseElim ctx (CaseBox qty box ret body loc) = do - tbox <- computeElimType ctx SOne box -- [fixme] is there any way to avoid this? - (pi, tinner) <- wrapExpect `(expectBOX) ctx loc tbox +eraseElim defs ctx (CaseBox qty box ret body loc) = do + tbox <- computeElimType defs ctx SOne box + (pi, tinner) <- wrapExpect `(expectBOX) defs ctx loc tbox let ctx' = extendTy (pi * qty) body.name tinner ctx bty = sub1 (ret // shift 1) $ Ann (Box (BVT 0 loc) loc) (weakT 1 tbox) loc case isErased $ qty * pi of Kept => do - ebox <- eraseElim ctx box - ebody <- eraseTerm ctx' bty body.term + ebox <- eraseElim defs ctx box + ebody <- eraseTerm defs ctx' bty body.term pure $ EraRes (sub1 ret box) $ Let body.name ebox.term ebody loc Erased => do - body' <- eraseTerm ctx' bty body.term + body' <- eraseTerm defs ctx' bty body.term pure $ EraRes (sub1 ret box) $ body' // one (Erased loc) -- f ⤋ f' ⇒ Eq (𝑖 ⇒ A) l r -- ------------------------------ -- f @r ⤋ f' ⇒ A‹r/𝑖› -eraseElim ctx (DApp fun arg loc) = do - efun <- eraseElim ctx fun - a <- fst <$> wrapExpect `(expectEq) ctx loc efun.type +eraseElim defs ctx (DApp fun arg loc) = do + efun <- eraseElim defs ctx fun + a <- fst <$> wrapExpect `(expectEq) defs ctx loc efun.type pure $ EraRes (dsub1 a arg) efun.term -- s ⤋ s' ⇐ A -- ---------------- -- s ∷ A ⤋ s' ⇒ A -eraseElim ctx (Ann tm ty loc) = - EraRes ty <$> eraseTerm ctx ty tm +eraseElim defs ctx (Ann tm ty loc) = + EraRes ty <$> eraseTerm defs ctx ty tm -- s ⤋ s' ⇐ A‹p/𝑖› -- ----------------------------------- -- coe (𝑖 ⇒ A) @p @q s ⤋ s' ⇒ A‹q/𝑖› -eraseElim ctx (Coe ty p q val loc) = do - val <- eraseTerm ctx (dsub1 ty p) val +eraseElim defs ctx (Coe ty p q val loc) = do + val <- eraseTerm defs ctx (dsub1 ty p) val pure $ EraRes (dsub1 ty q) val -- s ⤋ s' ⇐ A -- -------------------------------- -- comp A @p @q s @r {⋯} ⤋ s' ⇒ A -eraseElim ctx (Comp ty p q val r zero one loc) = - EraRes ty <$> eraseTerm ctx ty val +eraseElim defs ctx (Comp ty p q val r zero one loc) = + EraRes ty <$> eraseTerm defs ctx ty val -eraseElim ctx t@(TypeCase ty ret arms def loc) = +eraseElim defs ctx t@(TypeCase ty ret arms def loc) = throw $ CompileTimeOnly ctx $ E t -eraseElim ctx (CloE (Sub term th)) = - eraseElim ctx $ pushSubstsWith' id th term +eraseElim defs ctx (CloE (Sub term th)) = + eraseElim defs ctx $ pushSubstsWith' id th term -eraseElim ctx (DCloE (Sub term th)) = - eraseElim ctx $ pushSubstsWith' th id term +eraseElim defs ctx (DCloE (Sub term th)) = + eraseElim defs ctx $ pushSubstsWith' th id term export @@ -539,8 +550,8 @@ trimLets (Erased loc) = Erased loc export covering -eraseDef : Name -> Q.Definition -> Eff Erase U.Definition -eraseDef name def@(MkDef qty type body scheme isMain loc) = +eraseDef : Q.Definitions -> Name -> Q.Definition -> Eff Erase U.Definition +eraseDef defs name def@(MkDef qty type body scheme isMain loc) = wrapErr (WhileErasing name def) $ case isErased qty.qty of Erased => do @@ -552,4 +563,4 @@ eraseDef name def@(MkDef qty type body scheme isMain loc) = Nothing => case body of Postulate => throw $ Postulate loc name Concrete body => KeptDef isMain . trimLets <$> - eraseTerm empty type body + eraseTerm defs empty type body diff --git a/lib/Quox/Whnf/Coercion.idr b/lib/Quox/Whnf/Coercion.idr index 0e35665..4b900bf 100644 --- a/lib/Quox/Whnf/Coercion.idr +++ b/lib/Quox/Whnf/Coercion.idr @@ -28,7 +28,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT} export covering piCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val, s : Term d n) -> Loc -> - Eff Whnf (Subset (Elim d n) (No . isRedexE defs sg)) + Eff Whnf (Subset (Elim d n) (No . isRedexE defs ctx sg)) piCoe sty@(S [< i] ty) p q val s loc = do -- (coe [i ⇒ π.(x : A) → B] @p @q t) s ⇝ -- coe [i ⇒ B[𝒔‹i›/x] @p @q ((t ∷ (π.(x : A) → B)‹p/i›) 𝒔‹p›) @@ -49,7 +49,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT} sigCoe : (qty : Qty) -> (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) -> (ret : ScopeTerm d n) -> (body : ScopeTermN 2 d n) -> Loc -> - Eff Whnf (Subset (Elim d n) (No . isRedexE defs sg)) + Eff Whnf (Subset (Elim d n) (No . isRedexE defs ctx sg)) sigCoe qty sty@(S [< i] ty) p q val ret body loc = do -- caseπ (coe [i ⇒ (x : A) × B] @p @q s) return z ⇒ C of { (a, b) ⇒ e } -- ⇝ @@ -74,7 +74,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT} ||| reduce a pair projection `Fst (Coe ty p q val) loc` export covering fstCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) -> - Loc -> Eff Whnf (Subset (Elim d n) (No . isRedexE defs sg)) + Loc -> Eff Whnf (Subset (Elim d n) (No . isRedexE defs ctx sg)) fstCoe sty@(S [< i] ty) p q val loc = do -- fst (coe (𝑖 ⇒ (x : A) × B) @p @q s) -- ⇝ @@ -91,7 +91,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT} ||| reduce a pair projection `Snd (Coe ty p q val) loc` export covering sndCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) -> - Loc -> Eff Whnf (Subset (Elim d n) (No . isRedexE defs sg)) + Loc -> Eff Whnf (Subset (Elim d n) (No . isRedexE defs ctx sg)) sndCoe sty@(S [< i] ty) p q val loc = do -- snd (coe (𝑖 ⇒ (x : A) × B) @p @q s) -- ⇝ @@ -115,7 +115,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT} export covering eqCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) -> (r : Dim d) -> Loc -> - Eff Whnf (Subset (Elim d n) (No . isRedexE defs sg)) + Eff Whnf (Subset (Elim d n) (No . isRedexE defs ctx sg)) eqCoe sty@(S [< j] ty) p q val r loc = do -- (coe [j ⇒ Eq [i ⇒ A] L R] @p @q eq) @r -- ⇝ @@ -133,7 +133,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT} boxCoe : (qty : Qty) -> (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) -> (ret : ScopeTerm d n) -> (body : ScopeTerm d n) -> Loc -> - Eff Whnf (Subset (Elim d n) (No . isRedexE defs sg)) + Eff Whnf (Subset (Elim d n) (No . isRedexE defs ctx sg)) boxCoe qty sty@(S [< i] ty) p q val ret body loc = do -- caseπ (coe [i ⇒ [ρ. A]] @p @q s) return z ⇒ C of { [a] ⇒ e } -- ⇝ @@ -151,7 +151,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT} pushCoe : BindName -> (ty : Term (S d) n) -> (p, q : Dim d) -> (s : Term d n) -> Loc -> (0 pc : So (canPushCoe sg ty s)) => - Eff Whnf (NonRedex Elim d n defs sg) + Eff Whnf (NonRedex Elim d n defs ctx sg) pushCoe i ty p q s loc = case ty of -- (coe ★ᵢ @_ @_ s) ⇝ (s ∷ ★ᵢ) diff --git a/lib/Quox/Whnf/ComputeElimType.idr b/lib/Quox/Whnf/ComputeElimType.idr index 441ba1e..87891ff 100644 --- a/lib/Quox/Whnf/ComputeElimType.idr +++ b/lib/Quox/Whnf/ComputeElimType.idr @@ -14,8 +14,8 @@ export covering computeElimType : CanWhnf Term Interface.isRedexT => CanWhnf Elim Interface.isRedexE => - (defs : Definitions) -> WhnfContext d n -> (0 sg : SQty) -> - (e : Elim d n) -> (0 ne : No (isRedexE defs sg e)) => + (defs : Definitions) -> (ctx : WhnfContext d n) -> (0 sg : SQty) -> + (e : Elim d n) -> (0 ne : No (isRedexE defs ctx sg e)) => Eff Whnf (Term d n) @@ -24,8 +24,8 @@ export covering computeWhnfElimType0 : CanWhnf Term Interface.isRedexT => CanWhnf Elim Interface.isRedexE => - (defs : Definitions) -> WhnfContext d n -> (0 sg : SQty) -> - (e : Elim d n) -> (0 ne : No (isRedexE defs sg e)) => + (defs : Definitions) -> (ctx : WhnfContext d n) -> (0 sg : SQty) -> + (e : Elim d n) -> (0 ne : No (isRedexE defs ctx sg e)) => Eff Whnf (Term d n) computeElimType defs ctx sg e = @@ -36,7 +36,7 @@ computeElimType defs ctx sg e = pure $ def.typeWithAt ctx.dimLen ctx.termLen u B i _ => - pure $ ctx.tctx !! i + pure (ctx.tctx !! i).type App f s loc => case !(computeWhnfElimType0 defs ctx sg f {ne = noOr1 ne}) of diff --git a/lib/Quox/Whnf/Interface.idr b/lib/Quox/Whnf/Interface.idr index 15a8978..d2b11da 100644 --- a/lib/Quox/Whnf/Interface.idr +++ b/lib/Quox/Whnf/Interface.idr @@ -18,13 +18,14 @@ Whnf = [Except Error, NameGen] public export 0 RedexTest : TermLike -> Type -RedexTest tm = {0 d, n : Nat} -> Definitions -> SQty -> tm d n -> Bool +RedexTest tm = + {0 d, n : Nat} -> Definitions -> WhnfContext d n -> SQty -> tm d n -> Bool public export interface CanWhnf (0 tm : TermLike) (0 isRedex : RedexTest tm) | tm where whnf : (defs : Definitions) -> (ctx : WhnfContext d n) -> (q : SQty) -> - tm d n -> Eff Whnf (Subset (tm d n) (No . isRedex defs q)) + tm d n -> Eff Whnf (Subset (tm d n) (No . isRedex defs ctx q)) -- having isRedex be part of the class header, and needing to be explicitly -- quantified on every use since idris can't infer its type, is a little ugly. -- but none of the alternatives i've thought of so far work. e.g. in some @@ -37,19 +38,20 @@ whnf0 defs ctx q t = fst <$> whnf defs ctx q t public export 0 IsRedex, NotRedex : {isRedex : RedexTest tm} -> CanWhnf tm isRedex => - Definitions -> SQty -> Pred (tm d n) -IsRedex defs q = So . isRedex defs q -NotRedex defs q = No . isRedex defs q + Definitions -> WhnfContext d n -> SQty -> Pred (tm d n) +IsRedex defs ctx q = So . isRedex defs ctx q +NotRedex defs ctx q = No . isRedex defs ctx q public export 0 NonRedex : (tm : TermLike) -> {isRedex : RedexTest tm} -> CanWhnf tm isRedex => (d, n : Nat) -> - (defs : Definitions) -> SQty -> Type -NonRedex tm d n defs q = Subset (tm d n) (NotRedex defs q) + Definitions -> WhnfContext d n -> SQty -> Type +NonRedex tm d n defs ctx q = Subset (tm d n) (NotRedex defs ctx q) public export %inline nred : {0 isRedex : RedexTest tm} -> (0 _ : CanWhnf tm isRedex) => - (t : tm d n) -> (0 nr : NotRedex defs q t) => NonRedex tm d n defs q + (t : tm d n) -> (0 nr : NotRedex defs ctx q t) => + NonRedex tm d n defs ctx q nred t = Element t nr @@ -193,50 +195,52 @@ mutual ||| a reducible elimination ||| ||| 1. a free variable, if its definition is known - ||| 2. an elimination whose head is reducible - ||| 3. an "active" elimination: + ||| 2. a bound variable pointing to a `let` + ||| 3. an elimination whose head is reducible + ||| 4. an "active" elimination: ||| an application whose head is an annotated lambda, ||| a case expression whose head is an annotated constructor form, etc - ||| 4. a redundant annotation, or one whose term or type is reducible - ||| 5. a coercion `coe (𝑖 ⇒ A) @p @q s` where: + ||| 5. a redundant annotation, or one whose term or type is reducible + ||| 6. a coercion `coe (𝑖 ⇒ A) @p @q s` where: ||| a. `A` is reducible or a type constructor, or ||| b. `𝑖` is not mentioned in `A` ||| ([fixme] should be A‹0/𝑖› = A‹1/𝑖›), or ||| c. `p = q` - ||| 6. a composition `comp A @p @q s @r {⋯}` + ||| 7. a composition `comp A @p @q s @r {⋯}` ||| where `p = q`, `r = 0`, or `r = 1` - ||| 7. a closure + ||| 8. a closure public export isRedexE : RedexTest Elim - isRedexE defs sg (F {x, u, _}) = isJust $ lookupElim0 x u defs - isRedexE _ sg (B {}) = False - isRedexE defs sg (App {fun, _}) = - isRedexE defs sg fun || isLamHead fun - isRedexE defs sg (CasePair {pair, _}) = - isRedexE defs sg pair || isPairHead pair || isYes (sg `decEq` SZero) - isRedexE defs sg (Fst pair _) = - isRedexE defs sg pair || isPairHead pair - isRedexE defs sg (Snd pair _) = - isRedexE defs sg pair || isPairHead pair - isRedexE defs sg (CaseEnum {tag, _}) = - isRedexE defs sg tag || isTagHead tag - isRedexE defs sg (CaseNat {nat, _}) = - isRedexE defs sg nat || isNatHead nat - isRedexE defs sg (CaseBox {box, _}) = - isRedexE defs sg box || isBoxHead box - isRedexE defs sg (DApp {fun, arg, _}) = - isRedexE defs sg fun || isDLamHead fun || isK arg - isRedexE defs sg (Ann {tm, ty, _}) = - isE tm || isRedexT defs sg tm || isRedexT defs SZero ty - isRedexE defs sg (Coe {ty = S _ (N _), _}) = True - isRedexE defs sg (Coe {ty = S _ (Y ty), p, q, val, _}) = - isRedexT defs SZero ty || canPushCoe sg ty val || isYes (p `decEqv` q) - isRedexE defs sg (Comp {ty, p, q, r, _}) = + isRedexE defs ctx sg (F {x, u, _}) = isJust $ lookupElim0 x u defs + isRedexE _ ctx sg (B {i, _}) = isJust (ctx.tctx !! i).term + isRedexE defs ctx sg (App {fun, _}) = + isRedexE defs ctx sg fun || isLamHead fun + isRedexE defs ctx sg (CasePair {pair, _}) = + isRedexE defs ctx sg pair || isPairHead pair || isYes (sg `decEq` SZero) + isRedexE defs ctx sg (Fst pair _) = + isRedexE defs ctx sg pair || isPairHead pair + isRedexE defs ctx sg (Snd pair _) = + isRedexE defs ctx sg pair || isPairHead pair + isRedexE defs ctx sg (CaseEnum {tag, _}) = + isRedexE defs ctx sg tag || isTagHead tag + isRedexE defs ctx sg (CaseNat {nat, _}) = + isRedexE defs ctx sg nat || isNatHead nat + isRedexE defs ctx sg (CaseBox {box, _}) = + isRedexE defs ctx sg box || isBoxHead box + isRedexE defs ctx sg (DApp {fun, arg, _}) = + isRedexE defs ctx sg fun || isDLamHead fun || isK arg + isRedexE defs ctx sg (Ann {tm, ty, _}) = + isE tm || isRedexT defs ctx sg tm || isRedexT defs ctx SZero ty + isRedexE defs ctx sg (Coe {ty = S _ (N _), _}) = True + isRedexE defs ctx sg (Coe {ty = S [< i] (Y ty), p, q, val, _}) = + isRedexT defs (extendDim i ctx) SZero ty || + canPushCoe sg ty val || isYes (p `decEqv` q) + isRedexE defs ctx sg (Comp {ty, p, q, r, _}) = isYes (p `decEqv` q) || isK r - isRedexE defs sg (TypeCase {ty, ret, _}) = - isRedexE defs sg ty || isRedexT defs sg ret || isAnnTyCon ty - isRedexE _ _ (CloE {}) = True - isRedexE _ _ (DCloE {}) = True + isRedexE defs ctx sg (TypeCase {ty, ret, _}) = + isRedexE defs ctx sg ty || isRedexT defs ctx sg ret || isAnnTyCon ty + isRedexE _ _ _ (CloE {}) = True + isRedexE _ _ _ (DCloE {}) = True ||| a reducible term ||| @@ -248,9 +252,9 @@ mutual ||| 5. a `let` expression public export isRedexT : RedexTest Term - isRedexT _ _ (CloT {}) = True - isRedexT _ _ (DCloT {}) = True - isRedexT _ _ (Let {}) = True - isRedexT defs sg (E {e, _}) = isAnn e || isRedexE defs sg e - isRedexT _ _ (Succ p {}) = isNatConst p - isRedexT _ _ _ = False + isRedexT _ _ _ (CloT {}) = True + isRedexT _ _ _ (DCloT {}) = True + isRedexT _ _ _ (Let {}) = True + isRedexT defs ctx sg (E {e, _}) = isAnn e || isRedexE defs ctx sg e + isRedexT _ _ _ (Succ p {}) = isNatConst p + isRedexT _ _ _ _ = False diff --git a/lib/Quox/Whnf/Main.idr b/lib/Quox/Whnf/Main.idr index 9248edf..2c0a67a 100644 --- a/lib/Quox/Whnf/Main.idr +++ b/lib/Quox/Whnf/Main.idr @@ -20,7 +20,10 @@ CanWhnf Elim Interface.isRedexE where _ | Just y = whnf defs ctx sg $ setLoc loc $ injElim ctx y _ | Nothing = pure $ Element (F x u loc) $ rewrite eq in Ah - whnf _ _ _ (B i loc) = pure $ nred $ B i loc + whnf defs ctx sg (B i loc) with (ctx.tctx !! i) proof eq1 + _ | l with (l.term) proof eq2 + _ | Just y = whnf defs ctx sg $ Ann y l.type loc + _ | Nothing = pure $ Element (B i loc) $ rewrite eq1 in rewrite eq2 in Ah -- ((λ x ⇒ t) ∷ (π.x : A) → B) s ⇝ t[s∷A/x] ∷ B[s∷A/x] whnf defs ctx sg (App f s appLoc) = do diff --git a/lib/Quox/Whnf/TypeCase.idr b/lib/Quox/Whnf/TypeCase.idr index a42c0d9..2fc34ba 100644 --- a/lib/Quox/Whnf/TypeCase.idr +++ b/lib/Quox/Whnf/TypeCase.idr @@ -35,7 +35,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT} ||| for an elim returns a pair of type-cases that will reduce to that; ||| for other intro forms error export covering - tycasePi : (t : Term d n) -> (0 tnf : No (isRedexT defs SZero t)) => + tycasePi : (t : Term d n) -> (0 tnf : No (isRedexT defs ctx SZero t)) => Eff Whnf (Term d n, ScopeTerm d n) tycasePi (Pi {arg, res, _}) = pure (arg, res) tycasePi (E e) {tnf} = do @@ -53,7 +53,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT} ||| for an elim returns a pair of type-cases that will reduce to that; ||| for other intro forms error export covering - tycaseSig : (t : Term d n) -> (0 tnf : No (isRedexT defs SZero t)) => + tycaseSig : (t : Term d n) -> (0 tnf : No (isRedexT defs ctx SZero t)) => Eff Whnf (Term d n, ScopeTerm d n) tycaseSig (Sig {fst, snd, _}) = pure (fst, snd) tycaseSig (E e) {tnf} = do @@ -71,7 +71,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT} ||| for an elim returns a type-case that will reduce to that; ||| for other intro forms error export covering - tycaseBOX : (t : Term d n) -> (0 tnf : No (isRedexT defs SZero t)) => + tycaseBOX : (t : Term d n) -> (0 tnf : No (isRedexT defs ctx SZero t)) => Eff Whnf (Term d n) tycaseBOX (BOX {ty, _}) = pure ty tycaseBOX (E e) {tnf} = do @@ -83,7 +83,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT} ||| for an elim returns five type-cases that will reduce to that; ||| for other intro forms error export covering - tycaseEq : (t : Term d n) -> (0 tnf : No (isRedexT defs SZero t)) => + tycaseEq : (t : Term d n) -> (0 tnf : No (isRedexT defs ctx SZero t)) => Eff Whnf (Term d n, Term d n, DScopeTerm d n, Term d n, Term d n) tycaseEq (Eq {ty, l, r, _}) = pure (ty.zero, ty.one, ty, l, r) tycaseEq (E e) {tnf} = do @@ -107,7 +107,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT} reduceTypeCase : (ty : Term d n) -> (u : Universe) -> (ret : Term d n) -> (arms : TypeCaseArms d n) -> (def : Term d n) -> (0 _ : So (isTyCon ty)) => Loc -> - Eff Whnf (Subset (Elim d n) (No . isRedexE defs SZero)) + Eff Whnf (Subset (Elim d n) (No . isRedexE defs ctx SZero)) reduceTypeCase ty u ret arms def loc = case ty of -- (type-case ★ᵢ ∷ _ return Q of { ★ ⇒ s; ⋯ }) ⇝ s ∷ Q TYPE {} => diff --git a/tests/AstExtra.idr b/tests/AstExtra.idr index f9cc1fd..76257b7 100644 --- a/tests/AstExtra.idr +++ b/tests/AstExtra.idr @@ -24,13 +24,17 @@ anys : {n : Nat} -> QContext n anys {n = 0} = [<] anys {n = S n} = anys :< Any +public export +locals : Context (Term d) n -> Context (LocalVar d) n +locals = map $ \t => MkLocal t Nothing + public export ctx, ctx01 : {n : Nat} -> Context (\n => (BindName, Term 0 n)) n -> TyContext 0 n ctx tel = let (ns, ts) = unzip tel in - MkTyContext new [<] ts ns anys + MkTyContext new [<] (locals ts) ns anys ctx01 tel = let (ns, ts) = unzip tel in - MkTyContext ZeroIsOne [<] ts ns anys + MkTyContext ZeroIsOne [<] (locals ts) ns anys export mkDef : GQty -> Term 0 0 -> Term 0 0 -> Definition diff --git a/tests/Tests/Reduce.idr b/tests/Tests/Reduce.idr index 66595cb..40c071e 100644 --- a/tests/Tests/Reduce.idr +++ b/tests/Tests/Reduce.idr @@ -33,7 +33,7 @@ parameters {0 isRedex : RedexTest tm} {auto _ : CanWhnf tm isRedex} {d, n : Nat} private ctx : {n : Nat} -> Context (\n => (BindName, Term 0 n)) n -> WhnfContext 0 n -ctx xs = let (ns, ts) = unzip xs in MkWhnfContext [<] ns ts +ctx xs = let (ns, ts) = unzip xs in MkWhnfContext [<] ns (locals ts) export