This commit is contained in:
rhiannon morris 2023-09-27 16:07:28 +02:00
parent 84ef66de41
commit 5e0a557854
8 changed files with 464 additions and 65 deletions

View File

@ -424,7 +424,7 @@ namespace Elim
Eff EqualElim (Term 0 n)
computeElimTypeE defs ectx sg e =
let Val n = ectx.termLen in
lift $ computeElimType defs (toWhnfContext ectx) sg e
lift $ computeElimType defs (toWhnfContext ectx) e {ne}
private
putError : Has InnerErrEff fs => Error -> Eff fs ()

View File

@ -73,7 +73,7 @@ namespace TContext
zeroFor : Context tm n -> QOutput n
zeroFor ctx = Zero <$ ctx
private
public export
extendLen : Telescope a n1 n2 -> Singleton n1 -> Singleton n2
extendLen [<] x = x
extendLen (tel :< _) x = [|S $ extendLen tel x|]

335
lib/Quox/Untyped/Erase.idr Normal file
View File

@ -0,0 +1,335 @@
module Quox.Untyped.Erase
import Quox.Definition
import Quox.Syntax.Term.Base as Q
import Quox.Syntax.Term.Subst
import Quox.Untyped.Syntax as U
import Quox.Typing.Context
import Quox.Whnf
import Quox.EffExtra
import Data.Singleton
import Data.SnocVect
import Language.Reflection
%default total
%language ElabReflection
%hide TT.Name
public export
data IsErased = Erased | Kept
public export
isErased : Qty -> IsErased
isErased Zero = Erased
isErased One = Kept
isErased Any = Kept
public export
EContext : Nat -> Type
EContext = Context' IsErased
public export
record ErasureContext d n where
constructor MkEContexts
{auto dimLen : Singleton d}
{auto termLen : Singleton n}
dnames : BContext d
tnames : BContext n
tctx : TContext d n
erased : EContext n
%name ErasureContext ctx
public export
TypeError : Type
TypeError = Typing.Error.Error
%hide Typing.Error.Error
public export
data Error =
CompileTimeOnly (ErasureContext d n) (Q.Term d n)
| ErasedVar (ErasureContext d n) (Var n)
| NotInScope Name
| WrapTypeError TypeError
public export
Erase : List (Type -> Type)
Erase = [Except Error, DefsReader, NameGen]
export
toWhnfContext : ErasureContext d n -> WhnfContext d n
toWhnfContext (MkEContexts {dnames, tnames, tctx, _}) =
MkWhnfContext {dnames, tnames, tctx}
export
(.names) : ErasureContext d n -> NameContexts d n
ctx.names = MkNameContexts ctx.dnames ctx.tnames
export
liftWhnf : Eff Whnf a -> Eff Erase a
liftWhnf act = runEff act
[\x => send x, \case (Err e) => throw $ WrapTypeError e]
export covering
whnf0 : {0 isRedex : RedexTest tm} -> CanWhnf tm isRedex =>
ErasureContext d n -> SQty ->
tm d n -> Eff Erase (tm d n)
whnf0 ctx sg tm = do
let Val n = ctx.termLen; Val d = ctx.dimLen
liftWhnf $ whnf0 !(askAt DEFS) (toWhnfContext ctx) sg tm
export covering
computeElimType : ErasureContext d n -> SQty -> Elim d n -> Eff Erase (Term d n)
computeElimType ctx sg e = do
let Val n = ctx.termLen; Val d = ctx.dimLen
ctx = toWhnfContext ctx
defs <- askAt DEFS
Element e enf <- liftWhnf $ whnf defs ctx sg e
liftWhnf $ computeElimType defs ctx e {ne = enf}
parameters (ctx : ErasureContext d n) (loc : Loc)
private covering %macro
expect : (forall d, n. Loc -> NameContexts d n -> Term d n -> TypeError) ->
TTImp -> TTImp -> Elab (Term d n -> Eff Erase a)
expect k l r = do
f <- check `(\case ~(l) => Just ~(r); _ => Nothing)
pure $ \t =>
let err = throw $ WrapTypeError $ k loc ctx.names t in
maybe err pure . f =<< whnf0 ctx SZero t
export covering %inline
expectTYPE : Term d n -> Eff Erase Universe
expectTYPE = expect ExpectedTYPE `(TYPE {l, _}) `(l)
export covering %inline
expectPi : Term d n -> Eff Erase (Qty, Term d n, ScopeTerm d n)
expectPi = expect ExpectedPi `(Pi {qty, arg, res, _}) `((qty, arg, res))
export covering %inline
expectSig : Term d n -> Eff Erase (Term d n, ScopeTerm d n)
expectSig = expect ExpectedSig `(Sig {fst, snd, _}) `((fst, snd))
export covering %inline
expectEnum : Term d n -> Eff Erase (SortedSet TagVal)
expectEnum = expect ExpectedEnum `(Enum {cases, _}) `(cases)
export covering %inline
expectEq : Term d n -> Eff Erase (DScopeTerm d n, Term d n, Term d n)
expectEq = expect ExpectedEq `(Eq {ty, l, r, _}) `((ty, l, r))
export covering %inline
expectNat : Term d n -> Eff Erase ()
expectNat = expect ExpectedNat `(Nat {}) `(())
export covering %inline
expectBOX : Term d n -> Eff Erase (Qty, Term d n)
expectBOX = expect ExpectedBOX `(BOX {qty, ty, _}) `((qty, ty))
export
extendTyN : CtxExtension d n1 n2 ->
ErasureContext d n1 -> ErasureContext d n2
extendTyN tel (MkEContexts {termLen, dnames, tnames, tctx, erased}) =
let (qs, xs, ss) = unzip3 tel in
MkEContexts {
dnames,
tnames = tnames . xs,
tctx = tctx . ss,
erased = erased . map isErased qs,
termLen = extendLen tel termLen
}
export
extendTy : Qty -> BindName -> Term d n ->
ErasureContext d n -> ErasureContext d (S n)
extendTy pi x ty = extendTyN [< (pi, x, ty)]
export
extendDim : BindName -> ErasureContext d n -> ErasureContext (S d) n
extendDim i (MkEContexts {dimLen, dnames, tnames, tctx, erased}) =
MkEContexts {
tnames, erased,
dimLen = [|S dimLen|],
dnames = dnames :< i,
tctx = map (dweakT 1) tctx
}
export covering
eraseTerm : ErasureContext d n ->
(ty, tm : Q.Term d n) -> Eff Erase (U.Term n)
export covering
eraseElim : ErasureContext d n -> (tm : Q.Elim d n) ->
Eff Erase (Q.Term d n, U.Term n)
eraseTerm ctx _ s@(TYPE {}) =
throw $ CompileTimeOnly ctx s
eraseTerm ctx _ s@(Pi {}) =
throw $ CompileTimeOnly ctx s
eraseTerm ctx ty (Lam body loc) = do
(qty, arg, res) <- expectPi ctx loc ty
case isErased qty of
Erased => do
-- replace with a fake name like "<erased x>"
-- a little trap for future me :3c
let name = baseStr body.name.val
dummy = F (unq $ UN "<erased \{name}>") 0 noLoc
eraseTerm ctx (sub1 res dummy) (sub1 body dummy)
Kept => do
let x = body.name
body <- eraseTerm (extendTy qty x arg ctx) res.term body.term
pure $ U.Lam x body loc
eraseTerm ctx _ s@(Sig {}) =
throw $ CompileTimeOnly ctx s
eraseTerm ctx ty (Pair fst snd loc) = do
(a, b) <- expectSig ctx loc ty
let b = sub1 b (Ann fst a a.loc)
fst <- eraseTerm ctx a fst
snd <- eraseTerm ctx b snd
pure $ Pair fst snd loc
eraseTerm ctx _ s@(Enum {}) =
throw $ CompileTimeOnly ctx s
eraseTerm ctx _ (Tag tag loc) =
pure $ Tag tag loc
eraseTerm ctx ty s@(Eq {}) =
throw $ CompileTimeOnly ctx s
eraseTerm ctx ty (DLam body loc) = do
(a, _, _) <- expectEq ctx loc ty
eraseTerm (extendDim body.name ctx) a.term body.term
eraseTerm ctx _ s@(Nat {}) =
throw $ CompileTimeOnly ctx s
eraseTerm _ _ (Zero loc) =
pure $ Zero loc
eraseTerm ctx ty (Succ p loc) = do
p <- eraseTerm ctx ty p
pure $ Succ p loc
eraseTerm ctx ty s@(BOX {}) =
throw $ CompileTimeOnly ctx s
eraseTerm ctx ty (Box val loc) = do
(qty, a) <- expectBOX ctx loc ty
case isErased qty of
Erased => pure $ ErasedBox loc -- lmao
Kept => eraseTerm ctx a val
eraseTerm ctx ty (E e) =
snd <$> eraseElim ctx e
eraseTerm ctx ty (CloT (Sub term th)) =
eraseTerm ctx ty $ pushSubstsWith' id th term
eraseTerm ctx ty (DCloT (Sub term th)) =
eraseTerm ctx ty $ pushSubstsWith' th id term
eraseElim ctx e@(F x u loc) = do
Just def <- asksAt DEFS $ lookup x
| Nothing => throw $ NotInScope x
case isErased def.qty.qty of
Erased => throw $ CompileTimeOnly ctx $ E e
Kept =>
let Val d = ctx.dimLen; Val n = ctx.termLen in
pure (def.type, F x loc)
eraseElim ctx e@(B i loc) = do
case ctx.erased !!! i of
Erased => throw $ CompileTimeOnly ctx $ E e
Kept => pure (ctx.tctx !! i, B i loc)
eraseElim ctx (App fun arg loc) = do
(tfun, fun) <- eraseElim ctx fun
(qty, targ, tres) <- expectPi ctx loc tfun
let ty = sub1 tres (Ann arg targ arg.loc)
case isErased qty of
Erased => pure (ty, fun)
Kept => do arg <- eraseTerm ctx targ arg
pure (ty, App fun arg loc)
eraseElim ctx (CasePair qty pair ret body loc) = do
ty <- computeElimType ctx SOne pair
eraseElim ctx $
Ann (subN body [< Fst pair pair.loc, Snd pair pair.loc])
(sub1 ret pair) loc
eraseElim ctx (Fst pair loc) = do
(ty, pair) <- eraseElim ctx pair
(a, _) <- expectSig ctx loc ty
pure (a, Fst pair loc)
eraseElim ctx (Snd pair0 loc) = do
(ty, pair) <- eraseElim ctx pair0
(_, b) <- expectSig ctx loc ty
pure (sub1 b (Fst pair0 loc), Snd pair loc)
eraseElim ctx e@(CaseEnum qty tag ret arms loc) =
case isErased qty of
Erased => case SortedMap.toList arms of
[] => pure (sub1 ret tag, Absurd loc)
[(_, arm)] => let ty = sub1 ret tag in (ty,) <$> eraseTerm ctx ty arm
_ => throw $ CompileTimeOnly ctx $ E e
Kept => do
let ty = sub1 ret tag
(tagTy, tag) <- eraseElim ctx tag
arms <- for (SortedMap.toList arms) $ \(t, rhs) =>
(t,) <$> eraseTerm ctx (sub1 ret $ Ann (Tag t loc) tagTy loc) rhs
pure (ty, CaseEnum tag arms loc)
eraseElim ctx (CaseNat qty qtyIH nat ret zero succ loc) = do
let ty = sub1 ret nat
(_, nat) <- eraseElim ctx nat
zero <- eraseTerm ctx (sub1 ret (Ann (Zero loc) (Nat loc) loc)) zero
let [< p, ih] = succ.names
succ <- eraseTerm
(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))
succ.term
pure (ty, CaseNat nat zero p ih succ loc)
eraseElim ctx (CaseBox qty box ret body loc) = do
let ty = sub1 ret box
body <- eraseTerm ctx ty $ sub1 body box
pure (ty, body)
eraseElim ctx (DApp fun arg loc) = do
(tfun, fun) <- eraseElim ctx fun
(a, _, _) <- expectEq ctx loc tfun
pure (dsub1 a arg, fun)
eraseElim ctx (Ann tm ty loc) =
(ty,) <$> eraseTerm ctx ty tm
eraseElim ctx (Coe ty p q val loc) = do
val <- eraseTerm ctx (dsub1 ty q) val
pure (dsub1 ty p, val)
eraseElim ctx (Comp ty p q val r zero one loc) =
(ty,) <$> eraseTerm ctx ty val
eraseElim 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 ctx (DCloE (Sub term th)) =
eraseElim ctx $ pushSubstsWith' th id term

View File

@ -4,10 +4,12 @@ import Quox.Var
import Quox.Context
import Quox.Name
import Quox.Pretty
import Quox.Syntax.Subst
import Data.Vect
import Data.DPair
import Data.SortedMap
import Data.SnocVect
import Derive.Prelude
%hide TT.Name
@ -15,40 +17,54 @@ import Derive.Prelude
%language ElabReflection
public export
data Binder = Bind BaseName
Eq Binder where _ == _ = True
Ord Binder where compare _ _ = EQ
%runElab derive "Binder" [Show]
public export
data Term : Nat -> Type where
F : (x : Name) -> Term n
B : (i : Var n) -> Term n
F : (x : Name) -> Loc -> Term n
B : (i : Var n) -> Loc -> Term n
Lam : (x : Binder) -> (body : Term (S n)) -> Term n
App : (fun, arg : Term n) -> Term n
Lam : (x : BindName) -> (body : Term (S n)) -> Loc -> Term n
App : (fun, arg : Term n) -> Loc -> Term n
Pair : (fst, snd : Term n) -> Term n
Fst : (pair : Term n) -> Term n
Snd : (pair : Term n) -> Term n
Pair : (fst, snd : Term n) -> Loc -> Term n
Fst : (pair : Term n) -> Loc -> Term n
Snd : (pair : Term n) -> Loc -> Term n
Tag : (tag : String) -> Term n
CaseEnum : (tag : Term n) -> (cases : List (String, Term n)) -> Term n
Tag : (tag : String) -> Loc -> Term n
CaseEnum : (tag : Term n) -> (cases : List (String, Term n)) -> Loc -> Term n
||| empty match with an erased head
Absurd : Loc -> Term n
Zero : Term n
Succ : (nat : Term n) -> Term n
Zero : Loc -> Term n
Succ : (nat : Term n) -> Loc -> Term n
CaseNat : (nat : Term n) ->
(zer : Term n) ->
(x, ih : Binder) -> (suc : Term (2 + n)) ->
(x, ih : BindName) -> (suc : Term (2 + n)) ->
Loc ->
Term n
||| replacement for terms of type [0.A], for now
ErasedBox : Term n
Erased : Loc -> Term n
%name Term s, t, u
%runElab deriveIndexed "Term" [Eq, Ord, Show]
export
Located (Term n) where
(F x loc).loc = loc
(B i loc).loc = loc
(Lam x body loc).loc = loc
(App fun arg loc).loc = loc
(Pair fst snd loc).loc = loc
(Fst pair loc).loc = loc
(Snd pair loc).loc = loc
(Tag tag loc).loc = loc
(CaseEnum tag cases loc).loc = loc
(Absurd loc).loc = loc
(Zero loc).loc = loc
(Succ nat loc).loc = loc
(CaseNat nat zer x ih suc loc).loc = loc
(Erased loc).loc = loc
public export
0 Definitions : Type
Definitions = SortedMap Name $ Term 0
@ -56,37 +72,33 @@ Definitions = SortedMap Name $ Term 0
parameters {opts : LayoutOpts}
export
prettyBind : Binder -> Eff Pretty (Doc opts)
prettyBind (Bind x) = hl TVar $ text $ baseStr x
prettyTerm : BContext n -> Term n -> Eff Pretty (Doc opts)
export
prettyTerm : Context' Binder n -> Term n -> Eff Pretty (Doc opts)
export
prettyArg : Context' Binder n -> Term n -> Eff Pretty (Doc opts)
prettyArg : BContext n -> Term n -> Eff Pretty (Doc opts)
prettyArg xs arg = withPrec Arg $ prettyTerm xs arg
export
prettyApp' : Context' Binder n -> Doc opts -> Term n -> Eff Pretty (Doc opts)
prettyApp' : Context' BindName n -> Doc opts -> Term n -> Eff Pretty (Doc opts)
prettyApp' xs fun arg =
parensIfM App =<< do
arg <- prettyArg xs arg
pure $ sep [fun, arg]
export
prettyApp : Context' Binder n -> Term n -> Term n -> Eff Pretty (Doc opts)
prettyApp : Context' BindName n -> Term n -> Term n -> Eff Pretty (Doc opts)
prettyApp xs fun arg = prettyApp' xs !(prettyArg xs fun) arg
public export
PrettyCaseArm : Nat -> Type
PrettyCaseArm n = Exists $ \s => (Vect s Binder, Term (s + n))
PrettyCaseArm n = Exists $ \s => (Vect s BindName, Term (s + n))
export %inline
caseArm : Vect s Binder -> Term (s + n) -> PrettyCaseArm n
caseArm : Vect s BindName -> Term (s + n) -> PrettyCaseArm n
caseArm xs t = Evidence _ (xs, t)
export
prettyCase : Context' Binder n ->
prettyCase : Context' BindName n ->
(a -> Eff Pretty (Doc opts)) ->
Term n -> List (a, PrettyCaseArm n) ->
Eff Pretty (Doc opts)
@ -100,30 +112,81 @@ parameters {opts : LayoutOpts}
body <- braces $ separateLoose !semiD cases
pure $ sep [header, body]
prettyTerm _ (F x) = prettyFree x
prettyTerm xs (B i) = prettyBind $ xs !!! i
prettyTerm xs (Lam x body) =
prettyTerm _ (F x _) = prettyFree x
prettyTerm xs (B i _) = prettyTBind $ xs !!! i
prettyTerm xs (Lam x body _) =
parensIfM Outer =<< do
header <- hsep <$> sequence [lamD, prettyBind x, darrowD]
header <- hsep <$> sequence [lamD, prettyTBind x, darrowD]
body <- withPrec Outer $ prettyTerm (xs :< x) body
hangDSingle header body
prettyTerm xs (App fun arg) = prettyApp xs fun arg
prettyTerm xs (Pair fst snd) =
prettyTerm xs (App fun arg _) = prettyApp xs fun arg
prettyTerm xs (Pair fst snd _) =
parens =<< separateTight !commaD <$>
sequence {t = List} [prettyTerm xs fst, prettyTerm xs snd]
prettyTerm xs (Fst pair) = prettyApp' xs !fstD pair
prettyTerm xs (Snd pair) = prettyApp' xs !sndD pair
prettyTerm xs (Tag tag) = prettyTag tag
prettyTerm xs (CaseEnum tag cases) = assert_total
prettyTerm xs (Fst pair _) = prettyApp' xs !fstD pair
prettyTerm xs (Snd pair _) = prettyApp' xs !sndD pair
prettyTerm xs (Tag tag _) = prettyTag tag
prettyTerm xs (CaseEnum tag cases _) = assert_total
prettyCase xs prettyTag tag $ map (mapSnd $ caseArm []) cases
prettyTerm xs Zero = zeroD
prettyTerm xs (Succ nat) = prettyApp' xs !succD nat
prettyTerm xs (CaseNat nat zer x ih suc) = assert_total
prettyTerm xs (Absurd _) = hl Syntax "absurd"
prettyTerm xs (Zero _) = zeroD
prettyTerm xs (Succ nat _) = prettyApp' xs !succD nat
prettyTerm xs (CaseNat nat zer x ih suc _) = assert_total
prettyCase xs pure nat
[(!zeroD, caseArm [] zer),
(!sucPat, caseArm [x, ih] suc)]
where
sucPat = separateTight {t = List} !commaD <$>
sequence [[|succD <++> prettyBind x|], prettyBind ih]
prettyTerm _ ErasedBox =
sequence [[|succD <++> prettyTBind x|], prettyTBind ih]
prettyTerm _ (Erased _) =
hl Syntax =<< ifUnicode "" "[]"
public export
USubst : Nat -> Nat -> Type
USubst = Subst Term
public export FromVar Term where fromVarLoc = B
public export
CanSubstSelf Term where
s // th = case s of
F x loc =>
F x loc
B i loc =>
getLoc th i loc
Lam x body loc =>
Lam x (assert_total $ body // push th) loc
App fun arg loc =>
App (fun // th) (arg // th) loc
Pair fst snd loc =>
Pair (fst // th) (snd // th) loc
Fst pair loc =>
Fst (pair // th) loc
Snd pair loc =>
Snd (pair // th) loc
Tag tag loc =>
Tag tag loc
CaseEnum tag cases loc =>
CaseEnum (tag // th) (map (assert_total mapSnd (// th)) cases) loc
Absurd loc =>
Absurd loc
Zero loc =>
Zero loc
Succ nat loc =>
Succ (nat // th) loc
CaseNat nat zer x ih suc loc =>
CaseNat (nat // th) (zer // th)
x ih (assert_total $ suc // pushN 2 th) loc
Erased loc =>
Erased loc
public export
subN : SnocVect s (Term n) -> Term (s + n) -> Term n
subN th t = t // fromSnocVect th
public export
sub1 : Term n -> Term (S n) -> Term n
sub1 e = subN [< e]

View File

@ -14,8 +14,8 @@ export covering
computeElimType : CanWhnf Term Interface.isRedexT =>
CanWhnf Elim Interface.isRedexE =>
{d, n : Nat} ->
(defs : Definitions) -> WhnfContext d n -> (pi : SQty) ->
(e : Elim d n) -> (0 ne : No (isRedexE defs pi e)) =>
(defs : Definitions) -> WhnfContext d n ->
(e : Elim d n) -> (0 ne : No (isRedexE defs sg e)) =>
Eff Whnf (Term d n)
@ -24,11 +24,11 @@ export covering
computeWhnfElimType0 : CanWhnf Term Interface.isRedexT =>
CanWhnf Elim Interface.isRedexE =>
{d, n : Nat} ->
(defs : Definitions) -> WhnfContext d n -> (pi : SQty) ->
(e : Elim d n) -> (0 ne : No (isRedexE defs pi e)) =>
(defs : Definitions) -> WhnfContext d n ->
(e : Elim d n) -> (0 ne : No (isRedexE defs sg e)) =>
Eff Whnf (Term d n)
computeElimType defs ctx pi e {ne} =
computeElimType defs ctx e {ne} =
case e of
F x u loc => do
let Just def = lookup x defs
@ -39,7 +39,7 @@ computeElimType defs ctx pi e {ne} =
pure $ ctx.tctx !! i
App f s loc =>
case !(computeWhnfElimType0 defs ctx pi f {ne = noOr1 ne}) of
case !(computeWhnfElimType0 defs ctx f {ne = noOr1 ne}) of
Pi {arg, res, _} => pure $ sub1 res $ Ann s arg loc
ty => throw $ ExpectedPi loc ctx.names ty
@ -47,12 +47,12 @@ computeElimType defs ctx pi e {ne} =
pure $ sub1 ret pair
Fst pair loc =>
case !(computeWhnfElimType0 defs ctx pi pair {ne = noOr1 ne}) of
case !(computeWhnfElimType0 defs ctx pair {ne = noOr1 ne}) of
Sig {fst, _} => pure fst
ty => throw $ ExpectedSig loc ctx.names ty
Snd pair loc =>
case !(computeWhnfElimType0 defs ctx pi pair {ne = noOr1 ne}) of
case !(computeWhnfElimType0 defs ctx pair {ne = noOr1 ne}) of
Sig {snd, _} => pure $ sub1 snd $ Fst pair loc
ty => throw $ ExpectedSig loc ctx.names ty
@ -66,7 +66,7 @@ computeElimType defs ctx pi e {ne} =
pure $ sub1 ret box
DApp {fun = f, arg = p, loc} =>
case !(computeWhnfElimType0 defs ctx pi f {ne = noOr1 ne}) of
case !(computeWhnfElimType0 defs ctx f {ne = noOr1 ne}) of
Eq {ty, _} => pure $ dsub1 ty p
t => throw $ ExpectedEq loc ctx.names t
@ -82,5 +82,5 @@ computeElimType defs ctx pi e {ne} =
TypeCase {ret, _} =>
pure ret
computeWhnfElimType0 defs ctx pi e =
computeElimType defs ctx pi e >>= whnf0 defs ctx pi
computeWhnfElimType0 defs ctx e =
computeElimType defs ctx e {sg} >>= whnf0 defs ctx SZero

View File

@ -157,7 +157,7 @@ CanWhnf Elim Interface.isRedexE where
eqCoe defs ctx sg ty p' q' val p appLoc
Right ndlh => case p of
K e _ => do
Eq {l, r, ty, _} <- computeWhnfElimType0 defs ctx sg f
Eq {l, r, ty, _} <- computeWhnfElimType0 defs ctx f {ne = fnf}
| ty => throw $ ExpectedEq ty.loc ctx.names ty
whnf defs ctx sg $
ends (Ann (setLoc appLoc l) ty.zero appLoc)

View File

@ -39,7 +39,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
Eff Whnf (Term d n, ScopeTerm d n)
tycasePi (Pi {arg, res, _}) = pure (arg, res)
tycasePi (E e) {tnf} = do
ty <- computeElimType defs ctx SZero e {ne = noOr2 tnf}
ty <- computeElimType defs ctx e {ne = noOr2 tnf}
let loc = e.loc
narg = mnb "Arg" loc; nret = mnb "Ret" loc
arg = E $ typeCase1Y e ty KPi [< !narg, !nret] (BVT 1 loc) loc
@ -57,7 +57,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
Eff Whnf (Term d n, ScopeTerm d n)
tycaseSig (Sig {fst, snd, _}) = pure (fst, snd)
tycaseSig (E e) {tnf} = do
ty <- computeElimType defs ctx SZero e {ne = noOr2 tnf}
ty <- computeElimType defs ctx e {ne = noOr2 tnf}
let loc = e.loc
nfst = mnb "Fst" loc; nsnd = mnb "Snd" loc
fst = E $ typeCase1Y e ty KSig [< !nfst, !nsnd] (BVT 1 loc) loc
@ -75,7 +75,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
Eff Whnf (Term d n)
tycaseBOX (BOX {ty, _}) = pure ty
tycaseBOX (E e) {tnf} = do
ty <- computeElimType defs ctx SZero e {ne = noOr2 tnf}
ty <- computeElimType defs ctx e {ne = noOr2 tnf}
pure $ E $ typeCase1Y e ty KBOX [< !(mnb "Ty" e.loc)] (BVT 0 e.loc) e.loc
tycaseBOX t = throw $ ExpectedBOX t.loc ctx.names t
@ -87,7 +87,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
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
ty <- computeElimType defs ctx SZero e {ne = noOr2 tnf}
ty <- computeElimType defs ctx e {ne = noOr2 tnf}
let loc = e.loc
names = traverse' (\x => mnb x loc) [< "A0", "A1", "A", "L", "R"]
a0 = E $ typeCase1Y e ty KEq !names (BVT 4 loc) loc

View File

@ -57,4 +57,5 @@ modules =
Quox.Parser.FromParser,
Quox.Parser.FromParser.Error,
Quox.Parser,
Quox.Untyped.Syntax
Quox.Untyped.Syntax,
Quox.Untyped.Erase