add logging to core
This commit is contained in:
parent
861bd55f94
commit
3b6ae36e4e
14 changed files with 353 additions and 132 deletions
|
@ -2,6 +2,7 @@ module Quox.Whnf.ComputeElimType
|
|||
|
||||
import Quox.Whnf.Interface
|
||||
import Quox.Displace
|
||||
import Quox.Pretty
|
||||
|
||||
%default total
|
||||
|
||||
|
@ -18,7 +19,6 @@ computeElimType :
|
|||
(e : Elim d n) -> (0 ne : No (isRedexE defs ctx sg e)) =>
|
||||
Eff Whnf (Term d n)
|
||||
|
||||
|
||||
||| computes a type and then reduces it to whnf
|
||||
export covering
|
||||
computeWhnfElimType0 :
|
||||
|
@ -28,7 +28,16 @@ computeWhnfElimType0 :
|
|||
(e : Elim d n) -> (0 ne : No (isRedexE defs ctx sg e)) =>
|
||||
Eff Whnf (Term d n)
|
||||
|
||||
computeElimType defs ctx sg e =
|
||||
|
||||
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)
|
||||
|
||||
computeElimTypeNoLog defs ctx sg e =
|
||||
case e of
|
||||
F x u loc => do
|
||||
let Just def = lookup x defs
|
||||
|
@ -39,7 +48,7 @@ computeElimType defs ctx sg e =
|
|||
pure (ctx.tctx !! i).type
|
||||
|
||||
App f s loc =>
|
||||
case !(computeWhnfElimType0 defs ctx sg f {ne = noOr1 ne}) of
|
||||
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
|
||||
|
||||
|
@ -47,12 +56,12 @@ computeElimType defs ctx sg e =
|
|||
pure $ sub1 ret pair
|
||||
|
||||
Fst pair loc =>
|
||||
case !(computeWhnfElimType0 defs ctx sg pair {ne = noOr1 ne}) of
|
||||
case !(computeWhnfElimType0NoLog defs ctx sg pair {ne = noOr1 ne}) of
|
||||
Sig {fst, _} => pure fst
|
||||
ty => throw $ ExpectedSig loc ctx.names ty
|
||||
|
||||
Snd pair loc =>
|
||||
case !(computeWhnfElimType0 defs ctx sg pair {ne = noOr1 ne}) of
|
||||
case !(computeWhnfElimType0NoLog defs ctx sg pair {ne = noOr1 ne}) of
|
||||
Sig {snd, _} => pure $ sub1 snd $ Fst pair loc
|
||||
ty => throw $ ExpectedSig loc ctx.names ty
|
||||
|
||||
|
@ -66,7 +75,7 @@ computeElimType defs ctx sg e =
|
|||
pure $ sub1 ret box
|
||||
|
||||
DApp {fun = f, arg = p, loc} =>
|
||||
case !(computeWhnfElimType0 defs ctx sg f {ne = noOr1 ne}) of
|
||||
case !(computeWhnfElimType0NoLog defs ctx sg f {ne = noOr1 ne}) of
|
||||
Eq {ty, _} => pure $ dsub1 ty p
|
||||
t => throw $ ExpectedEq loc ctx.names t
|
||||
|
||||
|
@ -82,5 +91,20 @@ computeElimType defs ctx sg e =
|
|||
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]]
|
||||
res <- computeElimTypeNoLog defs ctx sg e {ne}
|
||||
say "whnf" 91 e.loc $
|
||||
hsep ["computeElimType ⇝",
|
||||
runPretty $ prettyTerm ctx.dnames ctx.tnames res]
|
||||
pure res
|
||||
|
||||
computeWhnfElimType0 defs ctx sg e =
|
||||
computeElimType defs ctx sg e >>= whnf0 defs ctx SZero
|
||||
|
||||
computeWhnfElimType0NoLog defs ctx sg e {ne} =
|
||||
computeElimTypeNoLog defs ctx sg e {ne} >>= whnf0 defs ctx SZero
|
||||
|
|
|
@ -1,6 +1,7 @@
|
|||
module Quox.Whnf.Interface
|
||||
|
||||
import public Quox.No
|
||||
import public Quox.Log
|
||||
import public Quox.Syntax
|
||||
import public Quox.Definition
|
||||
import public Quox.Typing.Context
|
||||
|
@ -13,7 +14,7 @@ import public Control.Eff
|
|||
|
||||
public export
|
||||
Whnf : List (Type -> Type)
|
||||
Whnf = [Except Error, NameGen]
|
||||
Whnf = [Except Error, NameGen, Log]
|
||||
|
||||
|
||||
public export
|
||||
|
@ -24,17 +25,20 @@ RedexTest tm =
|
|||
public export
|
||||
interface CanWhnf (0 tm : TermLike) (0 isRedex : RedexTest tm) | tm
|
||||
where
|
||||
whnf : (defs : Definitions) -> (ctx : WhnfContext d n) -> (q : SQty) ->
|
||||
tm d n -> Eff Whnf (Subset (tm d n) (No . isRedex defs ctx q))
|
||||
whnf, whnfNoLog :
|
||||
(defs : Definitions) -> (ctx : WhnfContext d n) -> (q : SQty) ->
|
||||
tm d n -> Eff Whnf (Subset (tm d n) (No . isRedex defs ctx 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
|
||||
-- cases idris can't tell that `isRedex` and `isRedexT` are the same thing
|
||||
|
||||
public export %inline
|
||||
whnf0 : {0 isRedex : RedexTest tm} -> CanWhnf tm isRedex =>
|
||||
Definitions -> WhnfContext d n -> SQty -> tm d n -> Eff Whnf (tm d n)
|
||||
whnf0 defs ctx q t = fst <$> whnf defs ctx q t
|
||||
whnf0, whnfNoLog0 :
|
||||
{0 isRedex : RedexTest tm} -> CanWhnf tm isRedex =>
|
||||
Definitions -> WhnfContext d n -> SQty -> tm d n -> Eff Whnf (tm 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 =>
|
||||
|
|
|
@ -4,6 +4,7 @@ import Quox.Whnf.Interface
|
|||
import Quox.Whnf.ComputeElimType
|
||||
import Quox.Whnf.TypeCase
|
||||
import Quox.Whnf.Coercion
|
||||
import Quox.Pretty
|
||||
import Quox.Displace
|
||||
import Data.SnocVect
|
||||
|
||||
|
@ -14,19 +15,43 @@ export covering CanWhnf Term Interface.isRedexT
|
|||
export covering CanWhnf Elim Interface.isRedexE
|
||||
|
||||
|
||||
-- the String is what to call the "s" argument in logs (maybe "s", or "e")
|
||||
private %inline
|
||||
whnfDefault :
|
||||
{0 isRedex : RedexTest tm} ->
|
||||
(CanWhnf tm isRedex, Located2 tm) =>
|
||||
String ->
|
||||
(forall d, n. WhnfContext d n -> tm d n -> Eff Pretty LogDoc) ->
|
||||
(defs : Definitions) ->
|
||||
(ctx : WhnfContext d n) ->
|
||||
(sg : SQty) ->
|
||||
(s : tm d n) ->
|
||||
Eff Whnf (Subset (tm 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],
|
||||
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]
|
||||
pure res
|
||||
|
||||
covering
|
||||
CanWhnf Elim Interface.isRedexE where
|
||||
whnf defs ctx sg (F x u loc) with (lookupElim0 x u defs) proof eq
|
||||
whnf = whnfDefault "e" $ \ctx, e => prettyElim ctx.dnames ctx.tnames 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
|
||||
_ | Nothing = pure $ Element (F x u loc) $ rewrite eq in Ah
|
||||
|
||||
whnf defs ctx sg (B i loc) with (ctx.tctx !! i) proof eq1
|
||||
whnfNoLog defs ctx sg (B i loc) with (ctx.tctx !! i) proof eq1
|
||||
_ | l with (l.term) proof eq2
|
||||
_ | Just y = whnf defs ctx sg $ Ann y l.type loc
|
||||
_ | Nothing = pure $ Element (B i loc) $ rewrite eq1 in rewrite eq2 in Ah
|
||||
|
||||
-- ((λ x ⇒ t) ∷ (π.x : A) → B) s ⇝ t[s∷A/x] ∷ B[s∷A/x]
|
||||
whnf defs ctx sg (App f s appLoc) = do
|
||||
whnfNoLog defs ctx sg (App f s appLoc) = do
|
||||
Element f fnf <- whnf defs ctx sg f
|
||||
case nchoose $ isLamHead f of
|
||||
Left _ => case f of
|
||||
|
@ -41,7 +66,7 @@ CanWhnf Elim Interface.isRedexE where
|
|||
--
|
||||
-- 0 · case e return p ⇒ C of { (a, b) ⇒ u } ⇝
|
||||
-- u[fst e/a, snd e/b] ∷ C[e/p]
|
||||
whnf defs ctx sg (CasePair pi pair ret body caseLoc) = do
|
||||
whnfNoLog defs ctx sg (CasePair pi pair ret body caseLoc) = do
|
||||
Element pair pairnf <- whnf defs ctx sg pair
|
||||
case nchoose $ isPairHead pair of
|
||||
Left _ => case pair of
|
||||
|
@ -64,7 +89,7 @@ CanWhnf Elim Interface.isRedexE where
|
|||
(pairnf `orNo` np `orNo` notYesNo n0)
|
||||
|
||||
-- fst ((s, t) ∷ (x : A) × B) ⇝ s ∷ A
|
||||
whnf defs ctx sg (Fst pair fstLoc) = do
|
||||
whnfNoLog defs ctx sg (Fst pair fstLoc) = do
|
||||
Element pair pairnf <- whnf defs ctx sg pair
|
||||
case nchoose $ isPairHead pair of
|
||||
Left _ => case pair of
|
||||
|
@ -76,7 +101,7 @@ CanWhnf Elim Interface.isRedexE where
|
|||
pure $ Element (Fst pair fstLoc) (pairnf `orNo` np)
|
||||
|
||||
-- snd ((s, t) ∷ (x : A) × B) ⇝ t ∷ B[(s ∷ A)/x]
|
||||
whnf defs ctx sg (Snd pair sndLoc) = do
|
||||
whnfNoLog defs ctx sg (Snd pair sndLoc) = do
|
||||
Element pair pairnf <- whnf defs ctx sg pair
|
||||
case nchoose $ isPairHead pair of
|
||||
Left _ => case pair of
|
||||
|
@ -89,7 +114,7 @@ CanWhnf Elim Interface.isRedexE where
|
|||
|
||||
-- case 'a ∷ {a,…} return p ⇒ C of { 'a ⇒ u } ⇝
|
||||
-- u ∷ C['a∷{a,…}/p]
|
||||
whnf defs ctx sg (CaseEnum pi tag ret arms caseLoc) = do
|
||||
whnfNoLog defs ctx sg (CaseEnum pi tag ret arms caseLoc) = do
|
||||
Element tag tagnf <- whnf defs ctx sg tag
|
||||
case nchoose $ isTagHead tag of
|
||||
Left _ => case tag of
|
||||
|
@ -110,7 +135,7 @@ 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 sg (CaseNat pi piIH nat ret zer suc caseLoc) = do
|
||||
whnfNoLog defs ctx sg (CaseNat pi piIH nat ret zer suc caseLoc) = do
|
||||
Element nat natnf <- whnf defs ctx sg nat
|
||||
case nchoose $ isNatHead nat of
|
||||
Left _ =>
|
||||
|
@ -137,7 +162,7 @@ CanWhnf Elim Interface.isRedexE where
|
|||
|
||||
-- case [t] ∷ [π.A] return p ⇒ C of { [x] ⇒ u } ⇝
|
||||
-- u[t∷A/x] ∷ C[[t] ∷ [π.A]/p]
|
||||
whnf defs ctx sg (CaseBox pi box ret body caseLoc) = do
|
||||
whnfNoLog defs ctx sg (CaseBox pi box ret body caseLoc) = do
|
||||
Element box boxnf <- whnf defs ctx sg box
|
||||
case nchoose $ isBoxHead box of
|
||||
Left _ => case box of
|
||||
|
@ -153,7 +178,7 @@ CanWhnf Elim Interface.isRedexE where
|
|||
-- e : Eq (𝑗 ⇒ A) t u ⊢ e @1 ⇝ u ∷ A‹1/𝑗›
|
||||
--
|
||||
-- ((δ 𝑖 ⇒ s) ∷ Eq (𝑗 ⇒ A) t u) @𝑘 ⇝ s‹𝑘/𝑖› ∷ A‹𝑘/𝑗›
|
||||
whnf defs ctx sg (DApp f p appLoc) = do
|
||||
whnfNoLog defs ctx sg (DApp f p appLoc) = do
|
||||
Element f fnf <- whnf defs ctx sg f
|
||||
case nchoose $ isDLamHead f of
|
||||
Left _ => case f of
|
||||
|
@ -173,7 +198,7 @@ CanWhnf Elim Interface.isRedexE where
|
|||
B {} => pure $ Element (DApp f p appLoc) (fnf `orNo` ndlh `orNo` Ah)
|
||||
|
||||
-- e ∷ A ⇝ e
|
||||
whnf defs ctx sg (Ann s a annLoc) = do
|
||||
whnfNoLog defs ctx sg (Ann s a annLoc) = do
|
||||
Element s snf <- whnf defs ctx sg s
|
||||
case nchoose $ isE s of
|
||||
Left _ => let E e = s in pure $ Element e $ noOr2 snf
|
||||
|
@ -181,7 +206,7 @@ CanWhnf Elim Interface.isRedexE where
|
|||
Element a anf <- whnf defs ctx SZero a
|
||||
pure $ Element (Ann s a annLoc) (ne `orNo` snf `orNo` anf)
|
||||
|
||||
whnf defs ctx sg (Coe sty p q val coeLoc) =
|
||||
whnfNoLog defs ctx sg (Coe sty p q val coeLoc) =
|
||||
-- 𝑖 ∉ fv(A)
|
||||
-- -------------------------------
|
||||
-- coe (𝑖 ⇒ A) @p @q s ⇝ s ∷ A
|
||||
|
@ -201,7 +226,7 @@ CanWhnf Elim Interface.isRedexE where
|
|||
(_, Right ty) =>
|
||||
whnf defs ctx sg $ Ann val ty coeLoc
|
||||
|
||||
whnf defs ctx sg (Comp ty p q val r zero one compLoc) =
|
||||
whnfNoLog defs ctx sg (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 sg $ Ann val ty compLoc
|
||||
|
@ -213,7 +238,7 @@ CanWhnf Elim Interface.isRedexE where
|
|||
B {} => pure $ Element (Comp ty p q val r zero one compLoc)
|
||||
(notYesNo npq `orNo` Ah)
|
||||
|
||||
whnf defs ctx sg (TypeCase ty ret arms def tcLoc) =
|
||||
whnfNoLog defs ctx sg (TypeCase ty ret arms def tcLoc) =
|
||||
case sg `decEq` SZero of
|
||||
Yes Refl => do
|
||||
Element ty tynf <- whnf defs ctx SZero ty
|
||||
|
@ -226,48 +251,50 @@ CanWhnf Elim Interface.isRedexE where
|
|||
No _ =>
|
||||
throw $ ClashQ tcLoc sg.qty Zero
|
||||
|
||||
whnf defs ctx sg (CloE (Sub el th)) =
|
||||
whnf defs ctx sg $ pushSubstsWith' id th el
|
||||
whnf defs ctx sg (DCloE (Sub el th)) =
|
||||
whnf defs ctx sg $ pushSubstsWith' th id el
|
||||
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
|
||||
|
||||
covering
|
||||
CanWhnf Term Interface.isRedexT where
|
||||
whnf _ _ _ t@(TYPE {}) = pure $ nred t
|
||||
whnf _ _ _ t@(IOState {}) = 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@(Nat {}) = pure $ nred t
|
||||
whnf _ _ _ t@(STRING {}) = pure $ nred t
|
||||
whnf _ _ _ t@(Str {}) = pure $ nred t
|
||||
whnf _ _ _ t@(BOX {}) = pure $ nred t
|
||||
whnf _ _ _ t@(Box {}) = pure $ nred t
|
||||
whnf = whnfDefault "e" $ \ctx, s => prettyTerm ctx.dnames ctx.tnames s
|
||||
|
||||
whnf _ _ _ (Succ p loc) =
|
||||
whnfNoLog _ _ _ t@(TYPE {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(IOState {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(Pi {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(Lam {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(Sig {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(Pair {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(Enum {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(Tag {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(Eq {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(DLam {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(NAT {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(Nat {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(STRING {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(Str {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(BOX {}) = pure $ nred t
|
||||
whnfNoLog _ _ _ t@(Box {}) = pure $ nred t
|
||||
|
||||
whnfNoLog _ _ _ (Succ p loc) =
|
||||
case nchoose $ isNatConst p of
|
||||
Left _ => case p of
|
||||
Nat p _ => pure $ nred $ Nat (S p) loc
|
||||
E (Ann (Nat p _) _ _) => pure $ nred $ Nat (S p) loc
|
||||
Right nc => pure $ nred $ Succ p loc
|
||||
|
||||
whnf defs ctx sg (Let _ rhs body _) =
|
||||
whnfNoLog defs ctx sg (Let _ rhs body _) =
|
||||
whnf defs ctx sg $ sub1 body rhs
|
||||
|
||||
-- s ∷ A ⇝ s (in term context)
|
||||
whnf defs ctx sg (E e) = do
|
||||
whnfNoLog defs ctx sg (E e) = do
|
||||
Element e enf <- whnf defs ctx sg 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 sg (CloT (Sub tm th)) =
|
||||
whnf defs ctx sg $ pushSubstsWith' id th tm
|
||||
whnf defs ctx sg (DCloT (Sub tm th)) =
|
||||
whnf defs ctx sg $ pushSubstsWith' th id tm
|
||||
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
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue