quox/lib/Quox/Reduce.idr

951 lines
37 KiB
Idris
Raw Normal View History

module Quox.Reduce
2022-04-23 18:21:30 -04:00
import Quox.No
import Quox.Syntax
import Quox.Definition
2023-05-21 14:09:34 -04:00
import Quox.Displace
2023-04-15 09:13:01 -04:00
import Quox.Typing.Context
import Quox.Typing.Error
2023-03-26 10:09:47 -04:00
import Data.SnocVect
import Data.Maybe
import Data.List
2023-05-01 21:06:25 -04:00
import Control.Eff
2022-04-23 18:21:30 -04:00
2022-05-02 16:38:37 -04:00
%default total
2022-04-23 18:21:30 -04:00
2023-05-01 21:06:25 -04:00
public export
2023-05-21 14:11:01 -04:00
Whnf : List (Type -> Type)
Whnf = [NameGen, Except Error]
2023-05-01 21:06:25 -04:00
export
2023-05-21 14:11:01 -04:00
runWhnfWith : NameSuf -> Eff Whnf a -> (Either Error a, NameSuf)
2023-05-01 21:06:25 -04:00
runWhnfWith suf act = extract $ runStateAt GEN suf $ runExcept act
export
2023-05-21 14:11:01 -04:00
runWhnf : Eff Whnf a -> Either Error a
2023-05-01 21:06:25 -04:00
runWhnf = fst . runWhnfWith 0
public export
0 RedexTest : TermLike -> Type
RedexTest tm = {d, n : Nat} -> Definitions -> tm d n -> Bool
public export
2023-05-21 14:11:01 -04:00
interface CanWhnf (0 tm : TermLike) (0 isRedex : RedexTest tm) | tm
where
whnf : {d, n : Nat} -> (defs : Definitions) ->
2023-04-17 17:58:24 -04:00
(ctx : WhnfContext d n) ->
2023-05-21 14:11:01 -04:00
tm d n -> Eff Whnf (Subset (tm d n) (No . isRedex defs))
2023-04-15 09:13:01 -04:00
public export %inline
2023-05-21 14:11:01 -04:00
whnf0 : {d, n : Nat} -> {0 isRedex : RedexTest tm} -> CanWhnf tm isRedex =>
(defs : Definitions) -> WhnfContext d n -> tm d n -> Eff Whnf (tm d n)
whnf0 defs ctx t = fst <$> whnf defs ctx t
public export
2023-05-21 14:11:01 -04:00
0 IsRedex, NotRedex : {isRedex : RedexTest tm} -> CanWhnf tm isRedex =>
Definitions -> Pred (tm d n)
IsRedex defs = So . isRedex defs
NotRedex defs = No . isRedex defs
public export
0 NonRedex : (tm : TermLike) -> {isRedex : RedexTest tm} ->
2023-05-21 14:11:01 -04:00
CanWhnf tm isRedex => (d, n : Nat) -> (defs : Definitions) -> Type
NonRedex tm d n defs = Subset (tm d n) (NotRedex defs)
public export %inline
2023-05-21 14:11:01 -04:00
nred : {0 isRedex : RedexTest tm} -> (0 _ : CanWhnf tm isRedex) =>
(t : tm d n) -> (0 nr : NotRedex defs t) => NonRedex tm d n defs
nred t = Element t nr
2022-04-23 18:21:30 -04:00
2023-01-20 20:34:28 -05:00
public export %inline
isLamHead : Elim {} -> Bool
2023-05-01 21:06:25 -04:00
isLamHead (Ann (Lam {}) (Pi {}) _) = True
isLamHead (Coe {}) = True
isLamHead _ = False
2022-05-25 09:59:58 -04:00
public export %inline
isDLamHead : Elim {} -> Bool
2023-05-01 21:06:25 -04:00
isDLamHead (Ann (DLam {}) (Eq {}) _) = True
isDLamHead (Coe {}) = True
isDLamHead _ = False
2022-05-25 09:59:58 -04:00
2023-01-26 13:54:46 -05:00
public export %inline
isPairHead : Elim {} -> Bool
2023-05-01 21:06:25 -04:00
isPairHead (Ann (Pair {}) (Sig {}) _) = True
isPairHead (Coe {}) = True
isPairHead _ = False
2023-01-26 13:54:46 -05:00
2023-08-06 04:46:55 -04:00
public export %inline
isWHead : Elim {} -> Bool
isWHead (Ann (Sup {}) (W {}) _) = True
isWHead (Coe {}) = True
isWHead _ = False
public export %inline
isTagHead : Elim {} -> Bool
2023-05-01 21:06:25 -04:00
isTagHead (Ann (Tag {}) (Enum {}) _) = True
isTagHead (Coe {}) = True
isTagHead _ = False
2023-03-26 08:40:54 -04:00
public export %inline
isNatHead : Elim {} -> Bool
2023-05-01 21:06:25 -04:00
isNatHead (Ann (Zero {}) (Nat {}) _) = True
isNatHead (Ann (Succ {}) (Nat {}) _) = True
isNatHead (Coe {}) = True
isNatHead _ = False
2023-03-26 08:40:54 -04:00
2023-03-31 13:11:35 -04:00
public export %inline
isBoxHead : Elim {} -> Bool
2023-05-01 21:06:25 -04:00
isBoxHead (Ann (Box {}) (BOX {}) _) = True
isBoxHead (Coe {}) = True
isBoxHead _ = False
2023-03-31 13:11:35 -04:00
2022-05-25 09:59:58 -04:00
public export %inline
isE : Term {} -> Bool
2023-05-01 21:06:25 -04:00
isE (E {}) = True
isE _ = False
2022-05-25 09:59:58 -04:00
public export %inline
isAnn : Elim {} -> Bool
2023-05-01 21:06:25 -04:00
isAnn (Ann {}) = True
isAnn _ = False
2023-04-03 11:46:23 -04:00
||| true if a term is syntactically a type.
public export %inline
2023-04-15 09:13:01 -04:00
isTyCon : Term {} -> Bool
2023-05-01 21:06:25 -04:00
isTyCon (TYPE {}) = True
isTyCon (Pi {}) = True
isTyCon (Lam {}) = False
isTyCon (Sig {}) = True
isTyCon (Pair {}) = False
2023-08-13 14:55:14 -04:00
isTyCon (W {}) = True
isTyCon (Sup {}) = False
2023-05-01 21:06:25 -04:00
isTyCon (Enum {}) = True
isTyCon (Tag {}) = False
isTyCon (Eq {}) = True
isTyCon (DLam {}) = False
isTyCon (Nat {}) = True
isTyCon (Zero {}) = False
isTyCon (Succ {}) = False
isTyCon (BOX {}) = True
isTyCon (Box {}) = False
isTyCon (E {}) = False
isTyCon (CloT {}) = False
isTyCon (DCloT {}) = False
2023-04-03 11:46:23 -04:00
2023-08-13 14:55:14 -04:00
||| canPushCoe A s is true if a coercion along (𝑖 ⇒ A) on s can be pushed.
||| for a type with η like functions, or a ground type like ,
||| this is true for any s.
||| otherwise, like for pairs, it is only true if s is a constructor form.
||| if A isn't a type, or isn't in whnf, then the question is meaningless. but
||| it returns False anyway.
public export %inline
canPushCoe : Term (S d) n -> Term d n -> Bool
canPushCoe (TYPE {}) _ = True
canPushCoe (Pi {}) _ = True
canPushCoe (Lam {}) _ = False
canPushCoe (Sig {}) (Pair {}) = True
canPushCoe (Sig {}) _ = False
canPushCoe (Pair {}) _ = False
canPushCoe (W {}) (Sup {}) = True
canPushCoe (W {}) _ = False
canPushCoe (Sup {}) _ = False
canPushCoe (Enum {}) _ = True
canPushCoe (Tag {}) _ = False
canPushCoe (Eq {}) _ = True
canPushCoe (DLam {}) _ = False
canPushCoe (Nat {}) _ = True
canPushCoe (Zero {}) _ = False
canPushCoe (Succ {}) _ = False
canPushCoe (BOX {}) _ = True
canPushCoe (Box {}) _ = False
canPushCoe (E {}) _ = False
canPushCoe (CloT {}) _ = False
canPushCoe (DCloT {}) _ = False
2023-04-03 11:46:23 -04:00
||| true if a term is syntactically a type, or a neutral.
public export %inline
2023-04-15 09:13:01 -04:00
isTyConE : Term {} -> Bool
2023-04-03 11:46:23 -04:00
isTyConE s = isTyCon s || isE s
||| true if a term is syntactically a type.
public export %inline
2023-04-15 09:13:01 -04:00
isAnnTyCon : Elim {} -> Bool
2023-05-01 21:06:25 -04:00
isAnnTyCon (Ann ty (TYPE {}) _) = isTyCon ty
isAnnTyCon _ = False
2023-04-03 11:46:23 -04:00
2023-04-15 09:13:01 -04:00
public export %inline
isK : Dim d -> Bool
2023-05-01 21:06:25 -04:00
isK (K {}) = True
isK _ = False
2023-04-15 09:13:01 -04:00
mutual
public export
isRedexE : RedexTest Elim
2023-05-01 21:06:25 -04:00
isRedexE defs (F {x, _}) {d, n} =
isJust $ lookupElim x defs {d, n}
2023-05-01 21:06:25 -04:00
isRedexE _ (B {}) = False
isRedexE defs (App {fun, _}) =
isRedexE defs fun || isLamHead fun
isRedexE defs (CasePair {pair, _}) =
isRedexE defs pair || isPairHead pair
2023-08-06 04:46:55 -04:00
isRedexE defs (CaseW {tree, _}) =
isRedexE defs tree || isWHead tree
isRedexE defs (CaseEnum {tag, _}) =
isRedexE defs tag || isTagHead tag
isRedexE defs (CaseNat {nat, _}) =
isRedexE defs nat || isNatHead nat
isRedexE defs (CaseBox {box, _}) =
isRedexE defs box || isBoxHead box
2023-05-01 21:06:25 -04:00
isRedexE defs (DApp {fun, arg, _}) =
isRedexE defs fun || isDLamHead fun || isK arg
isRedexE defs (Ann {tm, ty, _}) =
isE tm || isRedexT defs tm || isRedexT defs ty
2023-08-13 14:55:14 -04:00
isRedexE defs (Coe {ty, val, _}) =
let ty = assert_smaller ty ty.term in
isRedexT defs ty || canPushCoe ty val
isRedexE defs (Comp {ty, r, _}) =
isRedexT defs ty || isK r
isRedexE defs (TypeCase {ty, ret, _}) =
isRedexE defs ty || isRedexT defs ret || isAnnTyCon ty
isRedexE _ (CloE {}) = True
isRedexE _ (DCloE {}) = True
public export
isRedexT : RedexTest Term
isRedexT _ (CloT {}) = True
isRedexT _ (DCloT {}) = True
2023-05-01 21:06:25 -04:00
isRedexT defs (E {e, _}) = isAnn e || isRedexE defs e
isRedexT _ _ = False
2023-04-15 09:13:01 -04:00
public export
tycaseRhs : (k : TyConKind) -> TypeCaseArms d n ->
Maybe (ScopeTermN (arity k) d n)
tycaseRhs k arms = lookupPrecise k arms
public export
tycaseRhsDef : Term d n -> (k : TyConKind) -> TypeCaseArms d n ->
ScopeTermN (arity k) d n
tycaseRhsDef def k arms = fromMaybe (SN def) $ tycaseRhs k arms
public export
tycaseRhs0 : (k : TyConKind) -> TypeCaseArms d n ->
(0 eq : arity k = 0) => Maybe (Term d n)
tycaseRhs0 k arms {eq} with (tycaseRhs k arms) | (arity k)
tycaseRhs0 k arms {eq = Refl} | res | 0 = map (.term) res
public export
tycaseRhsDef0 : Term d n -> (k : TyConKind) -> TypeCaseArms d n ->
(0 eq : arity k = 0) => Term d n
tycaseRhsDef0 def k arms = fromMaybe def $ tycaseRhs0 k arms
2023-08-06 04:46:55 -04:00
export
2023-08-06 05:07:17 -04:00
weakAtT : (CanTSubst f, Located2 f) => (by, at : Nat) ->
2023-08-06 04:46:55 -04:00
f d (at + n) -> f d (at + (by + n))
2023-08-06 05:07:17 -04:00
weakAtT by at t = t `CanTSubst.(//)` pushN at (shift by) t.loc
2023-04-15 09:13:01 -04:00
2023-08-06 04:46:55 -04:00
export
2023-08-06 05:07:17 -04:00
weakAtD : (CanDSubst f, Located2 f) => (by, at : Nat) ->
2023-08-06 04:46:55 -04:00
f (at + d) n -> f (at + (by + d)) n
2023-08-06 05:07:17 -04:00
weakAtD by at t = t `CanDSubst.(//)` pushN at (shift by) t.loc
2023-08-06 04:46:55 -04:00
parameters {s : Nat}
export
weakS : (by : Nat) -> ScopeTermN s d n -> ScopeTermN s d (by + n)
weakS by (S names (Y body)) = S names $ Y $ weakAtT by s body
weakS by (S names (N body)) = S names $ N $ weakT by body
export
weakDS : (by : Nat) -> DScopeTermN s d n -> DScopeTermN s d (by + n)
weakDS by (S names (Y body)) = S names $ Y $ weakT by body
weakDS by (S names (N body)) = S names $ N $ weakT by body
export
dweakS : (by : Nat) -> ScopeTermN s d n -> ScopeTermN s (by + d) n
dweakS by (S names (Y body)) = S names $ Y $ dweakT by body
dweakS by (S names (N body)) = S names $ N $ dweakT by body
export
dweakDS : (by : Nat) -> DScopeTermN s d n -> DScopeTermN s (by + d) n
dweakDS by (S names (Y body)) = S names $ Y $ weakAtD by s body
dweakDS by (S names (N body)) = S names $ N $ dweakT by body
2023-04-15 09:13:01 -04:00
private
2023-05-01 21:06:25 -04:00
coeScoped : {s : Nat} -> DScopeTerm d n -> Dim d -> Dim d -> Loc ->
2023-04-15 09:13:01 -04:00
ScopeTermN s d n -> ScopeTermN s d n
2023-05-01 21:06:25 -04:00
coeScoped ty p q loc (S names (Y body)) =
ST names $ E $ Coe (weakDS s ty) p q body loc
2023-05-01 21:06:25 -04:00
coeScoped ty p q loc (S names (N body)) =
S names $ N $ E $ Coe ty p q body loc
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
export covering
2023-05-21 14:11:01 -04:00
CanWhnf Term Reduce.isRedexT
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
export covering
2023-05-21 14:11:01 -04:00
CanWhnf Elim Reduce.isRedexE
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
parameters {d, n : Nat} (defs : Definitions) (ctx : WhnfContext d n)
||| performs the minimum work required to recompute the type of an elim.
|||
||| ⚠ **assumes the elim is already typechecked.** ⚠
export covering
2023-05-01 21:06:25 -04:00
computeElimType : (e : Elim d n) -> (0 ne : No (isRedexE defs e)) =>
2023-05-21 14:11:01 -04:00
Eff Whnf (Term d n)
2023-05-21 14:09:34 -04:00
computeElimType (F {x, u, loc}) = do
2023-05-01 21:06:25 -04:00
let Just def = lookup x defs | Nothing => throw $ NotInScope loc x
2023-05-21 14:09:34 -04:00
pure $ displace u def.type
2023-05-01 21:06:25 -04:00
computeElimType (B {i, _}) = pure $ ctx.tctx !! i
computeElimType (App {fun = f, arg = s, loc}) {ne} = do
Pi {arg, res, _} <- whnf0 defs ctx =<< computeElimType f {ne = noOr1 ne}
| t => throw $ ExpectedPi loc ctx.names t
pure $ sub1 res $ Ann s arg loc
computeElimType (CasePair {pair, ret, _}) = pure $ sub1 ret pair
2023-08-06 04:46:55 -04:00
computeElimType (CaseW {tree, ret, _}) = pure $ sub1 ret tree
2023-05-01 21:06:25 -04:00
computeElimType (CaseEnum {tag, ret, _}) = pure $ sub1 ret tag
computeElimType (CaseNat {nat, ret, _}) = pure $ sub1 ret nat
computeElimType (CaseBox {box, ret, _}) = pure $ sub1 ret box
computeElimType (DApp {fun = f, arg = p, loc}) {ne} = do
Eq {ty, _} <- whnf0 defs ctx =<< computeElimType f {ne = noOr1 ne}
| t => throw $ ExpectedEq loc ctx.names t
pure $ dsub1 ty p
2023-05-21 14:09:34 -04:00
computeElimType (Ann {ty, _}) = pure ty
2023-05-01 21:06:25 -04:00
computeElimType (Coe {ty, q, _}) = pure $ dsub1 ty q
computeElimType (Comp {ty, _}) = pure ty
computeElimType (TypeCase {ret, _}) = pure ret
parameters {d, n : Nat} (defs : Definitions) (ctx : WhnfContext (S d) n)
||| for π.(x : A) → B, returns (A, B);
||| for an elim returns a pair of type-cases that will reduce to that;
||| for other intro forms error
private covering
tycasePi : (t : Term (S d) n) -> (0 tnf : No (isRedexT defs t)) =>
2023-05-21 14:11:01 -04:00
Eff Whnf (Term (S d) n, ScopeTerm (S d) n)
2023-05-01 21:06:25 -04:00
tycasePi (Pi {arg, res, _}) = pure (arg, res)
tycasePi (E e) {tnf} = do
ty <- computeElimType defs ctx e @{noOr2 tnf}
let loc = e.loc
2023-08-06 04:46:55 -04:00
narg = mnb "Arg" loc; nret = mnb "Ret" loc
2023-05-01 21:06:25 -04:00
arg = E $ typeCase1Y e ty KPi [< !narg, !nret] (BVT 1 loc) loc
res' = typeCase1Y e (Arr Zero arg ty loc) KPi [< !narg, !nret]
(BVT 0 loc) loc
res = ST [< !narg] $ E $ App (weakE 1 res') (BVT 0 loc) loc
2023-05-01 21:06:25 -04:00
pure (arg, res)
tycasePi t = throw $ ExpectedPi t.loc ctx.names t
||| for (x : A) × B, returns (A, B);
||| for an elim returns a pair of type-cases that will reduce to that;
||| for other intro forms error
private covering
tycaseSig : (t : Term (S d) n) -> (0 tnf : No (isRedexT defs t)) =>
2023-05-21 14:11:01 -04:00
Eff Whnf (Term (S d) n, ScopeTerm (S d) n)
2023-05-01 21:06:25 -04:00
tycaseSig (Sig {fst, snd, _}) = pure (fst, snd)
tycaseSig (E e) {tnf} = do
ty <- computeElimType defs ctx e @{noOr2 tnf}
let loc = e.loc
2023-08-06 04:46:55 -04:00
nfst = mnb "Fst" loc; nsnd = mnb "Snd" loc
2023-05-01 21:06:25 -04:00
fst = E $ typeCase1Y e ty KSig [< !nfst, !nsnd] (BVT 1 loc) loc
snd' = typeCase1Y e (Arr Zero fst ty loc) KSig [< !nfst, !nsnd]
(BVT 0 loc) loc
snd = ST [< !nfst] $ E $ App (weakE 1 snd') (BVT 0 loc) loc
2023-05-01 21:06:25 -04:00
pure (fst, snd)
tycaseSig t = throw $ ExpectedSig t.loc ctx.names t
2023-08-06 04:46:55 -04:00
||| for (x : A) ⊲ B, returns (A, B);
||| for an elim returns a pair of type-cases that will reduce to that;
||| for other intro forms error
private covering
tycaseW : (t : Term (S d) n) -> (0 tnf : No (isRedexT defs t)) =>
Eff Whnf (Term (S d) n, ScopeTerm (S d) n)
tycaseW (W {shape, body, _}) = pure (shape, body)
tycaseW (E e) {tnf} = do
ty <- computeElimType defs ctx e @{noOr2 tnf}
let loc = e.loc
nshape = mnb "Shape" loc; nbody = mnb "Body" loc
shape = E $ typeCase1Y e ty KW [< !nshape, !nbody] (BVT 1 loc) loc
body' = typeCase1Y e (Arr Zero shape ty loc) KW [< !nshape, !nbody]
(BVT 0 loc) loc
body = ST [< !nshape] $ E $ App (weakE 1 body') (BVT 0 loc) loc
pure (shape, body)
tycaseW t = throw $ ExpectedW t.loc ctx.names t
2023-05-01 21:06:25 -04:00
||| for [π. A], returns A;
||| for an elim returns a type-case that will reduce to that;
||| for other intro forms error
private covering
tycaseBOX : (t : Term (S d) n) -> (0 tnf : No (isRedexT defs t)) =>
2023-05-21 14:11:01 -04:00
Eff Whnf (Term (S d) n)
2023-05-01 21:06:25 -04:00
tycaseBOX (BOX {ty, _}) = pure ty
tycaseBOX (E e) {tnf} = do
2023-08-06 04:46:55 -04:00
let loc = e.loc
2023-05-01 21:06:25 -04:00
ty <- computeElimType defs ctx e @{noOr2 tnf}
2023-08-06 04:46:55 -04:00
pure $ E $ typeCase1Y e ty KBOX [< !(mnb "Ty" loc)] (BVT 0 loc) loc
2023-05-01 21:06:25 -04:00
tycaseBOX t = throw $ ExpectedBOX t.loc ctx.names t
||| for Eq [i ⇒ A] l r, returns (A0/i, A1/i, A, l, r);
||| for an elim returns five type-cases that will reduce to that;
||| for other intro forms error
private covering
tycaseEq : (t : Term (S d) n) -> (0 tnf : No (isRedexT defs t)) =>
2023-05-21 14:11:01 -04:00
Eff Whnf (Term (S d) n, Term (S d) n, DScopeTerm (S d) n,
Term (S d) n, Term (S d) n)
2023-05-01 21:06:25 -04:00
tycaseEq (Eq {ty, l, r, _}) = pure (ty.zero, ty.one, ty, l, r)
tycaseEq (E e) {tnf} = do
ty <- computeElimType defs ctx e @{noOr2 tnf}
let loc = e.loc
2023-08-06 04:46:55 -04:00
names = traverse' (\x => mnb x loc) [< "A0", "A1", "A", "l", "r"]
2023-05-01 21:06:25 -04:00
a0 = E $ typeCase1Y e ty KEq !names (BVT 4 loc) loc
a1 = E $ typeCase1Y e ty KEq !names (BVT 3 loc) loc
a' = typeCase1Y e (Eq0 ty a0 a1 loc) KEq !names (BVT 2 loc) loc
2023-08-06 04:46:55 -04:00
a = DST [< !(mnb "i" loc)] $ E $ DApp (dweakE 1 a') (B VZ loc) loc
2023-05-01 21:06:25 -04:00
l = E $ typeCase1Y e a0 KEq !names (BVT 1 loc) loc
r = E $ typeCase1Y e a1 KEq !names (BVT 0 loc) loc
pure (a0, a1, a, l, r)
tycaseEq t = throw $ ExpectedEq t.loc ctx.names t
-- new block because the functions below might pass a different ctx
-- into the ones above
parameters {d, n : Nat} (defs : Definitions) (ctx : WhnfContext d n)
||| reduce a function application `App (Coe ty p q val) s loc`
private covering
piCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) ->
(val, s : Term d n) -> Loc ->
2023-05-21 14:11:01 -04:00
Eff Whnf (Subset (Elim d n) (No . isRedexE defs))
2023-05-01 21:06:25 -04:00
piCoe sty@(S [< i] ty) p q val s loc = do
2023-08-06 04:46:55 -04:00
-- 𝒮𝑘 ≔ π.(x : A𝑘/𝑖) → B𝑘/𝑖
-- 𝓈𝑘 ≔ coe [𝑖 ⇒ A] @q @𝑘 s
-- -------------------------------------------------------
-- (coe [𝑖𝒮𝑖] @p @q t) s
-- ⇝
-- coe [i ⇒ B[𝓈i/x] @p @q ((t ∷ 𝒮p) 𝓈p)
2023-03-31 13:11:35 -04:00
--
2023-05-01 21:06:25 -04:00
-- type-case is used to expose A,B if the type is neutral
let ctx1 = extendDim i ctx
Element ty tynf <- whnf defs ctx1 ty.term
(arg, res) <- tycasePi defs ctx1 ty
let s0 = CoeT i arg q p s s.loc
body = E $ App (Ann val (ty // one p) val.loc) (E s0) loc
s1 = CoeT i (arg // (BV 0 i.loc ::: shift 2)) (weakD 1 q) (BV 0 i.loc)
(s // shift 1) s.loc
whnf defs ctx $ CoeT i (sub1 res s1) p q body loc
||| reduce a pair elimination `CasePair pi (Coe ty p q val) ret body loc`
private covering
sigCoe : (qty : Qty) ->
(ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
(ret : ScopeTerm d n) -> (body : ScopeTermN 2 d n) -> Loc ->
2023-05-21 14:11:01 -04:00
Eff Whnf (Subset (Elim d n) (No . isRedexE defs))
2023-05-01 21:06:25 -04:00
sigCoe qty sty@(S [< i] ty) p q val ret body loc = do
2023-08-06 04:46:55 -04:00
-- caseπ (coe [i ⇒ (x : A) × B] @p @q s) return z ⇒ C of { (a, b) ⇒ u }
2023-05-01 21:06:25 -04:00
-- ⇝
-- caseπ s ∷ ((x : A) × B)p/i return z ⇒ C
-- of { (a, b) ⇒
2023-08-06 04:46:55 -04:00
-- u[(coe [i ⇒ A] @p @q a)/a,
2023-05-01 21:06:25 -04:00
-- (coe [i ⇒ B[(coe [j ⇒ Aj/i] @p @i a)/x]] @p @q b)/b] }
--
-- type-case is used to expose A,B if the type is neutral
let ctx1 = extendDim i ctx
Element ty tynf <- whnf defs ctx1 ty.term
(tfst, tsnd) <- tycaseSig defs ctx1 ty
let [< x, y] = body.names
2023-08-06 04:46:55 -04:00
a' = CoeT i (weakT 2 tfst) p q (BVT 1 x.loc) x.loc
2023-05-01 21:06:25 -04:00
tsnd' = tsnd.term //
2023-08-06 04:46:55 -04:00
(CoeT !(fresh i) (weakT 2 $ tfst // (B VZ i.loc ::: shift 2))
(weakD 1 p) (B VZ i.loc) (BVT 1 x.loc) y.loc ::: shift 2)
b' = CoeT i tsnd' p q (BVT 0 y.loc) y.loc
whnf defs ctx $ CasePair qty (Ann val (dsub1 sty p) val.loc) ret
2023-05-01 21:06:25 -04:00
(ST body.names $ body.term // (a' ::: b' ::: shift 2)) loc
2023-08-06 04:46:55 -04:00
||| reduce a w elimination `CaseW pi piIH (Coe ty p q val) ret body loc`
private covering
wCoe : (qty, qtyIH : Qty) ->
(ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
(ret : ScopeTerm d n) -> (body : ScopeTermN 3 d n) -> Loc ->
Eff Whnf (Subset (Elim d n) (No . isRedexE defs))
wCoe qty qtyIH sty@(S [< i] ty) p q val ret body loc = do
-- 𝒮𝑘 ≔ ((x : A) ⊲ B)𝑘/𝑖
-- 𝒶𝑘 ≔ coe [𝑖 ⇒ A] @p @𝑘 a
-- : A𝑘/𝑖
-- 𝒷𝑘 ≔ coe [𝑖 ⇒ ω.B[𝒶𝑖/x] → 𝒮𝑖] @p @𝑘 b
-- : ω.B𝑘/𝑖[𝒶𝑘/x] → 𝒮𝑘
2023-08-06 04:46:55 -04:00
-- 𝒾𝒽 ≔ coe [𝑖 ⇒ π.(z : B[𝒶𝑖/x]) → C[𝒷𝑖 z/p]] @p @q ih
-- : π.(z : Bq/𝑖[𝒶q/x]) → C[𝒷q z/p]
-- --------------------------------------------------------------------
-- caseπ (coe [𝑖𝒮𝑖] @p @q s) return z ⇒ C of { a ⋄ b, ς.ih ⇒ u }
-- ⇝
-- caseπ s ∷ 𝒮p return z ⇒ C
-- of { a ⋄ b, ς.ih ⇒ u[𝒶q/a, 𝒷q/b, 𝒾𝒽/ih] }
let ctx1 = extendDim i ctx
Element ty tynf <- whnf defs ctx1 ty.term
(shape, tbody) <- tycaseW defs ctx1 ty
let [< a, b, ih] = body.names
ai <- fresh i; bi <- fresh i; ihi <- fresh i; z <- mnb "z" ih.loc
let a', b' : forall d'. (by : Shift d d') => Dim d' -> Elim d' (3 + n)
a' k =
let shape' = shape // Shift (weak 1 by) // shift 3 in
CoeT ai shape' (p // by) k (BVT 2 a.loc) a.loc
b' k =
let tbody' = tbody.term // Shift (weak 1 by)
// (a' (BV 0 bi.loc) ::: shift 3)
ty' = ty // Shift (weak 1 by) // shift 3
in
CoeT bi (Arr Any tbody' ty' b.loc) (p // by) k (BVT 1 b.loc) b.loc
2023-08-06 04:46:55 -04:00
ih' : Elim d (3 + n) :=
let tbody' = tbody.term // (a' (BV 0 ihi.loc) ::: shift 3)
ret' = sub1 (weakS 4 $ dweakS 1 ret) $
App (weakE 1 $ b' (BV 0 ihi.loc)) (BVT 0 z.loc) ih.loc
ty = PiY qty z tbody' ret' ih.loc
in
CoeT ihi ty p q (BVT 0 ih.loc) ih.loc
whnf defs ctx $ CaseW qty qtyIH (Ann val (dsub1 sty p) val.loc) ret
(ST body.names $ body.term // (a' q ::: b' q ::: ih' ::: shift 3)) loc
2023-05-01 21:06:25 -04:00
||| reduce a dimension application `DApp (Coe ty p q val) r loc`
2023-04-15 09:13:01 -04:00
private covering
2023-05-01 21:06:25 -04:00
eqCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
(r : Dim d) -> Loc ->
2023-05-21 14:11:01 -04:00
Eff Whnf (Subset (Elim d n) (No . isRedexE defs))
2023-05-01 21:06:25 -04:00
eqCoe sty@(S [< j] ty) p q val r loc = do
-- (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)
-- @r { 0 j ⇒ L; 1 j ⇒ R }
let ctx1 = extendDim j ctx
Element ty tynf <- whnf defs ctx1 ty.term
(a0, a1, a, s, t) <- tycaseEq defs ctx1 ty
let a' = dsub1 a (weakD 1 r)
2023-08-06 04:46:55 -04:00
val' = E $ DApp (Ann val (dsub1 sty p) val.loc) r loc
2023-05-01 21:06:25 -04:00
whnf defs ctx $ CompH j a' p q val' r j s j t loc
||| reduce a pair elimination `CaseBox pi (Coe ty p q val) ret body`
2023-04-15 09:13:01 -04:00
private covering
2023-05-01 21:06:25 -04:00
boxCoe : (qty : Qty) ->
(ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
(ret : ScopeTerm d n) -> (body : ScopeTerm d n) -> Loc ->
2023-05-21 14:11:01 -04:00
Eff Whnf (Subset (Elim d n) (No . isRedexE defs))
2023-05-01 21:06:25 -04:00
boxCoe qty sty@(S [< i] ty) p q val ret body loc = do
-- caseπ (coe [i ⇒ [ρ. A]] @p @q s) return z ⇒ C of { [a] ⇒ e }
-- ⇝
-- caseπ s ∷ [ρ. A]p/i return z ⇒ C
-- of { [a] ⇒ e[(coe [i ⇒ A] p q a)/a] }
let ctx1 = extendDim i ctx
Element ty tynf <- whnf defs ctx1 ty.term
ta <- tycaseBOX defs ctx1 ty
2023-08-06 04:46:55 -04:00
let a' = CoeT i (weakT 1 ta) p q (BVT 0 body.name.loc) body.name.loc
whnf defs ctx $ CaseBox qty (Ann val (dsub1 sty p) val.loc) ret
2023-05-01 21:06:25 -04:00
(ST body.names $ body.term // (a' ::: shift 1)) loc
||| reduce a type-case applied to a type constructor
private covering
reduceTypeCase : {d, n : Nat} -> (defs : Definitions) -> WhnfContext d n ->
(ty : Term d n) -> (u : Universe) -> (ret : Term d n) ->
(arms : TypeCaseArms d n) -> (def : Term d n) ->
(0 _ : So (isTyCon ty)) => Loc ->
2023-05-21 14:11:01 -04:00
Eff Whnf (Subset (Elim d n) (No . isRedexE defs))
2023-05-01 21:06:25 -04:00
reduceTypeCase defs ctx ty u ret arms def loc = case ty of
-- (type-case ★ᵢ ∷ _ return Q of { ★ ⇒ s; ⋯ }) ⇝ s ∷ Q
TYPE {} =>
whnf defs ctx $ Ann (tycaseRhsDef0 def KTYPE arms) ret loc
-- (type-case π.(x : A) → B ∷ ★ᵢ return Q of { (a → b) ⇒ s; ⋯ }) ⇝
-- s[(A ∷ ★ᵢ)/a, ((λ x ⇒ B) ∷ 0.A → ★ᵢ)/b] ∷ Q
Pi {arg, res, loc = piLoc, _} =>
2023-08-06 04:46:55 -04:00
let arg' = Ann arg (TYPE u arg.loc) arg.loc
2023-05-01 21:06:25 -04:00
res' = Ann (Lam res res.loc)
2023-08-06 04:46:55 -04:00
(Arr Zero arg (TYPE u arg.loc) arg.loc) res.loc
2023-05-01 21:06:25 -04:00
in
whnf defs ctx $
Ann (subN (tycaseRhsDef def KPi arms) [< arg', res']) ret loc
2023-08-06 04:46:55 -04:00
-- (type-case (x : A) ⊲ π.B ∷ ★ᵢ return Q of { (a ⊲ b) ⇒ s; ⋯ }) ⇝
-- s[(A ∷ ★ᵢ)/a, ((λ x ⇒ B) ∷ 0.A → ★ᵢ)/b] ∷ Q
W {shape, body, loc = wLoc, _} =>
let shape' = Ann shape (TYPE u shape.loc) shape.loc
body' = Ann (Lam body body.loc)
(Arr Zero shape (TYPE u shape.loc) shape.loc) body.loc
in
whnf defs ctx $
Ann (subN (tycaseRhsDef def KW arms) [< shape', body']) ret loc
2023-05-01 21:06:25 -04:00
-- (type-case (x : A) × B ∷ ★ᵢ return Q of { (a × b) ⇒ s; ⋯ }) ⇝
-- s[(A ∷ ★ᵢ)/a, ((λ x ⇒ B) ∷ 0.A → ★ᵢ)/b] ∷ Q
Sig {fst, snd, loc = sigLoc, _} =>
2023-08-06 04:46:55 -04:00
let fst' = Ann fst (TYPE u fst.loc) fst.loc
2023-05-01 21:06:25 -04:00
snd' = Ann (Lam snd snd.loc)
2023-08-06 04:46:55 -04:00
(Arr Zero fst (TYPE u fst.loc) fst.loc) snd.loc
2023-05-01 21:06:25 -04:00
in
whnf defs ctx $
Ann (subN (tycaseRhsDef def KSig arms) [< fst', snd']) ret loc
-- (type-case {⋯} ∷ _ return Q of { {} ⇒ s; ⋯ }) ⇝ s ∷ Q
Enum {} =>
whnf defs ctx $ Ann (tycaseRhsDef0 def KEnum arms) ret loc
-- (type-case Eq [i ⇒ A] L R ∷ ★ᵢ return Q
-- of { Eq a₀ a₁ a l r ⇒ s; ⋯ }) ⇝
-- s[(A0/i ∷ ★ᵢ)/a₀, (A1/i ∷ ★ᵢ)/a₁,
-- ((δ i ⇒ A) ∷ Eq [★ᵢ] A0/i A1/i)/a,
-- (L ∷ A0/i)/l, (R ∷ A1/i)/r] ∷ Q
Eq {ty = a, l, r, loc = eqLoc, _} =>
let a0 = a.zero; a1 = a.one in
whnf defs ctx $ Ann
(subN (tycaseRhsDef def KEq arms)
2023-08-06 04:46:55 -04:00
[< Ann a0 (TYPE u a.loc) a.loc, Ann a1 (TYPE u a.loc) a.loc,
Ann (DLam a a.loc) (Eq0 (TYPE u a.loc) a0 a1 a.loc) a.loc,
2023-05-01 21:06:25 -04:00
Ann l a0 l.loc, Ann r a1 r.loc])
ret loc
-- (type-case ∷ _ return Q of { ⇒ s; ⋯ }) ⇝ s ∷ Q
Nat {} =>
whnf defs ctx $ Ann (tycaseRhsDef0 def KNat arms) ret loc
-- (type-case [π.A] ∷ ★ᵢ return Q of { [a] ⇒ s; ⋯ }) ⇝ s[(A ∷ ★ᵢ)/a] ∷ Q
BOX {ty = a, loc = boxLoc, _} =>
whnf defs ctx $ Ann
2023-08-06 04:46:55 -04:00
(sub1 (tycaseRhsDef def KBOX arms) (Ann a (TYPE u a.loc) a.loc))
2023-05-01 21:06:25 -04:00
ret loc
||| pushes a coercion inside a whnf-ed term
private covering
2023-05-21 14:09:34 -04:00
pushCoe : {d, n : Nat} -> (defs : Definitions) -> WhnfContext d n ->
2023-08-13 14:55:14 -04:00
BindName -> (ty : Term (S d) n) ->
2023-05-01 21:06:25 -04:00
Dim d -> Dim d ->
2023-08-13 14:55:14 -04:00
(s : Term d n) -> (0 snf : No (isRedexT defs s)) =>
(0 pc : So (canPushCoe ty s)) => Loc ->
2023-05-21 14:11:01 -04:00
Eff Whnf (NonRedex Elim d n defs)
2023-05-01 21:06:25 -04:00
pushCoe defs ctx i ty p q s loc =
if p == q then whnf defs ctx $ Ann s (ty // one q) loc else
2023-08-13 14:55:14 -04:00
case ty of
-- (coe ★ᵢ @p @q s) ⇝ (s ∷ ★ᵢ)
--
-- no η (what would that even mean), but ground type
TYPE {l, loc = tyLoc} =>
whnf defs ctx $ Ann s (TYPE l tyLoc) loc
2023-05-01 21:06:25 -04:00
-- just η expand it. then whnf for App will handle it later
-- this is how @xtt does it
--
2023-08-13 14:55:14 -04:00
-- (coe (𝑖 ⇒ π.(x : A) → B) @p @q s) ⇝
-- (λ y ⇒ (coe (𝑖 ⇒ π.(x : A) → B) @p @q s) y)
-- ∷ (π.(x : A) → B)q/𝑖
Pi {} => do
y <- mnb "y" loc
let s' = Coe (SY [< i] ty) p q s loc
body = SY [< y] $ E $ App (weakE 1 s') (BVT 0 y.loc) s.loc
ret = ty // one q
whnf defs ctx $ Ann (Lam body loc) ret loc
2023-05-01 21:06:25 -04:00
-- (coe [i ⇒ (x : A) × B] @p @q (s, t)) ⇝
-- (coe [i ⇒ A] @p @q s,
-- coe [i ⇒ B[(coe [j ⇒ Aj/i] @p @i s)/x]] @p @q t)
-- ∷ (x : Aq/i) × Bq/i
--
2023-08-13 14:55:14 -04:00
-- no η so only reduce on an actual pair 🍐
Sig {fst = tfst, snd = tsnd, loc = tyLoc} => do
let Pair fst snd sLoc = s
fst' = E $ CoeT i tfst p q fst fst.loc
2023-08-06 04:46:55 -04:00
tfst' = tfst // (B VZ i.loc ::: shift 2)
2023-05-01 21:06:25 -04:00
tsnd' = sub1 tsnd $
2023-08-06 04:46:55 -04:00
CoeT !(fresh i) tfst' (weakD 1 p) (B VZ snd.loc)
(dweakT 1 fst) snd.loc
2023-05-01 21:06:25 -04:00
snd' = E $ CoeT i tsnd' p q snd snd.loc
2023-08-13 14:55:14 -04:00
pure $ nred $
Ann (Pair fst' snd' sLoc)
(Sig (tfst // one q) (tsnd // one q) tyLoc) loc
2023-05-01 21:06:25 -04:00
2023-08-06 04:46:55 -04:00
-- (coe [i ⇒ (x : A) ⊲ π.B] @p @q (s ⋄ t) ⇝
-- (coe [i ⇒ A] @p @q s ⋄
-- coe [i ⇒ ω.B[coe [j ⇒ Aj/i] @p @i s/x] → (x : A) ⊲ B] t)
2023-08-06 04:46:55 -04:00
-- ∷ ((x : A) ⊲ B)q/i
2023-08-13 14:55:14 -04:00
--
-- again, no η
W {shape, body, loc = tyLoc} => do
let Sup root sub sLoc = s
root' = E $ CoeT i shape p q root root.loc
shape' = shape // (B VZ i.loc ::: shift 2)
coeRoot =
CoeT (setLoc shape.loc !(fresh i)) shape'
(weakD 1 p) (B VZ i.loc) (dweakT 1 root) root.loc
tsub' = Arr Any (sub1 body coeRoot) ty sub.loc
sub' = E $ CoeT i tsub' p q sub sub.loc
pure $ nred $
Ann (Sup root' sub' sLoc)
(W (shape // one q) (body // one q) tyLoc) loc
-- (coe {𝗮, …} @p @q s) ⇝ (s ∷ {𝗮, …})
--
-- no η, but ground type
Enum {cases, loc = tyLoc} =>
whnf defs ctx $ Ann s (Enum cases tyLoc) loc
2023-08-06 04:46:55 -04:00
2023-05-01 21:06:25 -04:00
-- η expand, like for Lam
--
2023-08-13 14:55:14 -04:00
-- (coe (𝑖 ⇒ Eq (𝑗 ⇒ A) l r) @p @q s) ⇝
-- (δ k ⇒ (coe (𝑖 ⇒ Eq (𝑗 ⇒ A) l r) @p @q s) @k)
-- ∷ (Eq (𝑗 ⇒ A) l r)q/𝑖
Eq {} => do
k <- mnb "k" loc
let s' = Coe (SY [< i] ty) p q s loc
term = DLam (SY [< k] $ E $ DApp (dweakE 1 s') (BV 0 k.loc) loc) loc
ret = ty // one q
whnf defs ctx $ Ann term ret loc
-- (coe @p @q s) ⇝ (s ∷ )
--
-- no η, but ground type
Nat {loc = tyLoc} =>
whnf defs ctx $ Ann s (Nat tyLoc) loc
-- (coe (𝑖 ⇒ [π. A]) @p @q s) ⇝
-- [coe (𝑖 ⇒ A) @p @q (case1 s ∷ [π. Ap/𝑖] return Ap/𝑖 of {[x] ⇒ x})]
-- ∷ [π. Aq/𝑖]
--
-- [todo] box should probably have an η rule
BOX {qty, ty = innerTy, loc = tyLoc} => do
let s' = Ann s (BOX qty (innerTy // one p) tyLoc) s.loc
inner' = CaseBox One s' (SN $ innerTy // one p)
(SY [< !(mnb "x" s.loc)] $ BVT 0 s.loc) s.loc
inner = Box (E $ CoeT i innerTy p q (E inner') loc) loc
ret = BOX qty (innerTy // one q) tyLoc
whnf defs ctx $ Ann inner ret loc
2023-05-01 21:06:25 -04:00
export covering
2023-05-21 14:11:01 -04:00
CanWhnf Elim Reduce.isRedexE where
2023-05-21 14:09:34 -04:00
whnf defs ctx (F x u loc) with (lookupElim x defs) proof eq
_ | Just y = whnf defs ctx $ setLoc loc $ displace u y
_ | Nothing = pure $ Element (F x u loc) $ rewrite eq in Ah
2023-05-01 21:06:25 -04:00
whnf _ _ (B i loc) = pure $ nred $ B i loc
-- ((λ x ⇒ t) ∷ (π.x : A) → B) s ⇝ t[s∷A/x] ∷ B[s∷A/x]
whnf defs ctx (App f s appLoc) = do
Element f fnf <- whnf defs ctx f
case nchoose $ isLamHead f of
Left _ => case f of
Ann (Lam {body, _}) (Pi {arg, res, _}) floc =>
let s = Ann s arg s.loc in
whnf defs ctx $ Ann (sub1 body s) (sub1 res s) appLoc
Coe ty p q val _ => piCoe defs ctx ty p q val s appLoc
Right nlh => pure $ Element (App f s appLoc) $ fnf `orNo` nlh
2023-08-06 04:46:55 -04:00
-- s' ≔ s ∷ A
-- t' ≔ t ∷ B[s'/x]
-- st' ≔ (s, t) ∷ (x : A) × B
-- C' ≔ C[st'/p]
-- ---------------------------------------------------------------
-- case st' return p ⇒ C of { (a, b) ⇒ u } ⇝ u[s'/a, t'/x]] ∷ C'
2023-05-01 21:06:25 -04:00
whnf defs ctx (CasePair pi pair ret body caseLoc) = do
Element pair pairnf <- whnf defs ctx pair
case nchoose $ isPairHead pair of
Left _ => case pair of
Ann (Pair {fst, snd, _}) (Sig {fst = tfst, snd = tsnd, _}) pairLoc =>
let fst = Ann fst tfst fst.loc
snd = Ann snd (sub1 tsnd fst) snd.loc
in
whnf defs ctx $ Ann (subN body [< fst, snd]) (sub1 ret pair) caseLoc
Coe ty p q val _ => do
sigCoe defs ctx pi ty p q val ret body caseLoc
Right np =>
pure $ Element (CasePair pi pair ret body caseLoc) $ pairnf `orNo` np
2023-08-06 04:46:55 -04:00
-- s' ≔ s ∷ A
-- t' ≔ t ∷ ω.B[s'/x] → (x : A) ⊲ B
2023-08-06 04:46:55 -04:00
-- ih' ≔ (λ x ⇒ caseπ t x return p ⇒ C of { (a ⋄ b), ς.ih ⇒ u }) ∷
-- π.(y : B[s'/x]) → C[t' y/p]
-- st' ≔ s ⋄ t ∷ (x : A) ⊲ B
-- C' ≔ C[st'/p]
-- --------------------------------------------------------------------
-- caseπ st' return p ⇒ C of { a ⋄ b, ς.ih ⇒ u }
-- ⇝
-- u[s'/a, t'/b, ih'/ih] ∷ C'
whnf defs ctx (CaseW qty qtyIH tree ret body caseLoc) = do
Element tree treenf <- whnf defs ctx tree
case nchoose $ isWHead tree of
Left _ => case tree of
Ann (Sup {root, sub, _})
w@(W {shape, body = wbody, _}) treeLoc =>
let root = Ann root shape root.loc
wbody' = sub1 wbody root
tsub = Arr Any wbody' w sub.loc
2023-08-06 04:46:55 -04:00
sub = Ann sub tsub sub.loc
ih' = LamY !(mnb "y" caseLoc) -- [todo] better name
(E $ CaseW qty qtyIH
(App (weakE 1 sub) (BVT 0 sub.loc) sub.loc)
(weakS 1 ret) (weakS 1 body) caseLoc) sub.loc
ihret = sub1 (weakS 1 ret)
(App (weakE 1 sub) (BVT 0 sub.loc) caseLoc)
tih = PiY qty !(mnb "y" caseLoc)
wbody' ihret caseLoc
ih = Ann ih' tih caseLoc
in
whnf defs ctx $
Ann (subN body [< root, sub, ih]) (sub1 ret tree) tree.loc
Coe ty p q val _ => do
wCoe defs ctx qty qtyIH ty p q val ret body caseLoc
Right nw => pure $
Element (CaseW qty qtyIH tree ret body caseLoc) $ treenf `orNo` nw
2023-05-01 21:06:25 -04:00
-- case 'a ∷ {a,…} return p ⇒ C of { 'a ⇒ u } ⇝
-- u ∷ C['a∷{a,…}/p]
whnf defs ctx (CaseEnum pi tag ret arms caseLoc) = do
Element tag tagnf <- whnf defs ctx tag
case nchoose $ isTagHead tag of
Left _ => case tag of
Ann (Tag t _) (Enum ts _) _ =>
let ty = sub1 ret tag in
case lookup t arms of
Just arm => whnf defs ctx $ Ann arm ty arm.loc
Nothing => throw $ MissingEnumArm caseLoc t (keys arms)
Coe ty p q val _ =>
-- there is nowhere an equality can be hiding inside an enum type
whnf defs ctx $
CaseEnum pi (Ann val (dsub1 ty q) val.loc) ret arms caseLoc
Right nt =>
pure $ Element (CaseEnum pi tag ret arms caseLoc) $ tagnf `orNo` nt
-- case zero ∷ return p ⇒ C of { zero ⇒ u; … } ⇝
-- u ∷ C[zero∷/p]
--
-- case succ n ∷ return p ⇒ C of { succ n', π.ih ⇒ u; … } ⇝
-- u[n∷/n', (case n ∷ ⋯)/ih] ∷ C[succ n ∷ /p]
whnf defs ctx (CaseNat pi piIH nat ret zer suc caseLoc) = do
Element nat natnf <- whnf defs ctx nat
case nchoose $ isNatHead nat of
Left _ =>
let ty = sub1 ret nat in
case nat of
Ann (Zero _) (Nat _) _ =>
whnf defs ctx $ Ann zer ty zer.loc
Ann (Succ n succLoc) (Nat natLoc) _ =>
let nn = Ann n (Nat natLoc) succLoc
tm = subN suc [< nn, CaseNat pi piIH nn ret zer suc caseLoc]
in
whnf defs ctx $ Ann tm ty caseLoc
Coe ty p q val _ =>
-- same deal as Enum
whnf defs ctx $
CaseNat pi piIH (Ann val (dsub1 ty q) val.loc) ret zer suc caseLoc
Right nn => pure $
Element (CaseNat pi piIH nat ret zer suc caseLoc) $ natnf `orNo` nn
-- case [t] ∷ [π.A] return p ⇒ C of { [x] ⇒ u } ⇝
-- u[t∷A/x] ∷ C[[t] ∷ [π.A]/p]
whnf defs ctx (CaseBox pi box ret body caseLoc) = do
Element box boxnf <- whnf defs ctx box
case nchoose $ isBoxHead box of
Left _ => case box of
Ann (Box val boxLoc) (BOX q bty tyLoc) _ =>
let ty = sub1 ret box in
whnf defs ctx $ Ann (sub1 body (Ann val bty val.loc)) ty caseLoc
Coe ty p q val _ =>
boxCoe defs ctx pi ty p q val ret body caseLoc
Right nb =>
pure $ Element (CaseBox pi box ret body caseLoc) $ boxnf `orNo` nb
2023-05-21 14:09:34 -04:00
-- e : Eq (𝑗 ⇒ A) t u ⊢ e @0 ⇝ t ∷ A0/𝑗
-- e : Eq (𝑗 ⇒ A) t u ⊢ e @1 ⇝ u ∷ A1/𝑗
--
-- ((δ 𝑖 ⇒ s) ∷ Eq (𝑗 ⇒ A) t u) @𝑘 ⇝ s𝑘/𝑖 ∷ A𝑘/𝑗
2023-05-01 21:06:25 -04:00
whnf defs ctx (DApp f p appLoc) = do
Element f fnf <- whnf defs ctx f
case nchoose $ isDLamHead f of
Left _ => case f of
Ann (DLam {body, _}) (Eq {ty, l, r, _}) _ =>
whnf defs ctx $
Ann (endsOr (setLoc appLoc l) (setLoc appLoc r) (dsub1 body p) p)
(dsub1 ty p) appLoc
Coe ty p' q' val _ =>
eqCoe defs ctx ty p' q' val p appLoc
Right ndlh => case p of
K e _ => do
Eq {l, r, ty, _} <- whnf0 defs ctx =<< computeElimType defs ctx f
| ty => throw $ ExpectedEq ty.loc ctx.names ty
whnf defs ctx $
ends (Ann (setLoc appLoc l) ty.zero appLoc)
(Ann (setLoc appLoc r) ty.one appLoc) e
B {} => pure $ Element (DApp f p appLoc) $ fnf `orNo` ndlh `orNo` Ah
-- e ∷ A ⇝ e
whnf defs ctx (Ann s a annLoc) = do
Element s snf <- whnf defs ctx s
case nchoose $ isE s of
Left _ => let E e = s in pure $ Element e $ noOr2 snf
Right ne => do
Element a anf <- whnf defs ctx a
pure $ Element (Ann s a annLoc) $ ne `orNo` snf `orNo` anf
whnf defs ctx (Coe (S _ (N ty)) _ _ val coeLoc) =
whnf defs ctx $ Ann val ty coeLoc
whnf defs ctx (Coe (S [< i] (Y ty)) p q val coeLoc) = do
Element ty tynf <- whnf defs (extendDim i ctx) ty
2023-05-01 21:06:25 -04:00
Element val valnf <- whnf defs ctx val
2023-08-13 14:55:14 -04:00
case nchoose $ canPushCoe ty val of
Right n => pure $ Element (Coe (SY [< i] ty) p q val coeLoc) $
tynf `orNo` n
Left y => pushCoe defs ctx i ty p q val coeLoc
2023-05-01 21:06:25 -04:00
whnf defs ctx (Comp ty p q val r zero one compLoc) =
-- comp [A] @p @p s { ⋯ } ⇝ s ∷ A
if p == q then whnf defs ctx $ Ann val ty compLoc else
case nchoose (isK r) of
-- comp [A] @p @q s @0 { 0 j ⇒ t; ⋯ } ⇝ tq/j ∷ A
-- comp [A] @p @q s @1 { 1 j ⇒ t; ⋯ } ⇝ tq/j ∷ A
Left y => case r of
K Zero _ => whnf defs ctx $ Ann (dsub1 zero q) ty compLoc
K One _ => whnf defs ctx $ Ann (dsub1 one q) ty compLoc
Right nk => do
Element ty tynf <- whnf defs ctx ty
pure $ Element (Comp ty p q val r zero one compLoc) $ tynf `orNo` nk
whnf defs ctx (TypeCase ty ret arms def tcLoc) = do
Element ty tynf <- whnf defs ctx ty
Element ret retnf <- whnf defs ctx ret
case nchoose $ isAnnTyCon ty of
Left y =>
let Ann ty (TYPE u _) _ = ty in
reduceTypeCase defs ctx ty u ret arms def tcLoc
Right nt => pure $
Element (TypeCase ty ret arms def tcLoc) (tynf `orNo` retnf `orNo` nt)
whnf defs ctx (CloE (Sub el th)) = whnf defs ctx $ pushSubstsWith' id th el
whnf defs ctx (DCloE (Sub el th)) = whnf defs ctx $ pushSubstsWith' th id el
export covering
2023-05-21 14:11:01 -04:00
CanWhnf Term Reduce.isRedexT where
2023-05-01 21:06:25 -04:00
whnf _ _ t@(TYPE {}) = pure $ nred t
whnf _ _ t@(Pi {}) = pure $ nred t
whnf _ _ t@(Lam {}) = pure $ nred t
whnf _ _ t@(Sig {}) = pure $ nred t
whnf _ _ t@(Pair {}) = pure $ nred t
2023-08-06 04:46:55 -04:00
whnf _ _ t@(W {}) = pure $ nred t
whnf _ _ t@(Sup {}) = pure $ nred t
2023-05-01 21:06:25 -04:00
whnf _ _ t@(Enum {}) = pure $ nred t
whnf _ _ t@(Tag {}) = pure $ nred t
whnf _ _ t@(Eq {}) = pure $ nred t
whnf _ _ t@(DLam {}) = pure $ nred t
whnf _ _ t@(Nat {}) = pure $ nred t
whnf _ _ t@(Zero {}) = pure $ nred t
whnf _ _ t@(Succ {}) = pure $ nred t
whnf _ _ t@(BOX {}) = pure $ nred t
whnf _ _ t@(Box {}) = pure $ nred t
-- s ∷ A ⇝ s (in term context)
whnf defs ctx (E e) = do
Element e enf <- whnf defs ctx e
case nchoose $ isAnn e of
Left _ => let Ann {tm, _} = e in pure $ Element tm $ noOr1 $ noOr2 enf
Right na => pure $ Element (E e) $ na `orNo` enf
whnf defs ctx (CloT (Sub tm th)) = whnf defs ctx $ pushSubstsWith' id th tm
whnf defs ctx (DCloT (Sub tm th)) = whnf defs ctx $ pushSubstsWith' th id tm