add local bindings to context

- without this, inside the body of `let x = e in …`, the typechecker
  would forget that `x = e`
- now bound variables can reduce, if they have a definition, so RedexTest
  needs to take the context too
This commit is contained in:
rhiannon morris 2023-12-07 01:35:39 +01:00
parent cdf1ec6deb
commit 03c197bd04
13 changed files with 300 additions and 211 deletions

View file

@ -88,16 +88,16 @@ parameters {opts : LayoutOpts} (showContext : Bool)
public export
Erase : List (Type -> Type)
Erase = [Except Error, DefsReader, NameGen]
Erase = [Except Error, NameGen]
export
liftWhnf : Eff Whnf a -> Eff Erase a
liftWhnf act = lift $ wrapErr WrapTypeError act
export covering
computeElimType : ErasureContext d n -> SQty -> Elim d n -> Eff Erase (Term d n)
computeElimType ctx sg e = do
defs <- askAt DEFS
computeElimType : Q.Definitions -> ErasureContext d n -> SQty ->
Elim d n -> Eff Erase (Term d n)
computeElimType defs ctx sg e = do
let ctx = toWhnfContext ctx
liftWhnf $ do
Element e _ <- whnf defs ctx sg e
@ -106,10 +106,11 @@ computeElimType ctx sg e = do
private %macro
wrapExpect : TTImp ->
Elab (TyContext d n -> Loc -> Term d n -> Eff Erase a)
Elab (Q.Definitions -> TyContext d n -> Loc ->
Term d n -> Eff Erase a)
wrapExpect f_ = do
f <- check `(\x => ~(f_) x)
pure $ \ctx, loc, s => liftWhnf $ f !(askAt DEFS) ctx SZero loc s
pure $ \defs, ctx, loc, s => liftWhnf $ f defs ctx SZero loc s
public export
@ -119,27 +120,36 @@ record EraseElimResult d n where
term : U.Term n
export covering
eraseTerm' : (defs : Q.Definitions) -> (ctx : ErasureContext d n) ->
(ty, tm : Q.Term d n) ->
(0 _ : NotRedex defs (toWhnfContext ctx) SZero ty) =>
Eff Erase (U.Term n)
-- "Ψ | Γ | Σ ⊢ s ⤋ s' ⇐ A" for `s' <- eraseTerm (Ψ,Γ,Σ) A s`
--
-- in the below comments, Ψ, Γ, Σ are implicit and
-- only their extensions are written
export covering
eraseTerm : ErasureContext d n ->
eraseTerm : Q.Definitions -> ErasureContext d n ->
(ty, tm : Q.Term d n) -> Eff Erase (U.Term n)
eraseTerm defs ctx ty tm = do
Element ty _ <- liftWhnf $ Interface.whnf defs (toWhnfContext ctx) SZero ty
eraseTerm' defs ctx ty tm
-- "Ψ | Γ | Σ ⊢ e ⤋ e' ⇒ A" for `EraRes A e' <- eraseElim (Ψ,Γ,Σ) e`
export covering
eraseElim : ErasureContext d n -> (tm : Q.Elim d n) ->
eraseElim : Q.Definitions -> ErasureContext d n -> (tm : Q.Elim d n) ->
Eff Erase (EraseElimResult d n)
eraseTerm ctx _ s@(TYPE {}) =
eraseTerm' defs ctx _ s@(TYPE {}) =
throw $ CompileTimeOnly ctx s
eraseTerm ctx _ s@(IOState {}) =
eraseTerm' defs ctx _ s@(IOState {}) =
throw $ CompileTimeOnly ctx s
eraseTerm ctx _ s@(Pi {}) =
eraseTerm' defs ctx _ s@(Pi {}) =
throw $ CompileTimeOnly ctx s
-- x : A | 0.x ⊢ s ⤋ s' ⇐ B
@ -149,66 +159,66 @@ eraseTerm ctx _ s@(Pi {}) =
-- x : A | π.x ⊢ s ⤋ s' ⇐ B π ≠ 0
-- ----------------------------------------
-- (λ x ⇒ s) ⤋ (λ x ⇒ s') ⇐ π.(x : A) → B
eraseTerm ctx ty (Lam body loc) = do
eraseTerm' defs ctx ty (Lam body loc) = do
let x = body.name
(qty, arg, res) <- wrapExpect `(expectPi) ctx loc ty
body <- eraseTerm (extendTy qty x arg ctx) res.term body.term
(qty, arg, res) <- wrapExpect `(expectPi) defs ctx loc ty
body <- eraseTerm defs (extendTy qty x arg ctx) res.term body.term
pure $ case isErased qty of
Kept => U.Lam x body loc
Erased => sub1 (Erased loc) body
eraseTerm ctx _ s@(Sig {}) =
eraseTerm' defs ctx _ s@(Sig {}) =
throw $ CompileTimeOnly ctx s
-- s ⤋ s' ⇐ A t ⤋ t' ⇐ B[s/x]
-- ---------------------------------
-- (s, t) ⤋ (s', t') ⇐ (x : A) × B
eraseTerm ctx ty (Pair fst snd loc) = do
(a, b) <- wrapExpect `(expectSig) ctx loc ty
eraseTerm' defs ctx ty (Pair fst snd loc) = do
(a, b) <- wrapExpect `(expectSig) defs ctx loc ty
let b = sub1 b (Ann fst a a.loc)
fst <- eraseTerm ctx a fst
snd <- eraseTerm ctx b snd
fst <- eraseTerm defs ctx a fst
snd <- eraseTerm defs ctx b snd
pure $ Pair fst snd loc
eraseTerm ctx _ s@(Enum {}) =
eraseTerm' defs ctx _ s@(Enum {}) =
throw $ CompileTimeOnly ctx s
-- '𝐚 ⤋ '𝐚 ⇐ {⋯}
eraseTerm ctx _ (Tag tag loc) =
eraseTerm' defs ctx _ (Tag tag loc) =
pure $ Tag tag loc
eraseTerm ctx ty s@(Eq {}) =
eraseTerm' defs ctx ty s@(Eq {}) =
throw $ CompileTimeOnly ctx s
-- 𝑖 ⊢ s ⤋ s' ⇐ A
-- ---------------------------------
-- (δ 𝑖 ⇒ s) ⤋ s' ⇐ Eq (𝑖 ⇒ A) l r
eraseTerm ctx ty (DLam body loc) = do
a <- fst <$> wrapExpect `(expectEq) ctx loc ty
eraseTerm (extendDim body.name ctx) a.term body.term
eraseTerm' defs ctx ty (DLam body loc) = do
a <- fst <$> wrapExpect `(expectEq) defs ctx loc ty
eraseTerm defs (extendDim body.name ctx) a.term body.term
eraseTerm ctx _ s@(NAT {}) =
eraseTerm' defs ctx _ s@(NAT {}) =
throw $ CompileTimeOnly ctx s
-- n ⤋ n ⇐
eraseTerm _ _ (Nat n loc) =
eraseTerm' _ _ _ (Nat n loc) =
pure $ Nat n loc
-- s ⤋ s' ⇐
-- -----------------------
-- succ s ⤋ succ s' ⇐
eraseTerm ctx ty (Succ p loc) = do
p <- eraseTerm ctx ty p
eraseTerm' defs ctx ty (Succ p loc) = do
p <- eraseTerm defs ctx ty p
pure $ Succ p loc
eraseTerm ctx ty s@(STRING {}) =
eraseTerm' defs ctx ty s@(STRING {}) =
throw $ CompileTimeOnly ctx s
-- s ⤋ s ⇐ String
eraseTerm _ _ (Str s loc) =
eraseTerm' _ _ _ (Str s loc) =
pure $ Str s loc
eraseTerm ctx ty s@(BOX {}) =
eraseTerm' defs ctx ty s@(BOX {}) =
throw $ CompileTimeOnly ctx s
-- [s] ⤋ ⌷ ⇐ [0.A]
@ -216,48 +226,49 @@ eraseTerm ctx ty s@(BOX {}) =
-- π ≠ 0 s ⤋ s' ⇐ A
-- --------------------
-- [s] ⤋ s' ⇐ [π.A]
eraseTerm ctx ty (Box val loc) = do
(qty, a) <- wrapExpect `(expectBOX) ctx loc ty
eraseTerm' defs ctx ty (Box val loc) = do
(qty, a) <- wrapExpect `(expectBOX) defs ctx loc ty
case isErased qty of
Erased => pure $ Erased loc
Kept => eraseTerm ctx a val
Kept => eraseTerm defs ctx a val
-- s ⤋ s' ⇐ A
-- ---------------------------------
-- let0 x = e in s ⤋ s'[⌷/x] ⇐ A
--
-- e ⤋ e' s ⤋ s' ⇐ A π ≠ 0
-- e ⤋ e' ⇒ E π ≠ 0
-- x : E ≔ e ⊢ s ⤋ s' ⇐ A
-- -------------------------------------
-- letπ x = e in s ⤋ let x = e' in s'
eraseTerm ctx ty (Let pi e s loc) = do
eraseTerm' defs ctx ty (Let pi e s loc) = do
let x = s.name
case isErased pi of
Erased => do
ety <- computeElimType ctx SZero e
s' <- eraseTerm (extendTy pi x ety ctx) (weakT 1 ty) s.term
ety <- computeElimType defs ctx SZero e
s' <- eraseTerm defs (extendTyLet pi x ety (E e) ctx) (weakT 1 ty) s.term
pure $ sub1 (Erased e.loc) s'
Kept => do
EraRes ety e' <- eraseElim ctx e
s' <- eraseTerm (extendTy pi x ety ctx) (weakT 1 ty) s.term
EraRes ety e' <- eraseElim defs ctx e
s' <- eraseTerm defs (extendTyLet pi x ety (E e) ctx) (weakT 1 ty) s.term
pure $ Let x e' s' loc
-- e ⤋ e' ⇒ B
-- ------------
-- e ⤋ e' ⇐ A
eraseTerm ctx ty (E e) =
term <$> eraseElim ctx e
eraseTerm' defs ctx ty (E e) =
term <$> eraseElim defs ctx e
eraseTerm ctx ty (CloT (Sub term th)) =
eraseTerm ctx ty $ pushSubstsWith' id th term
eraseTerm' defs ctx ty (CloT (Sub term th)) =
eraseTerm defs ctx ty $ pushSubstsWith' id th term
eraseTerm ctx ty (DCloT (Sub term th)) =
eraseTerm ctx ty $ pushSubstsWith' th id term
eraseTerm' defs ctx ty (DCloT (Sub term th)) =
eraseTerm defs ctx ty $ pushSubstsWith' th id term
-- defω x : A = s
-- ----------------
-- x ⤋ x ⇒ A
eraseElim ctx e@(F x u loc) = do
Just def <- asksAt DEFS $ lookup x
eraseElim defs ctx e@(F x u loc) = do
let Just def = lookup x defs
| Nothing => throw $ notInScope loc x
case isErased def.qty.qty of
Erased => throw $ CompileTimeOnly ctx $ E e
@ -266,10 +277,10 @@ eraseElim ctx e@(F x u loc) = do
-- π.x ∈ Σ π ≠ 0
-- -----------------
-- x ⤋ x ⇒ A
eraseElim ctx e@(B i loc) = do
eraseElim defs ctx e@(B i loc) = do
case isErased $ ctx.qtys !!! i of
Erased => throw $ CompileTimeOnly ctx $ E e
Kept => pure $ EraRes (ctx.tctx !! i) $ B i loc
Kept => pure $ EraRes (ctx.tctx !! i).type $ B i loc
-- f ⤋ f' ⇒ π.(x : A) → B s ⤋ s' ⇒ A π ≠ 0
-- ---------------------------------------------
@ -278,13 +289,13 @@ eraseElim ctx e@(B i loc) = do
-- f ⤋ f' ⇒ 0.(x : A) → B
-- -------------------------
-- f s ⤋ f' ⇒ B[s/x]
eraseElim ctx (App fun arg loc) = do
efun <- eraseElim ctx fun
(qty, targ, tres) <- wrapExpect `(expectPi) ctx loc efun.type
eraseElim defs ctx (App fun arg loc) = do
efun <- eraseElim defs ctx fun
(qty, targ, tres) <- wrapExpect `(expectPi) defs ctx loc efun.type
let ty = sub1 tres (Ann arg targ arg.loc)
case isErased qty of
Erased => pure $ EraRes ty efun.term
Kept => do arg <- eraseTerm ctx targ arg
Kept => do arg <- eraseTerm defs ctx targ arg
pure $ EraRes ty $ App efun.term arg loc
-- e ⇒ (x : A) × B
@ -298,16 +309,16 @@ eraseElim ctx (App fun arg loc) = do
-- (caseρ e return z ⇒ R of {(x, y) ⇒ s}) ⤋
-- ⤋
-- let xy = e' in let x = fst xy in let y = snd xy in s' ⇒ R[e/z]
eraseElim ctx (CasePair qty pair ret body loc) = do
eraseElim defs ctx (CasePair qty pair ret body loc) = do
let [< x, y] = body.names
case isErased qty of
Kept => do
EraRes ety eterm <- eraseElim ctx pair
EraRes ety eterm <- eraseElim defs ctx pair
let ty = sub1 (ret // shift 2) $
Ann (Pair (BVT 0 loc) (BVT 1 loc) loc) (weakT 2 ety) loc
(tfst, tsnd) <- wrapExpect `(expectSig) ctx loc ety
(tfst, tsnd) <- wrapExpect `(expectSig) defs ctx loc ety
let ctx' = extendTyN [< (qty, x, tfst), (qty, y, tsnd.term)] ctx
body' <- eraseTerm ctx' ty body.term
body' <- eraseTerm defs ctx' ty body.term
p <- mnb "p" loc
pure $ EraRes (sub1 ret pair) $
Let p eterm
@ -316,28 +327,28 @@ eraseElim ctx (CasePair qty pair ret body loc) = do
(body' // (B VZ loc ::: B (VS VZ) loc ::: shift 3))
loc) loc) loc
Erased => do
ety <- computeElimType ctx SOne pair
ety <- computeElimType defs ctx SOne pair
let ty = sub1 (ret // shift 2) $
Ann (Pair (BVT 0 loc) (BVT 1 loc) loc) (weakT 2 ety) loc
(tfst, tsnd) <- wrapExpect `(expectSig) ctx loc ety
(tfst, tsnd) <- wrapExpect `(expectSig) defs ctx loc ety
let ctx' = extendTyN0 [< (x, tfst), (y, tsnd.term)] ctx
body' <- eraseTerm ctx' ty body.term
body' <- eraseTerm defs ctx' ty body.term
pure $ EraRes (sub1 ret pair) $ subN [< Erased loc, Erased loc] body'
-- e ⤋ e' ⇒ (x : A) × B
-- ----------------------
-- fst e ⤋ fst e' ⇒ A
eraseElim ctx (Fst pair loc) = do
epair <- eraseElim ctx pair
a <- fst <$> wrapExpect `(expectSig) ctx loc epair.type
eraseElim defs ctx (Fst pair loc) = do
epair <- eraseElim defs ctx pair
a <- fst <$> wrapExpect `(expectSig) defs ctx loc epair.type
pure $ EraRes a $ Fst epair.term loc
-- e ⤋ e' ⇒ (x : A) × B
-- -----------------------------
-- snd e ⤋ snd e' ⇒ B[fst e/x]
eraseElim ctx (Snd pair loc) = do
epair <- eraseElim ctx pair
b <- snd <$> wrapExpect `(expectSig) ctx loc epair.type
eraseElim defs ctx (Snd pair loc) = do
epair <- eraseElim defs ctx pair
b <- snd <$> wrapExpect `(expectSig) defs ctx loc epair.type
pure $ EraRes (sub1 b (Fst pair loc)) $ Snd epair.term loc
-- caseρ e return z ⇒ R of {} ⤋ absurd ⇒ R[e/z]
@ -349,23 +360,23 @@ eraseElim ctx (Snd pair loc) = do
-- e ⤋ e' ⇒ A sᵢ ⤋ s'ᵢ ⇐ R[𝐚ᵢ/z] ρ ≠ 0 i ≠ 0
-- -------------------------------------------------------------------
-- caseρ e return z ⇒ R of {𝐚ᵢ ⇒ sᵢ} ⤋ case e of {𝐚ᵢ ⇒ s'ᵢ} ⇒ R[e/z]
eraseElim ctx e@(CaseEnum qty tag ret arms loc) = do
eraseElim defs ctx e@(CaseEnum qty tag ret arms loc) = do
let ty = sub1 ret tag
case isErased qty of
Erased => case SortedMap.toList arms of
[] => pure $ EraRes ty $ Absurd loc
[(t, rhs)] => do
let ty' = sub1 ret (Ann (Tag t loc) (enum [t] loc) loc)
rhs' <- eraseTerm ctx ty' rhs
rhs' <- eraseTerm defs ctx ty' rhs
pure $ EraRes ty rhs'
_ => throw $ CompileTimeOnly ctx $ E e
Kept => case List1.fromList $ SortedMap.toList arms of
Nothing => pure $ EraRes ty $ Absurd loc
Just arms => do
etag <- eraseElim ctx tag
etag <- eraseElim defs ctx tag
arms <- for arms $ \(t, rhs) => do
let ty' = sub1 ret (Ann (Tag t loc) etag.type loc)
rhs' <- eraseTerm ctx ty' rhs
rhs' <- eraseTerm defs ctx ty' rhs
pure (t, rhs')
pure $ EraRes ty $ CaseEnum etag.term arms loc
@ -382,12 +393,12 @@ eraseElim ctx e@(CaseEnum qty tag ret arms loc) = do
-- caseρ n return z ⇒ R of {0 ⇒ z; succ m, 0.ih ⇒ s}
-- ⤋
-- case n' of {0 ⇒ z'; succ m ⇒ s'[⌷/ih]} ⇒ R[n/z]
eraseElim ctx (CaseNat qty qtyIH nat ret zero succ loc) = do
eraseElim defs ctx (CaseNat qty qtyIH nat ret zero succ loc) = do
let ty = sub1 ret nat
enat <- eraseElim ctx nat
zero <- eraseTerm ctx (sub1 ret (Ann (Zero loc) (NAT loc) loc)) zero
enat <- eraseElim defs ctx nat
zero <- eraseTerm defs ctx (sub1 ret (Ann (Zero loc) (NAT loc) loc)) zero
let [< p, ih] = succ.names
succ' <- eraseTerm
succ' <- eraseTerm defs
(extendTyN [< (qty, p, NAT loc),
(qtyIH, ih, sub1 (ret // shift 1) (BV 0 loc))] ctx)
(sub1 (ret // shift 2) (Ann (Succ (BVT 1 loc) loc) (NAT loc) loc))
@ -404,56 +415,56 @@ eraseElim ctx (CaseNat qty qtyIH nat ret zero succ loc) = do
-- b ⇒ [π.A] x : A | 0.x ⊢ s ⤋ s' ⇐ R[[x]∷[0.A]/z] πρ = 0
-- -------------------------------------------------------------
-- caseρ b return z ⇒ R of {[x] ⇒ s} ⤋ s'[⌷/x] ⇒ R[b/z]
eraseElim ctx (CaseBox qty box ret body loc) = do
tbox <- computeElimType ctx SOne box -- [fixme] is there any way to avoid this?
(pi, tinner) <- wrapExpect `(expectBOX) ctx loc tbox
eraseElim defs ctx (CaseBox qty box ret body loc) = do
tbox <- computeElimType defs ctx SOne box
(pi, tinner) <- wrapExpect `(expectBOX) defs ctx loc tbox
let ctx' = extendTy (pi * qty) body.name tinner ctx
bty = sub1 (ret // shift 1) $
Ann (Box (BVT 0 loc) loc) (weakT 1 tbox) loc
case isErased $ qty * pi of
Kept => do
ebox <- eraseElim ctx box
ebody <- eraseTerm ctx' bty body.term
ebox <- eraseElim defs ctx box
ebody <- eraseTerm defs ctx' bty body.term
pure $ EraRes (sub1 ret box) $ Let body.name ebox.term ebody loc
Erased => do
body' <- eraseTerm ctx' bty body.term
body' <- eraseTerm defs ctx' bty body.term
pure $ EraRes (sub1 ret box) $ body' // one (Erased loc)
-- f ⤋ f' ⇒ Eq (𝑖 ⇒ A) l r
-- ------------------------------
-- f @r ⤋ f' ⇒ Ar/𝑖
eraseElim ctx (DApp fun arg loc) = do
efun <- eraseElim ctx fun
a <- fst <$> wrapExpect `(expectEq) ctx loc efun.type
eraseElim defs ctx (DApp fun arg loc) = do
efun <- eraseElim defs ctx fun
a <- fst <$> wrapExpect `(expectEq) defs ctx loc efun.type
pure $ EraRes (dsub1 a arg) efun.term
-- s ⤋ s' ⇐ A
-- ----------------
-- s ∷ A ⤋ s' ⇒ A
eraseElim ctx (Ann tm ty loc) =
EraRes ty <$> eraseTerm ctx ty tm
eraseElim defs ctx (Ann tm ty loc) =
EraRes ty <$> eraseTerm defs ctx ty tm
-- s ⤋ s' ⇐ Ap/𝑖
-- -----------------------------------
-- coe (𝑖 ⇒ A) @p @q s ⤋ s' ⇒ Aq/𝑖
eraseElim ctx (Coe ty p q val loc) = do
val <- eraseTerm ctx (dsub1 ty p) val
eraseElim defs ctx (Coe ty p q val loc) = do
val <- eraseTerm defs ctx (dsub1 ty p) val
pure $ EraRes (dsub1 ty q) val
-- s ⤋ s' ⇐ A
-- --------------------------------
-- comp A @p @q s @r {⋯} ⤋ s' ⇒ A
eraseElim ctx (Comp ty p q val r zero one loc) =
EraRes ty <$> eraseTerm ctx ty val
eraseElim defs ctx (Comp ty p q val r zero one loc) =
EraRes ty <$> eraseTerm defs ctx ty val
eraseElim ctx t@(TypeCase ty ret arms def loc) =
eraseElim defs ctx t@(TypeCase ty ret arms def loc) =
throw $ CompileTimeOnly ctx $ E t
eraseElim ctx (CloE (Sub term th)) =
eraseElim ctx $ pushSubstsWith' id th term
eraseElim defs ctx (CloE (Sub term th)) =
eraseElim defs ctx $ pushSubstsWith' id th term
eraseElim ctx (DCloE (Sub term th)) =
eraseElim ctx $ pushSubstsWith' th id term
eraseElim defs ctx (DCloE (Sub term th)) =
eraseElim defs ctx $ pushSubstsWith' th id term
export
@ -539,8 +550,8 @@ trimLets (Erased loc) = Erased loc
export covering
eraseDef : Name -> Q.Definition -> Eff Erase U.Definition
eraseDef name def@(MkDef qty type body scheme isMain loc) =
eraseDef : Q.Definitions -> Name -> Q.Definition -> Eff Erase U.Definition
eraseDef defs name def@(MkDef qty type body scheme isMain loc) =
wrapErr (WhileErasing name def) $
case isErased qty.qty of
Erased => do
@ -552,4 +563,4 @@ eraseDef name def@(MkDef qty type body scheme isMain loc) =
Nothing => case body of
Postulate => throw $ Postulate loc name
Concrete body => KeptDef isMain . trimLets <$>
eraseTerm empty type body
eraseTerm defs empty type body