do scope checking in FromParser where it belongs

This commit is contained in:
rhiannon morris 2023-04-18 22:55:23 +02:00
parent 55c0bf9974
commit b666bc20cf
11 changed files with 305 additions and 309 deletions

View file

@ -55,40 +55,21 @@ isZero g = g.qty.fst == Zero
public export public export
data DefEnvTag = NS | DEFS data DefEnvTag = DEFS
public export public export
Definitions : Type Definitions : Type
Definitions = SortedMap Name Definition Definitions = SortedMap Name Definition
public export
0 CurNSReader : Type -> Type
CurNSReader = ReaderL NS Mods
public export public export
0 DefsReader : Type -> Type 0 DefsReader : Type -> Type
DefsReader = ReaderL DEFS Definitions DefsReader = ReaderL DEFS Definitions
export
ns : Has CurNSReader fs => Eff fs Mods
ns = askAt NS
export export
defs : Has DefsReader fs => Eff fs Definitions defs : Has DefsReader fs => Eff fs Definitions
defs = askAt DEFS defs = askAt DEFS
public export
lookupNS : Mods -> Name -> Definitions -> Maybe Definition
lookupNS [<] x defs = lookup x defs
lookupNS mms@(ms :< _) x defs =
lookup (addMods mms x) defs <|> lookupNS ms x defs
parameters {d, n : Nat}
public export %inline public export %inline
lookupElim : Name -> Definitions -> Maybe (Elim d n) lookupElim : {d, n : Nat} -> Name -> Definitions -> Maybe (Elim d n)
lookupElim x defs = toElim !(lookup x defs) lookupElim x defs = toElim !(lookup x defs)
public export %inline
lookupElimNS : Mods -> Name -> Definitions -> Maybe (Elim d n)
lookupElimNS mods x defs = toElim !(lookupNS mods x defs)

View file

@ -81,19 +81,19 @@ sameTyCon (E {}) _ = False
||| * a box type is a subsingleton if its content is ||| * a box type is a subsingleton if its content is
public export covering public export covering
isSubSing : Has ErrorEff fs => {n : Nat} -> isSubSing : Has ErrorEff fs => {n : Nat} ->
Mods -> Definitions -> EqContext n -> Term 0 n -> Eff fs Bool Definitions -> EqContext n -> Term 0 n -> Eff fs Bool
isSubSing ns defs ctx ty0 = do isSubSing defs ctx ty0 = do
Element ty0 nc <- whnf ns defs ctx ty0 Element ty0 nc <- whnf defs ctx ty0
case ty0 of case ty0 of
TYPE _ => pure False TYPE _ => pure False
Pi _ arg res => isSubSing ns defs (extendTy Zero res.name arg ctx) res.term Pi _ arg res => isSubSing defs (extendTy Zero res.name arg ctx) res.term
Sig fst snd => isSubSing ns defs ctx fst `andM` Sig fst snd => isSubSing defs ctx fst `andM`
isSubSing ns defs (extendTy Zero snd.name fst ctx) snd.term isSubSing defs (extendTy Zero snd.name fst ctx) snd.term
Enum tags => pure $ length (SortedSet.toList tags) <= 1 Enum tags => pure $ length (SortedSet.toList tags) <= 1
Eq {} => pure True Eq {} => pure True
Nat => pure False Nat => pure False
BOX _ ty => isSubSing ns defs ctx ty BOX _ ty => isSubSing defs ctx ty
E (s :# _) => isSubSing ns defs ctx s E (s :# _) => isSubSing defs ctx s
E _ => pure False E _ => pure False
Lam _ => pure False Lam _ => pure False
Pair _ _ => pure False Pair _ _ => pure False
@ -116,14 +116,14 @@ ensureTyCon ctx t = case nchoose $ isTyConE t of
||| |||
||| ⚠ **assumes the elim is already typechecked.** ⚠ ||| ⚠ **assumes the elim is already typechecked.** ⚠
private covering private covering
computeElimTypeE : (ns : Mods) -> (defs : Definitions) -> EqContext n -> computeElimTypeE : (defs : Definitions) -> EqContext n ->
(e : Elim 0 n) -> (0 ne : NotRedex ns defs e) => (e : Elim 0 n) -> (0 ne : NotRedex defs e) =>
Equal (Term 0 n) Equal (Term 0 n)
computeElimTypeE ns defs ectx e = computeElimTypeE defs ectx e =
let Val n = ectx.termLen in let Val n = ectx.termLen in
rethrow $ computeElimType ns defs (toWhnfContext ectx) e rethrow $ computeElimType defs (toWhnfContext ectx) e
parameters (ns : Mods) (defs : Definitions) parameters (defs : Definitions)
mutual mutual
namespace Term namespace Term
||| `compare0 ctx ty s t` compares `s` and `t` at type `ty`, according to ||| `compare0 ctx ty s t` compares `s` and `t` at type `ty`, according to
@ -135,9 +135,9 @@ parameters (ns : Mods) (defs : Definitions)
compare0 ctx ty s t = compare0 ctx ty s t =
wrapErr (WhileComparingT ctx !mode ty s t) $ do wrapErr (WhileComparingT ctx !mode ty s t) $ do
let Val n = ctx.termLen let Val n = ctx.termLen
Element ty _ <- whnf ns defs ctx ty Element ty _ <- whnf defs ctx ty
Element s _ <- whnf ns defs ctx s Element s _ <- whnf defs ctx s
Element t _ <- whnf ns defs ctx t Element t _ <- whnf defs ctx t
tty <- ensureTyCon ctx ty tty <- ensureTyCon ctx ty
compare0' ctx ty s t compare0' ctx ty s t
@ -150,8 +150,8 @@ parameters (ns : Mods) (defs : Definitions)
private covering private covering
compare0' : EqContext n -> compare0' : EqContext n ->
(ty, s, t : Term 0 n) -> (ty, s, t : Term 0 n) ->
(0 _ : NotRedex ns defs ty) => (0 _ : So (isTyConE ty)) => (0 _ : NotRedex defs ty) => (0 _ : So (isTyConE ty)) =>
(0 _ : NotRedex ns defs s) => (0 _ : NotRedex ns defs t) => (0 _ : NotRedex defs s) => (0 _ : NotRedex defs t) =>
Equal () Equal ()
compare0' ctx (TYPE _) s t = compareType ctx s t compare0' ctx (TYPE _) s t = compareType ctx s t
@ -275,8 +275,8 @@ parameters (ns : Mods) (defs : Definitions)
compareType : EqContext n -> (s, t : Term 0 n) -> Equal () compareType : EqContext n -> (s, t : Term 0 n) -> Equal ()
compareType ctx s t = do compareType ctx s t = do
let Val n = ctx.termLen let Val n = ctx.termLen
Element s _ <- whnf ns defs ctx s Element s _ <- whnf defs ctx s
Element t _ <- whnf ns defs ctx t Element t _ <- whnf defs ctx t
ts <- ensureTyCon ctx s ts <- ensureTyCon ctx s
tt <- ensureTyCon ctx t tt <- ensureTyCon ctx t
st <- either pure (const $ clashTy ctx s t) $ nchoose $ sameTyCon s t st <- either pure (const $ clashTy ctx s t) $ nchoose $ sameTyCon s t
@ -284,8 +284,8 @@ parameters (ns : Mods) (defs : Definitions)
private covering private covering
compareType' : EqContext n -> (s, t : Term 0 n) -> compareType' : EqContext n -> (s, t : Term 0 n) ->
(0 _ : NotRedex ns defs s) => (0 _ : So (isTyConE s)) => (0 _ : NotRedex defs s) => (0 _ : So (isTyConE s)) =>
(0 _ : NotRedex ns defs t) => (0 _ : So (isTyConE t)) => (0 _ : NotRedex defs t) => (0 _ : So (isTyConE t)) =>
(0 _ : So (sameTyCon s t)) => (0 _ : So (sameTyCon s t)) =>
Equal () Equal ()
-- equality is the same as subtyping, except with the -- equality is the same as subtyping, except with the
@ -349,10 +349,10 @@ parameters (ns : Mods) (defs : Definitions)
private covering private covering
replaceEnd : EqContext n -> replaceEnd : EqContext n ->
(e : Elim 0 n) -> DimConst -> (0 ne : NotRedex ns defs e) -> (e : Elim 0 n) -> DimConst -> (0 ne : NotRedex defs e) ->
Equal (Elim 0 n) Equal (Elim 0 n)
replaceEnd ctx e p ne = do replaceEnd ctx e p ne = do
(ty, l, r) <- expectEq ns defs ctx !(computeElimTypeE ns defs ctx e) (ty, l, r) <- expectEq defs ctx !(computeElimTypeE defs ctx e)
pure $ ends l r p :# dsub1 ty (K p) pure $ ends l r p :# dsub1 ty (K p)
namespace Elim namespace Elim
@ -368,16 +368,16 @@ parameters (ns : Mods) (defs : Definitions)
compare0 ctx e f = compare0 ctx e f =
wrapErr (WhileComparingE ctx !mode e f) $ do wrapErr (WhileComparingE ctx !mode e f) $ do
let Val n = ctx.termLen let Val n = ctx.termLen
Element e ne <- whnf ns defs ctx e Element e ne <- whnf defs ctx e
Element f nf <- whnf ns defs ctx f Element f nf <- whnf defs ctx 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 ns defs ctx !(computeElimTypeE ns defs ctx e)) $ unless !(isSubSing defs ctx !(computeElimTypeE defs ctx e)) $
compare0' ctx e f ne nf compare0' ctx e f ne nf
private covering private covering
compare0' : EqContext n -> compare0' : EqContext n ->
(e, f : Elim 0 n) -> (e, f : Elim 0 n) ->
(0 ne : NotRedex ns defs e) -> (0 nf : NotRedex ns defs f) -> (0 ne : NotRedex defs e) -> (0 nf : NotRedex defs f) ->
Equal () Equal ()
-- replace applied equalities with the appropriate end first -- replace applied equalities with the appropriate end first
-- e.g. e : Eq [i ⇒ A] s t ⊢ e 𝟎 = s : A𝟎/i -- e.g. e : Eq [i ⇒ A] s t ⊢ e 𝟎 = s : A𝟎/i
@ -397,8 +397,8 @@ parameters (ns : Mods) (defs : Definitions)
compare0' ctx (e :@ s) (f :@ t) ne nf = compare0' ctx (e :@ s) (f :@ t) ne nf =
local_ Equal $ do local_ Equal $ do
compare0 ctx e f compare0 ctx e f
(_, arg, _) <- expectPi ns defs ctx =<< (_, arg, _) <- expectPi defs ctx =<<
computeElimTypeE ns defs ctx e @{noOr1 ne} computeElimTypeE defs ctx e @{noOr1 ne}
Term.compare0 ctx arg s t Term.compare0 ctx arg s t
compare0' ctx e@(_ :@ _) f _ _ = clashE ctx e f compare0' ctx e@(_ :@ _) f _ _ = clashE ctx e f
@ -406,9 +406,9 @@ parameters (ns : Mods) (defs : Definitions)
(CasePair fpi f fret fbody) ne nf = (CasePair fpi f fret fbody) ne nf =
local_ Equal $ do local_ Equal $ do
compare0 ctx e f compare0 ctx e f
ety <- computeElimTypeE ns defs ctx e @{noOr1 ne} ety <- computeElimTypeE defs ctx e @{noOr1 ne}
compareType (extendTy Zero eret.name ety ctx) eret.term fret.term compareType (extendTy Zero eret.name ety ctx) eret.term fret.term
(fst, snd) <- expectSig ns defs ctx ety (fst, snd) <- expectSig defs ctx ety
let [< x, y] = ebody.names let [< x, y] = ebody.names
Term.compare0 (extendTyN [< (epi, x, fst), (epi, y, snd.term)] ctx) Term.compare0 (extendTyN [< (epi, x, fst), (epi, y, snd.term)] ctx)
(substCasePairRet ety eret) (substCasePairRet ety eret)
@ -420,9 +420,9 @@ parameters (ns : Mods) (defs : Definitions)
(CaseEnum fpi f fret farms) ne nf = (CaseEnum fpi f fret farms) ne nf =
local_ Equal $ do local_ Equal $ do
compare0 ctx e f compare0 ctx e f
ety <- computeElimTypeE ns defs ctx e @{noOr1 ne} ety <- computeElimTypeE defs ctx e @{noOr1 ne}
compareType (extendTy Zero eret.name ety ctx) eret.term fret.term compareType (extendTy Zero eret.name ety ctx) eret.term fret.term
for_ !(expectEnum ns defs ctx ety) $ \t => for_ !(expectEnum defs ctx ety) $ \t =>
compare0 ctx (sub1 eret $ Tag t :# ety) compare0 ctx (sub1 eret $ Tag t :# ety)
!(lookupArm t earms) !(lookupArm t farms) !(lookupArm t earms) !(lookupArm t farms)
expectEqualQ epi fpi expectEqualQ epi fpi
@ -437,7 +437,7 @@ parameters (ns : Mods) (defs : Definitions)
(CaseNat fpi fpi' f fret fzer fsuc) ne nf = (CaseNat fpi fpi' f fret fzer fsuc) ne nf =
local_ Equal $ do local_ Equal $ do
compare0 ctx e f compare0 ctx e f
ety <- computeElimTypeE ns defs ctx e @{noOr1 ne} ety <- computeElimTypeE defs ctx e @{noOr1 ne}
compareType (extendTy Zero eret.name ety ctx) eret.term fret.term compareType (extendTy Zero eret.name ety ctx) eret.term fret.term
compare0 ctx (sub1 eret (Zero :# Nat)) ezer fzer compare0 ctx (sub1 eret (Zero :# Nat)) ezer fzer
let [< p, ih] = esuc.names let [< p, ih] = esuc.names
@ -452,9 +452,9 @@ parameters (ns : Mods) (defs : Definitions)
(CaseBox fpi f fret fbody) ne nf = (CaseBox fpi f fret fbody) ne nf =
local_ Equal $ do local_ Equal $ do
compare0 ctx e f compare0 ctx e f
ety <- computeElimTypeE ns defs ctx e @{noOr1 ne} ety <- computeElimTypeE defs ctx e @{noOr1 ne}
compareType (extendTy Zero eret.name ety ctx) eret.term fret.term compareType (extendTy Zero eret.name ety ctx) eret.term fret.term
(q, ty) <- expectBOX ns defs ctx ety (q, ty) <- expectBOX defs ctx ety
compare0 (extendTy (epi * q) ebody.name ty ctx) compare0 (extendTy (epi * q) ebody.name ty ctx)
(substCaseBoxRet ety eret) (substCaseBoxRet ety eret)
ebody.term fbody.term ebody.term fbody.term
@ -480,8 +480,7 @@ parameters (ns : Mods) (defs : Definitions)
ne _ = ne _ =
local_ Equal $ do local_ Equal $ do
compare0 ctx ty1 ty2 compare0 ctx ty1 ty2
u <- expectTYPE ns defs ctx =<< u <- expectTYPE defs ctx =<< computeElimTypeE defs ctx ty1 @{noOr1 ne}
computeElimTypeE ns defs ctx ty1 @{noOr1 ne}
compareType ctx ret1 ret2 compareType ctx ret1 ret2
compareType ctx def1 def2 compareType ctx def1 def2
for_ allKinds $ \k => for_ allKinds $ \k =>
@ -550,8 +549,7 @@ parameters (ns : Mods) (defs : Definitions)
compare0 ctx (weakT 1 ret) b1.term b1.term compare0 ctx (weakT 1 ret) b1.term b1.term
parameters {auto _ : (Has DefsReader fs, Has CurNSReader fs, Has ErrorEff fs)} parameters {auto _ : (Has DefsReader fs, Has ErrorEff fs)} (ctx : TyContext d n)
(ctx : TyContext d n)
-- [todo] only split on the dvars that are actually used anywhere in -- [todo] only split on the dvars that are actually used anywhere in
-- the calls to `splits` -- the calls to `splits`
@ -563,7 +561,7 @@ parameters {auto _ : (Has DefsReader fs, Has CurNSReader fs, Has ErrorEff fs)}
map fst $ runState @{Z} mode $ map fst $ runState @{Z} mode $
for_ (splits ctx.dctx) $ \th => for_ (splits ctx.dctx) $ \th =>
let ectx = makeEqContext ctx th in let ectx = makeEqContext ctx th in
lift $ compare0 !ns !defs ectx (ty // th) (s // th) (t // th) lift $ compare0 !defs ectx (ty // th) (s // th) (t // th)
export covering export covering
compareType : (s, t : Term d n) -> Eff fs () compareType : (s, t : Term d n) -> Eff fs ()
@ -571,7 +569,7 @@ parameters {auto _ : (Has DefsReader fs, Has CurNSReader fs, Has ErrorEff fs)}
map fst $ runState @{Z} mode $ map fst $ runState @{Z} mode $
for_ (splits ctx.dctx) $ \th => for_ (splits ctx.dctx) $ \th =>
let ectx = makeEqContext ctx th in let ectx = makeEqContext ctx th in
lift $ compareType !ns !defs ectx (s // th) (t // th) lift $ compareType !defs ectx (s // th) (t // th)
namespace Elim namespace Elim
||| you don't have to pass the type in but the arguments must still be ||| you don't have to pass the type in but the arguments must still be
@ -582,7 +580,7 @@ parameters {auto _ : (Has DefsReader fs, Has CurNSReader fs, Has ErrorEff fs)}
map fst $ runState @{Z} mode $ map fst $ runState @{Z} mode $
for_ (splits ctx.dctx) $ \th => for_ (splits ctx.dctx) $ \th =>
let ectx = makeEqContext ctx th in let ectx = makeEqContext ctx th in
lift $ compare0 !ns !defs ectx (e // th) (f // th) lift $ compare0 !defs ectx (e // th) (f // th)
namespace Term namespace Term
export covering %inline export covering %inline

View file

@ -20,8 +20,6 @@ import public Quox.Parser.FromParser.Error as Quox.Parser.FromParser
%hide Typing.Error %hide Typing.Error
%hide Lexer.Error %hide Lexer.Error
%hide Parser.Error %hide Parser.Error
%hide Definition.DEFS
%hide Definition.NS
public export public export
@ -38,7 +36,12 @@ SeenFiles = SortedSet String
public export public export
data StateTag = DEFS | NS | SEEN data StateTag = NS | SEEN
public export
0 FromParserPure : List (Type -> Type)
FromParserPure =
[Except Error, StateL DEFS Definitions, StateL NS Mods]
public export public export
0 FromParserEff : List (Type -> Type) 0 FromParserEff : List (Type -> Type)
@ -70,16 +73,25 @@ fromPDimWith ds (V i) =
private private
avoidDim : Has (Except Error) fs => avoidDim : Has (Except Error) fs =>
Context' BName d -> PName -> Eff fs (Term d n) Context' BName d -> PName -> Eff fs Name
avoidDim ds x = avoidDim ds x =
fromName (const $ throw $ DimNameInTerm x.base) (pure . FT . fromPName) ds x fromName (const $ throw $ DimNameInTerm x.base) (pure . fromPName) ds x
private
resolveName : Mods -> Name -> Eff FromParserPure (Term d n)
resolveName ns x =
let here = addMods ns x in
if isJust $ lookup here !(getAt DEFS) then
pure $ FT here
else do
let ns :< _ = ns
| _ => throw $ TermNotInScope x
resolveName ns x
mutual mutual
export export
fromPTermWith : Has (Except Error) fs => fromPTermWith : Context' BName d -> Context' BName n ->
Context' BName d -> Context' BName n -> PTerm -> Eff FromParserPure (Term d n)
PTerm -> Eff fs (Term d n)
fromPTermWith ds ns t0 = case t0 of fromPTermWith ds ns t0 = case t0 of
TYPE k => TYPE k =>
pure $ TYPE $ k pure $ TYPE $ k
@ -154,7 +166,7 @@ mutual
map E $ (:%) <$> fromPTermElim ds ns s <*> fromPDimWith ds p map E $ (:%) <$> fromPTermElim ds ns s <*> fromPDimWith ds p
V x => V x =>
fromName (pure . E . B) (avoidDim ds) ns x fromName (pure . E . B) (resolveName !(getAt NS) <=< avoidDim ds) ns x
s :# a => s :# a =>
map E $ (:#) <$> fromPTermWith ds ns s <*> fromPTermWith ds ns a map E $ (:#) <$> fromPTermWith ds ns s <*> fromPTermWith ds ns a
@ -177,27 +189,29 @@ mutual
<*> fromPTermDScope ds ns [< j1] val1 <*> fromPTermDScope ds ns [< j1] val1
private private
fromPTermEnumArms : Has (Except Error) fs => fromPTermEnumArms : Context' BName d -> Context' BName n ->
Context' BName d -> Context' BName n -> List (TagVal, PTerm) ->
List (TagVal, PTerm) -> Eff fs (CaseEnumArms d n) Eff FromParserPure (CaseEnumArms d n)
fromPTermEnumArms ds ns = fromPTermEnumArms ds ns =
map SortedMap.fromList . traverse (traverse $ fromPTermWith ds ns) map SortedMap.fromList . traverse (traverse $ fromPTermWith ds ns)
private private
fromPTermElim : Has (Except Error) fs => fromPTermElim : Context' BName d -> Context' BName n ->
Context' BName d -> Context' BName n -> PTerm -> Eff FromParserPure (Elim d n)
PTerm -> Eff fs (Elim d n)
fromPTermElim ds ns e = fromPTermElim ds ns e =
case !(fromPTermWith ds ns e) of case !(fromPTermWith ds ns e) of
E e => pure e E e => pure e
_ => throw $ AnnotationNeeded e t => let ctx = MkNameContexts (map name ds) (map name ns) in
throw $ AnnotationNeeded ctx t
where
name : BName -> BaseName
name = maybe Unused UN
-- [todo] use SN if the var is named but still unused -- [todo] use SN if the var is named but still unused
private private
fromPTermTScope : {s : Nat} -> Has (Except Error) fs => fromPTermTScope : {s : Nat} -> Context' BName d -> Context' BName n ->
Context' BName d -> Context' BName n -> SnocVect s BName -> PTerm ->
SnocVect s BName -> Eff FromParserPure (ScopeTermN s d n)
PTerm -> Eff fs (ScopeTermN s d n)
fromPTermTScope ds ns xs t = fromPTermTScope ds ns xs t =
if all isNothing xs then if all isNothing xs then
SN <$> fromPTermWith ds ns t SN <$> fromPTermWith ds ns t
@ -206,10 +220,9 @@ mutual
<$> fromPTermWith ds (ns ++ xs) t <$> fromPTermWith ds (ns ++ xs) t
private private
fromPTermDScope : {s : Nat} -> Has (Except Error) fs => fromPTermDScope : {s : Nat} -> Context' BName d -> Context' BName n ->
Context' BName d -> Context' BName n -> SnocVect s BName -> PTerm ->
SnocVect s BName -> Eff FromParserPure (DScopeTermN s d n)
PTerm -> Eff fs (DScopeTermN s d n)
fromPTermDScope ds ns xs t = fromPTermDScope ds ns xs t =
if all isNothing xs then if all isNothing xs then
SN <$> fromPTermWith ds ns t SN <$> fromPTermWith ds ns t
@ -219,7 +232,7 @@ mutual
export %inline export %inline
fromPTerm : Has (Except Error) fs => PTerm -> Eff fs (Term 0 0) fromPTerm : PTerm -> Eff FromParserPure (Term 0 0)
fromPTerm = fromPTermWith [<] [<] fromPTerm = fromPTermWith [<] [<]
@ -236,10 +249,9 @@ fromPNameNS : Has (StateL NS Mods) fs => PName -> Eff fs Name
fromPNameNS name = pure $ addMods !(getAt NS) $ fromPName name fromPNameNS name = pure $ addMods !(getAt NS) $ fromPName name
private private
injTC : (Has (StateL DEFS Definitions) fs, Has (StateL NS Mods) fs, injTC : (Has (StateL DEFS Definitions) fs, Has (Except Error) fs) =>
Has (Except Error) fs) =>
TC a -> Eff fs a TC a -> Eff fs a
injTC act = rethrow $ mapFst TypeError $ runTC !(getAt NS) !(getAt DEFS) act injTC act = rethrow $ mapFst TypeError $ runTC !(getAt DEFS) act
export covering export covering
fromPDef : (Has (StateL DEFS Definitions) fs, fromPDef : (Has (StateL DEFS Definitions) fs,
@ -251,8 +263,8 @@ fromPDef (MkPDef qty pname ptype pterm) = do
qtyGlobal <- globalPQty qty qtyGlobal <- globalPQty qty
let gqty = Element qty qtyGlobal let gqty = Element qty qtyGlobal
let sqty = globalToSubj gqty let sqty = globalToSubj gqty
type <- traverse fromPTerm ptype type <- lift $ traverse fromPTerm ptype
term <- fromPTerm pterm term <- lift $ fromPTerm pterm
case type of case type of
Just type => do Just type => do
injTC $ checkTypeC empty type Nothing injTC $ checkTypeC empty type Nothing
@ -261,7 +273,7 @@ fromPDef (MkPDef qty pname ptype pterm) = do
modifyAt DEFS $ insert name def modifyAt DEFS $ insert name def
pure (name, def) pure (name, def)
Nothing => do Nothing => do
let E elim = term | _ => throw $ AnnotationNeeded pterm let E elim = term | t => throw $ AnnotationNeeded empty t
res <- injTC $ inferC empty sqty elim res <- injTC $ inferC empty sqty elim
let def = mkDef gqty res.type term let def = mkDef gqty res.type term
modifyAt DEFS $ insert name def modifyAt DEFS $ insert name def
@ -315,6 +327,15 @@ parameters {auto _ : (Has IO fs,
fromPTopLevel (PD decl) = fromPDecl decl fromPTopLevel (PD decl) = fromPDecl decl
fromPTopLevel (PLoad file) = loadProcessFile file fromPTopLevel (PLoad file) = loadProcessFile file
export
fromParserPure : Definitions ->
Eff FromParserPure a ->
Either Error (a, Definitions)
fromParserPure defs act =
extract $
runExcept $
evalStateAt NS [<] $
runStateAt DEFS defs act
export export
fromParserIO : (MonadRec io, HasIO io) => fromParserIO : (MonadRec io, HasIO io) =>

View file

@ -9,8 +9,9 @@ import Quox.Pretty
public export public export
data Error = data Error =
AnnotationNeeded PTerm AnnotationNeeded (NameContexts d n) (Term d n)
| DuplicatesInEnum (List TagVal) | DuplicatesInEnum (List TagVal)
| TermNotInScope Name
| DimNotInScope PBaseName | DimNotInScope PBaseName
| QtyNotGlobal Qty | QtyNotGlobal Qty
| DimNameInTerm PBaseName | DimNameInTerm PBaseName
@ -25,17 +26,19 @@ data Error =
parameters (unicode, showContext : Bool) parameters (unicode, showContext : Bool)
export export
prettyError : Error -> Doc HL prettyError : Error -> Doc HL
prettyError (AnnotationNeeded tm) = prettyError (AnnotationNeeded ctx tm) =
sep ["the term", "«PTerm»", -- pretty0 unicode tm, sep ["the term", prettyTerm unicode ctx.dnames ctx.tnames tm,
"needs a type annotation"] "needs a type annotation"]
-- [todo] print the original PTerm instead
prettyError (DuplicatesInEnum tags) = prettyError (DuplicatesInEnum tags) =
sep ["duplicate tags in enum type", sep ["duplicate tags in enum type", braces $ fillSep $ map pretty tags]
braces $ fillSep $ map pretty tags]
prettyError (DimNotInScope i) = prettyError (DimNotInScope i) =
sep ["dimension", pretty0 unicode $ DV $ fromString i, sep ["dimension", pretty0 unicode $ DV $ fromString i, "not in scope"]
"not in scope"]
prettyError (TermNotInScope x) =
sep ["term variable", pretty0 unicode $ F x {d = 0, n = 0}, "not in scope"]
prettyError (QtyNotGlobal pi) = prettyError (QtyNotGlobal pi) =
sep ["quantity", pretty0 unicode pi, sep ["quantity", pretty0 unicode pi,
@ -46,12 +49,10 @@ parameters (unicode, showContext : Bool)
"used in a term context"] "used in a term context"]
prettyError (TypeError err) = prettyError (TypeError err) =
Typing.prettyError unicode showContext $ Typing.prettyError unicode showContext $ trimContext 2 err
trimContext 2 err
prettyError (LoadError str err) = prettyError (LoadError str err) =
vsep [hsep ["couldn't load file", pretty str], vsep [hsep ["couldn't load file", pretty str], fromString $ show err]
"«FileError»"]
prettyError (ParseError err) = prettyError (ParseError err) =
pretty $ show err pretty $ show err

View file

@ -14,37 +14,36 @@ import Data.List
public export public export
0 RedexTest : TermLike -> Type 0 RedexTest : TermLike -> Type
RedexTest tm = {d, n : Nat} -> Mods -> Definitions -> tm d n -> Bool RedexTest tm = {d, n : Nat} -> Definitions -> tm d n -> Bool
public export public export
interface Whnf (0 tm : TermLike) (0 isRedex : RedexTest tm) | tm interface Whnf (0 tm : TermLike) (0 isRedex : RedexTest tm) | tm
where where
whnf : {d, n : Nat} -> (ns : Mods) -> (defs : Definitions) -> whnf : {d, n : Nat} -> (defs : Definitions) ->
(ctx : WhnfContext d n) -> (ctx : WhnfContext d n) ->
tm d n -> Either Error (Subset (tm d n) (No . isRedex ns defs)) tm d n -> Either Error (Subset (tm d n) (No . isRedex defs))
public export %inline public export %inline
whnf0 : {d, n : Nat} -> {0 isRedex : RedexTest tm} -> Whnf tm isRedex => whnf0 : {d, n : Nat} -> {0 isRedex : RedexTest tm} -> Whnf tm isRedex =>
(ns : Mods) -> (defs : Definitions) -> (defs : Definitions) ->
WhnfContext d n -> tm d n -> Either Error (tm d n) WhnfContext d n -> tm d n -> Either Error (tm d n)
whnf0 ns defs ctx t = fst <$> whnf ns defs ctx t whnf0 defs ctx t = fst <$> whnf defs ctx t
public export public export
0 IsRedex, NotRedex : {isRedex : RedexTest tm} -> Whnf tm isRedex => 0 IsRedex, NotRedex : {isRedex : RedexTest tm} -> Whnf tm isRedex =>
Mods -> Definitions -> Pred (tm d n) Definitions -> Pred (tm d n)
IsRedex ns defs = So . isRedex ns defs IsRedex defs = So . isRedex defs
NotRedex ns defs = No . isRedex ns defs NotRedex defs = No . isRedex defs
public export public export
0 NonRedex : (tm : TermLike) -> {isRedex : RedexTest tm} -> 0 NonRedex : (tm : TermLike) -> {isRedex : RedexTest tm} ->
Whnf tm isRedex => (d, n : Nat) -> Whnf tm isRedex => (d, n : Nat) ->
(ns : Mods) -> (defs : Definitions) -> Type (defs : Definitions) -> Type
NonRedex tm d n ns defs = Subset (tm d n) (NotRedex ns defs) NonRedex tm d n defs = Subset (tm d n) (NotRedex defs)
public export %inline public export %inline
nred : {0 isRedex : RedexTest tm} -> (0 _ : Whnf tm isRedex) => nred : {0 isRedex : RedexTest tm} -> (0 _ : Whnf tm isRedex) =>
(t : tm d n) -> (0 nr : NotRedex ns defs t) => (t : tm d n) -> (0 nr : NotRedex defs t) => NonRedex tm d n defs
NonRedex tm d n ns defs
nred t = Element t nr nred t = Element t nr
@ -136,38 +135,38 @@ isK _ = False
mutual mutual
public export public export
isRedexE : RedexTest Elim isRedexE : RedexTest Elim
isRedexE ns defs (F x) {d, n} = isRedexE defs (F x) {d, n} =
isJust $ lookupElimNS ns x defs {d, n} isJust $ lookupElim x defs {d, n}
isRedexE _ _ (B _) = False isRedexE _ (B _) = False
isRedexE ns defs (f :@ _) = isRedexE defs (f :@ _) =
isRedexE ns defs f || isLamHead f isRedexE defs f || isLamHead f
isRedexE ns defs (CasePair {pair, _}) = isRedexE defs (CasePair {pair, _}) =
isRedexE ns defs pair || isPairHead pair isRedexE defs pair || isPairHead pair
isRedexE ns defs (CaseEnum {tag, _}) = isRedexE defs (CaseEnum {tag, _}) =
isRedexE ns defs tag || isTagHead tag isRedexE defs tag || isTagHead tag
isRedexE ns defs (CaseNat {nat, _}) = isRedexE defs (CaseNat {nat, _}) =
isRedexE ns defs nat || isNatHead nat isRedexE defs nat || isNatHead nat
isRedexE ns defs (CaseBox {box, _}) = isRedexE defs (CaseBox {box, _}) =
isRedexE ns defs box || isBoxHead box isRedexE defs box || isBoxHead box
isRedexE ns defs (f :% p) = isRedexE defs (f :% p) =
isRedexE ns defs f || isDLamHead f || isK p isRedexE defs f || isDLamHead f || isK p
isRedexE ns defs (t :# a) = isRedexE defs (t :# a) =
isE t || isRedexT ns defs t || isRedexT ns defs a isE t || isRedexT defs t || isRedexT defs a
isRedexE ns defs (Coe {val, _}) = isRedexE defs (Coe {val, _}) =
isRedexT ns defs val || not (isE val) isRedexT defs val || not (isE val)
isRedexE ns defs (Comp {ty, r, _}) = isRedexE defs (Comp {ty, r, _}) =
isRedexT ns defs ty || isK r isRedexT defs ty || isK r
isRedexE ns defs (TypeCase {ty, ret, _}) = isRedexE defs (TypeCase {ty, ret, _}) =
isRedexE ns defs ty || isRedexT ns defs ret || isAnnTyCon ty isRedexE defs ty || isRedexT defs ret || isAnnTyCon ty
isRedexE _ _ (CloE {}) = True isRedexE _ (CloE {}) = True
isRedexE _ _ (DCloE {}) = True isRedexE _ (DCloE {}) = True
public export public export
isRedexT : RedexTest Term isRedexT : RedexTest Term
isRedexT _ _ (CloT {}) = True isRedexT _ (CloT {}) = True
isRedexT _ _ (DCloT {}) = True isRedexT _ (DCloT {}) = True
isRedexT ns defs (E e) = isAnn e || isRedexE ns defs e isRedexT defs (E e) = isAnn e || isRedexE defs e
isRedexT _ _ _ = False isRedexT _ _ = False
public export public export
@ -213,23 +212,21 @@ coeScoped ty p q (S names (N body)) =
mutual mutual
parameters {d, n : Nat} (ns : Mods) (defs : Definitions) parameters {d, n : Nat} (defs : Definitions) (ctx : WhnfContext d n)
(ctx : WhnfContext d n)
||| performs the minimum work required to recompute the type of an elim. ||| performs the minimum work required to recompute the type of an elim.
||| |||
||| ⚠ **assumes the elim is already typechecked.** ⚠ ||| ⚠ **assumes the elim is already typechecked.** ⚠
export covering export covering
computeElimType : (e : Elim d n) -> (0 ne : No (isRedexE ns defs e)) => computeElimType : (e : Elim d n) -> (0 ne : No (isRedexE defs e)) =>
Either Error (Term d n) Either Error (Term d n)
computeElimType (F x) = do computeElimType (F x) = do
let Just def = lookupNS ns x defs | Nothing => Left $ NotInScope x let Just def = lookup x defs | Nothing => Left $ NotInScope x
pure $ def.type pure $ def.type
computeElimType (B i) = pure $ ctx.tctx !! i computeElimType (B i) = pure $ ctx.tctx !! i
computeElimType (f :@ s) {ne} = do computeElimType (f :@ s) {ne} = do
-- can't use `expectPi` (or `expectEq` below) without making this -- can't use `expectPi` (or `expectEq` below) without making this
-- mutual block from hell even worse lol -- mutual block from hell even worse lol
Pi {arg, res, _} <- whnf0 ns defs ctx =<< Pi {arg, res, _} <- whnf0 defs ctx =<< computeElimType f {ne = noOr1 ne}
computeElimType f {ne = noOr1 ne}
| t => Left $ ExpectedPi ctx.names t | t => Left $ ExpectedPi ctx.names t
pure $ sub1 res (s :# arg) pure $ sub1 res (s :# arg)
computeElimType (CasePair {pair, ret, _}) = pure $ sub1 ret pair computeElimType (CasePair {pair, ret, _}) = pure $ sub1 ret pair
@ -237,7 +234,7 @@ mutual
computeElimType (CaseNat {nat, ret, _}) = pure $ sub1 ret nat computeElimType (CaseNat {nat, ret, _}) = pure $ sub1 ret nat
computeElimType (CaseBox {box, ret, _}) = pure $ sub1 ret box computeElimType (CaseBox {box, ret, _}) = pure $ sub1 ret box
computeElimType (f :% p) {ne} = do computeElimType (f :% p) {ne} = do
Eq {ty, _} <- whnf0 ns defs ctx =<< computeElimType f {ne = noOr1 ne} Eq {ty, _} <- whnf0 defs ctx =<< computeElimType f {ne = noOr1 ne}
| t => Left $ ExpectedEq ctx.names t | t => Left $ ExpectedEq ctx.names t
pure $ dsub1 ty p pure $ dsub1 ty p
computeElimType (Coe {ty, q, _}) = pure $ dsub1 ty q computeElimType (Coe {ty, q, _}) = pure $ dsub1 ty q
@ -245,17 +242,16 @@ mutual
computeElimType (TypeCase {ret, _}) = pure ret computeElimType (TypeCase {ret, _}) = pure ret
computeElimType (_ :# ty) = pure ty computeElimType (_ :# ty) = pure ty
parameters {d, n : Nat} (ns : Mods) (defs : Definitions) parameters {d, n : Nat} (defs : Definitions) (ctx : WhnfContext (S d) n)
(ctx : WhnfContext (S d) n)
||| for π.(x : A) → B, returns (A, B); ||| for π.(x : A) → B, returns (A, B);
||| for an elim returns a pair of type-cases that will reduce to that; ||| for an elim returns a pair of type-cases that will reduce to that;
||| for other intro forms error ||| for other intro forms error
private covering private covering
tycasePi : (t : Term (S d) n) -> (0 tnf : No (isRedexT ns defs t)) => tycasePi : (t : Term (S d) n) -> (0 tnf : No (isRedexT defs t)) =>
Either Error (Term (S d) n, ScopeTerm (S d) n) Either Error (Term (S d) n, ScopeTerm (S d) n)
tycasePi (Pi {arg, res, _}) = pure (arg, res) tycasePi (Pi {arg, res, _}) = pure (arg, res)
tycasePi (E e) {tnf} = do tycasePi (E e) {tnf} = do
ty <- computeElimType ns defs ctx e @{noOr2 tnf} ty <- computeElimType defs ctx e @{noOr2 tnf}
let arg = E $ typeCase1Y e ty KPi [< "Arg", "Ret"] (BVT 1) let arg = E $ typeCase1Y e ty KPi [< "Arg", "Ret"] (BVT 1)
res' = typeCase1Y e (Arr Zero arg ty) KPi [< "Arg", "Ret"] (BVT 0) res' = typeCase1Y e (Arr Zero arg ty) KPi [< "Arg", "Ret"] (BVT 0)
res = SY [< "Arg"] $ E $ weakE 1 res' :@ BVT 0 res = SY [< "Arg"] $ E $ weakE 1 res' :@ BVT 0
@ -266,11 +262,11 @@ mutual
||| for an elim returns a pair of type-cases that will reduce to that; ||| for an elim returns a pair of type-cases that will reduce to that;
||| for other intro forms error ||| for other intro forms error
private covering private covering
tycaseSig : (t : Term (S d) n) -> (0 tnf : No (isRedexT ns defs t)) => tycaseSig : (t : Term (S d) n) -> (0 tnf : No (isRedexT defs t)) =>
Either Error (Term (S d) n, ScopeTerm (S d) n) Either Error (Term (S d) n, ScopeTerm (S d) n)
tycaseSig (Sig {fst, snd, _}) = pure (fst, snd) tycaseSig (Sig {fst, snd, _}) = pure (fst, snd)
tycaseSig (E e) {tnf} = do tycaseSig (E e) {tnf} = do
ty <- computeElimType ns defs ctx e @{noOr2 tnf} ty <- computeElimType defs ctx e @{noOr2 tnf}
let fst = E $ typeCase1Y e ty KSig [< "Fst", "Snd"] (BVT 1) let fst = E $ typeCase1Y e ty KSig [< "Fst", "Snd"] (BVT 1)
snd' = typeCase1Y e (Arr Zero fst ty) KSig [< "Fst", "Snd"] (BVT 0) snd' = typeCase1Y e (Arr Zero fst ty) KSig [< "Fst", "Snd"] (BVT 0)
snd = SY [< "Fst"] $ E $ weakE 1 snd' :@ BVT 0 snd = SY [< "Fst"] $ E $ weakE 1 snd' :@ BVT 0
@ -281,11 +277,11 @@ mutual
||| for an elim returns a type-case that will reduce to that; ||| for an elim returns a type-case that will reduce to that;
||| for other intro forms error ||| for other intro forms error
private covering private covering
tycaseBOX : (t : Term (S d) n) -> (0 tnf : No (isRedexT ns defs t)) => tycaseBOX : (t : Term (S d) n) -> (0 tnf : No (isRedexT defs t)) =>
Either Error (Term (S d) n) Either Error (Term (S d) n)
tycaseBOX (BOX _ a) = pure a tycaseBOX (BOX _ a) = pure a
tycaseBOX (E e) {tnf} = do tycaseBOX (E e) {tnf} = do
ty <- computeElimType ns defs ctx e @{noOr2 tnf} ty <- computeElimType defs ctx e @{noOr2 tnf}
pure $ E $ typeCase1Y e ty KBOX [< "Ty"] (BVT 0) pure $ E $ typeCase1Y e ty KBOX [< "Ty"] (BVT 0)
tycaseBOX t = Left $ ExpectedBOX ctx.names t tycaseBOX t = Left $ ExpectedBOX ctx.names t
@ -293,12 +289,12 @@ mutual
||| for an elim returns five type-cases that will reduce to that; ||| for an elim returns five type-cases that will reduce to that;
||| for other intro forms error ||| for other intro forms error
private covering private covering
tycaseEq : (t : Term (S d) n) -> (0 tnf : No (isRedexT ns defs t)) => tycaseEq : (t : Term (S d) n) -> (0 tnf : No (isRedexT defs t)) =>
Either Error (Term (S d) n, Term (S d) n, DScopeTerm (S d) n, Either Error (Term (S d) n, Term (S d) n, DScopeTerm (S d) n,
Term (S d) n, Term (S d) n) Term (S d) n, Term (S d) n)
tycaseEq (Eq {ty, l, r}) = pure (ty.zero, ty.one, ty, l, r) tycaseEq (Eq {ty, l, r}) = pure (ty.zero, ty.one, ty, l, r)
tycaseEq (E e) {tnf} = do tycaseEq (E e) {tnf} = do
ty <- computeElimType ns defs ctx e @{noOr2 tnf} ty <- computeElimType defs ctx e @{noOr2 tnf}
let names = [< "A0", "A1", "A", "L", "R"] let names = [< "A0", "A1", "A", "L", "R"]
a0 = E $ typeCase1Y e ty KEq names (BVT 4) a0 = E $ typeCase1Y e ty KEq names (BVT 4)
a1 = E $ typeCase1Y e ty KEq names (BVT 3) a1 = E $ typeCase1Y e ty KEq names (BVT 3)
@ -311,12 +307,11 @@ mutual
-- new block because the functions below might pass a different ctx -- new block because the functions below might pass a different ctx
-- into the ones above -- into the ones above
parameters {d, n : Nat} (ns : Mods) (defs : Definitions) parameters {d, n : Nat} (defs : Definitions) (ctx : WhnfContext d n)
(ctx : WhnfContext d n)
||| reduce a function application `Coe ty p q val :@ s` ||| reduce a function application `Coe ty p q val :@ s`
private covering private covering
piCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val, s : Term d n) -> piCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val, s : Term d n) ->
Either Error (Subset (Elim d n) (No . isRedexE ns defs)) Either Error (Subset (Elim d n) (No . isRedexE defs))
piCoe sty@(S [< i] ty) p q val s = do piCoe sty@(S [< i] ty) p q val s = do
-- (coe [i ⇒ π.(x : A) → B] @p @q t) s ⇝ -- (coe [i ⇒ π.(x : A) → B] @p @q t) s ⇝
-- coe [i ⇒ B[𝒔i/x] @p @q ((t ∷ (π.(x : A) → B)p/i) 𝒔p) -- coe [i ⇒ B[𝒔i/x] @p @q ((t ∷ (π.(x : A) → B)p/i) 𝒔p)
@ -324,20 +319,20 @@ mutual
-- --
-- type-case is used to expose A,B if the type is neutral -- type-case is used to expose A,B if the type is neutral
let ctx1 = extendDim i ctx let ctx1 = extendDim i ctx
Element ty tynf <- whnf ns defs ctx1 ty.term Element ty tynf <- whnf defs ctx1 ty.term
(arg, res) <- tycasePi ns defs ctx1 ty (arg, res) <- tycasePi defs ctx1 ty
let s0 = CoeT i arg q p s let s0 = CoeT i arg q p s
body = E $ (val :# (ty // one p)) :@ E s0 body = E $ (val :# (ty // one p)) :@ E s0
s1 = CoeT i (arg // (BV 0 ::: shift 2)) (weakD 1 q) (BV 0) s1 = CoeT i (arg // (BV 0 ::: shift 2)) (weakD 1 q) (BV 0)
(s // shift 1) (s // shift 1)
whnf ns defs ctx $ CoeT i (sub1 res s1) p q body whnf defs ctx $ CoeT i (sub1 res s1) p q body
||| reduce a pair elimination `CasePair pi (Coe ty p q val) ret body` ||| reduce a pair elimination `CasePair pi (Coe ty p q val) ret body`
private covering private covering
sigCoe : (qty : Qty) -> sigCoe : (qty : Qty) ->
(ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) -> (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
(ret : ScopeTerm d n) -> (body : ScopeTermN 2 d n) -> (ret : ScopeTerm d n) -> (body : ScopeTermN 2 d n) ->
Either Error (Subset (Elim d n) (No . isRedexE ns defs)) Either Error (Subset (Elim d n) (No . isRedexE defs))
sigCoe qty sty@(S [< i] ty) p q val ret body = do sigCoe qty sty@(S [< i] ty) p q val ret body = do
-- caseπ (coe [i ⇒ (x : A) × B] @p @q s) return z ⇒ C of { (a, b) ⇒ e } -- caseπ (coe [i ⇒ (x : A) × B] @p @q s) return z ⇒ C of { (a, b) ⇒ e }
-- ⇝ -- ⇝
@ -348,101 +343,101 @@ mutual
-- --
-- type-case is used to expose A,B if the type is neutral -- type-case is used to expose A,B if the type is neutral
let ctx1 = extendDim i ctx let ctx1 = extendDim i ctx
Element ty tynf <- whnf ns defs ctx1 ty.term Element ty tynf <- whnf defs ctx1 ty.term
(tfst, tsnd) <- tycaseSig ns defs ctx1 ty (tfst, tsnd) <- tycaseSig defs ctx1 ty
let a' = CoeT i (weakT 2 tfst) p q (BVT 1) let a' = CoeT i (weakT 2 tfst) p q (BVT 1)
tsnd' = tsnd.term // tsnd' = tsnd.term //
(CoeT i (weakT 2 $ tfst // (BV 0 ::: shift 2)) (CoeT i (weakT 2 $ tfst // (BV 0 ::: shift 2))
(weakD 1 p) (BV 0) (BVT 1) ::: shift 2) (weakD 1 p) (BV 0) (BVT 1) ::: shift 2)
b' = CoeT i tsnd' p q (BVT 0) b' = CoeT i tsnd' p q (BVT 0)
whnf ns defs ctx $ CasePair qty (val :# (ty // one p)) ret $ whnf defs ctx $ CasePair qty (val :# (ty // one p)) ret $
ST body.names $ body.term // (a' ::: b' ::: shift 2) ST body.names $ body.term // (a' ::: b' ::: shift 2)
||| reduce a dimension application `Coe ty p q val :% r` ||| reduce a dimension application `Coe ty p q val :% r`
private covering private covering
eqCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) -> eqCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
(r : Dim d) -> (r : Dim d) ->
Either Error (Subset (Elim d n) (No . isRedexE ns defs)) Either Error (Subset (Elim d n) (No . isRedexE defs))
eqCoe sty@(S [< j] ty) p q val r = do eqCoe sty@(S [< j] ty) p q val r = do
-- (coe [j ⇒ Eq [i ⇒ A] L R] @p @q eq) @r -- (coe [j ⇒ Eq [i ⇒ A] L R] @p @q eq) @r
-- ⇝ -- ⇝
-- comp [j ⇒ Ar/i] @p @q (eq ∷ (Eq [i ⇒ A] L R)p/j) -- comp [j ⇒ Ar/i] @p @q (eq ∷ (Eq [i ⇒ A] L R)p/j)
-- { (r=0) j ⇒ L; (r=1) j ⇒ R } -- { (r=0) j ⇒ L; (r=1) j ⇒ R }
let ctx1 = extendDim j ctx let ctx1 = extendDim j ctx
Element ty tynf <- whnf ns defs ctx1 ty.term Element ty tynf <- whnf defs ctx1 ty.term
(a0, a1, a, s, t) <- tycaseEq ns defs ctx1 ty (a0, a1, a, s, t) <- tycaseEq defs ctx1 ty
let a' = dsub1 a (weakD 1 r) let a' = dsub1 a (weakD 1 r)
val' = E $ (val :# (ty // one p)) :% r val' = E $ (val :# (ty // one p)) :% r
whnf ns defs ctx $ CompH j a' p q val' r j s j t whnf defs ctx $ CompH j a' p q val' r j s j t
||| reduce a pair elimination `CaseBox pi (Coe ty p q val) ret body` ||| reduce a pair elimination `CaseBox pi (Coe ty p q val) ret body`
private covering private covering
boxCoe : (qty : Qty) -> boxCoe : (qty : Qty) ->
(ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) -> (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
(ret : ScopeTerm d n) -> (body : ScopeTerm d n) -> (ret : ScopeTerm d n) -> (body : ScopeTerm d n) ->
Either Error (Subset (Elim d n) (No . isRedexE ns defs)) Either Error (Subset (Elim d n) (No . isRedexE defs))
boxCoe qty sty@(S [< i] ty) p q val ret body = do boxCoe qty sty@(S [< i] ty) p q val ret body = do
-- caseπ (coe [i ⇒ [ρ. A]] @p @q s) return z ⇒ C of { [a] ⇒ e } -- caseπ (coe [i ⇒ [ρ. A]] @p @q s) return z ⇒ C of { [a] ⇒ e }
-- ⇝ -- ⇝
-- caseπ s ∷ [ρ. A]p/i return z ⇒ C -- caseπ s ∷ [ρ. A]p/i return z ⇒ C
-- of { [a] ⇒ e[(coe [i ⇒ A] p q a)/a] } -- of { [a] ⇒ e[(coe [i ⇒ A] p q a)/a] }
let ctx1 = extendDim i ctx let ctx1 = extendDim i ctx
Element ty tynf <- whnf ns defs ctx1 ty.term Element ty tynf <- whnf defs ctx1 ty.term
ta <- tycaseBOX ns defs ctx1 ty ta <- tycaseBOX defs ctx1 ty
let a' = CoeT i (weakT 1 ta) p q $ BVT 0 let a' = CoeT i (weakT 1 ta) p q $ BVT 0
whnf ns defs ctx $ CaseBox qty (val :# (ty // one p)) ret $ whnf defs ctx $ CaseBox qty (val :# (ty // one p)) ret $
ST body.names $ body.term // (a' ::: shift 1) ST body.names $ body.term // (a' ::: shift 1)
export covering export covering
Whnf Elim Reduce.isRedexE where Whnf Elim Reduce.isRedexE where
whnf ns defs ctx (F x) with (lookupElimNS ns x defs) proof eq whnf defs ctx (F x) with (lookupElim x defs) proof eq
_ | Just y = whnf ns defs ctx y _ | Just y = whnf defs ctx y
_ | Nothing = pure $ Element (F x) $ rewrite eq in Ah _ | Nothing = pure $ Element (F x) $ rewrite eq in Ah
whnf _ _ _ (B i) = pure $ nred $ B i whnf _ _ (B i) = pure $ nred $ B i
-- ((λ x ⇒ t) ∷ (π.x : A) → B) s ⇝ t[s∷A/x] ∷ B[s∷A/x] -- ((λ x ⇒ t) ∷ (π.x : A) → B) s ⇝ t[s∷A/x] ∷ B[s∷A/x]
whnf ns defs ctx (f :@ s) = do whnf defs ctx (f :@ s) = do
Element f fnf <- whnf ns defs ctx f Element f fnf <- whnf defs ctx f
case nchoose $ isLamHead f of case nchoose $ isLamHead f of
Left _ => case f of Left _ => case f of
Lam body :# Pi {arg, res, _} => Lam body :# Pi {arg, res, _} =>
let s = s :# arg in let s = s :# arg in
whnf ns defs ctx $ sub1 body s :# sub1 res s whnf defs ctx $ sub1 body s :# sub1 res s
Coe ty p q val => piCoe ns defs ctx ty p q val s Coe ty p q val => piCoe defs ctx ty p q val s
Right nlh => pure $ Element (f :@ s) $ fnf `orNo` nlh Right nlh => pure $ Element (f :@ s) $ fnf `orNo` nlh
-- case (s, t) ∷ (x : A) × B return p ⇒ C of { (a, b) ⇒ u } ⇝ -- case (s, t) ∷ (x : A) × B return p ⇒ C of { (a, b) ⇒ u } ⇝
-- u[s∷A/a, t∷B[s∷A/x]] ∷ C[(s, t)∷((x : A) × B)/p] -- u[s∷A/a, t∷B[s∷A/x]] ∷ C[(s, t)∷((x : A) × B)/p]
whnf ns defs ctx (CasePair pi pair ret body) = do whnf defs ctx (CasePair pi pair ret body) = do
Element pair pairnf <- whnf ns defs ctx pair Element pair pairnf <- whnf defs ctx pair
case nchoose $ isPairHead pair of case nchoose $ isPairHead pair of
Left _ => case pair of Left _ => case pair of
Pair {fst, snd} :# Sig {fst = tfst, snd = tsnd, _} => Pair {fst, snd} :# Sig {fst = tfst, snd = tsnd, _} =>
let fst = fst :# tfst let fst = fst :# tfst
snd = snd :# sub1 tsnd fst snd = snd :# sub1 tsnd fst
in in
whnf ns defs ctx $ subN body [< fst, snd] :# sub1 ret pair whnf defs ctx $ subN body [< fst, snd] :# sub1 ret pair
Coe ty p q val => do Coe ty p q val => do
sigCoe ns defs ctx pi ty p q val ret body sigCoe defs ctx pi ty p q val ret body
Right np => Right np =>
pure $ Element (CasePair pi pair ret body) $ pairnf `orNo` np pure $ Element (CasePair pi pair ret body) $ pairnf `orNo` np
-- case 'a ∷ {a,…} return p ⇒ C of { 'a ⇒ u } ⇝ -- case 'a ∷ {a,…} return p ⇒ C of { 'a ⇒ u } ⇝
-- u ∷ C['a∷{a,…}/p] -- u ∷ C['a∷{a,…}/p]
whnf ns defs ctx (CaseEnum pi tag ret arms) = do whnf defs ctx (CaseEnum pi tag ret arms) = do
Element tag tagnf <- whnf ns defs ctx tag Element tag tagnf <- whnf defs ctx tag
case nchoose $ isTagHead tag of case nchoose $ isTagHead tag of
Left t => case tag of Left t => case tag of
Tag t :# Enum ts => Tag t :# Enum ts =>
let ty = sub1 ret tag in let ty = sub1 ret tag in
case lookup t arms of case lookup t arms of
Just arm => whnf ns defs ctx $ arm :# ty Just arm => whnf defs ctx $ arm :# ty
Nothing => Left $ MissingEnumArm t (keys arms) Nothing => Left $ MissingEnumArm t (keys arms)
Coe ty p q val => Coe ty p q val =>
-- there is nowhere an equality can be hiding inside an Enum -- there is nowhere an equality can be hiding inside an Enum
whnf ns defs ctx $ CaseEnum pi (val :# dsub1 ty q) ret arms whnf defs ctx $ CaseEnum pi (val :# dsub1 ty q) ret arms
Right nt => Right nt =>
pure $ Element (CaseEnum pi tag ret arms) $ tagnf `orNo` nt pure $ Element (CaseEnum pi tag ret arms) $ tagnf `orNo` nt
@ -451,35 +446,35 @@ mutual
-- --
-- case succ n ∷ return p ⇒ C of { succ n', π.ih ⇒ u; … } ⇝ -- case succ n ∷ return p ⇒ C of { succ n', π.ih ⇒ u; … } ⇝
-- u[n∷/n', (case n ∷ ⋯)/ih] ∷ C[succ n ∷ /p] -- u[n∷/n', (case n ∷ ⋯)/ih] ∷ C[succ n ∷ /p]
whnf ns defs ctx (CaseNat pi piIH nat ret zer suc) = do whnf defs ctx (CaseNat pi piIH nat ret zer suc) = do
Element nat natnf <- whnf ns defs ctx nat Element nat natnf <- whnf defs ctx nat
case nchoose $ isNatHead nat of case nchoose $ isNatHead nat of
Left _ => Left _ =>
let ty = sub1 ret nat in let ty = sub1 ret nat in
case nat of case nat of
Zero :# Nat => whnf ns defs ctx (zer :# ty) Zero :# Nat => whnf defs ctx (zer :# ty)
Succ n :# Nat => Succ n :# Nat =>
let nn = n :# Nat let nn = n :# Nat
tm = subN suc [< nn, CaseNat pi piIH nn ret zer suc] tm = subN suc [< nn, CaseNat pi piIH nn ret zer suc]
in in
whnf ns defs ctx $ tm :# ty whnf defs ctx $ tm :# ty
Coe ty p q val => Coe ty p q val =>
-- same deal as Enum -- same deal as Enum
whnf ns defs ctx $ CaseNat pi piIH (val :# dsub1 ty q) ret zer suc whnf defs ctx $ CaseNat pi piIH (val :# dsub1 ty q) ret zer suc
Right nn => Right nn =>
pure $ Element (CaseNat pi piIH nat ret zer suc) $ natnf `orNo` nn pure $ Element (CaseNat pi piIH nat ret zer suc) $ natnf `orNo` nn
-- case [t] ∷ [π.A] return p ⇒ C of { [x] ⇒ u } ⇝ -- case [t] ∷ [π.A] return p ⇒ C of { [x] ⇒ u } ⇝
-- u[t∷A/x] ∷ C[[t] ∷ [π.A]/p] -- u[t∷A/x] ∷ C[[t] ∷ [π.A]/p]
whnf ns defs ctx (CaseBox pi box ret body) = do whnf defs ctx (CaseBox pi box ret body) = do
Element box boxnf <- whnf ns defs ctx box Element box boxnf <- whnf defs ctx box
case nchoose $ isBoxHead box of case nchoose $ isBoxHead box of
Left _ => case box of Left _ => case box of
Box val :# BOX q bty => Box val :# BOX q bty =>
let ty = sub1 ret box in let ty = sub1 ret box in
whnf ns defs ctx $ sub1 body (val :# bty) :# ty whnf defs ctx $ sub1 body (val :# bty) :# ty
Coe ty p q val => Coe ty p q val =>
boxCoe ns defs ctx pi ty p q val ret body boxCoe defs ctx pi ty p q val ret body
Right nb => Right nb =>
pure $ Element (CaseBox pi box ret body) $ boxnf `orNo` nb pure $ Element (CaseBox pi box ret body) $ boxnf `orNo` nb
@ -487,104 +482,102 @@ mutual
-- ((δ 𝑖 ⇒ s) ∷ Eq [𝑗 ⇒ A] t u) @1 ⇝ u ∷ A1/𝑗 -- ((δ 𝑖 ⇒ s) ∷ Eq [𝑗 ⇒ A] t u) @1 ⇝ u ∷ A1/𝑗
-- ((δ 𝑖 ⇒ s) ∷ Eq [𝑗 ⇒ A] t u) @𝑘 ⇝ s𝑘/𝑖 ∷ A𝑘/𝑗 -- ((δ 𝑖 ⇒ s) ∷ Eq [𝑗 ⇒ A] t u) @𝑘 ⇝ s𝑘/𝑖 ∷ A𝑘/𝑗
-- (if 𝑘 is a variable) -- (if 𝑘 is a variable)
whnf ns defs ctx (f :% p) = do whnf defs ctx (f :% p) = do
Element f fnf <- whnf ns defs ctx f Element f fnf <- whnf defs ctx f
case nchoose $ isDLamHead f of case nchoose $ isDLamHead f of
Left _ => case f of Left _ => case f of
DLam body :# Eq {ty = ty, l, r, _} => DLam body :# Eq {ty = ty, l, r, _} =>
let body = endsOr l r (dsub1 body p) p in let body = endsOr l r (dsub1 body p) p in
whnf ns defs ctx $ body :# dsub1 ty p whnf defs ctx $ body :# dsub1 ty p
Coe ty p' q' val => Coe ty p' q' val =>
eqCoe ns defs ctx ty p' q' val p eqCoe defs ctx ty p' q' val p
Right ndlh => case p of Right ndlh => case p of
K e => do K e => do
Eq {l, r, ty, _} <- whnf0 ns defs ctx =<< Eq {l, r, ty, _} <- whnf0 defs ctx =<< computeElimType defs ctx f
computeElimType ns defs ctx f
| ty => Left $ ExpectedEq ctx.names ty | ty => Left $ ExpectedEq ctx.names ty
whnf ns defs ctx $ ends (l :# ty.zero) (r :# ty.one) e whnf defs ctx $ ends (l :# ty.zero) (r :# ty.one) e
B _ => pure $ Element (f :% p) $ fnf `orNo` ndlh `orNo` Ah B _ => pure $ Element (f :% p) $ fnf `orNo` ndlh `orNo` Ah
-- e ∷ A ⇝ e -- e ∷ A ⇝ e
whnf ns defs ctx (s :# a) = do whnf defs ctx (s :# a) = do
Element s snf <- whnf ns defs ctx s Element s snf <- whnf defs ctx s
case nchoose $ isE s of case nchoose $ isE s of
Left _ => let E e = s in pure $ Element e $ noOr2 snf Left _ => let E e = s in pure $ Element e $ noOr2 snf
Right ne => do Right ne => do
Element a anf <- whnf ns defs ctx a Element a anf <- whnf defs ctx a
pure $ Element (s :# a) $ ne `orNo` snf `orNo` anf pure $ Element (s :# a) $ ne `orNo` snf `orNo` anf
whnf ns defs ctx (Coe (S _ (N ty)) _ _ val) = whnf defs ctx (Coe (S _ (N ty)) _ _ val) =
whnf ns defs ctx $ val :# ty whnf defs ctx $ val :# ty
whnf ns defs ctx (Coe (S [< i] ty) p q val) = do whnf defs ctx (Coe (S [< i] ty) p q val) = do
Element ty tynf <- whnf ns defs (extendDim i ctx) ty.term Element ty tynf <- whnf defs (extendDim i ctx) ty.term
Element val valnf <- whnf ns defs ctx val Element val valnf <- whnf defs ctx val
pushCoe ns defs ctx i ty p q val pushCoe defs ctx i ty p q val
whnf ns defs ctx (Comp ty p q val r zero one) = whnf defs ctx (Comp ty p q val r zero one) =
-- comp [A] @p @p s { ⋯ } ⇝ s ∷ A -- comp [A] @p @p s { ⋯ } ⇝ s ∷ A
if p == q then whnf ns defs ctx $ val :# ty else if p == q then whnf defs ctx $ val :# ty else
case nchoose (isK r) of case nchoose (isK r) of
-- comp [A] @p @q s { (0=0) j ⇒ t; ⋯ } ⇝ tq/j ∷ A -- comp [A] @p @q s { (0=0) j ⇒ t; ⋯ } ⇝ tq/j ∷ A
-- comp [A] @p @q s { (1=1) j ⇒ t; ⋯ } ⇝ tq/j ∷ A -- comp [A] @p @q s { (1=1) j ⇒ t; ⋯ } ⇝ tq/j ∷ A
Left y => case r of Left y => case r of
K Zero => whnf ns defs ctx $ dsub1 zero q :# ty K Zero => whnf defs ctx $ dsub1 zero q :# ty
K One => whnf ns defs ctx $ dsub1 one q :# ty K One => whnf defs ctx $ dsub1 one q :# ty
Right nk => do Right nk => do
Element ty tynf <- whnf ns defs ctx ty Element ty tynf <- whnf defs ctx ty
pure $ Element (Comp ty p q val r zero one) $ tynf `orNo` nk pure $ Element (Comp ty p q val r zero one) $ tynf `orNo` nk
-- [todo] anything other than just the boundaries?? -- [todo] anything other than just the boundaries??
whnf ns defs ctx (TypeCase ty ret arms def) = do whnf defs ctx (TypeCase ty ret arms def) = do
Element ty tynf <- whnf ns defs ctx ty Element ty tynf <- whnf defs ctx ty
Element ret retnf <- whnf ns defs ctx ret Element ret retnf <- whnf defs ctx ret
case nchoose $ isAnnTyCon ty of case nchoose $ isAnnTyCon ty of
Left y => let ty :# TYPE u = ty in Left y => let ty :# TYPE u = ty in
reduceTypeCase ns defs ctx ty u ret arms def reduceTypeCase defs ctx ty u ret arms def
Right nt => pure $ Element (TypeCase ty ret arms def) $ Right nt => pure $ Element (TypeCase ty ret arms def) $
tynf `orNo` retnf `orNo` nt tynf `orNo` retnf `orNo` nt
whnf ns defs ctx (CloE el th) = whnf ns defs ctx $ pushSubstsWith' id th el whnf defs ctx (CloE el th) = whnf defs ctx $ pushSubstsWith' id th el
whnf ns defs ctx (DCloE el th) = whnf ns defs ctx $ pushSubstsWith' th id el whnf defs ctx (DCloE el th) = whnf defs ctx $ pushSubstsWith' th id el
export covering export covering
Whnf Term Reduce.isRedexT where Whnf Term Reduce.isRedexT where
whnf _ _ _ t@(TYPE {}) = pure $ nred t whnf _ _ t@(TYPE {}) = pure $ nred t
whnf _ _ _ t@(Pi {}) = pure $ nred t whnf _ _ t@(Pi {}) = pure $ nred t
whnf _ _ _ t@(Lam {}) = pure $ nred t whnf _ _ t@(Lam {}) = pure $ nred t
whnf _ _ _ t@(Sig {}) = pure $ nred t whnf _ _ t@(Sig {}) = pure $ nred t
whnf _ _ _ t@(Pair {}) = pure $ nred t whnf _ _ t@(Pair {}) = pure $ nred t
whnf _ _ _ t@(Enum {}) = pure $ nred t whnf _ _ t@(Enum {}) = pure $ nred t
whnf _ _ _ t@(Tag {}) = pure $ nred t whnf _ _ t@(Tag {}) = pure $ nred t
whnf _ _ _ t@(Eq {}) = pure $ nred t whnf _ _ t@(Eq {}) = pure $ nred t
whnf _ _ _ t@(DLam {}) = pure $ nred t whnf _ _ t@(DLam {}) = pure $ nred t
whnf _ _ _ Nat = pure $ nred Nat whnf _ _ Nat = pure $ nred Nat
whnf _ _ _ Zero = pure $ nred Zero whnf _ _ Zero = pure $ nred Zero
whnf _ _ _ t@(Succ {}) = pure $ nred t whnf _ _ t@(Succ {}) = pure $ nred t
whnf _ _ _ t@(BOX {}) = pure $ nred t whnf _ _ t@(BOX {}) = pure $ nred t
whnf _ _ _ t@(Box {}) = pure $ nred t whnf _ _ t@(Box {}) = pure $ nred t
-- s ∷ A ⇝ s (in term context) -- s ∷ A ⇝ s (in term context)
whnf ns defs ctx (E e) = do whnf defs ctx (E e) = do
Element e enf <- whnf ns defs ctx e Element e enf <- whnf defs ctx e
case nchoose $ isAnn e of case nchoose $ isAnn e of
Left _ => let tm :# _ = e in pure $ Element tm $ noOr1 $ noOr2 enf Left _ => let tm :# _ = e in pure $ Element tm $ noOr1 $ noOr2 enf
Right na => pure $ Element (E e) $ na `orNo` enf Right na => pure $ Element (E e) $ na `orNo` enf
whnf ns defs ctx (CloT tm th) = whnf ns defs ctx $ pushSubstsWith' id th tm whnf defs ctx (CloT tm th) = whnf defs ctx $ pushSubstsWith' id th tm
whnf ns defs ctx (DCloT tm th) = whnf ns defs ctx $ pushSubstsWith' th id tm whnf defs ctx (DCloT tm th) = whnf defs ctx $ pushSubstsWith' th id tm
||| reduce a type-case applied to a type constructor ||| reduce a type-case applied to a type constructor
private covering private covering
reduceTypeCase : {d, n : Nat} -> (ns : Mods) -> (defs : Definitions) -> reduceTypeCase : {d, n : Nat} -> (defs : Definitions) -> WhnfContext d n ->
WhnfContext d n ->
(ty : Term d n) -> (u : Universe) -> (ret : Term d n) -> (ty : Term d n) -> (u : Universe) -> (ret : Term d n) ->
(arms : TypeCaseArms d n) -> (def : Term d n) -> (arms : TypeCaseArms d n) -> (def : Term d n) ->
(0 _ : So (isTyCon ty)) => (0 _ : So (isTyCon ty)) =>
Either Error (Subset (Elim d n) (No . isRedexE ns defs)) Either Error (Subset (Elim d n) (No . isRedexE defs))
reduceTypeCase ns defs ctx ty u ret arms def = case ty of reduceTypeCase defs ctx ty u ret arms def = case ty of
-- (type-case ★ᵢ ∷ _ return Q of { ★ ⇒ s; ⋯ }) ⇝ s ∷ Q -- (type-case ★ᵢ ∷ _ return Q of { ★ ⇒ s; ⋯ }) ⇝ s ∷ Q
TYPE _ => TYPE _ =>
whnf ns defs ctx $ tycaseRhsDef0 def KTYPE arms :# ret whnf defs ctx $ tycaseRhsDef0 def KTYPE arms :# ret
-- (type-case π.(x : A) → B ∷ ★ᵢ return Q of { (a → b) ⇒ s; ⋯ }) ⇝ -- (type-case π.(x : A) → B ∷ ★ᵢ return Q of { (a → b) ⇒ s; ⋯ }) ⇝
-- s[(A ∷ ★ᵢ)/a, ((λ x ⇒ B) ∷ 0.A → ★ᵢ)/b] ∷ ★ᵢ -- s[(A ∷ ★ᵢ)/a, ((λ x ⇒ B) ∷ 0.A → ★ᵢ)/b] ∷ ★ᵢ
@ -592,7 +585,7 @@ mutual
let arg = arg :# TYPE u let arg = arg :# TYPE u
res = Lam res :# Arr Zero (TYPE u) (TYPE u) res = Lam res :# Arr Zero (TYPE u) (TYPE u)
in in
whnf ns defs ctx $ subN (tycaseRhsDef def KPi arms) [< arg, res] :# ret whnf defs ctx $ subN (tycaseRhsDef def KPi arms) [< arg, res] :# ret
-- (type-case (x : A) × B ∷ ★ᵢ return Q of { (a × b) ⇒ s; ⋯ }) ⇝ -- (type-case (x : A) × B ∷ ★ᵢ return Q of { (a × b) ⇒ s; ⋯ }) ⇝
-- s[(A ∷ ★ᵢ)/a, ((λ x ⇒ B) ∷ 0.A → ★ᵢ)/b] ∷ ★ᵢ -- s[(A ∷ ★ᵢ)/a, ((λ x ⇒ B) ∷ 0.A → ★ᵢ)/b] ∷ ★ᵢ
@ -600,11 +593,11 @@ mutual
let fst = fst :# TYPE u let fst = fst :# TYPE u
snd = Lam snd :# Arr Zero (TYPE u) (TYPE u) snd = Lam snd :# Arr Zero (TYPE u) (TYPE u)
in in
whnf ns defs ctx $ subN (tycaseRhsDef def KSig arms) [< fst, snd] :# ret whnf defs ctx $ subN (tycaseRhsDef def KSig arms) [< fst, snd] :# ret
-- (type-case {⋯} ∷ _ return Q of { {} ⇒ s; ⋯ }) ⇝ s ∷ Q -- (type-case {⋯} ∷ _ return Q of { {} ⇒ s; ⋯ }) ⇝ s ∷ Q
Enum _ => Enum _ =>
whnf ns defs ctx $ tycaseRhsDef0 def KEnum arms :# ret whnf defs ctx $ tycaseRhsDef0 def KEnum arms :# ret
-- (type-case Eq [i ⇒ A] L R ∷ ★ᵢ return Q -- (type-case Eq [i ⇒ A] L R ∷ ★ᵢ return Q
-- of { Eq a₀ a₁ a l r ⇒ s; ⋯ }) ⇝ -- of { Eq a₀ a₁ a l r ⇒ s; ⋯ }) ⇝
@ -613,7 +606,7 @@ mutual
-- (L ∷ A0/i)/l, (R ∷ A1/i)/r] ∷ Q -- (L ∷ A0/i)/l, (R ∷ A1/i)/r] ∷ Q
Eq a l r => Eq a l r =>
let a0 = a.zero; a1 = a.one in let a0 = a.zero; a1 = a.one in
whnf ns defs ctx $ whnf defs ctx $
subN (tycaseRhsDef def KEq arms) subN (tycaseRhsDef def KEq arms)
[< a0 :# TYPE u, a1 :# TYPE u, [< a0 :# TYPE u, a1 :# TYPE u,
DLam a :# Eq0 (TYPE u) a0 a1, l :# a0, r :# a1] DLam a :# Eq0 (TYPE u) a0 a1, l :# a0, r :# a1]
@ -621,23 +614,22 @@ mutual
-- (type-case ∷ _ return Q of { ⇒ s; ⋯ }) ⇝ s ∷ Q -- (type-case ∷ _ return Q of { ⇒ s; ⋯ }) ⇝ s ∷ Q
Nat => Nat =>
whnf ns defs ctx $ tycaseRhsDef0 def KNat arms :# ret whnf defs ctx $ tycaseRhsDef0 def KNat arms :# ret
-- (type-case [π.A] ∷ ★ᵢ return Q of { [a] ⇒ s; ⋯ }) ⇝ s[(A ∷ ★ᵢ)/a] ∷ Q -- (type-case [π.A] ∷ ★ᵢ return Q of { [a] ⇒ s; ⋯ }) ⇝ s[(A ∷ ★ᵢ)/a] ∷ Q
BOX _ s => BOX _ s =>
whnf ns defs ctx $ sub1 (tycaseRhsDef def KBOX arms) (s :# TYPE u) :# ret whnf defs ctx $ sub1 (tycaseRhsDef def KBOX arms) (s :# TYPE u) :# ret
||| pushes a coercion inside a whnf-ed term ||| pushes a coercion inside a whnf-ed term
private covering private covering
pushCoe : {n, d : Nat} -> (ns : Mods) -> (defs : Definitions) -> pushCoe : {n, d : Nat} -> (defs : Definitions) -> WhnfContext d n ->
WhnfContext d n ->
BaseName -> BaseName ->
(ty : Term (S d) n) -> (0 tynf : No (isRedexT ns defs ty)) => (ty : Term (S d) n) -> (0 tynf : No (isRedexT defs ty)) =>
Dim d -> Dim d -> Dim d -> Dim d ->
(s : Term d n) -> (0 snf : No (isRedexT ns defs s)) => (s : Term d n) -> (0 snf : No (isRedexT defs s)) =>
Either Error (NonRedex Elim d n ns defs) Either Error (NonRedex Elim d n defs)
pushCoe ns defs ctx i ty p q s = pushCoe defs ctx i ty p q s =
if p == q then whnf ns defs ctx $ s :# (ty // one q) else if p == q then whnf defs ctx $ s :# (ty // one q) else
case s of case s of
-- (coe [_ ⇒ ★ᵢ] @_ @_ ty) ⇝ (ty ∷ ★ᵢ) -- (coe [_ ⇒ ★ᵢ] @_ @_ ty) ⇝ (ty ∷ ★ᵢ)
TYPE {} => pure $ nred $ s :# TYPE !(unwrapTYPE ty) TYPE {} => pure $ nred $ s :# TYPE !(unwrapTYPE ty)
@ -654,7 +646,7 @@ mutual
let body' = CoeT i ty p q $ Lam body let body' = CoeT i ty p q $ Lam body
term' = [< "y"] :\\ E (weakE 1 body' :@ BVT 0) term' = [< "y"] :\\ E (weakE 1 body' :@ BVT 0)
type' = ty // one q type' = ty // one q
whnf ns defs ctx $ term' :# type' whnf defs ctx $ term' :# type'
{- {-
-- maybe: -- maybe:
@ -692,7 +684,7 @@ mutual
let body' = CoeT i ty p q $ DLam body let body' = CoeT i ty p q $ DLam body
term' = [< "j"] :\\% E (dweakE 1 body' :% BV 0) term' = [< "j"] :\\% E (dweakE 1 body' :% BV 0)
type' = ty // one q type' = ty // one q
whnf ns defs ctx $ term' :# type' whnf defs ctx $ term' :# type'
-- (coe [_ ⇒ {⋯}] @_ @_ t) ⇝ (t ∷ {⋯}) -- (coe [_ ⇒ {⋯}] @_ @_ t) ⇝ (t ∷ {⋯})
Tag tag => do Tag tag => do

View file

@ -13,16 +13,16 @@ import Quox.EffExtra
public export public export
0 TCEff : List (Type -> Type) 0 TCEff : List (Type -> Type)
TCEff = [ErrorEff, DefsReader, CurNSReader] TCEff = [ErrorEff, DefsReader]
public export public export
0 TC : Type -> Type 0 TC : Type -> Type
TC = Eff TCEff TC = Eff TCEff
export export
runTC : Mods -> Definitions -> TC a -> Either Error a runTC : Definitions -> TC a -> Either Error a
runTC ns defs = runTC defs =
extract . runExcept . runReaderAt DEFS defs . runReaderAt NS ns extract . runExcept . runReaderAt DEFS defs
export export
@ -149,7 +149,7 @@ mutual
(subj : Term d n) -> (0 nc : NotClo subj) => Term d n -> (subj : Term d n) -> (0 nc : NotClo subj) => Term d n ->
TC (CheckResult' n) TC (CheckResult' n)
toCheckType ctx sg t ty = do toCheckType ctx sg t ty = do
u <- expectTYPE !ns !defs ctx ty u <- expectTYPE !defs ctx ty
expectEqualQ Zero sg.fst expectEqualQ Zero sg.fst
checkTypeNoWrap ctx t (Just u) checkTypeNoWrap ctx t (Just u)
pure $ zeroFor ctx pure $ zeroFor ctx
@ -164,7 +164,7 @@ mutual
check' ctx sg t@(Pi {}) ty = toCheckType ctx sg t ty check' ctx sg t@(Pi {}) ty = toCheckType ctx sg t ty
check' ctx sg (Lam body) ty = do check' ctx sg (Lam body) ty = do
(qty, arg, res) <- expectPi !ns !defs ctx ty (qty, arg, res) <- expectPi !defs ctx ty
-- if Ψ | Γ, x : A ⊢ σ · t ⇐ B ⊳ Σ, ρ·x -- if Ψ | Γ, x : A ⊢ σ · t ⇐ B ⊳ Σ, ρ·x
-- with ρ ≤ σπ -- with ρ ≤ σπ
let qty' = sg.fst * qty let qty' = sg.fst * qty
@ -175,7 +175,7 @@ mutual
check' ctx sg t@(Sig {}) ty = toCheckType ctx sg t ty check' ctx sg t@(Sig {}) ty = toCheckType ctx sg t ty
check' ctx sg (Pair fst snd) ty = do check' ctx sg (Pair fst snd) ty = do
(tfst, tsnd) <- expectSig !ns !defs ctx ty (tfst, tsnd) <- expectSig !defs ctx ty
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ₁ -- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ₁
qfst <- checkC ctx sg fst tfst qfst <- checkC ctx sg fst tfst
let tsnd = sub1 tsnd (fst :# tfst) let tsnd = sub1 tsnd (fst :# tfst)
@ -187,7 +187,7 @@ mutual
check' ctx sg t@(Enum _) ty = toCheckType ctx sg t ty check' ctx sg t@(Enum _) ty = toCheckType ctx sg t ty
check' ctx sg (Tag t) ty = do check' ctx sg (Tag t) ty = do
tags <- expectEnum !ns !defs ctx ty tags <- expectEnum !defs ctx ty
-- if t ∈ ts -- if t ∈ ts
unless (t `elem` tags) $ throw $ TagNotIn t tags unless (t `elem` tags) $ throw $ TagNotIn t tags
-- then Ψ | Γ ⊢ σ · t ⇐ {ts} ⊳ 𝟎 -- then Ψ | Γ ⊢ σ · t ⇐ {ts} ⊳ 𝟎
@ -196,7 +196,7 @@ mutual
check' ctx sg t@(Eq {}) ty = toCheckType ctx sg t ty check' ctx sg t@(Eq {}) ty = toCheckType ctx sg t ty
check' ctx sg (DLam body) ty = do check' ctx sg (DLam body) ty = do
(ty, l, r) <- expectEq !ns !defs ctx ty (ty, l, r) <- expectEq !defs ctx ty
let ctx' = extendDim body.name ctx let ctx' = extendDim body.name ctx
ty = ty.term ty = ty.term
body = body.term body = body.term
@ -212,17 +212,17 @@ mutual
check' ctx sg Nat ty = toCheckType ctx sg Nat ty check' ctx sg Nat ty = toCheckType ctx sg Nat ty
check' ctx sg Zero ty = do check' ctx sg Zero ty = do
expectNat !ns !defs ctx ty expectNat !defs ctx ty
pure $ zeroFor ctx pure $ zeroFor ctx
check' ctx sg (Succ n) ty = do check' ctx sg (Succ n) ty = do
expectNat !ns !defs ctx ty expectNat !defs ctx ty
checkC ctx sg n Nat checkC ctx sg n Nat
check' ctx sg t@(BOX {}) ty = toCheckType ctx sg t ty check' ctx sg t@(BOX {}) ty = toCheckType ctx sg t ty
check' ctx sg (Box val) ty = do check' ctx sg (Box val) ty = do
(q, ty) <- expectBOX !ns !defs ctx ty (q, ty) <- expectBOX !defs ctx ty
-- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ -- if Ψ | Γ ⊢ σ · s ⇐ A ⊳ Σ
valout <- checkC ctx sg val ty valout <- checkC ctx sg val ty
-- then Ψ | Γ ⊢ σ · [s] ⇐ [π.A] ⊳ πΣ -- then Ψ | Γ ⊢ σ · [s] ⇐ [π.A] ⊳ πΣ
@ -301,7 +301,7 @@ mutual
case u of case u of
Just u => subtype ctx infres.type (TYPE u) Just u => subtype ctx infres.type (TYPE u)
Nothing => ignore $ Nothing => ignore $
expectTYPE !ns !defs ctx infres.type expectTYPE !defs ctx infres.type
-- then Ψ | Γ ⊢₀ E ⇐ Type 𝓀 -- then Ψ | Γ ⊢₀ E ⇐ Type 𝓀
private covering private covering
@ -338,7 +338,7 @@ mutual
infer' ctx sg (F x) = do infer' ctx sg (F x) = do
-- if π·x : A {≔ s} in global context -- if π·x : A {≔ s} in global context
g <- lookupFree !ns x !defs g <- lookupFree x !defs
-- if σ ≤ π -- if σ ≤ π
expectCompatQ sg.fst g.qty.fst expectCompatQ sg.fst g.qty.fst
-- then Ψ | Γ ⊢ σ · x ⇒ A ⊳ 𝟎 -- then Ψ | Γ ⊢ σ · x ⇒ A ⊳ 𝟎
@ -360,7 +360,7 @@ mutual
infer' ctx sg (fun :@ arg) = do infer' ctx sg (fun :@ arg) = do
-- if Ψ | Γ ⊢ σ · f ⇒ (π·x : A) → B ⊳ Σ₁ -- if Ψ | Γ ⊢ σ · f ⇒ (π·x : A) → B ⊳ Σ₁
funres <- inferC ctx sg fun funres <- inferC ctx sg fun
(qty, argty, res) <- expectPi !ns !defs ctx funres.type (qty, argty, res) <- expectPi !defs ctx funres.type
-- if Ψ | Γ ⊢ σ ⨴ π · s ⇐ A ⊳ Σ₂ -- if Ψ | Γ ⊢ σ ⨴ π · s ⇐ A ⊳ Σ₂
argout <- checkC ctx (subjMult sg qty) arg argty argout <- checkC ctx (subjMult sg qty) arg argty
-- then Ψ | Γ ⊢ σ · f s ⇒ B[s] ⊳ Σ₁ + Σ₂ -- then Ψ | Γ ⊢ σ · f s ⇒ B[s] ⊳ Σ₁ + Σ₂
@ -377,7 +377,7 @@ mutual
pairres <- inferC ctx sg pair pairres <- inferC ctx sg pair
-- if Ψ | Γ, p : (x : A) × B ⊢₀ ret ⇐ Type -- if Ψ | Γ, p : (x : A) × B ⊢₀ ret ⇐ Type
checkTypeC (extendTy Zero ret.name pairres.type ctx) ret.term Nothing checkTypeC (extendTy Zero ret.name pairres.type ctx) ret.term Nothing
(tfst, tsnd) <- expectSig !ns !defs ctx pairres.type (tfst, tsnd) <- expectSig !defs ctx pairres.type
-- if Ψ | Γ, x : A, y : B ⊢ σ · body ⇐ -- if Ψ | Γ, x : A, y : B ⊢ σ · body ⇐
-- ret[(x, y) ∷ (x : A) × B/p] ⊳ Σ₂, ρ₁·x, ρ₂·y -- ret[(x, y) ∷ (x : A) × B/p] ⊳ Σ₂, ρ₁·x, ρ₂·y
-- with ρ₁, ρ₂ ≤ πσ -- with ρ₁, ρ₂ ≤ πσ
@ -395,7 +395,7 @@ mutual
infer' ctx sg (CaseEnum pi t ret arms) {d, n} = do infer' ctx sg (CaseEnum pi t ret arms) {d, n} = do
-- if Ψ | Γ ⊢ σ · t ⇒ {ts} ⊳ Σ₁ -- if Ψ | Γ ⊢ σ · t ⇒ {ts} ⊳ Σ₁
tres <- inferC ctx sg t tres <- inferC ctx sg t
ttags <- expectEnum !ns !defs ctx tres.type ttags <- expectEnum !defs ctx tres.type
-- if 1 ≤ π, OR there is only zero or one option -- if 1 ≤ π, OR there is only zero or one option
unless (length (SortedSet.toList ttags) <= 1) $ expectCompatQ One pi unless (length (SortedSet.toList ttags) <= 1) $ expectCompatQ One pi
-- if Ψ | Γ, x : {ts} ⊢₀ A ⇐ Type -- if Ψ | Γ, x : {ts} ⊢₀ A ⇐ Type
@ -421,7 +421,7 @@ mutual
expectCompatQ One pi expectCompatQ One pi
-- if Ψ | Γ ⊢ σ · n ⇒ ⊳ Σn -- if Ψ | Γ ⊢ σ · n ⇒ ⊳ Σn
nres <- inferC ctx sg n nres <- inferC ctx sg n
expectNat !ns !defs ctx nres.type expectNat !defs ctx nres.type
-- if Ψ | Γ, n : ⊢₀ A ⇐ Type -- if Ψ | Γ, n : ⊢₀ A ⇐ Type
checkTypeC (extendTy Zero ret.name Nat ctx) ret.term Nothing checkTypeC (extendTy Zero ret.name Nat ctx) ret.term Nothing
-- if Ψ | Γ ⊢ σ · zer ⇐ A[0 ∷ /n] ⊳ Σz -- if Ψ | Γ ⊢ σ · zer ⇐ A[0 ∷ /n] ⊳ Σz
@ -445,7 +445,7 @@ mutual
infer' ctx sg (CaseBox pi box ret body) = do infer' ctx sg (CaseBox pi box ret body) = do
-- if Ψ | Γ ⊢ σ · b ⇒ [ρ.A] ⊳ Σ₁ -- if Ψ | Γ ⊢ σ · b ⇒ [ρ.A] ⊳ Σ₁
boxres <- inferC ctx sg box boxres <- inferC ctx sg box
(q, ty) <- expectBOX !ns !defs ctx boxres.type (q, ty) <- expectBOX !defs ctx boxres.type
-- if Ψ | Γ, x : [ρ.A] ⊢₀ R ⇐ Type -- if Ψ | Γ, x : [ρ.A] ⊢₀ R ⇐ Type
checkTypeC (extendTy Zero ret.name boxres.type ctx) ret.term Nothing checkTypeC (extendTy Zero ret.name boxres.type ctx) ret.term Nothing
-- if Ψ | Γ, x : A ⊢ t ⇐ R[[x] ∷ [ρ.A/x]] ⊳ Σ₂, ς·x -- if Ψ | Γ, x : A ⊢ t ⇐ R[[x] ∷ [ρ.A/x]] ⊳ Σ₂, ς·x
@ -463,7 +463,7 @@ mutual
infer' ctx sg (fun :% dim) = do infer' ctx sg (fun :% dim) = do
-- if Ψ | Γ ⊢ σ · f ⇒ Eq [𝑖 ⇒ A] l r ⊳ Σ -- if Ψ | Γ ⊢ σ · f ⇒ Eq [𝑖 ⇒ A] l r ⊳ Σ
InfRes {type, qout} <- inferC ctx sg fun InfRes {type, qout} <- inferC ctx sg fun
ty <- fst <$> expectEq !ns !defs ctx type ty <- fst <$> expectEq !defs ctx type
-- then Ψ | Γ ⊢ σ · f p ⇒ Ap/𝑖 ⊳ Σ -- then Ψ | Γ ⊢ σ · f p ⇒ Ap/𝑖 ⊳ Σ
pure $ InfRes {type = dsub1 ty dim, qout} pure $ InfRes {type = dsub1 ty dim, qout}
@ -496,7 +496,7 @@ mutual
-- if σ = 0 -- if σ = 0
expectEqualQ Zero sg.fst expectEqualQ Zero sg.fst
-- if Ψ, Γ ⊢₀ e ⇒ Type u -- if Ψ, Γ ⊢₀ e ⇒ Type u
u <- expectTYPE !ns !defs ctx . type =<< inferC ctx szero ty u <- expectTYPE !defs ctx . type =<< inferC ctx szero ty
-- if Ψ, Γ ⊢₀ C ⇐ Type (non-dependent return type) -- if Ψ, Γ ⊢₀ C ⇐ Type (non-dependent return type)
checkTypeC ctx ret Nothing checkTypeC ctx ret Nothing
-- if Ψ, Γ' ⊢₀ A ⇐ C for each rhs A -- if Ψ, Γ' ⊢₀ A ⇐ C for each rhs A

View file

@ -39,8 +39,8 @@ InferResult eqs d n = IfConsistent eqs $ InferResult' d n
export export
lookupFree : Has ErrorEff fs => Mods -> Name -> Definitions -> Eff fs Definition lookupFree : Has ErrorEff fs => Name -> Definitions -> Eff fs Definition
lookupFree ns x defs = maybe (throw $ NotInScope x) pure $ lookupNS ns x defs lookupFree x defs = maybe (throw $ NotInScope x) pure $ lookup x defs
public export public export
@ -59,14 +59,14 @@ substCaseBoxRet dty retty =
retty.term // (Box (BVT 0) :# weakT 1 dty ::: shift 1) retty.term // (Box (BVT 0) :# weakT 1 dty ::: shift 1)
parameters (ns : Mods) (defs : Definitions) {auto _ : Has ErrorEff fs} parameters (defs : Definitions) {auto _ : Has ErrorEff fs}
namespace TyContext namespace TyContext
parameters (ctx : TyContext d n) parameters (ctx : TyContext d n)
export covering export covering
whnf : {0 isRedex : RedexTest tm} -> Whnf tm isRedex => whnf : {0 isRedex : RedexTest tm} -> Whnf tm isRedex =>
tm d n -> Eff fs (NonRedex tm d n ns defs) tm d n -> Eff fs (NonRedex tm d n defs)
whnf = let Val n = ctx.termLen; Val d = ctx.dimLen in whnf = let Val n = ctx.termLen; Val d = ctx.dimLen in
rethrow . whnf ns defs (toWhnfContext ctx) rethrow . whnf defs (toWhnfContext ctx)
private covering %macro private covering %macro
expect : (forall d, n. NameContexts d n -> Term d n -> Error) -> expect : (forall d, n. NameContexts d n -> Term d n -> Error) ->
@ -108,8 +108,8 @@ parameters (ns : Mods) (defs : Definitions) {auto _ : Has ErrorEff fs}
parameters (ctx : EqContext n) parameters (ctx : EqContext n)
export covering export covering
whnf : {0 isRedex : RedexTest tm} -> Whnf tm isRedex => whnf : {0 isRedex : RedexTest tm} -> Whnf tm isRedex =>
tm 0 n -> Eff fs (NonRedex tm 0 n ns defs) tm 0 n -> Eff fs (NonRedex tm 0 n defs)
whnf = let Val n = ctx.termLen in rethrow . whnf ns defs (toWhnfContext ctx) whnf = let Val n = ctx.termLen in rethrow . whnf defs (toWhnfContext ctx)
private covering %macro private covering %macro
expect : (forall d, n. NameContexts d n -> Term d n -> Error) -> expect : (forall d, n. NameContexts d n -> Term d n -> Error) ->

View file

@ -21,10 +21,10 @@ defGlobals = fromList
parameters (label : String) (act : Lazy (TC ())) parameters (label : String) (act : Lazy (TC ()))
{default defGlobals globals : Definitions} {default defGlobals globals : Definitions}
testEq : Test testEq : Test
testEq = test label $ runTC [<] globals act testEq = test label $ runTC globals act
testNeq : Test testNeq : Test
testNeq = testThrows label (const True) $ runTC [<] globals act $> "()" testNeq = testThrows label (const True) $ runTC globals act $> "()"
parameters (0 d : Nat) (ctx : TyContext d n) parameters (0 d : Nat) (ctx : TyContext d n)

View file

@ -61,8 +61,9 @@ parameters {c : Bool} {auto _ : Show b}
FromString BName where fromString = Just . fromString FromString BName where fromString = Just . fromString
runFromParser : Eff [Except FromParser.Error] a -> Either FromParser.Error a runFromParser : {default empty defs : Definitions} ->
runFromParser = extract . runExcept Eff FromParserPure a -> Either FromParser.Error a
runFromParser = map fst . fromParserPure defs
export export
tests : Test tests : Test
@ -79,7 +80,9 @@ tests = "PTerm → Term" :- [
], ],
"terms" :- "terms" :-
let fromPTerm = runFromParser . let defs = fromList [("f", mkDef gany Nat Zero)]
-- doesn't have to be well typed yet, just well scoped
fromPTerm = runFromParser {defs} .
fromPTermWith [< "𝑖", "𝑗"] [< "A", "x", "y", "z"] fromPTermWith [< "𝑖", "𝑗"] [< "A", "x", "y", "z"]
in [ in [
note "dim ctx: [𝑖, 𝑗]; term ctx: [A, x, y, z]", note "dim ctx: [𝑖, 𝑗]; term ctx: [A, x, y, z]",

View file

@ -13,7 +13,7 @@ parameters {0 isRedex : RedexTest tm} {auto _ : Whnf tm isRedex} {d, n : Nat}
private private
testWhnf : String -> WhnfContext d n -> tm d n -> tm d n -> Test testWhnf : String -> WhnfContext d n -> tm d n -> tm d n -> Test
testWhnf label ctx from to = test "\{label} (whnf)" $ do testWhnf label ctx from to = test "\{label} (whnf)" $ do
result <- bimap toInfo fst $ whnf [<] defs ctx from result <- bimap toInfo fst $ whnf defs ctx from
unless (result == to) $ Left [("exp", show to), ("got", show result)] unless (result == to) $ Left [("exp", show to), ("got", show result)]
private private

View file

@ -28,7 +28,7 @@ ToInfo Error' where
M = Eff [Except Error', DefsReader] M = Eff [Except Error', DefsReader]
inj : TC a -> M a inj : TC a -> M a
inj = rethrow . mapFst TCError <=< lift . runExcept . runReaderAt NS [<] inj = rethrow . mapFst TCError <=< lift . runExcept
reflTy : Term d n reflTy : Term d n