more typed equality, uip, etc

This commit is contained in:
rhiannon morris 2023-02-11 18:15:50 +01:00
parent 7fd7a31635
commit 7d2c3b5a8e
8 changed files with 381 additions and 217 deletions

View file

@ -43,17 +43,31 @@ clashE e f = throwError $ ClashE !mode e f
public export %inline
isType : (t : Term {}) -> Bool
isType (TYPE {}) = True
isType (Pi {}) = True
isType (Lam {}) = False
isType (Sig {}) = True
isType (Pair {}) = False
isType (Eq {}) = True
isType (DLam {}) = False
isType (E {}) = True
isType (CloT {}) = False
isType (DCloT {}) = False
isTyCon : (t : Term {}) -> Bool
isTyCon (TYPE {}) = True
isTyCon (Pi {}) = True
isTyCon (Lam {}) = False
isTyCon (Sig {}) = True
isTyCon (Pair {}) = False
isTyCon (Eq {}) = True
isTyCon (DLam {}) = False
isTyCon (E {}) = True
isTyCon (CloT {}) = False
isTyCon (DCloT {}) = False
private
isSubSing : Term {} -> Bool
isSubSing ty =
let Element ty _ = pushSubsts ty in
case ty of
TYPE _ => False
Pi {res, _} => isSubSing res.term
Lam {} => False
Sig {fst, snd, _} => isSubSing fst && isSubSing snd.term
Pair {} => False
Eq {} => True
DLam {} => False
E e => False
parameters {auto _ : HasErr q m}
@ -64,8 +78,8 @@ parameters {auto _ : HasErr q m}
Right n => throwError $ e t
export %inline
ensureType : (t : Term q d n) -> m (So (isType t))
ensureType = ensure NotType isType
ensureType : (t : Term q d n) -> m (So (isTyCon t))
ensureType = ensure NotType isTyCon
parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
mutual
@ -88,7 +102,7 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
private covering
compare0' : TContext q 0 n ->
(ty, s, t : Term q 0 n) ->
(0 nty : NotRedex defs ty) => (0 tty : So (isType ty)) =>
(0 nty : NotRedex defs ty) => (0 tty : So (isTyCon ty)) =>
(0 ns : NotRedex defs s) => (0 nt : NotRedex defs t) =>
m ()
compare0' ctx (TYPE _) s t = compareType ctx s t
@ -103,7 +117,7 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
(Lam _ b1, Lam _ b2) => compare0 ctx' res.term b1.term b2.term
(E e, Lam _ b) => eta e b
(Lam _ b, E e) => eta e b
(E e, E f) => ignore $ compare0 ctx e f
(E e, E f) => compare0 ctx e f
_ => throwError $ WrongType ty s t
compare0' ctx ty@(Sig {fst, snd, _}) s t = local {mode := Equal} $
@ -123,7 +137,7 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
-- e.g. an abstract value in an abstract type, bound variables, …
E e <- pure s | _ => throwError $ WrongType ty s t
E f <- pure t | _ => throwError $ WrongType ty s t
ignore $ compare0 ctx e f
compare0 ctx e f
export covering
compareType : TContext q 0 n -> (s, t : Term q 0 n) -> m ()
@ -136,8 +150,8 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
private covering
compareType' : TContext q 0 n -> (s, t : Term q 0 n) ->
(0 ns : NotRedex defs s) => (0 ts : So (isType s)) =>
(0 nt : NotRedex defs t) => (0 tt : So (isType t)) =>
(0 ns : NotRedex defs s) => (0 ts : So (isTyCon s)) =>
(0 nt : NotRedex defs t) => (0 tt : So (isTyCon t)) =>
m ()
compareType' ctx s t = do
let err : m () = clashT (TYPE UAny) s t
@ -170,80 +184,90 @@ parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
E f <- pure t | _ => err
-- no fanciness needed here cos anything other than a neutral
-- has been inlined by whnfD
ignore $ compare0 ctx e f
compare0 ctx e f
||| assumes the elim is already typechecked! only does the work necessary
||| to calculate the overall type
private covering
computeElimType : TContext q 0 n -> (e : Elim q 0 n) ->
(0 ne : NotRedex defs e) =>
m (Term q 0 n)
computeElimType ctx (F x) = do
defs <- lookupFree' defs x
pure $ defs.type.get
computeElimType ctx (B i) = do
pure $ ctx !! i
computeElimType ctx (f :@ s) {ne} = do
(_, arg, res) <- computeElimType ctx f {ne = noOr1 ne} >>= expectPi defs
pure $ sub1 res (s :# arg)
computeElimType ctx (CasePair {pair, ret, _}) = do
pure $ sub1 ret pair
computeElimType ctx (f :% p) {ne} = do
(ty, _, _) <- computeElimType ctx f {ne = noOr1 ne} >>= expectEq defs
pure $ dsub1 ty p
computeElimType ctx (_ :# ty) = do
pure ty
private covering
replaceEnd : TContext q 0 n ->
(e : Elim q 0 n) -> DimConst -> (0 ne : NotRedex defs e) ->
m (Elim q 0 n)
replaceEnd ctx e p ne = do
(ty, l, r) <- computeElimType ctx e >>= expectEq defs
pure $ ends l r p :# dsub1 ty (K p)
namespace Elim
-- [fixme] the following code ends up repeating a lot of work in the
-- computeElimType calls. the results should be shared better
export covering %inline
compare0 : TContext q 0 n -> (e, f : Elim q 0 n) -> m (Term q 0 n)
compare0 : TContext q 0 n -> (e, f : Elim q 0 n) -> m ()
compare0 ctx e f =
let Element e ne = whnfD defs e
Element f nf = whnfD defs f
in
compare0' ctx e f
private
isSubSing : Term {} -> Bool
isSubSing (TYPE _) = False
isSubSing (Pi {res, _}) = isSubSing res.term
isSubSing (Lam {}) = False
isSubSing (Sig {fst, snd, _}) = isSubSing fst && isSubSing snd.term
isSubSing (Pair {}) = False
isSubSing (Eq {}) = True
isSubSing (DLam {}) = False
isSubSing (E e) = False
isSubSing (CloT tm th) = False
isSubSing (DCloT tm th) = False
-- [fixme] there is a better way to do this "isSubSing" stuff for sure
unless (isSubSing !(computeElimType ctx e)) $ compare0' ctx e f
private covering
compare0' : TContext q 0 n ->
(e, f : Elim q 0 n) ->
(0 ne : NotRedex defs e) => (0 nf : NotRedex defs f) =>
m (Term q 0 n)
compare0' _ e@(F x) f@(F y) = do
d <- lookupFree' defs x
let ty = d.type
-- [fixme] there is a better way to do this for sure
unless (isSubSing ty.get0 || x == y) $ clashE e f
pure ty.get
m ()
-- replace applied equalities with the appropriate end first
-- e.g. e : Eq [i ⇒ A] s t ⊢ e 0 = s : A0/i
compare0' ctx (e :% K p) f {ne} =
compare0 ctx !(replaceEnd ctx e p $ noOr1 ne) f
compare0' ctx e (f :% K q) {nf} =
compare0 ctx e !(replaceEnd ctx f q $ noOr1 nf)
compare0' _ e@(F x) f@(F y) = unless (x == y) $ clashE e f
compare0' _ e@(F _) f = clashE e f
compare0' ctx e@(B i) f@(B j) = do
let ty = ctx !! i
-- [fixme] there is a better way to do this for sure
unless (isSubSing ty || i == j) $ clashE e f
pure ty
compare0' ctx e@(B i) f@(B j) = unless (i == j) $ clashE e f
compare0' _ e@(B _) f = clashE e f
compare0' ctx (e :@ s) (f :@ t) = local {mode := Equal} $ do
Pi {arg, res, _} <- compare0 ctx e f
| ty => throwError $ ExpectedPi ty
compare0' ctx (e :@ s) (f :@ t) {ne} = local {mode := Equal} $ do
compare0 ctx e f
(_, arg, _) <- computeElimType ctx e {ne = noOr1 ne} >>= expectPi defs
compare0 ctx arg s t
pure $ sub1 res (s :# arg)
compare0' _ e@(_ :@ _) f = clashE e f
compare0' ctx (CasePair epi e _ eret _ _ ebody)
(CasePair fpi f _ fret _ _ fbody) =
(CasePair fpi f _ fret _ _ fbody) {ne} =
local {mode := Equal} $ do
ty@(Sig {fst, snd, _}) <- compare0 ctx e f
| ty => throwError $ ExpectedSig ty
unless (epi == fpi) $ throwError $ ClashQ epi fpi
compareType (ctx :< ty) eret.term fret.term
compare0 (ctx :< fst :< snd.term) (substCasePairRet ty eret)
compare0 ctx e f
ety <- computeElimType ctx e {ne = noOr1 ne}
compareType (ctx :< ety) eret.term fret.term
(fst, snd) <- expectSig defs ety
compare0 (ctx :< fst :< snd.term) (substCasePairRet ety eret)
ebody.term fbody.term
pure $ sub1 eret e
unless (epi == fpi) $ throwError $ ClashQ epi fpi
compare0' _ e@(CasePair {}) f = clashE e f
compare0' ctx (e :% p) (f :% q) = local {mode := Equal} $ do
Eq {ty, _} <- compare0 ctx e f
| ty => throwError $ ExpectedEq ty
unless (p == q) $ throwError $ ClashD p q
pure $ dsub1 ty p
compare0' _ e@(_ :% _) f = clashE e f
compare0' ctx (s :# a) (t :# b) = do
compareType ctx a b
compare0 ctx a s t
pure b
compare0' _ e@(_ :# _) f = clashE e f
@ -268,14 +292,15 @@ parameters {auto _ : (HasDefs' q _ m, HasErr q m, Eq q)}
compareType defs (map (/// th) ctx) (s /// th) (t /// th)
namespace Elim
-- can't return the type since it might be different in each dsubst ☹
||| you don't have to pass the type in but the arguments must still be
||| of the same type!!
export covering %inline
compare : (e, f : Elim q d n) -> m ()
compare e f = do
defs <- ask
runReaderT {m} (MakeEnv {mode}) $
for_ (splits eq) $ \th =>
ignore $ compare0 defs (map (/// th) ctx) (e /// th) (f /// th)
compare0 defs (map (/// th) ctx) (e /// th) (f /// th)
namespace Term
export covering %inline

View file

@ -19,37 +19,6 @@ public export
CanTC q = CanTC' q IsGlobal
private covering %inline
expectTYPE : CanTC' q _ m => Term q d n -> m Universe
expectTYPE s =
case whnfD !ask s of
Element (TYPE l) _ => pure l
_ => throwError $ ExpectedTYPE s
private covering %inline
expectPi : CanTC' q _ m => Term q d n ->
m (q, Term q d n, ScopeTerm q d n)
expectPi ty =
case whnfD !ask ty of
Element (Pi qty _ arg res) _ => pure (qty, arg, res)
_ => throwError $ ExpectedPi ty
private covering %inline
expectSig : CanTC' q _ m => Term q d n ->
m (Term q d n, ScopeTerm q d n)
expectSig ty =
case whnfD !ask ty of
Element (Sig _ arg res) _ => pure (arg, res)
_ => throwError $ ExpectedSig ty
private covering %inline
expectEq : CanTC' q _ m => Term q d n ->
m (DScopeTerm q d n, Term q d n, Term q d n)
expectEq ty =
case whnfD !ask ty of
Element (Eq _ ty l r) _ => pure (ty, l, r)
_ => throwError $ ExpectedEq ty
private
popQs : HasErr q m => IsQty q =>
@ -133,13 +102,13 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
check' ctx sg (TYPE l) _ ty = do
-- if < ' then Ψ | Γ ⊢ Type · 0 ⇐ Type ' ⊳ 𝟎
l' <- expectTYPE ty
l' <- expectTYPE !ask ty
expectEqualQ zero sg.fst
unless (l < l') $ throwError $ BadUniverse l l'
pure $ zeroFor ctx
check' ctx sg (Pi qty _ arg res) _ ty = do
l <- expectTYPE ty
l <- expectTYPE !ask ty
expectEqualQ zero sg.fst
-- if Ψ | Γ ⊢ A · 0 ⇐ Type 𝟎
ignore $ check0 ctx arg (TYPE l)
@ -151,14 +120,14 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
pure $ zeroFor ctx
check' ctx sg (Lam _ body) _ ty = do
(qty, arg, res) <- expectPi ty
(qty, arg, res) <- expectPi !ask ty
-- if Ψ | Γ, x · πσ : A ⊢ t · σ ⇐ B ⊳ Σ, x · σπ
qout <- check (extendTy arg (sg.fst * qty) ctx) sg body.term res.term
-- then Ψ | Γ ⊢ λx. t · σ ⇐ (x · π : A) → B ⊳ Σ
popQ (sg.fst * qty) qout
check' ctx sg (Sig _ fst snd) _ ty = do
l <- expectTYPE ty
l <- expectTYPE !ask ty
expectEqualQ zero sg.fst
-- if Ψ | Γ ⊢ A · 0 ⇐ Type 𝟎
ignore $ check0 ctx fst (TYPE l)
@ -170,7 +139,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
pure $ zeroFor ctx
check' ctx sg (Pair fst snd) _ ty = do
(tfst, tsnd) <- expectSig ty
(tfst, tsnd) <- expectSig !ask ty
-- if Ψ | Γ ⊢ s · σ ⇐ A ⊳ Σ₁
qfst <- check ctx sg fst tfst
let tsnd = sub1 tsnd (fst :# tfst)
@ -180,7 +149,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
pure $ qfst + qsnd
check' ctx sg (Eq i t l r) _ ty = do
u <- expectTYPE ty
u <- expectTYPE !ask ty
expectEqualQ zero sg.fst
-- if Ψ, i | Γ ⊢ A · 0 ⇐ Type 𝟎
case t of
@ -194,7 +163,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
pure $ zeroFor ctx
check' ctx sg (DLam i body) _ ty = do
(ty, l, r) <- expectEq ty
(ty, l, r) <- expectEq !ask ty
-- if Ψ, i | Γ ⊢ t · σ ⇐ A ⊳ Σ
qout <- check (extendDim ctx) sg body.term ty.term
let eqs = makeDimEq ctx.dctx
@ -234,7 +203,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
infer' ctx sg (fun :@ arg) _ = do
-- if Ψ | Γ ⊢ f · σ ⇒ (x · π : A) → B ⊳ Σ₁
funres <- infer ctx sg fun
(qty, argty, res) <- expectPi funres.type
(qty, argty, res) <- expectPi !ask funres.type
-- if Ψ | Γ ⊢ s · σ∧π ⇐ A ⊳ Σ₂
-- (0∧π = σ∧0 = 0; σ∧π = σ otherwise)
argout <- check ctx (subjMult sg qty) arg argty
@ -250,7 +219,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
-- if Ψ | Γ ⊢ pair · 1 ⇒ (x : A) × B ⊳ Σ₁
pairres <- infer ctx sone pair
ignore $ check0 (extendTy pairres.type zero ctx) ret.term (TYPE UAny)
(tfst, tsnd) <- expectSig pairres.type
(tfst, tsnd) <- expectSig !ask pairres.type
-- if Ψ | Γ, x · π : A, y · π : B ⊢ σ body ⇐ ret[(x, y)]
-- ⊳ Σ₂, x · π₁, y · π₂
-- if π₁, π₂ ≤ π
@ -263,11 +232,10 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
qout = pi * pairres.qout + bodyout
}
infer' ctx sg (fun :% dim) _ = do
-- if Ψ | Γ ⊢ f · σ ⇒ Eq [i ⇒ A] l r ⊳ Σ
InfRes {type, qout} <- infer ctx sg fun
(ty, _, _) <- expectEq type
(ty, _, _) <- expectEq !ask type
-- then Ψ | Γ ⊢ f p · σ ⇒ Ap ⊳ Σ
pure $ InfRes {type = dsub1 ty dim, qout}

View file

@ -177,3 +177,33 @@ substCasePairRet : Term q d n -> ScopeTerm q d n -> Term q d (2 + n)
substCasePairRet dty retty =
let arg = Pair (BVT 0) (BVT 1) :# (dty // fromNat 2) in
retty.term // (arg ::: shift 2)
parameters {auto _ : HasErr q m} (defs : Definitions' q _)
export covering %inline
expectTYPE : Term q d n -> m Universe
expectTYPE s =
case fst $ whnfD defs s of
TYPE l => pure l
_ => throwError $ ExpectedTYPE s
export covering %inline
expectPi : Term q d n -> m (q, Term q d n, ScopeTerm q d n)
expectPi s =
case fst $ whnfD defs s of
Pi {qty, arg, res, _} => pure (qty, arg, res)
_ => throwError $ ExpectedPi s
export covering %inline
expectSig : Term q d n -> m (Term q d n, ScopeTerm q d n)
expectSig s =
case fst $ whnfD defs s of
Sig {fst, snd, _} => pure (fst, snd)
_ => throwError $ ExpectedSig s
export covering %inline
expectEq : Term q d n -> m (DScopeTerm q d n, Term q d n, Term q d n)
expectEq s =
case fst $ whnfD defs s of
Eq {ty, l, r, _} => pure (ty, l, r)
_ => throwError $ ExpectedEq s