pass the subject quantity through equality etc

in preparation for non-linear η laws
This commit is contained in:
rhiannon morris 2023-09-18 18:21:30 +02:00
parent 3fe9b96f05
commit e6c06a5c81
17 changed files with 654 additions and 605 deletions

View file

@ -23,12 +23,12 @@ where
parameters {auto _ : CanWhnf Term Interface.isRedexT}
{auto _ : CanWhnf Elim Interface.isRedexE}
{d, n : Nat} (defs : Definitions) (ctx : WhnfContext d n)
{d, n : Nat} (defs : Definitions) (ctx : WhnfContext d n) (pi : SQty)
||| reduce a function application `App (Coe ty p q val) s loc`
export covering
piCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) ->
(val, s : Term d n) -> Loc ->
Eff Whnf (Subset (Elim d n) (No . isRedexE defs))
Eff Whnf (Subset (Elim d n) (No . isRedexE defs pi))
piCoe sty@(S [< i] ty) p q val s loc = do
-- (coe [i ⇒ π.(x : A) → B] @p @q t) s ⇝
-- coe [i ⇒ B[𝒔i/x] @p @q ((t ∷ (π.(x : A) → B)p/i) 𝒔p)
@ -36,20 +36,20 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
--
-- type-case is used to expose A,B if the type is neutral
let ctx1 = extendDim i ctx
Element ty tynf <- whnf defs ctx1 $ getTerm ty
Element ty tynf <- whnf defs ctx1 SZero $ getTerm ty
(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
whnf defs ctx pi $ CoeT i (sub1 res s1) p q body loc
||| reduce a pair elimination `CasePair pi (Coe ty p q val) ret body loc`
export 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 ->
Eff Whnf (Subset (Elim d n) (No . isRedexE defs))
Eff Whnf (Subset (Elim d n) (No . isRedexE defs pi))
sigCoe qty sty@(S [< i] ty) p q val ret body loc = do
-- caseπ (coe [i ⇒ (x : A) × B] @p @q s) return z ⇒ C of { (a, b) ⇒ e }
-- ⇝
@ -60,7 +60,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
--
-- type-case is used to expose A,B if the type is neutral
let ctx1 = extendDim i ctx
Element ty tynf <- whnf defs ctx1 $ getTerm ty
Element ty tynf <- whnf defs ctx1 SZero $ getTerm ty
(tfst, tsnd) <- tycaseSig defs ctx1 ty
let [< x, y] = body.names
a' = CoeT i (weakT 2 tfst) p q (BVT 1 noLoc) x.loc
@ -68,41 +68,41 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
(CoeT i (weakT 2 $ tfst // (B VZ noLoc ::: shift 2))
(weakD 1 p) (B VZ noLoc) (BVT 1 noLoc) y.loc ::: shift 2)
b' = CoeT i tsnd' p q (BVT 0 noLoc) y.loc
whnf defs ctx $ CasePair qty (Ann val (ty // one p) val.loc) ret
whnf defs ctx pi $ CasePair qty (Ann val (ty // one p) val.loc) ret
(ST body.names $ body.term // (a' ::: b' ::: shift 2)) loc
||| reduce a dimension application `DApp (Coe ty p q val) r loc`
export covering
eqCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
(r : Dim d) -> Loc ->
Eff Whnf (Subset (Elim d n) (No . isRedexE defs))
Eff Whnf (Subset (Elim d n) (No . isRedexE defs pi))
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 $ getTerm ty
Element ty tynf <- whnf defs ctx1 SZero $ getTerm ty
(a0, a1, a, s, t) <- tycaseEq defs ctx1 ty
let a' = dsub1 a (weakD 1 r)
val' = E $ DApp (Ann val (ty // one p) val.loc) r loc
whnf defs ctx $ CompH j a' p q val' r j s j t loc
whnf defs ctx pi $ 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`
export covering
boxCoe : (qty : Qty) ->
(ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
(ret : ScopeTerm d n) -> (body : ScopeTerm d n) -> Loc ->
Eff Whnf (Subset (Elim d n) (No . isRedexE defs))
Eff Whnf (Subset (Elim d n) (No . isRedexE defs pi))
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 $ getTerm ty
Element ty tynf <- whnf defs ctx1 SZero $ getTerm ty
ta <- tycaseBOX defs ctx1 ty
let a' = CoeT i (weakT 1 ta) p q (BVT 0 noLoc) body.name.loc
whnf defs ctx $ CaseBox qty (Ann val (ty // one p) val.loc) ret
whnf defs ctx pi $ CaseBox qty (Ann val (ty // one p) val.loc) ret
(ST body.names $ body.term // (a' ::: shift 1)) loc
@ -110,13 +110,13 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
export covering
pushCoe : BindName ->
(ty : Term (S d) n) -> (p, q : Dim d) -> (s : Term d n) -> Loc ->
(0 pc : So (canPushCoe ty s)) =>
Eff Whnf (NonRedex Elim d n defs)
(0 pc : So (canPushCoe pi ty s)) =>
Eff Whnf (NonRedex Elim d n defs pi)
pushCoe i ty p q s loc =
case ty of
-- (coe ★ᵢ @_ @_ s) ⇝ (s ∷ ★ᵢ)
TYPE l tyLoc =>
whnf defs ctx $ Ann s (TYPE l tyLoc) loc
whnf defs ctx pi $ Ann s (TYPE l tyLoc) loc
-- η expand it so that whnf for App can deal with it
--
@ -125,7 +125,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
-- (λ y ⇒ (coe (𝑖 ⇒ π.(x : A) → B) @p @q s) y) ∷ (π.(x : A) → B)q/𝑖
Pi {} =>
let inner = Coe (SY [< i] ty) p q s loc in
whnf defs ctx $
whnf defs ctx pi $
Ann (LamY !(mnb "y" loc)
(E $ App (weakE 1 inner) (BVT 0 loc) loc) loc)
(ty // one q) loc
@ -147,12 +147,12 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
(tfst // (BV 0 loc ::: shift 2))
(weakD 1 p) (BV 0 loc) (dweakT 1 s) fst.loc
snd' = CoeT i (sub1 tsnd fstInSnd) p q snd snd.loc
whnf defs ctx $
whnf defs ctx pi $
Ann (Pair (E fst') (E snd') sLoc) (ty // one q) loc
-- (coe {𝐚̄} @_ @_ s) ⇝ (s ∷ {𝐚̄})
Enum cases tyLoc =>
whnf defs ctx $ Ann s (Enum cases tyLoc) loc
whnf defs ctx pi $ Ann s (Enum cases tyLoc) loc
-- η expand, same as for Π
--
@ -161,14 +161,14 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
-- (δ 𝑘 ⇒ (coe (𝑖 ⇒ Eq (𝑗 ⇒ A) l r) @p @q s) @𝑘) ∷ (Eq (𝑗 ⇒ A) l r)q/𝑖
Eq {} =>
let inner = Coe (SY [< i] ty) p q s loc in
whnf defs ctx $
whnf defs ctx pi $
Ann (DLamY !(mnb "k" loc)
(E $ DApp (dweakE 1 inner) (BV 0 loc) loc) loc)
(ty // one q) loc
-- (coe @_ @_ s) ⇝ (s ∷ )
Nat tyLoc =>
whnf defs ctx $ Ann s (Nat tyLoc) loc
whnf defs ctx pi $ Ann s (Nat tyLoc) loc
-- η expand
--
@ -185,4 +185,4 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
loc
}
in
whnf defs ctx $ Ann (Box (E inner) loc) (ty // one q) loc
whnf defs ctx pi $ Ann (Box (E inner) loc) (ty // one q) loc

View file

@ -14,8 +14,8 @@ export covering
computeElimType : CanWhnf Term Interface.isRedexT =>
CanWhnf Elim Interface.isRedexE =>
{d, n : Nat} ->
(defs : Definitions) -> WhnfContext d n ->
(e : Elim d n) -> (0 ne : No (isRedexE defs e)) =>
(defs : Definitions) -> WhnfContext d n -> (pi : SQty) ->
(e : Elim d n) -> (0 ne : No (isRedexE defs pi e)) =>
Eff Whnf (Term d n)
@ -24,11 +24,11 @@ export covering
computeWhnfElimType0 : CanWhnf Term Interface.isRedexT =>
CanWhnf Elim Interface.isRedexE =>
{d, n : Nat} ->
(defs : Definitions) -> WhnfContext d n ->
(e : Elim d n) -> (0 ne : No (isRedexE defs e)) =>
(defs : Definitions) -> WhnfContext d n -> (pi : SQty) ->
(e : Elim d n) -> (0 ne : No (isRedexE defs pi e)) =>
Eff Whnf (Term d n)
computeElimType defs ctx e {ne} =
computeElimType defs ctx pi e {ne} =
case e of
F x u loc => do
let Just def = lookup x defs
@ -39,7 +39,7 @@ computeElimType defs ctx e {ne} =
pure $ ctx.tctx !! i
App f s loc =>
case !(computeWhnfElimType0 defs ctx f {ne = noOr1 ne}) of
case !(computeWhnfElimType0 defs ctx pi f {ne = noOr1 ne}) of
Pi {arg, res, _} => pure $ sub1 res $ Ann s arg loc
t => throw $ ExpectedPi loc ctx.names t
@ -56,7 +56,7 @@ computeElimType defs ctx e {ne} =
pure $ sub1 ret box
DApp {fun = f, arg = p, loc} =>
case !(computeWhnfElimType0 defs ctx f {ne = noOr1 ne}) of
case !(computeWhnfElimType0 defs ctx pi f {ne = noOr1 ne}) of
Eq {ty, _} => pure $ dsub1 ty p
t => throw $ ExpectedEq loc ctx.names t
@ -72,5 +72,5 @@ computeElimType defs ctx e {ne} =
TypeCase {ret, _} =>
pure ret
computeWhnfElimType0 defs ctx e =
computeElimType defs ctx e >>= whnf0 defs ctx
computeWhnfElimType0 defs ctx pi e =
computeElimType defs ctx pi e >>= whnf0 defs ctx pi

View file

@ -18,14 +18,14 @@ Whnf = [NameGen, Except Error]
public export
0 RedexTest : TermLike -> Type
RedexTest tm = {d, n : Nat} -> Definitions -> tm d n -> Bool
RedexTest tm = {d, n : Nat} -> Definitions -> SQty -> tm d n -> Bool
public export
interface CanWhnf (0 tm : TermLike) (0 isRedex : RedexTest tm) | tm
where
whnf : {d, n : Nat} -> (defs : Definitions) ->
(ctx : WhnfContext d n) ->
tm d n -> Eff Whnf (Subset (tm d n) (No . isRedex defs))
(ctx : WhnfContext d n) -> (q : SQty) ->
tm d n -> Eff Whnf (Subset (tm d n) (No . isRedex defs q))
-- having isRedex be part of the class header, and needing to be explicitly
-- quantified on every use since idris can't infer its type, is a little ugly.
-- but none of the alternatives i've thought of so far work. e.g. in some
@ -33,23 +33,24 @@ where
public export %inline
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
Definitions -> WhnfContext d n -> SQty -> tm d n -> Eff Whnf (tm d n)
whnf0 defs ctx q t = fst <$> whnf defs ctx q t
public export
0 IsRedex, NotRedex : {isRedex : RedexTest tm} -> CanWhnf tm isRedex =>
Definitions -> Pred (tm d n)
IsRedex defs = So . isRedex defs
NotRedex defs = No . isRedex defs
Definitions -> SQty -> Pred (tm d n)
IsRedex defs q = So . isRedex defs q
NotRedex defs q = No . isRedex defs q
public export
0 NonRedex : (tm : TermLike) -> {isRedex : RedexTest tm} ->
CanWhnf tm isRedex => (d, n : Nat) -> (defs : Definitions) -> Type
NonRedex tm d n defs = Subset (tm d n) (NotRedex defs)
CanWhnf tm isRedex => (d, n : Nat) ->
(defs : Definitions) -> SQty -> Type
NonRedex tm d n defs q = Subset (tm d n) (NotRedex defs q)
public export %inline
nred : {0 isRedex : RedexTest tm} -> (0 _ : CanWhnf tm isRedex) =>
(t : tm d n) -> (0 nr : NotRedex defs t) => NonRedex tm d n defs
(t : tm d n) -> (0 nr : NotRedex defs q t) => NonRedex tm d n defs q
nred t = Element t nr
@ -153,25 +154,25 @@ isK _ = False
||| - `ty` has η
||| - `val` is a constructor form
public export %inline
canPushCoe : (ty, val : Term {}) -> Bool
canPushCoe (TYPE {}) _ = True
canPushCoe (Pi {}) _ = True
canPushCoe (Lam {}) _ = False
canPushCoe (Sig {}) (Pair {}) = True
canPushCoe (Sig {}) _ = False
canPushCoe (Pair {}) _ = 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
canPushCoe : SQty -> (ty, val : Term {}) -> Bool
canPushCoe pi (TYPE {}) _ = True
canPushCoe pi (Pi {}) _ = True
canPushCoe pi (Lam {}) _ = False
canPushCoe pi (Sig {}) (Pair {}) = True
canPushCoe pi (Sig {}) _ = False
canPushCoe pi (Pair {}) _ = False
canPushCoe pi (Enum {}) _ = True
canPushCoe pi (Tag {}) _ = False
canPushCoe pi (Eq {}) _ = True
canPushCoe pi (DLam {}) _ = False
canPushCoe pi (Nat {}) _ = True
canPushCoe pi (Zero {}) _ = False
canPushCoe pi (Succ {}) _ = False
canPushCoe pi (BOX {}) _ = True
canPushCoe pi (Box {}) _ = False
canPushCoe pi (E {}) _ = False
canPushCoe pi (CloT {}) _ = False
canPushCoe pi (DCloT {}) _ = False
mutual
@ -183,42 +184,42 @@ mutual
||| an application whose head is an annotated lambda,
||| a case expression whose head is an annotated constructor form, etc
||| 4. a redundant annotation, or one whose term or type is reducible
||| 5. a coercion `coe (𝑖 ⇒ A) @p @q s` where:
||| 5. a coercion `coe (𝑖 ⇒ A) @p @pi s` where:
||| a. `A` is reducible or a type constructor, or
||| b. `𝑖` is not mentioned in `A`
||| ([fixme] should be A0/𝑖 = A1/𝑖), or
||| c. `p = q`
||| 6. a composition `comp A @p @q s @r {⋯}`
||| where `p = q`, `r = 0`, or `r = 1`
||| c. `p = pi`
||| 6. a composition `comp A @p @pi s @r {⋯}`
||| where `p = pi`, `r = 0`, or `r = 1`
||| 7. a closure
public export
isRedexE : RedexTest Elim
isRedexE defs (F {x, u, _}) {d, n} =
isRedexE defs pi (F {x, u, _}) {d, n} =
isJust $ lookupElim x u defs {d, n}
isRedexE _ (B {}) = False
isRedexE defs (App {fun, _}) =
isRedexE defs fun || isLamHead fun
isRedexE defs (CasePair {pair, _}) =
isRedexE defs pair || isPairHead pair
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
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
isRedexE defs (Coe {ty = S _ (N _), _}) = True
isRedexE defs (Coe {ty = S _ (Y ty), p, q, val, _}) =
isRedexT defs ty || canPushCoe ty val || isYes (p `decEqv` q)
isRedexE defs (Comp {ty, p, q, r, _}) =
isRedexE _ pi (B {}) = False
isRedexE defs pi (App {fun, _}) =
isRedexE defs pi fun || isLamHead fun
isRedexE defs pi (CasePair {pair, _}) =
isRedexE defs pi pair || isPairHead pair
isRedexE defs pi (CaseEnum {tag, _}) =
isRedexE defs pi tag || isTagHead tag
isRedexE defs pi (CaseNat {nat, _}) =
isRedexE defs pi nat || isNatHead nat
isRedexE defs pi (CaseBox {box, _}) =
isRedexE defs pi box || isBoxHead box
isRedexE defs pi (DApp {fun, arg, _}) =
isRedexE defs pi fun || isDLamHead fun || isK arg
isRedexE defs pi (Ann {tm, ty, _}) =
isE tm || isRedexT defs pi tm || isRedexT defs SZero ty
isRedexE defs pi (Coe {ty = S _ (N _), _}) = True
isRedexE defs pi (Coe {ty = S _ (Y ty), p, q, val, _}) =
isRedexT defs SZero ty || canPushCoe pi ty val || isYes (p `decEqv` q)
isRedexE defs pi (Comp {ty, p, q, r, _}) =
isYes (p `decEqv` q) || isK r
isRedexE defs (TypeCase {ty, ret, _}) =
isRedexE defs ty || isRedexT defs ret || isAnnTyCon ty
isRedexE _ (CloE {}) = True
isRedexE _ (DCloE {}) = True
isRedexE defs pi (TypeCase {ty, ret, _}) =
isRedexE defs pi ty || isRedexT defs pi ret || isAnnTyCon ty
isRedexE _ _ (CloE {}) = True
isRedexE _ _ (DCloE {}) = True
||| a reducible term
|||
@ -228,7 +229,7 @@ mutual
||| 3. a closure
public export
isRedexT : RedexTest Term
isRedexT _ (CloT {}) = True
isRedexT _ (DCloT {}) = True
isRedexT defs (E {e, _}) = isAnn e || isRedexE defs e
isRedexT _ _ = False
isRedexT _ _ (CloT {}) = True
isRedexT _ _ (DCloT {}) = True
isRedexT defs pi (E {e, _}) = isAnn e || isRedexE defs pi e
isRedexT _ _ _ = False

View file

@ -16,53 +16,53 @@ export covering CanWhnf Elim Interface.isRedexE
covering
CanWhnf Elim Interface.isRedexE where
whnf defs ctx (F x u loc) with (lookupElim x u defs) proof eq
_ | Just y = whnf defs ctx $ setLoc loc y
whnf defs ctx rh (F x u loc) with (lookupElim x u defs) proof eq
_ | Just y = whnf defs ctx rh $ setLoc loc y
_ | Nothing = pure $ Element (F x u loc) $ rewrite eq in Ah
whnf _ _ (B i loc) = pure $ nred $ B i loc
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
whnf defs ctx rh (App f s appLoc) = do
Element f fnf <- whnf defs ctx rh 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
whnf defs ctx rh $ Ann (sub1 body s) (sub1 res s) appLoc
Coe ty p q val _ => piCoe defs ctx rh ty p q val s appLoc
Right nlh => pure $ Element (App f s appLoc) $ fnf `orNo` nlh
-- 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]
whnf defs ctx (CasePair pi pair ret body caseLoc) = do
Element pair pairnf <- whnf defs ctx pair
whnf defs ctx rh (CasePair pi pair ret body caseLoc) = do
Element pair pairnf <- whnf defs ctx rh 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
whnf defs ctx rh $ 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
sigCoe defs ctx rh pi ty p q val ret body caseLoc
Right np =>
pure $ Element (CasePair pi pair ret body caseLoc) $ pairnf `orNo` np
-- 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
whnf defs ctx rh (CaseEnum pi tag ret arms caseLoc) = do
Element tag tagnf <- whnf defs ctx rh 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
Just arm => whnf defs ctx rh $ 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 $
whnf defs ctx rh $
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
@ -72,37 +72,37 @@ CanWhnf Elim Interface.isRedexE where
--
-- 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
whnf defs ctx rh (CaseNat pi piIH nat ret zer suc caseLoc) = do
Element nat natnf <- whnf defs ctx rh 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
whnf defs ctx rh $ 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
whnf defs ctx rh $ Ann tm ty caseLoc
Coe ty p q val _ =>
-- same deal as Enum
whnf defs ctx $
whnf defs ctx rh $
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
whnf defs ctx rh (CaseBox pi box ret body caseLoc) = do
Element box boxnf <- whnf defs ctx rh 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
whnf defs ctx rh $ 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
boxCoe defs ctx rh pi ty p q val ret body caseLoc
Right nb =>
pure $ Element (CaseBox pi box ret body caseLoc) (boxnf `orNo` nb)
@ -110,35 +110,35 @@ CanWhnf Elim Interface.isRedexE where
-- e : Eq (𝑗 ⇒ A) t u ⊢ e @1 ⇝ u ∷ A1/𝑗
--
-- ((δ 𝑖 ⇒ s) ∷ Eq (𝑗 ⇒ A) t u) @𝑘 ⇝ s𝑘/𝑖 ∷ A𝑘/𝑗
whnf defs ctx (DApp f p appLoc) = do
Element f fnf <- whnf defs ctx f
whnf defs ctx rh (DApp f p appLoc) = do
Element f fnf <- whnf defs ctx rh f
case nchoose $ isDLamHead f of
Left _ => case f of
Ann (DLam {body, _}) (Eq {ty, l, r, _}) _ =>
whnf defs ctx $
whnf defs ctx rh $
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
eqCoe defs ctx rh ty p' q' val p appLoc
Right ndlh => case p of
K e _ => do
Eq {l, r, ty, _} <- computeWhnfElimType0 defs ctx f
Eq {l, r, ty, _} <- computeWhnfElimType0 defs ctx rh f
| ty => throw $ ExpectedEq ty.loc ctx.names ty
whnf defs ctx $
whnf defs ctx rh $
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
whnf defs ctx rh (Ann s a annLoc) = do
Element s snf <- whnf defs ctx rh 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
Element a anf <- whnf defs ctx SZero a
pure $ Element (Ann s a annLoc) (ne `orNo` snf `orNo` anf)
whnf defs ctx (Coe sty p q val coeLoc) =
whnf defs ctx rh (Coe sty p q val coeLoc) =
-- 𝑖 ∉ fv(A)
-- -------------------------------
-- coe (𝑖 ⇒ A) @p @q s ⇝ s ∷ A
@ -148,63 +148,71 @@ CanWhnf Elim Interface.isRedexE where
([< i], Left ty) =>
case p `decEqv` q of
-- coe (𝑖 ⇒ A) @p @p s ⇝ (s ∷ Ap/𝑖)
Yes _ => whnf defs ctx $ Ann val (dsub1 sty p) coeLoc
Yes _ => whnf defs ctx rh $ Ann val (dsub1 sty p) coeLoc
No npq => do
Element ty tynf <- whnf defs (extendDim i ctx) ty
case nchoose $ canPushCoe ty val of
Left pc => pushCoe defs ctx i ty p q val coeLoc
Element ty tynf <- whnf defs (extendDim i ctx) SZero ty
case nchoose $ canPushCoe rh ty val of
Left pc => pushCoe defs ctx rh i ty p q val coeLoc
Right npc => pure $ Element (Coe (SY [< i] ty) p q val coeLoc)
(tynf `orNo` npc `orNo` notYesNo npq)
(_, Right ty) =>
whnf defs ctx $ Ann val ty coeLoc
whnf defs ctx rh $ Ann val ty coeLoc
whnf defs ctx (Comp ty p q val r zero one compLoc) =
whnf defs ctx rh (Comp ty p q val r zero one compLoc) =
case p `decEqv` q of
-- comp [A] @p @p s @r { ⋯ } ⇝ s ∷ A
Yes y => whnf defs ctx $ Ann val ty compLoc
Yes y => whnf defs ctx rh $ Ann val ty compLoc
No npq => case r of
-- comp [A] @p @q s @0 { 0 𝑗 ⇒ t₀; ⋯ } ⇝ t₀q/𝑗 ∷ A
K Zero _ => whnf defs ctx $ Ann (dsub1 zero q) ty compLoc
K Zero _ => whnf defs ctx rh $ Ann (dsub1 zero q) ty compLoc
-- comp [A] @p @q s @1 { 1 𝑗 ⇒ t₁; ⋯ } ⇝ t₁q/𝑗 ∷ A
K One _ => whnf defs ctx $ Ann (dsub1 one q) ty compLoc
K One _ => whnf defs ctx rh $ Ann (dsub1 one q) ty compLoc
B {} => pure $ Element (Comp ty p q val r zero one compLoc)
(notYesNo npq `orNo` Ah)
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 rh (TypeCase ty ret arms def tcLoc) =
case rh `decEq` SZero of
Yes Refl => do
Element ty tynf <- whnf defs ctx SZero ty
Element ret retnf <- whnf defs ctx SZero 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)
No _ =>
throw $ ClashQ tcLoc rh.qty Zero
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
whnf defs ctx rh (CloE (Sub el th)) =
whnf defs ctx rh $ pushSubstsWith' id th el
whnf defs ctx rh (DCloE (Sub el th)) =
whnf defs ctx rh $ pushSubstsWith' th id el
covering
CanWhnf Term Interface.isRedexT where
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 _ _ 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
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 _ _ _ 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
whnf defs ctx rh (E e) = do
Element e enf <- whnf defs ctx rh 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
whnf defs ctx rh (CloT (Sub tm th)) =
whnf defs ctx rh $ pushSubstsWith' id th tm
whnf defs ctx rh (DCloT (Sub tm th)) =
whnf defs ctx rh $ pushSubstsWith' th id tm

View file

@ -35,11 +35,11 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
||| for an elim returns a pair of type-cases that will reduce to that;
||| for other intro forms error
export covering
tycasePi : (t : Term d n) -> (0 tnf : No (isRedexT defs t)) =>
tycasePi : (t : Term d n) -> (0 tnf : No (isRedexT defs SZero t)) =>
Eff Whnf (Term d n, ScopeTerm d n)
tycasePi (Pi {arg, res, _}) = pure (arg, res)
tycasePi (E e) {tnf} = do
ty <- computeElimType defs ctx e {ne = noOr2 tnf}
ty <- computeElimType defs ctx SZero e {ne = noOr2 tnf}
let loc = e.loc
narg = mnb "Arg" loc; nret = mnb "Ret" loc
arg = E $ typeCase1Y e ty KPi [< !narg, !nret] (BVT 1 loc) loc
@ -53,11 +53,11 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
||| for an elim returns a pair of type-cases that will reduce to that;
||| for other intro forms error
export covering
tycaseSig : (t : Term d n) -> (0 tnf : No (isRedexT defs t)) =>
tycaseSig : (t : Term d n) -> (0 tnf : No (isRedexT defs SZero t)) =>
Eff Whnf (Term d n, ScopeTerm d n)
tycaseSig (Sig {fst, snd, _}) = pure (fst, snd)
tycaseSig (E e) {tnf} = do
ty <- computeElimType defs ctx e {ne = noOr2 tnf}
ty <- computeElimType defs ctx SZero e {ne = noOr2 tnf}
let loc = e.loc
nfst = mnb "Fst" loc; nsnd = mnb "Snd" loc
fst = E $ typeCase1Y e ty KSig [< !nfst, !nsnd] (BVT 1 loc) loc
@ -71,11 +71,11 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
||| for an elim returns a type-case that will reduce to that;
||| for other intro forms error
export covering
tycaseBOX : (t : Term d n) -> (0 tnf : No (isRedexT defs t)) =>
tycaseBOX : (t : Term d n) -> (0 tnf : No (isRedexT defs SZero t)) =>
Eff Whnf (Term d n)
tycaseBOX (BOX {ty, _}) = pure ty
tycaseBOX (E e) {tnf} = do
ty <- computeElimType defs ctx e {ne = noOr2 tnf}
ty <- computeElimType defs ctx SZero e {ne = noOr2 tnf}
pure $ E $ typeCase1Y e ty KBOX [< !(mnb "Ty" e.loc)] (BVT 0 e.loc) e.loc
tycaseBOX t = throw $ ExpectedBOX t.loc ctx.names t
@ -83,11 +83,11 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
||| for an elim returns five type-cases that will reduce to that;
||| for other intro forms error
export covering
tycaseEq : (t : Term d n) -> (0 tnf : No (isRedexT defs t)) =>
tycaseEq : (t : Term d n) -> (0 tnf : No (isRedexT defs SZero t)) =>
Eff Whnf (Term d n, Term d n, DScopeTerm d n, Term d n, Term d n)
tycaseEq (Eq {ty, l, r, _}) = pure (ty.zero, ty.one, ty, l, r)
tycaseEq (E e) {tnf} = do
ty <- computeElimType defs ctx e {ne = noOr2 tnf}
ty <- computeElimType defs ctx SZero e {ne = noOr2 tnf}
let loc = e.loc
names = traverse' (\x => mnb x loc) [< "A0", "A1", "A", "L", "R"]
a0 = E $ typeCase1Y e ty KEq !names (BVT 4 loc) loc
@ -108,11 +108,11 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
reduceTypeCase : (ty : Term d n) -> (u : Universe) -> (ret : Term d n) ->
(arms : TypeCaseArms d n) -> (def : Term d n) ->
(0 _ : So (isTyCon ty)) => Loc ->
Eff Whnf (Subset (Elim d n) (No . isRedexE defs))
Eff Whnf (Subset (Elim d n) (No . isRedexE defs SZero))
reduceTypeCase 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
whnf defs ctx SZero $ 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
@ -121,7 +121,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
res' = Ann (Lam res res.loc)
(Arr Zero arg (TYPE u noLoc) arg.loc) res.loc
in
whnf defs ctx $
whnf defs ctx SZero $
Ann (subN (tycaseRhsDef def KPi arms) [< arg', res']) ret loc
-- (type-case (x : A) × B ∷ ★ᵢ return Q of { (a × b) ⇒ s; ⋯ }) ⇝
@ -131,12 +131,12 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
snd' = Ann (Lam snd snd.loc)
(Arr Zero fst (TYPE u noLoc) fst.loc) snd.loc
in
whnf defs ctx $
whnf defs ctx SZero $
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
whnf defs ctx SZero $ Ann (tycaseRhsDef0 def KEnum arms) ret loc
-- (type-case Eq [i ⇒ A] L R ∷ ★ᵢ return Q
-- of { Eq a₀ a₁ a l r ⇒ s; ⋯ }) ⇝
@ -145,7 +145,7 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
-- (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
whnf defs ctx SZero $ Ann
(subN (tycaseRhsDef def KEq arms)
[< Ann a0 (TYPE u noLoc) a.loc, Ann a1 (TYPE u noLoc) a.loc,
Ann (DLam a a.loc) (Eq0 (TYPE u noLoc) a0 a1 a.loc) a.loc,
@ -154,10 +154,10 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
-- (type-case ∷ _ return Q of { ⇒ s; ⋯ }) ⇝ s ∷ Q
Nat {} =>
whnf defs ctx $ Ann (tycaseRhsDef0 def KNat arms) ret loc
whnf defs ctx SZero $ 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
whnf defs ctx SZero $ Ann
(sub1 (tycaseRhsDef def KBOX arms) (Ann a (TYPE u noLoc) a.loc))
ret loc