mirror of
https://github.com/idris-lang/Idris2.git
synced 2025-01-05 15:08:00 +03:00
Merge pull request #235 from edwinb/reflect-binders
Make lambda elaborator dependent
This commit is contained in:
commit
1e393b6ffc
@ -20,9 +20,9 @@ data Elab : Type -> Type where
|
|||||||
Quote : val -> Elab TTImp
|
Quote : val -> Elab TTImp
|
||||||
|
|
||||||
-- Elaborate under a lambda
|
-- Elaborate under a lambda
|
||||||
Lambda : (0 x : Type) -> (x -> Elab ty) -> Elab (x -> ty)
|
Lambda : (0 x : Type) ->
|
||||||
-- Elaborate under a forall
|
{0 ty : x -> Type} ->
|
||||||
ForAll : (0 x : Type) -> (x -> Elab ty) -> Elab (x -> ty)
|
((val : x) -> Elab (ty val)) -> Elab ((val : x) -> (ty val))
|
||||||
|
|
||||||
-- Get the current goal type, if known
|
-- Get the current goal type, if known
|
||||||
-- (it might need to be inferred from the solution)
|
-- (it might need to be inferred from the solution)
|
||||||
@ -90,13 +90,11 @@ quote : val -> Elab TTImp
|
|||||||
quote = Quote
|
quote = Quote
|
||||||
|
|
||||||
export
|
export
|
||||||
lambda : (0 x : Type) -> (x -> Elab ty) -> Elab (x -> ty)
|
lambda : (0 x : Type) ->
|
||||||
|
{0 ty : x -> Type} ->
|
||||||
|
((val : x) -> Elab (ty val)) -> Elab ((val : x) -> (ty val))
|
||||||
lambda = Lambda
|
lambda = Lambda
|
||||||
|
|
||||||
export
|
|
||||||
forAll : (0 x : Type) -> (x -> Elab ty) -> Elab (x -> ty)
|
|
||||||
forAll = ForAll
|
|
||||||
|
|
||||||
export
|
export
|
||||||
goal : Elab (Maybe TTImp)
|
goal : Elab (Maybe TTImp)
|
||||||
goal = Goal
|
goal = Goal
|
||||||
|
@ -93,7 +93,7 @@ elabScript fc nest env (NDCon nfc nm t ar args) exp
|
|||||||
defs <- get Ctxt
|
defs <- get Ctxt
|
||||||
empty <- clearDefs defs
|
empty <- clearDefs defs
|
||||||
scriptRet !(unelabUniqueBinders env !(quote empty env tm'))
|
scriptRet !(unelabUniqueBinders env !(quote empty env tm'))
|
||||||
elabCon defs "Lambda" [_, x, scope]
|
elabCon defs "Lambda" [x, _, scope]
|
||||||
= do empty <- clearDefs defs
|
= do empty <- clearDefs defs
|
||||||
NBind bfc x (Lam c p ty) sc <- evalClosure defs scope
|
NBind bfc x (Lam c p ty) sc <- evalClosure defs scope
|
||||||
| _ => throw (GenericMsg fc "Not a lambda")
|
| _ => throw (GenericMsg fc "Not a lambda")
|
||||||
@ -114,27 +114,6 @@ elabScript fc nest env (NDCon nfc nm t ar args) exp
|
|||||||
quotePi Implicit = pure Implicit
|
quotePi Implicit = pure Implicit
|
||||||
quotePi AutoImplicit = pure AutoImplicit
|
quotePi AutoImplicit = pure AutoImplicit
|
||||||
quotePi (DefImplicit t) = throw (GenericMsg fc "Can't add default lambda")
|
quotePi (DefImplicit t) = throw (GenericMsg fc "Can't add default lambda")
|
||||||
elabCon defs "ForAll" [_, x, scope]
|
|
||||||
= do empty <- clearDefs defs
|
|
||||||
NBind bfc x (Pi c p ty) sc <- evalClosure defs scope
|
|
||||||
| _ => throw (GenericMsg fc "Not a lambda")
|
|
||||||
n <- genVarName "x"
|
|
||||||
sc' <- sc defs (toClosure withAll env (Ref bfc Bound n))
|
|
||||||
qsc <- quote empty env sc'
|
|
||||||
let lamsc = refToLocal n x qsc
|
|
||||||
qp <- quotePi p
|
|
||||||
qty <- quote empty env ty
|
|
||||||
let env' = Pi c qp qty :: env
|
|
||||||
|
|
||||||
runsc <- elabScript fc (weaken nest) env'
|
|
||||||
!(nf defs env' lamsc) Nothing -- (map weaken exp)
|
|
||||||
nf empty env (Bind bfc x (Pi c qp qty) !(quote empty env' runsc))
|
|
||||||
where
|
|
||||||
quotePi : PiInfo (NF vars) -> Core (PiInfo (Term vars))
|
|
||||||
quotePi Explicit = pure Explicit
|
|
||||||
quotePi Implicit = pure Implicit
|
|
||||||
quotePi AutoImplicit = pure AutoImplicit
|
|
||||||
quotePi (DefImplicit t) = throw (GenericMsg fc "Can't add default lambda")
|
|
||||||
elabCon defs "Goal" []
|
elabCon defs "Goal" []
|
||||||
= do let Just gty = exp
|
= do let Just gty = exp
|
||||||
| Nothing => nfOpts withAll defs env
|
| Nothing => nfOpts withAll defs env
|
||||||
|
Loading…
Reference in New Issue
Block a user