wip make qtys into polynomials
This commit is contained in:
parent
1d8a6bb325
commit
4c008577b4
22 changed files with 1650 additions and 1254 deletions
|
@ -9,140 +9,147 @@ import Quox.Whnf.TypeCase
|
|||
|
||||
|
||||
private
|
||||
coeScoped : {s : Nat} -> DScopeTerm d n -> Dim d -> Dim d -> Loc ->
|
||||
ScopeTermN s d n -> ScopeTermN s d n
|
||||
coeScoped ty p q loc (S names (N body)) =
|
||||
S names $ N $ E $ Coe ty p q body loc
|
||||
coeScoped ty p q loc (S names (Y body)) =
|
||||
ST names $ E $ Coe (weakDS s ty) p q body loc
|
||||
coeScoped : {s, q : Nat} -> DScopeTerm q d n -> Dim d -> Dim d -> Loc ->
|
||||
ScopeTermN s q d n -> ScopeTermN s q d n
|
||||
coeScoped ty p p' loc (S names (N body)) =
|
||||
S names $ N $ E $ Coe ty p p' body loc
|
||||
coeScoped ty p p' loc (S names (Y body)) =
|
||||
ST names $ E $ Coe (weakDS s ty) p p' body loc
|
||||
where
|
||||
weakDS : (by : Nat) -> DScopeTerm d n -> DScopeTerm d (by + n)
|
||||
weakDS : (by : Nat) -> DScopeTerm q d n -> DScopeTerm q 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
|
||||
|
||||
|
||||
parameters {auto _ : CanWhnf Term Interface.isRedexT}
|
||||
{auto _ : CanWhnf Elim Interface.isRedexE}
|
||||
(defs : Definitions) (ctx : WhnfContext d n) (sg : SQty)
|
||||
||| reduce a function application `App (Coe ty p q val) s loc`
|
||||
(defs : Definitions) (ctx : WhnfContext q d n) (sg : SQty)
|
||||
||| reduce a function application `App (Coe ty p p' 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 ctx sg))
|
||||
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›)
|
||||
-- where 𝒔‹j› ≔ coe [i ⇒ A] @q @j s
|
||||
piCoe : (ty : DScopeTerm q d n) -> (p, p' : Dim d) ->
|
||||
(val, s : Term q d n) -> Loc ->
|
||||
Eff Whnf (Subset (Elim q d n) (No . isRedexE defs ctx sg))
|
||||
piCoe sty@(S [< i] ty) p p' val s loc = do
|
||||
-- (coe [i ⇒ π.(x : A) → B] @p @p' t) s ⇝
|
||||
-- coe [i ⇒ B[𝒔‹i›/x] @p @p' ((t ∷ (π.(x : A) → B)‹p/i›) 𝒔‹p›)
|
||||
-- where 𝒔‹j› ≔ coe [i ⇒ A] @p' @j s
|
||||
--
|
||||
-- type-case is used to expose A,B if the type is neutral
|
||||
let ctx1 = extendDim i ctx
|
||||
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 sg $ CoeT i (sub1 res s1) p q body loc
|
||||
let Val q = ctx.qtyLen
|
||||
s0 = CoeT i arg p' 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 p') (BV 0 i.loc)
|
||||
(s // shift 1) s.loc
|
||||
whnf defs ctx sg $ CoeT i (sub1 res s1) p p' body loc
|
||||
|
||||
||| reduce a pair elimination `CasePair pi (Coe ty p q val) ret body loc`
|
||||
||| reduce a pair elimination `CasePair pi (Coe ty p p' 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 ctx sg))
|
||||
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 }
|
||||
sigCoe : (qty : Qty q) ->
|
||||
(ty : DScopeTerm q d n) -> (p, p' : Dim d) -> (val : Term q d n) ->
|
||||
(ret : ScopeTerm q d n) -> (body : ScopeTermN 2 q d n) -> Loc ->
|
||||
Eff Whnf (Subset (Elim q d n) (No . isRedexE defs ctx sg))
|
||||
sigCoe qty sty@(S [< i] ty) p p' val ret body loc = do
|
||||
-- caseπ (coe [i ⇒ (x : A) × B] @p @p' 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 ⇒ A‹j/i›] @p @i a)/x]] @p @q b)/b] }
|
||||
-- e[(coe [i ⇒ A] @p @p' a)/a,
|
||||
-- (coe [i ⇒ B[(coe [j ⇒ A‹j/i›] @p @i a)/x]] @p @p' 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 SZero $ getTerm ty
|
||||
(tfst, tsnd) <- tycaseSig defs ctx1 ty
|
||||
let [< x, y] = body.names
|
||||
a' = CoeT i (weakT 2 tfst) p q (BVT 1 x.loc) x.loc
|
||||
let Val q = ctx.qtyLen
|
||||
[< x, y] = body.names
|
||||
a' = CoeT i (weakT 2 tfst) p p' (BVT 1 x.loc) x.loc
|
||||
tsnd' = tsnd.term //
|
||||
(CoeT i (weakT 2 $ tfst // (B VZ tsnd.loc ::: shift 2))
|
||||
(weakD 1 p) (B VZ i.loc) (BVT 1 tsnd.loc) y.loc ::: shift 2)
|
||||
b' = CoeT i tsnd' p q (BVT 0 y.loc) y.loc
|
||||
b' = CoeT i tsnd' p p' (BVT 0 y.loc) y.loc
|
||||
whnf defs ctx sg $ CasePair qty (Ann val (ty // one p) val.loc) ret
|
||||
(ST body.names $ body.term // (a' ::: b' ::: shift 2)) loc
|
||||
|
||||
||| reduce a pair projection `Fst (Coe ty p q val) loc`
|
||||
||| reduce a pair projection `Fst (Coe ty p p' val) loc`
|
||||
export covering
|
||||
fstCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
|
||||
Loc -> Eff Whnf (Subset (Elim d n) (No . isRedexE defs ctx sg))
|
||||
fstCoe sty@(S [< i] ty) p q val loc = do
|
||||
-- fst (coe (𝑖 ⇒ (x : A) × B) @p @q s)
|
||||
fstCoe : (ty : DScopeTerm q d n) -> (p, p' : Dim d) -> (val : Term q d n) ->
|
||||
Loc -> Eff Whnf (Subset (Elim q d n) (No . isRedexE defs ctx sg))
|
||||
fstCoe sty@(S [< i] ty) p p' val loc = do
|
||||
-- fst (coe (𝑖 ⇒ (x : A) × B) @p @p' s)
|
||||
-- ⇝
|
||||
-- coe (𝑖 ⇒ A) @p @q (fst (s ∷ (x : A‹p/𝑖›) × B‹p/𝑖›))
|
||||
-- coe (𝑖 ⇒ A) @p @p' (fst (s ∷ (x : A‹p/𝑖›) × B‹p/𝑖›))
|
||||
--
|
||||
-- type-case is used to expose A,B if the type is neutral
|
||||
let ctx1 = extendDim i ctx
|
||||
Element ty tynf <- whnf defs ctx1 SZero $ getTerm ty
|
||||
(tfst, _) <- tycaseSig defs ctx1 ty
|
||||
let Val q = ctx.qtyLen
|
||||
whnf defs ctx sg $
|
||||
Coe (ST [< i] tfst) p q
|
||||
Coe (ST [< i] tfst) p p'
|
||||
(E (Fst (Ann val (ty // one p) val.loc) val.loc)) loc
|
||||
|
||||
||| reduce a pair projection `Snd (Coe ty p q val) loc`
|
||||
export covering
|
||||
sndCoe : (ty : DScopeTerm d n) -> (p, q : Dim d) -> (val : Term d n) ->
|
||||
Loc -> Eff Whnf (Subset (Elim d n) (No . isRedexE defs ctx sg))
|
||||
sndCoe sty@(S [< i] ty) p q val loc = do
|
||||
-- snd (coe (𝑖 ⇒ (x : A) × B) @p @q s)
|
||||
sndCoe : (ty : DScopeTerm q d n) -> (p, p' : Dim d) -> (val : Term q d n) ->
|
||||
Loc -> Eff Whnf (Subset (Elim q d n) (No . isRedexE defs ctx sg))
|
||||
sndCoe sty@(S [< i] ty) p p' val loc = do
|
||||
-- snd (coe (𝑖 ⇒ (x : A) × B) @p @p' s)
|
||||
-- ⇝
|
||||
-- coe (𝑖 ⇒ B[coe (𝑗 ⇒ A‹𝑗/𝑖›) @p @𝑖 (fst (s ∷ (x : A) × B))/x]) @p @q
|
||||
-- coe (𝑖 ⇒ B[coe (𝑗 ⇒ A‹𝑗/𝑖›) @p @𝑖 (fst (s ∷ (x : A) × B))/x]) @p @p'
|
||||
-- (snd (s ∷ (x : A‹p/𝑖›) × B‹p/𝑖›))
|
||||
--
|
||||
-- type-case is used to expose A,B if the type is neutral
|
||||
let ctx1 = extendDim i ctx
|
||||
Element ty tynf <- whnf defs ctx1 SZero $ getTerm ty
|
||||
(tfst, tsnd) <- tycaseSig defs ctx1 ty
|
||||
let Val q = ctx.qtyLen
|
||||
whnf defs ctx sg $
|
||||
Coe (ST [< i] $ sub1 tsnd $
|
||||
Coe (ST [< !(fresh i)] $ tfst // (BV 0 i.loc ::: shift 2))
|
||||
(weakD 1 p) (BV 0 loc)
|
||||
(E (Fst (Ann (dweakT 1 val) ty val.loc) val.loc)) loc)
|
||||
p q
|
||||
p p'
|
||||
(E (Snd (Ann val (ty // one p) val.loc) val.loc))
|
||||
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) ->
|
||||
eqCoe : (ty : DScopeTerm q d n) -> (p, p' : Dim d) -> (val : Term q d n) ->
|
||||
(r : Dim d) -> Loc ->
|
||||
Eff Whnf (Subset (Elim d n) (No . isRedexE defs ctx sg))
|
||||
eqCoe sty@(S [< j] ty) p q val r loc = do
|
||||
-- (coe [j ⇒ Eq [i ⇒ A] L R] @p @q eq) @r
|
||||
Eff Whnf (Subset (Elim q d n) (No . isRedexE defs ctx sg))
|
||||
eqCoe sty@(S [< j] ty) p p' val r loc = do
|
||||
-- (coe [j ⇒ Eq [i ⇒ A] L R] @p @p' eq) @r
|
||||
-- ⇝
|
||||
-- comp [j ⇒ A‹r/i›] @p @q ((eq ∷ (Eq [i ⇒ A] L R)‹p/j›) @r)
|
||||
-- comp [j ⇒ A‹r/i›] @p @p' ((eq ∷ (Eq [i ⇒ A] L R)‹p/j›) @r)
|
||||
-- @r { 0 j ⇒ L; 1 j ⇒ R }
|
||||
let ctx1 = extendDim j ctx
|
||||
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 sg $ CompH j a' p q val' r j s j t loc
|
||||
let Val q = ctx.qtyLen
|
||||
a' = dsub1 a (weakD 1 r)
|
||||
val' = E $ DApp (Ann val (ty // one p) val.loc) r loc
|
||||
whnf defs ctx sg $ CompH j a' p p' 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 ctx sg))
|
||||
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 }
|
||||
boxCoe : (qty : Qty q) ->
|
||||
(ty : DScopeTerm q d n) -> (p, p' : Dim d) -> (val : Term q d n) ->
|
||||
(ret : ScopeTerm q d n) -> (body : ScopeTerm q d n) -> Loc ->
|
||||
Eff Whnf (Subset (Elim q d n) (No . isRedexE defs ctx sg))
|
||||
boxCoe qty sty@(S [< i] ty) p p' val ret body loc = do
|
||||
-- caseπ (coe [i ⇒ [ρ. A]] @p @p' 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
|
||||
-- caseπ s ∷ [ρ. A]‹p/i› return z ⇒ C of { [a] ⇒ e[(coe [i ⇒ A] p p' a)/a] }
|
||||
let ctx1 = extendDim i ctx
|
||||
Val q = ctx.qtyLen
|
||||
Element ty tynf <- whnf defs ctx1 SZero $ getTerm ty
|
||||
ta <- tycaseBOX defs ctx1 ty
|
||||
let xloc = body.name.loc
|
||||
let a' = CoeT i (weakT 1 ta) p q (BVT 0 xloc) xloc
|
||||
let a' = CoeT i (weakT 1 ta) p p' (BVT 0 xloc) xloc
|
||||
whnf defs ctx sg $ CaseBox qty (Ann val (ty // one p) val.loc) ret
|
||||
(ST body.names $ body.term // (a' ::: shift 1)) loc
|
||||
|
||||
|
@ -150,14 +157,14 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
|
|||
-- new params block to call the above functions at different `n`
|
||||
parameters {auto _ : CanWhnf Term Interface.isRedexT}
|
||||
{auto _ : CanWhnf Elim Interface.isRedexE}
|
||||
(defs : Definitions) (ctx : WhnfContext d n) (sg : SQty)
|
||||
(defs : Definitions) (ctx : WhnfContext q d n) (sg : SQty)
|
||||
||| pushes a coercion inside a whnf-ed term
|
||||
export covering
|
||||
pushCoe : BindName ->
|
||||
(ty : Term (S d) n) -> (p, q : Dim d) -> (s : Term d n) -> Loc ->
|
||||
(0 pc : So (canPushCoe sg ty s)) =>
|
||||
Eff Whnf (NonRedex Elim d n defs ctx sg)
|
||||
pushCoe i ty p q s loc =
|
||||
(ty : Term q (S d) n) -> (p, p' : Dim d) -> (s : Term q d n) ->
|
||||
Loc -> (0 pc : So (canPushCoe sg ty s)) =>
|
||||
Eff Whnf (NonRedex Elim q d n defs ctx sg)
|
||||
pushCoe i ty p p' s loc =
|
||||
case ty of
|
||||
-- (coe ★ᵢ @_ @_ s) ⇝ (s ∷ ★ᵢ)
|
||||
TYPE l tyLoc =>
|
||||
|
@ -169,40 +176,40 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
|
|||
|
||||
-- η expand, then simplify the Coe/App in the body
|
||||
--
|
||||
-- (coe (𝑖 ⇒ π.(x : A) → B) @p @q s)
|
||||
-- (coe (𝑖 ⇒ π.(x : A) → B) @p @p' s)
|
||||
-- ⇝
|
||||
-- (λ y ⇒ (coe (𝑖 ⇒ π.(x : A) → B) @p @q s) y) ∷ (π.(x : A) → B)‹q/𝑖›
|
||||
-- ⇝ ‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾
|
||||
-- (λ y ⇒ ⋯) ∷ (π.(x : A) → B)‹q/𝑖› -- see `piCoe`
|
||||
-- (λ y ⇒ (coe (𝑖 ⇒ π.(x : A) → B) @p @p' s) y) ∷ (π.(x : A) → B)‹p'/𝑖›
|
||||
-- ⇝ ‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾
|
||||
-- (λ y ⇒ ⋯) ∷ (π.(x : A) → B)‹p'/𝑖› -- see `piCoe`
|
||||
--
|
||||
-- do the piCoe step here because otherwise equality checking keeps
|
||||
-- doing the η forever
|
||||
Pi {arg, res = S [< x] _, _} => do
|
||||
let ctx' = extendTy x (arg // one p) ctx
|
||||
body <- piCoe defs ctx' sg
|
||||
(weakDS 1 $ SY [< i] ty) p q (weakT 1 s) (BVT 0 loc) loc
|
||||
(weakDS 1 $ SY [< i] ty) p p' (weakT 1 s) (BVT 0 loc) loc
|
||||
whnf defs ctx sg $
|
||||
Ann (LamY x (E body.fst) loc) (ty // one q) loc
|
||||
Ann (LamY x (E body.fst) loc) (ty // one p') loc
|
||||
|
||||
-- no η!!!
|
||||
-- push into a pair constructor, otherwise still stuck
|
||||
--
|
||||
-- s̃‹𝑘› ≔ coe (𝑗 ⇒ A‹𝑗/𝑖›) @p @𝑘 s
|
||||
-- -----------------------------------------------
|
||||
-- (coe (𝑖 ⇒ (x : A) × B) @p @q (s, t))
|
||||
-- (coe (𝑖 ⇒ (x : A) × B) @p @p' (s, t))
|
||||
-- ⇝
|
||||
-- (s̃‹q›, coe (𝑖 ⇒ B[s̃‹𝑖›/x]) @p @q t)
|
||||
-- ∷ ((x : A) × B)‹q/𝑖›
|
||||
-- (s̃‹p'›, coe (𝑖 ⇒ B[s̃‹𝑖›/x]) @p @p' t)
|
||||
-- ∷ ((x : A) × B)‹p'/𝑖›
|
||||
Sig tfst tsnd tyLoc => do
|
||||
let Pair fst snd sLoc = s
|
||||
fst' = CoeT i tfst p q fst fst.loc
|
||||
let Val q = ctx.qtyLen
|
||||
Pair fst snd sLoc = s
|
||||
fst' = CoeT i tfst p p' fst fst.loc
|
||||
fstInSnd =
|
||||
CoeT !(fresh i)
|
||||
(tfst // (BV 0 loc ::: shift 2))
|
||||
(weakD 1 p) (BV 0 loc) (dweakT 1 fst) fst.loc
|
||||
snd' = CoeT i (sub1 tsnd fstInSnd) p q snd snd.loc
|
||||
snd' = CoeT i (sub1 tsnd fstInSnd) p p' snd snd.loc
|
||||
whnf defs ctx sg $
|
||||
Ann (Pair (E fst') (E snd') sLoc) (ty // one q) loc
|
||||
Ann (Pair (E fst') (E snd') sLoc) (ty // one p') loc
|
||||
|
||||
-- (coe {𝐚̄} @_ @_ s) ⇝ (s ∷ {𝐚̄})
|
||||
Enum cases tyLoc =>
|
||||
|
@ -210,21 +217,21 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
|
|||
|
||||
-- η expand/simplify, same as for Π
|
||||
--
|
||||
-- (coe (𝑖 ⇒ Eq (𝑗 ⇒ A) l r) @p @q s)
|
||||
-- (coe (𝑖 ⇒ Eq (𝑗 ⇒ A) l r) @p @p' s)
|
||||
-- ⇝
|
||||
-- (δ 𝑘 ⇒ (coe (𝑖 ⇒ Eq (𝑗 ⇒ A) l r) @p @q s) @𝑘) ∷ (Eq (𝑗 ⇒ A) l r)‹q/𝑖›
|
||||
-- ⇝ ‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾
|
||||
-- (δ 𝑘 ⇒ ⋯) ∷ (Eq (𝑗 ⇒ A) l r)‹q/𝑖› -- see `eqCoe`
|
||||
-- (δ 𝑘 ⇒ (coe (𝑖 ⇒ Eq (𝑗 ⇒ A) l r) @p @p' s) @𝑘) ∷ (Eq (𝑗 ⇒ A) l r)‹p'/𝑖›
|
||||
-- ⇝ ‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾‾
|
||||
-- (δ 𝑘 ⇒ ⋯) ∷ (Eq (𝑗 ⇒ A) l r)‹p'/𝑖› -- see `eqCoe`
|
||||
--
|
||||
-- do the eqCoe step here because otherwise equality checking keeps
|
||||
-- doing the η forever
|
||||
Eq {ty = S [< j] _, _} => do
|
||||
let ctx' = extendDim j ctx
|
||||
body <- eqCoe defs ctx' sg
|
||||
(dweakDS 1 $ S [< i] $ Y ty) (weakD 1 p) (weakD 1 q)
|
||||
(dweakDS 1 $ S [< i] $ Y ty) (weakD 1 p) (weakD 1 p')
|
||||
(dweakT 1 s) (BV 0 loc) loc
|
||||
whnf defs ctx sg $
|
||||
Ann (DLamY i (E body.fst) loc) (ty // one q) loc
|
||||
Ann (DLamY i (E body.fst) loc) (ty // one p') loc
|
||||
|
||||
-- (coe ℕ @_ @_ s) ⇝ (s ∷ ℕ)
|
||||
NAT tyLoc =>
|
||||
|
@ -236,17 +243,17 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
|
|||
|
||||
-- η expand/simplify
|
||||
--
|
||||
-- (coe (𝑖 ⇒ [π.A]) @p @q s)
|
||||
-- (coe (𝑖 ⇒ [π.A]) @p @p' s)
|
||||
-- ⇝
|
||||
-- [case coe (𝑖 ⇒ [π.A]) @p @q s return A‹q/𝑖› of {[x] ⇒ x}]
|
||||
-- [case coe (𝑖 ⇒ [π.A]) @p @p' s return A‹p'/𝑖› of {[x] ⇒ x}]
|
||||
-- ⇝
|
||||
-- [case1 s ∷ [π.A]‹p/𝑖› ⋯] ∷ [π.A]‹q/𝑖› -- see `boxCoe`
|
||||
-- [case1 s ∷ [π.A]‹p/𝑖› ⋯] ∷ [π.A]‹p'/𝑖› -- see `boxCoe`
|
||||
--
|
||||
-- do the eqCoe step here because otherwise equality checking keeps
|
||||
-- do the boxCoe step here because otherwise equality checking keeps
|
||||
-- doing the η forever
|
||||
BOX qty inner tyLoc => do
|
||||
body <- boxCoe defs ctx sg qty
|
||||
(SY [< i] ty) p q s
|
||||
(SN $ inner // one q)
|
||||
(SY [< i] ty) p p' s
|
||||
(SN $ inner // one p')
|
||||
(SY [< !(mnb "inner" loc)] (BVT 0 loc)) loc
|
||||
whnf defs ctx sg $ Ann (Box (E body.fst) loc) (ty // one q) loc
|
||||
whnf defs ctx sg $ Ann (Box (E body.fst) loc) (ty // one p') loc
|
||||
|
|
|
@ -15,46 +15,44 @@ export covering
|
|||
computeElimType :
|
||||
CanWhnf Term Interface.isRedexT =>
|
||||
CanWhnf Elim Interface.isRedexE =>
|
||||
(defs : Definitions) -> (ctx : WhnfContext d n) -> (0 sg : SQty) ->
|
||||
(e : Elim d n) -> (0 ne : No (isRedexE defs ctx sg e)) =>
|
||||
Eff Whnf (Term d n)
|
||||
(defs : Definitions) -> (ctx : WhnfContext q d n) -> (0 sg : SQty) ->
|
||||
(e : Elim q d n) -> (0 ne : No (isRedexE defs ctx sg e)) =>
|
||||
Eff Whnf (Term q d n)
|
||||
|
||||
||| computes a type and then reduces it to whnf
|
||||
export covering
|
||||
computeWhnfElimType0 :
|
||||
CanWhnf Term Interface.isRedexT =>
|
||||
CanWhnf Elim Interface.isRedexE =>
|
||||
(defs : Definitions) -> (ctx : WhnfContext d n) -> (0 sg : SQty) ->
|
||||
(e : Elim d n) -> (0 ne : No (isRedexE defs ctx sg e)) =>
|
||||
Eff Whnf (Term d n)
|
||||
(defs : Definitions) -> (ctx : WhnfContext q d n) -> (0 sg : SQty) ->
|
||||
(e : Elim q d n) -> (0 ne : No (isRedexE defs ctx sg e)) =>
|
||||
Eff Whnf (Term q d n)
|
||||
|
||||
|
||||
private covering
|
||||
computeElimTypeNoLog, computeWhnfElimType0NoLog :
|
||||
CanWhnf Term Interface.isRedexT =>
|
||||
CanWhnf Elim Interface.isRedexE =>
|
||||
(defs : Definitions) -> WhnfContext d n -> (0 sg : SQty) ->
|
||||
(e : Elim d n) -> (0 ne : No (isRedexE defs ctx sg e)) =>
|
||||
Eff Whnf (Term d n)
|
||||
(defs : Definitions) -> WhnfContext q d n -> (0 sg : SQty) ->
|
||||
(e : Elim q d n) -> (0 ne : No (isRedexE defs ctx sg e)) =>
|
||||
Eff Whnf (Term q d n)
|
||||
|
||||
computeElimTypeNoLog defs ctx sg e =
|
||||
case e of
|
||||
F x u loc => do
|
||||
let Just def = lookup x defs
|
||||
| Nothing => throw $ NotInScope loc x
|
||||
pure $ def.typeWithAt ctx.dimLen ctx.termLen u
|
||||
let polyTy = def.typeWithAt ctx.dimLen ctx.termLen u
|
||||
Val q = ctx.qtyLen
|
||||
pure $ polyTy.type // ?freeQSubst
|
||||
|
||||
B i _ =>
|
||||
pure (ctx.tctx !! i).type
|
||||
B i _ => pure (ctx.tctx !! i).type
|
||||
|
||||
App f s loc =>
|
||||
case !(computeWhnfElimType0NoLog defs ctx sg f {ne = noOr1 ne}) of
|
||||
Pi {arg, res, _} => pure $ sub1 res $ Ann s arg loc
|
||||
ty => throw $ ExpectedPi loc ctx.names ty
|
||||
|
||||
CasePair {pair, ret, _} =>
|
||||
pure $ sub1 ret pair
|
||||
|
||||
Fst pair loc =>
|
||||
case !(computeWhnfElimType0NoLog defs ctx sg pair {ne = noOr1 ne}) of
|
||||
Sig {fst, _} => pure fst
|
||||
|
@ -65,46 +63,34 @@ computeElimTypeNoLog defs ctx sg e =
|
|||
Sig {snd, _} => pure $ sub1 snd $ Fst pair loc
|
||||
ty => throw $ ExpectedSig loc ctx.names ty
|
||||
|
||||
CaseEnum {tag, ret, _} =>
|
||||
pure $ sub1 ret tag
|
||||
|
||||
CaseNat {nat, ret, _} =>
|
||||
pure $ sub1 ret nat
|
||||
|
||||
CaseBox {box, ret, _} =>
|
||||
pure $ sub1 ret box
|
||||
CasePair {pair, ret, _} => pure $ sub1 ret pair
|
||||
CaseEnum {tag, ret, _} => pure $ sub1 ret tag
|
||||
CaseNat {nat, ret, _} => pure $ sub1 ret nat
|
||||
CaseBox {box, ret, _} => pure $ sub1 ret box
|
||||
|
||||
DApp {fun = f, arg = p, loc} =>
|
||||
case !(computeWhnfElimType0NoLog defs ctx sg f {ne = noOr1 ne}) of
|
||||
Eq {ty, _} => pure $ dsub1 ty p
|
||||
t => throw $ ExpectedEq loc ctx.names t
|
||||
|
||||
Ann {ty, _} =>
|
||||
pure ty
|
||||
|
||||
Coe {ty, q, _} =>
|
||||
pure $ dsub1 ty q
|
||||
|
||||
Comp {ty, _} =>
|
||||
pure ty
|
||||
|
||||
TypeCase {ret, _} =>
|
||||
pure ret
|
||||
Ann {ty, _} => pure ty
|
||||
Coe {ty, p', _} => pure $ dsub1 ty p'
|
||||
Comp {ty, _} => pure ty
|
||||
TypeCase {ret, _} => pure ret
|
||||
|
||||
computeElimType defs ctx sg e {ne} = do
|
||||
let Val n = ctx.termLen
|
||||
sayMany "whnf" e.loc
|
||||
[90 :> "computeElimType",
|
||||
95 :> hsep ["ctx =", runPretty $ prettyWhnfContext ctx],
|
||||
90 :> hsep ["e =", runPretty $ prettyElim ctx.dnames ctx.tnames e]]
|
||||
90 :> hsep ["e =", runPretty $ prettyElim ctx.names e]]
|
||||
res <- computeElimTypeNoLog defs ctx sg e {ne}
|
||||
say "whnf" 91 e.loc $
|
||||
hsep ["computeElimType ⇝",
|
||||
runPretty $ prettyTerm ctx.dnames ctx.tnames res]
|
||||
hsep ["computeElimType ⇝", runPretty $ prettyTerm ctx.names res]
|
||||
pure res
|
||||
|
||||
computeWhnfElimType0 defs ctx sg e =
|
||||
computeElimType defs ctx sg e >>= whnf0 defs ctx SZero
|
||||
whnf0 defs ctx SZero !(computeElimType defs ctx sg e)
|
||||
|
||||
computeWhnfElimType0NoLog defs ctx sg e {ne} =
|
||||
computeElimTypeNoLog defs ctx sg e {ne} >>= whnf0 defs ctx SZero
|
||||
whnf0 defs ctx SZero !(computeElimTypeNoLog defs ctx sg e {ne})
|
||||
|
|
|
@ -20,14 +20,15 @@ Whnf = [Except Error, NameGen, Log]
|
|||
public export
|
||||
0 RedexTest : TermLike -> Type
|
||||
RedexTest tm =
|
||||
{0 d, n : Nat} -> Definitions -> WhnfContext d n -> SQty -> tm d n -> Bool
|
||||
{0 q, d, n : Nat} ->
|
||||
Definitions -> WhnfContext q d n -> SQty -> tm q d n -> Bool
|
||||
|
||||
public export
|
||||
interface CanWhnf (0 tm : TermLike) (0 isRedex : RedexTest tm) | tm
|
||||
where
|
||||
whnf, whnfNoLog :
|
||||
(defs : Definitions) -> (ctx : WhnfContext d n) -> (q : SQty) ->
|
||||
tm d n -> Eff Whnf (Subset (tm d n) (No . isRedex defs ctx q))
|
||||
(defs : Definitions) -> (ctx : WhnfContext q d n) -> (sg : SQty) ->
|
||||
tm q d n -> Eff Whnf (Subset (tm q d n) (No . isRedex defs ctx sg))
|
||||
-- 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
|
||||
|
@ -36,26 +37,27 @@ where
|
|||
public export %inline
|
||||
whnf0, whnfNoLog0 :
|
||||
{0 isRedex : RedexTest tm} -> CanWhnf tm isRedex =>
|
||||
Definitions -> WhnfContext d n -> SQty -> tm d n -> Eff Whnf (tm d n)
|
||||
Definitions -> WhnfContext q d n -> SQty -> tm q d n -> Eff Whnf (tm q d n)
|
||||
whnf0 defs ctx q t = fst <$> whnf defs ctx q t
|
||||
whnfNoLog0 defs ctx q t = fst <$> whnfNoLog defs ctx q t
|
||||
|
||||
public export
|
||||
0 IsRedex, NotRedex : {isRedex : RedexTest tm} -> CanWhnf tm isRedex =>
|
||||
Definitions -> WhnfContext d n -> SQty -> Pred (tm d n)
|
||||
Definitions -> WhnfContext q d n -> SQty ->
|
||||
Pred (tm q d n)
|
||||
IsRedex defs ctx q = So . isRedex defs ctx q
|
||||
NotRedex defs ctx q = No . isRedex defs ctx q
|
||||
|
||||
public export
|
||||
0 NonRedex : (tm : TermLike) -> {isRedex : RedexTest tm} ->
|
||||
CanWhnf tm isRedex => (d, n : Nat) ->
|
||||
Definitions -> WhnfContext d n -> SQty -> Type
|
||||
NonRedex tm d n defs ctx q = Subset (tm d n) (NotRedex defs ctx q)
|
||||
CanWhnf tm isRedex => (q, d, n : Nat) ->
|
||||
Definitions -> WhnfContext q d n -> SQty -> Type
|
||||
NonRedex tm q d n defs ctx sg = Subset (tm q d n) (NotRedex defs ctx sg)
|
||||
|
||||
public export %inline
|
||||
nred : {0 isRedex : RedexTest tm} -> (0 _ : CanWhnf tm isRedex) =>
|
||||
(t : tm d n) -> (0 nr : NotRedex defs ctx q t) =>
|
||||
NonRedex tm d n defs ctx q
|
||||
(t : tm q d n) -> (0 nr : NotRedex defs ctx sg t) =>
|
||||
NonRedex tm q d n defs ctx sg
|
||||
nred t = Element t nr
|
||||
|
||||
|
||||
|
@ -97,7 +99,7 @@ isNatHead _ = False
|
|||
|
||||
||| a natural constant, with or without an annotation
|
||||
public export %inline
|
||||
isNatConst : Term d n -> Bool
|
||||
isNatConst : Term {} -> Bool
|
||||
isNatConst (Nat {}) = True
|
||||
isNatConst (E (Ann (Nat {}) _ _)) = True
|
||||
isNatConst _ = False
|
||||
|
@ -145,6 +147,7 @@ isTyCon (Let {}) = False
|
|||
isTyCon (E {}) = False
|
||||
isTyCon (CloT {}) = False
|
||||
isTyCon (DCloT {}) = False
|
||||
isTyCon (QCloT {}) = False
|
||||
|
||||
||| a syntactic type, or a neutral
|
||||
public export %inline
|
||||
|
@ -195,6 +198,7 @@ canPushCoe sg (Let {}) _ = False
|
|||
canPushCoe sg (E {}) _ = False
|
||||
canPushCoe sg (CloT {}) _ = False
|
||||
canPushCoe sg (DCloT {}) _ = False
|
||||
canPushCoe sg (QCloT {}) _ = False
|
||||
|
||||
|
||||
mutual
|
||||
|
@ -238,15 +242,16 @@ mutual
|
|||
isRedexE defs ctx sg (Ann {tm, ty, _}) =
|
||||
isE tm || isRedexT defs ctx sg tm || isRedexT defs ctx SZero ty
|
||||
isRedexE defs ctx sg (Coe {ty = S _ (N _), _}) = True
|
||||
isRedexE defs ctx sg (Coe {ty = S [< i] (Y ty), p, q, val, _}) =
|
||||
isRedexE defs ctx sg (Coe {ty = S [< i] (Y ty), p, p', val, _}) =
|
||||
isRedexT defs (extendDim i ctx) SZero ty ||
|
||||
canPushCoe sg ty val || isYes (p `decEqv` q)
|
||||
isRedexE defs ctx sg (Comp {ty, p, q, r, _}) =
|
||||
isYes (p `decEqv` q) || isK r
|
||||
canPushCoe sg ty val || isYes (p `decEqv` p')
|
||||
isRedexE defs ctx sg (Comp {ty, p, p', r, _}) =
|
||||
isYes (p `decEqv` p') || isK r
|
||||
isRedexE defs ctx sg (TypeCase {ty, ret, _}) =
|
||||
isRedexE defs ctx sg ty || isRedexT defs ctx sg ret || isAnnTyCon ty
|
||||
isRedexE _ _ _ (CloE {}) = True
|
||||
isRedexE _ _ _ (DCloE {}) = True
|
||||
isRedexE _ _ _ (QCloE {}) = True
|
||||
|
||||
||| a reducible term
|
||||
|||
|
||||
|
@ -260,6 +265,7 @@ mutual
|
|||
isRedexT : RedexTest Term
|
||||
isRedexT _ _ _ (CloT {}) = True
|
||||
isRedexT _ _ _ (DCloT {}) = True
|
||||
isRedexT _ _ _ (QCloT {}) = True
|
||||
isRedexT _ _ _ (Let {}) = True
|
||||
isRedexT defs ctx sg (E {e, _}) = isAnn e || isRedexE defs ctx sg e
|
||||
isRedexT _ _ _ (Succ p {}) = isNatConst p
|
||||
|
|
|
@ -19,19 +19,19 @@ export covering CanWhnf Elim Interface.isRedexE
|
|||
private %inline
|
||||
whnfDefault :
|
||||
{0 isRedex : RedexTest tm} ->
|
||||
(CanWhnf tm isRedex, Located2 tm) =>
|
||||
(CanWhnf tm isRedex, Located3 tm) =>
|
||||
String ->
|
||||
(forall d, n. WhnfContext d n -> tm d n -> Eff Pretty LogDoc) ->
|
||||
(forall d, n. WhnfContext q d n -> tm q d n -> Eff Pretty LogDoc) ->
|
||||
(defs : Definitions) ->
|
||||
(ctx : WhnfContext d n) ->
|
||||
(ctx : WhnfContext q d n) ->
|
||||
(sg : SQty) ->
|
||||
(s : tm d n) ->
|
||||
Eff Whnf (Subset (tm d n) (No . isRedex defs ctx sg))
|
||||
(s : tm q d n) ->
|
||||
Eff Whnf (Subset (tm q d n) (No . isRedex defs ctx sg))
|
||||
whnfDefault name ppr defs ctx sg s = do
|
||||
sayMany "whnf" s.loc
|
||||
[10 :> "whnf",
|
||||
95 :> hsep ["ctx =", runPretty $ prettyWhnfContext ctx],
|
||||
95 :> hsep ["sg =", runPretty $ prettyQty sg.qty],
|
||||
95 :> hsep ["sg =", runPretty $ prettyQConst sg.qconst],
|
||||
10 :> hsep [text name, "=", runPretty $ ppr ctx s]]
|
||||
res <- whnfNoLog defs ctx sg s
|
||||
say "whnf" 11 s.loc $ hsep ["whnf ⇝", runPretty $ ppr ctx res.fst]
|
||||
|
@ -39,10 +39,10 @@ whnfDefault name ppr defs ctx sg s = do
|
|||
|
||||
covering
|
||||
CanWhnf Elim Interface.isRedexE where
|
||||
whnf = whnfDefault "e" $ \ctx, e => prettyElim ctx.dnames ctx.tnames e
|
||||
whnf = whnfDefault "e" $ \ctx, e => prettyElim ctx.names e
|
||||
|
||||
whnfNoLog defs ctx sg (F x u loc) with (lookupElim0 x u defs) proof eq
|
||||
_ | Just y = whnf defs ctx sg $ setLoc loc $ injElim ctx y
|
||||
_ | Just y = whnf defs ctx sg $ setLoc loc $ ?whnfFreeQSubst
|
||||
_ | Nothing = pure $ Element (F x u loc) $ rewrite eq in Ah
|
||||
|
||||
whnfNoLog defs ctx sg (B i loc) with (ctx.tctx !! i) proof eq1
|
||||
|
@ -186,8 +186,8 @@ CanWhnf Elim Interface.isRedexE where
|
|||
whnf defs ctx sg $
|
||||
Ann (endsOr (setLoc appLoc l) (setLoc appLoc r) (dsub1 body p) p)
|
||||
(dsub1 ty p) appLoc
|
||||
Coe ty p' q' val _ =>
|
||||
eqCoe defs ctx sg ty p' q' val p appLoc
|
||||
Coe ty r r' val _ =>
|
||||
eqCoe defs ctx sg ty r r' val p appLoc
|
||||
Right ndlh => case p of
|
||||
K e _ => do
|
||||
Eq {l, r, ty, _} <- computeWhnfElimType0 defs ctx sg f
|
||||
|
@ -206,36 +206,37 @@ CanWhnf Elim Interface.isRedexE where
|
|||
Element a anf <- whnf defs ctx SZero a
|
||||
pure $ Element (Ann s a annLoc) (ne `orNo` snf `orNo` anf)
|
||||
|
||||
whnfNoLog defs ctx sg (Coe sty p q val coeLoc) =
|
||||
whnfNoLog defs ctx sg (Coe sty p p' val coeLoc) = do
|
||||
-- 𝑖 ∉ fv(A)
|
||||
-- -------------------------------
|
||||
-- coe (𝑖 ⇒ A) @p @q s ⇝ s ∷ A
|
||||
-- coe (𝑖 ⇒ A) @p @p' s ⇝ s ∷ A
|
||||
--
|
||||
-- [fixme] needs a real equality check between A‹0/𝑖› and A‹1/𝑖›
|
||||
case dsqueeze sty {f = Term} of
|
||||
let Val q = ctx.qtyLen
|
||||
case dsqueeze sty {f = Term q} of
|
||||
([< i], Left ty) =>
|
||||
case p `decEqv` q of
|
||||
case p `decEqv` p' of
|
||||
-- coe (𝑖 ⇒ A) @p @p s ⇝ (s ∷ A‹p/𝑖›)
|
||||
Yes _ => whnf defs ctx sg $ Ann val (dsub1 sty p) coeLoc
|
||||
No npq => do
|
||||
Element ty tynf <- whnf defs (extendDim i ctx) SZero ty
|
||||
case nchoose $ canPushCoe sg ty val of
|
||||
Left pc => pushCoe defs ctx sg i ty p q val coeLoc
|
||||
Right npc => pure $ Element (Coe (SY [< i] ty) p q val coeLoc)
|
||||
Left pc => pushCoe defs ctx sg i ty p p' val coeLoc
|
||||
Right npc => pure $ Element (Coe (SY [< i] ty) p p' val coeLoc)
|
||||
(tynf `orNo` npc `orNo` notYesNo npq)
|
||||
(_, Right ty) =>
|
||||
whnf defs ctx sg $ Ann val ty coeLoc
|
||||
|
||||
whnfNoLog defs ctx sg (Comp ty p q val r zero one compLoc) =
|
||||
case p `decEqv` q of
|
||||
whnfNoLog defs ctx sg (Comp ty p p' val r zero one compLoc) =
|
||||
case p `decEqv` p' of
|
||||
-- comp [A] @p @p s @r { ⋯ } ⇝ s ∷ A
|
||||
Yes y => whnf defs ctx sg $ 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 sg $ Ann (dsub1 zero q) ty compLoc
|
||||
-- comp [A] @p @q s @1 { 1 𝑗 ⇒ t₁; ⋯ } ⇝ t₁‹q/𝑗› ∷ A
|
||||
K One _ => whnf defs ctx sg $ Ann (dsub1 one q) ty compLoc
|
||||
B {} => pure $ Element (Comp ty p q val r zero one compLoc)
|
||||
-- comp [A] @p @p' s @0 { 0 𝑗 ⇒ t₀; ⋯ } ⇝ t₀‹p'/𝑗› ∷ A
|
||||
K Zero _ => whnf defs ctx sg $ Ann (dsub1 zero p') ty compLoc
|
||||
-- comp [A] @p @p' s @1 { 1 𝑗 ⇒ t₁; ⋯ } ⇝ t₁‹p'/𝑗› ∷ A
|
||||
K One _ => whnf defs ctx sg $ Ann (dsub1 one p') ty compLoc
|
||||
B {} => pure $ Element (Comp ty p p' val r zero one compLoc)
|
||||
(notYesNo npq `orNo` Ah)
|
||||
|
||||
whnfNoLog defs ctx sg (TypeCase ty ret arms def tcLoc) =
|
||||
|
@ -248,17 +249,23 @@ CanWhnf Elim Interface.isRedexE where
|
|||
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 sg.qty Zero
|
||||
No _ => do
|
||||
let Val q = ctx.qtyLen
|
||||
throw $ ClashQ tcLoc ctx.qnames (toQty tcLoc sg) (zero tcLoc)
|
||||
|
||||
whnfNoLog defs ctx sg (CloE (Sub el th)) =
|
||||
whnfNoLog defs ctx sg $ pushSubstsWith' id th el
|
||||
whnfNoLog defs ctx sg (DCloE (Sub el th)) =
|
||||
whnfNoLog defs ctx sg $ pushSubstsWith' th id el
|
||||
whnfNoLog defs ctx sg (CloE (Sub el th)) = do
|
||||
let Val q = ctx.qtyLen
|
||||
whnfNoLog defs ctx sg $ pushSubstsWith' id id th el
|
||||
whnfNoLog defs ctx sg (DCloE (Sub el th)) = do
|
||||
let Val q = ctx.qtyLen
|
||||
whnfNoLog defs ctx sg $ pushSubstsWith' id th id el
|
||||
whnfNoLog defs ctx sg (QCloE (SubR el th)) = do
|
||||
let Val q = ctx.qtyLen
|
||||
whnfNoLog defs ctx sg $ pushSubstsWith' th id id el
|
||||
|
||||
covering
|
||||
CanWhnf Term Interface.isRedexT where
|
||||
whnf = whnfDefault "e" $ \ctx, s => prettyTerm ctx.dnames ctx.tnames s
|
||||
whnf = whnfDefault "e" $ \ctx, s => prettyTerm ctx.names s
|
||||
|
||||
whnfNoLog _ _ _ t@(TYPE {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(IOState {}) = pure $ nred t
|
||||
|
@ -294,7 +301,14 @@ CanWhnf Term Interface.isRedexT where
|
|||
Left _ => let Ann {tm, _} = e in pure $ Element tm $ noOr1 $ noOr2 enf
|
||||
Right na => pure $ Element (E e) $ na `orNo` enf
|
||||
|
||||
whnfNoLog defs ctx sg (CloT (Sub tm th)) =
|
||||
whnfNoLog defs ctx sg $ pushSubstsWith' id th tm
|
||||
whnfNoLog defs ctx sg (DCloT (Sub tm th)) =
|
||||
whnfNoLog defs ctx sg $ pushSubstsWith' th id tm
|
||||
whnfNoLog defs ctx sg (CloT (Sub tm th)) = do
|
||||
let Val q = ctx.qtyLen
|
||||
whnfNoLog defs ctx sg $ pushSubstsWith' id id th tm
|
||||
|
||||
whnfNoLog defs ctx sg (DCloT (Sub tm th)) = do
|
||||
let Val q = ctx.qtyLen
|
||||
whnfNoLog defs ctx sg $ pushSubstsWith' id th id tm
|
||||
|
||||
whnfNoLog defs ctx sg (QCloT (SubR tm th)) = do
|
||||
let Val q = ctx.qtyLen
|
||||
whnfNoLog defs ctx sg $ pushSubstsWith' th id id tm
|
||||
|
|
|
@ -8,44 +8,45 @@ import Data.SnocVect
|
|||
|
||||
|
||||
private
|
||||
tycaseRhs : (k : TyConKind) -> TypeCaseArms d n ->
|
||||
Maybe (ScopeTermN (arity k) d n)
|
||||
tycaseRhs : (k : TyConKind) -> TypeCaseArms q d n ->
|
||||
Maybe (ScopeTermN (arity k) q d n)
|
||||
tycaseRhs k arms = lookupPrecise k arms
|
||||
|
||||
private
|
||||
tycaseRhsDef : Term d n -> (k : TyConKind) -> TypeCaseArms d n ->
|
||||
ScopeTermN (arity k) d n
|
||||
tycaseRhsDef : (def : Term q d n) -> (k : TyConKind) -> TypeCaseArms q d n ->
|
||||
ScopeTermN (arity k) q d n
|
||||
tycaseRhsDef def k arms = fromMaybe (SN def) $ tycaseRhs k arms
|
||||
|
||||
private
|
||||
tycaseRhs0 : (k : TyConKind) -> TypeCaseArms d n ->
|
||||
(0 eq : arity k = 0) => Maybe (Term d n)
|
||||
tycaseRhs0 : (k : TyConKind) -> TypeCaseArms q d n ->
|
||||
(0 eq : arity k = 0) => Maybe (Term q d n)
|
||||
tycaseRhs0 k arms = map (.term0) $ rewrite sym eq in tycaseRhs k arms
|
||||
|
||||
private
|
||||
tycaseRhsDef0 : Term d n -> (k : TyConKind) -> TypeCaseArms d n ->
|
||||
(0 eq : arity k = 0) => Term d n
|
||||
tycaseRhsDef0 : (def : Term q d n) -> (k : TyConKind) -> TypeCaseArms q d n ->
|
||||
(0 eq : arity k = 0) => Term q d n
|
||||
tycaseRhsDef0 def k arms = fromMaybe def $ tycaseRhs0 k arms
|
||||
|
||||
|
||||
parameters {auto _ : CanWhnf Term Interface.isRedexT}
|
||||
{auto _ : CanWhnf Elim Interface.isRedexE}
|
||||
(defs : Definitions) (ctx : WhnfContext d n)
|
||||
(defs : Definitions) (ctx : WhnfContext q 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
|
||||
export covering
|
||||
tycasePi : (t : Term d n) -> (0 tnf : No (isRedexT defs ctx SZero t)) =>
|
||||
Eff Whnf (Term d n, ScopeTerm d n)
|
||||
tycasePi : (t : Term q d n) -> (0 tnf : No (isRedexT defs ctx SZero t)) =>
|
||||
Eff Whnf (Term q d n, ScopeTerm q d n)
|
||||
tycasePi (Pi {arg, res, _}) = pure (arg, res)
|
||||
tycasePi (E e) {tnf} = do
|
||||
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
|
||||
res' = typeCase1Y e (Arr Zero arg ty loc) KPi [< !narg, !nret]
|
||||
let Val q = ctx.qtyLen
|
||||
loc = e.loc
|
||||
narg = mnb "Arg" loc; nret = mnb "Ret" loc
|
||||
arg = E $ typeCase1Y e ty KPi [< !narg, !nret] (BVT 1 loc) loc
|
||||
res' = typeCase1Y e (Arr (zero loc) arg ty loc) KPi [< !narg, !nret]
|
||||
(BVT 0 loc) loc
|
||||
res = ST [< !narg] $ E $ App (weakE 1 res') (BVT 0 loc) loc
|
||||
res = ST [< !narg] $ E $ App (weakE 1 res') (BVT 0 loc) loc
|
||||
pure (arg, res)
|
||||
tycasePi t = throw $ ExpectedPi t.loc ctx.names t
|
||||
|
||||
|
@ -53,17 +54,18 @@ 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 ctx SZero t)) =>
|
||||
Eff Whnf (Term d n, ScopeTerm d n)
|
||||
tycaseSig : (t : Term q d n) -> (0 tnf : No (isRedexT defs ctx SZero t)) =>
|
||||
Eff Whnf (Term q d n, ScopeTerm q d n)
|
||||
tycaseSig (Sig {fst, snd, _}) = pure (fst, snd)
|
||||
tycaseSig (E e) {tnf} = do
|
||||
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
|
||||
snd' = typeCase1Y e (Arr Zero fst ty loc) KSig [< !nfst, !nsnd]
|
||||
let Val q = ctx.qtyLen
|
||||
loc = e.loc
|
||||
nfst = mnb "Fst" loc; nsnd = mnb "Snd" loc
|
||||
fst = E $ typeCase1Y e ty KSig [< !nfst, !nsnd] (BVT 1 loc) loc
|
||||
snd' = typeCase1Y e (Arr (zero loc) fst ty loc) KSig [< !nfst, !nsnd]
|
||||
(BVT 0 loc) loc
|
||||
snd = ST [< !nfst] $ E $ App (weakE 1 snd') (BVT 0 loc) loc
|
||||
snd = ST [< !nfst] $ E $ App (weakE 1 snd') (BVT 0 loc) loc
|
||||
pure (fst, snd)
|
||||
tycaseSig t = throw $ ExpectedSig t.loc ctx.names t
|
||||
|
||||
|
@ -71,8 +73,8 @@ 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 ctx SZero t)) =>
|
||||
Eff Whnf (Term d n)
|
||||
tycaseBOX : (t : Term q d n) -> (0 tnf : No (isRedexT defs ctx SZero t)) =>
|
||||
Eff Whnf (Term q d n)
|
||||
tycaseBOX (BOX {ty, _}) = pure ty
|
||||
tycaseBOX (E e) {tnf} = do
|
||||
ty <- computeElimType defs ctx SZero e {ne = noOr2 tnf}
|
||||
|
@ -83,12 +85,14 @@ 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 ctx SZero t)) =>
|
||||
Eff Whnf (Term d n, Term d n, DScopeTerm d n, Term d n, Term d n)
|
||||
tycaseEq : (t : Term q d n) -> (0 tnf : No (isRedexT defs ctx SZero t)) =>
|
||||
Eff Whnf (Term q d n, Term q d n, DScopeTerm q d n,
|
||||
Term q d n, Term q d n)
|
||||
tycaseEq (Eq {ty, l, r, _}) = pure (ty.zero, ty.one, ty, l, r)
|
||||
tycaseEq (E e) {tnf} = do
|
||||
ty <- computeElimType defs ctx SZero e {ne = noOr2 tnf}
|
||||
let loc = e.loc
|
||||
let Val q = ctx.qtyLen
|
||||
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
|
||||
a1 = E $ typeCase1Y e ty KEq !names (BVT 3 loc) loc
|
||||
|
@ -104,10 +108,10 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
|
|||
||| `reduceTypeCase A i Q arms def _` reduces an expression
|
||||
||| `type-case A ∷ ★ᵢ return Q of { arms; _ ⇒ def }`
|
||||
export covering
|
||||
reduceTypeCase : (ty : Term d n) -> (u : Universe) -> (ret : Term d n) ->
|
||||
(arms : TypeCaseArms d n) -> (def : Term d n) ->
|
||||
reduceTypeCase : (ty : Term q d n) -> (u : Universe) -> (ret : Term q d n) ->
|
||||
(arms : TypeCaseArms q d n) -> (def : Term q d n) ->
|
||||
(0 _ : So (isTyCon ty)) => Loc ->
|
||||
Eff Whnf (Subset (Elim d n) (No . isRedexE defs ctx SZero))
|
||||
Eff Whnf (Subset (Elim q d n) (No . isRedexE defs ctx SZero))
|
||||
reduceTypeCase ty u ret arms def loc = case ty of
|
||||
-- (type-case ★ᵢ ∷ _ return Q of { ★ ⇒ s; ⋯ }) ⇝ s ∷ Q
|
||||
TYPE {} =>
|
||||
|
@ -120,9 +124,10 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
|
|||
-- (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, _} =>
|
||||
let arg' = Ann arg (TYPE u arg.loc) arg.loc
|
||||
res' = Ann (Lam res res.loc)
|
||||
(Arr Zero arg (TYPE u arg.loc) arg.loc) res.loc
|
||||
let Val q = ctx.qtyLen
|
||||
arg' = Ann arg (TYPE u arg.loc) arg.loc
|
||||
res' = Ann (Lam res res.loc)
|
||||
(Arr (zero loc) arg (TYPE u arg.loc) arg.loc) res.loc
|
||||
in
|
||||
whnf defs ctx SZero $
|
||||
Ann (subN (tycaseRhsDef def KPi arms) [< arg', res']) ret loc
|
||||
|
@ -130,9 +135,10 @@ parameters {auto _ : CanWhnf Term Interface.isRedexT}
|
|||
-- (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, _} =>
|
||||
let fst' = Ann fst (TYPE u fst.loc) fst.loc
|
||||
snd' = Ann (Lam snd snd.loc)
|
||||
(Arr Zero fst (TYPE u fst.loc) fst.loc) snd.loc
|
||||
let Val q = ctx.qtyLen
|
||||
fst' = Ann fst (TYPE u fst.loc) fst.loc
|
||||
snd' = Ann (Lam snd snd.loc)
|
||||
(Arr (zero loc) fst (TYPE u fst.loc) fst.loc) snd.loc
|
||||
in
|
||||
whnf defs ctx SZero $
|
||||
Ann (subN (tycaseRhsDef def KSig arms) [< fst', snd']) ret loc
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue