more work on case
This commit is contained in:
@@ -69,11 +69,17 @@ parameters (ctx: Context)
|
||||
Just x => goSpine ren lvl (Bnd fc $ cast x) sp
|
||||
go ren lvl (VRef fc nm def sp) = goSpine ren lvl (Ref fc nm def) sp
|
||||
go ren lvl (VMeta fc ix sp) = if ix == meta
|
||||
then error emptyFC "meta occurs check"
|
||||
-- REVIEW is this the right fc?
|
||||
then error fc "meta occurs check"
|
||||
else goSpine ren lvl (Meta fc ix) sp
|
||||
go ren lvl (VLam fc n t) = pure (Lam fc n !(go (lvl :: ren) (S lvl) !(t $$ VVar emptyFC lvl [<])))
|
||||
go ren lvl (VPi fc n icit ty tm) = pure (Pi fc n icit !(go ren lvl ty) !(go (lvl :: ren) (S lvl) !(tm $$ VVar emptyFC lvl [<])))
|
||||
go ren lvl (VU fc) = pure (U fc)
|
||||
go ren lvl (VCase fc sc alts) = error fc "Case in solution"
|
||||
-- This seems dicey, for VLam we eval and then process the levels.
|
||||
-- Here we have raw Tm so we haven't even done occurs check. I'm thinking
|
||||
-- we don't allow solutions with Case in them
|
||||
-- pure (Case fc !(go ren lvl sc) alts)
|
||||
|
||||
lams : Nat -> Tm -> Tm
|
||||
lams 0 tm = tm
|
||||
|
||||
@@ -9,6 +9,7 @@ import Data.List
|
||||
import Lib.Types -- Name / Tm
|
||||
import Lib.TopContext
|
||||
import Lib.TT -- lookupMeta
|
||||
import Lib.Util
|
||||
|
||||
public export
|
||||
data CExp : Type
|
||||
@@ -31,13 +32,6 @@ data CExp : Type where
|
||||
CRef : Name -> CExp
|
||||
CMeta : Nat -> CExp
|
||||
|
||||
funArgs : Tm -> (Tm, List Tm)
|
||||
funArgs tm = go tm []
|
||||
where
|
||||
go : Tm -> List Tm -> (Tm, List Tm)
|
||||
go (App _ t u) args = go t (u :: args)
|
||||
go t args = (t, args)
|
||||
|
||||
||| I'm counting Lam in the term for arity. This matches what I need in
|
||||
||| code gen.
|
||||
export
|
||||
|
||||
@@ -2,19 +2,19 @@ module Lib.ProcessDecl
|
||||
|
||||
import Data.IORef
|
||||
|
||||
import Lib.Types
|
||||
import Lib.Parser
|
||||
import Lib.TT
|
||||
import Lib.TopContext
|
||||
import Lib.Check
|
||||
import Lib.Parser
|
||||
import Lib.Syntax
|
||||
import Lib.TopContext
|
||||
import Lib.TT
|
||||
import Lib.Types
|
||||
import Lib.Util
|
||||
|
||||
getArity : Tm -> Nat
|
||||
getArity (Pi x str icit t u) = S (getArity u)
|
||||
-- Ref or App (of type constructor) are valid
|
||||
getArity _ = Z
|
||||
|
||||
|
||||
-- Can metas live in context for now?
|
||||
|
||||
export
|
||||
@@ -84,10 +84,19 @@ processDecl (Data fc nm ty cons) = do
|
||||
-- expecting tm to be a Pi type
|
||||
(TypeSig fc nm' tm) => do
|
||||
ctx <- get
|
||||
-- TODO check pi type ending in full tyty application
|
||||
-- TODO count arity
|
||||
dty <- check (mkCtx ctx.metas) tm (VU fc)
|
||||
debug "dty \{nm'} is \{pprint [] dty}"
|
||||
|
||||
-- We only check that codomain uses the right type constructor
|
||||
-- We know it's in U because it's part of a checked Pi type
|
||||
let (codomain, tele) = splitTele dty
|
||||
-- for printing
|
||||
let tnames = reverse $ map (\(MkBind _ nm _ _) => nm) tele
|
||||
let (Ref _ hn _, args) := funArgs codomain
|
||||
| (tm, _) => error (getFC tm) "expected \{nm} got \{pprint tnames tm}"
|
||||
when (hn /= nm) $
|
||||
error (getFC codomain) "Constructor codomain is \{pprint tnames codomain} rather than \{nm}"
|
||||
|
||||
modify $ setDef nm' dty (DCon (getArity dty) nm)
|
||||
pure nm'
|
||||
_ => throwError $ E (0,0) "expected constructor declaration"
|
||||
|
||||
@@ -106,6 +106,8 @@ lookup ctx nm = go ctx.types
|
||||
export
|
||||
eval : Env -> Mode -> Tm -> M Val
|
||||
|
||||
|
||||
|
||||
-- REVIEW everything is evalutated whether it's needed or not
|
||||
-- It would be nice if the environment were lazy.
|
||||
-- e.g. case is getting evaluated when passed to a function because
|
||||
@@ -130,6 +132,24 @@ vappSpine : Val -> SnocList Val -> M Val
|
||||
vappSpine t [<] = pure t
|
||||
vappSpine t (xs :< x) = vapp !(vappSpine t xs) x
|
||||
|
||||
-- So we need:
|
||||
-- - a Neutral case statement
|
||||
-- - split out data / type constructors from VRef application
|
||||
-- - should we sort out what the case tree really looks like first?
|
||||
|
||||
-- Technically I don't need this now, as a neutral would be fine.
|
||||
|
||||
evalAlt : Env -> Mode -> Val -> List CaseAlt -> M (Maybe Val)
|
||||
-- FIXME spine length? Should this be VRef or do we specialize?
|
||||
evalAlt env mode (VRef _ nm y sp) ((CaseCons name args t) :: xs) =
|
||||
if nm == name
|
||||
-- Here we bind the args and push on. Do we have enough? Too many?
|
||||
then ?evalAlt_success
|
||||
-- here we need to know if we've got a mismatched constructor or some function app
|
||||
else ?evalAlt_what
|
||||
evalAlt env mode sc (CaseDefault u :: xs) = pure $ Just !(eval (sc :: env) mode u)
|
||||
evalAlt env mode sc _ = pure Nothing -- stuck
|
||||
|
||||
bind : Val -> Env -> Env
|
||||
bind v env = v :: env
|
||||
|
||||
@@ -154,16 +174,24 @@ eval env mode (Meta fc i) =
|
||||
(Solved k t) => pure $ t
|
||||
eval env mode (Lam fc x t) = pure $ VLam fc x (MkClosure env t)
|
||||
eval env mode (Pi fc x icit a b) = pure $ VPi fc x icit !(eval env mode a) (MkClosure env b)
|
||||
-- Here, we assume env has everything. We push levels onto it during type checking.
|
||||
-- I think we could pass in an l and assume everything outside env is free and
|
||||
-- translate to a level
|
||||
eval env mode (Bnd fc i) = case getAt i env of
|
||||
Just rval => pure rval
|
||||
Nothing => error' "Bad deBruin index \{show i}"
|
||||
|
||||
-- We need a neutral and some code to run the case tree
|
||||
eval env mode tm@(Case{}) = ?todo_eval_case
|
||||
|
||||
eval env mode tm@(Case fc sc alts) = pure $ VCase fc !(eval env mode sc) alts
|
||||
-- case !(evalAlt env mode !(eval env mode sc) alts) of
|
||||
-- Just foo => ?goodAlt
|
||||
-- Nothing => ?stuckAlt
|
||||
|
||||
export
|
||||
quote : (lvl : Nat) -> Val -> M Tm
|
||||
|
||||
|
||||
quoteSp : (lvl : Nat) -> Tm -> SnocList Val -> M Tm
|
||||
quoteSp lvl t [<] = pure t
|
||||
quoteSp lvl t (xs :< x) =
|
||||
@@ -178,6 +206,7 @@ quote l (VLam fc x t) = pure $ Lam fc x !(quote (S l) !(t $$ VVar emptyFC l [<])
|
||||
quote l (VPi fc x icit a b) = pure $ Pi fc x icit !(quote l a) !(quote (S l) !(b $$ VVar emptyFC l [<]))
|
||||
quote l (VU fc) = pure (U fc)
|
||||
quote l (VRef fc n def sp) = quoteSp l (Ref fc n def) sp
|
||||
quote l (VCase fc sc alts) = pure $ Case fc !(quote l sc) alts
|
||||
|
||||
-- Can we assume closed terms?
|
||||
-- ezoo only seems to use it at [], but essentially does this:
|
||||
|
||||
@@ -180,6 +180,8 @@ data Val : Type where
|
||||
-- I wanted the Maybe Tm in here, but for now we're always expanding.
|
||||
-- Maybe this is where I glue
|
||||
VRef : FC -> (nm : String) -> Def -> (sp : SnocList Val) -> Val
|
||||
-- neutral case
|
||||
VCase : FC -> (sc : Val) -> List CaseAlt -> Val
|
||||
-- we'll need to look this up in ctx with IO
|
||||
VMeta : FC -> (ix : Nat) -> (sp : SnocList Val) -> Val
|
||||
VLam : FC -> Name -> Closure -> Val
|
||||
@@ -187,6 +189,7 @@ data Val : Type where
|
||||
VU : FC -> Val
|
||||
|
||||
|
||||
|
||||
Show Closure
|
||||
|
||||
covering export
|
||||
@@ -197,6 +200,7 @@ Show Val where
|
||||
show (VLam _ str x) = "(%lam \{str} \{show x})"
|
||||
show (VPi fc str Implicit x y) = "(%pi {\{str} : \{show x}}. \{show y})"
|
||||
show (VPi fc str Explicit x y) = "(%pi (\{str} : \{show x}). \{show y})"
|
||||
show (VCase fc sc alts) = "(%case \{show sc} ...)"
|
||||
show (VU _) = "U"
|
||||
|
||||
-- Not used - I was going to change context to have a List Binder
|
||||
|
||||
36
src/Lib/Util.idr
Normal file
36
src/Lib/Util.idr
Normal file
@@ -0,0 +1,36 @@
|
||||
module Lib.Util
|
||||
|
||||
import Lib.Types
|
||||
|
||||
export
|
||||
funArgs : Tm -> (Tm, List Tm)
|
||||
funArgs tm = go tm []
|
||||
where
|
||||
go : Tm -> List Tm -> (Tm, List Tm)
|
||||
go (App _ t u) args = go t (u :: args)
|
||||
go t args = (t, args)
|
||||
|
||||
|
||||
public export
|
||||
data Binder : Type where
|
||||
MkBind : FC -> String -> Icit -> Tm -> Binder
|
||||
|
||||
-- I don't have a show for terms without a name list
|
||||
export
|
||||
Show Binder where
|
||||
show (MkBind _ nm icit t) = "[\{nm} \{show icit} : ...]"
|
||||
|
||||
data Foo : Type -> Type where
|
||||
MkFoo : Nat -> Foo a
|
||||
|
||||
export
|
||||
splitTele : Tm -> (Tm, List Binder)
|
||||
splitTele = go []
|
||||
where
|
||||
go : List Binder -> Tm -> (Tm, List Binder)
|
||||
go ts (Pi fc nm icit t u) = go (MkBind fc nm icit t :: ts) u
|
||||
go ts tm = (tm, reverse ts)
|
||||
|
||||
|
||||
-- splitTele (Pi x str icit t u) = ?splitTele_rhs_6
|
||||
-- splitTele tm =
|
||||
Reference in New Issue
Block a user