quox/lib/Quox/Reduce.idr

719 lines
28 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-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
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
public export
0 RedexTest : TermLike -> Type
2023-04-17 17:58:24 -04:00
RedexTest tm = {d, n : Nat} -> Mods -> Definitions -> tm d n -> Bool
public export
2023-04-15 09:13:01 -04:00
interface Whnf (0 tm : TermLike) (0 isRedex : RedexTest tm) | tm
where
2023-04-17 17:58:24 -04:00
whnf : {d, n : Nat} -> (ns : Mods) -> (defs : Definitions) ->
(ctx : WhnfContext d n) ->
tm d n -> Either Error (Subset (tm d n) (No . isRedex ns defs))
2023-04-15 09:13:01 -04:00
public export %inline
whnf0 : {d, n : Nat} -> {0 isRedex : RedexTest tm} -> Whnf tm isRedex =>
2023-04-17 17:58:24 -04:00
(ns : Mods) -> (defs : Definitions) ->
WhnfContext d n -> tm d n -> Either Error (tm d n)
whnf0 ns defs ctx t = fst <$> whnf ns defs ctx t
public export
2023-04-15 09:13:01 -04:00
0 IsRedex, NotRedex : {isRedex : RedexTest tm} -> Whnf tm isRedex =>
2023-04-17 17:58:24 -04:00
Mods -> Definitions -> Pred (tm d n)
IsRedex ns defs = So . isRedex ns defs
NotRedex ns defs = No . isRedex ns defs
public export
0 NonRedex : (tm : TermLike) -> {isRedex : RedexTest tm} ->
2023-04-17 17:58:24 -04:00
Whnf tm isRedex => (d, n : Nat) ->
(ns : Mods) -> (defs : Definitions) -> Type
NonRedex tm d n ns defs = Subset (tm d n) (NotRedex ns defs)
public export %inline
2023-04-15 09:13:01 -04:00
nred : {0 isRedex : RedexTest tm} -> (0 _ : Whnf tm isRedex) =>
2023-04-17 17:58:24 -04:00
(t : tm d n) -> (0 nr : NotRedex ns defs t) =>
NonRedex tm d n ns 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
isLamHead (Lam {} :# Pi {}) = True
2023-04-15 09:13:01 -04:00
isLamHead (Coe {}) = True
isLamHead _ = False
2022-05-25 09:59:58 -04:00
public export %inline
isDLamHead : Elim {} -> Bool
isDLamHead (DLam {} :# Eq {}) = True
2023-04-15 09:13:01 -04:00
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
isPairHead (Pair {} :# Sig {}) = True
2023-04-15 09:13:01 -04:00
isPairHead (Coe {}) = True
2023-01-26 13:54:46 -05:00
isPairHead _ = False
public export %inline
isTagHead : Elim {} -> Bool
isTagHead (Tag t :# Enum _) = True
2023-04-15 09:13:01 -04:00
isTagHead (Coe {}) = True
isTagHead _ = False
2023-03-26 08:40:54 -04:00
public export %inline
isNatHead : Elim {} -> Bool
isNatHead (Zero :# Nat) = True
isNatHead (Succ n :# Nat) = True
2023-04-15 09:13:01 -04:00
isNatHead (Coe {}) = True
2023-03-26 08:40:54 -04:00
isNatHead _ = False
2023-03-31 13:11:35 -04:00
public export %inline
isBoxHead : Elim {} -> Bool
isBoxHead (Box {} :# BOX {}) = True
2023-04-15 09:13:01 -04:00
isBoxHead (Coe {}) = True
2023-03-31 13:11:35 -04:00
isBoxHead _ = False
2022-05-25 09:59:58 -04:00
public export %inline
isE : Term {} -> Bool
isE (E _) = True
isE _ = False
2022-05-25 09:59:58 -04:00
public export %inline
isAnn : Elim {} -> Bool
isAnn (_ :# _) = 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
isTyCon (TYPE {}) = True
isTyCon (Pi {}) = True
isTyCon (Lam {}) = False
isTyCon (Sig {}) = True
isTyCon (Pair {}) = False
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
||| 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-04-03 11:46:23 -04:00
isAnnTyCon (ty :# TYPE _) = isTyCon ty
isAnnTyCon _ = False
2023-04-15 09:13:01 -04:00
public export %inline
isK : Dim d -> Bool
isK (K _) = True
isK _ = False
mutual
public export
isRedexE : RedexTest Elim
2023-04-17 17:58:24 -04:00
isRedexE ns defs (F x) {d, n} =
isJust $ lookupElimNS ns x defs {d, n}
isRedexE _ _ (B _) = False
isRedexE ns defs (f :@ _) =
isRedexE ns defs f || isLamHead f
isRedexE ns defs (CasePair {pair, _}) =
isRedexE ns defs pair || isPairHead pair
isRedexE ns defs (CaseEnum {tag, _}) =
isRedexE ns defs tag || isTagHead tag
isRedexE ns defs (CaseNat {nat, _}) =
isRedexE ns defs nat || isNatHead nat
isRedexE ns defs (CaseBox {box, _}) =
isRedexE ns defs box || isBoxHead box
isRedexE ns defs (f :% p) =
isRedexE ns defs f || isDLamHead f || isK p
isRedexE ns defs (t :# a) =
isE t || isRedexT ns defs t || isRedexT ns defs a
isRedexE ns defs (Coe {val, _}) =
isRedexT ns defs val || not (isE val)
isRedexE ns defs (Comp {ty, r, _}) =
isRedexT ns defs ty || isK r
isRedexE ns defs (TypeCase {ty, ret, _}) =
isRedexE ns defs ty || isRedexT ns defs ret || isAnnTyCon ty
isRedexE _ _ (CloE {}) = True
isRedexE _ _ (DCloE {}) = True
public export
isRedexT : RedexTest Term
2023-04-17 17:58:24 -04:00
isRedexT _ _ (CloT {}) = True
isRedexT _ _ (DCloT {}) = True
isRedexT ns defs (E e) = isAnn e || isRedexE ns 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
private
weakDS : (by : Nat) -> DScopeTerm d n -> DScopeTerm 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
private
dweakS : (by : Nat) -> ScopeTerm d n -> ScopeTerm (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
private
coeScoped : {s : Nat} -> DScopeTerm d n -> Dim d -> Dim d ->
ScopeTermN s d n -> ScopeTermN s d n
coeScoped ty p q (S names (Y body)) =
S names $ Y $ E $ Coe (weakDS s ty) p q body
coeScoped ty p q (S names (N body)) =
S names $ N $ E $ Coe ty p q body
mutual
2023-04-17 17:58:24 -04:00
parameters {d, n : Nat} (ns : Mods) (defs : Definitions)
(ctx : WhnfContext d n)
2023-04-15 09:13:01 -04:00
||| performs the minimum work required to recompute the type of an elim.
|||
||| ⚠ **assumes the elim is already typechecked.** ⚠
export covering
2023-04-17 17:58:24 -04:00
computeElimType : (e : Elim d n) -> (0 ne : No (isRedexE ns defs e)) =>
2023-04-15 09:13:01 -04:00
Either Error (Term d n)
computeElimType (F x) = do
2023-04-17 17:58:24 -04:00
let Just def = lookupNS ns x defs | Nothing => Left $ NotInScope x
2023-04-15 09:13:01 -04:00
pure $ def.type
computeElimType (B i) = pure $ ctx.tctx !! i
computeElimType (f :@ s) {ne} = do
-- can't use `expectPi` (or `expectEq` below) without making this
-- mutual block from hell even worse lol
2023-04-17 17:58:24 -04:00
Pi {arg, res, _} <- whnf0 ns defs ctx =<<
computeElimType f {ne = noOr1 ne}
2023-04-15 09:13:01 -04:00
| t => Left $ ExpectedPi ctx.names t
pure $ sub1 res (s :# arg)
computeElimType (CasePair {pair, ret, _}) = pure $ sub1 ret pair
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 (f :% p) {ne} = do
2023-04-17 17:58:24 -04:00
Eq {ty, _} <- whnf0 ns defs ctx =<< computeElimType f {ne = noOr1 ne}
2023-04-15 09:13:01 -04:00
| t => Left $ ExpectedEq ctx.names t
pure $ dsub1 ty p
computeElimType (Coe {ty, q, _}) = pure $ dsub1 ty q
computeElimType (Comp {ty, _}) = pure ty
computeElimType (TypeCase {ret, _}) = pure ret
computeElimType (_ :# ty) = pure ty
2023-04-17 17:58:24 -04:00
parameters {d, n : Nat} (ns : Mods) (defs : Definitions)
(ctx : WhnfContext (S d) n)
2023-04-15 09:13:01 -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
2023-04-17 17:58:24 -04:00
tycasePi : (t : Term (S d) n) -> (0 tnf : No (isRedexT ns defs t)) =>
2023-04-15 09:13:01 -04:00
Either Error (Term (S d) n, ScopeTerm (S d) n)
tycasePi (Pi {arg, res, _}) = pure (arg, res)
tycasePi (E e) {tnf} = do
2023-04-17 17:58:24 -04:00
ty <- computeElimType ns defs ctx e @{noOr2 tnf}
let arg = E $ typeCase1Y e ty KPi [< "Arg", "Ret"] (BVT 1)
res' = typeCase1Y e (Arr Zero arg ty) KPi [< "Arg", "Ret"] (BVT 0)
2023-04-15 09:13:01 -04:00
res = SY [< "Arg"] $ E $ weakE 1 res' :@ BVT 0
pure (arg, res)
tycasePi t = Left $ ExpectedPi 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
2023-04-17 17:58:24 -04:00
tycaseSig : (t : Term (S d) n) -> (0 tnf : No (isRedexT ns defs t)) =>
2023-04-15 09:13:01 -04:00
Either Error (Term (S d) n, ScopeTerm (S d) n)
tycaseSig (Sig {fst, snd, _}) = pure (fst, snd)
tycaseSig (E e) {tnf} = do
2023-04-17 17:58:24 -04:00
ty <- computeElimType ns defs ctx e @{noOr2 tnf}
let fst = E $ typeCase1Y e ty KSig [< "Fst", "Snd"] (BVT 1)
snd' = typeCase1Y e (Arr Zero fst ty) KSig [< "Fst", "Snd"] (BVT 0)
2023-04-15 09:13:01 -04:00
snd = SY [< "Fst"] $ E $ weakE 1 snd' :@ BVT 0
pure (fst, snd)
tycaseSig t = Left $ ExpectedSig ctx.names t
||| for [π. A], returns A;
||| for an elim returns a type-case that will reduce to that;
||| for other intro forms error
private covering
2023-04-17 17:58:24 -04:00
tycaseBOX : (t : Term (S d) n) -> (0 tnf : No (isRedexT ns defs t)) =>
2023-04-15 09:13:01 -04:00
Either Error (Term (S d) n)
tycaseBOX (BOX _ a) = pure a
tycaseBOX (E e) {tnf} = do
2023-04-17 17:58:24 -04:00
ty <- computeElimType ns defs ctx e @{noOr2 tnf}
pure $ E $ typeCase1Y e ty KBOX [< "Ty"] (BVT 0)
2023-04-15 09:13:01 -04:00
tycaseBOX t = Left $ ExpectedBOX 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
2023-04-17 17:58:24 -04:00
tycaseEq : (t : Term (S d) n) -> (0 tnf : No (isRedexT ns defs t)) =>
2023-04-15 09:13:01 -04:00
Either Error (Term (S d) n, Term (S d) n, DScopeTerm (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 (E e) {tnf} = do
2023-04-17 17:58:24 -04:00
ty <- computeElimType ns defs ctx e @{noOr2 tnf}
2023-04-15 09:13:01 -04:00
let names = [< "A0", "A1", "A", "L", "R"]
a0 = E $ typeCase1Y e ty KEq names (BVT 4)
a1 = E $ typeCase1Y e ty KEq names (BVT 3)
a' = typeCase1Y e (Eq0 ty a0 a1) KEq names (BVT 2)
2023-04-15 09:13:01 -04:00
a = SY [< "i"] $ E $ dweakE 1 a' :% BV 0
l = E $ typeCase1Y e a0 KEq names (BVT 1)
r = E $ typeCase1Y e a1 KEq names (BVT 0)
2023-04-15 09:13:01 -04:00
pure (a0, a1, a, l, r)
tycaseEq t = Left $ ExpectedEq ctx.names t
2023-04-17 17:58:24 -04:00
-- new block because the functions below might pass a different ctx
-- into the ones above
parameters {d, n : Nat} (ns : Mods) (defs : Definitions)
(ctx : WhnfContext d n)
2023-04-15 09:13:01 -04:00
||| reduce a function application `Coe ty p q val :@ s`
private covering
piCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val, s : Term d n) ->
2023-04-17 17:58:24 -04:00
Either Error (Subset (Elim d n) (No . isRedexE ns defs))
piCoe sty@(S [< i] ty) p q val s = do
2023-04-15 09:13:01 -04:00
-- (coe [i ⇒ π.(x : A) → B] @p @q t) s ⇝
-- coe [i ⇒ B[𝒔i/x] @p @q ((t ∷ (π.(x : A) → B)p/i) 𝒔p)
-- where 𝒔j ≔ coe [i ⇒ A] @q @j s
--
-- type-case is used to expose A,B if the type is neutral
let ctx1 = extendDim i ctx
2023-04-17 17:58:24 -04:00
Element ty tynf <- whnf ns defs ctx1 ty.term
(arg, res) <- tycasePi ns defs ctx1 ty
let s0 = CoeT i arg q p s
2023-04-15 09:13:01 -04:00
body = E $ (val :# (ty // one p)) :@ E s0
s1 = CoeT i (arg // (BV 0 ::: shift 2)) (weakD 1 q) (BV 0)
(s // shift 1)
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ CoeT i (sub1 res s1) p q body
2023-04-15 09:13:01 -04:00
||| reduce a pair elimination `CasePair pi (Coe ty p q val) ret body`
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) ->
2023-04-17 17:58:24 -04:00
Either Error (Subset (Elim d n) (No . isRedexE ns defs))
sigCoe qty sty@(S [< i] ty) p q val ret body = do
2023-04-15 09:13:01 -04:00
-- caseπ (coe [i ⇒ (x : A) × B] @p @q s) return z ⇒ C of { (a, b) ⇒ e }
-- ⇝
-- caseπ s ∷ ((x : A) × B)p/i return z ⇒ C
-- of { (a, b) ⇒
-- e[(coe [i ⇒ A] @p @q a)/a,
-- (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
2023-04-17 17:58:24 -04:00
Element ty tynf <- whnf ns defs ctx1 ty.term
(tfst, tsnd) <- tycaseSig ns defs ctx1 ty
let a' = CoeT i (weakT 2 tfst) p q (BVT 1)
2023-04-15 09:13:01 -04:00
tsnd' = tsnd.term //
(CoeT i (weakT 2 $ tfst // (BV 0 ::: shift 2))
2023-04-15 09:13:01 -04:00
(weakD 1 p) (BV 0) (BVT 1) ::: shift 2)
b' = CoeT i tsnd' p q (BVT 0)
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ CasePair qty (val :# (ty // one p)) ret $
ST body.names $ body.term // (a' ::: b' ::: shift 2)
2023-04-15 09:13:01 -04:00
||| reduce a dimension application `Coe ty p q val :% r`
private covering
eqCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
(r : Dim d) ->
2023-04-17 17:58:24 -04:00
Either Error (Subset (Elim d n) (No . isRedexE ns defs))
eqCoe sty@(S [< j] ty) p q val r = do
2023-04-15 09:13:01 -04:00
-- (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; (r=1) j ⇒ R }
let ctx1 = extendDim j ctx
2023-04-17 17:58:24 -04:00
Element ty tynf <- whnf ns defs ctx1 ty.term
(a0, a1, a, s, t) <- tycaseEq ns defs ctx1 ty
let a' = dsub1 a (weakD 1 r)
2023-04-15 09:13:01 -04:00
val' = E $ (val :# (ty // one p)) :% r
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ CompH j a' p q val' r j s j t
2023-04-15 09:13:01 -04:00
||| reduce a pair elimination `CaseBox pi (Coe ty p q val) ret body`
private covering
boxCoe : (qty : Qty) ->
(ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
(ret : ScopeTerm d n) -> (body : ScopeTerm d n) ->
2023-04-17 17:58:24 -04:00
Either Error (Subset (Elim d n) (No . isRedexE ns defs))
boxCoe qty sty@(S [< i] ty) p q val ret body = do
2023-04-15 09:13:01 -04:00
-- 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
2023-04-17 17:58:24 -04:00
Element ty tynf <- whnf ns defs ctx1 ty.term
ta <- tycaseBOX ns defs ctx1 ty
let a' = CoeT i (weakT 1 ta) p q $ BVT 0
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ CaseBox qty (val :# (ty // one p)) ret $
ST body.names $ body.term // (a' ::: shift 1)
2023-04-15 09:13:01 -04:00
export covering
2023-04-15 09:13:01 -04:00
Whnf Elim Reduce.isRedexE where
2023-04-17 17:58:24 -04:00
whnf ns defs ctx (F x) with (lookupElimNS ns x defs) proof eq
_ | Just y = whnf ns defs ctx y
_ | Nothing = pure $ Element (F x) $ rewrite eq in Ah
2023-04-17 17:58:24 -04:00
whnf _ _ _ (B i) = pure $ nred $ B i
2023-03-31 13:11:35 -04:00
-- ((λ x ⇒ t) ∷ (π.x : A) → B) s ⇝ t[s∷A/x] ∷ B[s∷A/x]
2023-04-17 17:58:24 -04:00
whnf ns defs ctx (f :@ s) = do
Element f fnf <- whnf ns defs ctx f
case nchoose $ isLamHead f of
2023-04-15 09:13:01 -04:00
Left _ => case f of
Lam body :# Pi {arg, res, _} =>
let s = s :# arg in
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ sub1 body s :# sub1 res s
Coe ty p q val => piCoe ns defs ctx ty p q val s
Right nlh => pure $ Element (f :@ s) $ fnf `orNo` nlh
2023-03-31 13:11:35 -04:00
-- 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]
2023-04-17 17:58:24 -04:00
whnf ns defs ctx (CasePair pi pair ret body) = do
Element pair pairnf <- whnf ns defs ctx pair
case nchoose $ isPairHead pair of
2023-04-15 09:13:01 -04:00
Left _ => case pair of
Pair {fst, snd} :# Sig {fst = tfst, snd = tsnd, _} =>
let fst = fst :# tfst
snd = snd :# sub1 tsnd fst
in
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ subN body [< fst, snd] :# sub1 ret pair
2023-04-15 09:13:01 -04:00
Coe ty p q val => do
2023-04-17 17:58:24 -04:00
sigCoe ns defs ctx pi ty p q val ret body
Right np =>
2023-04-15 09:13:01 -04:00
pure $ Element (CasePair pi pair ret body) $ pairnf `orNo` np
2023-03-31 13:11:35 -04:00
-- case 'a ∷ {a,…} return p ⇒ C of { 'a ⇒ u } ⇝
-- u ∷ C['a∷{a,…}/p]
2023-04-17 17:58:24 -04:00
whnf ns defs ctx (CaseEnum pi tag ret arms) = do
Element tag tagnf <- whnf ns defs ctx tag
case nchoose $ isTagHead tag of
2023-04-15 09:13:01 -04:00
Left t => case tag of
Tag t :# Enum ts =>
let ty = sub1 ret tag in
case lookup t arms of
2023-04-17 17:58:24 -04:00
Just arm => whnf ns defs ctx $ arm :# ty
2023-04-15 09:13:01 -04:00
Nothing => Left $ MissingEnumArm t (keys arms)
Coe ty p q val =>
-- there is nowhere an equality can be hiding inside an Enum
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ CaseEnum pi (val :# dsub1 ty q) ret arms
Right nt =>
pure $ Element (CaseEnum pi tag ret arms) $ tagnf `orNo` nt
2023-03-31 13:11:35 -04:00
-- case zero ∷ return p ⇒ C of { zero ⇒ u; … } ⇝
-- u ∷ C[zero∷/p]
--
2023-04-15 09:13:01 -04:00
-- case succ n ∷ return p ⇒ C of { succ n', π.ih ⇒ u; … } ⇝
2023-03-31 13:11:35 -04:00
-- u[n∷/n', (case n ∷ ⋯)/ih] ∷ C[succ n ∷ /p]
2023-04-17 17:58:24 -04:00
whnf ns defs ctx (CaseNat pi piIH nat ret zer suc) = do
Element nat natnf <- whnf ns defs ctx nat
2023-03-26 08:40:54 -04:00
case nchoose $ isNatHead nat of
Left _ =>
let ty = sub1 ret nat in
case nat of
2023-04-17 17:58:24 -04:00
Zero :# Nat => whnf ns defs ctx (zer :# ty)
2023-03-26 08:40:54 -04:00
Succ n :# Nat =>
let nn = n :# Nat
2023-03-26 10:09:47 -04:00
tm = subN suc [< nn, CaseNat pi piIH nn ret zer suc]
2023-03-26 08:40:54 -04:00
in
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ tm :# ty
2023-04-15 09:13:01 -04:00
Coe ty p q val =>
-- same deal as Enum
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ CaseNat pi piIH (val :# dsub1 ty q) ret zer suc
2023-03-26 08:40:54 -04:00
Right nn =>
pure $ Element (CaseNat pi piIH nat ret zer suc) $ natnf `orNo` nn
2023-03-31 13:11:35 -04:00
-- case [t] ∷ [π.A] return p ⇒ C of { [x] ⇒ u } ⇝
-- u[t∷A/x] ∷ C[[t] ∷ [π.A]/p]
2023-04-17 17:58:24 -04:00
whnf ns defs ctx (CaseBox pi box ret body) = do
Element box boxnf <- whnf ns defs ctx box
2023-03-31 13:11:35 -04:00
case nchoose $ isBoxHead box of
2023-04-15 09:13:01 -04:00
Left _ => case box of
Box val :# BOX q bty =>
let ty = sub1 ret box in
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ sub1 body (val :# bty) :# ty
2023-04-15 09:13:01 -04:00
Coe ty p q val =>
2023-04-17 17:58:24 -04:00
boxCoe ns defs ctx pi ty p q val ret body
2023-03-31 13:11:35 -04:00
Right nb =>
pure $ Element (CaseBox pi box ret body) $ boxnf `orNo` nb
-- ((δ 𝑖 ⇒ s) ∷ Eq [𝑗 ⇒ A] t u) @0 ⇝ t ∷ A0/𝑗
-- ((δ 𝑖 ⇒ s) ∷ Eq [𝑗 ⇒ A] t u) @1 ⇝ u ∷ A1/𝑗
-- ((δ 𝑖 ⇒ s) ∷ Eq [𝑗 ⇒ A] t u) @𝑘 ⇝ s𝑘/𝑖 ∷ A𝑘/𝑗
-- (if 𝑘 is a variable)
2023-04-17 17:58:24 -04:00
whnf ns defs ctx (f :% p) = do
Element f fnf <- whnf ns defs ctx f
case nchoose $ isDLamHead f of
2023-04-15 09:13:01 -04:00
Left _ => case f of
DLam body :# Eq {ty = ty, l, r, _} =>
let body = endsOr l r (dsub1 body p) p in
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ body :# dsub1 ty p
2023-04-15 09:13:01 -04:00
Coe ty p' q' val =>
2023-04-17 17:58:24 -04:00
eqCoe ns defs ctx ty p' q' val p
2023-04-15 09:13:01 -04:00
Right ndlh => case p of
K e => do
2023-04-17 17:58:24 -04:00
Eq {l, r, ty, _} <- whnf0 ns defs ctx =<<
computeElimType ns defs ctx f
2023-04-15 09:13:01 -04:00
| ty => Left $ ExpectedEq ctx.names ty
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ ends (l :# ty.zero) (r :# ty.one) e
2023-04-15 09:13:01 -04:00
B _ => pure $ Element (f :% p) $ fnf `orNo` ndlh `orNo` Ah
2023-03-31 13:11:35 -04:00
-- e ∷ A ⇝ e
2023-04-17 17:58:24 -04:00
whnf ns defs ctx (s :# a) = do
Element s snf <- whnf ns defs ctx s
case nchoose $ isE s of
Left _ => let E e = s in pure $ Element e $ noOr2 snf
Right ne => do
2023-04-17 17:58:24 -04:00
Element a anf <- whnf ns defs ctx a
pure $ Element (s :# a) $ ne `orNo` snf `orNo` anf
2023-04-17 17:58:24 -04:00
whnf ns defs ctx (Coe (S _ (N ty)) _ _ val) =
whnf ns defs ctx $ val :# ty
whnf ns defs ctx (Coe (S [< i] ty) p q val) = do
Element ty tynf <- whnf ns defs (extendDim i ctx) ty.term
Element val valnf <- whnf ns defs ctx val
pushCoe ns defs ctx i ty p q val
2023-04-15 09:13:01 -04:00
2023-04-17 17:58:24 -04:00
whnf ns defs ctx (Comp ty p q val r zero one) =
2023-04-15 09:13:01 -04:00
-- comp [A] @p @p s { ⋯ } ⇝ s ∷ A
2023-04-17 17:58:24 -04:00
if p == q then whnf ns defs ctx $ val :# ty else
2023-04-15 09:13:01 -04:00
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
2023-04-17 17:58:24 -04:00
K Zero => whnf ns defs ctx $ dsub1 zero q :# ty
K One => whnf ns defs ctx $ dsub1 one q :# ty
2023-04-15 09:13:01 -04:00
Right nk => do
2023-04-17 17:58:24 -04:00
Element ty tynf <- whnf ns defs ctx ty
2023-04-15 09:13:01 -04:00
pure $ Element (Comp ty p q val r zero one) $ tynf `orNo` nk
-- [todo] anything other than just the boundaries??
2023-04-17 17:58:24 -04:00
whnf ns defs ctx (TypeCase ty ret arms def) = do
Element ty tynf <- whnf ns defs ctx ty
Element ret retnf <- whnf ns defs ctx ret
2023-04-15 09:13:01 -04:00
case nchoose $ isAnnTyCon ty of
Left y => let ty :# TYPE u = ty in
2023-04-17 17:58:24 -04:00
reduceTypeCase ns defs ctx ty u ret arms def
2023-04-15 09:13:01 -04:00
Right nt => pure $ Element (TypeCase ty ret arms def) $
tynf `orNo` retnf `orNo` nt
2023-04-17 17:58:24 -04:00
whnf ns defs ctx (CloE el th) = whnf ns defs ctx $ pushSubstsWith' id th el
whnf ns defs ctx (DCloE el th) = whnf ns defs ctx $ pushSubstsWith' th id el
2023-04-15 09:13:01 -04:00
export covering
Whnf Term Reduce.isRedexT where
2023-04-17 17:58:24 -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
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 _ _ _ Nat = pure $ nred Nat
whnf _ _ _ Zero = pure $ nred Zero
whnf _ _ _ t@(Succ {}) = pure $ nred t
whnf _ _ _ t@(BOX {}) = pure $ nred t
whnf _ _ _ t@(Box {}) = pure $ nred t
2023-04-15 09:13:01 -04:00
-- s ∷ A ⇝ s (in term context)
2023-04-17 17:58:24 -04:00
whnf ns defs ctx (E e) = do
Element e enf <- whnf ns defs ctx e
2023-04-15 09:13:01 -04:00
case nchoose $ isAnn e of
Left _ => let tm :# _ = e in pure $ Element tm $ noOr1 $ noOr2 enf
Right na => pure $ Element (E e) $ na `orNo` enf
2023-04-17 17:58:24 -04:00
whnf ns defs ctx (CloT tm th) = whnf ns defs ctx $ pushSubstsWith' id th tm
whnf ns defs ctx (DCloT tm th) = whnf ns defs ctx $ pushSubstsWith' th id tm
2023-04-15 09:13:01 -04:00
||| reduce a type-case applied to a type constructor
private covering
2023-04-17 17:58:24 -04:00
reduceTypeCase : {d, n : Nat} -> (ns : Mods) -> (defs : Definitions) ->
WhnfContext d n ->
2023-04-15 09:13:01 -04:00
(ty : Term d n) -> (u : Universe) -> (ret : Term d n) ->
(arms : TypeCaseArms d n) -> (def : Term d n) ->
(0 _ : So (isTyCon ty)) =>
2023-04-17 17:58:24 -04:00
Either Error (Subset (Elim d n) (No . isRedexE ns defs))
reduceTypeCase ns defs ctx ty u ret arms def = case ty of
2023-04-03 11:46:23 -04:00
-- (type-case ★ᵢ ∷ _ return Q of { ★ ⇒ s; ⋯ }) ⇝ s ∷ Q
2023-04-15 09:13:01 -04:00
TYPE _ =>
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ tycaseRhsDef0 def KTYPE arms :# ret
2023-04-15 09:13:01 -04:00
2023-04-03 11:46:23 -04:00
-- (type-case π.(x : A) → B ∷ ★ᵢ return Q of { (a → b) ⇒ s; ⋯ }) ⇝
-- s[(A ∷ ★ᵢ)/a, ((λ x ⇒ B) ∷ 0.A → ★ᵢ)/b] ∷ ★ᵢ
2023-04-15 09:13:01 -04:00
Pi _ arg res =>
let arg = arg :# TYPE u
res = Lam res :# Arr Zero (TYPE u) (TYPE u)
in
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ subN (tycaseRhsDef def KPi arms) [< arg, res] :# ret
2023-04-15 09:13:01 -04:00
2023-04-03 11:46:23 -04:00
-- (type-case (x : A) × B ∷ ★ᵢ return Q of { (a × b) ⇒ s; ⋯ }) ⇝
-- s[(A ∷ ★ᵢ)/a, ((λ x ⇒ B) ∷ 0.A → ★ᵢ)/b] ∷ ★ᵢ
2023-04-15 09:13:01 -04:00
Sig fst snd =>
let fst = fst :# TYPE u
snd = Lam snd :# Arr Zero (TYPE u) (TYPE u)
in
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ subN (tycaseRhsDef def KSig arms) [< fst, snd] :# ret
2023-04-15 09:13:01 -04:00
2023-04-03 11:46:23 -04:00
-- (type-case {⋯} ∷ _ return Q of { {} ⇒ s; ⋯ }) ⇝ s ∷ Q
2023-04-15 09:13:01 -04:00
Enum _ =>
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ tycaseRhsDef0 def KEnum arms :# ret
2023-04-15 09:13:01 -04:00
-- (type-case Eq [i ⇒ A] L R ∷ ★ᵢ return Q
-- of { Eq a₀ a₁ a l r ⇒ s; ⋯ }) ⇝
2023-04-03 11:46:23 -04:00
-- s[(A0/i ∷ ★ᵢ)/a₀, (A1/i ∷ ★ᵢ)/a₁,
-- ((δ i ⇒ A) ∷ Eq [★ᵢ] A0/i A1/i)/a,
-- (L ∷ A0/i)/l, (R ∷ A1/i)/r] ∷ Q
2023-04-15 09:13:01 -04:00
Eq a l r =>
let a0 = a.zero; a1 = a.one in
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $
2023-04-15 09:13:01 -04:00
subN (tycaseRhsDef def KEq arms)
[< a0 :# TYPE u, a1 :# TYPE u,
DLam a :# Eq0 (TYPE u) a0 a1, l :# a0, r :# a1]
:# ret
2023-04-15 09:13:01 -04:00
-- (type-case ∷ _ return Q of { ⇒ s; ⋯ }) ⇝ s ∷ Q
Nat =>
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ tycaseRhsDef0 def KNat arms :# ret
2023-04-15 09:13:01 -04:00
-- (type-case [π.A] ∷ ★ᵢ return Q of { [a] ⇒ s; ⋯ }) ⇝ s[(A ∷ ★ᵢ)/a] ∷ Q
BOX _ s =>
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ sub1 (tycaseRhsDef def KBOX arms) (s :# TYPE u) :# ret
2023-04-15 09:13:01 -04:00
||| pushes a coercion inside a whnf-ed term
private covering
2023-04-17 17:58:24 -04:00
pushCoe : {n, d : Nat} -> (ns : Mods) -> (defs : Definitions) ->
WhnfContext d n ->
2023-04-15 09:13:01 -04:00
BaseName ->
2023-04-17 17:58:24 -04:00
(ty : Term (S d) n) -> (0 tynf : No (isRedexT ns defs ty)) =>
2023-04-15 09:13:01 -04:00
Dim d -> Dim d ->
2023-04-17 17:58:24 -04:00
(s : Term d n) -> (0 snf : No (isRedexT ns defs s)) =>
Either Error (NonRedex Elim d n ns defs)
pushCoe ns defs ctx i ty p q s =
if p == q then whnf ns defs ctx $ s :# (ty // one q) else
2023-04-15 09:13:01 -04:00
case s of
-- (coe [_ ⇒ ★ᵢ] @_ @_ ty) ⇝ (ty ∷ ★ᵢ)
TYPE {} => pure $ nred $ s :# TYPE !(unwrapTYPE ty)
Pi {} => pure $ nred $ s :# TYPE !(unwrapTYPE ty)
Sig {} => pure $ nred $ s :# TYPE !(unwrapTYPE ty)
Enum {} => pure $ nred $ s :# TYPE !(unwrapTYPE ty)
Eq {} => pure $ nred $ s :# TYPE !(unwrapTYPE ty)
Nat => pure $ nred $ s :# TYPE !(unwrapTYPE ty)
BOX {} => pure $ nred $ s :# TYPE !(unwrapTYPE ty)
-- just η expand it. then whnf for (:@) will handle it later
-- this is how @xtt does it
Lam body => do
let body' = CoeT i ty p q $ Lam body
term' = [< "y"] :\\ E (weakE 1 body' :@ BVT 0)
2023-04-15 09:13:01 -04:00
type' = ty // one q
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ term' :# type'
2023-04-15 09:13:01 -04:00
{-
-- maybe:
-- (coe [i ⇒ π.(x : A) → B] @p @q (λ x ⇒ s)) ⇝
-- (λ x ⇒ coe [i ⇒ B[(coe [j ⇒ Aj/i] @q @i x)/x]] @p @q s)
-- ∷ (π.(x: Aq/i) → Bq/i)
Lam body => do
let Pi {qty, arg, res} = ty
| _ => Left $ ?err
let arg' = ST [< "j"] $ weakT 1 $ arg // (BV 0 ::: shift 2)
res' = ST [< i] $ res.term //
2023-04-15 09:13:01 -04:00
(Coe arg' (weakD 1 q) (BV 0) (BVT 0) ::: shift 1)
body = ST body.names $ E $ Coe res' p q body.term
2023-04-15 09:13:01 -04:00
pure $ Element (Lam body :# Pi qty (arg // one q) (res // one q)) Ah
-}
-- (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
--
-- can't use η here because... it doesn't exist
Pair fst snd => do
let Sig {fst = tfst, snd = tsnd} = ty
| _ => Left $ ExpectedSig (extendDim i ctx.names) ty
let fst' = E $ CoeT i tfst p q fst
tfst' = tfst // (BV 0 ::: shift 2)
tsnd' = sub1 tsnd $ CoeT "j" tfst' (weakD 1 p) (BV 0) $ dweakT 1 fst
snd' = E $ CoeT i tsnd' p q snd
2023-04-15 09:13:01 -04:00
pure $
Element (Pair fst' snd' :# Sig (tfst // one q) (tsnd // one q)) Ah
-- η expand like λ
DLam body => do
let body' = CoeT i ty p q $ DLam body
term' = [< "j"] :\\% E (dweakE 1 body' :% BV 0)
2023-04-15 09:13:01 -04:00
type' = ty // one q
2023-04-17 17:58:24 -04:00
whnf ns defs ctx $ term' :# type'
2023-04-15 09:13:01 -04:00
-- (coe [_ ⇒ {⋯}] @_ @_ t) ⇝ (t ∷ {⋯})
Tag tag => do
let Enum ts = ty
| _ => Left $ ExpectedEnum (extendDim i ctx.names) ty
pure $ Element (Tag tag :# Enum ts) Ah
-- (coe [_ ⇒ ] @_ @_ n) ⇝ (n ∷ )
Zero => pure $ Element (Zero :# Nat) Ah
Succ t => pure $ Element (Succ t :# Nat) Ah
-- (coe [i ⇒ [π.A]] @p @q [s]) ⇝
-- [coe [i ⇒ A] @p @q s] ∷ [π. Aq/i]
Box val => do
let BOX {qty, ty = a} = ty
| _ => Left $ ExpectedBOX (extendDim i ctx.names) ty
pure $ Element (Box (E $ CoeT i a p q s) :# BOX qty (a // one q)) Ah
2023-04-15 09:13:01 -04:00
E e => pure $ Element (CoeT i ty p q (E e)) (snf `orNo` Ah)
2023-04-15 09:13:01 -04:00
where
unwrapTYPE : Term (S d) n -> Either Error Universe
unwrapTYPE (TYPE u) = pure u
unwrapTYPE ty = Left $ ExpectedTYPE (extendDim i ctx.names) ty