diff --git a/src/Compiler/ANF.idr b/src/Compiler/ANF.idr index fa694850a8..c74be2229b 100644 --- a/src/Compiler/ANF.idr +++ b/src/Compiler/ANF.idr @@ -6,6 +6,7 @@ import Core.CompileExpr import Core.Context import Core.Core import Core.TT +import Core.Name.ScopedList import Data.List import Data.Vect @@ -136,9 +137,9 @@ Show ANFDef where show args ++ " -> " ++ show ret show (MkAError exp) = "Error: " ++ show exp -data AVars : List Name -> Type where - Nil : AVars [] - (::) : Int -> AVars xs -> AVars (x :: xs) +data AVars : ScopedList Name -> Type where + Nil : AVars SLNil + (::) : Int -> AVars xs -> AVars (x :%: xs) data Next : Type where @@ -194,7 +195,7 @@ mutual List (Lifted vars) -> (List AVar -> ANF) -> Core ANF anfArgs fc vs args f = do args' <- traverse (anf vs) args - letBind fc args' f + letBind fc (toList args') f anf : {vars : _} -> {auto v : Ref Next Int} -> @@ -244,10 +245,10 @@ mutual = do (is, vs') <- bindArgs args vs pure $ MkAConAlt n ci t is !(anf vs' sc) where - bindArgs : (args : List Name) -> AVars vars' -> - Core (List Int, AVars (args ++ vars')) - bindArgs [] vs = pure ([], vs) - bindArgs (n :: ns) vs + bindArgs : (args : ScopedList Name) -> AVars vars' -> + Core (List Int, AVars (args +%+ vars')) + bindArgs SLNil vs = pure ([], vs) + bindArgs (n :%: ns) vs = do i <- nextVar (is, vs') <- bindArgs ns vs pure (i :: is, i :: vs') @@ -269,10 +270,10 @@ toANF (MkLFun args scope sc) pure $ MkAFun (iargs ++ reverse iargs') !(anf vs sc) where bindArgs : {auto v : Ref Next Int} -> - (args : List Name) -> AVars vars' -> - Core (List Int, AVars (args ++ vars')) - bindArgs [] vs = pure ([], vs) - bindArgs (n :: ns) vs + (args : ScopedList Name) -> AVars vars' -> + Core (List Int, AVars (args +%+ vars')) + bindArgs SLNil vs = pure ([], vs) + bindArgs (n :%: ns) vs = do i <- nextVar (is, vs') <- bindArgs ns vs pure (i :: is, i :: vs') diff --git a/src/Compiler/CaseOpts.idr b/src/Compiler/CaseOpts.idr index 2f20494c05..dda4d92313 100644 --- a/src/Compiler/CaseOpts.idr +++ b/src/Compiler/CaseOpts.idr @@ -32,14 +32,14 @@ case t of shiftUnder : {args : _} -> {idx : _} -> - (0 p : IsVar n idx (x :: args ++ vars)) -> - NVar n (args ++ x :: vars) + (0 p : IsVar n idx (x :%: args +%+ vars)) -> + NVar n (args +%+ x :%: vars) shiftUnder First = weakenNVar (mkSizeOf args) (MkNVar First) shiftUnder (Later p) = insertNVar (mkSizeOf args) (MkNVar p) shiftVar : {outer, args : Scope} -> - NVar n (outer ++ (x :: args ++ vars)) -> - NVar n (outer ++ (args ++ x :: vars)) + NVar n (outer +%+ (x :%: args +%+ vars)) -> + NVar n (outer +%+ (args +%+ x :%: vars)) shiftVar nvar = let out = mkSizeOf outer in case locateNVar out nvar of @@ -49,21 +49,21 @@ shiftVar nvar mutual shiftBinder : {outer, args : _} -> (new : Name) -> - CExp (outer ++ old :: (args ++ vars)) -> - CExp (outer ++ (args ++ new :: vars)) + CExp (outer +%+ old :%: (args +%+ vars)) -> + CExp (outer +%+ (args +%+ new :%: vars)) shiftBinder new (CLocal fc p) = case shiftVar (MkNVar p) of MkNVar p' => CLocal fc (renameVar p') where - renameVar : IsVar x i (outer ++ (args ++ (old :: rest))) -> - IsVar x i (outer ++ (args ++ (new :: rest))) + renameVar : IsVar x i (outer +%+ (args +%+ (old :%: rest))) -> + IsVar x i (outer +%+ (args +%+ (new :%: rest))) renameVar = believe_me -- it's the same index, so just the identity at run time shiftBinder new (CRef fc n) = CRef fc n shiftBinder {outer} new (CLam fc n sc) - = CLam fc n $ shiftBinder {outer = n :: outer} new sc + = CLam fc n $ shiftBinder {outer = n :%: outer} new sc shiftBinder new (CLet fc n inlineOK val sc) = CLet fc n inlineOK (shiftBinder new val) - $ shiftBinder {outer = n :: outer} new sc + $ shiftBinder {outer = n :%: outer} new sc shiftBinder new (CApp fc f args) = CApp fc (shiftBinder new f) $ map (shiftBinder new) args shiftBinder new (CCon fc ci c tag args) @@ -87,34 +87,34 @@ mutual shiftBinderConAlt : {outer, args : _} -> (new : Name) -> - CConAlt (outer ++ (x :: args ++ vars)) -> - CConAlt (outer ++ (args ++ new :: vars)) + CConAlt (outer +%+ (x :%: args +%+ vars)) -> + CConAlt (outer +%+ (args +%+ new :%: vars)) shiftBinderConAlt new (MkConAlt n ci t args' sc) - = let sc' : CExp ((args' ++ outer) ++ (x :: args ++ vars)) - = rewrite sym (appendAssociative args' outer (x :: args ++ vars)) in sc in + = let sc' : CExp ((args' +%+ outer) +%+ (x :%: args +%+ vars)) + = rewrite sym (appendAssociative args' outer (x :%: args +%+ vars)) in sc in MkConAlt n ci t args' $ - rewrite (appendAssociative args' outer (args ++ new :: vars)) - in shiftBinder new {outer = args' ++ outer} sc' + rewrite (appendAssociative args' outer (args +%+ new :%: vars)) + in shiftBinder new {outer = args' +%+ outer} sc' shiftBinderConstAlt : {outer, args : _} -> (new : Name) -> - CConstAlt (outer ++ (x :: args ++ vars)) -> - CConstAlt (outer ++ (args ++ new :: vars)) + CConstAlt (outer +%+ (x :%: args +%+ vars)) -> + CConstAlt (outer +%+ (args +%+ new :%: vars)) shiftBinderConstAlt new (MkConstAlt c sc) = MkConstAlt c $ shiftBinder new sc -- If there's a lambda inside a case, move the variable so that it's bound -- outside the case block so that we can bind it just once outside the block liftOutLambda : {args : _} -> (new : Name) -> - CExp (old :: args ++ vars) -> - CExp (args ++ new :: vars) -liftOutLambda = shiftBinder {outer = []} + CExp (old :%: args +%+ vars) -> + CExp (args +%+ new :%: vars) +liftOutLambda = shiftBinder {outer = SLNil} -- If all the alternatives start with a lambda, we can have a single lambda -- binding outside tryLiftOut : (new : Name) -> List (CConAlt vars) -> - Maybe (List (CConAlt (new :: vars))) + Maybe (List (CConAlt (new :%: vars))) tryLiftOut new [] = Just [] tryLiftOut new (MkConAlt n ci t args (CLam fc x sc) :: as) = do as' <- tryLiftOut new as @@ -124,20 +124,20 @@ tryLiftOut _ _ = Nothing tryLiftOutConst : (new : Name) -> List (CConstAlt vars) -> - Maybe (List (CConstAlt (new :: vars))) + Maybe (List (CConstAlt (new :%: vars))) tryLiftOutConst new [] = Just [] tryLiftOutConst new (MkConstAlt c (CLam fc x sc) :: as) = do as' <- tryLiftOutConst new as - let sc' = liftOutLambda {args = []} new sc + let sc' = liftOutLambda {args = SLNil} new sc pure (MkConstAlt c sc' :: as') tryLiftOutConst _ _ = Nothing tryLiftDef : (new : Name) -> Maybe (CExp vars) -> - Maybe (Maybe (CExp (new :: vars))) + Maybe (Maybe (CExp (new :%: vars))) tryLiftDef new Nothing = Just Nothing tryLiftDef new (Just (CLam fc x sc)) - = let sc' = liftOutLambda {args = []} new sc in + = let sc' = liftOutLambda {args = SLNil} new sc in pure (Just sc') tryLiftDef _ _ = Nothing diff --git a/src/Compiler/Common.idr b/src/Compiler/Common.idr index e113a34d0a..086e14e823 100644 --- a/src/Compiler/Common.idr +++ b/src/Compiler/Common.idr @@ -80,7 +80,7 @@ Ord UsePhase where public export record CompileData where constructor MkCompileData - mainExpr : CExp [] -- main expression to execute. This also appears in + mainExpr : CExp SLNil -- main expression to execute. This also appears in -- the definitions below as MN "__mainExpression" 0 -- For incremental compilation and for compiling exported -- names only, this can be set to 'erased'. @@ -152,7 +152,7 @@ getMinimalDef (Coded ns bin) name <- fromBuf b let def = MkGlobalDef fc name (Erased fc Placeholder) [] [] [] [] mul - [] (specified Public) (MkTotality Unchecked IsCovering) False + SLNil (specified Public) (MkTotality Unchecked IsCovering) False [] Nothing refsR False False True None cdef Nothing [] Nothing pure (def, Just (ns, bin)) @@ -351,8 +351,8 @@ getCompileDataWith exports doLazyAnnots phase_in tm_in traverse (lambdaLift doLazyAnnots) cseDefs else pure [] - let lifted = (mainname, MkLFun [] [] liftedtm) :: - ldefs ++ concat lifted_in + let lifted = (mainname, MkLFun SLNil SLNil liftedtm) :: + (ldefs ++ concat lifted_in) anf <- if phase >= ANF then logTime 2 "Get ANF" $ traverse (\ (n, d) => pure (n, !(toANF d))) lifted @@ -408,7 +408,7 @@ getCompileData = getCompileDataWith [] export compileTerm : {auto c : Ref Ctxt Defs} -> - ClosedTerm -> Core (CExp []) + ClosedTerm -> Core (CExp SLNil) compileTerm tm_in = do tm <- toFullNames tm_in fixArityExp !(compileExp tm) diff --git a/src/Compiler/CompileExpr.idr b/src/Compiler/CompileExpr.idr index 249773ace0..34dffb1700 100644 --- a/src/Compiler/CompileExpr.idr +++ b/src/Compiler/CompileExpr.idr @@ -6,6 +6,7 @@ import Core.Context import Core.Context.Log import Core.Env import Core.Name +import Core.Name.ScopedList as SL import Core.Normalise import Core.Options import Core.TT @@ -39,16 +40,16 @@ numArgs defs (Ref _ _ n) _ => pure (Arity 0) numArgs _ tm = pure (Arity 0) -mkSub : Nat -> (ns : List Name) -> List Nat -> (ns' ** Thin ns' ns) +mkSub : Nat -> (ns : ScopedList Name) -> List Nat -> (ns' ** Thin ns' ns) mkSub i _ [] = (_ ** Refl) -mkSub i [] ns = (_ ** Refl) -mkSub i (x :: xs) es +mkSub i SLNil ns = (_ ** Refl) +mkSub i (x :%: xs) es = let (ns' ** p) = mkSub (S i) xs es in if i `elem` es then (ns' ** Drop p) - else (x :: ns' ** Keep p) + else (x :%: ns' ** Keep p) -weakenVar : Var ns -> Var (a :: ns) +weakenVar : Var ns -> Var (a :%: ns) weakenVar (MkVar p) = (MkVar (Later p)) etaExpand : {vars : _} -> @@ -79,7 +80,7 @@ expandToArity Z f args = applyAll f args where applyAll : CExp vars -> List (CExp vars) -> CExp vars applyAll fn [] = fn - applyAll fn (a :: args) = applyAll (CApp (getFC fn) fn [a]) args + applyAll fn (a :: args) = applyAll (CApp (getFC fn) fn []) args expandToArity (S k) f (a :: args) = expandToArity k (addArg f a) args where addArg : CExp vars -> CExp vars -> CExp vars @@ -132,15 +133,15 @@ eraseConArgs arity epos fn args else fn' mkDropSubst : Nat -> List Nat -> - (rest : List Name) -> - (vars : List Name) -> - (vars' ** Thin (vars' ++ rest) (vars ++ rest)) -mkDropSubst i es rest [] = ([] ** Refl) -mkDropSubst i es rest (x :: xs) + (rest : ScopedList Name) -> + (vars : ScopedList Name) -> + (vars' ** Thin (vars' +%+ rest) (vars +%+ rest)) +mkDropSubst i es rest SLNil = (SLNil ** Refl) +mkDropSubst i es rest (x :%: xs) = let (vs ** sub) = mkDropSubst (1 + i) es rest xs in if i `elem` es then (vs ** Drop sub) - else (x :: vs ** Keep sub) + else (x :%: vs ** Keep sub) -- Rewrite applications of Nat-like constructors and functions to more optimal -- versions using Integer @@ -226,12 +227,12 @@ natBranch (MkConAlt n SUCC _ _ _) = True natBranch _ = False trySBranch : CExp vars -> CConAlt vars -> Maybe (CExp vars) -trySBranch n (MkConAlt nm SUCC _ [arg] sc) +trySBranch n (MkConAlt nm SUCC _ (arg :%: SLNil) sc) = Just (CLet (getFC n) arg YesInline (magic__natUnsuc (getFC n) (getFC n) [n]) sc) trySBranch _ _ = Nothing tryZBranch : CConAlt vars -> Maybe (CExp vars) -tryZBranch (MkConAlt n ZERO _ [] sc) = Just sc +tryZBranch (MkConAlt n ZERO _ SLNil sc) = Just sc tryZBranch _ = Nothing getSBranch : CExp vars -> List (CConAlt vars) -> Maybe (CExp vars) @@ -271,7 +272,7 @@ enumTree (CConCase fc sc alts def) CConstCase fc sc alts' def where toEnum : CConAlt vars -> Maybe (CConstAlt vars) - toEnum (MkConAlt nm (ENUM n) (Just tag) [] sc) + toEnum (MkConAlt nm (ENUM n) (Just tag) SLNil sc) = pure $ MkConstAlt (enumTag n tag) sc toEnum _ = Nothing enumTree t = t @@ -279,7 +280,7 @@ enumTree t = t -- remove pattern matches on unit unitTree : {auto u : Ref NextMN Int} -> CExp vars -> Core (CExp vars) unitTree exp@(CConCase fc sc alts def) = fromMaybe (pure exp) - $ do let [MkConAlt _ UNIT _ [] e] = alts + $ do let [MkConAlt _ UNIT _ SLNil e] = alts | _ => Nothing Just $ case sc of -- TODO: Check scrutinee has no effect, and skip let binding CLocal _ _ => pure e @@ -343,8 +344,7 @@ toCExpTm n (Bind fc x (Let _ rig val _) sc) rig toCExpTm n (Bind fc x (Pi _ c e ty) sc) = pure $ CCon fc (UN (Basic "->")) TYCON Nothing - [ !(toCExp n ty) - , CLam fc x !(toCExp n sc)] + [!(toCExp n ty), CLam fc x !(toCExp n sc)] toCExpTm n (Bind fc x b tm) = pure $ CErased fc -- We'd expect this to have been dealt with in toCExp, but for completeness... toCExpTm n (App fc tm arg) @@ -460,12 +460,12 @@ mutual pure $ Just (substs s env !(toCExpTree n sc)) else -- let bind the scrutinee, and substitute the -- name into the RHS - let (s, env) : (_, SubstCEnv args (MN "eff" 0 :: vars)) + let (s, env) : (_, SubstCEnv args (MN "eff" 0 :%: vars)) = mkSubst 0 (CLocal fc First) pos args in do sc' <- toCExpTree n sc let scope = insertNames {outer=args} {inner=vars} - {ns = [MN "eff" 0]} + {ns = (MN "eff" 0 :%: SLNil)} (mkSizeOf _) (mkSizeOf _) sc' let tm = CLet fc (MN "eff" 0) NotInline scr (substs s env scope) log "compiler.newtype.world" 50 "Kept the scrutinee \{show tm}" @@ -473,9 +473,9 @@ mutual _ => pure Nothing -- there's a normal match to do where mkSubst : Nat -> CExp vs -> - Nat -> (args : List Name) -> (SizeOf args, SubstCEnv args vs) - mkSubst _ _ _ [] = (zero, []) - mkSubst i scr pos (a :: as) + Nat -> (args : ScopedList Name) -> (SizeOf args, SubstCEnv args vs) + mkSubst _ _ _ SLNil = (zero, []) + mkSubst i scr pos (a :%: as) = let (s, env) = mkSubst (1 + i) scr pos as in if i == pos then (suc s, scr :: env) @@ -545,9 +545,9 @@ mutual -- Need this for ensuring that argument list matches up to operator arity for -- builtins -data ArgList : Nat -> List Name -> Type where - NoArgs : ArgList Z [] - ConsArg : (a : Name) -> ArgList k as -> ArgList (S k) (a :: as) +data ArgList : Nat -> ScopedList Name -> Type where + NoArgs : ArgList Z SLNil + ConsArg : (a : Name) -> ArgList k as -> ArgList (S k) (a :%: as) mkArgList : Int -> (n : Nat) -> (ns ** ArgList n ns) mkArgList i Z = (_ ** NoArgs) @@ -556,35 +556,35 @@ mkArgList i (S k) (_ ** ConsArg (MN "arg" i) rec) data NArgs : Type where - User : Name -> List (Closure []) -> NArgs - Struct : String -> List (String, Closure []) -> NArgs + User : Name -> List (Closure SLNil) -> NArgs + Struct : String -> List (String, Closure SLNil) -> NArgs NUnit : NArgs NPtr : NArgs NGCPtr : NArgs NBuffer : NArgs NForeignObj : NArgs - NIORes : Closure [] -> NArgs + NIORes : Closure SLNil -> NArgs getPArgs : {auto c : Ref Ctxt Defs} -> - Defs -> Closure [] -> Core (String, Closure []) + Defs -> Closure SLNil -> Core (String, Closure SLNil) getPArgs defs cl = do NDCon fc _ _ _ args <- evalClosure defs cl | nf => throw (GenericMsg (getLoc nf) "Badly formed struct type") case reverse (map snd args) of - (tydesc :: n :: _) => + (tydesc :%: n :%: _) => do NPrimVal _ (Str n') <- evalClosure defs n | nf => throw (GenericMsg (getLoc nf) "Unknown field name") pure (n', tydesc) _ => throw (GenericMsg fc "Badly formed struct type") getFieldArgs : {auto c : Ref Ctxt Defs} -> - Defs -> Closure [] -> Core (List (String, Closure [])) + Defs -> Closure SLNil -> Core (List (String, Closure SLNil)) getFieldArgs defs cl = do NDCon fc _ _ _ args <- evalClosure defs cl | nf => throw (GenericMsg (getLoc nf) "Badly formed struct type") case map snd args of -- cons - [_, t, rest] => + (_ :%: t :%: rest :%: SLNil) => do rest' <- getFieldArgs defs rest (n, ty) <- getPArgs defs t pure ((n, ty) :: rest') @@ -592,23 +592,23 @@ getFieldArgs defs cl _ => pure [] getNArgs : {auto c : Ref Ctxt Defs} -> - Defs -> Name -> List (Closure []) -> Core NArgs -getNArgs defs (NS _ (UN $ Basic "IORes")) [arg] = pure $ NIORes arg -getNArgs defs (NS _ (UN $ Basic "Ptr")) [arg] = pure NPtr -getNArgs defs (NS _ (UN $ Basic "AnyPtr")) [] = pure NPtr -getNArgs defs (NS _ (UN $ Basic "GCPtr")) [arg] = pure NGCPtr -getNArgs defs (NS _ (UN $ Basic "GCAnyPtr")) [] = pure NGCPtr -getNArgs defs (NS _ (UN $ Basic "Buffer")) [] = pure NBuffer -getNArgs defs (NS _ (UN $ Basic "ForeignObj")) [] = pure NForeignObj -getNArgs defs (NS _ (UN $ Basic "Unit")) [] = pure NUnit -getNArgs defs (NS _ (UN $ Basic "Struct")) [n, args] + Defs -> Name -> ScopedList (Closure SLNil) -> Core NArgs +getNArgs defs (NS _ (UN $ Basic "IORes")) (arg :%: SLNil) = pure $ NIORes arg +getNArgs defs (NS _ (UN $ Basic "Ptr")) (arg :%: SLNil) = pure NPtr +getNArgs defs (NS _ (UN $ Basic "AnyPtr")) SLNil = pure NPtr +getNArgs defs (NS _ (UN $ Basic "GCPtr")) (arg :%: SLNil) = pure NGCPtr +getNArgs defs (NS _ (UN $ Basic "GCAnyPtr")) SLNil = pure NGCPtr +getNArgs defs (NS _ (UN $ Basic "Buffer")) SLNil = pure NBuffer +getNArgs defs (NS _ (UN $ Basic "ForeignObj")) SLNil = pure NForeignObj +getNArgs defs (NS _ (UN $ Basic "Unit")) SLNil = pure NUnit +getNArgs defs (NS _ (UN $ Basic "Struct")) (n :%: args :%: SLNil) = do NPrimVal _ (Str n') <- evalClosure defs n | nf => throw (GenericMsg (getLoc nf) "Unknown name for struct") pure (Struct n' !(getFieldArgs defs args)) -getNArgs defs n args = pure $ User n args +getNArgs defs n args = pure $ User n (toList args) nfToCFType : {auto c : Ref Ctxt Defs} -> - FC -> (inStruct : Bool) -> NF [] -> Core CFType + FC -> (inStruct : Bool) -> NF SLNil -> Core CFType nfToCFType _ _ (NPrimVal _ $ PrT IntType) = pure CFInt nfToCFType _ _ (NPrimVal _ $ PrT IntegerType) = pure CFInteger nfToCFType _ _ (NPrimVal _ $ PrT Bits8Type) = pure CFUnsigned8 @@ -669,7 +669,7 @@ nfToCFType fc s t show !(toFullNames ty))) getCFTypes : {auto c : Ref Ctxt Defs} -> - List CFType -> NF [] -> + List CFType -> NF SLNil -> Core (List CFType, CFType) getCFTypes args (NBind fc _ (Pi _ _ _ ty) sc) = do defs <- get Ctxt @@ -679,38 +679,38 @@ getCFTypes args (NBind fc _ (Pi _ _ _ ty) sc) getCFTypes args t = pure (reverse args, !(nfToCFType (getLoc t) False t)) -lamRHSenv : Int -> FC -> (ns : List Name) -> (SizeOf ns, SubstCEnv ns []) -lamRHSenv i fc [] = (zero, []) -lamRHSenv i fc (n :: ns) +lamRHSenv : Int -> FC -> (ns : ScopedList Name) -> (SizeOf ns, SubstCEnv ns SLNil) +lamRHSenv i fc SLNil = (zero, []) +lamRHSenv i fc (n :%: ns) = let (s, env) = lamRHSenv (i + 1) fc ns in (suc s, CRef fc (MN "x" i) :: env) -mkBounds : (xs : _) -> Bounds xs -mkBounds [] = None -mkBounds (x :: xs) = Add x x (mkBounds xs) +mkBounds : (xs : _) -> SL.Bounds xs +mkBounds SLNil = None +mkBounds (x :%: xs) = Add x x (mkBounds xs) getNewArgs : {done : _} -> - SubstCEnv done args -> List Name -getNewArgs [] = [] -getNewArgs (CRef _ n :: xs) = n :: getNewArgs xs -getNewArgs {done = x :: xs} (_ :: sub) = x :: getNewArgs sub + SubstCEnv done args -> ScopedList Name +getNewArgs [] = SLNil +getNewArgs (CRef _ n :: xs) = n :%: getNewArgs xs +getNewArgs {done = x :%: xs} (_ :: sub) = x :%: getNewArgs sub -- If a name is declared in one module and defined in another, -- we have to assume arity 0 for incremental compilation because -- we have no idea how it's defined, and when we made calls to the -- function, they had arity 0. -lamRHS : (ns : List Name) -> CExp ns -> CExp [] +lamRHS : (ns : ScopedList Name) -> CExp ns -> CExp SLNil lamRHS ns tm = let (s, env) = lamRHSenv 0 (getFC tm) ns tmExp = substs s env (rewrite appendNilRightNeutral ns in tm) newArgs = reverse $ getNewArgs env bounds = mkBounds newArgs - expLocs = mkLocals zero {vars = []} bounds tmExp in + expLocs = mkLocals zero {vars = SLNil} bounds tmExp in lamBind (getFC tm) _ expLocs where - lamBind : FC -> (ns : List Name) -> CExp ns -> CExp [] - lamBind fc [] tm = tm - lamBind fc (n :: ns) tm = lamBind fc ns (CLam fc n tm) + lamBind : FC -> (ns : ScopedList Name) -> CExp ns -> CExp SLNil + lamBind fc SLNil tm = tm + lamBind fc (n :%: ns) tm = lamBind fc ns (CLam fc n tm) toCDef : {auto c : Ref Ctxt Defs} -> Name -> ClosedTerm -> List Nat -> Def -> @@ -726,7 +726,7 @@ toCDef n ty erased (PMDef pi args _ tree _) else MkFun args' (shrinkCExp p comptree) where toLam : Bool -> CDef -> CDef - toLam True (MkFun args rhs) = MkFun [] (lamRHS args rhs) + toLam True (MkFun args rhs) = MkFun SLNil (lamRHS args rhs) toLam _ d = d toCDef n ty _ (ExternDef arity) = let (ns ** args) = mkArgList 0 arity in @@ -755,7 +755,7 @@ toCDef n ty _ (Builtin {arity} op) toCDef n _ _ (DCon tag arity pos) = do let nt = snd <$> pos defs <- get Ctxt - args <- numArgs {vars = []} defs (Ref EmptyFC (DataCon tag arity) n) + args <- numArgs {vars = SLNil} defs (Ref EmptyFC (DataCon tag arity) n) let arity' = case args of NewTypeBy ar _ => ar EraseArgs ar erased => ar `minus` length erased @@ -780,7 +780,7 @@ toCDef n ty _ def export compileExp : {auto c : Ref Ctxt Defs} -> - ClosedTerm -> Core (CExp []) + ClosedTerm -> Core (CExp SLNil) compileExp tm = do s <- newRef NextMN 0 exp <- toCExp (UN $ Basic "main") tm diff --git a/src/Compiler/ES/ToAst.idr b/src/Compiler/ES/ToAst.idr index c9d5661fe5..1d65dfff63 100644 --- a/src/Compiler/ES/ToAst.idr +++ b/src/Compiler/ES/ToAst.idr @@ -151,11 +151,11 @@ mutual -- list, to the surrounding scope. stmt e (NmApp _ x xs) = do (mbx, vx) <- liftFun x - (mbxs, args) <- liftArgs xs + (mbxs, args) <- liftArgs (toList xs) pure . prepend (mbx ++ mbxs) $ assign e (EApp vx args) stmt e (NmCon _ n ci tg xs) = do - (mbxs, args) <- liftArgs xs + (mbxs, args) <- liftArgs (toList xs) pure . prepend mbxs $ assign e (ECon (tag n tg) ci args) stmt e o@(NmOp _ x xs) = @@ -166,7 +166,7 @@ mutual pure . prepend mbxs $ assign e (EOp x args) stmt e (NmExtPrim _ n xs) = do - (mbxs, args) <- liftArgs xs + (mbxs, args) <- liftArgs (toList xs) pure . prepend mbxs $ assign e (EExtPrim n args) stmt e (NmForce _ _ x) = do diff --git a/src/Compiler/Inline.idr b/src/Compiler/Inline.idr index 535a9ca2b7..5472ed62d6 100644 --- a/src/Compiler/Inline.idr +++ b/src/Compiler/Inline.idr @@ -13,6 +13,7 @@ import Core.FC import Core.Hash import Core.Options import Core.TT +import Core.Name.ScopedList import Data.Maybe import Data.List @@ -23,17 +24,17 @@ import Libraries.Data.WithDefault %default covering -data EEnv : List Name -> List Name -> Type where - Nil : EEnv free [] - (::) : CExp free -> EEnv free vars -> EEnv free (x :: vars) +data EEnv : ScopedList Name -> ScopedList Name -> Type where + Nil : EEnv free SLNil + (::) : CExp free -> EEnv free vars -> EEnv free (x :%: vars) -extend : EEnv free vars -> (args : List (CExp free)) -> (args' : List Name) -> - LengthMatch args args' -> EEnv free (args' ++ vars) -extend env [] [] NilMatch = env -extend env (a :: xs) (n :: ns) (ConsMatch w) +extend : EEnv free vars -> (args : ScopedList (CExp free)) -> (args' : ScopedList Name) -> + LengthMatch args args' -> EEnv free (args' +%+ vars) +extend env SLNil SLNil NilMatch = env +extend env (a :%: xs) (n :%: ns) (ConsMatch w) = a :: extend env xs ns w -Stack : List Name -> Type +Stack : ScopedList Name -> Type Stack vars = List (CExp vars) unload : Stack vars -> CExp vars -> CExp vars @@ -49,12 +50,12 @@ getArity (MkCon _ arity _) = arity getArity (MkForeign _ args _) = length args getArity (MkError _) = 0 -takeFromStack : EEnv free vars -> Stack free -> (args : List Name) -> - Maybe (EEnv free (args ++ vars), Stack free) -takeFromStack env (e :: es) (a :: as) +takeFromStack : EEnv free vars -> Stack free -> (args : ScopedList Name) -> + Maybe (EEnv free (args +%+ vars), Stack free) +takeFromStack env (e :: es) (a :%: as) = do (env', stk') <- takeFromStack env es as pure (e :: env', stk') -takeFromStack env stk [] = pure (env, stk) +takeFromStack env stk SLNil = pure (env, stk) takeFromStack env [] args = Nothing data LVar : Type where @@ -66,7 +67,7 @@ genName n put LVar (i + 1) pure (MN n i) -refToLocal : Name -> (x : Name) -> CExp vars -> CExp (x :: vars) +refToLocal : Name -> (x : Name) -> CExp vars -> CExp (x :%: vars) refToLocal x new tm = refsToLocals (Add new x None) tm largest : Ord a => a -> List a -> a @@ -118,15 +119,15 @@ mutual {auto l : Ref LVar Int} -> FC -> List Name -> Stack free -> EEnv free vars -> - {idx : Nat} -> (0 p : IsVar x idx (vars ++ free)) -> + {idx : Nat} -> (0 p : IsVar x idx (vars +%+ free)) -> Core (CExp free) - evalLocal {vars = []} fc rec stk env p + evalLocal {vars = SLNil} fc rec stk env p = pure $ unload stk (CLocal fc p) - evalLocal {vars = x :: xs} fc rec stk (v :: env) First + evalLocal {vars = x :%: xs} fc rec stk (v :: env) First = case stk of [] => pure v _ => eval rec env stk (weakenNs (mkSizeOf xs) v) - evalLocal {vars = x :: xs} fc rec stk (_ :: env) (Later p) + evalLocal {vars = x :%: xs} fc rec stk (_ :: env) (Later p) = evalLocal fc rec stk env p tryApply : {vars, free : _} -> @@ -139,14 +140,14 @@ mutual | Nothing => pure Nothing res <- eval rec env' stk' (rewrite sym (appendAssociative args vars free) in - embed {outer = vars ++ free} exp) + embed {outer = vars +%+ free} exp) pure (Just res) tryApply rec stk env _ = pure Nothing eval : {vars, free : _} -> {auto c : Ref Ctxt Defs} -> {auto l : Ref LVar Int} -> - List Name -> EEnv free vars -> Stack free -> CExp (vars ++ free) -> + List Name -> EEnv free vars -> Stack free -> CExp (vars +%+ free) -> Core (CExp free) eval rec env stk (CLocal fc p) = evalLocal fc rec stk env p -- This is hopefully a temporary hack, giving a special case for io_bind. @@ -244,14 +245,14 @@ mutual def' where updateLoc : {idx, vs : _} -> - (0 p : IsVar x idx (vs ++ free)) -> + (0 p : IsVar x idx (vs +%+ free)) -> EEnv free vs -> CExp free -> EEnv free vs - updateLoc {vs = []} p env val = env - updateLoc {vs = (x::xs)} First (e :: env) val = val :: env - updateLoc {vs = (y::xs)} (Later p) (e :: env) val = e :: updateLoc p env val + updateLoc {vs = SLNil} p env val = env + updateLoc {vs = (x:%:xs)} First (e :: env) val = val :: env + updateLoc {vs = (y:%:xs)} (Later p) (e :: env) val = e :: updateLoc p env val update : {vs : _} -> - CExp (vs ++ free) -> EEnv free vs -> CExp free -> EEnv free vs + CExp (vs +%+ free) -> EEnv free vs -> CExp free -> EEnv free vs update (CLocal _ p) env sc = updateLoc p env sc update _ env _ = env @@ -267,10 +268,10 @@ mutual eval rec env stk (CCrash fc str) = pure $ unload stk $ CCrash fc str extendLoc : {auto l : Ref LVar Int} -> - FC -> EEnv free vars -> (args' : List Name) -> - Core (Bounds args', EEnv free (args' ++ vars)) - extendLoc fc env [] = pure (None, env) - extendLoc fc env (n :: ns) + FC -> EEnv free vars -> (args' : ScopedList Name) -> + Core (Bounds args', EEnv free (args' +%+ vars)) + extendLoc fc env SLNil = pure (None, env) + extendLoc fc env (n :%: ns) = do xn <- genName "cv" (bs', env') <- extendLoc fc env ns pure (Add n xn bs', CRef fc xn :: env') @@ -278,7 +279,7 @@ mutual evalAlt : {vars, free : _} -> {auto c : Ref Ctxt Defs} -> {auto l : Ref LVar Int} -> - FC -> List Name -> EEnv free vars -> Stack free -> CConAlt (vars ++ free) -> + FC -> List Name -> EEnv free vars -> Stack free -> CConAlt (vars +%+ free) -> Core (CConAlt free) evalAlt {free} {vars} fc rec env stk (MkConAlt n ci t args sc) = do (bs, env') <- extendLoc fc env args @@ -289,7 +290,7 @@ mutual evalConstAlt : {vars, free : _} -> {auto c : Ref Ctxt Defs} -> {auto l : Ref LVar Int} -> - List Name -> EEnv free vars -> Stack free -> CConstAlt (vars ++ free) -> + List Name -> EEnv free vars -> Stack free -> CConstAlt (vars +%+ free) -> Core (CConstAlt free) evalConstAlt rec env stk (MkConstAlt c sc) = MkConstAlt c <$> eval rec env stk sc @@ -298,18 +299,20 @@ mutual {auto c : Ref Ctxt Defs} -> {auto l : Ref LVar Int} -> List Name -> EEnv free vars -> Stack free -> - CExp free -> List (CConAlt (vars ++ free)) -> - Maybe (CExp (vars ++ free)) -> + CExp free -> List (CConAlt (vars +%+ free)) -> + Maybe (CExp (vars +%+ free)) -> Core (Maybe (CExp free)) pickAlt rec env stk (CCon fc n ci t args) [] def = traverseOpt (eval rec env stk) def pickAlt {vars} {free} rec env stk con@(CCon fc n ci t args) (MkConAlt n' _ t' args' sc :: alts) def - = if matches n t n' t' - then case checkLengthMatch args args' of + = + let args'' = fromList args + in if matches n t n' t' + then case checkLengthMatch args'' args' of Nothing => pure Nothing Just m => - do let env' : EEnv free (args' ++ vars) - = extend env args args' m + do let env' : EEnv free (args' +%+ vars) + = extend env args'' args' m pure $ Just !(eval rec env' stk (rewrite sym (appendAssociative args' vars free) in sc)) @@ -325,8 +328,8 @@ mutual {auto c : Ref Ctxt Defs} -> {auto l : Ref LVar Int} -> List Name -> EEnv free vars -> Stack free -> - CExp free -> List (CConstAlt (vars ++ free)) -> - Maybe (CExp (vars ++ free)) -> + CExp free -> List (CConstAlt (vars +%+ free)) -> + Maybe (CExp (vars +%+ free)) -> Core (Maybe (CExp free)) pickConstAlt rec env stk (CPrimVal fc c) [] def = traverseOpt (eval rec env stk) def @@ -415,32 +418,32 @@ fixArity d = pure d -- TODO: get rid of this `done` by making the return `args'` runtime irrelevant? getLams : {done : _} -> SizeOf done -> - Int -> SubstCEnv done args -> CExp (done ++ args) -> - (args' ** (SizeOf args', SubstCEnv args' args, CExp (args' ++ args))) + Int -> SubstCEnv done args -> CExp (done +%+ args) -> + (args' ** (SizeOf args', SubstCEnv args' args, CExp (args' +%+ args))) getLams {done} d i env (CLam fc x sc) - = getLams {done = x :: done} (suc d) (i + 1) (CRef fc (MN "ext" i) :: env) sc + = getLams {done = x :%: done} (suc d) (i + 1) (CRef fc (MN "ext" i) :: env) sc getLams {done} d i env sc = (done ** (d, env, sc)) -mkBounds : (xs : _) -> Bounds xs -mkBounds [] = None -mkBounds (x :: xs) = Add x x (mkBounds xs) +mkBounds : (xs : _) -> Core.Name.ScopedList.Bounds xs +mkBounds SLNil = None +mkBounds (x :%: xs) = Add x x (mkBounds xs) getNewArgs : {done : _} -> - SubstCEnv done args -> List Name -getNewArgs [] = [] -getNewArgs (CRef _ n :: xs) = n :: getNewArgs xs -getNewArgs {done = x :: xs} (_ :: sub) = x :: getNewArgs sub + SubstCEnv done args -> ScopedList Name +getNewArgs [] = SLNil +getNewArgs (CRef _ n :: xs) = n :%: getNewArgs xs +getNewArgs {done = x :%: xs} (_ :: sub) = x :%: getNewArgs sub -- Move any lambdas in the body of the definition into the lhs list of vars. -- Annoyingly, the indices will need fixing up because the order in the top -- level definition goes left to right (i.e. first argument has lowest index, -- not the highest, as you'd expect if they were all lambdas). -mergeLambdas : (args : List Name) -> CExp args -> (args' ** CExp args') +mergeLambdas : (args : ScopedList Name) -> CExp args -> (args' ** CExp args') mergeLambdas args (CLam fc x sc) = let (args' ** (s, env, exp')) = getLams zero 0 [] (CLam fc x sc) expNs = substs s env exp' newArgs = reverse $ getNewArgs env - expLocs = mkLocals (mkSizeOf args) {vars = []} (mkBounds newArgs) + expLocs = mkLocals (mkSizeOf args) {vars = SLNil} (mkBounds newArgs) (rewrite appendNilRightNeutral args in expNs) in (_ ** expLocs) mergeLambdas args exp = (args ** exp) diff --git a/src/Compiler/Interpreter/VMCode.idr b/src/Compiler/Interpreter/VMCode.idr index b5279d6176..7bbc715877 100644 --- a/src/Compiler/Interpreter/VMCode.idr +++ b/src/Compiler/Interpreter/VMCode.idr @@ -113,7 +113,7 @@ indexMaybe (x :: xs) idx = if idx <= 0 then Just x else indexMaybe xs (idx - 1) callPrim : Ref State InterpState => Stack -> PrimFn ar -> Vect ar Object -> Core Object callPrim stk BelieveMe [_, _, obj] = pure obj callPrim stk fn args = case the (Either Object (Vect ar Constant)) $ traverse getConst args of - Right args' => case getOp {vars=[]} fn (NPrimVal EmptyFC <$> args') of + Right args' => case getOp {vars=SLNil} fn (NPrimVal EmptyFC <$> args') of Just (NPrimVal _ res) => pure $ Const res _ => interpError stk $ "OP: Error calling " ++ show (opName fn) ++ " with operands: " ++ show args' Left obj => interpError stk $ "OP: Expected Constant, found " ++ showType obj diff --git a/src/Compiler/LambdaLift.idr b/src/Compiler/LambdaLift.idr index ec53c35ac1..6ba449580b 100644 --- a/src/Compiler/LambdaLift.idr +++ b/src/Compiler/LambdaLift.idr @@ -14,6 +14,7 @@ import Core.CompileExpr import Core.Context import Core.Core import Core.TT +import Core.Name.ScopedList import Data.List import Data.Vect @@ -92,7 +93,7 @@ mutual ||| @ expr is the expression to bind `x` to. ||| @ body is the expression to evaluate after binding. LLet : FC -> (x : Name) -> (expr : Lifted vars) -> - (body : Lifted (x :: vars)) -> Lifted vars + (body : Lifted (x :%: vars)) -> Lifted vars ||| Use of a constructor to construct a compound data type value. ||| @@ -169,7 +170,7 @@ mutual ||| @ vars is the list of names accessible within the current scope of the ||| lambda-lifted code. public export - data LiftedConAlt : (vars : List Name) -> Type where + data LiftedConAlt : (vars : ScopedList Name) -> Type where ||| Constructs a branch of an "LCon" (constructor tag) case statement. ||| @@ -187,7 +188,7 @@ mutual ||| @ body is the expression that is evaluated as the consequence of ||| this branch matching. MkLConAlt : (n : Name) -> (info : ConInfo) -> (tag : Maybe Int) -> - (args : List Name) -> (body : Lifted (args ++ vars)) -> + (args : ScopedList Name) -> (body : Lifted (args +%+ vars)) -> LiftedConAlt vars ||| A branch of an "LConst" (constant expression) case statement. @@ -195,7 +196,7 @@ mutual ||| @ vars is the list of names accessible within the current scope of the ||| lambda-lifted code. public export - data LiftedConstAlt : (vars : List Name) -> Type where + data LiftedConstAlt : (vars : ScopedList Name) -> Type where ||| Constructs a branch of an "LConst" (constant expression) case ||| statement. @@ -224,8 +225,8 @@ data LiftedDef : Type where -- (Sorry for the awkward API - it's to do with how the indices are -- arranged for the variables, and it could be expensive to reshuffle them! -- See Compiler.ANF for an example of how they get resolved to names) - MkLFun : (args : List Name) -> (scope : List Name) -> - (body : Lifted (scope ++ args)) -> LiftedDef + MkLFun : (args : ScopedList Name) -> (scope : ScopedList Name) -> + (body : Lifted (scope +%+ args)) -> LiftedDef ||| Constructs a definition of a constructor for a compound data type. ||| @@ -258,7 +259,7 @@ data LiftedDef : Type where ||| `LCrash` rather than `prim_crash`. ||| ||| @ expl : an explanation of the error. - MkLError : (expl : Lifted []) -> LiftedDef + MkLError : (expl : Lifted SLNil) -> LiftedDef showLazy : Maybe LazyReason -> String showLazy = maybe "" $ (" " ++) . show @@ -269,26 +270,26 @@ mutual {vs : _} -> Show (Lifted vs) where show (LLocal {idx} _ p) = "!" ++ show (nameAt p) show (LAppName fc lazy n args) - = show n ++ showLazy lazy ++ "(" ++ showSep ", " (map show args) ++ ")" + = show n ++ showLazy lazy ++ "(" ++ showSep ", " (toList $ map show args) ++ ")" show (LUnderApp fc n m args) = "<" ++ show n ++ " underapp " ++ show m ++ ">(" ++ - showSep ", " (map show args) ++ ")" + showSep ", " (toList $ map show args) ++ ")" show (LApp fc lazy c arg) = show c ++ showLazy lazy ++ " @ (" ++ show arg ++ ")" show (LLet fc x val sc) = "%let " ++ show x ++ " = " ++ show val ++ " in " ++ show sc show (LCon fc n _ t args) - = "%con " ++ show n ++ "(" ++ showSep ", " (map show args) ++ ")" + = "%con " ++ show n ++ "(" ++ showSep ", " (toList $ map show args) ++ ")" show (LOp fc lazy op args) = "%op " ++ show op ++ showLazy lazy ++ "(" ++ showSep ", " (toList (map show args)) ++ ")" show (LExtPrim fc lazy p args) - = "%extprim " ++ show p ++ showLazy lazy ++ "(" ++ showSep ", " (map show args) ++ ")" + = "%extprim " ++ show p ++ showLazy lazy ++ "(" ++ showSep ", " (toList $ map show args) ++ ")" show (LConCase fc sc alts def) = "%case " ++ show sc ++ " of { " - ++ showSep "| " (map show alts) ++ " " ++ show def + ++ showSep "| " (toList $ map show alts) ++ " " ++ show def show (LConstCase fc sc alts def) = "%case " ++ show sc ++ " of { " - ++ showSep "| " (map show alts) ++ " " ++ show def + ++ showSep "| " (toList $ map show alts) ++ " " ++ show def show (LPrimVal _ x) = show x show (LErased _) = "___" show (LCrash _ x) = "%CRASH(" ++ show x ++ ")" @@ -298,7 +299,7 @@ mutual {vs : _} -> Show (LiftedConAlt vs) where show (MkLConAlt n _ t args sc) = "%conalt " ++ show n ++ - "(" ++ showSep ", " (map show args) ++ ") => " ++ show sc + "(" ++ showSep ", " (toList $ map show args) ++ ") => " ++ show sc export covering @@ -349,7 +350,7 @@ unload fc _ f [] = pure f -- only outermost LApp must be lazy as rest will be closures unload fc lazy f (a :: as) = unload fc Nothing (LApp fc lazy f a) as -record Used (vars : List Name) where +record Used (vars : ScopedList Name) where constructor MkUsed used : Vect (length vars) Bool @@ -363,19 +364,19 @@ lengthDistributesOverAppend [] ys = Refl lengthDistributesOverAppend (x :: xs) ys = cong S $ lengthDistributesOverAppend xs ys -weakenUsed : {outer : _} -> Used vars -> Used (outer ++ vars) +weakenUsed : {outer : _} -> Used vars -> Used (outer +%+ vars) weakenUsed {outer} (MkUsed xs) = MkUsed (rewrite lengthDistributesOverAppend outer vars in (replicate (length outer) False ++ xs)) -contractUsed : (Used (x::vars)) -> Used vars +contractUsed : (Used (x:%:vars)) -> Used vars contractUsed (MkUsed xs) = MkUsed (tail xs) contractUsedMany : {remove : _} -> - (Used (remove ++ vars)) -> + (Used (remove +%+ vars)) -> Used vars -contractUsedMany {remove=[]} x = x -contractUsedMany {remove=(r::rs)} x = contractUsedMany {remove=rs} (contractUsed x) +contractUsedMany {remove=SLNil} x = x +contractUsedMany {remove=(r:%:rs)} x = contractUsedMany {remove=rs} (contractUsed x) markUsed : {vars : _} -> (idx : Nat) -> @@ -397,21 +398,21 @@ getUnused : Used vars -> getUnused (MkUsed uv) = map not uv total -dropped : (vars : List Name) -> +dropped : (vars : ScopedList Name) -> (drop : Vect (length vars) Bool) -> - List Name -dropped [] _ = [] -dropped (x::xs) (False::us) = x::(dropped xs us) -dropped (x::xs) (True::us) = dropped xs us + ScopedList Name +dropped SLNil _ = SLNil +dropped (x:%:xs) (False::us) = x:%:(dropped xs us) +dropped (x:%:xs) (True::us) = dropped xs us mutual makeLam : {auto l : Ref Lifts LDefs} -> {vars : _} -> {doLazyAnnots : Bool} -> {default Nothing lazy : Maybe LazyReason} -> - FC -> (bound : List Name) -> - CExp (bound ++ vars) -> Core (Lifted vars) - makeLam fc bound (CLam _ x sc') = makeLam fc {doLazyAnnots} {lazy} (x :: bound) sc' + FC -> (bound : ScopedList Name) -> + CExp (bound +%+ vars) -> Core (Lifted vars) + makeLam fc bound (CLam _ x sc') = makeLam fc {doLazyAnnots} {lazy} (x :%: bound) sc' makeLam {vars} fc bound sc = do scl <- liftExp {doLazyAnnots} {lazy} sc -- Find out which variables aren't used in the new definition, and @@ -425,15 +426,15 @@ mutual pure $ LUnderApp fc n (length bound) (allVars fc vars unused) where - allPrfs : (vs : List Name) -> (unused : Vect (length vs) Bool) -> List (Var vs) - allPrfs [] _ = [] - allPrfs (v :: vs) (False::uvs) = MkVar First :: map weaken (allPrfs vs uvs) - allPrfs (v :: vs) (True::uvs) = map weaken (allPrfs vs uvs) + allPrfs : (vs : ScopedList Name) -> (unused : Vect (length vs) Bool) -> List (Var vs) + allPrfs SLNil _ = [] + allPrfs (v :%: vs) (False::uvs) = MkVar First :: map weaken (allPrfs vs uvs) + allPrfs (v :%: vs) (True::uvs) = map weaken (allPrfs vs uvs) -- apply to all the variables. 'First' will be first in the last, which -- is good, because the most recently bound name is the first argument to -- the resulting function - allVars : FC -> (vs : List Name) -> (unused : Vect (length vs) Bool) -> List (Lifted vs) + allVars : FC -> (vs : ScopedList Name) -> (unused : Vect (length vs) Bool) -> List (Lifted vs) allVars fc vs unused = map (\ (MkVar p) => LLocal fc p) (allPrfs vs unused) -- if doLazyAnnots = True then annotate function application with laziness @@ -445,7 +446,7 @@ mutual CExp vars -> Core (Lifted vars) liftExp (CLocal fc prf) = pure $ LLocal fc prf liftExp (CRef fc n) = pure $ LAppName fc lazy n [] -- probably shouldn't happen! - liftExp (CLam fc x sc) = makeLam {doLazyAnnots} {lazy} fc [x] sc + liftExp (CLam fc x sc) = makeLam {doLazyAnnots} {lazy} fc (x :%: SLNil) sc liftExp (CLet fc x _ val sc) = pure $ LLet fc x !(liftExp {doLazyAnnots} val) !(liftExp {doLazyAnnots} sc) liftExp (CApp fc (CRef _ n) args) -- names are applied exactly in compileExp = pure $ LAppName fc lazy n !(traverse (liftExp {doLazyAnnots}) args) @@ -497,7 +498,7 @@ mutual usedVars used (LApp fc lazy c arg) = usedVars (usedVars used arg) c usedVars used (LLet fc x val sc) = - let innerUsed = contractUsed $ usedVars (weakenUsed {outer=[x]} used) sc in + let innerUsed = contractUsed $ usedVars (weakenUsed {outer=(x :%: SLNil)} used) sc in usedVars innerUsed val usedVars used (LCon fc n ci tag args) = foldl (usedVars {vars}) used args @@ -529,24 +530,24 @@ mutual dropIdx : {vars : _} -> {idx : _} -> - (outer : List Name) -> + (outer : ScopedList Name) -> (unused : Vect (length vars) Bool) -> - (0 p : IsVar x idx (outer ++ vars)) -> - Var (outer ++ (dropped vars unused)) - dropIdx [] (False::_) First = MkVar First - dropIdx [] (True::_) First = assert_total $ + (0 p : IsVar x idx (outer +%+ vars)) -> + Var (outer +%+ (dropped vars unused)) + dropIdx SLNil (False::_) First = MkVar First + dropIdx SLNil (True::_) First = assert_total $ idris_crash "INTERNAL ERROR: Referenced variable marked as unused" - dropIdx [] (False::rest) (Later p) = Var.later $ dropIdx [] rest p - dropIdx [] (True::rest) (Later p) = dropIdx [] rest p - dropIdx (_::xs) unused First = MkVar First - dropIdx (_::xs) unused (Later p) = Var.later $ dropIdx xs unused p + dropIdx SLNil (False::rest) (Later p) = Var.later $ dropIdx SLNil rest p + dropIdx SLNil (True::rest) (Later p) = dropIdx SLNil rest p + dropIdx (_:%:xs) unused First = MkVar First + dropIdx (_:%:xs) unused (Later p) = Var.later $ dropIdx xs unused p dropUnused : {vars : _} -> {auto _ : Ref Lifts LDefs} -> - {outer : List Name} -> + {outer : ScopedList Name} -> (unused : Vect (length vars) Bool) -> - (l : Lifted (outer ++ vars)) -> - Lifted (outer ++ (dropped vars unused)) + (l : Lifted (outer +%+ vars)) -> + Lifted (outer +%+ (dropped vars unused)) dropUnused _ (LPrimVal fc val) = LPrimVal fc val dropUnused _ (LErased fc) = LErased fc dropUnused _ (LCrash fc msg) = LCrash fc msg @@ -557,7 +558,7 @@ mutual LCon fc n ci tag args' dropUnused {outer} unused (LLet fc n val sc) = let val' = dropUnused unused val - sc' = dropUnused {outer=n::outer} (unused) sc in + sc' = dropUnused {outer=n:%:outer} (unused) sc in LLet fc n val' sc' dropUnused unused (LApp fc lazy c arg) = let c' = dropUnused unused c @@ -579,18 +580,18 @@ mutual let alts' = map dropConCase alts in LConCase fc (dropUnused unused sc) alts' (map (dropUnused unused) def) where - dropConCase : LiftedConAlt (outer ++ vars) -> - LiftedConAlt (outer ++ (dropped vars unused)) + dropConCase : LiftedConAlt (outer +%+ vars) -> + LiftedConAlt (outer +%+ (dropped vars unused)) dropConCase (MkLConAlt n ci t args sc) = let sc' = (rewrite sym $ appendAssociative args outer vars in sc) - droppedSc = dropUnused {vars=vars} {outer=args++outer} unused sc' in + droppedSc = dropUnused {vars=vars} {outer=args+%+outer} unused sc' in MkLConAlt n ci t args (rewrite appendAssociative args outer (dropped vars unused) in droppedSc) dropUnused {vars} {outer} unused (LConstCase fc sc alts def) = let alts' = map dropConstCase alts in LConstCase fc (dropUnused unused sc) alts' (map (dropUnused unused) def) where - dropConstCase : LiftedConstAlt (outer ++ vars) -> - LiftedConstAlt (outer ++ (dropped vars unused)) + dropConstCase : LiftedConstAlt (outer +%+ vars) -> + LiftedConstAlt (outer +%+ (dropped vars unused)) dropConstCase (MkLConstAlt c val) = MkLConstAlt c (dropUnused unused val) export @@ -606,7 +607,7 @@ export lambdaLiftDef : (doLazyAnnots : Bool) -> Name -> CDef -> Core (List (Name, LiftedDef)) lambdaLiftDef doLazyAnnots n (MkFun args exp) = do (expl, defs) <- liftBody {doLazyAnnots} n exp - pure ((n, MkLFun args [] expl) :: defs) + pure ((n, MkLFun args SLNil expl) :: defs) lambdaLiftDef _ n (MkCon t a nt) = pure [(n, MkLCon t a nt)] lambdaLiftDef _ n (MkForeign ccs fargs ty) = pure [(n, MkLForeign ccs fargs ty)] lambdaLiftDef doLazyAnnots n (MkError exp) diff --git a/src/Compiler/Opts/CSE.idr b/src/Compiler/Opts/CSE.idr index cf1ce54e03..c68a534837 100644 --- a/src/Compiler/Opts/CSE.idr +++ b/src/Compiler/Opts/CSE.idr @@ -48,7 +48,7 @@ import Libraries.Data.SortedMap ||| whether it was encountered in delayed subexpression. public export UsageMap : Type -UsageMap = SortedMap (Integer, CExp []) (Name, Integer, Bool) +UsageMap = SortedMap (Integer, CExp SLNil) (Name, Integer, Bool) ||| Number of appearances of a closed expression. ||| @@ -76,7 +76,7 @@ Show Count where ||| some delayed expression. public export ReplaceMap : Type -ReplaceMap = SortedMap Name (CExp [], Count, Bool) +ReplaceMap = SortedMap Name (CExp SLNil, Count, Bool) toReplaceMap : UsageMap -> ReplaceMap toReplaceMap = SortedMap.fromList @@ -99,7 +99,7 @@ record St where -- returning a new machine generated name to be used -- if the expression should be lifted to the toplevel. -- Very small expressions are being ignored. -store : Ref Sts St => Integer -> CExp [] -> Core (Maybe Name) +store : Ref Sts St => Integer -> CExp SLNil -> Core (Maybe Name) store sz exp = if sz < 5 then pure Nothing @@ -118,13 +118,13 @@ store sz exp = -- Strengthening of Expressions -------------------------------------------------------------------------------- -dropVar : (pre : List Name) +dropVar : (pre : ScopedList Name) -> (n : Nat) - -> (0 p : IsVar x n (pre ++ ns)) + -> (0 p : IsVar x n (pre +%+ ns)) -> Maybe (IsVar x n pre) -dropVar [] _ _ = Nothing -dropVar (y :: xs) 0 First = Just First -dropVar (y :: xs) (S k) (Later p) = +dropVar SLNil _ _ = Nothing +dropVar (y :%: xs) 0 First = Just First +dropVar (y :%: xs) (S k) (Later p) = case dropVar xs k p of Just p' => Just $ Later p' Nothing => Nothing @@ -133,7 +133,7 @@ mutual -- tries to 'strengthen' an expression by removing -- a prefix of bound variables. typically, this is invoked -- with `{pre = []}`. - dropEnv : {pre : List Name} -> CExp (pre ++ ns) -> Maybe (CExp pre) + dropEnv : {pre : ScopedList Name} -> CExp (pre +%+ ns) -> Maybe (CExp pre) dropEnv (CLocal {idx} fc p) = (\q => CLocal fc q) <$> dropVar pre idx p dropEnv (CRef fc x) = Just (CRef fc x) dropEnv (CLam fc x y) = CLam fc x <$> dropEnv y @@ -161,14 +161,14 @@ mutual dropEnv (CErased fc) = Just $ CErased fc dropEnv (CCrash fc x) = Just $ CCrash fc x - dropConAlt : {pre : List Name} - -> CConAlt (pre ++ ns) + dropConAlt : {pre : ScopedList Name} + -> CConAlt (pre +%+ ns) -> Maybe (CConAlt pre) dropConAlt (MkConAlt x y tag args z) = MkConAlt x y tag args . embed <$> dropEnv z - dropConstAlt : {pre : List Name} - -> CConstAlt (pre ++ ns) + dropConstAlt : {pre : ScopedList Name} + -> CConstAlt (pre +%+ ns) -> Maybe (CConstAlt pre) dropConstAlt (MkConstAlt x y) = MkConstAlt x <$> dropEnv y @@ -204,7 +204,7 @@ mutual analyze exp = do (sze, exp') <- analyzeSubExp exp - case dropEnv {pre = []} exp' of + case dropEnv {pre = SLNil} exp' of Just e0 => do Just nm <- store sze e0 | Nothing => pure (sze, exp') @@ -474,12 +474,12 @@ replaceDef (n, fc, d@(MkError _)) = pure (n, fc, d) newToplevelDefs : ReplaceMap -> List (Name, FC, CDef) newToplevelDefs rm = mapMaybe toDef $ SortedMap.toList rm - where toDef : (Name,(CExp[],Count,Bool)) -> Maybe (Name, FC, CDef) - toDef (nm,(exp,Many,False)) = Just (nm, EmptyFC, MkFun [] exp) - toDef (nm,(exp,Many,True)) = Just (nm, EmptyFC, MkFun [] (CDelay EmptyFC LLazy exp)) + where toDef : (Name,(CExp SLNil,Count,Bool)) -> Maybe (Name, FC, CDef) + toDef (nm,(exp,Many,False)) = Just (nm, EmptyFC, MkFun SLNil exp) + toDef (nm,(exp,Many,True)) = Just (nm, EmptyFC, MkFun SLNil (CDelay EmptyFC LLazy exp)) toDef _ = Nothing -undefinedCount : (Name, (CExp [], Count)) -> Bool +undefinedCount : (Name, (CExp SLNil, Count)) -> Bool undefinedCount (_, _, Once) = False undefinedCount (_, _, Many) = False undefinedCount (_, _, C x) = True diff --git a/src/Compiler/Opts/ConstantFold.idr b/src/Compiler/Opts/ConstantFold.idr index ebb5d187a0..68398c54ac 100644 --- a/src/Compiler/Opts/ConstantFold.idr +++ b/src/Compiler/Opts/ConstantFold.idr @@ -6,6 +6,7 @@ import Core.Context.Log import Core.Primitives import Core.Value import Core.Name +import Core.Name.ScopedList import Data.List import Data.Vect @@ -24,29 +25,29 @@ foldableOp (Cast from to) = isJust (intKind from) && isJust (intKind to) foldableOp _ = True -data Subst : List Name -> List Name -> Type where - Nil : Subst [] vars - (::) : CExp vars -> Subst ds vars -> Subst (d :: ds) vars - Wk : SizeOf ws -> Subst ds vars -> Subst (ws ++ ds) (ws ++ vars) +data Subst : ScopedList Name -> ScopedList Name -> Type where + Nil : Subst SLNil vars + (::) : CExp vars -> Subst ds vars -> Subst (d :%: ds) vars + Wk : SizeOf ws -> Subst ds vars -> Subst (ws +%+ ds) (ws +%+ vars) -initSubst : (vars : List Name) -> Subst vars vars +initSubst : (vars : ScopedList Name) -> Subst vars vars initSubst vars = rewrite sym $ appendNilRightNeutral vars in Wk (mkSizeOf vars) [] -wk : SizeOf out -> Subst ds vars -> Subst (out ++ ds) (out ++ vars) +wk : SizeOf out -> Subst ds vars -> Subst (out +%+ ds) (out +%+ vars) wk sout (Wk {ws, ds, vars} sws rho) = rewrite appendAssociative out ws ds in rewrite appendAssociative out ws vars in Wk (sout + sws) rho wk ws rho = Wk ws rho -record WkCExp (vars : List Name) where +record WkCExp (vars : ScopedList Name) where constructor MkWkCExp - {0 outer, supp : List Name} + {0 outer, supp : ScopedList Name} size : SizeOf outer - 0 prf : vars === outer ++ supp + 0 prf : vars === outer +%+ supp expr : CExp supp Weaken WkCExp where @@ -86,19 +87,19 @@ constFold : {vars' : _} -> constFold rho (CLocal fc p) = lookup fc (MkVar p) rho constFold rho e@(CRef fc x) = CRef fc x constFold rho (CLam fc x y) - = CLam fc x $ constFold (wk (mkSizeOf [x]) rho) y + = CLam fc x $ constFold (wk (mkSizeOf (x :%: SLNil)) rho) y constFold rho (CLet fc x inl y z) = let val = constFold rho y in if replace val then constFold (val :: rho) z - else CLet fc x inl val (constFold (wk (mkSizeOf [x]) rho) z) + else CLet fc x inl val (constFold (wk (mkSizeOf (x :%: SLNil)) rho) z) constFold rho (CApp fc (CRef fc2 n) [x]) = if n == NS typesNS (UN $ Basic "prim__integerToNat") then case constFold rho x of CPrimVal fc3 (BI v) => if v >= 0 then CPrimVal fc3 (BI v) else CPrimVal fc3 (BI 0) - v => CApp fc (CRef fc2 n) [v] - else CApp fc (CRef fc2 n) [constFold rho x] + v => CApp fc (CRef fc2 n) (v :: Nil) + else CApp fc (CRef fc2 n) (constFold rho x :: Nil) constFold rho (CApp fc x xs) = CApp fc (constFold rho x) (constFold rho <$> xs) constFold rho (CCon fc x y tag xs) = CCon fc x y tag $ constFold rho <$> xs constFold rho (COp fc BelieveMe [CErased _, CErased _ , x]) = constFold rho x diff --git a/src/Compiler/Opts/Identity.idr b/src/Compiler/Opts/Identity.idr index 0349381e65..c4e8a31e88 100644 --- a/src/Compiler/Opts/Identity.idr +++ b/src/Compiler/Opts/Identity.idr @@ -3,20 +3,21 @@ module Compiler.Opts.Identity import Compiler.CompileExpr import Core.Context import Core.Context.Log +import Core.Name.ScopedList import Data.List import Data.Vect -makeArgs : (args : List Name) -> List (Var (args ++ vars)) +makeArgs : (args : ScopedList Name) -> List (Var (args +%+ vars)) makeArgs args = makeArgs' args id where - makeArgs' : (args : List Name) -> (Var (args ++ vars) -> a) -> List a - makeArgs' [] f = [] - makeArgs' (x :: xs) f = f (MkVar First) :: makeArgs' xs (f . weaken) + makeArgs' : (args : ScopedList Name) -> (Var (args +%+ vars) -> a) -> List a + makeArgs' SLNil f = [] + makeArgs' (x :%: xs) f = f (MkVar First) :: makeArgs' xs (f . weaken) parameters (fn1 : Name) (idIdx : Nat) mutual -- special case for matching on 'Nat'-shaped things - isUnsucc : Var vars -> CExp vars -> Maybe (Constant, Var (x :: vars)) + isUnsucc : Var vars -> CExp vars -> Maybe (Constant, Var (x :%: vars)) isUnsucc var (COp _ (Sub _) [CLocal _ p, CPrimVal _ c]) = if var == MkVar p then Just (c, MkVar First) @@ -113,13 +114,13 @@ checkIdentity fn (v :: vs) exp idx = if cexpIdentity fn idx v Nothing Nothing ex else checkIdentity fn vs exp (S idx) calcIdentity : (fullName : Name) -> CDef -> Maybe Nat -calcIdentity fn (MkFun args exp) = checkIdentity fn (makeArgs {vars=[]} args) (rewrite appendNilRightNeutral args in exp) Z +calcIdentity fn (MkFun args exp) = checkIdentity fn (makeArgs {vars=SLNil} args) (rewrite appendNilRightNeutral args in exp) Z calcIdentity _ _ = Nothing -getArg : FC -> Nat -> (args : List Name) -> Maybe (CExp args) -getArg _ _ [] = Nothing -getArg fc Z (a :: _) = Just $ CLocal fc First -getArg fc (S k) (_ :: as) = weaken <$> getArg fc k as +getArg : FC -> Nat -> (args : ScopedList Name) -> Maybe (CExp args) +getArg _ _ SLNil = Nothing +getArg fc Z (a :%: _) = Just $ CLocal fc First +getArg fc (S k) (_ :%: as) = weaken <$> getArg fc k as idCDef : Nat -> CDef -> Maybe CDef idCDef idx (MkFun args exp) = MkFun args <$> getArg (getFC exp) idx args diff --git a/src/Compiler/VMCode.idr b/src/Compiler/VMCode.idr index b66d0a4043..c693bb4b40 100644 --- a/src/Compiler/VMCode.idr +++ b/src/Compiler/VMCode.idr @@ -173,7 +173,7 @@ toVM t res (AConCase fc (ALocal scr) [MkAConAlt n ci mt args code] Nothing) -- e used = foldMap collectUsed body in projectArgs scr 0 used args ++ body toVM t res (AConCase fc (ALocal scr) alts def) - = [CASE (Loc scr) (map toVMConAlt alts) (map (toVM t res) def)] + = [CASE (Loc scr) (toList $ map toVMConAlt alts) (map (toVM t res) def)] where toVMConAlt : AConAlt -> (Either Int Name, List VMInst) toVMConAlt (MkAConAlt n ci tag args code) @@ -181,7 +181,7 @@ toVM t res (AConCase fc (ALocal scr) alts def) used = foldMap collectUsed body in (maybe (Right n) Left tag, projectArgs scr 0 used args ++ body) toVM t res (AConstCase fc (ALocal scr) alts def) - = [CONSTCASE (Loc scr) (map toVMConstAlt alts) (map (toVM t res) def)] + = [CONSTCASE (Loc scr) (toList $ map toVMConstAlt alts) (map (toVM t res) def)] where toVMConstAlt : AConstAlt -> (Constant, List VMInst) toVMConstAlt (MkAConstAlt c code) diff --git a/src/Core/AutoSearch.idr b/src/Core/AutoSearch.idr index 2c4441c6d3..6b67a15a7f 100644 --- a/src/Core/AutoSearch.idr +++ b/src/Core/AutoSearch.idr @@ -38,7 +38,7 @@ tryNoDefaultsFirst : {auto c : Ref Ctxt Defs} -> (Bool -> Core a) -> Core a tryNoDefaultsFirst f = tryUnifyUnambig {preferLeftError=True} (f False) (f True) -SearchEnv : List Name -> Type +SearchEnv : ScopedList Name -> Type SearchEnv vars = List (NF vars, Closure vars) searchType : {vars : _} -> @@ -52,7 +52,7 @@ searchType : {vars : _} -> Env Term vars -> (target : Term vars) -> Core (Term vars) public export -record ArgInfo (vars : List Name) where +record ArgInfo (vars : ScopedList Name) where constructor MkArgInfo holeID : Int argRig : RigCount @@ -203,17 +203,17 @@ getUsableEnv : {vars : _} -> FC -> RigCount -> SizeOf done -> Env Term vars -> - List (Term (done ++ vars), Term (done ++ vars)) + List (Term (done +%+ vars), Term (done +%+ vars)) getUsableEnv fc rigc p [] = [] -getUsableEnv {vars = v :: vs} {done} fc rigc p (b :: env) +getUsableEnv {vars = v :%: vs} {done} fc rigc p (b :: env) = let rest = getUsableEnv fc rigc (sucR p) env in if (multiplicity b == top || isErased rigc) then let 0 var = mkIsVar (hasLength p) in (Local (binderLoc b) Nothing _ var, - rewrite appendAssociative done [v] vs in + rewrite appendAssociative done (v :%: SLNil) vs in weakenNs (sucR p) (binderType b)) :: - rewrite appendAssociative done [v] vs in rest - else rewrite appendAssociative done [v] vs in rest + rewrite appendAssociative done (v :%: SLNil) vs in rest + else rewrite appendAssociative done (v :%: SLNil) vs in rest -- A local is usable if it contains no holes in a determining argument position usableLocal : {vars : _} -> @@ -227,7 +227,7 @@ usableLocal loc defaults env (NApp fc (NMeta _ _ _) args) = pure False usableLocal {vars} loc defaults env (NTCon _ n _ _ args) = do sd <- getSearchData loc (not defaults) n - usableLocalArg 0 (detArgs sd) (map snd args) + usableLocalArg 0 (detArgs sd) (toList $ map snd args) -- usable if none of the determining arguments of the local's type are -- holes where @@ -318,7 +318,7 @@ searchLocalWith {vars} fc rigc defaults trying depth def top env (prf, ty) targe NF vars -> -- local's type (target : NF vars) -> Core (Term vars) - findPos defs f nty@(NTCon pfc pn _ _ [(_, xty), (_, yty)]) target + findPos defs f nty@(NTCon pfc pn _ _ ((_, xty) :%: (_, yty) :%: SLNil)) target = tryUnifyUnambig (findDirect defs f nty target) $ do fname <- maybe (throw (CantSolveGoal fc (gamma defs) [] top Nothing)) pure @@ -459,11 +459,11 @@ concreteDets {vars} fc defaults env top pos dets (arg :: args) concrete defs argnf True concreteDets fc defaults env top (1 + pos) dets args where - drop : Nat -> List Nat -> List t -> List t - drop i ns [] = [] - drop i ns (x :: xs) + drop : Nat -> List Nat -> ScopedList t -> ScopedList t + drop i ns SLNil = SLNil + drop i ns (x :%: xs) = if i `elem` ns - then x :: drop (1+i) ns xs + then x :%: drop (1+i) ns xs else drop (1+i) ns xs concrete : Defs -> NF vars -> (atTop : Bool) -> Core () @@ -501,19 +501,19 @@ checkConcreteDets fc defaults env top (NTCon tfc tyn t a args) = do defs <- get Ctxt if !(isPairType tyn) then case args of - [(_, aty), (_, bty)] => + ((_, aty) :%: (_, bty) :%: SLNil) => do anf <- evalClosure defs aty bnf <- evalClosure defs bty checkConcreteDets fc defaults env top anf checkConcreteDets fc defaults env top bnf _ => do sd <- getSearchData fc defaults tyn - concreteDets fc defaults env top 0 (detArgs sd) (map snd args) + concreteDets fc defaults env top 0 (detArgs sd) (toList $ map snd args) else do sd <- getSearchData fc defaults tyn log "auto.determining" 10 $ "Determining arguments for " ++ show !(toFullNames tyn) ++ " " ++ show (detArgs sd) - concreteDets fc defaults env top 0 (detArgs sd) (map snd args) + concreteDets fc defaults env top 0 (detArgs sd) (toList $ map snd args) checkConcreteDets fc defaults env top _ = pure () diff --git a/src/Core/Binary/Prims.idr b/src/Core/Binary/Prims.idr index ddd0fccbd2..0bb6d85273 100644 --- a/src/Core/Binary/Prims.idr +++ b/src/Core/Binary/Prims.idr @@ -7,6 +7,7 @@ import Data.Buffer import Data.List import Data.List.Elem import Data.List1 +import Core.Name.ScopedList import Data.Nat import Data.String import Data.Vect @@ -489,3 +490,29 @@ hashFileWith (Just sha256sum) fileName osEscape = if isWindows then id else escapeStringUnix + +export +TTC a => TTC (ScopedList a) where + toBuf b xs + = do toBuf b (TailRec_length xs) + traverse_ (toBuf b) xs + where + ||| Tail-recursive length as buffer sizes can get large + ||| + ||| Once we port to Idris2, can use Data.List.TailRec.length instead + length_aux : ScopedList a -> Int -> Int + length_aux SLNil len = len + length_aux (_ :%: xs) len = length_aux xs (1 + len) + + TailRec_length : ScopedList a -> Int + TailRec_length xs = length_aux xs 0 + + fromBuf b + = do len <- fromBuf b {a = Int} + readElems SLNil (integerToNat (cast len)) + where + readElems : ScopedList a -> Nat -> Core (ScopedList a) + readElems xs Z = pure (reverse xs) + readElems xs (S k) + = do val <- fromBuf b + readElems (val :%: xs) k diff --git a/src/Core/Case/CaseBuilder.idr b/src/Core/Case/CaseBuilder.idr index 9f0286a778..c2326674c0 100644 --- a/src/Core/Case/CaseBuilder.idr +++ b/src/Core/Case/CaseBuilder.idr @@ -9,6 +9,7 @@ import Core.Env import Core.Normalise import Core.Options import Core.TT +import Core.Name.ScopedList import Core.Value import Idris.Pretty.Annotations @@ -35,7 +36,7 @@ Eq Phase where RunTime == RunTime = True _ == _ = False -data ArgType : List Name -> Type where +data ArgType : ScopedList Name -> Type where Known : RigCount -> (ty : Term vars) -> ArgType vars -- arg has type 'ty' Stuck : (fty : Term vars) -> ArgType vars -- ^ arg will have argument type of 'fty' when we know enough to @@ -59,7 +60,7 @@ covering show (Stuck t) = "Stuck " ++ show t show Unknown = "Unknown" -record PatInfo (pvar : Name) (vars : List Name) where +record PatInfo (pvar : Name) (vars : ScopedList Name) where constructor MkInfo {idx : Nat} {name : Name} @@ -94,16 +95,16 @@ NamedPats always have the same 'Elem' proof, though this isn't expressed in a type anywhere. -} -data NamedPats : List Name -> -- pattern variables still to process - List Name -> -- the pattern variables still to process, +data NamedPats : ScopedList Name -> -- pattern variables still to process + ScopedList Name -> -- the pattern variables still to process, -- in order Type where - Nil : NamedPats vars [] + Nil : NamedPats vars SLNil (::) : PatInfo pvar vars -> -- ^ a pattern, where its variable appears in the vars list, -- and its type. The type has no variable names; any names it -- refers to are explicit - NamedPats vars ns -> NamedPats vars (pvar :: ns) + NamedPats vars ns -> NamedPats vars (pvar :%: ns) getPatInfo : NamedPats vars todo -> List Pat getPatInfo [] = [] @@ -114,7 +115,7 @@ updatePats : {vars, todo : _} -> Env Term vars -> NF vars -> NamedPats vars todo -> Core (NamedPats vars todo) updatePats env nf [] = pure [] -updatePats {todo = pvar :: ns} env (NBind fc _ (Pi _ c _ farg) fsc) (p :: ps) +updatePats {todo = pvar :%: ns} env (NBind fc _ (Pi _ c _ farg) fsc) (p :: ps) = case argType p of Unknown => do defs <- get Ctxt @@ -193,9 +194,9 @@ covering where showAll : {vs, ts : _} -> NamedPats vs ts -> String showAll [] = "" - showAll {ts = t :: _ } [x] + showAll {ts = t :%: _ } [x] = show t ++ " " ++ show (pat x) ++ " [" ++ show (argType x) ++ "]" - showAll {ts = t :: _ } (x :: xs) + showAll {ts = t :%: _ } (x :: xs) = show t ++ " " ++ show (pat x) ++ " [" ++ show (argType x) ++ "]" ++ ", " ++ showAll xs @@ -204,7 +205,7 @@ covering where prettyAll : {vs, ts : _} -> NamedPats vs ts -> List (Doc IdrisSyntax) prettyAll [] = [] - prettyAll {ts = t :: _ } (x :: xs) + prettyAll {ts = t :%: _ } (x :: xs) = parens (pretty0 t <++> equals <++> pretty (pat x)) :: prettyAll xs @@ -222,29 +223,29 @@ Weaken (PatInfo p) where -- FIXME: perhaps 'vars' should be second argument so we can use Weaken interface weaken : {x, vars : _} -> - NamedPats vars todo -> NamedPats (x :: vars) todo + NamedPats vars todo -> NamedPats (x :%: vars) todo weaken [] = [] weaken (p :: ps) = weaken p :: weaken ps weakenNs : SizeOf ns -> NamedPats vars todo -> - NamedPats (ns ++ vars) todo + NamedPats (ns +%+ vars) todo weakenNs ns [] = [] weakenNs ns (p :: ps) = weakenNs ns p :: weakenNs ns ps -(++) : NamedPats vars ms -> NamedPats vars ns -> NamedPats vars (ms ++ ns) +(++) : NamedPats vars ms -> NamedPats vars ns -> NamedPats vars (ms +%+ ns) (++) [] ys = ys (++) (x :: xs) ys = x :: xs ++ ys -tail : NamedPats vars (p :: ps) -> NamedPats vars ps +tail : NamedPats vars (p :%: ps) -> NamedPats vars ps tail (x :: xs) = xs -take : (as : List Name) -> NamedPats vars (as ++ bs) -> NamedPats vars as -take [] ps = [] -take (x :: xs) (p :: ps) = p :: take xs ps +take : (as : ScopedList Name) -> NamedPats vars (as +%+ bs) -> NamedPats vars as +take SLNil ps = [] +take (x :%: xs) (p :: ps) = p :: take xs ps -data PatClause : (vars : List Name) -> (todo : List Name) -> Type where +data PatClause : (vars : ScopedList Name) -> (todo : ScopedList Name) -> Type where MkPatClause : List Name -> -- names matched so far (from original lhs) NamedPats vars todo -> Int -> (rhs : Term vars) -> PatClause vars todo @@ -271,8 +272,8 @@ HasNames (PatClause vars todo) where substInClause : {a, vars, todo : _} -> {auto c : Ref Ctxt Defs} -> - FC -> PatClause vars (a :: todo) -> - Core (PatClause vars (a :: todo)) + FC -> PatClause vars (a :%: todo) -> + Core (PatClause vars (a :%: todo)) substInClause {vars} {a} fc (MkPatClause pvars (MkInfo pat pprf fty :: pats) pid rhs) = do pats' <- substInPats fc a (mkTerm vars pat) pats pure (MkPatClause pvars (MkInfo pat pprf fty :: pats') pid rhs) @@ -316,7 +317,7 @@ namesFrom (PDelay _ _ t p) = namesFrom t ++ namesFrom p namesFrom (PLoc _ n) = [n] namesFrom _ = [] -clauseType : Phase -> PatClause vars (a :: as) -> ClauseType +clauseType : Phase -> PatClause vars (a :%: as) -> ClauseType -- If it's irrelevant, a constructor, and there's no names we haven't seen yet -- and don't see later, treat it as a variable -- Or, if we're compiling for runtime we won't be able to split on it, so @@ -327,7 +328,7 @@ clauseType phase (MkPatClause pvars (MkInfo arg _ ty :: rest) pid rhs) where -- used when we are tempted to split on a constructor: is -- this actually a fully applied one? - splitCon : Nat -> List Pat -> ClauseType + splitCon : Nat -> ScopedList Pat -> ClauseType splitCon arity xs = if arity == length xs then ConClause else VarClause @@ -353,7 +354,7 @@ clauseType phase (MkPatClause pvars (MkInfo arg _ ty :: rest) pid rhs) getClauseType phase l _ = clauseType' l partition : {a, as, vars : _} -> - Phase -> (ps : List (PatClause vars (a :: as))) -> Partitions ps + Phase -> (ps : List (PatClause vars (a :%: as))) -> Partitions ps partition phase [] = NoClauses partition phase (x :: xs) with (partition phase xs) partition phase (x :: (cs ++ ps)) | (ConClauses cs rest) @@ -384,16 +385,16 @@ conTypeEq CDelay CDelay = Just Refl conTypeEq (CConst x) (CConst y) = (\xy => cong CConst xy) <$> constantEq x y conTypeEq _ _ = Nothing -data Group : List Name -> -- variables in scope - List Name -> -- pattern variables still to process +data Group : ScopedList Name -> -- variables in scope + ScopedList Name -> -- pattern variables still to process Type where ConGroup : {newargs : _} -> Name -> (tag : Int) -> - List (PatClause (newargs ++ vars) (newargs ++ todo)) -> + List (PatClause (newargs +%+ vars) (newargs +%+ todo)) -> Group vars todo DelayGroup : {tyarg, valarg : _} -> - List (PatClause (tyarg :: valarg :: vars) - (tyarg :: valarg :: todo)) -> + List (PatClause (tyarg :%: valarg :%: vars) + (tyarg :%: valarg :%: todo)) -> Group vars todo ConstGroup : Constant -> List (PatClause vars todo) -> Group vars todo @@ -404,17 +405,17 @@ covering show (DelayGroup cs) = "Delay: " ++ show cs show (ConstGroup c cs) = "Const " ++ show c ++ ": " ++ show cs -data GroupMatch : ConType -> List Pat -> Group vars todo -> Type where +data GroupMatch : ConType -> ScopedList Pat -> Group vars todo -> Type where ConMatch : {tag : Int} -> LengthMatch ps newargs -> GroupMatch (CName n tag) ps (ConGroup {newargs} n tag (MkPatClause pvs pats pid rhs :: rest)) - DelayMatch : GroupMatch CDelay [] + DelayMatch : GroupMatch CDelay SLNil (DelayGroup {tyarg} {valarg} (MkPatClause pvs pats pid rhs :: rest)) - ConstMatch : GroupMatch (CConst c) [] + ConstMatch : GroupMatch (CConst c) SLNil (ConstGroup c (MkPatClause pvs pats pid rhs :: rest)) NoMatch : GroupMatch ct ps g -checkGroupMatch : (c : ConType) -> (ps : List Pat) -> (g : Group vars todo) -> +checkGroupMatch : (c : ConType) -> (ps : ScopedList Pat) -> (g : Group vars todo) -> GroupMatch c ps g checkGroupMatch (CName x tag) ps (ConGroup {newargs} x' tag' (MkPatClause pvs pats pid rhs :: rest)) = case checkLengthMatch ps newargs of @@ -423,9 +424,9 @@ checkGroupMatch (CName x tag) ps (ConGroup {newargs} x' tag' (MkPatClause pvs pa (Just Refl, Yes Refl) => ConMatch prf _ => NoMatch checkGroupMatch (CName x tag) ps _ = NoMatch -checkGroupMatch CDelay [] (DelayGroup (MkPatClause pvs pats pid rhs :: rest)) +checkGroupMatch CDelay SLNil (DelayGroup (MkPatClause pvs pats pid rhs :: rest)) = DelayMatch -checkGroupMatch (CConst c) [] (ConstGroup c' (MkPatClause pvs pats pid rhs :: rest)) +checkGroupMatch (CConst c) SLNil (ConstGroup c' (MkPatClause pvs pats pid rhs :: rest)) = case constantEq c c' of Nothing => NoMatch Just Refl => ConstMatch @@ -443,10 +444,10 @@ nextName root nextNames : {vars : _} -> {auto i : Ref PName Int} -> {auto c : Ref Ctxt Defs} -> - FC -> String -> List Pat -> Maybe (NF vars) -> - Core (args ** (SizeOf args, NamedPats (args ++ vars) args)) -nextNames fc root [] fty = pure ([] ** (zero, [])) -nextNames {vars} fc root (p :: pats) fty + FC -> String -> ScopedList Pat -> Maybe (NF vars) -> + Core (args ** (SizeOf args, NamedPats (args +%+ vars) args)) +nextNames fc root SLNil fty = pure (SLNil ** (zero, [])) +nextNames {vars} fc root (p :%: pats) fty = do defs <- get Ctxt empty <- clearDefs defs n <- nextName root @@ -469,28 +470,32 @@ nextNames {vars} fc root (p :: pats) fty Unknown => Unknown Known rig t => Known rig (weakenNs (suc l) t) Stuck t => Stuck (weakenNs (suc l) t) - pure (n :: args ** (suc l, MkInfo p First argTy :: weaken ps)) + pure (n :%: args ** (suc l, MkInfo p First argTy :: weaken ps)) -- replace the prefix of patterns with 'pargs' -newPats : (pargs : List Pat) -> LengthMatch pargs ns -> - NamedPats vars (ns ++ todo) -> +newPats : (pargs : ScopedList Pat) -> LengthMatch pargs ns -> + NamedPats vars (ns +%+ todo) -> NamedPats vars ns -newPats [] NilMatch rest = [] -newPats (newpat :: xs) (ConsMatch w) (pi :: rest) +newPats SLNil NilMatch rest = [] +newPats (newpat :%: xs) (ConsMatch w) (pi :: rest) = { pat := newpat} pi :: newPats xs w rest -updateNames : List (Name, Pat) -> List (Name, Name) +updateNames : ScopedList (Name, Pat) -> ScopedList (Name, Name) updateNames = mapMaybe update where update : (Name, Pat) -> Maybe (Name, Name) update (n, PLoc fc p) = Just (p, n) update _ = Nothing -updatePatNames : List (Name, Name) -> NamedPats vars todo -> NamedPats vars todo +updatePatNames : ScopedList (Name, Name) -> NamedPats vars todo -> NamedPats vars todo updatePatNames _ [] = [] updatePatNames ns (pi :: ps) = { pat $= update } pi :: updatePatNames ns ps where + lookup : Name -> ScopedList (Name, Name) -> Maybe Name + lookup n SLNil = Nothing + lookup n ((x, n') :%: ns) = if x == n then Just n' else lookup n ns + update : Pat -> Pat update (PAs fc n p) = case lookup n ns of @@ -511,14 +516,14 @@ groupCons : {a, vars, todo : _} -> {auto ct : Ref Ctxt Defs} -> FC -> Name -> List Name -> - List (PatClause vars (a :: todo)) -> + List (PatClause vars (a :%: todo)) -> Core (List (Group vars todo)) groupCons fc fn pvars cs = gc [] cs where addConG : {vars', todo' : _} -> Name -> (tag : Int) -> - List Pat -> NamedPats vars' todo' -> + ScopedList Pat -> NamedPats vars' todo' -> Int -> (rhs : Term vars') -> (acc : List (Group vars' todo')) -> Core (List (Group vars' todo')) @@ -540,7 +545,7 @@ groupCons fc fn pvars cs -- explicit dependencies in types accurate) let pats' = updatePatNames (updateNames (zip patnames pargs)) (weakenNs l pats) - let clause = MkPatClause {todo = patnames ++ todo'} + let clause = MkPatClause {todo = patnames +%+ todo'} pvars (newargs ++ pats') pid (weakenNs l rhs) @@ -553,7 +558,7 @@ groupCons fc fn pvars cs let l = mkSizeOf newargs let pats' = updatePatNames (updateNames (zip newargs pargs)) (weakenNs l pats) - let newclause : PatClause (newargs ++ vars') (newargs ++ todo') + let newclause : PatClause (newargs +%+ vars') (newargs +%+ todo') = MkPatClause pvars (newps ++ pats') pid @@ -581,27 +586,27 @@ groupCons fc fn pvars cs do a' <- evalClosure d a pure (NBind fc (MN "x" 0) (Pi fc top Explicit a) (\dv, av => pure (NDelayed fc LUnknown a')))) - ([tyname, argname] ** (l, newargs)) <- nextNames {vars=vars'} fc "e" [pty, parg] + ((tyname :%: argname :%: SLNil) ** (l, newargs)) <- nextNames {vars=vars'} fc "e" (pty :%: parg :%: SLNil) (Just dty) | _ => throw (InternalError "Error compiling Delay pattern match") - let pats' = updatePatNames (updateNames [(tyname, pty), - (argname, parg)]) + let pats' = updatePatNames (updateNames ((tyname, pty) :%: + (argname, parg) :%: SLNil)) (weakenNs l pats) - let clause = MkPatClause {todo = tyname :: argname :: todo'} + let clause = MkPatClause {todo = tyname :%: argname :%: todo'} pvars (newargs ++ pats') pid (weakenNs l rhs) pure [DelayGroup [clause]] - addDelayG {vars'} {todo'} pty parg pats pid rhs (g :: gs) with (checkGroupMatch CDelay [] g) + addDelayG {vars'} {todo'} pty parg pats pid rhs (g :: gs) with (checkGroupMatch CDelay SLNil g) addDelayG {vars'} {todo'} pty parg pats pid rhs ((DelayGroup {tyarg} {valarg} ((MkPatClause pvars ps tid tm) :: rest)) :: gs) | (DelayMatch {tyarg} {valarg}) - = do let l = mkSizeOf [tyarg, valarg] - let newps = newPats [pty, parg] (ConsMatch (ConsMatch NilMatch)) ps - let pats' = updatePatNames (updateNames [(tyarg, pty), - (valarg, parg)]) + = do let l = mkSizeOf (tyarg :%: valarg :%: SLNil) + let newps = newPats (pty :%: parg :%: SLNil) (ConsMatch (ConsMatch NilMatch)) ps + let pats' = updatePatNames (updateNames ((tyarg, pty) :%: + (valarg, parg) :%: SLNil)) (weakenNs l pats) - let newclause : PatClause (tyarg :: valarg :: vars') - (tyarg :: valarg :: todo') + let newclause : PatClause (tyarg :%: valarg :%: vars') + (tyarg :%: valarg :%: todo') = MkPatClause pvars (newps ++ pats') pid (weakenNs l rhs) pure ((DelayGroup (MkPatClause pvars ps tid tm :: rest ++ [newclause])) @@ -617,7 +622,7 @@ groupCons fc fn pvars cs Core (List (Group vars' todo')) addConstG c pats pid rhs [] = pure [ConstGroup c [MkPatClause pvars pats pid rhs]] - addConstG {todo'} {vars'} c pats pid rhs (g :: gs) with (checkGroupMatch (CConst c) [] g) + addConstG {todo'} {vars'} c pats pid rhs (g :: gs) with (checkGroupMatch (CConst c) SLNil g) addConstG {todo'} {vars'} c pats pid rhs ((ConstGroup c ((MkPatClause pvars ps tid tm) :: rest)) :: gs) | ConstMatch = let newclause : PatClause vars' todo' @@ -646,7 +651,7 @@ groupCons fc fn pvars cs then addConG n 0 pargs pats pid rhs acc else throw (CaseCompile cfc fn (NotFullyApplied n)) addGroup (PArrow _ _ s t) pprf pats pid rhs acc - = addConG (UN $ Basic "->") 0 [s, t] pats pid rhs acc + = addConG (UN $ Basic "->") 0 (s :%: t :%: SLNil) pats pid rhs acc -- Go inside the delay; we'll flag the case as needing to force its -- scrutinee (need to check in 'caseGroups below) addGroup (PDelay _ _ pty parg) pprf pats pid rhs acc @@ -658,49 +663,49 @@ groupCons fc fn pvars cs gc : {a, vars, todo : _} -> List (Group vars todo) -> - List (PatClause vars (a :: todo)) -> + List (PatClause vars (a :%: todo)) -> Core (List (Group vars todo)) gc acc [] = pure acc gc {a} acc ((MkPatClause pvars (MkInfo pat pprf fty :: pats) pid rhs) :: cs) = do acc' <- addGroup pat pprf pats pid rhs acc gc acc' cs -getFirstPat : NamedPats ns (p :: ps) -> Pat +getFirstPat : NamedPats ns (p :%: ps) -> Pat getFirstPat (p :: _) = pat p -getFirstArgType : NamedPats ns (p :: ps) -> ArgType ns +getFirstArgType : NamedPats ns (p :%: ps) -> ArgType ns getFirstArgType (p :: _) = argType p ||| Store scores alongside rows of named patterns. These scores are used to determine ||| which column of patterns to switch on first. One score per column. -data ScoredPats : List Name -> List Name -> Type where - Scored : List (NamedPats ns (p :: ps)) -> Vect (length (p :: ps)) Int -> ScoredPats ns (p :: ps) +data ScoredPats : ScopedList Name -> ScopedList Name -> Type where + Scored : List (NamedPats ns (p :%: ps)) -> Vect (length (p :%: ps)) Int -> ScoredPats ns (p :%: ps) {ps : _} -> Show (ScoredPats ns ps) where show (Scored xs ys) = (show ps) ++ "//" ++ (show ys) -zeroedScore : {ps : _} -> List (NamedPats ns (p :: ps)) -> ScoredPats ns (p :: ps) +zeroedScore : {ps : _} -> List (NamedPats ns (p :%: ps)) -> ScoredPats ns (p :%: ps) zeroedScore nps = Scored nps (replicate (S $ length ps) 0) ||| Proof that a value `v` inserted in the middle of a list with ||| prefix `ps` and suffix `qs` can equivalently be snoced with ||| `ps` or consed with `qs` before appending `qs` to `ps`. -elemInsertedMiddle : (v : a) -> (ps,qs : List a) -> (ps ++ (v :: qs)) = ((ps `snoc` v) ++ qs) -elemInsertedMiddle v [] qs = Refl -elemInsertedMiddle v (x :: xs) qs = rewrite elemInsertedMiddle v xs qs in Refl +elemInsertedMiddle : (v : a) -> (ps,qs : ScopedList a) -> (ps +%+ (v :%: qs)) = ((ps `snoc` v) +%+ qs) +elemInsertedMiddle v SLNil qs = Refl +elemInsertedMiddle v (x :%: xs) qs = rewrite elemInsertedMiddle v xs qs in Refl ||| Helper to find a single highest scoring name (or none at all) while ||| retaining the context of all names processed. -highScore : {prev : List Name} -> - (names : List Name) -> +highScore : {prev : ScopedList Name} -> + (names : ScopedList Name) -> (scores : Vect (length names) Int) -> (highVal : Int) -> - (highIdx : (n ** NVar n (prev ++ names))) -> + (highIdx : (n ** NVar n (prev +%+ names))) -> (duped : Bool) -> - Maybe (n ** NVar n (prev ++ names)) -highScore [] [] high idx True = Nothing -highScore [] [] high idx False = Just idx -highScore (x :: xs) (y :: ys) high idx duped = + Maybe (n ** NVar n (prev +%+ names)) +highScore SLNil [] high idx True = Nothing +highScore SLNil [] high idx False = Just idx +highScore (x :%: xs) (y :: ys) high idx duped = let next = highScore {prev = prev `snoc` x} xs ys prf = elemInsertedMiddle x prev xs in rewrite prf in @@ -714,8 +719,8 @@ highScore (x :: xs) (y :: ys) high idx duped = ||| the result is Nothing indicating we need to apply more scoring ||| to break the tie. ||| Suggested heuristic application order: f, b, a. -highScoreIdx : {p : _} -> {ps : _} -> ScoredPats ns (p :: ps) -> Maybe (n ** NVar n (p :: ps)) -highScoreIdx (Scored xs (y :: ys)) = highScore {prev = []} (p :: ps) (y :: ys) (y - 1) (p ** MkNVar First) False +highScoreIdx : {p : _} -> {ps : _} -> ScoredPats ns (p :%: ps) -> Maybe (n ** NVar n (p :%: ps)) +highScoreIdx (Scored xs (y :: ys)) = highScore {prev = SLNil} (p :%: ps) (y :: ys) (y - 1) (p ** MkNVar First) False ||| Apply the penalty function to the head constructor's ||| arity. Produces 0 for all non-head-constructors. @@ -740,21 +745,21 @@ consScoreHeuristic scorePat (Scored xs ys) = where -- also returns NamePats of remaining columns while its in there -- scoring the first column. - scoreFirstColumn : (nps : List (NamedPats ns (p' :: ps'))) -> (res : List (NamedPats ns ps') ** (LengthMatch nps res, Vect (length nps) Int)) + scoreFirstColumn : (nps : List (NamedPats ns (p' :%: ps'))) -> (res : List (NamedPats ns ps') ** (LengthMatch nps res, Vect (length nps) Int)) scoreFirstColumn [] = ([] ** (NilMatch, [])) scoreFirstColumn ((w :: ws) :: nps) = let (ws' ** (prf, scores)) = scoreFirstColumn nps in (ws :: ws' ** (ConsMatch prf, scorePat (pat w) :: scores)) scoreColumns : {ps' : _} -> (nps : List (NamedPats ns ps')) -> Vect (length ps') (Vect (length nps) Int) - scoreColumns {ps' = []} nps = [] - scoreColumns {ps' = (w :: ws)} nps = + scoreColumns {ps' = SLNil} nps = [] + scoreColumns {ps' = (w :%: ws)} nps = let (rest ** (prf, firstColScore)) = scoreFirstColumn nps in firstColScore :: (rewrite lengthsMatch prf in scoreColumns rest) ||| Add 1 to each non-default pat in the first row. ||| This favors constructive matching first and reduces tree depth on average. -heuristicF : {ps : _} -> ScoredPats ns (p :: ps) -> ScoredPats ns (p :: ps) +heuristicF : {ps : _} -> ScoredPats ns (p :%: ps) -> ScoredPats ns (p :%: ps) heuristicF sps@(Scored [] _) = sps heuristicF (Scored (x :: xs) ys) = let columnScores = scores x @@ -781,9 +786,9 @@ heuristicA = consScoreHeuristic (headConsPenalty (negate . cast)) applyHeuristics : {p : _} -> {ps : _} -> - ScoredPats ns (p :: ps) -> - List (ScoredPats ns (p :: ps) -> ScoredPats ns (p :: ps)) -> - Maybe (n ** NVar n (p :: ps)) + ScoredPats ns (p :%: ps) -> + List (ScoredPats ns (p :%: ps) -> ScoredPats ns (p :%: ps)) -> + Maybe (n ** NVar n (p :%: ps)) applyHeuristics x [] = highScoreIdx x applyHeuristics x (f :: fs) = highScoreIdx x <|> applyHeuristics (f x) fs @@ -796,8 +801,8 @@ nextIdxByScore : {p : _} -> {ps : _} -> (useHeuristics : Bool) -> Phase -> - List (NamedPats ns (p :: ps)) -> - (n ** NVar n (p :: ps)) + List (NamedPats ns (p :%: ps)) -> + (n ** NVar n (p :%: ps)) nextIdxByScore False _ _ = (_ ** (MkNVar First)) nextIdxByScore _ (CompileTime _) _ = (_ ** (MkNVar First)) nextIdxByScore True RunTime xs = @@ -811,7 +816,7 @@ sameType : {ns : _} -> {auto i : Ref PName Int} -> {auto c : Ref Ctxt Defs} -> FC -> Phase -> Name -> - Env Term ns -> List (NamedPats ns (p :: ps)) -> + Env Term ns -> List (NamedPats ns (p :%: ps)) -> Core () sameType fc phase fn env [] = pure () sameType {ns} fc phase fn env (p :: xs) @@ -822,7 +827,7 @@ sameType {ns} fc phase fn env (p :: xs) (map getFirstArgType xs) ty => throw (CaseCompile fc fn DifferingTypes) where - firstPat : NamedPats ns (np :: nps) -> Pat + firstPat : NamedPats ns (np :%: nps) -> Pat firstPat (pinf :: _) = pat pinf headEq : NF ns -> NF ns -> Phase -> Bool @@ -846,7 +851,7 @@ sameType {ns} fc phase fn env (p :: xs) -- Check whether all the initial patterns are the same, or are all a variable. -- If so, we'll match it to refine later types and move on -samePat : List (NamedPats ns (p :: ps)) -> Bool +samePat : List (NamedPats ns (p :%: ps)) -> Bool samePat [] = True samePat (pi :: xs) = samePatAs (dropAs (getFirstPat pi)) @@ -871,11 +876,11 @@ samePat (pi :: xs) samePatAs (PLoc fc n) (PLoc _ _ :: ps) = samePatAs (PLoc fc n) ps samePatAs x y = False -getFirstCon : NamedPats ns (p :: ps) -> Pat +getFirstCon : NamedPats ns (p :%: ps) -> Pat getFirstCon (p :: _) = pat p -- Count the number of distinct constructors in the initial pattern -countDiff : List (NamedPats ns (p :: ps)) -> Nat +countDiff : List (NamedPats ns (p :%: ps)) -> Nat countDiff xs = length (distinct [] (map getFirstCon xs)) where isVar : Pat -> Bool @@ -909,7 +914,7 @@ getScore : {ns : _} -> {auto i : Ref PName Int} -> {auto c : Ref Ctxt Defs} -> FC -> Phase -> Name -> - List (NamedPats ns (p :: ps)) -> + List (NamedPats ns (p :%: ps)) -> Core (Either CaseError ()) getScore fc phase name npss = do catch (do sameType fc phase name (mkEnv fc ns) npss @@ -923,16 +928,16 @@ getScore fc phase name npss pickNextViable : {p, ns, ps : _} -> {auto i : Ref PName Int} -> {auto c : Ref Ctxt Defs} -> - FC -> Phase -> Name -> List (NamedPats ns (p :: ps)) -> - Core (n ** NVar n (p :: ps)) + FC -> Phase -> Name -> List (NamedPats ns (p :%: ps)) -> + Core (n ** NVar n (p :%: ps)) -- last possible variable -pickNextViable {ps = []} fc phase fn npss +pickNextViable {ps = SLNil} fc phase fn npss = if samePat npss then pure (_ ** MkNVar First) else do Right () <- getScore fc phase fn npss | Left err => throw (CaseCompile fc fn err) pure (_ ** MkNVar First) -pickNextViable {ps = q :: qs} fc phase fn npss +pickNextViable {ps = q :%: qs} fc phase fn npss = if samePat npss then pure (_ ** MkNVar First) else case !(getScore fc phase fn npss) of @@ -941,11 +946,11 @@ pickNextViable {ps = q :: qs} fc phase fn npss pure (_ ** MkNVar (Later var)) moveFirst : {idx : Nat} -> (0 el : IsVar nm idx ps) -> NamedPats ns ps -> - NamedPats ns (nm :: dropIsVar ps el) + NamedPats ns (nm :%: dropIsVar ps el) moveFirst el nps = getPat el nps :: dropPat el nps shuffleVars : {idx : Nat} -> (0 el : IsVar nm idx todo) -> PatClause vars todo -> - PatClause vars (nm :: dropIsVar todo el) + PatClause vars (nm :%: dropIsVar todo el) shuffleVars First orig@(MkPatClause pvars lhs pid rhs) = orig -- no-op shuffleVars el (MkPatClause pvars lhs pid rhs) = MkPatClause pvars (moveFirst el lhs) pid rhs @@ -966,7 +971,7 @@ mutual -- Before 'partition', reorder the arguments so that the one we -- inspect next has a concrete type that is the same in all cases, and -- has the most distinct constructors (via pickNextViable) - match {todo = (_ :: _)} fc fn phase clauses err + match {todo = (_ :%: _)} fc fn phase clauses err = do let nps = getNPs <$> clauses let (_ ** (MkNVar next)) = nextIdxByScore (caseTreeHeuristics !getSession) phase nps let prioritizedClauses = shuffleVars next <$> clauses @@ -985,12 +990,12 @@ mutual Just m => do log "compile.casetree.intermediate" 25 $ "match: new case tree " ++ show m Core.pure m - match {todo = []} fc fn phase [] err + match {todo = SLNil} fc fn phase [] err = maybe (pure (Unmatched "No patterns")) pure err - match {todo = []} fc fn phase ((MkPatClause pvars [] pid (Erased _ Impossible)) :: _) err + match {todo = SLNil} fc fn phase ((MkPatClause pvars [] pid (Erased _ Impossible)) :: _) err = pure Impossible - match {todo = []} fc fn phase ((MkPatClause pvars [] pid rhs) :: _) err + match {todo = SLNil} fc fn phase ((MkPatClause pvars [] pid rhs) :: _) err = pure $ STerm pid rhs caseGroups : {pvar, vars, todo : _} -> @@ -1013,7 +1018,7 @@ mutual cs' <- altGroups cs pure (ConCase cn tag newargs crest :: cs') altGroups (DelayGroup {tyarg} {valarg} rest :: cs) - = do crest <- match fc fn phase rest (map (weakenNs (mkSizeOf [tyarg, valarg])) errorCase) + = do crest <- match fc fn phase rest (map (weakenNs (mkSizeOf (tyarg :%: valarg :%: SLNil))) errorCase) cs' <- altGroups cs pure (DelayCase tyarg valarg crest :: cs') altGroups (ConstGroup c rest :: cs) @@ -1025,7 +1030,7 @@ mutual {auto i : Ref PName Int} -> {auto c : Ref Ctxt Defs} -> FC -> Name -> Phase -> - List (PatClause vars (a :: todo)) -> + List (PatClause vars (a :%: todo)) -> Maybe (CaseTree vars) -> Core (CaseTree vars) conRule fc fn phase [] err = maybe (pure (Unmatched "No constructor clauses")) pure err @@ -1045,14 +1050,14 @@ mutual {auto i : Ref PName Int} -> {auto c : Ref Ctxt Defs} -> FC -> Name -> Phase -> - List (PatClause vars (a :: todo)) -> + List (PatClause vars (a :%: todo)) -> Maybe (CaseTree vars) -> Core (CaseTree vars) varRule {vars} {a} fc fn phase cs err = do alts' <- traverse updateVar cs match fc fn phase alts' err where - updateVar : PatClause vars (a :: todo) -> Core (PatClause vars todo) + updateVar : PatClause vars (a :%: todo) -> Core (PatClause vars todo) -- replace the name with the relevant variable on the rhs updateVar (MkPatClause pvars (MkInfo (PLoc pfc n) prf fty :: pats) pid rhs) = pure $ MkPatClause (n :: pvars) @@ -1073,7 +1078,7 @@ mutual mixture : {a, vars, todo : _} -> {auto i : Ref PName Int} -> {auto c : Ref Ctxt Defs} -> - {ps : List (PatClause vars (a :: todo))} -> + {ps : List (PatClause vars (a :%: todo))} -> FC -> Name -> Phase -> Partitions ps -> Maybe (CaseTree vars) -> @@ -1088,8 +1093,8 @@ mutual = pure err export -mkPat : {auto c : Ref Ctxt Defs} -> List Pat -> ClosedTerm -> ClosedTerm -> Core Pat -mkPat [] orig (Ref fc Bound n) = pure $ PLoc fc n +mkPat : {auto c : Ref Ctxt Defs} -> ScopedList Pat -> ClosedTerm -> ClosedTerm -> Core Pat +mkPat SLNil orig (Ref fc Bound n) = pure $ PLoc fc n mkPat args orig (Ref fc (DataCon t a) n) = pure $ PCon fc n t a args mkPat args orig (Ref fc (TyCon t a) n) = pure $ PTyCon fc n a args mkPat args orig (Ref fc Func n) @@ -1099,7 +1104,7 @@ mkPat args orig (Ref fc Func n) Just tm => if tm /= orig -- check we made progress; if there's an -- unresolved interface, we might be stuck -- and we'd loop forever - then mkPat [] tm tm + then mkPat SLNil tm tm else -- Possibly this should be an error instead? pure $ PUnmatchable (getLoc orig) orig Nothing => @@ -1109,21 +1114,21 @@ mkPat args orig (Ref fc Func n) mkPat args orig (Bind fc x (Pi _ _ _ s) t) -- For (b:Nat) -> b, the codomain looks like b [__], but we want `b` as the pattern = case subst (Erased fc Placeholder) t of - App _ t'@(Ref fc Bound n) (Erased _ _) => pure $ PArrow fc x !(mkPat [] s s) !(mkPat [] t' t') - t' => pure $ PArrow fc x !(mkPat [] s s) !(mkPat [] t' t') + App _ t'@(Ref fc Bound n) (Erased _ _) => pure $ PArrow fc x !(mkPat SLNil s s) !(mkPat SLNil t' t') + t' => pure $ PArrow fc x !(mkPat SLNil s s) !(mkPat SLNil t' t') mkPat args orig (App fc fn arg) - = do parg <- mkPat [] arg arg - mkPat (parg :: args) orig fn + = do parg <- mkPat SLNil arg arg + mkPat (parg :%: args) orig fn mkPat args orig (As fc _ (Ref _ Bound n) ptm) - = pure $ PAs fc n !(mkPat [] ptm ptm) + = pure $ PAs fc n !(mkPat SLNil ptm ptm) mkPat args orig (As fc _ _ ptm) - = mkPat [] orig ptm + = mkPat SLNil orig ptm mkPat args orig (TDelay fc r ty p) - = pure $ PDelay fc r !(mkPat [] orig ty) !(mkPat [] orig p) + = pure $ PDelay fc r !(mkPat SLNil orig ty) !(mkPat SLNil orig p) mkPat args orig (PrimVal fc $ PrT c) -- Primitive type constant - = pure $ PTyCon fc (UN (Basic $ show c)) 0 [] + = pure $ PTyCon fc (UN (Basic $ show c)) 0 SLNil mkPat args orig (PrimVal fc c) = pure $ PConst fc c -- Non-type constant -mkPat args orig (TType fc _) = pure $ PTyCon fc (UN $ Basic "Type") 0 [] +mkPat args orig (TType fc _) = pure $ PTyCon fc (UN $ Basic "Type") 0 SLNil mkPat args orig tm = do log "compile.casetree" 10 $ "Catchall: marking " ++ show tm ++ " as unmatchable" @@ -1131,13 +1136,13 @@ mkPat args orig tm export argToPat : {auto c : Ref Ctxt Defs} -> ClosedTerm -> Core Pat -argToPat tm = mkPat [] tm tm +argToPat tm = mkPat SLNil tm tm mkPatClause : {auto c : Ref Ctxt Defs} -> FC -> Name -> - (args : List Name) -> ClosedTerm -> - Int -> (List Pat, ClosedTerm) -> + (args : ScopedList Name) -> ClosedTerm -> + Int -> (ScopedList Pat, ClosedTerm) -> Core (PatClause args args) mkPatClause fc fn args ty pid (ps, rhs) = maybe (throw (CaseCompile fc fn DifferingArgNumbers)) @@ -1153,11 +1158,11 @@ mkPatClause fc fn args ty pid (ps, rhs) (weakenNs (mkSizeOf args) rhs)))) (checkLengthMatch args ps) where - mkNames : (vars : List Name) -> (ps : List Pat) -> - LengthMatch vars ps -> Maybe (NF []) -> + mkNames : (vars : ScopedList Name) -> (ps : ScopedList Pat) -> + LengthMatch vars ps -> Maybe (NF SLNil) -> Core (NamedPats vars vars) - mkNames [] [] NilMatch fty = pure [] - mkNames (arg :: args) (p :: ps) (ConsMatch eq) fty + mkNames SLNil SLNil NilMatch fty = pure [] + mkNames (arg :%: args) (p :%: ps) (ConsMatch eq) fty = do defs <- get Ctxt empty <- clearDefs defs fa_tys <- the (Core (Maybe _, ArgType _)) $ @@ -1165,11 +1170,11 @@ mkPatClause fc fn args ty pid (ps, rhs) Nothing => pure (Nothing, CaseBuilder.Unknown) Just (NBind pfc _ (Pi _ c _ farg) fsc) => pure (Just !(fsc defs (toClosure defaultOpts [] (Ref pfc Bound arg))), - Known c (embed {outer = arg :: args} + Known c (embed {outer = arg :%: args} !(quote empty [] farg))) Just t => pure (Nothing, - Stuck (embed {outer = arg :: args} + Stuck (embed {outer = arg :%: args} !(quote empty [] t))) pure (MkInfo p First (Builtin.snd fa_tys) :: weaken !(mkNames args ps eq (Builtin.fst fa_tys))) @@ -1177,12 +1182,12 @@ mkPatClause fc fn args ty pid (ps, rhs) export patCompile : {auto c : Ref Ctxt Defs} -> FC -> Name -> Phase -> - ClosedTerm -> List (List Pat, ClosedTerm) -> - Maybe (CaseTree []) -> + ClosedTerm -> List (ScopedList Pat, ClosedTerm) -> + Maybe (CaseTree SLNil) -> Core (args ** CaseTree args) patCompile fc fn phase ty [] def - = maybe (pure ([] ** Unmatched "No definition")) - (\e => pure ([] ** e)) + = maybe (pure (SLNil ** Unmatched "No definition")) + (\e => pure (SLNil ** e)) def patCompile fc fn phase ty (p :: ps) def = do let (ns ** n) = getNames 0 (fst p) @@ -1200,8 +1205,8 @@ patCompile fc fn phase ty (p :: ps) def map (weakenNs n) def) pure (_ ** cases) where - mkPatClausesFrom : Int -> (args : List Name) -> - List (List Pat, ClosedTerm) -> + mkPatClausesFrom : Int -> (args : ScopedList Name) -> + List (ScopedList Pat, ClosedTerm) -> Core (List (PatClause args args)) mkPatClausesFrom i ns [] = pure [] mkPatClausesFrom i ns (p :: ps) @@ -1209,15 +1214,15 @@ patCompile fc fn phase ty (p :: ps) def ps' <- mkPatClausesFrom (i + 1) ns ps pure (p' :: ps') - getNames : Int -> List Pat -> (ns : List Name ** SizeOf ns) - getNames i [] = ([] ** zero) - getNames i (x :: xs) = + getNames : Int -> ScopedList Pat -> (ns : ScopedList Name ** SizeOf ns) + getNames i SLNil = (SLNil ** zero) + getNames i (x :%: xs) = let (ns ** n) = getNames (i + 1) xs - in (MN "arg" i :: ns ** suc n) + in (MN "arg" i :%: ns ** suc n) toPatClause : {auto c : Ref Ctxt Defs} -> FC -> Name -> (ClosedTerm, ClosedTerm) -> - Core (List Pat, ClosedTerm) + Core (ScopedList Pat, ClosedTerm) toPatClause fc n (lhs, rhs) = case getFnArgs lhs of (Ref ffc Func fn, args) @@ -1225,7 +1230,7 @@ toPatClause fc n (lhs, rhs) (np, _) <- getPosition n (gamma defs) (fnp, _) <- getPosition fn (gamma defs) if np == fnp - then pure (!(traverse argToPat args), rhs) + then pure (!(traverse argToPat (fromList args)), rhs) else throw (GenericMsg ffc ("Wrong function name in pattern LHS " ++ show (n, fn))) (f, args) => throw (GenericMsg fc "Not a function name in pattern LHS") @@ -1234,7 +1239,7 @@ toPatClause fc n (lhs, rhs) -- the names of the top level variables we created are returned in 'args' export simpleCase : {auto c : Ref Ctxt Defs} -> - FC -> Phase -> Name -> ClosedTerm -> (def : Maybe (CaseTree [])) -> + FC -> Phase -> Name -> ClosedTerm -> (def : Maybe (CaseTree SLNil)) -> (clauses : List (ClosedTerm, ClosedTerm)) -> Core (args ** CaseTree args) simpleCase fc phase fn ty def clauses @@ -1345,12 +1350,12 @@ getPMDef fc phase fn ty [] defs <- get Ctxt pure (!(getArgs 0 !(nf defs [] ty)) ** (Unmatched "No clauses", [])) where - getArgs : Int -> NF [] -> Core (List Name) + getArgs : Int -> NF SLNil -> Core (ScopedList Name) getArgs i (NBind fc x (Pi _ _ _ _) sc) = do defs <- get Ctxt sc' <- sc defs (toClosure defaultOpts [] (Erased fc Placeholder)) - pure (MN "arg" i :: !(getArgs i sc')) - getArgs i _ = pure [] + pure (MN "arg" i :%: !(getArgs i sc')) + getArgs i _ = pure SLNil getPMDef fc phase fn ty clauses = do defs <- get Ctxt let cs = map (toClosed defs) (labelPat 0 clauses) diff --git a/src/Core/Case/CaseTree.idr b/src/Core/Case/CaseTree.idr index c11b001f5b..2b26039ade 100644 --- a/src/Core/Case/CaseTree.idr +++ b/src/Core/Case/CaseTree.idr @@ -1,6 +1,7 @@ module Core.Case.CaseTree import Core.TT +import Core.Name.ScopedList import Data.List import Data.String @@ -16,7 +17,7 @@ mutual ||| Case trees in A-normal forms ||| i.e. we may only dispatch on variables, not expressions public export - data CaseTree : List Name -> Type where + data CaseTree : ScopedList Name -> Type where ||| case x return scTy of { p1 => e1 ; ... } Case : {name : _} -> (idx : Nat) -> @@ -35,13 +36,13 @@ mutual ||| Case alternatives. Unlike arbitrary patterns, they can be at most ||| one constructor deep. public export - data CaseAlt : List Name -> Type where + data CaseAlt : ScopedList Name -> Type where ||| Constructor for a data type; bind the arguments and subterms. - ConCase : Name -> (tag : Int) -> (args : List Name) -> - CaseTree (args ++ vars) -> CaseAlt vars + ConCase : Name -> (tag : Int) -> (args : ScopedList Name) -> + CaseTree (args +%+ vars) -> CaseAlt vars ||| Lazy match for the Delay type use for codata types DelayCase : (ty : Name) -> (arg : Name) -> - CaseTree (ty :: arg :: vars) -> CaseAlt vars + CaseTree (ty :%: arg :%: vars) -> CaseAlt vars ||| Match against a literal ConstCase : Constant -> CaseTree vars -> CaseAlt vars ||| Catch-all case @@ -96,14 +97,14 @@ public export data Pat : Type where PAs : FC -> Name -> Pat -> Pat PCon : FC -> Name -> (tag : Int) -> (arity : Nat) -> - List Pat -> Pat - PTyCon : FC -> Name -> (arity : Nat) -> List Pat -> Pat + ScopedList Pat -> Pat + PTyCon : FC -> Name -> (arity : Nat) -> ScopedList Pat -> Pat PConst : FC -> (c : Constant) -> Pat PArrow : FC -> (x : Name) -> Pat -> Pat -> Pat PDelay : FC -> LazyReason -> Pat -> Pat -> Pat -- TODO: Matching on lazy types PLoc : FC -> Name -> Pat - PUnmatchable : FC -> Term [] -> Pat + PUnmatchable : FC -> Term SLNil -> Pat export isPConst : Pat -> Maybe Constant @@ -117,14 +118,14 @@ showCT indent (Case {name} idx prf ty alts) = "case " ++ show name ++ "[" ++ show idx ++ "] : " ++ show ty ++ " of" ++ "\n" ++ indent ++ " { " ++ showSep ("\n" ++ indent ++ " | ") - (assert_total (map (showCA (" " ++ indent)) alts)) + (assert_total (toList (map (showCA (" " ++ indent)) alts))) ++ "\n" ++ indent ++ " }" showCT indent (STerm i tm) = "[" ++ show i ++ "] " ++ show tm showCT indent (Unmatched msg) = "Error: " ++ show msg showCT indent Impossible = "Impossible" showCA indent (ConCase n tag args sc) - = showSep " " (map show (n :: args)) ++ " => " ++ + = showSep " " (toList (map show (n :%: args))) ++ " => " ++ showCT indent sc showCA indent (DelayCase _ arg sc) = "Delay " ++ show arg ++ " => " ++ showCT indent sc @@ -182,9 +183,9 @@ export Pretty IdrisSyntax Pat where prettyPrec d (PAs _ n p) = pretty0 n <++> keyword "@" <+> parens (pretty p) prettyPrec d (PCon _ n _ _ args) = - parenthesise (d > Open) $ hsep (pretty0 n :: map (prettyPrec App) args) + parenthesise (d > Open) $ hsep (pretty0 n :: map (prettyPrec App) (toList args)) prettyPrec d (PTyCon _ n _ args) = - parenthesise (d > Open) $ hsep (pretty0 n :: map (prettyPrec App) args) + parenthesise (d > Open) $ hsep (pretty0 n :: map (prettyPrec App) (toList args)) prettyPrec d (PConst _ c) = pretty c prettyPrec d (PArrow _ _ p q) = parenthesise (d > Open) $ pretty p <++> arrow <++> pretty q @@ -195,8 +196,8 @@ Pretty IdrisSyntax Pat where mutual insertCaseNames : SizeOf outer -> SizeOf ns -> - CaseTree (outer ++ inner) -> - CaseTree (outer ++ (ns ++ inner)) + CaseTree (outer +%+ inner) -> + CaseTree (outer +%+ (ns +%+ inner)) insertCaseNames outer ns (Case idx prf scTy alts) = let MkNVar prf' = insertNVarNames outer ns (MkNVar prf) in Case _ prf' (insertNames outer ns scTy) @@ -207,11 +208,11 @@ mutual insertCaseAltNames : SizeOf outer -> SizeOf ns -> - CaseAlt (outer ++ inner) -> - CaseAlt (outer ++ (ns ++ inner)) + CaseAlt (outer +%+ inner) -> + CaseAlt (outer +%+ (ns +%+ inner)) insertCaseAltNames p q (ConCase x tag args ct) = ConCase x tag args - (rewrite appendAssociative args outer (ns ++ inner) in + (rewrite appendAssociative args outer (ns +%+ inner) in insertCaseNames (mkSizeOf args + p) q {inner} (rewrite sym (appendAssociative args outer inner) in ct)) @@ -262,17 +263,17 @@ getMetas : CaseTree vars -> NameMap Bool getMetas = getNames (addMetas False) empty export -mkTerm : (vars : List Name) -> Pat -> Term vars +mkTerm : (vars : ScopedList Name) -> Pat -> Term vars mkTerm vars (PAs fc x y) = mkTerm vars y mkTerm vars (PCon fc x tag arity xs) = apply fc (Ref fc (DataCon tag arity) x) - (map (mkTerm vars) xs) + (toList $ map (mkTerm vars) xs) mkTerm vars (PTyCon fc x arity xs) = apply fc (Ref fc (TyCon 0 arity) x) - (map (mkTerm vars) xs) + (toList $ map (mkTerm vars) xs) mkTerm vars (PConst fc c) = PrimVal fc c mkTerm vars (PArrow fc x s t) - = Bind fc x (Pi fc top Explicit (mkTerm vars s)) (mkTerm (x :: vars) t) + = Bind fc x (Pi fc top Explicit (mkTerm vars s)) (mkTerm (x :%: vars) t) mkTerm vars (PDelay fc r ty p) = TDelay fc r (mkTerm vars ty) (mkTerm vars p) mkTerm vars (PLoc fc n) diff --git a/src/Core/Case/CaseTree/Pretty.idr b/src/Core/Case/CaseTree/Pretty.idr index 493ed50d31..8e5067e177 100644 --- a/src/Core/Case/CaseTree/Pretty.idr +++ b/src/Core/Case/CaseTree/Pretty.idr @@ -30,7 +30,7 @@ namespace Raw prettyTree Impossible = "Impossible" prettyAlt (ConCase n tag args sc) - = hsep (annotate (DCon (Just n)) (pretty0 n) :: map pretty0 args) + = hsep (annotate (DCon (Just n)) (pretty0 n) :: map pretty0 (toList args)) <++> fatArrow <+> let sc = prettyTree sc in Union (spaces 1 <+> sc) (nest 2 (hardline <+> sc)) @@ -84,11 +84,11 @@ namespace Resugared prettyAlt env (ConCase n tag args sc) = do con <- prettyName n sc <- prettyTree (mkEnvOnto emptyFC args env) sc - pure $ hsep (annotate (DCon (Just n)) con :: map pretty0 args) + pure $ hsep (annotate (DCon (Just n)) con :: map pretty0 (toList args)) <++> fatArrow <+> Union (spaces 1 <+> sc) (nest 2 (hardline <+> sc)) prettyAlt env (DelayCase _ arg sc) = do - sc <- prettyTree (mkEnvOnto emptyFC [_,_] env) sc + sc <- prettyTree (mkEnvOnto emptyFC (_ :%: _ :%: SLNil) env) sc pure $ keyword "Delay" <++> pretty0 arg <++> fatArrow <+> Union (spaces 1 <+> sc) (nest 2 (hardline <+> sc)) diff --git a/src/Core/CompileExpr.idr b/src/Core/CompileExpr.idr index 2935bd19e7..7798b45184 100644 --- a/src/Core/CompileExpr.idr +++ b/src/Core/CompileExpr.idr @@ -5,11 +5,13 @@ module Core.CompileExpr import Core.FC import Core.Name import Core.TT +import Core.Name.ScopedList import Data.List import Data.Vect import Libraries.Data.SnocList.SizeOf +import Core.Name.ScopedList.SizeOf %default covering @@ -71,15 +73,15 @@ Eq InlineOk where mutual ||| CExp - an expression ready for compiling. public export - data CExp : List Name -> Type where + data CExp : ScopedList Name -> Type where CLocal : {idx : Nat} -> FC -> (0 p : IsVar x idx vars) -> CExp vars CRef : FC -> Name -> CExp vars -- Lambda expression - CLam : FC -> (x : Name) -> CExp (x :: vars) -> CExp vars + CLam : FC -> (x : Name) -> CExp (x :%: vars) -> CExp vars -- Let bindings CLet : FC -> (x : Name) -> InlineOk -> -- Don't inline if set - CExp vars -> CExp (x :: vars) -> CExp vars + CExp vars -> CExp (x :%: vars) -> CExp vars -- Application of a defined function. The length of the argument list is -- exactly the same length as expected by its definition (so saturate with -- lambdas if necessary, or overapply with additional CApps) @@ -109,14 +111,14 @@ mutual CCrash : FC -> String -> CExp vars public export - data CConAlt : List Name -> Type where + data CConAlt : ScopedList Name -> Type where -- If no tag, then match by constructor name. Back ends might want to -- convert names to a unique integer for performance. - MkConAlt : Name -> ConInfo -> (tag : Maybe Int) -> (args : List Name) -> - CExp (args ++ vars) -> CConAlt vars + MkConAlt : Name -> ConInfo -> (tag : Maybe Int) -> (args : ScopedList Name) -> + CExp (args +%+ vars) -> CConAlt vars public export - data CConstAlt : List Name -> Type where + data CConstAlt : ScopedList Name -> Type where MkConstAlt : Constant -> CExp vars -> CConstAlt vars mutual @@ -199,7 +201,7 @@ data CFType : Type where public export data CDef : Type where -- Normal function definition - MkFun : (args : List Name) -> CExp args -> CDef + MkFun : (args : ScopedList Name) -> CExp args -> CDef -- Constructor MkCon : (tag : Maybe Int) -> (arity : Nat) -> (nt : Maybe Nat) -> CDef -- Foreign definition @@ -209,7 +211,7 @@ data CDef : Type where CDef -- A function which will fail at runtime (usually due to being a hole) so needs -- to run, discarding arguments, no matter how many arguments are passed - MkError : CExp [] -> CDef + MkError : CExp SLNil -> CDef public export data NamedDef : Type where @@ -271,9 +273,9 @@ mutual = "(%constcase " ++ show x ++ " " ++ show exp ++ ")" export -data Names : List Name -> Type where - Nil : Names [] - (::) : Name -> Names xs -> Names (x :: xs) +data Names : ScopedList Name -> Type where + Nil : Names SLNil + (::) : Name -> Names xs -> Names (x :%: xs) elem : Name -> Names xs -> Bool elem n [] = False @@ -296,25 +298,25 @@ getLocName Z (x :: xs) First = x getLocName (S k) (x :: xs) (Later p) = getLocName k xs p export -addLocs : (args : List Name) -> Names vars -> Names (args ++ vars) -addLocs [] ns = ns -addLocs (x :: xs) ns +addLocs : (args : ScopedList Name) -> Names vars -> Names (args +%+ vars) +addLocs SLNil ns = ns +addLocs (x :%: xs) ns = let rec = addLocs xs ns in uniqueName x rec :: rec -conArgs : (args : List Name) -> Names (args ++ vars) -> List Name -conArgs [] ns = [] -conArgs (a :: as) (n :: ns) = n :: conArgs as ns +conArgs : (args : ScopedList Name) -> Names (args +%+ vars) -> List Name +conArgs SLNil ns = [] +conArgs (a :%: as) (n :: ns) = n :: conArgs as ns mutual forgetExp : Names vars -> CExp vars -> NamedCExp forgetExp locs (CLocal fc p) = NmLocal fc (getLocName _ locs p) forgetExp locs (CRef fc n) = NmRef fc n forgetExp locs (CLam fc x sc) - = let locs' = addLocs [x] locs in + = let locs' = addLocs (x :%: SLNil) locs in NmLam fc (getLocName _ locs' First) (forgetExp locs' sc) forgetExp locs (CLet fc x _ val sc) - = let locs' = addLocs [x] locs in + = let locs' = addLocs (x :%: SLNil) locs in NmLet fc (getLocName _ locs' First) (forgetExp locs val) (forgetExp locs' sc) @@ -359,7 +361,7 @@ export forgetDef : CDef -> NamedDef forgetDef (MkFun args def) = let ns = addLocs args [] - args' = conArgs {vars = []} args ns in + args' = conArgs {vars = SLNil} args ns in MkNmFun args' (forget def) forgetDef (MkCon t a nt) = MkNmCon t a nt forgetDef (MkForeign ccs fargs ty) = MkNmForeign ccs fargs ty @@ -395,7 +397,7 @@ Show CFType where show (CFFun s t) = show s ++ " -> " ++ show t show (CFIORes t) = "IORes " ++ show t show (CFStruct n args) = "struct " ++ show n ++ " " ++ showSep " " (map show args) - show (CFUser n args) = show n ++ " " ++ showSep " " (map show args) + show (CFUser n args) = show n ++ " " ++ showSep " " (toList $ map show args) export covering @@ -425,8 +427,8 @@ mutual export insertNames : SizeOf outer -> SizeOf ns -> - CExp (outer ++ inner) -> - CExp (outer ++ (ns ++ inner)) + CExp (outer +%+ inner) -> + CExp (outer +%+ (ns +%+ inner)) insertNames outer ns (CLocal fc prf) = let MkNVar var' = insertNVarNames outer ns (MkNVar prf) in CLocal fc var' @@ -459,19 +461,19 @@ mutual insertNamesConAlt : SizeOf outer -> SizeOf ns -> - CConAlt (outer ++ inner) -> - CConAlt (outer ++ (ns ++ inner)) + CConAlt (outer +%+ inner) -> + CConAlt (outer +%+ (ns +%+ inner)) insertNamesConAlt {outer} {ns} p q (MkConAlt x ci tag args sc) - = let sc' : CExp ((args ++ outer) ++ inner) + = let sc' : CExp ((args +%+ outer) +%+ inner) = rewrite sym (appendAssociative args outer inner) in sc in MkConAlt x ci tag args - (rewrite appendAssociative args outer (ns ++ inner) in + (rewrite appendAssociative args outer (ns +%+ inner) in insertNames (mkSizeOf args + p) q sc') insertNamesConstAlt : SizeOf outer -> SizeOf ns -> - CConstAlt (outer ++ inner) -> - CConstAlt (outer ++ (ns ++ inner)) + CConstAlt (outer +%+ inner) -> + CConstAlt (outer +%+ (ns +%+ inner)) insertNamesConstAlt outer ns (MkConstAlt x sc) = MkConstAlt x (insertNames outer ns sc) export @@ -572,7 +574,7 @@ mutual = MkConAlt x ci tag args (rewrite appendAssociative args outer vars in substEnv (mkSizeOf args + p) q env - (rewrite sym (appendAssociative args outer (dropped ++ vars)) in + (rewrite sym (appendAssociative args outer (dropped +%+ vars)) in sc)) substConstAlt : Substitutable CExp CConstAlt @@ -581,15 +583,15 @@ mutual export substs : {dropped, vars : _} -> SizeOf dropped -> - SubstCEnv dropped vars -> CExp (dropped ++ vars) -> CExp vars + SubstCEnv dropped vars -> CExp (dropped +%+ vars) -> CExp vars substs = substEnv zero mutual export mkLocals : SizeOf outer -> Bounds bound -> - CExp (outer ++ vars) -> - CExp (outer ++ (bound ++ vars)) + CExp (outer +%+ vars) -> + CExp (outer +%+ (bound +%+ vars)) mkLocals later bs (CLocal {idx} {x} fc p) = let MkNVar p' = addVars later bs (MkNVar p) in CLocal {x} fc p' mkLocals later bs (CRef fc var) @@ -628,23 +630,23 @@ mutual mkLocalsConAlt : SizeOf outer -> Bounds bound -> - CConAlt (outer ++ vars) -> - CConAlt (outer ++ (bound ++ vars)) + CConAlt (outer +%+ vars) -> + CConAlt (outer +%+ (bound +%+ vars)) mkLocalsConAlt {bound} {outer} {vars} p bs (MkConAlt x ci tag args sc) - = let sc' : CExp ((args ++ outer) ++ vars) + = let sc' : CExp ((args +%+ outer) +%+ vars) = rewrite sym (appendAssociative args outer vars) in sc in MkConAlt x ci tag args - (rewrite appendAssociative args outer (bound ++ vars) in + (rewrite appendAssociative args outer (bound +%+ vars) in mkLocals (mkSizeOf args + p) bs sc') mkLocalsConstAlt : SizeOf outer -> Bounds bound -> - CConstAlt (outer ++ vars) -> - CConstAlt (outer ++ (bound ++ vars)) + CConstAlt (outer +%+ vars) -> + CConstAlt (outer +%+ (bound +%+ vars)) mkLocalsConstAlt later bs (MkConstAlt x sc) = MkConstAlt x (mkLocals later bs sc) export -refsToLocals : Bounds bound -> CExp vars -> CExp (bound ++ vars) +refsToLocals : Bounds bound -> CExp vars -> CExp (bound +%+ vars) refsToLocals None tm = tm refsToLocals bs y = mkLocals zero bs y diff --git a/src/Core/CompileExpr/Pretty.idr b/src/Core/CompileExpr/Pretty.idr index ba7ae1fc73..0b9d1985ec 100644 --- a/src/Core/CompileExpr/Pretty.idr +++ b/src/Core/CompileExpr/Pretty.idr @@ -17,7 +17,7 @@ import Libraries.Data.String.Extra %hide Vect.catMaybes %hide Vect.(++) -%hide SizeOf.map +%hide Libraries.Data.List.SizeOf.map %hide Core.Name.prettyOp @@ -120,9 +120,9 @@ prettyCExp : {args : _} -> CExp args -> Doc IdrisSyntax prettyCExp = prettyNamedCExp . forget prettyCDef : CDef -> Doc IdrisDocAnn -prettyCDef (MkFun [] exp) = reAnnotate Syntax $ prettyCExp exp +prettyCDef (MkFun SLNil exp) = reAnnotate Syntax $ prettyCExp exp prettyCDef (MkFun args exp) = reAnnotate Syntax $ - keyword "\\" <++> concatWith (\ x, y => x <+> keyword "," <++> y) (map prettyName args) + keyword "\\" <++> concatWith (\ x, y => x <+> keyword "," <++> y) (map prettyName $ toList args) <++> fatArrow <++> prettyCExp exp prettyCDef (MkCon mtag arity nt) = vcat $ header (maybe "Data" (const "Type") mtag <++> "Constructor") :: map (indent 2) diff --git a/src/Core/Context.idr b/src/Core/Context.idr index ce7256c79e..53e2994a13 100644 --- a/src/Core/Context.idr +++ b/src/Core/Context.idr @@ -323,7 +323,7 @@ commitCtxt ctxt ||| @vis Visibility, defaulting to private ||| @def actual definition export -newDef : (fc : FC) -> (n : Name) -> (rig : RigCount) -> (vars : List Name) -> +newDef : (fc : FC) -> (n : Name) -> (rig : RigCount) -> (vars : ScopedList Name) -> (ty : ClosedTerm) -> (vis : WithDefault Visibility Private) -> (def : Def) -> GlobalDef newDef fc n rig vars ty vis def = MkGlobalDef @@ -1353,7 +1353,7 @@ addBuiltin n ty tot op , specArgs = [] , inferrable = [] , multiplicity = top - , localVars = [] + , localVars = SLNil , visibility = specified Public , totality = tot , isEscapeHatch = False diff --git a/src/Core/Context/Context.idr b/src/Core/Context/Context.idr index 6abf0e08f8..2ac584819d 100644 --- a/src/Core/Context/Context.idr +++ b/src/Core/Context/Context.idr @@ -4,6 +4,7 @@ import Core.Case.CaseTree import Core.CompileExpr import Core.Env import public Core.Name +import Core.Name.ScopedList import public Core.Options.Log import public Core.TT @@ -70,7 +71,7 @@ public export data Def : Type where None : Def -- Not yet defined PMDef : (pminfo : PMDefInfo) -> - (args : List Name) -> + (args : ScopedList Name) -> (treeCT : CaseTree args) -> (treeRT : CaseTree args) -> (pats : List (vs ** (Env Term vs, Term vs, Term vs))) -> @@ -307,7 +308,7 @@ record GlobalDef where specArgs : List Nat -- arguments to specialise by inferrable : List Nat -- arguments which can be inferred from elsewhere in the type multiplicity : RigCount - localVars : List Name -- environment name is defined in + localVars : ScopedList Name -- environment name is defined in visibility : WithDefault Visibility Private totality : Totality isEscapeHatch : Bool diff --git a/src/Core/Context/Data.idr b/src/Core/Context/Data.idr index 97a5dcd2ea..2dd08197a6 100644 --- a/src/Core/Context/Data.idr +++ b/src/Core/Context/Data.idr @@ -90,7 +90,7 @@ paramPos tyn dcons = do export addData : {auto c : Ref Ctxt Defs} -> - List Name -> Visibility -> Int -> DataDef -> Core Int + ScopedList Name -> Visibility -> Int -> DataDef -> Core Int addData vars vis tidx (MkData (MkCon dfc tyn arity tycon) datacons) = do defs <- get Ctxt tag <- getNextTypeTag diff --git a/src/Core/Context/Pretty.idr b/src/Core/Context/Pretty.idr index 068a5b5d74..acd1b3355b 100644 --- a/src/Core/Context/Pretty.idr +++ b/src/Core/Context/Pretty.idr @@ -43,7 +43,7 @@ namespace Raw prettyDef (PMDef _ args ct _ pats) = let ct = prettyTree ct in vcat - [ "Arguments" <++> cast (prettyList args) + [ "Arguments" <++> cast (prettyList $ toList args) , header "Compile time tree" <++> reAnnotate Syntax ct ] prettyDef (DCon tag arity nt) = @@ -95,7 +95,7 @@ namespace Resugared prettyDef (PMDef _ args ct _ pats) = do ct <- prettyTree (mkEnv emptyFC _) ct pure $ vcat - [ "Arguments" <++> cast (prettyList args) + [ "Arguments" <++> cast (prettyList $ toList args) , header "Compile time tree" <++> reAnnotate Syntax ct ] prettyDef (DCon tag arity nt) = pure $ diff --git a/src/Core/Core.idr b/src/Core/Core.idr index 878201c01d..39402ec68f 100644 --- a/src/Core/Core.idr +++ b/src/Core/Core.idr @@ -3,6 +3,7 @@ module Core.Core import Core.Context.Context import Core.Env import Core.TT +import Core.Name.ScopedList import Data.List1 import Data.Vect @@ -256,7 +257,7 @@ Show Error where case cov of IsCovering => "Oh yes it is (Internal error!)" MissingCases cs => "Missing cases:\n\t" ++ - showSep "\n\t" (map show cs) + showSep "\n\t" (toList $ map show cs) NonCoveringCall ns => "Calls non covering function" ++ (case ns of [fn] => " " ++ show fn @@ -760,6 +761,18 @@ export traverse : (a -> Core b) -> List a -> Core (List b) traverse f xs = traverse' f xs [] +namespace ScopedList + -- Traversable (specialised) + traverse' : (a -> Core b) -> ScopedList a -> ScopedList b -> Core (ScopedList b) + traverse' f SLNil acc = pure (reverse acc) + traverse' f (x :%: xs) acc + = traverse' f xs (!(f x) :%: acc) + + %inline + export + traverse : (a -> Core b) -> ScopedList a -> Core (ScopedList b) + traverse f xs = traverse' f xs SLNil + export mapMaybeM : (a -> Core (Maybe b)) -> List a -> Core (List b) mapMaybeM f = go [<] where @@ -830,6 +843,14 @@ traverseList1_ f xxs ignore (f x) traverse_ f xs +namespace ScopedList + export + traverse_ : (a -> Core b) -> ScopedList a -> Core () + traverse_ f SLNil = pure () + traverse_ f (x :%: xs) + = Core.do ignore (f x) + traverse_ f xs + namespace PiInfo export traverse : (a -> Core b) -> PiInfo a -> Core (PiInfo b) @@ -876,6 +897,14 @@ anyM f (x :: xs) then pure True else anyM f xs +export +anyMScoped : (a -> Core Bool) -> ScopedList a -> Core Bool +anyMScoped f SLNil = pure False +anyMScoped f (x :%: xs) + = if !(f x) + then pure True + else anyMScoped f xs + export allM : (a -> Core Bool) -> List a -> Core Bool allM f [] = pure True diff --git a/src/Core/Coverage.idr b/src/Core/Coverage.idr index 085046f65d..5411e33c70 100644 --- a/src/Core/Coverage.idr +++ b/src/Core/Coverage.idr @@ -83,10 +83,10 @@ conflict defs env nfty n _ => pure False where mutual - conflictArgs : Int -> List (Closure vars) -> List (Closure []) -> + conflictArgs : Int -> ScopedList (Closure vars) -> ScopedList (Closure SLNil) -> Core (Maybe (List (Name, Term vars))) - conflictArgs _ [] [] = pure (Just []) - conflictArgs i (c :: cs) (c' :: cs') + conflictArgs _ SLNil SLNil = pure (Just []) + conflictArgs i (c :%: cs) (c' :%: cs') = do cnf <- evalClosure defs c cnf' <- evalClosure defs c' Just ms <- conflictNF i cnf cnf' @@ -96,13 +96,13 @@ conflict defs env nfty n pure (Just (ms ++ ms')) conflictArgs _ _ _ = pure (Just []) - -- If the constructor type (the NF []) matches the variable type, + -- If the constructor type (the NF SLNil) matches the variable type, -- then there may be a way to construct it, so return the matches in -- the indices. -- If any of those matches clash, the constructor is not valid -- e.g, Eq x x matches Eq Z (S Z), with x = Z and x = S Z -- conflictNF returns the list of matches, for checking - conflictNF : Int -> NF vars -> NF [] -> + conflictNF : Int -> NF vars -> NF SLNil -> Core (Maybe (List (Name, Term vars))) conflictNF i t (NBind fc x b sc) -- invent a fresh name, in case a user has bound the same name @@ -111,7 +111,7 @@ conflict defs env nfty n = let x' = MN (show x) i in conflictNF (i + 1) t !(sc defs (toClosure defaultOpts [] (Ref fc Bound x'))) - conflictNF i nf (NApp _ (NRef Bound n) []) + conflictNF i nf (NApp _ (NRef Bound n) SLNil) = do empty <- clearDefs defs pure (Just [(n, !(quote empty env nf))]) conflictNF i (NDCon _ n t a args) (NDCon _ n' t' a' args') @@ -189,11 +189,11 @@ getMissingAlts fc defs nfty alts noneOf alts c = not $ any (altMatch c) alts -- Mapping of variable to constructor tag already matched for it -KnownVars : List Name -> Type -> Type +KnownVars : ScopedList Name -> Type -> Type KnownVars vars a = List (Var vars, a) -getName : {idx : Nat} -> {vars : List Name} -> (0 p : IsVar n idx vars) -> Name -getName {vars = v :: _} First = v +getName : {idx : Nat} -> {vars : ScopedList Name} -> (0 p : IsVar n idx vars) -> Name +getName {vars = v :%: _} First = v getName (Later p) = getName p showK : {ns : _} -> @@ -204,7 +204,7 @@ showK {a} xs = show (map aString xs) (Var vars, a) -> (Name, a) aString (MkVar v, t) = (nameAt v, t) -weakenNs : SizeOf args -> KnownVars vars a -> KnownVars (args ++ vars) a +weakenNs : SizeOf args -> KnownVars vars a -> KnownVars (args +%+ vars) a weakenNs args [] = [] weakenNs args ((v, t) :: xs) = (weakenNs args v, t) :: weakenNs args xs @@ -271,8 +271,8 @@ buildArgs : {auto c : Ref Ctxt Defs} -> KnownVars vars Int -> -- Things which have definitely match KnownVars vars (List Int) -> -- Things an argument *can't* be -- (because a previous case matches) - List ClosedTerm -> -- ^ arguments, with explicit names - CaseTree vars -> Core (List (List ClosedTerm)) + ScopedList ClosedTerm -> -- ^ arguments, with explicit names + CaseTree vars -> Core (List (ScopedList ClosedTerm)) buildArgs fc defs known not ps cs@(Case {name = var} idx el ty altsIn) -- If we've already matched on 'el' in this branch, restrict the alternatives -- to the tag we already know. Otherwise, add missing cases and filter out @@ -289,17 +289,17 @@ buildArgs fc defs known not ps cs@(Case {name = var} idx el ty altsIn) buildArgsAlt not altsN where buildArgAlt : KnownVars vars (List Int) -> - CaseAlt vars -> Core (List (List ClosedTerm)) + CaseAlt vars -> Core (List (ScopedList ClosedTerm)) buildArgAlt not' (ConCase n t args sc) = do let l = mkSizeOf args let con = Ref fc (DataCon t (size l)) n let ps' = map (substName var (apply fc - con (map (Ref fc Bound) args))) ps + con (toList $ map (Ref fc Bound) args))) ps buildArgs fc defs (weakenNs l ((MkVar el, t) :: known)) (weakenNs l not') ps' sc buildArgAlt not' (DelayCase t a sc) - = let l = mkSizeOf [t, a] + = let l = mkSizeOf (t :%: a :%: SLNil) ps' = map (substName var (TDelay fc LUnknown (Ref fc Bound t) (Ref fc Bound a))) ps in @@ -312,7 +312,7 @@ buildArgs fc defs known not ps cs@(Case {name = var} idx el ty altsIn) = buildArgs fc defs known not' ps sc buildArgsAlt : KnownVars vars (List Int) -> List (CaseAlt vars) -> - Core (List (List ClosedTerm)) + Core (List (ScopedList ClosedTerm)) buildArgsAlt not' [] = pure [] buildArgsAlt not' (c@(ConCase _ t _ _) :: cs) = pure $ !(buildArgAlt not' c) ++ @@ -345,7 +345,7 @@ getMissing fc n ctree logC "coverage.missing" 20 $ map (join "\n") $ flip traverse pats $ \ pat => show <$> toFullNames pat - pure (map (apply fc (Ref fc Func n)) patss) + pure (map (apply fc (Ref fc Func n) . toList) patss) -- For the given name, get the names it refers to which are not themselves -- covering. diff --git a/src/Core/Env.idr b/src/Core/Env.idr index b95a35cdb7..fa9790cc64 100644 --- a/src/Core/Env.idr +++ b/src/Core/Env.idr @@ -1,24 +1,26 @@ module Core.Env import Core.TT +import Core.Name.ScopedList +import Core.Name.ScopedList.SizeOf import Data.List %default total -- Environment containing types and values of local variables public export -data Env : (tm : List Name -> Type) -> List Name -> Type where - Nil : Env tm [] - (::) : Binder (tm vars) -> Env tm vars -> Env tm (x :: vars) +data Env : (tm : ScopedList Name -> Type) -> ScopedList Name -> Type where + Nil : Env tm SLNil + (::) : Binder (tm vars) -> Env tm vars -> Env tm (x :%: vars) %name Env rho export -extend : (x : Name) -> Binder (tm vars) -> Env tm vars -> Env tm (x :: vars) +extend : (x : Name) -> Binder (tm vars) -> Env tm vars -> Env tm (x :%: vars) extend x = (::) {x} export -(++) : {ns : _} -> Env Term ns -> Env Term vars -> Env Term (ns ++ vars) +(++) : {ns : _} -> Env Term ns -> Env Term vars -> Env Term (ns +%+ vars) (++) (b :: bs) e = extend _ (map embed b) (bs ++ e) (++) [] e = e @@ -40,13 +42,13 @@ lengthExplicitPi (Pi _ _ Explicit _ :: rho) = S (lengthExplicitPi rho) lengthExplicitPi (_ :: rho) = lengthExplicitPi rho export -namesNoLet : {xs : _} -> Env tm xs -> List Name -namesNoLet [] = [] +namesNoLet : {xs : _} -> Env tm xs -> ScopedList Name +namesNoLet [] = SLNil namesNoLet (Let _ _ _ _ :: xs) = namesNoLet xs -namesNoLet {xs = x :: _} (_ :: env) = x :: namesNoLet env +namesNoLet {xs = x :%: _} (_ :: env) = x :%: namesNoLet env public export -data IsDefined : Name -> List Name -> Type where +data IsDefined : Name -> ScopedList Name -> Type where MkIsDefined : {idx : Nat} -> RigCount -> (0 p : IsVar n idx vars) -> IsDefined n vars @@ -55,7 +57,7 @@ defined : {vars : _} -> (n : Name) -> Env Term vars -> Maybe (IsDefined n vars) defined n [] = Nothing -defined {vars = x :: xs} n (b :: env) +defined {vars = x :%: xs} n (b :: env) = case nameEq n x of Nothing => do MkIsDefined rig prf <- defined n env pure (MkIsDefined rig (Later prf)) @@ -72,20 +74,20 @@ bindEnv loc (b :: env) tm Explicit (binderType b)) tm) -revOnto : (xs, vs : List a) -> reverseOnto xs vs = reverse vs ++ xs -revOnto xs [] = Refl -revOnto xs (v :: vs) - = rewrite revOnto (v :: xs) vs in - rewrite appendAssociative (reverse vs) [v] xs in - rewrite revOnto [v] vs in Refl - -revNs : (vs, ns : List a) -> reverse ns ++ reverse vs = reverse (vs ++ ns) -revNs [] ns = rewrite appendNilRightNeutral (reverse ns) in Refl -revNs (v :: vs) ns - = rewrite revOnto [v] vs in - rewrite revOnto [v] (vs ++ ns) in +revOnto : (xs, vs : ScopedList a) -> reverseOnto xs vs = reverse vs +%+ xs +revOnto xs SLNil = Refl +revOnto xs (v :%: vs) + = rewrite revOnto (v :%: xs) vs in + rewrite appendAssociative (reverse vs) (v :%: SLNil) xs in + rewrite revOnto (v :%: SLNil) vs in Refl + +revNs : (vs, ns : ScopedList a) -> reverse ns +%+ reverse vs = reverse (vs +%+ ns) +revNs SLNil ns = rewrite appendNilRightNeutral (reverse ns) in Refl +revNs (v :%: vs) ns + = rewrite revOnto (v :%: SLNil) vs in + rewrite revOnto (v :%: SLNil) (vs +%+ ns) in rewrite sym (revNs vs ns) in - rewrite appendAssociative (reverse ns) (reverse vs) [v] in + rewrite appendAssociative (reverse ns) (reverse vs) (v :%: SLNil) in Refl -- Weaken by all the names at once at the end, to save multiple traversals @@ -94,19 +96,19 @@ revNs (v :: vs) ns -- when environments get fairly big. getBinderUnder : Weaken tm => {vars : _} -> {idx : Nat} -> - (ns : List Name) -> + (ns : ScopedList Name) -> (0 p : IsVar x idx vars) -> Env tm vars -> Binder (tm (reverseOnto vars ns)) -getBinderUnder {idx = Z} {vars = v :: vs} ns First (b :: env) - = rewrite revOnto vs (v :: ns) in map (weakenNs (reverse (mkSizeOf (v :: ns)))) b -getBinderUnder {idx = S k} {vars = v :: vs} ns (Later lp) (b :: env) - = getBinderUnder (v :: ns) lp env +getBinderUnder {idx = Z} {vars = v :%: vs} ns First (b :: env) + = rewrite revOnto vs (v :%: ns) in map (weakenNs (reverse (mkSizeOf ((v :%: ns))))) b +getBinderUnder {idx = S k} {vars = v :%: vs} ns (Later lp) (b :: env) + = getBinderUnder (v :%: ns) lp env export getBinder : Weaken tm => {vars : _} -> {idx : Nat} -> (0 p : IsVar x idx vars) -> Env tm vars -> Binder (tm vars) -getBinder el env = getBinderUnder [] el env +getBinder el env = getBinderUnder SLNil el env -- For getBinderLoc, we are not reusing getBinder because there is no need to -- needlessly weaken stuff; @@ -161,11 +163,11 @@ letToLam (b :: env) = b :: letToLam env mutual -- Quicker, if less safe, to store variables as a Nat, for quick comparison findUsed : {vars : _} -> - Env Term vars -> List Nat -> Term vars -> List Nat + Env Term vars -> ScopedList Nat -> Term vars -> ScopedList Nat findUsed env used (Local fc r idx p) = if elemBy eqNat idx used then used - else assert_total (findUsedInBinder env (idx :: used) + else assert_total (findUsedInBinder env (idx :%: used) (getBinder p env)) where eqNat : Nat -> Nat -> Bool @@ -173,7 +175,7 @@ mutual findUsed env used (Meta _ _ _ args) = findUsedArgs env used args where - findUsedArgs : Env Term vars -> List Nat -> List (Term vars) -> List Nat + findUsedArgs : Env Term vars -> ScopedList Nat -> List (Term vars) -> ScopedList Nat findUsedArgs env u [] = u findUsedArgs env u (a :: as) = findUsedArgs env (findUsed env u a) as @@ -183,10 +185,10 @@ mutual (map S (findUsedInBinder env used b)) tm) where - dropS : List Nat -> List Nat - dropS [] = [] - dropS (Z :: xs) = dropS xs - dropS (S p :: xs) = p :: dropS xs + dropS : ScopedList Nat -> ScopedList Nat + dropS SLNil = SLNil + dropS (Z :%: xs) = dropS xs + dropS (S p :%: xs) = p :%: dropS xs findUsed env used (App fc fn arg) = findUsed env (findUsed env used fn) arg findUsed env used (As fc s a p) @@ -200,56 +202,56 @@ mutual findUsed env used _ = used findUsedInBinder : {vars : _} -> - Env Term vars -> List Nat -> - Binder (Term vars) -> List Nat + Env Term vars -> ScopedList Nat -> + Binder (Term vars) -> ScopedList Nat findUsedInBinder env used (Let _ _ val ty) = findUsed env (findUsed env used val) ty findUsedInBinder env used (PLet _ _ val ty) = findUsed env (findUsed env used val) ty findUsedInBinder env used b = findUsed env used (binderType b) -toVar : (vars : List Name) -> Nat -> Maybe (Var vars) -toVar (v :: vs) Z = Just (MkVar First) -toVar (v :: vs) (S k) +toVar : (vars : ScopedList Name) -> Nat -> Maybe (Var vars) +toVar (v :%: vs) Z = Just (MkVar First) +toVar (v :%: vs) (S k) = do MkVar prf <- toVar vs k Just (MkVar (Later prf)) toVar _ _ = Nothing export findUsedLocs : {vars : _} -> - Env Term vars -> Term vars -> List (Var vars) + Env Term vars -> Term vars -> ScopedList (Var vars) findUsedLocs env tm - = mapMaybe (toVar _) (findUsed env [] tm) + = mapMaybe (toVar _) (findUsed env SLNil tm) -isUsed : Nat -> List (Var vars) -> Bool -isUsed n [] = False -isUsed n (v :: vs) = n == varIdx v || isUsed n vs +isUsed : Nat -> ScopedList (Var vars) -> Bool +isUsed n SLNil = False +isUsed n (v :%: vs) = n == varIdx v || isUsed n vs mkShrinkSub : {n : _} -> - (vars : _) -> List (Var (n :: vars)) -> - (newvars ** Thin newvars (n :: vars)) -mkShrinkSub [] els + (vars : _) -> ScopedList (Var (n :%: vars)) -> + (newvars ** Thin newvars (n :%: vars)) +mkShrinkSub SLNil els = if isUsed 0 els then (_ ** Keep Refl) else (_ ** Drop Refl) -mkShrinkSub (x :: xs) els +mkShrinkSub (x :%: xs) els = let (_ ** subRest) = mkShrinkSub xs (dropFirst els) in if isUsed 0 els then (_ ** Keep subRest) else (_ ** Drop subRest) mkShrink : {vars : _} -> - List (Var vars) -> + ScopedList (Var vars) -> (newvars ** Thin newvars vars) -mkShrink {vars = []} xs = (_ ** Refl) -mkShrink {vars = v :: vs} xs = mkShrinkSub _ xs +mkShrink {vars = SLNil} xs = (_ ** Refl) +mkShrink {vars = v :%: vs} xs = mkShrinkSub _ xs -- Find the smallest subset of the environment which is needed to type check -- the given term export findSubEnv : {vars : _} -> Env Term vars -> Term vars -> - (vars' : List Name ** Thin vars' vars) + (vars' : ScopedList Name ** Thin vars' vars) findSubEnv env tm = mkShrink (findUsedLocs env tm) export @@ -262,9 +264,9 @@ shrinkEnv (b :: env) (Keep p) pure (b' :: env') export -mkEnvOnto : FC -> (xs : List Name) -> Env Term ys -> Env Term (xs ++ ys) -mkEnvOnto fc [] vs = vs -mkEnvOnto fc (n :: ns) vs +mkEnvOnto : FC -> (xs : ScopedList Name) -> Env Term ys -> Env Term (xs +%+ ys) +mkEnvOnto fc SLNil vs = vs +mkEnvOnto fc (n :%: ns) vs = PVar fc top Explicit (Erased fc Placeholder) :: mkEnvOnto fc ns vs @@ -272,9 +274,9 @@ mkEnvOnto fc (n :: ns) vs -- and types of the contents. -- We use this when building and comparing case trees. export -mkEnv : FC -> (vs : List Name) -> Env Term vs -mkEnv fc [] = [] -mkEnv fc (n :: ns) = PVar fc top Explicit (Erased fc Placeholder) :: mkEnv fc ns +mkEnv : FC -> (vs : ScopedList Name) -> Env Term vs +mkEnv fc SLNil = [] +mkEnv fc (n :%: ns) = PVar fc top Explicit (Erased fc Placeholder) :: mkEnv fc ns -- Update an environment so that all names are guaranteed unique. In the -- case of a clash, the most recently bound is left unchanged. @@ -282,7 +284,7 @@ export uniqifyEnv : {vars : _} -> Env Term vars -> (vars' ** (Env Term vars', CompatibleVars vars vars')) -uniqifyEnv env = uenv [] env +uniqifyEnv env = uenv SLNil env where next : Name -> Name next (MN n i) = MN n (i + 1) @@ -290,7 +292,7 @@ uniqifyEnv env = uenv [] env next (NS ns n) = NS ns (next n) next n = MN (show n) 0 - uniqueLocal : List Name -> Name -> Name + uniqueLocal : ScopedList Name -> Name -> Name uniqueLocal vs n = if n `elem` vs -- we'll find a new name eventualy since the list of names @@ -302,18 +304,18 @@ uniqifyEnv env = uenv [] env else n uenv : {vars : _} -> - List Name -> Env Term vars -> + ScopedList Name -> Env Term vars -> (vars' ** (Env Term vars', CompatibleVars vars vars')) - uenv used [] = ([] ** ([], Pre)) - uenv used {vars = v :: vs} (b :: bs) + uenv used [] = (SLNil ** ([], Pre)) + uenv used {vars = v :%: vs} (b :: bs) = if v `elem` used then let v' = uniqueLocal used v - (vs' ** (env', compat)) = uenv (v' :: used) bs + (vs' ** (env', compat)) = uenv (v' :%: used) bs b' = map (compatNs compat) b in - (v' :: vs' ** (b' :: env', Ext compat)) - else let (vs' ** (env', compat)) = uenv (v :: used) bs + (v' :%: vs' ** (b' :: env', Ext compat)) + else let (vs' ** (env', compat)) = uenv (v :%: used) bs b' = map (compatNs compat) b in - (v :: vs' ** (b' :: env', Ext compat)) + (v :%: vs' ** (b' :: env', Ext compat)) export allVars : {vars : _} -> Env Term vars -> List (Var vars) @@ -334,7 +336,7 @@ close fc nm env tm where - mkSubstEnv : Int -> Env Term vs -> (SizeOf vs, SubstEnv vs []) + mkSubstEnv : Int -> Env Term vs -> (SizeOf vs, SubstEnv vs SLNil) mkSubstEnv i [] = (zero, []) mkSubstEnv i (v :: vs) = let (s, env) = mkSubstEnv (i + 1) vs in diff --git a/src/Core/GetType.idr b/src/Core/GetType.idr index 696e076680..39b61b0cf2 100644 --- a/src/Core/GetType.idr +++ b/src/Core/GetType.idr @@ -33,7 +33,7 @@ mutual chkMeta fc env !(nf defs env (embed mty)) args chk env (Bind fc nm b sc) = do bt <- chkBinder env b - sct <- chk {vars = nm :: _} (b :: env) sc + sct <- chk {vars = nm :%: _} (b :: env) sc pure $ gnf env (discharge fc nm b !(getTerm bt) !(getTerm sct)) chk env (App fc f a) = do fty <- chk env f @@ -85,7 +85,7 @@ mutual chkBinder env b = chk env (binderType b) discharge : FC -> (nm : Name) -> Binder (Term vars) -> - Term vars -> Term (nm :: vars) -> (Term vars) + Term vars -> Term (nm :%: vars) -> (Term vars) discharge fc n (Lam fc' c x ty) bindty scopety = Bind fc n (Pi fc' c x ty) scopety discharge fc n (Let fc' c val ty) bindty scopety diff --git a/src/Core/Hash.idr b/src/Core/Hash.idr index fedb7dbef2..d397076536 100644 --- a/src/Core/Hash.idr +++ b/src/Core/Hash.idr @@ -3,6 +3,7 @@ module Core.Hash import Core.Case.CaseTree import Core.CompileExpr import Core.TT +import Core.Name.ScopedList import Data.List1 import Libraries.Data.String.Iterator @@ -83,6 +84,11 @@ Hashable a => Hashable (List a) where hashWithSalt h [] = abs h hashWithSalt h (x :: xs) = hashWithSalt (h * 33 + hash x) xs +export +Hashable a => Hashable (ScopedList a) where + hashWithSalt h SLNil = abs h + hashWithSalt h (x :%: xs) = hashWithSalt (h * 33 + hash x) xs + export Hashable a => Hashable (List1 a) where hashWithSalt h xxs = hashWithSalt (h * 33 + hash (head xxs)) (tail xxs) diff --git a/src/Core/LinearCheck.idr b/src/Core/LinearCheck.idr index b250f22387..3d5213e70a 100644 --- a/src/Core/LinearCheck.idr +++ b/src/Core/LinearCheck.idr @@ -19,7 +19,7 @@ import Libraries.Data.SnocList.SizeOf -- List of variable usages - we'll count the contents of specific variables -- when discharging binders, to ensure that linear names are only used once -data Usage : List Name -> Type where +data Usage : ScopedList Name -> Type where Nil : Usage vars (::) : Var vars -> Usage vars -> Usage vars @@ -31,7 +31,7 @@ Show (Usage vars) where showAll [el] = show el showAll (x :: xs) = show x ++ ", " ++ show xs -doneScope : Usage (n :: vars) -> Usage vars +doneScope : Usage (n :%: vars) -> Usage vars doneScope [] = [] doneScope (MkVar First :: xs) = doneScope xs doneScope (MkVar (Later p) :: xs) = MkVar p :: doneScope xs @@ -185,9 +185,9 @@ mutual when (not erase) $ rigSafe rigb rig pure (Local fc x idx prf, gnf env ty, used rig) where - getName : {idx : _} -> (vs : List Name) -> (0 p : IsVar n idx vs) -> Name - getName (x :: _) First = x - getName (x :: xs) (Later p) = getName xs p + getName : {idx : _} -> (vs : ScopedList Name) -> (0 p : IsVar n idx vs) -> Name + getName (x :%: _) First = x + getName (x :%: xs) (Later p) = getName xs p rigSafe : RigCount -> RigCount -> Core () rigSafe l r = when (l < r) @@ -425,7 +425,7 @@ mutual discharge : {vars : _} -> Defs -> Env Term vars -> FC -> (nm : Name) -> Binder (Term vars) -> Glued vars -> - Term (nm :: vars) -> Glued (nm :: vars) -> Usage vars -> + Term (nm :%: vars) -> Glued (nm :%: vars) -> Usage vars -> Core (Term vars, Glued vars, Usage vars) discharge defs env fc nm (Lam fc' c x ty) gbindty scope gscopety used = do scty <- getTerm gscopety @@ -529,7 +529,7 @@ mutual List (Term (done <>> vars)) -> Term (done <>> vars) -> Core () checkEnvUsage s rig [] usage args tm = pure () - checkEnvUsage s rig {done} {vars = nm :: xs} (b :: env) usage args tm + checkEnvUsage s rig {done} {vars = nm :%: xs} (b :: env) usage args tm = do let pos = mkVarChiply s let used_in = count (varIdx pos) usage @@ -645,12 +645,12 @@ mutual RigCount -> (erase : Bool) -> Env Term vars -> Name -> Int -> Def -> List (Term vars) -> Core (Term vars, Glued vars, Usage vars) - expandMeta rig erase env n idx (PMDef _ [] (STerm _ fn) _ _) args + expandMeta rig erase env n idx (PMDef _ SLNil (STerm _ fn) _ _) args = do tm <- substMeta (embed fn) args zero [] lcheck rig erase env tm where substMeta : {drop, vs : _} -> - Term (drop ++ vs) -> List (Term vs) -> + Term (drop +%+ vs) -> List (Term vs) -> SizeOf drop -> SubstEnv drop vs -> Core (Term vs) substMeta (Bind bfc n (Lam _ c e ty) sc) (a :: as) drop env @@ -702,7 +702,7 @@ checkEnvUsage : {vars : _} -> Term (done <>> vars) -> Core () checkEnvUsage fc s rig [] usage tm = pure () -checkEnvUsage fc s rig {vars = nm :: xs} (b :: env) usage tm +checkEnvUsage fc s rig {vars = nm :%: xs} (b :: env) usage tm = do let pos = mkVarChiply s let used_in = count (varIdx pos) usage diff --git a/src/Core/Metadata.idr b/src/Core/Metadata.idr index cb7d867b48..811ea61cfd 100644 --- a/src/Core/Metadata.idr +++ b/src/Core/Metadata.idr @@ -220,7 +220,7 @@ addLHS loc outerenvlen env tm substEnv : {vars : _} -> FC -> Env Term vars -> (tm : Term vars) -> ClosedTerm substEnv loc [] tm = tm -substEnv {vars = x :: _} loc (b :: env) tm +substEnv {vars = x :%: _} loc (b :: env) tm = substEnv loc env (subst (Ref loc Bound x) tm) export diff --git a/src/Core/Name/CompatibleVars.idr b/src/Core/Name/CompatibleVars.idr new file mode 100644 index 0000000000..0fa5be410c --- /dev/null +++ b/src/Core/Name/CompatibleVars.idr @@ -0,0 +1,42 @@ +module Core.Name.CompatibleVars + +public export +data CompatibleVars : (xs, ys : List a) -> Type where + Pre : CompatibleVars xs xs + Ext : CompatibleVars xs ys -> CompatibleVars (n :: xs) (m :: ys) + +export +invertExt : CompatibleVars (n :: xs) (m :: ys) -> CompatibleVars xs ys +invertExt Pre = Pre +invertExt (Ext p) = p + +export +extendCompats : (args : List a) -> + CompatibleVars xs ys -> + CompatibleVars (args ++ xs) (args ++ ys) +extendCompats args Pre = Pre +extendCompats args prf = go args prf where + + go : (args : List a) -> + CompatibleVars xs ys -> + CompatibleVars (args ++ xs) (args ++ ys) + go [] prf = prf + go (x :: xs) prf = Ext (go xs prf) + +export +decCompatibleVars : (xs, ys : List a) -> Dec (CompatibleVars xs ys) +decCompatibleVars [] [] = Yes Pre +decCompatibleVars [] (x :: xs) = No (\case p impossible) +decCompatibleVars (x :: xs) [] = No (\case p impossible) +decCompatibleVars (x :: xs) (y :: ys) = case decCompatibleVars xs ys of + Yes prf => Yes (Ext prf) + No nprf => No (nprf . invertExt) + +export +areCompatibleVars : (xs, ys : List a) -> + Maybe (CompatibleVars xs ys) +areCompatibleVars [] [] = pure Pre +areCompatibleVars (x :: xs) (y :: ys) + = do compat <- areCompatibleVars xs ys + pure (Ext compat) +areCompatibleVars _ _ = Nothing diff --git a/src/Core/Name/Scoped.idr b/src/Core/Name/Scoped.idr index 8d66b57fce..948d03a74c 100644 --- a/src/Core/Name/Scoped.idr +++ b/src/Core/Name/Scoped.idr @@ -6,6 +6,8 @@ import public Data.List.HasLength import public Libraries.Data.List.SizeOf +import Core.Name.ScopedList + %default total ------------------------------------------------------------------------ @@ -20,8 +22,7 @@ import public Libraries.Data.List.SizeOf ||| Γ ⊢ λx. t : A → B public export Scope : Type -Scope = List Name --- TODO: make that a SnocList +Scope = ScopedList Name ||| A scoped definition is one indexed by a scope public export @@ -33,8 +34,8 @@ Scoped = Scope -> Type export scopeEq : (xs, ys : Scope) -> Maybe (xs = ys) -scopeEq [] [] = Just Refl -scopeEq (x :: xs) (y :: ys) +scopeEq SLNil SLNil = Just Refl +scopeEq (x :%: xs) (y :%: ys) = do Refl <- nameEq x y Refl <- scopeEq xs ys Just Refl @@ -50,116 +51,23 @@ mkFresh vs n then assert_total $ mkFresh vs (next n) else n - ------------------------------------------------------------------------- --- Compatible variables - -public export -data CompatibleVars : (xs, ys : List a) -> Type where - Pre : CompatibleVars xs xs - Ext : CompatibleVars xs ys -> CompatibleVars (n :: xs) (m :: ys) - -export -invertExt : CompatibleVars (n :: xs) (m :: ys) -> CompatibleVars xs ys -invertExt Pre = Pre -invertExt (Ext p) = p - -export -extendCompats : (args : List a) -> - CompatibleVars xs ys -> - CompatibleVars (args ++ xs) (args ++ ys) -extendCompats args Pre = Pre -extendCompats args prf = go args prf where - - go : (args : List a) -> - CompatibleVars xs ys -> - CompatibleVars (args ++ xs) (args ++ ys) - go [] prf = prf - go (x :: xs) prf = Ext (go xs prf) - -export -decCompatibleVars : (xs, ys : List a) -> Dec (CompatibleVars xs ys) -decCompatibleVars [] [] = Yes Pre -decCompatibleVars [] (x :: xs) = No (\case p impossible) -decCompatibleVars (x :: xs) [] = No (\case p impossible) -decCompatibleVars (x :: xs) (y :: ys) = case decCompatibleVars xs ys of - Yes prf => Yes (Ext prf) - No nprf => No (nprf . invertExt) - -export -areCompatibleVars : (xs, ys : List a) -> - Maybe (CompatibleVars xs ys) -areCompatibleVars [] [] = pure Pre -areCompatibleVars (x :: xs) (y :: ys) - = do compat <- areCompatibleVars xs ys - pure (Ext compat) -areCompatibleVars _ _ = Nothing - ------------------------------------------------------------------------- --- Thinnings - -public export -data Thin : List a -> List a -> Type where - Refl : Thin xs xs - Drop : Thin xs ys -> Thin xs (y :: ys) - Keep : Thin xs ys -> Thin (x :: xs) (x :: ys) - -export -none : {xs : List a} -> Thin [] xs -none {xs = []} = Refl -none {xs = _ :: _} = Drop none - -{- UNUSED: we actually sometimes want Refl vs. Keep! -||| Smart constructor. We should use this to maximise the length -||| of the Refl segment thus getting more short-circuiting behaviours -export -Keep : Thin xs ys -> Thin (x :: xs) (x :: ys) -Keep Refl = Refl -Keep p = Keep p --} - -export -keeps : (args : List a) -> Thin xs ys -> Thin (args ++ xs) (args ++ ys) -keeps [] th = th -keeps (x :: xs) th = Keep (keeps xs th) - -||| Compute the thinning getting rid of the listed de Bruijn indices. --- TODO: is the list of erased arguments guaranteed to be sorted? --- Should it? -export -removeByIndices : - (erasedArgs : List Nat) -> - (args : Scope) -> - (args' ** Thin args' args) -removeByIndices es = go 0 where - - go : (currentIdx : Nat) -> (args : Scope) -> - (args' ** Thin args' args) - go idx [] = ([] ** Refl) - go idx (x :: xs) = - let (vs ** th) = go (S idx) xs in - if idx `elem` es - then (vs ** Drop th) - else (x :: vs ** Keep th) - - ------------------------------------------------------------------------ -- Concepts public export 0 Weakenable : Scoped -> Type Weakenable tm = {0 vars, ns : Scope} -> - SizeOf ns -> tm vars -> tm (ns ++ vars) + SizeOf ns -> tm vars -> tm (ns +%+ vars) public export 0 Strengthenable : Scoped -> Type Strengthenable tm = {0 vars, ns : Scope} -> - SizeOf ns -> tm (ns ++ vars) -> Maybe (tm vars) + SizeOf ns -> tm (ns +%+ vars) -> Maybe (tm vars) public export 0 GenWeakenable : Scoped -> Type GenWeakenable tm = {0 outer, ns, local : Scope} -> - SizeOf local -> SizeOf ns -> tm (local ++ outer) -> tm (local ++ (ns ++ outer)) + SizeOf local -> SizeOf ns -> tm (local +%+ outer) -> tm (local +%+ (ns +%+ outer)) public export 0 Thinnable : Scoped -> Type @@ -171,7 +79,7 @@ Shrinkable tm = {0 xs, ys : Scope} -> tm xs -> Thin ys xs -> Maybe (tm ys) public export 0 Embeddable : Scoped -> Type -Embeddable tm = {0 outer, vars : Scope} -> tm vars -> tm (vars ++ outer) +Embeddable tm = {0 outer, vars : Scope} -> tm vars -> tm (vars +%+ outer) ------------------------------------------------------------------------ -- IsScoped interface @@ -180,7 +88,7 @@ public export interface Weaken (0 tm : Scoped) where constructor MkWeaken -- methods - weaken : tm vars -> tm (nm :: vars) + weaken : tm vars -> tm (nm :%: vars) weakenNs : Weakenable tm -- default implementations weaken = weakenNs (suc zero) @@ -193,7 +101,7 @@ interface GenWeaken (0 tm : Scoped) where export genWeaken : GenWeaken tm => - SizeOf local -> tm (local ++ outer) -> tm (local ++ n :: outer) + SizeOf local -> tm (local +%+ outer) -> tm (local +%+ n :%: outer) genWeaken l = genWeakenNs l (suc zero) public export @@ -203,7 +111,7 @@ interface Strengthen (0 tm : Scoped) where strengthenNs : Strengthenable tm export -strengthen : Strengthen tm => tm (nm :: vars) -> Maybe (tm vars) +strengthen : Strengthen tm => tm (nm :%: vars) -> Maybe (tm vars) strengthen = strengthenNs (suc zero) public export @@ -257,5 +165,5 @@ interface Weaken tm => IsScoped (0 tm : Scoped) where shrink : Shrinkable tm export -compat : IsScoped tm => tm (m :: xs) -> tm (n :: xs) +compat : IsScoped tm => tm (m :%: xs) -> tm (n :%: xs) compat = compatNs (Ext Pre) diff --git a/src/Core/Name/ScopedList.idr b/src/Core/Name/ScopedList.idr new file mode 100644 index 0000000000..cb3ffc7ac2 --- /dev/null +++ b/src/Core/Name/ScopedList.idr @@ -0,0 +1,397 @@ +module Core.Name.ScopedList + +import Data.Nat +import Data.SnocList +import Data.Vect +import Core.Name + +import public Data.Zippable + +import Libraries.Data.IntMap + +%default total + +mutual + export infixr 7 :%: + + public export + data ScopedList a = SLNil | (:%:) a (ScopedList a) + -- TODO: make that a SnocList + +export infixr 7 +%+ + +public export +length : ScopedList a -> Nat +length SLNil = Z +length (x :%: xs) = S (length xs) + +public export +take : (n : Nat) -> (xs : ScopedList a) -> ScopedList a +take (S k) (x :%: xs) = x :%: take k xs +take _ _ = SLNil + +public export +mapImpl : (a -> b) -> ScopedList a -> ScopedList b +mapImpl f SLNil = SLNil +mapImpl f (x :%: xs) = (f x) :%: (mapImpl f xs) + +public export %inline +Functor ScopedList where + map = mapImpl + +export +lengthMap : (xs : ScopedList a) -> length (mapImpl f xs) = length xs +lengthMap SLNil = Refl +lengthMap (x :%: xs) = cong S (lengthMap xs) + +public export +Zippable ScopedList where + zipWith _ SLNil _ = SLNil + zipWith _ _ SLNil = SLNil + zipWith f (x :%: xs) (y :%: ys) = f x y :%: zipWith f xs ys + + zipWith3 _ SLNil _ _ = SLNil + zipWith3 _ _ SLNil _ = SLNil + zipWith3 _ _ _ SLNil = SLNil + zipWith3 f (x :%: xs) (y :%: ys) (z :%: zs) = f x y z :%: zipWith3 f xs ys zs + + unzipWith f SLNil = (SLNil, SLNil) + unzipWith f (x :%: xs) = let (b, c) = f x + (bs, cs) = unzipWith f xs in + (b :%: bs, c :%: cs) + + unzipWith3 f SLNil = (SLNil, SLNil, SLNil) + unzipWith3 f (x :%: xs) = let (b, c, d) = f x + (bs, cs, ds) = unzipWith3 f xs in + (b :%: bs, c :%: cs, d :%: ds) + +public export +(+%+) : (xs, ys : ScopedList a) -> ScopedList a +(+%+) SLNil ys = ys +(+%+) (x :%: xs) ys = x :%: ((+%+) xs ys) + +public export +Semigroup (ScopedList a) where + (<+>) = (+%+) + +public export +Monoid (ScopedList a) where + neutral = SLNil + +public export +Foldable ScopedList where + foldr c n SLNil = n + foldr c n (x :%: xs) = c x (foldr c n xs) + + foldl f q SLNil = q + foldl f q (x :%: xs) = foldl f (f q x) xs + + null SLNil = True + null (_ :%: _) = False + + toList SLNil = [] + toList (a :%: ax) = a :: toList ax + + foldMap f = foldl (\acc, elem => acc <+> f elem) neutral + +||| Find the first element of the list that satisfies the predicate. +public export +fromList : List a -> ScopedList a +fromList [] = SLNil +fromList (a :: ax) = a :%: fromList ax + +public export +data Thin : ScopedList a -> ScopedList a -> Type where + Refl : Thin xs xs + Drop : Thin xs ys -> Thin xs (y :%: ys) + Keep : Thin xs ys -> Thin (x :%: xs) (x :%: ys) + +export +keeps : (args : ScopedList a) -> Thin xs ys -> Thin (args +%+ xs) (args +%+ ys) +keeps SLNil th = th +keeps (x :%: xs) th = Keep (keeps xs th) + +||| Ensure that the list's length is the provided natural number +public export +data HasLength : Nat -> ScopedList a -> Type where + Z : HasLength Z SLNil + S : HasLength n xs -> HasLength (S n) (x :%: xs) + +export +hasLengthAppend : HasLength m xs -> HasLength n ys -> HasLength (m + n) (xs +%+ ys) +hasLengthAppend Z ys = ys +hasLengthAppend (S xs) ys = S (hasLengthAppend xs ys) + +export +mkHasLength : (xs : ScopedList a) -> HasLength (length xs) xs +mkHasLength SLNil = Z +mkHasLength (_ :%: xs) = S (mkHasLength xs) + +public export +record SizeOf {a : Type} (xs : ScopedList a) where + constructor MkSizeOf + size : Nat + 0 hasLength : HasLength size xs + +public export +zero : SizeOf SLNil +zero = MkSizeOf Z Z + +public export +suc : SizeOf as -> SizeOf (a :%: as) +suc (MkSizeOf n p) = MkSizeOf (S n) (S p) + +public export +snoc : ScopedList a -> a -> ScopedList a +snoc xs x = xs +%+ (x :%: SLNil) + +namespace Stream + public export + take : (n : Nat) -> (xs : Stream a) -> ScopedList a + take Z xs = SLNil + take (S k) (x :: xs) = x :%: take k xs + +namespace HasLength + export + cast : {ys : _} -> (0 _ : Core.Name.ScopedList.length xs = Core.Name.ScopedList.length ys) -> HasLength m xs -> HasLength m ys + cast {ys = SLNil} eq Z = Z + cast {ys = y :%: ys} eq (S p) = S (cast (injective eq) p) + + export + take : (n : Nat) -> (xs : Stream a) -> HasLength n (take n xs) + take Z _ = Z + take (S n) (x :: xs) = S (take n xs) + + export + sucR : HasLength n xs -> HasLength (S n) (snoc xs x) + sucR Z = S Z + sucR (S n) = S (sucR n) + +namespace SizeOf + export + take : {n : Nat} -> {0 sx : Stream a} -> SizeOf (take n sx) + take = MkSizeOf n (Core.Name.ScopedList.HasLength.take n sx) + + export + sucR : SizeOf as -> SizeOf (as +%+ (a :%: SLNil)) + sucR (MkSizeOf n p) = MkSizeOf (S n) (sucR p) + + export + map : SizeOf xs -> SizeOf (mapImpl f xs) + map (MkSizeOf n p) = MkSizeOf n (cast (sym $ lengthMap xs) p) + +export +(+) : SizeOf xs -> SizeOf ys -> SizeOf (xs +%+ ys) +MkSizeOf m p + MkSizeOf n q = MkSizeOf (m + n) (hasLengthAppend p q) + +export +mkSizeOf : (xs : ScopedList a) -> SizeOf xs +mkSizeOf xs = MkSizeOf (length xs) (mkHasLength xs) + +public export +data CompatibleVars : (xs, ys : ScopedList a) -> Type where + Pre : CompatibleVars xs xs + Ext : CompatibleVars xs ys -> CompatibleVars (n :%: xs) (m :%: ys) + +public export +(<>>) : SnocList a -> ScopedList a -> ScopedList a +Lin <>> xs = xs +(sx :< x) <>> xs = sx <>> (x :%: xs) + +public export +Cast (SnocList a) (ScopedList a) where + cast sx = sx <>> SLNil + +||| 'fish': Action of lists on snoc-lists +public export +(<><) : SnocList a -> ScopedList a -> SnocList a +sx <>< SLNil = sx +sx <>< (x :%: xs) = sx :< x <>< xs + +public export +Cast (ScopedList a) (SnocList a) where + cast xs = Lin <>< xs + +||| Appending lists is associative. +export +appendAssociative : (l, c, r : ScopedList a) -> l +%+ (c +%+ r) = (l +%+ c) +%+ r +appendAssociative SLNil c r = Refl +appendAssociative (_ :%: xs) c r = rewrite appendAssociative xs c r in Refl + +export +chipsAsListAppend : (xs : SnocList a) -> (ys : ScopedList a) -> xs <>> ys = cast xs +%+ ys +chipsAsListAppend [<] ys = Refl +chipsAsListAppend (sx :< x) ys = do + rewrite chipsAsListAppend sx (x :%: ys) + rewrite chipsAsListAppend sx (x :%: SLNil) + rewrite sym $ appendAssociative (cast sx) (x :%: SLNil) ys + Refl + +public export +data Bounds : ScopedList Name -> Type where + None : Bounds SLNil + Add : (x : Name) -> Name -> Bounds xs -> Bounds (x :%: xs) + +export +sizeOf : Bounds xs -> SizeOf xs +sizeOf None = zero +sizeOf (Add _ _ b) = suc (sizeOf b) + +export +appendNilRightNeutral : (l : ScopedList a) -> l +%+ SLNil = l +appendNilRightNeutral SLNil = Refl +appendNilRightNeutral (_ :%: xs) = rewrite appendNilRightNeutral xs in Refl + +export +Show a => Show (ScopedList a) where + show xs = "[%>" ++ show' "" xs ++ "<%]" + where + show' : String -> ScopedList a -> String + show' acc SLNil = acc + show' acc (x :%: SLNil) = acc ++ show x + show' acc (x :%: xs) = show' (acc ++ show x ++ ", ") xs + + +||| Reverse the second list, prepending its elements to the first. +public export +reverseOnto : ScopedList a -> ScopedList a -> ScopedList a +reverseOnto acc SLNil = acc +reverseOnto acc (x:%:xs) = reverseOnto (x:%:acc) xs + +||| Reverses the given list. +public export +reverse : ScopedList a -> ScopedList a +reverse = reverseOnto SLNil + +hasLengthReverseOnto : HasLength m acc -> HasLength n xs -> HasLength (m + n) (reverseOnto acc xs) +hasLengthReverseOnto p Z = rewrite plusZeroRightNeutral m in p +hasLengthReverseOnto {n = S n} p (S q) = rewrite sym (plusSuccRightSucc m n) in hasLengthReverseOnto (S p) q + +export +hasLengthReverse : HasLength m acc -> HasLength m (reverse acc) +hasLengthReverse = hasLengthReverseOnto Z + +||| Apply a partial function to the elements of a list, keeping the ones at which it is defined. +public export +mapMaybe : (a -> Maybe b) -> ScopedList a -> ScopedList b +mapMaybe f SLNil = SLNil +mapMaybe f (x:%:xs) = + case f x of + Nothing => mapMaybe f xs + Just j => j :%: mapMaybe f xs + +public export +listBindOnto : (a -> ScopedList b) -> ScopedList b -> ScopedList a -> ScopedList b +listBindOnto f xs SLNil = reverse xs +listBindOnto f xs (y :%: ys) = listBindOnto f (reverseOnto xs (f y)) ys + +-- tail recursive O(n) implementation of `(>>=)` for `List` +public export +listBind : ScopedList a -> (a -> ScopedList b) -> ScopedList b +listBind as f = listBindOnto f SLNil as + +public export +Applicative ScopedList where + pure x = (x :%: SLNil) + fs <*> vs = listBind fs (\f => map f vs) + +public export +Traversable ScopedList where + traverse f SLNil = pure SLNil + traverse f (x:%:xs) = [| f x :%: traverse f xs |] + +||| List.length is distributive over appending. +export +lengthDistributesOverAppend : (xs, ys : ScopedList a) -> length (xs +%+ ys) = length xs + length ys +lengthDistributesOverAppend SLNil ys = Refl +lengthDistributesOverAppend (x :%: xs) ys = + cong S $ lengthDistributesOverAppend xs ys + +||| Boolean check for whether the list is the empty list. +public export +isNil : ScopedList a -> Bool +isNil SLNil = True +isNil _ = False + +public export +toVect : (n : Nat) -> ScopedList a -> Maybe (Vect n a) +toVect Z SLNil = Just [] +toVect (S k) (x :%: xs) + = do xs' <- toVect k xs + pure (x :: xs') +toVect _ _ = Nothing + +namespace SizedView + + public export + data SizedView : SizeOf as -> Type where + Z : SizedView (MkSizeOf Z Z) + S : (n : SizeOf as) -> SizedView (suc {a} n) + +export +sizedView : (p : SizeOf as) -> SizedView p +sizedView (MkSizeOf Z Z) = Z +sizedView (MkSizeOf (S n) (S p)) = S (MkSizeOf n p) + +||| Lookup a value at a given position +export +getAt : Nat -> ScopedList a -> Maybe a +getAt Z (x :%: xs) = Just x +getAt (S k) (x :%: xs) = getAt k xs +getAt _ SLNil = Nothing + +public export +drop : (n : Nat) -> (xs : ScopedList a) -> ScopedList a +drop Z xs = xs +drop (S n) SLNil = SLNil +drop (S n) (_:%:xs) = drop n xs + +public export +data LengthMatch : ScopedList a -> ScopedList b -> Type where + NilMatch : LengthMatch SLNil SLNil + ConsMatch : LengthMatch xs ys -> LengthMatch (x :%: xs) (y :%: ys) + +export +checkLengthMatch : (xs : ScopedList a) -> (ys : ScopedList b) -> Maybe (LengthMatch xs ys) +checkLengthMatch SLNil SLNil = Just NilMatch +checkLengthMatch SLNil (x :%: xs) = Nothing +checkLengthMatch (x :%: xs) SLNil = Nothing +checkLengthMatch (x :%: xs) (y :%: ys) + = Just (ConsMatch !(checkLengthMatch xs ys)) + +export +lengthsMatch : LengthMatch xs ys -> (length xs) = (length ys) +lengthsMatch NilMatch = Refl +lengthsMatch (ConsMatch x) = cong (S) (lengthsMatch x) + +namespace IntMap + export + toScopedList : IntMap v -> ScopedList (Int, v) + toScopedList = fromList . toList + +public export +Eq a => Eq (ScopedList a) where + SLNil == SLNil = True + x :%: xs == y :%: ys = x == y && xs == ys + _ == _ = False + +public export +Ord a => Ord (ScopedList a) where + compare SLNil SLNil = EQ + compare SLNil (x :%: xs) = LT + compare (x :%: xs) SLNil = GT + compare (x :%: xs) (y :%: ys) + = case compare x y of + EQ => compare xs ys + c => c + +||| Find the first element of the list that satisfies the predicate. +public export +find : (p : a -> Bool) -> (xs : ScopedList a) -> Maybe a +find p SLNil = Nothing +find p (x:%:xs) = if p x then Just x else find p xs + +export +none : {xs : ScopedList a} -> Thin SLNil xs +none {xs = SLNil} = Refl +none {xs = _ :%: _} = Drop none diff --git a/src/Core/Name/ScopedList/HasLength.idr b/src/Core/Name/ScopedList/HasLength.idr new file mode 100644 index 0000000000..15eb9beb81 --- /dev/null +++ b/src/Core/Name/ScopedList/HasLength.idr @@ -0,0 +1,24 @@ +module Core.Name.ScopedList.HasLength + +import Data.Nat + +import Libraries.Data.SnocList.HasLength + +--------------------------------------- +-- horrible hack +import Core.Name.ScopedList as L + +public export +LHasLength : Nat -> ScopedList a -> Type +LHasLength = L.HasLength +%hide L.HasLength +--------------------------------------- + +%hide Libraries.Data.SnocList.HasLength.LHasLength + +export +hlChips : HasLength m sx -> LHasLength n ys -> LHasLength (m + n) (sx <>> ys) +hlChips Z y = y +hlChips {m = S m} {n} (S x) y + = rewrite plusSuccRightSucc m n in + hlChips x (S y) diff --git a/src/Core/Name/ScopedList/Name.idr b/src/Core/Name/ScopedList/Name.idr new file mode 100644 index 0000000000..50aab0d8ff --- /dev/null +++ b/src/Core/Name/ScopedList/Name.idr @@ -0,0 +1,15 @@ +module Core.Name.ScopedList.Name + +import Core.Name +import Core.Name.ScopedList + +export +namesEq : (xs, ys : ScopedList Name) -> Maybe (xs = ys) +namesEq SLNil SLNil = Just Refl +namesEq (x :%: xs) (y :%: ys) + = do p <- nameEq x y + ps <- namesEq xs ys + rewrite p + rewrite ps + Just Refl +namesEq _ _ = Nothing diff --git a/src/Core/Name/ScopedList/SizeOf.idr b/src/Core/Name/ScopedList/SizeOf.idr new file mode 100644 index 0000000000..f4a573fe12 --- /dev/null +++ b/src/Core/Name/ScopedList/SizeOf.idr @@ -0,0 +1,28 @@ +module Core.Name.ScopedList.SizeOf + +import Data.Nat + +import Libraries.Data.SnocList.SizeOf + +import Core.Name.ScopedList.HasLength + +--------------------------------------- +-- horrible hack +import Core.Name.ScopedList as L + +public export +0 LSizeOf : {0 a : Type} -> ScopedList a -> Type +LSizeOf xs = L.SizeOf xs + +export +reverse : L.SizeOf xs -> L.SizeOf (reverse xs) +reverse (L.MkSizeOf n p) = L.MkSizeOf n (hasLengthReverse p) + +%hide L.SizeOf +--------------------------------------- + +%hide Libraries.Data.SnocList.SizeOf.LSizeOf + +public export +(<>>) : SizeOf {a} sx -> LSizeOf {a} ys -> LSizeOf (sx <>> ys) +MkSizeOf m p <>> MkSizeOf n q = MkSizeOf (m + n) (hlChips p q) diff --git a/src/Core/Name/Thin.idr b/src/Core/Name/Thin.idr new file mode 100644 index 0000000000..7eecedeb12 --- /dev/null +++ b/src/Core/Name/Thin.idr @@ -0,0 +1,48 @@ +module Core.Name.Thin + +-- Thinnings + +public export +data Thin : List a -> List a -> Type where + Refl : Thin xs xs + Drop : Thin xs ys -> Thin xs (y :: ys) + Keep : Thin xs ys -> Thin (x :: xs) (x :: ys) + +export +none : {xs : List a} -> Thin [] xs +none {xs = []} = Refl +none {xs = _ :: _} = Drop none + +{- UNUSED: we actually sometimes want Refl vs. Keep! +||| Smart constructor. We should use this to maximise the length +||| of the Refl segment thus getting more short-circuiting behaviours +export +Keep : Thin xs ys -> Thin (x :: xs) (x :: ys) +Keep Refl = Refl +Keep p = Keep p +-} + +export +keeps : (args : List a) -> Thin xs ys -> Thin (args ++ xs) (args ++ ys) +keeps [] th = th +keeps (x :: xs) th = Keep (keeps xs th) + +||| Compute the thinning getting rid of the listed de Bruijn indices. +-- TODO: is the list of erased arguments guaranteed to be sorted? +-- Should it? +export +removeByIndices : + (erasedArgs : List Nat) -> + (args : Scope) -> + (args' ** Thin args' args) +removeByIndices es = go 0 where + + go : (currentIdx : Nat) -> (args : Scope) -> + (args' ** Thin args' args) + go idx SLNil = (SLNil ** Refl) + go idx (x :%: xs) = + let (vs ** th) = go (S idx) xs in + if idx `elem` es + then (vs ** Drop th) + else (x :%: vs ** Keep th) + diff --git a/src/Core/Normalise.idr b/src/Core/Normalise.idr index 787d22dff8..986b335e00 100644 --- a/src/Core/Normalise.idr +++ b/src/Core/Normalise.idr @@ -18,7 +18,7 @@ import Core.Value -- reduce export normalisePis : {auto c : Ref Ctxt Defs} -> - {vars : List Name} -> + {vars : ScopedList Name} -> Defs -> Env Term vars -> Term vars -> Core (Term vars) normalisePis defs env tm = do tmnf <- nf defs env tm @@ -240,13 +240,13 @@ logEnv str n msg env where - dumpEnv : {vs : List Name} -> Env Term vs -> Core () + dumpEnv : {vs : ScopedList Name} -> Env Term vs -> Core () dumpEnv [] = pure () - dumpEnv {vs = x :: _} (Let _ c val ty :: bs) + dumpEnv {vs = x :%: _} (Let _ c val ty :: bs) = do logTermNF' str n (msg ++ ": let " ++ show x) bs val logTermNF' str n (msg ++ ":" ++ show c ++ " " ++ show x) bs ty dumpEnv bs - dumpEnv {vs = x :: _} (b :: bs) + dumpEnv {vs = x :%: _} (b :: bs) = do logTermNF' str n (msg ++ ":" ++ show (multiplicity b) ++ " " ++ show (piInfo b) ++ " " ++ show x) bs (binderType b) @@ -273,25 +273,25 @@ replace' {vars} tmpi defs env lhs parg tm sc' <- replace' (tmpi + 1) defs env lhs parg !(scfn defs (toClosure defaultOpts env (Ref fc Bound x'))) pure (Bind fc x b' (refsToLocals (Add x x' None) sc')) - repSub (NApp fc hd []) + repSub (NApp fc hd SLNil) = do empty <- clearDefs defs - quote empty env (NApp fc hd []) + quote empty env (NApp fc hd SLNil) repSub (NApp fc hd args) = do args' <- traverse (traversePair repArg) args pure $ applyStackWithFC - !(replace' tmpi defs env lhs parg (NApp fc hd [])) + !(replace' tmpi defs env lhs parg (NApp fc hd SLNil)) args' repSub (NDCon fc n t a args) = do args' <- traverse (traversePair repArg) args empty <- clearDefs defs pure $ applyStackWithFC - !(quote empty env (NDCon fc n t a [])) + !(quote empty env (NDCon fc n t a SLNil)) args' repSub (NTCon fc n t a args) = do args' <- traverse (traversePair repArg) args empty <- clearDefs defs pure $ applyStackWithFC - !(quote empty env (NTCon fc n t a [])) + !(quote empty env (NTCon fc n t a SLNil)) args' repSub (NAs fc s a p) = do a' <- repSub a @@ -336,7 +336,7 @@ normalisePrims : {auto c : Ref Ctxt Defs} -> {vs : _} -> List Name -> -- view of the potential redex (n : Name) -> -- function name - (args : List arg) -> -- arguments from inside out (arg1, ..., argk) + (args : ScopedList arg) -> -- arguments from inside out (arg1, ..., argk) -- actual term to evaluate if needed (tm : Term vs) -> -- original term (n arg1 ... argk) Env Term vs -> -- evaluation environment @@ -345,7 +345,7 @@ normalisePrims : {auto c : Ref Ctxt Defs} -> {vs : _} -> normalisePrims boundSafe viewConstant all prims n args tm env = do let True = isPrimName prims !(getFullName n) -- is a primitive | _ => pure Nothing - let (mc :: _) = reverse args -- with at least one argument + let (mc :%: _) = reverse args -- with at least one argument | _ => pure Nothing let (Just c) = viewConstant mc -- that is a constant | _ => pure Nothing diff --git a/src/Core/Normalise/Convert.idr b/src/Core/Normalise/Convert.idr index 4639d4b269..9eebf04f49 100644 --- a/src/Core/Normalise/Convert.idr +++ b/src/Core/Normalise/Convert.idr @@ -8,6 +8,7 @@ import Core.Context import Core.Core import Core.Env import Core.TT +import Core.Name.ScopedList import Core.Value import Data.List @@ -17,11 +18,11 @@ import Data.List public export interface Convert tm where convert : {auto c : Ref Ctxt Defs} -> - {vars : List Name} -> + {vars : ScopedList Name} -> Defs -> Env Term vars -> tm vars -> tm vars -> Core Bool convertInf : {auto c : Ref Ctxt Defs} -> - {vars : List Name} -> + {vars : ScopedList Name} -> Defs -> Env Term vars -> tm vars -> tm vars -> Core Bool @@ -42,15 +43,15 @@ interface Convert tm where convGen q True defs env tm tm' tryUpdate : {vars, vars' : _} -> - List (Var vars, Var vars') -> + ScopedList (Var vars, Var vars') -> Term vars -> Maybe (Term vars') tryUpdate ms (Local fc l idx p) = do MkVar p' <- findIdx ms (MkVar p) pure $ Local fc l _ p' where - findIdx : List (Var vars, Var vars') -> Var vars -> Maybe (Var vars') - findIdx [] _ = Nothing - findIdx ((old, v) :: ps) n + findIdx : ScopedList (Var vars, Var vars') -> Var vars -> Maybe (Var vars') + findIdx SLNil _ = Nothing + findIdx ((old, v) :%: ps) n = if old == n then Just v else findIdx ps n tryUpdate ms (Ref fc nt n) = pure $ Ref fc nt n tryUpdate ms (Meta fc n i args) = pure $ Meta fc n i !(traverse (tryUpdate ms) args) @@ -71,7 +72,7 @@ tryUpdate ms (Bind fc x b sc) tryUpdateB _ = Nothing weakenP : {n : _} -> (Var vars, Var vars') -> - (Var (n :: vars), Var (n :: vars')) + (Var (n :%: vars), Var (n :%: vars')) weakenP (v, vs) = (weaken v, weaken vs) tryUpdate ms (App fc f a) = pure $ App fc !(tryUpdate ms f) !(tryUpdate ms a) tryUpdate ms (As fc s a p) = pure $ As fc s !(tryUpdate ms a) !(tryUpdate ms p) @@ -86,9 +87,9 @@ mutual allConvNF : {auto c : Ref Ctxt Defs} -> {vars : _} -> Ref QVar Int -> Bool -> Defs -> Env Term vars -> - List (NF vars) -> List (NF vars) -> Core Bool - allConvNF q i defs env [] [] = pure True - allConvNF q i defs env (x :: xs) (y :: ys) + ScopedList (NF vars) -> ScopedList (NF vars) -> Core Bool + allConvNF q i defs env SLNil SLNil = pure True + allConvNF q i defs env (x :%: xs) (y :%: ys) = do ok <- allConvNF q i defs env xs ys if ok then convGen q i defs env x y else pure False @@ -97,9 +98,9 @@ mutual -- return False if anything differs at the head, to quickly find -- conversion failures without going deeply into all the arguments. -- True means they might still match - quickConv : List (NF vars) -> List (NF vars) -> Bool - quickConv [] [] = True - quickConv (x :: xs) (y :: ys) = quickConvArg x y && quickConv xs ys + quickConv : ScopedList (NF vars) -> ScopedList (NF vars) -> Bool + quickConv SLNil SLNil = True + quickConv (x :%: xs) (y :%: ys) = quickConvArg x y && quickConv xs ys where quickConvHead : NHead vars -> NHead vars -> Bool quickConvHead (NLocal _ _ _) (NLocal _ _ _) = True @@ -127,7 +128,7 @@ mutual allConv : {auto c : Ref Ctxt Defs} -> {vars : _} -> Ref QVar Int -> Bool -> Defs -> Env Term vars -> - List (Closure vars) -> List (Closure vars) -> Core Bool + ScopedList (Closure vars) -> ScopedList (Closure vars) -> Core Bool allConv q i defs env xs ys = do xsnf <- traverse (evalClosure defs) xs ysnf <- traverse (evalClosure defs) ys @@ -140,9 +141,9 @@ mutual getMatchingVarAlt : {auto c : Ref Ctxt Defs} -> {args, args' : _} -> Defs -> - List (Var args, Var args') -> + ScopedList (Var args, Var args') -> CaseAlt args -> CaseAlt args' -> - Core (Maybe (List (Var args, Var args'))) + Core (Maybe (ScopedList (Var args, Var args'))) getMatchingVarAlt defs ms (ConCase n tag cargs t) (ConCase n' tag' cargs' t') = if n == n' then do let Just ms' = extend cargs cargs' ms @@ -156,27 +157,27 @@ mutual where weakenP : {0 c, c' : _} -> {0 args, args' : Scope} -> (Var args, Var args') -> - (Var (c :: args), Var (c' :: args')) + (Var (c :%: args), Var (c' :%: args')) weakenP (v, vs) = (weaken v, weaken vs) - extend : (cs : List Name) -> (cs' : List Name) -> - (List (Var args, Var args')) -> - Maybe (List (Var (cs ++ args), Var (cs' ++ args'))) - extend [] [] ms = pure ms - extend (c :: cs) (c' :: cs') ms + extend : (cs : ScopedList Name) -> (cs' : ScopedList Name) -> + (ScopedList (Var args, Var args')) -> + Maybe (ScopedList (Var (cs +%+ args), Var (cs' +%+ args'))) + extend SLNil SLNil ms = pure ms + extend (c :%: cs) (c' :%: cs') ms = do rest <- extend cs cs' ms - pure ((MkVar First, MkVar First) :: map weakenP rest) + pure ((MkVar First, MkVar First) :%: map weakenP rest) extend _ _ _ = Nothing dropV : forall args . - (cs : List Name) -> Var (cs ++ args) -> Maybe (Var args) - dropV [] v = Just v - dropV (c :: cs) (MkVar First) = Nothing - dropV (c :: cs) (MkVar (Later x)) + (cs : ScopedList Name) -> Var (cs +%+ args) -> Maybe (Var args) + dropV SLNil v = Just v + dropV (c :%: cs) (MkVar First) = Nothing + dropV (c :%: cs) (MkVar (Later x)) = dropV cs (MkVar x) - dropP : (cs : List Name) -> (cs' : List Name) -> - (Var (cs ++ args), Var (cs' ++ args')) -> + dropP : (cs : ScopedList Name) -> (cs' : ScopedList Name) -> + (Var (cs +%+ args), Var (cs' +%+ args')) -> Maybe (Var args, Var args') dropP cs cs' (x, y) = pure (!(dropV cs x), !(dropV cs' y)) @@ -191,9 +192,9 @@ mutual getMatchingVarAlts : {auto c : Ref Ctxt Defs} -> {args, args' : _} -> Defs -> - List (Var args, Var args') -> + ScopedList (Var args, Var args') -> List (CaseAlt args) -> List (CaseAlt args') -> - Core (Maybe (List (Var args, Var args'))) + Core (Maybe (ScopedList (Var args, Var args'))) getMatchingVarAlts defs ms [] [] = pure (Just ms) getMatchingVarAlts defs ms (a :: as) (a' :: as') = do Just ms <- getMatchingVarAlt defs ms a a' @@ -204,11 +205,11 @@ mutual getMatchingVars : {auto c : Ref Ctxt Defs} -> {args, args' : _} -> Defs -> - List (Var args, Var args') -> + ScopedList (Var args, Var args') -> CaseTree args -> CaseTree args' -> - Core (Maybe (List (Var args, Var args'))) + Core (Maybe (ScopedList (Var args, Var args'))) getMatchingVars defs ms (Case _ p _ alts) (Case _ p' _ alts') - = getMatchingVarAlts defs ((MkVar p, MkVar p') :: ms) alts alts' + = getMatchingVarAlts defs ((MkVar p, MkVar p') :%: ms) alts alts' getMatchingVars defs ms (STerm i tm) (STerm i' tm') = do let Just tm'' = tryUpdate ms tm | Nothing => pure Nothing @@ -223,7 +224,7 @@ mutual {vars : _} -> Ref QVar Int -> Bool -> Defs -> Env Term vars -> Name -> Name -> - List (Closure vars) -> List (Closure vars) -> Core Bool + ScopedList (Closure vars) -> ScopedList (Closure vars) -> Core Bool chkSameDefs q i defs env n n' nargs nargs' = do Just (PMDef _ args ct rt _) <- lookupDefExact n (gamma defs) | _ => pure False @@ -233,22 +234,22 @@ mutual -- If the two case blocks match in structure, get which variables -- correspond. If corresponding variables convert, the two case -- blocks convert. - Just ms <- getMatchingVars defs [] ct ct' + Just ms <- getMatchingVars defs SLNil ct ct' | Nothing => pure False convertMatches ms where -- We've only got the index into the argument list, and the indices -- don't match up, which is annoying. But it'll always be there! - getArgPos : Nat -> List (Closure vars) -> Maybe (Closure vars) - getArgPos _ [] = Nothing - getArgPos Z (c :: cs) = pure c - getArgPos (S k) (c :: cs) = getArgPos k cs + getArgPos : Nat -> ScopedList (Closure vars) -> Maybe (Closure vars) + getArgPos _ SLNil = Nothing + getArgPos Z (c :%: cs) = pure c + getArgPos (S k) (c :%: cs) = getArgPos k cs convertMatches : {vs, vs' : _} -> - List (Var vs, Var vs') -> + ScopedList (Var vs, Var vs') -> Core Bool - convertMatches [] = pure True - convertMatches ((MkVar {varIdx = ix} p, MkVar {varIdx = iy} p') :: vs) + convertMatches SLNil = pure True + convertMatches ((MkVar {varIdx = ix} p, MkVar {varIdx = iy} p') :%: vs) = do let Just varg = getArgPos ix nargs | Nothing => pure False let Just varg' = getArgPos iy nargs' @@ -261,8 +262,8 @@ mutual chkConvCaseBlock : {auto c : Ref Ctxt Defs} -> {vars : _} -> FC -> Ref QVar Int -> Bool -> Defs -> Env Term vars -> - NHead vars -> List (Closure vars) -> - NHead vars -> List (Closure vars) -> Core Bool + NHead vars -> ScopedList (Closure vars) -> + NHead vars -> ScopedList (Closure vars) -> Core Bool chkConvCaseBlock fc q i defs env (NRef _ n) nargs (NRef _ n') nargs' = do NS _ (CaseBlock _ _) <- full (gamma defs) n | _ => pure False @@ -301,9 +302,9 @@ mutual findArgPos (Case idx p _ _) = Just idx findArgPos _ = Nothing - getScrutinee : Nat -> List (Closure vs) -> Maybe (Closure vs) - getScrutinee Z (x :: xs) = Just x - getScrutinee (S k) (x :: xs) = getScrutinee k xs + getScrutinee : Nat -> ScopedList (Closure vs) -> Maybe (Closure vs) + getScrutinee Z (x :%: xs) = Just x + getScrutinee (S k) (x :%: xs) = getScrutinee k xs getScrutinee _ _ = Nothing chkConvCaseBlock _ _ _ _ _ _ _ _ _ = pure False @@ -379,19 +380,19 @@ mutual else pure [] getInfPos _ = pure [] - dropInf : Nat -> List Nat -> List a -> List a + dropInf : Nat -> List Nat -> ScopedList a -> ScopedList a dropInf _ [] xs = xs - dropInf _ _ [] = [] - dropInf i ds (x :: xs) + dropInf _ _ SLNil = SLNil + dropInf i ds (x :%: xs) = if i `elem` ds then dropInf (S i) ds xs - else x :: dropInf (S i) ds xs + else x :%: dropInf (S i) ds xs -- Discard file context information irrelevant for conversion checking - args1 : List (Closure vars) + args1 : ScopedList (Closure vars) args1 = map snd args - args2 : List (Closure vars) + args2 : ScopedList (Closure vars) args2 = map snd args' convGen q i defs env (NDCon _ nm tag _ args) (NDCon _ nm' tag' _ args') diff --git a/src/Core/Normalise/Eval.idr b/src/Core/Normalise/Eval.idr index d141c0ff7a..5015252187 100644 --- a/src/Core/Normalise/Eval.idr +++ b/src/Core/Normalise/Eval.idr @@ -7,6 +7,7 @@ import Core.Core import Core.Env import Core.Primitives import Core.TT +import Core.Name.ScopedList import Core.Value import Data.List @@ -23,7 +24,7 @@ import Libraries.Data.WithDefault -- from a term (via 'gnf') or a normal form (via 'glueBack') but the other -- part will only be constructed when needed, because it's in Core. public export -data Glued : List Name -> Type where +data Glued : ScopedList Name -> Type where MkGlue : (fromTerm : Bool) -> -- is it built from the term; i.e. can -- we read the term straight back? Core (Term vars) -> (Ref Ctxt Defs -> Core (NF vars)) -> Glued vars @@ -41,14 +42,14 @@ getNF : {auto c : Ref Ctxt Defs} -> Glued vars -> Core (NF vars) getNF {c} (MkGlue _ _ nf) = nf c public export -Stack : List Name -> Type +Stack : ScopedList Name -> Type Stack vars = List (FC, Closure vars) evalWithOpts : {auto c : Ref Ctxt Defs} -> {free, vars : _} -> Defs -> EvalOpts -> Env Term free -> LocalEnv free vars -> - Term (vars ++ free) -> Stack free -> Core (NF free) + Term (vars +%+ free) -> Stack free -> Core (NF free) export evalClosure : {auto c : Ref Ctxt Defs} -> @@ -85,31 +86,33 @@ data CaseResult a | NoMatch -- case alternative didn't match anything | GotStuck -- alternative matched, but got stuck later -record TermWithEnv (free : List Name) where +record TermWithEnv (free : ScopedList Name) where constructor MkTermEnv - { varsEnv : List Name } + { varsEnv : ScopedList Name } locEnv : LocalEnv free varsEnv - term : Term $ varsEnv ++ free + term : Term $ varsEnv +%+ free parameters (defs : Defs, topopts : EvalOpts) mutual eval : {auto c : Ref Ctxt Defs} -> {free, vars : _} -> Env Term free -> LocalEnv free vars -> - Term (vars ++ free) -> Stack free -> Core (NF free) + Term (vars +%+ free) -> Stack free -> Core (NF free) eval env locs (Local fc mrig idx prf) stk = evalLocal env fc mrig idx prf stk locs eval env locs (Ref fc nt fn) stk - = evalRef env False fc nt fn stk (NApp fc (NRef nt fn) stk) + = + let scoped_list_stk = fromList stk + in evalRef env False fc nt fn scoped_list_stk (NApp fc (NRef nt fn) scoped_list_stk) eval {vars} {free} env locs (Meta fc name idx args) stk = evalMeta env fc name idx (closeArgs args) stk where -- Yes, it's just a map, but specialising it by hand since we -- use this a *lot* and it saves the run time overhead of making -- a closure and calling APPLY. - closeArgs : List (Term (vars ++ free)) -> List (Closure free) - closeArgs [] = [] - closeArgs (t :: ts) = MkClosure topopts locs env t :: closeArgs ts + closeArgs : List (Term (vars +%+ free)) -> ScopedList (Closure free) + closeArgs [] = SLNil + closeArgs (t :: ts) = MkClosure topopts locs env t :%: closeArgs ts eval env locs (Bind fc x (Lam _ r _ ty) scope) (thunk :: stk) = eval env (snd thunk :: locs) scope stk eval env locs (Bind fc x b@(Let _ r val ty) scope) stk @@ -146,7 +149,7 @@ parameters (defs : Defs, topopts : EvalOpts) case tm' of NDelay fc r _ arg => eval env (arg :: locs) (Local {name = UN (Basic "fvar")} fc Nothing _ First) stk - _ => pure (NForce fc r tm' stk) + _ => pure (NForce fc r tm' (fromList stk)) eval env locs (PrimVal fc c) stk = pure $ NPrimVal fc c eval env locs (Erased fc a) stk = NErased fc <$> traverse @{%search} @{CORE} (\ t => eval env locs t stk) a @@ -171,16 +174,19 @@ parameters (defs : Defs, topopts : EvalOpts) = pure (NBind fc x b (\defs', arg => applyToStack env !(sc defs' arg) stk)) applyToStack env (NApp fc (NRef nt fn) args) stk - = evalRef env False fc nt fn (args ++ stk) - (NApp fc (NRef nt fn) (args ++ stk)) + = + let scoped_list_stack = fromList stk + combined_stack = args +%+ scoped_list_stack + in evalRef env False fc nt fn combined_stack + (NApp fc (NRef nt fn) combined_stack) applyToStack env (NApp fc (NLocal mrig idx p) args) stk - = evalLocal env fc mrig _ p (args ++ stk) [] + = evalLocal env fc mrig _ p ((toList args) ++ stk) [] applyToStack env (NApp fc (NMeta n i args) args') stk - = evalMeta env fc n i args (args' ++ stk) + = evalMeta env fc n i args ((toList args') ++ stk) applyToStack env (NDCon fc n t a args) stk - = pure $ NDCon fc n t a (args ++ stk) + = pure $ NDCon fc n t a (args +%+ fromList stk) applyToStack env (NTCon fc n t a args) stk - = pure $ NTCon fc n t a (args ++ stk) + = pure $ NTCon fc n t a (args +%+ fromList stk) applyToStack env (NAs fc s p t) stk = if removeAs topopts then applyToStack env t stk @@ -197,7 +203,7 @@ parameters (defs : Defs, topopts : EvalOpts) case tm' of NDelay fc r _ arg => eval env [arg] (Local {name = UN (Basic "fvar")} fc Nothing _ First) stk - _ => pure (NForce fc r tm' (args ++ stk)) + _ => pure (NForce fc r tm' (args +%+ fromList stk)) applyToStack env nf@(NPrimVal fc _) _ = pure nf applyToStack env (NErased fc a) stk = NErased fc <$> traverse @{%search} @{CORE} (\ t => applyToStack env t stk) a @@ -219,7 +225,7 @@ parameters (defs : Defs, topopts : EvalOpts) {free : _} -> Env Term free -> FC -> Maybe Bool -> - (idx : Nat) -> (0 p : IsVar nm idx (vars ++ free)) -> + (idx : Nat) -> (0 p : IsVar nm idx (vars +%+ free)) -> Stack free -> LocalEnv free vars -> Core (NF free) @@ -232,16 +238,16 @@ parameters (defs : Defs, topopts : EvalOpts) then case getBinder prf env of Let _ _ val _ => eval env [] val stk - _ => pure $ NApp fc (NLocal mrig idx prf) stk - else pure $ NApp fc (NLocal mrig idx prf) stk + _ => pure $ NApp fc (NLocal mrig idx prf) (fromList stk) + else pure $ NApp fc (NLocal mrig idx prf) (fromList stk) evalLocal env fc mrig Z First stk (x :: locs) = evalLocClosure env fc mrig stk x - evalLocal {vars = x :: xs} {free} + evalLocal {vars = x :%: xs} {free} env fc mrig (S idx) (Later p) stk (_ :: locs) = evalLocal {vars = xs} env fc mrig idx p stk locs updateLocal : EvalOpts -> Env Term free -> - (idx : Nat) -> (0 p : IsVar nm idx (vars ++ free)) -> + (idx : Nat) -> (0 p : IsVar nm idx (vars +%+ free)) -> LocalEnv free vars -> NF free -> LocalEnv free vars updateLocal opts env Z First (x :: locs) nf @@ -253,14 +259,16 @@ parameters (defs : Defs, topopts : EvalOpts) evalMeta : {auto c : Ref Ctxt Defs} -> {free : _} -> Env Term free -> - FC -> Name -> Int -> List (Closure free) -> + FC -> Name -> Int -> ScopedList (Closure free) -> Stack free -> Core (NF free) evalMeta env fc nm i args stk - = let args' = if isNil stk then map (EmptyFC,) args - else map (EmptyFC,) args ++ stk + = let + scoped_list_stack = fromList stk + args' = if isNil stk then map (EmptyFC,) args + else (map (EmptyFC,) args) +%+ scoped_list_stack in evalRef env True fc Func (Resolved i) args' - (NApp fc (NMeta nm i args) stk) + (NApp fc (NMeta nm i args) scoped_list_stack) -- The commented out logging here might still be useful one day, but -- evalRef is used a lot and even these tiny checks turn out to be @@ -269,7 +277,7 @@ parameters (defs : Defs, topopts : EvalOpts) {free : _} -> Env Term free -> (isMeta : Bool) -> - FC -> NameType -> Name -> Stack free -> (def : Lazy (NF free)) -> + FC -> NameType -> Name -> ScopedList (FC, Closure free) -> (def : Lazy (NF free)) -> Core (NF free) evalRef env meta fc (DataCon tag arity) fn stk def = do -- logC "eval.ref.data" 50 $ do fn' <- toFullNames fn -- Can't use ! here, it gets lifted too far @@ -315,14 +323,14 @@ parameters (defs : Defs, topopts : EvalOpts) pure nf else pure def - getCaseBound : List (Closure free) -> - (args : List Name) -> + getCaseBound : ScopedList (Closure free) -> + (args : ScopedList Name) -> LocalEnv free more -> - Maybe (LocalEnv free (args ++ more)) - getCaseBound [] [] loc = Just loc - getCaseBound [] (_ :: _) loc = Nothing -- mismatched arg length - getCaseBound (arg :: args) [] loc = Nothing -- mismatched arg length - getCaseBound (arg :: args) (n :: ns) loc = (arg ::) <$> getCaseBound args ns loc + Maybe (LocalEnv free (args +%+ more)) + getCaseBound SLNil SLNil loc = Just loc + getCaseBound SLNil (_ :%: _) loc = Nothing -- mismatched arg length + getCaseBound (arg :%: args) SLNil loc = Nothing -- mismatched arg length + getCaseBound (arg :%: args) (n :%: ns) loc = (arg ::) <$> (getCaseBound args ns loc) -- Returns the case term from the matched pattern with the LocalEnv (arguments from constructor pattern ConCase) evalConAlt : {auto c : Ref Ctxt Defs} -> @@ -330,9 +338,9 @@ parameters (defs : Defs, topopts : EvalOpts) Env Term free -> LocalEnv free more -> EvalOpts -> FC -> Stack free -> - (args : List Name) -> - List (Closure free) -> - CaseTree (args ++ more) -> + (args : ScopedList Name) -> + ScopedList (Closure free) -> + CaseTree (args +%+ more) -> Core (CaseResult (TermWithEnv free)) evalConAlt env loc opts fc stk args args' sc = do let Just bound = getCaseBound args' args loc @@ -361,21 +369,20 @@ parameters (defs : Defs, topopts : EvalOpts) -- Primitive type matching, in typecase tryAlt env loc opts fc stk (NPrimVal _ c) (ConCase nm tag args sc) = case args of -- can't just test for it in the `if` for typing reasons - [] => if UN (Basic $ show c) == nm + SLNil => if UN (Basic $ show c) == nm then evalTree env loc opts fc stk sc else pure NoMatch _ => pure NoMatch -- Type of type matching, in typecase - tryAlt env loc opts fc stk (NType _ _) (ConCase (UN (Basic "Type")) tag [] sc) + tryAlt env loc opts fc stk (NType _ _) (ConCase (UN (Basic "Type")) tag SLNil sc) = evalTree env loc opts fc stk sc tryAlt env loc opts fc stk (NType _ _) (ConCase _ _ _ _) = pure NoMatch -- Arrow matching, in typecase tryAlt {more} - env loc opts fc stk (NBind pfc x (Pi fc' r e aty) scty) (ConCase (UN (Basic "->")) tag [s,t] sc) - = evalConAlt {more} env loc opts fc stk [s,t] - [aty, - MkNFClosure opts env (NBind pfc x (Lam fc' r e aty) scty)] + env loc opts fc stk (NBind pfc x (Pi fc' r e aty) scty) (ConCase (UN (Basic "->")) tag (s :%: t :%: SLNil) sc) + = evalConAlt {more} env loc opts fc stk (s :%: t :%: SLNil) + (aty :%: MkNFClosure opts env (NBind pfc x (Lam fc' r e aty) scty) :%: SLNil) sc tryAlt {more} env loc opts fc stk (NBind pfc x (Pi fc' r e aty) scty) (ConCase nm tag args sc) @@ -443,38 +450,38 @@ parameters (defs : Defs, topopts : EvalOpts) -- Take arguments from the stack, as long as there's enough. -- Returns the arguments, and the rest of the stack - takeFromStack : (arity : Nat) -> Stack free -> - Maybe (Vect arity (Closure free), Stack free) + takeFromStack : (arity : Nat) -> ScopedList (FC, Closure free) -> + Maybe (Vect arity (Closure free), ScopedList (FC, Closure free)) takeFromStack arity stk = takeStk arity stk [] where - takeStk : (remain : Nat) -> Stack free -> + takeStk : (remain : Nat) -> ScopedList (FC, Closure free) -> Vect got (Closure free) -> - Maybe (Vect (got + remain) (Closure free), Stack free) + Maybe (Vect (got + remain) (Closure free), ScopedList (FC, Closure free)) takeStk {got} Z stk acc = Just (rewrite plusZeroRightNeutral got in reverse acc, stk) - takeStk (S k) [] acc = Nothing - takeStk {got} (S k) (arg :: stk) acc + takeStk (S k) SLNil acc = Nothing + takeStk {got} (S k) (arg :%: stk) acc = rewrite sym (plusSuccRightSucc got k) in takeStk k stk (snd arg :: acc) - argsFromStack : (args : List Name) -> - Stack free -> - Maybe (LocalEnv free args, Stack free) - argsFromStack [] stk = Just ([], stk) - argsFromStack (n :: ns) [] = Nothing - argsFromStack (n :: ns) (arg :: args) + argsFromStack : (args : ScopedList Name) -> + ScopedList (FC, Closure free) -> + Maybe (LocalEnv free args, ScopedList (FC, Closure free)) + argsFromStack SLNil stk = Just ([], stk) + argsFromStack (n :%: ns) SLNil = Nothing + argsFromStack (n :%: ns) (arg :%: args) = do (loc', stk') <- argsFromStack ns args pure (snd arg :: loc', stk') evalOp : {auto c : Ref Ctxt Defs} -> {arity, free : _} -> (Vect arity (NF free) -> Maybe (NF free)) -> - Stack free -> (def : Lazy (NF free)) -> + ScopedList (FC, Closure free) -> (def : Lazy (NF free)) -> Core (NF free) evalOp {arity} fn stk def = case takeFromStack arity stk of -- Stack must be exactly the right height - Just (args, []) => + Just (args, SLNil) => do argsnf <- evalAll args pure $ case fn argsnf of Nothing => def @@ -491,7 +498,7 @@ parameters (defs : Defs, topopts : EvalOpts) Env Term free -> EvalOpts -> (isMeta : Bool) -> FC -> RigCount -> Def -> List DefFlag -> - Stack free -> (def : Lazy (NF free)) -> + ScopedList (FC, Closure free) -> (def : Lazy (NF free)) -> Core (NF free) evalDef env opts meta fc rigd (PMDef r args tree _ _) flags stk def -- If evaluating the definition fails (e.g. due to a case being @@ -514,18 +521,18 @@ parameters (defs : Defs, topopts : EvalOpts) pure "Cannot reduce under-applied \{show def}" pure def Just (locs', stk') => - do (Result (MkTermEnv newLoc res)) <- evalTree env locs' opts fc stk' tree + do (Result (MkTermEnv newLoc res)) <- evalTree env locs' opts fc (toList stk') tree | _ => do logC "eval.def.stuck" 50 $ do def <- toFullNames def pure "evalTree failed on \{show def}" pure def case fuel opts of - Nothing => evalWithOpts defs opts env newLoc res stk' + Nothing => evalWithOpts defs opts env newLoc res (toList stk') Just Z => log "eval.def.stuck" 50 "Recursion depth limit exceeded" >> pure def Just (S k) => do let opts' = { fuel := Just k } opts - evalWithOpts defs opts' env newLoc res stk' + evalWithOpts defs opts' env newLoc res (toList stk') else do -- logC "eval.def.stuck" 50 $ do -- def <- toFullNames def -- pure $ unlines [ "Refusing to reduce \{show def}:" diff --git a/src/Core/Normalise/Quote.idr b/src/Core/Normalise/Quote.idr index 798db02de5..3f028fbd95 100644 --- a/src/Core/Normalise/Quote.idr +++ b/src/Core/Normalise/Quote.idr @@ -5,6 +5,7 @@ import Core.Core import Core.Env import Core.Normalise.Eval import Core.TT +import Core.Name.ScopedList import Core.Value %default covering @@ -24,13 +25,13 @@ record QuoteOpts where public export interface Quote tm where quote : {auto c : Ref Ctxt Defs} -> - {vars : List Name} -> + {vars : ScopedList Name} -> Defs -> Env Term vars -> tm vars -> Core (Term vars) quoteLHS : {auto c : Ref Ctxt Defs} -> - {vars : List Name} -> + {vars : ScopedList Name} -> Defs -> Env Term vars -> tm vars -> Core (Term vars) quoteOpts : {auto c : Ref Ctxt Defs} -> - {vars : List Name} -> + {vars : ScopedList Name} -> QuoteOpts -> Defs -> Env Term vars -> tm vars -> Core (Term vars) quoteGen : {auto c : Ref Ctxt Defs} -> @@ -62,7 +63,7 @@ mutual {bound, free : _} -> Ref QVar Int -> QuoteOpts -> Defs -> Bounds bound -> Env Term free -> Closure free -> - Core (Term (bound ++ free)) + Core (Term (bound +%+ free)) quoteArg q opts defs bounds env a = quoteGenNF q opts defs bounds env !(evalClosure defs a) @@ -70,22 +71,22 @@ mutual {bound, free : _} -> Ref QVar Int -> QuoteOpts -> Defs -> Bounds bound -> Env Term free -> (FC, Closure free) -> - Core ((FC, Term (bound ++ free))) + Core ((FC, Term (bound +%+ free))) quoteArgWithFC q opts defs bounds env = traversePair (quoteArg q opts defs bounds env) quoteArgs : {auto c : Ref Ctxt Defs} -> {bound, free : _} -> Ref QVar Int -> QuoteOpts -> Defs -> Bounds bound -> - Env Term free -> List (Closure free) -> - Core (List (Term (bound ++ free))) + Env Term free -> ScopedList (Closure free) -> + Core (ScopedList (Term (bound +%+ free))) quoteArgs q opts defs bounds env = traverse (quoteArg q opts defs bounds env) quoteArgsWithFC : {auto c : Ref Ctxt Defs} -> {bound, free : _} -> Ref QVar Int -> QuoteOpts -> Defs -> Bounds bound -> - Env Term free -> List (FC, Closure free) -> - Core (List (FC, Term (bound ++ free))) + Env Term free -> ScopedList (FC, Closure free) -> + Core (ScopedList (FC, Term (bound +%+ free))) quoteArgsWithFC q opts defs bounds env = traverse (quoteArgWithFC q opts defs bounds env) @@ -93,16 +94,16 @@ mutual {bound, free : _} -> Ref QVar Int -> QuoteOpts -> Defs -> FC -> Bounds bound -> Env Term free -> NHead free -> - Core (Term (bound ++ free)) + Core (Term (bound +%+ free)) quoteHead {bound} q opts defs fc bounds env (NLocal mrig _ prf) = let MkVar prf' = addLater bound prf in pure $ Local fc mrig _ prf' where addLater : {idx : _} -> - (ys : List Name) -> (0 p : IsVar n idx xs) -> - Var (ys ++ xs) - addLater [] isv = MkVar isv - addLater (x :: xs) isv + (ys : ScopedList Name) -> (0 p : IsVar n idx xs) -> + Var (ys +%+ xs) + addLater SLNil isv = MkVar isv + addLater (x :%: xs) isv = let MkVar isv' = addLater xs isv in MkVar (Later isv') quoteHead q opts defs fc bounds env (NRef Bound (MN n i)) @@ -124,13 +125,13 @@ mutual quoteHead q opts defs fc bounds env (NRef nt n) = pure $ Ref fc nt n quoteHead q opts defs fc bounds env (NMeta n i args) = do args' <- quoteArgs q opts defs bounds env args - pure $ Meta fc n i args' + pure $ Meta fc n i (toList args') quotePi : {auto c : Ref Ctxt Defs} -> {bound, free : _} -> Ref QVar Int -> QuoteOpts -> Defs -> Bounds bound -> Env Term free -> PiInfo (Closure free) -> - Core (PiInfo (Term (bound ++ free))) + Core (PiInfo (Term (bound +%+ free))) quotePi q opts defs bounds env Explicit = pure Explicit quotePi q opts defs bounds env Implicit = pure Implicit quotePi q opts defs bounds env AutoImplicit = pure AutoImplicit @@ -142,7 +143,7 @@ mutual {bound, free : _} -> Ref QVar Int -> QuoteOpts -> Defs -> Bounds bound -> Env Term free -> Binder (Closure free) -> - Core (Binder (Term (bound ++ free))) + Core (Binder (Term (bound +%+ free))) quoteBinder q opts defs bounds env (Lam fc r p ty) = do ty' <- quoteGenNF q opts defs bounds env !(evalClosure defs ty) p' <- quotePi q opts defs bounds env p @@ -171,7 +172,7 @@ mutual {bound, vars : _} -> Ref QVar Int -> QuoteOpts -> Defs -> Bounds bound -> - Env Term vars -> NF vars -> Core (Term (bound ++ vars)) + Env Term vars -> NF vars -> Core (Term (bound +%+ vars)) quoteGenNF q opts defs bound env (NBind fc n b sc) = do var <- genName "qv" sc' <- quoteGenNF q opts defs (Add n var bound) env @@ -248,7 +249,7 @@ Quote Closure where quoteWithPiGen : {auto _ : Ref Ctxt Defs} -> {bound, vars : _} -> Ref QVar Int -> QuoteOpts -> Defs -> Bounds bound -> - Env Term vars -> NF vars -> Core (Term (bound ++ vars)) + Env Term vars -> NF vars -> Core (Term (bound +%+ vars)) quoteWithPiGen q opts defs bound env (NBind fc n (Pi bfc c p ty) sc) = do var <- genName "qv" empty <- clearDefs defs @@ -267,7 +268,7 @@ quoteWithPiGen q opts defs bound env tm -- are, don't reduce anything else export quoteWithPi : {auto c : Ref Ctxt Defs} -> - {vars : List Name} -> + {vars : ScopedList Name} -> Defs -> Env Term vars -> NF vars -> Core (Term vars) quoteWithPi defs env tm = do q <- newRef QVar 0 diff --git a/src/Core/Ord.idr b/src/Core/Ord.idr index e3476cff49..6ba2bed53d 100644 --- a/src/Core/Ord.idr +++ b/src/Core/Ord.idr @@ -3,6 +3,8 @@ module Core.Ord import Core.CompileExpr import Core.Name +import Core.Name.ScopedList +import Core.Name.ScopedList.Name import Core.TT import Data.Vect diff --git a/src/Core/Primitives.idr b/src/Core/Primitives.idr index fedb4120db..838f55732a 100644 --- a/src/Core/Primitives.idr +++ b/src/Core/Primitives.idr @@ -2,6 +2,7 @@ module Core.Primitives import Core.Context import Core.TT +import Core.Name.ScopedList import Core.Value import Libraries.Utils.String @@ -532,7 +533,7 @@ doubleTy : ClosedTerm doubleTy = predTy DoubleType DoubleType pi : (x : String) -> RigCount -> PiInfo (Term xs) -> Term xs -> - Term (UN (Basic x) :: xs) -> Term xs + Term (UN (Basic x) :%: xs) -> Term xs pi x rig plic ty sc = Bind emptyFC (UN (Basic x)) (Pi emptyFC rig plic ty) sc believeMeTy : ClosedTerm @@ -566,7 +567,7 @@ castTo WorldType = const Nothing export getOp : {0 arity : Nat} -> PrimFn arity -> - {vars : List Name} -> Vect arity (NF vars) -> Maybe (NF vars) + {vars : ScopedList Name} -> Vect arity (NF vars) -> Maybe (NF vars) getOp (Add ty) = binOp add getOp (Sub ty) = binOp sub getOp (Mul ty) = binOp mul diff --git a/src/Core/Reflect.idr b/src/Core/Reflect.idr index c464c454f9..eadf7f8c9f 100644 --- a/src/Core/Reflect.idr +++ b/src/Core/Reflect.idr @@ -7,6 +7,7 @@ import Core.Env import Core.Normalise import Core.TT import Core.Value +import Core.Name.ScopedList import Libraries.Data.WithDefault @@ -224,7 +225,7 @@ Reify Nat where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of (UN (Basic "Z"), _) => pure Z - (UN (Basic "S"), [(_, k)]) + (UN (Basic "S"), ((_, k) :%: SLNil)) => do k' <- reify defs !(evalClosure defs k) pure (S k') _ => cantReify val "Nat" @@ -242,13 +243,25 @@ Reify a => Reify (List a) where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of (UN (Basic "Nil"), _) => pure [] - (UN (Basic "::"), [_, (_, x), (_, xs)]) + (UN (Basic "::"), (_ :%: (_, x) :%: (_, xs) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) xs' <- reify defs !(evalClosure defs xs) pure (x' :: xs') _ => cantReify val "List" reify defs val = cantReify val "List" +export +Reify a => Reify (ScopedList a) where + reify defs val@(NDCon _ n _ _ args) + = case (dropAllNS !(full (gamma defs) n), args) of + (UN (Basic "SLNil"), _) => pure SLNil + (UN (Basic ":%:"), (_ :%: (_, x) :%: (_, xs) :%: SLNil)) + => do x' <- reify defs !(evalClosure defs x) + xs' <- reify defs !(evalClosure defs xs) + pure (x' :%: xs') + _ => cantReify val "ScopedList" + reify defs val = cantReify val "ScopedList" + export Reflect a => Reflect (List a) where reflect fc defs lhs env [] = appCon fc defs (basics "Nil") [Erased fc Placeholder] @@ -257,9 +270,17 @@ Reflect a => Reflect (List a) where xs' <- reflect fc defs lhs env xs appCon fc defs (basics "::") [Erased fc Placeholder, x', xs'] +export +Reflect a => Reflect (ScopedList a) where + reflect fc defs lhs env SLNil = appCon fc defs (basics "SLNil") [Erased fc Placeholder] + reflect fc defs lhs env (x :%: xs) + = do x' <- reflect fc defs lhs env x + xs' <- reflect fc defs lhs env xs + appCon fc defs (basics ":%:") [Erased fc Placeholder, x', xs'] + export Reify a => Reify (List1 a) where - reify defs val@(NDCon _ n _ _ [_, (_, x), (_, xs)]) + reify defs val@(NDCon _ n _ _ (_ :%: (_, x) :%: (_, xs) :%: SLNil)) = case dropAllNS !(full (gamma defs) n) of UN (Basic ":::") => do x' <- reify defs !(evalClosure defs x) @@ -281,7 +302,7 @@ Reify a => Reify (SnocList a) where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of (UN (Basic "Lin"), _) => pure [<] - (UN (Basic ":<"), [_, (_, sx), (_, x)]) + (UN (Basic ":<"), (_ :%: (_, sx) :%: (_, x) :%: SLNil)) => do sx' <- reify defs !(evalClosure defs sx) x' <- reify defs !(evalClosure defs x) pure (sx' :< x') @@ -301,7 +322,7 @@ Reify a => Reify (Maybe a) where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of (UN (Basic "Nothing"), _) => pure Nothing - (UN (Basic "Just"), [_, (_, x)]) + (UN (Basic "Just"), (_ :%: (_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (Just x') _ => cantReify val "Maybe" @@ -320,7 +341,7 @@ Reify a => Reify (WithDefault a def) where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of (UN (Basic "DefaultedValue"), _) => pure defaulted - (UN (Basic "SpecifiedValue"), [_, _, (_, x)]) + (UN (Basic "SpecifiedValue"), (_ :%: _ :%: (_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (specified x') _ => cantReify val "WithDefault" @@ -337,7 +358,7 @@ Reflect a => Reflect (WithDefault a def) where export (Reify a, Reify b) => Reify (a, b) where - reify defs val@(NDCon _ n _ _ [_, _, (_, x), (_, y)]) + reify defs val@(NDCon _ n _ _ (_ :%: _ :%: (_, x) :%: (_, y) :%: SLNil)) = case dropAllNS !(full (gamma defs) n) of UN (Basic "MkPair") => do x' <- reify defs !(evalClosure defs x) @@ -355,7 +376,7 @@ export export Reify Namespace where - reify defs val@(NDCon _ n _ _ [(_, ns)]) + reify defs val@(NDCon _ n _ _ ((_, ns) :%: SLNil)) = case dropAllNS !(full (gamma defs) n) of UN (Basic "MkNS") => do ns' <- reify defs !(evalClosure defs ns) @@ -371,7 +392,7 @@ Reflect Namespace where export Reify ModuleIdent where - reify defs val@(NDCon _ n _ _ [(_, ns)]) + reify defs val@(NDCon _ n _ _ ((_, ns) :%: SLNil)) = case dropAllNS !(full (gamma defs) n) of UN (Basic "MkMI") => do ns' <- reify defs !(evalClosure defs ns) @@ -389,13 +410,13 @@ export Reify UserName where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of - (UN (Basic "Basic"), [(_, str)]) + (UN (Basic "Basic"), ((_, str) :%: SLNil)) => do str' <- reify defs !(evalClosure defs str) pure (Basic str') - (UN (Basic "Field"), [(_, str)]) + (UN (Basic "Field"), ((_, str) :%: SLNil)) => do str' <- reify defs !(evalClosure defs str) pure (Field str') - (UN (Basic "Underscore"), []) + (UN (Basic "Underscore"), SLNil) => pure Underscore (NS _ (UN _), _) => cantReify val "Name, reifying it is unimplemented or intentionally internal" @@ -417,30 +438,30 @@ export Reify Name where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of - (UN (Basic "UN"), [(_, str)]) + (UN (Basic "UN"), ((_, str) :%: SLNil)) => do str' <- reify defs !(evalClosure defs str) pure (UN str') - (UN (Basic "MN"), [(_, str), (_, i)]) + (UN (Basic "MN"), ((_, str) :%: (_, i) :%: SLNil)) => do str' <- reify defs !(evalClosure defs str) i' <- reify defs !(evalClosure defs i) pure (MN str' i') - (UN (Basic "NS"), [(_, ns), (_, n)]) + (UN (Basic "NS"), ((_, ns) :%: (_, n) :%: SLNil)) => do ns' <- reify defs !(evalClosure defs ns) n' <- reify defs !(evalClosure defs n) pure (NS ns' n') - (UN (Basic "DN"), [(_, str), (_, n)]) + (UN (Basic "DN"), ((_, str) :%: (_, n) :%: SLNil)) => do str' <- reify defs !(evalClosure defs str) n' <- reify defs !(evalClosure defs n) pure (DN str' n') - (UN (Basic "Nested"), [(_, ix), (_, n)]) + (UN (Basic "Nested"), ((_, ix) :%: (_, n) :%: SLNil)) => do ix' <- reify defs !(evalClosure defs ix) n' <- reify defs !(evalClosure defs n) pure (Nested ix' n') - (UN (Basic "CaseBlock"), [(_, outer), (_, i)]) + (UN (Basic "CaseBlock"), ((_, outer) :%: (_, i) :%: SLNil)) => do outer' <- reify defs !(evalClosure defs outer) i' <- reify defs !(evalClosure defs i) pure (CaseBlock outer' i') - (UN (Basic "WithBlock"), [(_, outer), (_, i)]) + (UN (Basic "WithBlock"), ((_, outer) :%: (_, i) :%: SLNil)) => do outer' <- reify defs !(evalClosure defs outer) i' <- reify defs !(evalClosure defs i) pure (WithBlock outer' i') @@ -492,11 +513,11 @@ Reify NameType where = case (dropAllNS !(full (gamma defs) n), args) of (UN (Basic "Bound"), _) => pure Bound (UN (Basic "Func"), _) => pure Func - (UN (Basic "DataCon"), [(_, t), (_, i)]) + (UN (Basic "DataCon"), ((_, t) :%: (_, i) :%: SLNil)) => do t' <- reify defs !(evalClosure defs t) i' <- reify defs !(evalClosure defs i) pure (DataCon t' i') - (UN (Basic "TyCon"), [(_, t),(_, i)]) + (UN (Basic "TyCon"), ((_, t) :%: (_, i) :%: SLNil)) => do t' <- reify defs !(evalClosure defs t) i' <- reify defs !(evalClosure defs i) pure (TyCon t' i') @@ -520,33 +541,33 @@ export Reify PrimType where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of - (UN (Basic "IntType"), []) + (UN (Basic "IntType"), SLNil) => pure IntType - (UN (Basic "Int8Type"), []) + (UN (Basic "Int8Type"), SLNil) => pure Int8Type - (UN (Basic "Int16Type"), []) + (UN (Basic "Int16Type"), SLNil) => pure Int16Type - (UN (Basic "Int32Type"), []) + (UN (Basic "Int32Type"), SLNil) => pure Int32Type - (UN (Basic "Int64Type"), []) + (UN (Basic "Int64Type"), SLNil) => pure Int64Type - (UN (Basic "IntegerType"), []) + (UN (Basic "IntegerType"), SLNil) => pure IntegerType - (UN (Basic "Bits8Type"), []) + (UN (Basic "Bits8Type"), SLNil) => pure Bits8Type - (UN (Basic "Bits16Type"), []) + (UN (Basic "Bits16Type"), SLNil) => pure Bits16Type - (UN (Basic "Bits32Type"), []) + (UN (Basic "Bits32Type"), SLNil) => pure Bits32Type - (UN (Basic "Bits64Type"), []) + (UN (Basic "Bits64Type"), SLNil) => pure Bits64Type - (UN (Basic "StringType"), []) + (UN (Basic "StringType"), SLNil) => pure StringType - (UN (Basic "CharType"), []) + (UN (Basic "CharType"), SLNil) => pure CharType - (UN (Basic "DoubleType"), []) + (UN (Basic "DoubleType"), SLNil) => pure DoubleType - (UN (Basic "WorldType"), []) + (UN (Basic "WorldType"), SLNil) => pure WorldType _ => cantReify val "PrimType" reify defs val = cantReify val "PrimType" @@ -555,49 +576,49 @@ export Reify Constant where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of - (UN (Basic "I"), [(_, x)]) + (UN (Basic "I"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (I x') - (UN (Basic "I8"), [(_, x)]) + (UN (Basic "I8"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (I8 x') - (UN (Basic "I16"), [(_, x)]) + (UN (Basic "I16"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (I16 x') - (UN (Basic "I32"), [(_, x)]) + (UN (Basic "I32"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (I32 x') - (UN (Basic "I64"), [(_, x)]) + (UN (Basic "I64"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (I64 x') - (UN (Basic "BI"), [(_, x)]) + (UN (Basic "BI"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (BI x') - (UN (Basic "B8"), [(_, x)]) + (UN (Basic "B8"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (B8 x') - (UN (Basic "B16"), [(_, x)]) + (UN (Basic "B16"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (B16 x') - (UN (Basic "B32"), [(_, x)]) + (UN (Basic "B32"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (B32 x') - (UN (Basic "B64"), [(_, x)]) + (UN (Basic "B64"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (B64 x') - (UN (Basic "Str"), [(_, x)]) + (UN (Basic "Str"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (Str x') - (UN (Basic "Ch"), [(_, x)]) + (UN (Basic "Ch"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (Ch x') - (UN (Basic "Db"), [(_, x)]) + (UN (Basic "Db"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (Db x') - (UN (Basic "PrT"), [(_, x)]) + (UN (Basic "PrT"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (PrT x') - (UN (Basic "WorldVal"), []) + (UN (Basic "WorldVal"), SLNil) => pure WorldVal _ => cantReify val "Constant" reify defs val = cantReify val "Constant" @@ -737,7 +758,7 @@ Reify t => Reify (PiInfo t) where (UN (Basic "ImplicitArg"), _) => pure Implicit (UN (Basic "ExplicitArg"), _) => pure Explicit (UN (Basic "AutoImplicit"), _) => pure AutoImplicit - (UN (Basic "DefImplicit"), [_, (_, t)]) + (UN (Basic "DefImplicit"), (_ :%: (_, t) :%: SLNil)) => do t' <- reify defs !(evalClosure defs t) pure (DefImplicit t') _ => cantReify val "PiInfo" @@ -775,7 +796,7 @@ export Reify VirtualIdent where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of - (UN (Basic "Interactive"), []) + (UN (Basic "Interactive"), SLNil) => pure Interactive _ => cantReify val "VirtualIdent" reify defs val = cantReify val "VirtualIdent" @@ -793,11 +814,11 @@ export Reify BuiltinType where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of - (UN (Basic "BuiltinNatural"), []) + (UN (Basic "BuiltinNatural"), SLNil) => pure BuiltinNatural - (UN (Basic "NaturalToInteger"), []) + (UN (Basic "NaturalToInteger"), SLNil) => pure NaturalToInteger - (UN (Basic "IntegerToNatural"), []) + (UN (Basic "IntegerToNatural"), SLNil) => pure IntegerToNatural _ => cantReify val "BuiltinType" reify defs val = cantReify val "BuiltinType" @@ -811,13 +832,13 @@ export Reify OriginDesc where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of - (UN (Basic "PhysicalIdrSrc"), [(_, ident)]) + (UN (Basic "PhysicalIdrSrc"), ((_, ident) :%: SLNil)) => do ident' <- reify defs !(evalClosure defs ident) pure (PhysicalIdrSrc ident') - (UN (Basic "PhysicalPkgSrc"), [(_, fname)]) + (UN (Basic "PhysicalPkgSrc"), ((_, fname) :%: SLNil)) => do fname' <- reify defs !(evalClosure defs fname) pure (PhysicalPkgSrc fname') - (UN (Basic "Virtual"), [(_, ident)]) + (UN (Basic "Virtual"), ((_, ident) :%: SLNil)) => do ident' <- reify defs !(evalClosure defs ident) pure (Virtual ident') _ => cantReify val "OriginDesc" @@ -839,7 +860,7 @@ export Reify FC where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of - (UN (Basic "MkFC"), [(_, fn), (_, start), (_, end)]) + (UN (Basic "MkFC"), ((_, fn) :%: (_, start) :%: (_, end) :%: SLNil)) => do fn' <- reify defs !(evalClosure defs fn) start' <- reify defs !(evalClosure defs start) end' <- reify defs !(evalClosure defs end) diff --git a/src/Core/SchemeEval/Compile.idr b/src/Core/SchemeEval/Compile.idr index 2f98457acd..053326ba4c 100644 --- a/src/Core/SchemeEval/Compile.idr +++ b/src/Core/SchemeEval/Compile.idr @@ -78,9 +78,9 @@ getName (Bound x) = x getName (Free x) = x public export -data SchVars : List Name -> Type where - Nil : SchVars [] - (::) : SVar -> SchVars ns -> SchVars (n :: ns) +data SchVars : ScopedList Name -> Type where + Nil : SchVars SLNil + (::) : SVar -> SchVars ns -> SchVars (n :%: ns) Show (SchVars ns) where show xs = show (toList xs) @@ -300,10 +300,10 @@ getArgName pure (MN "carg" (cast i)) extend : Ref Sym Integer => - (args : List Name) -> SchVars vars -> - Core (List Name, SchVars (args ++ vars)) -extend [] svs = pure ([], svs) -extend (arg :: args) svs + (args : ScopedList Name) -> SchVars vars -> + Core (List Name, SchVars (args +%+ vars)) +extend SLNil svs = pure ([], svs) +extend (arg :%: args) svs = do n <- getArgName (args', svs') <- extend args svs pure (n :: args', Bound (schVarName n) :: svs') @@ -361,7 +361,7 @@ compileCase blk svs (Case idx p scTy xs) (Apply (Var "vector-ref") [Var var, IntegerVal (cast i)]) (project (i + 1) var ns body) - bindArgs : String -> (args : List Name) -> CaseTree (args ++ vars) -> + bindArgs : String -> (args : ScopedList Name) -> CaseTree (args +%+ vars) -> Core (SchemeObj Write) bindArgs var args sc = do (bind, svs') <- extend args svs @@ -396,7 +396,7 @@ compileCase blk svs (Case idx p scTy xs) (Apply (Var "vector-ref") [Var var, IntegerVal (cast i)]) (project (i + 1) var ns body) - bindArgs : String -> (args : List Name) -> CaseTree (args ++ vars) -> + bindArgs : String -> (args : ScopedList Name) -> CaseTree (args +%+ vars) -> Core (SchemeObj Write) bindArgs var args sc = do (bind, svs') <- extend args svs @@ -404,7 +404,7 @@ compileCase blk svs (Case idx p scTy xs) makeAlt : String -> CaseAlt vars -> Core (Maybe (SchemeObj Write, SchemeObj Write)) - makeAlt var (ConCase (UN (Basic "->")) t [_, _] sc) + makeAlt var (ConCase (UN (Basic "->")) t (_ :%: _ :%: SLNil) sc) = pure Nothing -- do this in 'addPiMatch' below, since the -- representation is different makeAlt var (ConCase n t args sc) @@ -417,7 +417,7 @@ compileCase blk svs (Case idx p scTy xs) -- t is a function type, and conveniently the scope of a pi -- binding is represented as a function. Lucky us! So we just need -- to extract it then evaluate the scope - addPiMatch var (ConCase (UN (Basic "->")) _ [s, t] sc :: _) def + addPiMatch var (ConCase (UN (Basic "->")) _ (s :%: t :%: SLNil) sc :: _) def = do sn <- getArgName tn <- getArgName let svs' = Bound (schVarName sn) :: @@ -477,9 +477,9 @@ varObjs : SchVars ns -> List (SchemeObj Write) varObjs [] = [] varObjs (x :: xs) = Var (show x) :: varObjs xs -mkArgs : (ns : List Name) -> Core (SchVars ns) -mkArgs [] = pure [] -mkArgs (x :: xs) +mkArgs : (ns : ScopedList Name) -> Core (SchVars ns) +mkArgs SLNil = pure [] +mkArgs (x :%: xs) = pure $ Bound (schVarName x) :: !(mkArgs xs) bindArgs : Name -> @@ -520,12 +520,12 @@ compileBody _ n (DCon tag arity newtypeArg) let body = Vector (cast tag) (toScheme n :: toScheme emptyFC :: - map (Var . schVarName) args) + map (Var . schVarName) (toList args)) pure (bindArgs n argvs [] body) where - mkArgNs : Int -> Nat -> List Name - mkArgNs i Z = [] - mkArgNs i (S k) = MN "arg" i :: mkArgNs (i+1) k + mkArgNs : Int -> Nat -> ScopedList Name + mkArgNs i Z = SLNil + mkArgNs i (S k) = MN "arg" i :%: mkArgNs (i+1) k compileBody _ n (TCon tag Z parampos detpos flags mutwith datacons detagabbleBy) = pure $ Vector (-1) [IntegerVal (cast tag), StringVal (show n), toScheme n, toScheme emptyFC] @@ -537,12 +537,12 @@ compileBody _ n (TCon tag arity parampos detpos flags mutwith datacons detagabbl (IntegerVal (cast tag) :: StringVal (show n) :: toScheme n :: toScheme emptyFC :: - map (Var . schVarName) args) + map (Var . schVarName) (toList args)) pure (bindArgs n argvs [] body) where - mkArgNs : Int -> Nat -> List Name - mkArgNs i Z = [] - mkArgNs i (S k) = MN "arg" i :: mkArgNs (i+1) k + mkArgNs : Int -> Nat -> ScopedList Name + mkArgNs i Z = SLNil + mkArgNs i (S k) = MN "arg" i :%: mkArgNs (i+1) k compileBody _ n (Hole numlocs x) = pure $ blockedMetaApp n compileBody _ n (BySearch x maxdepth defining) = pure $ blockedMetaApp n compileBody _ n (Guess guess envbind constraints) = pure $ blockedMetaApp n diff --git a/src/Core/SchemeEval/Evaluate.idr b/src/Core/SchemeEval/Evaluate.idr index 0afdff6d62..cb7d266eaf 100644 --- a/src/Core/SchemeEval/Evaluate.idr +++ b/src/Core/SchemeEval/Evaluate.idr @@ -12,7 +12,7 @@ import Libraries.Data.NameMap import Libraries.Utils.Scheme public export -data SObj : List Name -> Type where +data SObj : ScopedList Name -> Type where MkSObj : ForeignObj -> SchVars vars -> SObj vars -- Values, which we read off evaluated scheme objects. @@ -22,13 +22,13 @@ data SObj : List Name -> Type where -- recording a LocalEnv for example). mutual public export - data SHead : List Name -> Type where + data SHead : ScopedList Name -> Type where SLocal : (idx : Nat) -> (0 p : IsVar nm idx vars) -> SHead vars SRef : NameType -> Name -> SHead vars SMeta : Name -> Int -> List (Core (SNF vars)) -> SHead vars public export - data SNF : List Name -> Type where + data SNF : ScopedList Name -> Type where SBind : FC -> (x : Name) -> Binder (SNF vars) -> (SObj vars -> Core (SNF vars)) -> SNF vars SApp : FC -> SHead vars -> List (Core (SNF vars)) -> SNF vars @@ -147,9 +147,9 @@ mutual -- Instead, decode the ForeignObj directly, which is uglier but faster. quoteVector : Ref Sym Integer => Ref Ctxt Defs => - SchVars (outer ++ vars) -> + SchVars (outer +%+ vars) -> Integer -> List ForeignObj -> - Core (Term (outer ++ vars)) + Core (Term (outer +%+ vars)) quoteVector svs (-2) [_, fname_in, args_in] -- Blocked app = quoteOrInvalid fname_in $ \ fname => do let argList = getArgList args_in @@ -275,9 +275,9 @@ mutual quotePiInfo : Ref Sym Integer => Ref Ctxt Defs => - SchVars (outer ++ vars) -> + SchVars (outer +%+ vars) -> ForeignObj -> - Core (PiInfo (Term (outer ++ vars))) + Core (PiInfo (Term (outer +%+ vars))) quotePiInfo svs obj = if isInteger obj then case unsafeGetInteger obj of @@ -305,49 +305,49 @@ mutual quoteBinder : Ref Sym Integer => Ref Ctxt Defs => - SchVars (outer ++ vars) -> + SchVars (outer +%+ vars) -> (forall ty . FC -> RigCount -> PiInfo ty -> ty -> Binder ty) -> ForeignObj -> -- body of binder, represented as a function RigCount -> - PiInfo (Term (outer ++ vars)) -> - Term (outer ++ vars) -> -- decoded type + PiInfo (Term (outer +%+ vars)) -> + Term (outer +%+ vars) -> -- decoded type Name -> -- bound name - Core (Term (outer ++ vars)) + Core (Term (outer +%+ vars)) quoteBinder svs binder proc_in r pi ty name = do let Procedure proc = decodeObj proc_in | _ => invalid i <- nextName let n = show name ++ "-" ++ show i let sc = unsafeApply proc (makeSymbol n) - sc' <- quote' {outer = name :: outer} (Bound n :: svs) sc + sc' <- quote' {outer = name :%: outer} (Bound n :: svs) sc pure (Bind emptyFC name (binder emptyFC r pi ty) sc') quotePLet : Ref Sym Integer => Ref Ctxt Defs => - SchVars (outer ++ vars) -> + SchVars (outer +%+ vars) -> ForeignObj -> -- body of binder, represented as a function RigCount -> - Term (outer ++ vars) -> -- decoded type - Term (outer ++ vars) -> -- decoded value + Term (outer +%+ vars) -> -- decoded type + Term (outer +%+ vars) -> -- decoded value Name -> -- bound name - Core (Term (outer ++ vars)) + Core (Term (outer +%+ vars)) quotePLet svs proc_in r val ty name = do let Procedure proc = decodeObj proc_in | _ => invalid i <- nextName let n = show name ++ "-" ++ show i let sc = unsafeApply proc (makeSymbol n) - sc' <- quote' {outer = name :: outer} (Bound n :: svs) sc + sc' <- quote' {outer = name :%: outer} (Bound n :: svs) sc pure (Bind emptyFC name (PLet emptyFC r val ty) sc') quote' : Ref Sym Integer => Ref Ctxt Defs => - SchVars (outer ++ vars) -> ForeignObj -> - Core (Term (outer ++ vars)) + SchVars (outer +%+ vars) -> ForeignObj -> + Core (Term (outer +%+ vars)) quote' svs obj = if isVector obj then quoteVector svs (unsafeGetInteger (unsafeVectorRef obj 0)) @@ -386,7 +386,7 @@ quoteObj : {auto c : Ref Ctxt Defs} -> SObj vars -> Core (Term vars) quoteObj (MkSObj val schEnv) = do i <- newRef Sym 0 - quote' {outer = []} schEnv val + quote' {outer = SLNil} schEnv val mutual snfVector : Ref Ctxt Defs => diff --git a/src/Core/SchemeEval/Quote.idr b/src/Core/SchemeEval/Quote.idr index a84bc57d3d..ec85f94e90 100644 --- a/src/Core/SchemeEval/Quote.idr +++ b/src/Core/SchemeEval/Quote.idr @@ -12,7 +12,7 @@ mutual {bound, free : _} -> Ref Sym Integer -> Bounds bound -> Env Term free -> List (Core (SNF free)) -> - Core (List (Term (bound ++ free))) + Core (List (Term (bound +%+ free))) quoteArgs q bound env args = traverse (\arg => do arg' <- arg quoteGen q bound env arg') args @@ -21,7 +21,7 @@ mutual {bound, free : _} -> Ref Sym Integer -> Bounds bound -> Env Term free -> PiInfo (SNF free) -> - Core (PiInfo (Term (bound ++ free))) + Core (PiInfo (Term (bound +%+ free))) quotePi q bound env Explicit = pure Explicit quotePi q bound env Implicit = pure Implicit quotePi q bound env AutoImplicit = pure AutoImplicit @@ -33,7 +33,7 @@ mutual {bound, free : _} -> Ref Sym Integer -> Bounds bound -> Env Term free -> Binder (SNF free) -> - Core (Binder (Term (bound ++ free))) + Core (Binder (Term (bound +%+ free))) quoteBinder q bound env (Lam fc r p ty) = do ty' <- quoteGen q bound env ty p' <- quotePi q bound env p @@ -62,16 +62,16 @@ mutual {bound, free : _} -> Ref Sym Integer -> FC -> Bounds bound -> Env Term free -> SHead free -> - Core (Term (bound ++ free)) + Core (Term (bound +%+ free)) quoteHead {bound} q fc bounds env (SLocal idx prf) = let MkVar prf' = addLater bound prf in pure (Local fc Nothing _ prf') where addLater : {idx : _} -> - (ys : List Name) -> (0 p : IsVar n idx xs) -> - Var (ys ++ xs) - addLater [] isv = MkVar isv - addLater (x :: xs) isv + (ys : ScopedList Name) -> (0 p : IsVar n idx xs) -> + Var (ys +%+ xs) + addLater SLNil isv = MkVar isv + addLater (x :%: xs) isv = let MkVar isv' = addLater xs isv in MkVar (Later isv') quoteHead q fc bounds env (SRef nt n) @@ -93,7 +93,7 @@ mutual quoteGen : {auto c : Ref Ctxt Defs} -> {bound, vars : _} -> Ref Sym Integer -> Bounds bound -> - Env Term vars -> SNF vars -> Core (Term (bound ++ vars)) + Env Term vars -> SNF vars -> Core (Term (bound +%+ vars)) quoteGen q bound env (SBind fc n b sc) = do i <- nextName let var = UN (Basic ("b-" ++ show (fromInteger i))) diff --git a/src/Core/TT.idr b/src/Core/TT.idr index efa8521ac1..0276e4ae6f 100644 --- a/src/Core/TT.idr +++ b/src/Core/TT.idr @@ -3,6 +3,7 @@ module Core.TT import public Core.FC import public Core.Name import public Core.Name.Scoped +import public Core.Name.ScopedList import Idris.Pretty.Annotations @@ -319,7 +320,7 @@ Pretty Void Terminating where public export data Covering = IsCovering - | MissingCases (List (Term [])) + | MissingCases (List (Term SLNil)) | NonCoveringCall (List Name) export @@ -387,21 +388,21 @@ namespace Bounds Add : (x : Name) -> Name -> Bounds xs -> Bounds (x :: xs) export - sizeOf : Bounds xs -> SizeOf xs + sizeOf : Bounds xs -> Libraries.Data.List.SizeOf.SizeOf xs sizeOf None = zero sizeOf (Add _ _ b) = suc (sizeOf b) export addVars : SizeOf outer -> Bounds bound -> - NVar name (outer ++ vars) -> - NVar name (outer ++ (bound ++ vars)) + NVar name (outer +%+ vars) -> + NVar name (outer +%+ (bound +%+ vars)) addVars p = insertNVarNames p . sizeOf export resolveRef : SizeOf outer -> SizeOf done -> Bounds bound -> FC -> Name -> - Maybe (Var (outer ++ (done <>> bound ++ vars))) + Maybe (Var (outer +%+ (done <>> (bound +%+ vars)))) resolveRef _ _ None _ _ = Nothing resolveRef {outer} {vars} {done} p q (Add {xs} new old bs) fc n = if n == old @@ -409,7 +410,7 @@ resolveRef {outer} {vars} {done} p q (Add {xs} new old bs) fc n else resolveRef p (q :< new) bs fc n mkLocals : SizeOf outer -> Bounds bound -> - Term (outer ++ vars) -> Term (outer ++ (bound ++ vars)) + Term (outer +%+ vars) -> Term (outer +%+ (bound +%+ vars)) mkLocals outer bs (Local fc r idx p) = let MkNVar p' = addVars outer bs (MkNVar p) in Local fc r _ p' mkLocals outer bs (Ref fc Bound name) @@ -442,13 +443,13 @@ mkLocals outer bs (Erased fc (Dotted t)) = Erased fc (Dotted (mkLocals outer bs mkLocals outer bs (TType fc u) = TType fc u export -refsToLocals : Bounds bound -> Term vars -> Term (bound ++ vars) +refsToLocals : Bounds bound -> Term vars -> Term (bound +%+ vars) refsToLocals None y = y refsToLocals bs y = mkLocals zero bs y -- Replace any reference to 'x' with a locally bound name 'new' export -refToLocal : (x : Name) -> (new : Name) -> Term vars -> Term (new :: vars) +refToLocal : (x : Name) -> (new : Name) -> Term vars -> Term (new :%: vars) refToLocal x new tm = refsToLocals (Add new x None) tm -- Replace an explicit name with a term diff --git a/src/Core/TT/Subst.idr b/src/Core/TT/Subst.idr index abe5031188..437703af89 100644 --- a/src/Core/TT/Subst.idr +++ b/src/Core/TT/Subst.idr @@ -3,13 +3,14 @@ module Core.TT.Subst import Core.Name import Core.Name.Scoped import Core.TT.Var +import Core.Name.ScopedList %default total public export data Subst : Scoped -> Scope -> Scoped where - Nil : Subst tm [] vars - (::) : tm vars -> Subst tm ds vars -> Subst tm (d :: ds) vars + Nil : Subst tm SLNil vars + (::) : tm vars -> Subst tm ds vars -> Subst tm (d :%: ds) vars namespace Var @@ -23,7 +24,7 @@ export findDrop : (Var vars -> tm vars) -> SizeOf dropped -> - Var (dropped ++ vars) -> + Var (dropped +%+ vars) -> Subst tm dropped vars -> tm vars findDrop k s var sub = case locateVar s var of @@ -34,9 +35,9 @@ export find : Weaken tm => (forall vars. Var vars -> tm vars) -> SizeOf outer -> SizeOf dropped -> - Var (outer ++ (dropped ++ vars)) -> + Var (outer +%+ (dropped +%+ vars)) -> Subst tm dropped vars -> - tm (outer ++ vars) + tm (outer +%+ vars) find k outer dropped var sub = case locateVar outer var of Left var => k (embed var) Right var => weakenNs outer (findDrop k dropped var sub) @@ -48,5 +49,5 @@ Substitutable val tm SizeOf outer -> SizeOf dropped -> Subst val dropped vars -> - tm (outer ++ (dropped ++ vars)) -> - tm (outer ++ vars) + tm (outer +%+ (dropped +%+ vars)) -> + tm (outer +%+ vars) diff --git a/src/Core/TT/Term.idr b/src/Core/TT/Term.idr index 32f812a0f3..01762fa1be 100644 --- a/src/Core/TT/Term.idr +++ b/src/Core/TT/Term.idr @@ -9,6 +9,7 @@ import Core.Name.Scoped import Core.TT.Binder import Core.TT.Primitive import Core.TT.Var +import Core.Name.ScopedList import Data.List @@ -101,7 +102,7 @@ data Term : Scoped where Meta : FC -> Name -> Int -> List (Term vars) -> Term vars Bind : FC -> (x : Name) -> (b : Binder (Term vars)) -> - (scope : Term (x :: vars)) -> Term vars + (scope : Term (x :%: vars)) -> Term vars App : FC -> (fn : Term vars) -> (arg : Term vars) -> Term vars -- as patterns; since we check LHS patterns as terms before turning -- them into patterns, this helps us get it right. When normalising, @@ -125,7 +126,7 @@ data Term : Scoped where public export ClosedTerm : Type -ClosedTerm = Term [] +ClosedTerm = Term SLNil ------------------------------------------------------------------------ -- Weakening @@ -295,9 +296,9 @@ applySpineWithFC fn (args :< (fc, arg)) = App fc (applySpineWithFC fn args) arg -- Creates a chain of `App` nodes, each with its own file context export -applyStackWithFC : Term vars -> List (FC, Term vars) -> Term vars -applyStackWithFC fn [] = fn -applyStackWithFC fn ((fc, arg) :: args) = applyStackWithFC (App fc fn arg) args +applyStackWithFC : Term vars -> ScopedList (FC, Term vars) -> Term vars +applyStackWithFC fn SLNil = fn +applyStackWithFC fn ((fc, arg) :%: args) = applyStackWithFC (App fc fn arg) args -- Build a simple function type export @@ -474,15 +475,15 @@ Eq (Term vars) where mutual - resolveNamesBinder : (vars : List Name) -> Binder (Term vars) -> Binder (Term vars) + resolveNamesBinder : (vars : ScopedList Name) -> Binder (Term vars) -> Binder (Term vars) resolveNamesBinder vars b = assert_total $ map (resolveNames vars) b - resolveNamesTerms : (vars : List Name) -> List (Term vars) -> List (Term vars) + resolveNamesTerms : (vars : ScopedList Name) -> List (Term vars) -> List (Term vars) resolveNamesTerms vars ts = assert_total $ map (resolveNames vars) ts -- Replace any Ref Bound in a type with appropriate local export - resolveNames : (vars : List Name) -> Term vars -> Term vars + resolveNames : (vars : ScopedList Name) -> Term vars -> Term vars resolveNames vars (Ref fc Bound name) = case isNVar name vars of Just (MkNVar prf) => Local fc (Just False) _ prf @@ -490,7 +491,7 @@ mutual resolveNames vars (Meta fc n i xs) = Meta fc n i (resolveNamesTerms vars xs) resolveNames vars (Bind fc x b scope) - = Bind fc x (resolveNamesBinder vars b) (resolveNames (x :: vars) scope) + = Bind fc x (resolveNamesBinder vars b) (resolveNames (x :%: vars) scope) resolveNames vars (App fc fn arg) = App fc (resolveNames vars fn) (resolveNames vars arg) resolveNames vars (As fc s as pat) diff --git a/src/Core/TT/Term/Subst.idr b/src/Core/TT/Term/Subst.idr index 05e0b3e00d..3845b455e9 100644 --- a/src/Core/TT/Term/Subst.idr +++ b/src/Core/TT/Term/Subst.idr @@ -2,6 +2,7 @@ module Core.TT.Term.Subst import Core.Name import Core.Name.Scoped +import Core.Name.ScopedList import Core.TT.Binder import Core.TT.Subst @@ -47,9 +48,9 @@ substBinder outer dropped env b = assert_total $ map (substTerm outer dropped env) b export -substs : SizeOf dropped -> SubstEnv dropped vars -> Term (dropped ++ vars) -> Term vars +substs : SizeOf dropped -> SubstEnv dropped vars -> Term (dropped +%+ vars) -> Term vars substs dropped env tm = substTerm zero dropped env tm export -subst : Term vars -> Term (x :: vars) -> Term vars +subst : Term vars -> Term (x :%: vars) -> Term vars subst val tm = substs (suc zero) [val] tm diff --git a/src/Core/TT/Var.idr b/src/Core/TT/Var.idr index 163c885750..a248cae7af 100644 --- a/src/Core/TT/Var.idr +++ b/src/Core/TT/Var.idr @@ -15,6 +15,10 @@ import Libraries.Data.SnocList.SizeOf import Libraries.Data.Erased +import Core.Name.ScopedList +import Core.Name.ScopedList.HasLength +import Core.Name.ScopedList.SizeOf + %default total ------------------------------------------------------------------------ @@ -25,9 +29,9 @@ import Libraries.Data.Erased ||| is a position k ||| in the snoclist ns public export -data IsVar : a -> Nat -> List a -> Type where - First : IsVar n Z (n :: ns) - Later : IsVar n i ns -> IsVar n (S i) (m :: ns) +data IsVar : a -> Nat -> ScopedList a -> Type where + First : IsVar n Z (n :%: ns) + Later : IsVar n i ns -> IsVar n (S i) (m :%: ns) %name IsVar idx @@ -40,55 +44,57 @@ finIdx (Later l) = FS (finIdx l) ||| Recover the value pointed at by an IsVar proof ||| O(n) in the size of the index export -nameAt : {vars : List a} -> {idx : Nat} -> (0 p : IsVar n idx vars) -> a -nameAt {vars = n :: _} First = n +nameAt : {vars : ScopedList a} -> {idx : Nat} -> (0 p : IsVar n idx vars) -> a +nameAt {vars = n :%: _} First = n nameAt (Later p) = nameAt p ||| Inversion principle for Later export -dropLater : IsVar nm (S idx) (n :: ns) -> IsVar nm idx ns +dropLater : IsVar nm (S idx) (n :%: ns) -> IsVar nm idx ns dropLater (Later p) = p export -0 mkIsVar : HasLength m inner -> IsVar nm m (inner ++ nm :: outer) +0 mkIsVar : HasLength m inner -> IsVar nm m (inner +%+ nm :%: outer) mkIsVar Z = First mkIsVar (S x) = Later (mkIsVar x) +-- SLZ = Core.Name.ScopedList.LHasLength.Z + export -0 mkIsVarChiply : HasLength m inner -> IsVar nm m (inner <>> nm :: outer) +0 mkIsVarChiply : HasLength m inner -> IsVar nm m (inner <>> nm :%: outer) mkIsVarChiply hl - = rewrite chipsAsListAppend inner (nm :: outer) in + = rewrite chipsAsListAppend inner (nm :%: outer) in rewrite sym $ plusZeroRightNeutral m in mkIsVar (hlChips hl Z) ||| Compute the remaining scope once the target variable has been removed public export dropIsVar : - (ns : List a) -> + (ns : ScopedList a) -> {idx : Nat} -> (0 p : IsVar name idx ns) -> - List a -dropIsVar (_ :: xs) First = xs -dropIsVar (n :: xs) (Later p) = n :: dropIsVar xs p + ScopedList a +dropIsVar (_ :%: xs) First = xs +dropIsVar (n :%: xs) (Later p) = n :%: dropIsVar xs p ||| Throw in extra variables on the outer side of the context ||| This is essentially the identity function ||| This is slow so we ensure it's only used in a runtime irrelevant manner export -0 embedIsVar : IsVar x idx xs -> IsVar x idx (xs ++ outer) +0 embedIsVar : IsVar x idx xs -> IsVar x idx (xs +%+ outer) embedIsVar First = First embedIsVar (Later p) = Later (embedIsVar p) ||| Throw in extra variables on the local end of the context. ||| This is slow so we ensure it's only used in a runtime irrelevant manner export -0 weakenIsVar : (s : SizeOf ns) -> IsVar x idx xs -> IsVar x (size s + idx) (ns ++ xs) +0 weakenIsVar : (s : SizeOf ns) -> IsVar x idx xs -> IsVar x (size s + idx) (ns +%+ xs) weakenIsVar (MkSizeOf Z Z) p = p weakenIsVar (MkSizeOf (S k) (S l)) p = Later (weakenIsVar (MkSizeOf k l) p) 0 locateIsVarLT : (s : SizeOf local) -> So (idx < size s) -> - IsVar x idx (local ++ outer) -> + IsVar x idx (local +%+ outer) -> IsVar x idx local locateIsVarLT (MkSizeOf Z Z) so v = case v of First impossible @@ -100,7 +106,7 @@ locateIsVarLT (MkSizeOf (S k) (S l)) so v = case v of 0 locateIsVarGE : (s : SizeOf local) -> So (idx >= size s) -> - IsVar x idx (local ++ outer) -> + IsVar x idx (local +%+ outer) -> IsVar x (idx `minus` size s) outer locateIsVarGE (MkSizeOf Z Z) so v = rewrite minusZeroRight idx in v locateIsVarGE (MkSizeOf (S k) (S l)) so v = case v of @@ -108,7 +114,7 @@ locateIsVarGE (MkSizeOf (S k) (S l)) so v = case v of export locateIsVar : {idx : Nat} -> (s : SizeOf local) -> - (0 p : IsVar x idx (local ++ outer)) -> + (0 p : IsVar x idx (local +%+ outer)) -> Either (Erased (IsVar x idx local)) (Erased (IsVar x (idx `minus` size s) outer)) locateIsVar s p = case choose (idx < size s) of @@ -122,7 +128,7 @@ locateIsVar s p = case choose (idx < size s) of ||| and a proof that the name is at that position in the scope. ||| Everything but the De Bruijn index is erased. public export -record Var {0 a : Type} (vars : List a) where +record Var {0 a : Type} (vars : ScopedList a) where constructor MkVar {varIdx : Nat} {0 varNm : a} @@ -131,52 +137,51 @@ record Var {0 a : Type} (vars : List a) where namespace Var export - later : Var ns -> Var (n :: ns) + later : Var ns -> Var (n :%: ns) later (MkVar p) = MkVar (Later p) export - isLater : Var (n :: vs) -> Maybe (Var vs) + isLater : Var (n :%: vs) -> Maybe (Var vs) isLater (MkVar First) = Nothing isLater (MkVar (Later p)) = Just (MkVar p) export -mkVar : SizeOf inner -> Var (inner ++ nm :: outer) +mkVar : SizeOf inner -> Var (inner +%+ nm :%: outer) mkVar (MkSizeOf s p) = MkVar (mkIsVar p) export -mkVarChiply : SizeOf inner -> Var (inner <>> nm :: outer) +mkVarChiply : SizeOf inner -> Var (inner <>> nm :%: outer) mkVarChiply (MkSizeOf s p) = MkVar (mkIsVarChiply p) ||| Generate all variables export -allVars : (vars : Scope) -> List (Var vars) +allVars : (vars : Scope) -> ScopedList (Var vars) allVars = go [<] where - go : SizeOf local -> (vs : Scope) -> List (Var (local <>> vs)) - go s [] = [] - go s (v :: vs) = mkVarChiply s :: go (s :< v) vs + go : SizeOf local -> (vs : Scope) -> ScopedList (Var (local <>> vs)) + go s SLNil = SLNil + go s (v :%: vs) = mkVarChiply s :%: go (s :< v) vs export Eq (Var xs) where v == w = varIdx v == varIdx w - ||| Removing var 0, strengthening all the other ones export -dropFirst : List (Var (n :: vs)) -> List (Var vs) -dropFirst = List.mapMaybe isLater +dropFirst : ScopedList (Var (n :%: vs)) -> ScopedList (Var vs) +dropFirst = ScopedList.mapMaybe isLater ||| Manufacturing a thinning from a list of variables to keep export thinFromVars : - (vars : _) -> List (Var vars) -> + (vars : _) -> ScopedList (Var vars) -> (newvars ** Thin newvars vars) -thinFromVars [] els +thinFromVars SLNil els = (_ ** Refl) -thinFromVars (x :: xs) els +thinFromVars (x :%: xs) els = let (vs ** subRest) = thinFromVars xs (dropFirst els) in if MkVar First `elem` els - then (x :: vs ** Keep subRest) + then (x :%: vs ** Keep subRest) else (vs ** Drop subRest) export @@ -187,18 +192,18 @@ Show (Var ns) where -- Named variable in scope public export -record NVar {0 a : Type} (nm : a) (vars : List a) where +record NVar {0 a : Type} (nm : a) (vars : ScopedList a) where constructor MkNVar {nvarIdx : Nat} 0 nvarPrf : IsVar nm nvarIdx vars namespace NVar export - later : NVar nm ns -> NVar nm (n :: ns) + later : NVar nm ns -> NVar nm (n :%: ns) later (MkNVar p) = MkNVar (Later p) export - isLater : NVar nm (n :: ns) -> Maybe (NVar nm ns) + isLater : NVar nm (n :%: ns) -> Maybe (NVar nm ns) isLater (MkNVar First) = Nothing isLater (MkNVar (Later p)) = Just (MkNVar p) @@ -211,47 +216,47 @@ recoverName : (v : Var vars) -> NVar (varNm v) vars recoverName (MkVar p) = MkNVar p export -mkNVar : SizeOf inner -> NVar nm (inner ++ nm :: outer) +mkNVar : SizeOf inner -> NVar nm (inner +%+ nm :%: outer) mkNVar (MkSizeOf s p) = MkNVar (mkIsVar p) export -mkNVarChiply : SizeOf inner -> NVar nm (inner <>> nm :: outer) +mkNVarChiply : SizeOf inner -> NVar nm (inner <>> nm :%: outer) mkNVarChiply (MkSizeOf s p) = MkNVar (mkIsVarChiply p) export -locateNVar : SizeOf local -> NVar nm (local ++ outer) -> +locateNVar : SizeOf local -> NVar nm (local +%+ outer) -> Either (NVar nm local) (NVar nm outer) locateNVar s (MkNVar p) = case locateIsVar s p of Left p => Left (MkNVar (runErased p)) Right p => Right (MkNVar (runErased p)) public export -dropNVar : {ns : List a} -> NVar nm ns -> List a +dropNVar : {ns : ScopedList a} -> NVar nm ns -> ScopedList a dropNVar (MkNVar p) = dropIsVar ns p ------------------------------------------------------------------------ -- Scope checking export -isDeBruijn : Nat -> (vars : List Name) -> Maybe (Var vars) -isDeBruijn Z (_ :: _) = pure (MkVar First) -isDeBruijn (S k) (_ :: vs) = later <$> isDeBruijn k vs +isDeBruijn : Nat -> (vars : ScopedList Name) -> Maybe (Var vars) +isDeBruijn Z (_ :%: _) = pure (MkVar First) +isDeBruijn (S k) (_ :%: vs) = later <$> isDeBruijn k vs isDeBruijn _ _ = Nothing export -isNVar : (n : Name) -> (ns : List Name) -> Maybe (NVar n ns) -isNVar n [] = Nothing -isNVar n (m :: ms) +isNVar : (n : Name) -> (ns : ScopedList Name) -> Maybe (NVar n ns) +isNVar n SLNil = Nothing +isNVar n (m :%: ms) = case nameEq n m of Nothing => map later (isNVar n ms) Just Refl => pure (MkNVar First) export -isVar : (n : Name) -> (ns : List Name) -> Maybe (Var ns) +isVar : (n : Name) -> (ns : ScopedList Name) -> Maybe (Var ns) isVar n ns = forgetName <$> isNVar n ns export -locateVar : SizeOf local -> Var (local ++ outer) -> +locateVar : SizeOf local -> Var (local +%+ outer) -> Either (Var local) (Var outer) locateVar s v = bimap forgetName forgetName $ locateNVar s (recoverName v) @@ -260,18 +265,18 @@ locateVar s v = bimap forgetName forgetName -- Weakening export -weakenNVar : SizeOf ns -> NVar name outer -> NVar name (ns ++ outer) +weakenNVar : SizeOf ns -> NVar name outer -> NVar name (ns +%+ outer) weakenNVar s (MkNVar {nvarIdx} p) = MkNVar {nvarIdx = plus (size s) nvarIdx} (weakenIsVar s p) export -embedNVar : NVar name ns -> NVar name (ns ++ outer) +embedNVar : NVar name ns -> NVar name (ns +%+ outer) embedNVar (MkNVar p) = MkNVar (embedIsVar p) export insertNVar : SizeOf local -> - NVar nm (local ++ outer) -> - NVar nm (local ++ n :: outer) + NVar nm (local +%+ outer) -> + NVar nm (local +%+ n :%: outer) insertNVar p v = case locateNVar p v of Left v => embedNVar v Right v => weakenNVar p (later v) @@ -279,9 +284,9 @@ insertNVar p v = case locateNVar p v of export insertNVarChiply : SizeOf local -> NVar nm (local <>> outer) -> - NVar nm (local <>> n :: outer) + NVar nm (local <>> n :%: outer) insertNVarChiply p v - = rewrite chipsAsListAppend local (n :: outer) in + = rewrite chipsAsListAppend local (n :%: outer) in insertNVar (p <>> zero) $ replace {p = NVar nm} (chipsAsListAppend local outer) v @@ -296,19 +301,19 @@ insertNVarNames p q v = case locateNVar p v of ||| The (partial) inverse to insertNVar export removeNVar : SizeOf local -> - NVar nm (local ++ n :: outer) -> - Maybe (NVar nm (local ++ outer)) + NVar nm (local +%+ n :%: outer) -> + Maybe (NVar nm (local +%+ outer)) removeNVar s var = case locateNVar s var of Left v => pure (embedNVar v) Right v => weakenNVar s <$> isLater v export insertVar : SizeOf local -> - Var (local ++ outer) -> - Var (local ++ n :: outer) + Var (local +%+ outer) -> + Var (local +%+ n :%: outer) insertVar p v = forgetName $ insertNVar p (recoverName v) -weakenVar : SizeOf ns -> Var outer -> Var (ns ++ outer) +weakenVar : SizeOf ns -> Var outer -> Var (ns +%+ outer) weakenVar p v = forgetName $ weakenNVar p (recoverName v) insertVarNames : GenWeakenable Var @@ -317,8 +322,8 @@ insertVarNames p q v = forgetName $ insertNVarNames p q (recoverName v) ||| The (partial) inverse to insertVar export removeVar : SizeOf local -> - Var (local ++ n :: outer) -> - Maybe (Var (local ++ outer)) + Var (local +%+ n :%: outer) -> + Maybe (Var (local +%+ outer)) removeVar s var = forgetName <$> removeNVar s (recoverName var) ------------------------------------------------------------------------ @@ -326,20 +331,20 @@ removeVar s var = forgetName <$> removeNVar s (recoverName var) export strengthenIsVar : {n : Nat} -> (s : SizeOf inner) -> - (0 p : IsVar x n (inner ++ vars)) -> + (0 p : IsVar x n (inner +%+ vars)) -> Maybe (Erased (IsVar x (n `minus` size s) vars)) strengthenIsVar s p = case locateIsVar s p of Left _ => Nothing Right p => pure p strengthenVar : SizeOf inner -> - Var (inner ++ vars) -> Maybe (Var vars) + Var (inner +%+ vars) -> Maybe (Var vars) strengthenVar s (MkVar p) = do MkErased p <- strengthenIsVar s p pure (MkVar p) strengthenNVar : SizeOf inner -> - NVar x (inner ++ vars) -> Maybe (NVar x vars) + NVar x (inner +%+ vars) -> Maybe (NVar x vars) strengthenNVar s (MkNVar p) = do MkErased p <- strengthenIsVar s p pure (MkNVar p) @@ -443,7 +448,7 @@ FreelyEmbeddable (NVar {a = Name} nm) where export shiftUnderNs : SizeOf {a = Name} inner -> {idx : _} -> - (0 p : IsVar n idx (x :: inner ++ outer)) -> - NVar n (inner ++ x :: outer) + (0 p : IsVar n idx (x :%: inner +%+ outer)) -> + NVar n (inner +%+ x :%: outer) shiftUnderNs s First = weakenNs s (MkNVar First) shiftUnderNs s (Later p) = insertNVar s (MkNVar p) diff --git a/src/Core/TTC.idr b/src/Core/TTC.idr index 2d7e912b3c..06c39ccd70 100644 --- a/src/Core/TTC.idr +++ b/src/Core/TTC.idr @@ -260,13 +260,13 @@ TTC NameType where -- (Indeed, we're expecting the whole IsVar proof to be erased because -- we have the idx...) mkPrf : (idx : Nat) -> IsVar n idx ns -mkPrf {n} {ns} Z = believe_me (First {n} {ns = n :: ns}) +mkPrf {n} {ns} Z = believe_me (First {n} {ns = n :%: ns}) mkPrf {n} {ns} (S k) = believe_me (Later {m=n} (mkPrf {n} {ns} k)) -getName : (idx : Nat) -> List Name -> Maybe Name -getName Z (x :: xs) = Just x -getName (S k) (x :: xs) = getName k xs -getName _ [] = Nothing +getName : (idx : Nat) -> ScopedList Name -> Maybe Name +getName Z (x :%: xs) = Just x +getName (S k) (x :%: xs) = getName k xs +getName _ SLNil = Nothing mutual export @@ -488,8 +488,8 @@ export = do toBuf b bnd; toBuf b env -- Length has to correspond to length of 'vars' - fromBuf {vars = []} b = pure Nil - fromBuf {vars = x :: xs} b + fromBuf {vars = SLNil} b = pure Nil + fromBuf {vars = x :%: xs} b = do bnd <- fromBuf b env <- fromBuf b pure (bnd :: env) @@ -1151,7 +1151,7 @@ TTC GlobalDef where mul vars vis tot hatch fl refs refsR inv c True def cdef Nothing sc Nothing) else pure (MkGlobalDef loc name (Erased loc Placeholder) [] [] [] [] - mul [] (specified Public) unchecked False [] refs refsR + mul SLNil (specified Public) unchecked False [] refs refsR False False True def cdef Nothing [] Nothing) export diff --git a/src/Core/Termination/Positivity.idr b/src/Core/Termination/Positivity.idr index 403a7272cd..7d4099b6e2 100644 --- a/src/Core/Termination/Positivity.idr +++ b/src/Core/Termination/Positivity.idr @@ -21,7 +21,7 @@ isAssertTotal (NRef _ fn_in) = isAssertTotal _ = pure False nameIn : {auto c : Ref Ctxt Defs} -> - Defs -> List Name -> NF [] -> Core Bool + Defs -> List Name -> NF SLNil -> Core Bool nameIn defs tyns (NBind fc x b sc) = if !(nameIn defs tyns !(evalClosure defs (binderType b))) then pure True @@ -33,25 +33,25 @@ nameIn defs tyns (NApp _ nh args) = do False <- isAssertTotal nh | True => pure False anyM (nameIn defs tyns) - !(traverse (evalClosure defs . snd) args) + !(traverse (evalClosure defs . snd) (toList args)) nameIn defs tyns (NTCon _ n _ _ args) = if n `elem` tyns then pure True - else do args' <- traverse (evalClosure defs . snd) args + else do args' <- traverse (evalClosure defs . snd) (toList args) anyM (nameIn defs tyns) args' nameIn defs tyns (NDCon _ n _ _ args) = anyM (nameIn defs tyns) - !(traverse (evalClosure defs . snd) args) + !(traverse (evalClosure defs . snd) (toList args)) nameIn defs tyns (NDelayed fc lr ty) = nameIn defs tyns ty nameIn defs tyns _ = pure False -- Check an argument type doesn't contain a negative occurrence of any of -- the given type names posArg : {auto c : Ref Ctxt Defs} -> - Defs -> List Name -> NF [] -> Core Terminating + Defs -> List Name -> NF SLNil -> Core Terminating posArgs : {auto c : Ref Ctxt Defs} -> - Defs -> List Name -> List (Closure []) -> Core Terminating + Defs -> List Name -> List (Closure SLNil) -> Core Terminating posArgs defs tyn [] = pure IsTerminating posArgs defs tyn (x :: xs) = do xNF <- evalClosure defs x @@ -66,18 +66,18 @@ posArg defs tyns nf@(NTCon loc tc _ _ args) = do logNF "totality.positivity" 50 "Found a type constructor" [] nf testargs <- case !(lookupDefExact tc (gamma defs)) of Just (TCon _ _ params _ _ _ _ _) => do - log "totality.positivity" 50 $ - unwords [show tc, "has", show (length params), "parameters"] - pure $ splitParams 0 params (map snd args) + logC "totality.positivity" 50 $ + do pure $ unwords [show tc, "has", show (length params), "parameters"] + pure $ splitParams 0 params (map snd (toList args)) _ => throw (GenericMsg loc (show tc ++ " not a data type")) let (params, indices) = testargs False <- anyM (nameIn defs tyns) !(traverse (evalClosure defs) indices) | True => pure (NotTerminating NotStrictlyPositive) posArgs defs tyns params where - splitParams : Nat -> List Nat -> List (Closure []) -> - ( List (Closure []) -- parameters (to be checked for strict positivity) - , List (Closure []) -- indices (to be checked for no mention at all) + splitParams : Nat -> List Nat -> List (Closure SLNil) -> + ( List (Closure SLNil) -- parameters (to be checked for strict positivity) + , List (Closure SLNil) -- indices (to be checked for no mention at all) ) splitParams i ps [] = ([], []) splitParams i ps (x :: xs) @@ -98,7 +98,7 @@ posArg defs tyns nf@(NApp fc nh args) | True => do logNF "totality.positivity" 50 "Trusting an assertion" [] nf pure IsTerminating logNF "totality.positivity" 50 "Found an application" [] nf - args <- traverse (evalClosure defs . snd) args + args <- traverse (evalClosure defs . snd) (toList args) pure $ if !(anyM (nameIn defs tyns) args) then NotTerminating NotStrictlyPositive else IsTerminating @@ -108,7 +108,7 @@ posArg defs tyn nf pure IsTerminating checkPosArgs : {auto c : Ref Ctxt Defs} -> - Defs -> List Name -> NF [] -> Core Terminating + Defs -> List Name -> NF SLNil -> Core Terminating checkPosArgs defs tyns (NBind fc x (Pi _ _ e ty) sc) = case !(posArg defs tyns !(evalClosure defs ty)) of IsTerminating => diff --git a/src/Core/Transform.idr b/src/Core/Transform.idr index c7c22819ed..a06df8a05c 100644 --- a/src/Core/Transform.idr +++ b/src/Core/Transform.idr @@ -13,7 +13,7 @@ unload [] fn = fn unload ((fc, arg) :: args) fn = unload args (App fc fn arg) -- List of matches on LHS -data MatchVars : List Name -> List Name -> Type where +data MatchVars : ScopedList Name -> ScopedList Name -> Type where None : MatchVars lhsvars vs Match : (idx : Nat) -> (0 p : IsVar n idx lhsvars) -> Term vs -> MatchVars lhsvars vs -> MatchVars lhsvars vs diff --git a/src/Core/Unify.idr b/src/Core/Unify.idr index 3d1ccfd06b..b026d941b5 100644 --- a/src/Core/Unify.idr +++ b/src/Core/Unify.idr @@ -112,7 +112,7 @@ solvedHole n = MkUnifyResult [] True [n] NoLazy public export interface Unify tm where -- Unify returns a list of ids referring to newly added constraints - unifyD : {vars : List Name} -> + unifyD : {vars : ScopedList Name} -> Ref Ctxt Defs -> Ref UST UState -> UnifyInfo -> @@ -264,10 +264,10 @@ unifyArgs : (Unify tm, Quote tm) => {auto c : Ref Ctxt Defs} -> {auto u : Ref UST UState} -> UnifyInfo -> FC -> Env Term vars -> - List (tm vars) -> List (tm vars) -> + ScopedList (tm vars) -> ScopedList (tm vars) -> Core UnifyResult -unifyArgs mode loc env [] [] = pure success -unifyArgs mode loc env (cx :: cxs) (cy :: cys) +unifyArgs mode loc env SLNil SLNil = pure success +unifyArgs mode loc env (cx :%: cxs) (cy :%: cys) = do -- Do later arguments first, since they may depend on earlier -- arguments and use their solutions. cs <- unifyArgs mode loc env cxs cys @@ -280,13 +280,13 @@ unifyArgs mode loc env _ _ = ufail loc "" -- We use this to check that the pattern unification rule is applicable -- when solving a metavariable applied to arguments getVars : {vars : _} -> - List Nat -> List (NF vars) -> Maybe (List (Var vars)) -getVars got [] = Just [] -getVars got (NErased fc (Dotted t) :: xs) = getVars got (t :: xs) -getVars got (NApp fc (NLocal r idx v) [] :: xs) + List Nat -> ScopedList (NF vars) -> Maybe (ScopedList (Var vars)) +getVars got SLNil = Just SLNil +getVars got (NErased fc (Dotted t) :%: xs) = getVars got (t :%: xs) +getVars got (NApp fc (NLocal r idx v) SLNil :%: xs) = if inArgs idx got then Nothing else do xs' <- getVars (idx :: got) xs - pure (MkVar v :: xs') + pure (MkVar v :%: xs') where -- Save the overhead of the call to APPLY, and the fact that == on -- Nat is linear time in Idris 1! @@ -294,16 +294,16 @@ getVars got (NApp fc (NLocal r idx v) [] :: xs) inArgs n [] = False inArgs n (n' :: ns) = natToInteger n == natToInteger n' || inArgs n ns -getVars got (NAs _ _ _ p :: xs) = getVars got (p :: xs) -getVars _ (_ :: xs) = Nothing +getVars got (NAs _ _ _ p :%: xs) = getVars got (p :%: xs) +getVars _ (_ :%: xs) = Nothing -- Make a sublist representing the variables used in the application. -- We'll use this to ensure that local variables which appear in a term -- are all arguments to a metavariable application for pattern unification -toThin : (vars : List Name) -> List (Var vars) -> +toThin : (vars : ScopedList Name) -> ScopedList (Var vars) -> (newvars ** Thin newvars vars) -toThin [] xs = ([] ** Refl) -toThin (n :: ns) xs +toThin SLNil xs = (SLNil ** Refl) +toThin (n :%: ns) xs -- If there's a proof 'First' in 'xs', then 'n' should be kept, -- otherwise dropped -- (Remember: 'n' might be shadowed; looking for 'First' ensures we @@ -313,15 +313,15 @@ toThin (n :: ns) xs then (_ ** Keep svs) else (_ ** Drop svs) where - anyFirst : List (Var (n :: ns)) -> Bool - anyFirst [] = False - anyFirst (MkVar First :: xs) = True - anyFirst (MkVar (Later p) :: xs) = anyFirst xs + anyFirst : ScopedList (Var (n :%: ns)) -> Bool + anyFirst SLNil = False + anyFirst (MkVar First :%: xs) = True + anyFirst (MkVar (Later p) :%: xs) = anyFirst xs -- Update the variable list to point into the sub environment -- (All of these will succeed because the Thin we have comes from -- the list of variable uses! It's not stated in the type, though.) -updateVars : List (Var {a = Name} vars) -> Thin newvars vars -> List (Var newvars) +updateVars : ScopedList (Var {a = Name} vars) -> Thin newvars vars -> ScopedList (Var newvars) updateVars vs th = mapMaybe (\ v => shrink v th) vs {- Applying the pattern unification rule is okay if: @@ -342,8 +342,8 @@ updateVars vs th = mapMaybe (\ v => shrink v th) vs patternEnv : {auto c : Ref Ctxt Defs} -> {auto u : Ref UST UState} -> {vars : _} -> - Env Term vars -> List (Closure vars) -> - Core (Maybe (newvars ** (List (Var newvars), + Env Term vars -> ScopedList (Closure vars) -> + Core (Maybe (newvars ** (ScopedList (Var newvars), Thin newvars vars))) patternEnv {vars} env args = do defs <- get Ctxt @@ -356,12 +356,12 @@ patternEnv {vars} env args let (newvars ** svs) = toThin _ vs in Just (newvars ** (updateVars vs svs, svs)) -getVarsTm : List Nat -> List (Term vars) -> Maybe (List (Var vars)) -getVarsTm got [] = Just [] +getVarsTm : List Nat -> List (Term vars) -> Maybe (ScopedList (Var vars)) +getVarsTm got [] = Just SLNil getVarsTm got (Local fc r idx v :: xs) = if idx `elem` got then Nothing else do xs' <- getVarsTm (idx :: got) xs - pure (MkVar v :: xs') + pure (MkVar v :%: xs') getVarsTm _ (_ :: xs) = Nothing export @@ -369,7 +369,7 @@ patternEnvTm : {auto c : Ref Ctxt Defs} -> {auto u : Ref UST UState} -> {vars : _} -> Env Term vars -> List (Term vars) -> - Core (Maybe (newvars ** (List (Var newvars), + Core (Maybe (newvars ** (ScopedList (Var newvars), Thin newvars vars))) patternEnvTm {vars} env args = do defs <- get Ctxt @@ -417,10 +417,10 @@ occursCheck fc env mode mname tm -- How the variables in a metavariable definition map to the variables in -- the solution term (the Var newvars) -data IVars : List Name -> List Name -> Type where - INil : IVars [] newvars +data IVars : ScopedList Name -> ScopedList Name -> Type where + INil : IVars SLNil newvars ICons : Maybe (Var newvars) -> IVars vs newvars -> - IVars (v :: vs) newvars + IVars (v :%: vs) newvars Weaken (IVars vs) where weakenNs s INil = INil @@ -440,7 +440,7 @@ tryInstantiate : {auto c : Ref Ctxt Defs} -> FC -> UnifyInfo -> Env Term vars -> (metavar : Name) -> (mref : Int) -> (numargs : Nat) -> (mdef : GlobalDef) -> - List (Var newvars) -> -- Variable each argument maps to + ScopedList (Var newvars) -> -- Variable each argument maps to Term vars -> -- original, just for error message Term newvars -> -- shrunk environment Core Bool -- postpone if the type is yet unknown @@ -472,7 +472,7 @@ tryInstantiate {newvars} loc mode env mname mref num mdef locs otm tm (not (isUserName mname) && isSimple rhs) False let newdef = { definition := - PMDef simpleDef [] (STerm 0 rhs) (STerm 0 rhs) [] + PMDef simpleDef SLNil (STerm 0 rhs) (STerm 0 rhs) [] } mdef ignore $ addDef (Resolved mref) newdef removeHole mref @@ -572,10 +572,10 @@ tryInstantiate {newvars} loc mode env mname mref num mdef locs otm tm updateIVars ivs (TType fc u) = Just (TType fc u) mkDef : {vs, newvars : _} -> - List (Var newvars) -> + ScopedList (Var newvars) -> IVars vs newvars -> Term newvars -> Term vs -> Core (Maybe (Term vs)) - mkDef (v :: vs) vars soln (Bind bfc x (Pi fc c _ ty) sc) + mkDef (v :%: vs) vars soln (Bind bfc x (Pi fc c _ ty) sc) = do sc' <- mkDef vs (ICons (Just v) vars) soln sc pure $ (Bind bfc x (Lam fc c Explicit (Erased bfc Placeholder)) <$> sc') mkDef vs vars soln (Bind bfc x b@(Let _ c val ty) sc) @@ -584,7 +584,7 @@ tryInstantiate {newvars} loc mode env mname mref num mdef locs otm tm case shrink sc' (Drop Refl) of Just scs => pure scs Nothing => pure $ Bind bfc x b sc' - mkDef [] vars soln _ + mkDef SLNil vars soln _ = do let Just soln' = updateIVars vars soln | Nothing => ufail loc ("Can't make solution for " ++ show mname ++ " " ++ show (getIVars vars, soln)) @@ -653,13 +653,13 @@ mutual getArgTypes : {vars : _} -> {auto c : Ref Ctxt Defs} -> - Defs -> (fnType : NF vars) -> List (Closure vars) -> - Core (Maybe (List (NF vars))) - getArgTypes defs (NBind _ n (Pi _ _ _ ty) sc) (a :: as) + Defs -> (fnType : NF vars) -> ScopedList (Closure vars) -> + Core (Maybe (ScopedList (NF vars))) + getArgTypes defs (NBind _ n (Pi _ _ _ ty) sc) (a :%: as) = do Just scTys <- getArgTypes defs !(sc defs a) as | Nothing => pure Nothing - pure (Just (!(evalClosure defs ty) :: scTys)) - getArgTypes _ _ [] = pure (Just []) + pure (Just (!(evalClosure defs ty) :%: scTys)) + getArgTypes _ _ SLNil = pure (Just SLNil) getArgTypes _ _ _ = pure Nothing headsConvert : {vars : _} -> @@ -667,11 +667,11 @@ mutual {auto u : Ref UST UState} -> UnifyInfo -> FC -> Env Term vars -> - Maybe (List (NF vars)) -> Maybe (List (NF vars)) -> + Maybe (ScopedList (NF vars)) -> Maybe (ScopedList (NF vars)) -> Core Bool headsConvert mode fc env (Just vs) (Just ns) = case (reverse vs, reverse ns) of - (v :: _, n :: _) => + (v :%: _, n :%: _) => do logNF "unify.head" 10 "Unifying head" env v logNF "unify.head" 10 ".........with" env n res <- unify mode fc env v n @@ -689,11 +689,11 @@ mutual (swaporder : Bool) -> UnifyInfo -> FC -> Env Term vars -> (metaname : Name) -> (metaref : Int) -> - (margs : List (Closure vars)) -> - (margs' : List (Closure vars)) -> + (margs : ScopedList (Closure vars)) -> + (margs' : ScopedList (Closure vars)) -> Maybe ClosedTerm -> - (List (FC, Closure vars) -> NF vars) -> - List (FC, Closure vars) -> + (ScopedList (FC, Closure vars) -> NF vars) -> + ScopedList (FC, Closure vars) -> Core UnifyResult unifyInvertible swap mode fc env mname mref margs margs' nty con args' = do defs <- get Ctxt @@ -701,7 +701,7 @@ mutual -- argument types match up Just vty <- lookupTyExact (Resolved mref) (gamma defs) | Nothing => ufail fc ("No such metavariable " ++ show mname) - vargTys <- getArgTypes defs !(nf defs env (embed vty)) (margs ++ margs') + vargTys <- getArgTypes defs !(nf defs env (embed vty)) (margs +%+ margs') nargTys <- maybe (pure Nothing) (\ty => getArgTypes defs !(nf defs env (embed ty)) $ map snd args') nty @@ -712,7 +712,7 @@ mutual -- Unify the rightmost arguments, with the goal of turning the -- hole application into a pattern form case (reverse margs', reverse args') of - (h :: hargs, f :: fargs) => + (h :%: hargs, f :%: fargs) => tryUnify (if not swap then do log "unify.invertible" 10 "Unifying invertible" @@ -731,7 +731,7 @@ mutual (NApp fc (NMeta mname mref margs) (reverse $ map (EmptyFC,) hargs)) pure (union ures uargs)) (postponeS swap fc mode "Postponing hole application [1]" env - (NApp fc (NMeta mname mref margs) $ map (EmptyFC,) margs') + (NApp fc (NMeta mname mref margs) $ (map (EmptyFC,) margs')) (con args')) _ => postponeS swap fc mode "Postponing hole application [2]" env (NApp fc (NMeta mname mref margs) (map (EmptyFC,) margs')) @@ -749,8 +749,8 @@ mutual (swaporder : Bool) -> UnifyInfo -> FC -> Env Term vars -> (metaname : Name) -> (metaref : Int) -> - (margs : List (Closure vars)) -> - (margs' : List (Closure vars)) -> + (margs : ScopedList (Closure vars)) -> + (margs' : ScopedList (Closure vars)) -> NF vars -> Core UnifyResult unifyHoleApp swap mode loc env mname mref margs margs' (NTCon nfc n t a args') @@ -789,8 +789,8 @@ mutual (swaporder : Bool) -> UnifyInfo -> FC -> Env Term vars -> (metaname : Name) -> (metaref : Int) -> - (margs : List (Closure vars)) -> - (margs' : List (Closure vars)) -> + (margs : ScopedList (Closure vars)) -> + (margs' : ScopedList (Closure vars)) -> (soln : NF vars) -> Core UnifyResult postponePatVar swap mode loc env mname mref margs margs' tm @@ -806,9 +806,9 @@ mutual {newvars, vars : _} -> FC -> UnifyInfo -> Env Term vars -> (metaname : Name) -> (metaref : Int) -> - (margs : List (Closure vars)) -> - (margs' : List (Closure vars)) -> - List (Var newvars) -> + (margs : ScopedList (Closure vars)) -> + (margs' : ScopedList (Closure vars)) -> + ScopedList (Var newvars) -> Thin newvars vars -> (solfull : Term vars) -> -- Original solution (soln : Term newvars) -> -- Solution with shrunk environment @@ -851,14 +851,14 @@ mutual (swaporder : Bool) -> UnifyInfo -> FC -> Env Term vars -> FC -> (metaname : Name) -> (metaref : Int) -> - (args : List (Closure vars)) -> - (args' : List (Closure vars)) -> + (args : ScopedList (Closure vars)) -> + (args' : ScopedList (Closure vars)) -> (soln : NF vars) -> Core UnifyResult unifyHole swap mode loc env fc mname mref margs margs' tmnf = do defs <- get Ctxt empty <- clearDefs defs - let args = if isNil margs' then margs else margs ++ margs' + let args = if isNil margs' then margs else margs +%+ margs' logC "unify.hole" 10 (do args' <- traverse (evalArg empty) args qargs <- traverse (quote empty env) args' @@ -917,7 +917,7 @@ mutual (swaporder : Bool) -> -- swap the order when postponing -- (this is to preserve second arg being expected type) UnifyInfo -> FC -> Env Term vars -> FC -> - NHead vars -> List (FC, Closure vars) -> NF vars -> + NHead vars -> ScopedList (FC, Closure vars) -> NF vars -> Core UnifyResult unifyApp swap mode loc env fc (NMeta n i margs) args tm = unifyHole swap mode loc env fc n i margs (map snd args) tm @@ -932,12 +932,12 @@ mutual if not swap then unifyIfEq True loc mode env (NApp fc (NRef nt n) args) tm else unifyIfEq True loc mode env tm (NApp fc (NRef nt n) args) - unifyApp swap mode loc env xfc (NLocal rx x xp) [] (NApp yfc (NLocal ry y yp) []) + unifyApp swap mode loc env xfc (NLocal rx x xp) SLNil (NApp yfc (NLocal ry y yp) SLNil) = do gam <- get Ctxt if x == y then pure success else postponeS swap loc mode "Postponing var" - env (NApp xfc (NLocal rx x xp) []) - (NApp yfc (NLocal ry y yp) []) + env (NApp xfc (NLocal rx x xp) SLNil) + (NApp yfc (NLocal ry y yp) SLNil) -- A local against something canonical (binder or constructor) is bad unifyApp swap mode loc env xfc (NLocal rx x xp) args y@(NBind _ _ _ _) = convertErrorS swap loc env (NApp xfc (NLocal rx x xp) args) y @@ -968,14 +968,14 @@ mutual {auto u : Ref UST UState} -> {vars : _} -> UnifyInfo -> FC -> Env Term vars -> - FC -> NHead vars -> List (FC, Closure vars) -> - FC -> NHead vars -> List (FC, Closure vars) -> + FC -> NHead vars -> ScopedList (FC, Closure vars) -> + FC -> NHead vars -> ScopedList (FC, Closure vars) -> Core UnifyResult - unifyBothApps mode loc env xfc (NLocal xr x xp) [] yfc (NLocal yr y yp) [] + unifyBothApps mode loc env xfc (NLocal xr x xp) SLNil yfc (NLocal yr y yp) SLNil = if x == y then pure success - else convertError loc env (NApp xfc (NLocal xr x xp) []) - (NApp yfc (NLocal yr y yp) []) + else convertError loc env (NApp xfc (NLocal xr x xp) SLNil) + (NApp yfc (NLocal yr y yp) SLNil) -- Locally bound things, in a term (not LHS). Since we have to unify -- for *all* possible values, we can safely unify the arguments. unifyBothApps mode@(MkUnifyInfo p InTerm) loc env xfc (NLocal xr x xp) xargs yfc (NLocal yr y yp) yargs @@ -994,8 +994,8 @@ mutual if xi == yi && (invx || umode mode == InSearch) -- Invertible, (from auto implicit search) -- so we can also unify the arguments. - then unifyArgs mode loc env (xargs ++ map snd xargs') - (yargs ++ map snd yargs') + then unifyArgs mode loc env (xargs +%+ map snd xargs') + (yargs +%+ map snd yargs') else do xlocs <- localsIn xargs ylocs <- localsIn yargs -- Solve the one with the bigger context, and if they're @@ -1014,9 +1014,9 @@ mutual pv (PV _ _) = True pv _ = False - localsIn : List (Closure vars) -> Core Nat - localsIn [] = pure 0 - localsIn (c :: cs) + localsIn : ScopedList (Closure vars) -> Core Nat + localsIn SLNil = pure 0 + localsIn (c :%: cs) = do defs <- get Ctxt case !(evalClosure defs c) of NApp _ (NLocal _ _ _) _ => pure $ S !(localsIn cs) @@ -1070,7 +1070,7 @@ mutual pure ("Unifying arg types " ++ show tx' ++ " and " ++ show ty')) ct <- unify (lower mode) loc env tx ty xn <- genVarName "x" - let env' : Env Term (x :: _) + let env' : Env Term (x :%: _) = Pi fcy cy Explicit tx' :: env case constraints ct of [] => -- No constraints, check the scope @@ -1107,7 +1107,7 @@ mutual ct <- unify (lower mode) loc env tx ty xn <- genVarName "x" txtm <- quote empty env tx - let env' : Env Term (x :: _) + let env' : Env Term (x :%: _) = Lam fcx cx Explicit txtm :: env tscx <- scx defs (toClosure defaultOpts env (Ref loc Bound xn)) @@ -1192,7 +1192,7 @@ mutual unifyNoEta mode loc env (NDelayed xfc _ x) (NDelayed yfc _ y) = unify (lower mode) loc env x y unifyNoEta mode loc env (NDelay xfc _ xty x) (NDelay yfc _ yty y) - = unifyArgs mode loc env [xty, x] [yty, y] + = unifyArgs mode loc env (xty :%: x :%: SLNil) (yty :%: y :%: SLNil) unifyNoEta mode loc env (NForce xfc _ x axs) (NForce yfc _ y ays) = do cs <- unify (lower mode) loc env x y cs' <- unifyArgs mode loc env (map snd axs) (map snd ays) @@ -1439,7 +1439,7 @@ retryGuess mode smode (hid, (loc, hname)) handleUnify (do tm <- search loc rig (smode == Defaults) depth defining (type def) [] - let gdef = { definition := PMDef defaultPI [] (STerm 0 tm) (STerm 0 tm) [] } def + let gdef = { definition := PMDef defaultPI SLNil (STerm 0 tm) (STerm 0 tm) [] } def logTermNF "unify.retry" 5 ("Solved " ++ show hname) [] tm ignore $ addDef (Resolved hid) gdef removeGuess hid @@ -1475,7 +1475,7 @@ retryGuess mode smode (hid, (loc, hname)) logTerm "unify.retry" 5 "Retry Delay" tm pure $ delayMeta r envb !(getTerm ty) tm let gdef = { definition := PMDef (MkPMDefInfo NotHole True False) - [] (STerm 0 tm') (STerm 0 tm') [] } def + SLNil (STerm 0 tm') (STerm 0 tm') [] } def logTerm "unify.retry" 5 ("Resolved " ++ show hname) tm' ignore $ addDef (Resolved hid) gdef removeGuess hid @@ -1501,7 +1501,7 @@ retryGuess mode smode (hid, (loc, hname)) -- proper definition and remove it from the -- hole list [] => do let gdef = { definition := PMDef (MkPMDefInfo NotHole True False) - [] (STerm 0 tm) (STerm 0 tm) [] } def + SLNil (STerm 0 tm) (STerm 0 tm) [] } def logTerm "unify.retry" 5 ("Resolved " ++ show hname) tm ignore $ addDef (Resolved hid) gdef removeGuess hid @@ -1564,7 +1564,7 @@ checkArgsSame : {auto u : Ref UST UState} -> checkArgsSame [] = pure False checkArgsSame (x :: xs) = do defs <- get Ctxt - Just (PMDef _ [] (STerm 0 def) _ _) <- + Just (PMDef _ SLNil (STerm 0 def) _ _) <- lookupDefExact (Resolved x) (gamma defs) | _ => checkArgsSame xs s <- anySame def xs @@ -1572,11 +1572,11 @@ checkArgsSame (x :: xs) then pure True else checkArgsSame xs where - anySame : Term [] -> List Int -> Core Bool + anySame : Term SLNil -> List Int -> Core Bool anySame tm [] = pure False anySame tm (t :: ts) = do defs <- get Ctxt - Just (PMDef _ [] (STerm 0 def) _ _) <- + Just (PMDef _ SLNil (STerm 0 def) _ _) <- lookupDefExact (Resolved t) (gamma defs) | _ => anySame tm ts if !(convert defs [] tm def) @@ -1594,7 +1594,7 @@ checkDots hs <- getCurrentHoles update UST { dotConstraints := [] } where - getHoleName : Term [] -> Core (Maybe Name) + getHoleName : Term SLNil -> Core (Maybe Name) getHoleName tm = do defs <- get Ctxt NApp _ (NMeta n' i args) _ <- nf defs [] tm diff --git a/src/Core/UnifyState.idr b/src/Core/UnifyState.idr index 84744155b9..dde413779f 100644 --- a/src/Core/UnifyState.idr +++ b/src/Core/UnifyState.idr @@ -11,6 +11,7 @@ import Core.Options import Core.TT import Core.TTC import Core.Value +import Core.Name.ScopedList import Data.List import Libraries.Data.IntMap @@ -316,15 +317,15 @@ addPolyConstraint fc env arg x@(NApp _ (NMeta _ _ _) _) y addPolyConstraint fc env arg x y = pure () -mkLocal : {wkns : SnocList Name} -> FC -> Binder (Term vars) -> Term (wkns <>> x :: (vars ++ done)) +mkLocal : {wkns : SnocList Name} -> FC -> Binder (Term vars) -> Term (wkns <>> x :%: (vars +%+ done)) mkLocal fc b = Local fc (Just (isLet b)) _ (mkIsVarChiply (mkHasLength wkns)) mkConstantAppArgs : {vars : _} -> Bool -> FC -> Env Term vars -> (wkns : SnocList Name) -> - List (Term (wkns <>> (vars ++ done))) + List (Term (wkns <>> (vars +%+ done))) mkConstantAppArgs lets fc [] wkns = [] -mkConstantAppArgs {done} {vars = x :: xs} lets fc (b :: env) wkns +mkConstantAppArgs {done} {vars = x :%: xs} lets fc (b :: env) wkns = let rec = mkConstantAppArgs {done} lets fc env (wkns :< x) in if lets || not (isLet b) then mkLocal fc b :: rec @@ -334,15 +335,15 @@ mkConstantAppArgsSub : {vars : _} -> Bool -> FC -> Env Term vars -> Thin smaller vars -> (wkns : SnocList Name) -> - List (Term (wkns <>> (vars ++ done))) + List (Term (wkns <>> (vars +%+ done))) mkConstantAppArgsSub lets fc [] p wkns = [] -mkConstantAppArgsSub {done} {vars = x :: xs} +mkConstantAppArgsSub {done} {vars = x :%: xs} lets fc (b :: env) Refl wkns = mkConstantAppArgs lets fc env (wkns :< x) -mkConstantAppArgsSub {done} {vars = x :: xs} +mkConstantAppArgsSub {done} {vars = x :%: xs} lets fc (b :: env) (Drop p) wkns = mkConstantAppArgsSub lets fc env p (wkns :< x) -mkConstantAppArgsSub {done} {vars = x :: xs} +mkConstantAppArgsSub {done} {vars = x :%: xs} lets fc (b :: env) (Keep p) wkns = let rec = mkConstantAppArgsSub {done} lets fc env p (wkns :< x) in if lets || not (isLet b) @@ -353,15 +354,15 @@ mkConstantAppArgsOthers : {vars : _} -> Bool -> FC -> Env Term vars -> Thin smaller vars -> (wkns : SnocList Name) -> - List (Term (wkns <>> (vars ++ done))) + List (Term (wkns <>> (vars +%+ done))) mkConstantAppArgsOthers lets fc [] p wkns = [] -mkConstantAppArgsOthers {done} {vars = x :: xs} +mkConstantAppArgsOthers {done} {vars = x :%: xs} lets fc (b :: env) Refl wkns = mkConstantAppArgsOthers lets fc env Refl (wkns :< x) -mkConstantAppArgsOthers {done} {vars = x :: xs} +mkConstantAppArgsOthers {done} {vars = x :%: xs} lets fc (b :: env) (Keep p) wkns = mkConstantAppArgsOthers lets fc env p (wkns :< x) -mkConstantAppArgsOthers {done} {vars = x :: xs} +mkConstantAppArgsOthers {done} {vars = x :%: xs} lets fc (b :: env) (Drop p) wkns = let rec = mkConstantAppArgsOthers {done} lets fc env p (wkns :< x) in if lets || not (isLet b) @@ -372,14 +373,14 @@ export applyTo : {vars : _} -> FC -> Term vars -> Env Term vars -> Term vars applyTo fc tm env - = let args = reverse (mkConstantAppArgs {done = []} False fc env [<]) in + = let args = reverse (mkConstantAppArgs {done = SLNil} False fc env [<]) in apply fc tm (rewrite sym (appendNilRightNeutral vars) in args) export applyToFull : {vars : _} -> FC -> Term vars -> Env Term vars -> Term vars applyToFull fc tm env - = let args = reverse (mkConstantAppArgs {done = []} True fc env [<]) in + = let args = reverse (mkConstantAppArgs {done = SLNil} True fc env [<]) in apply fc tm (rewrite sym (appendNilRightNeutral vars) in args) export @@ -387,7 +388,7 @@ applyToSub : {vars : _} -> FC -> Term vars -> Env Term vars -> Thin smaller vars -> Term vars applyToSub fc tm env sub - = let args = reverse (mkConstantAppArgsSub {done = []} True fc env sub [<]) in + = let args = reverse (mkConstantAppArgsSub {done = SLNil} True fc env sub [<]) in apply fc tm (rewrite sym (appendNilRightNeutral vars) in args) export @@ -395,7 +396,7 @@ applyToOthers : {vars : _} -> FC -> Term vars -> Env Term vars -> Thin smaller vars -> Term vars applyToOthers fc tm env sub - = let args = reverse (mkConstantAppArgsOthers {done = []} True fc env sub [<]) in + = let args = reverse (mkConstantAppArgsOthers {done = SLNil} True fc env sub [<]) in apply fc tm (rewrite sym (appendNilRightNeutral vars) in args) -- Create a new metavariable with the given name and return type, @@ -415,7 +416,7 @@ newMetaLets {vars} fc rig env n ty def nocyc lets = do let hty = if lets then abstractFullEnvType fc env ty else abstractEnvType fc env ty let hole = { noCycles := nocyc } - (newDef fc n rig [] hty (specified Public) def) + (newDef fc n rig SLNil hty (specified Public) def) log "unify.meta" 5 $ "Adding new meta " ++ show (n, fc, rig) logTerm "unify.meta" 10 ("New meta type " ++ show n) hty idx <- addDef n hole @@ -423,7 +424,7 @@ newMetaLets {vars} fc rig env n ty def nocyc lets pure (idx, Meta fc n idx envArgs) where envArgs : List (Term vars) - envArgs = let args = reverse (mkConstantAppArgs {done = []} lets fc env [<]) in + envArgs = let args = reverse (mkConstantAppArgs {done = SLNil} lets fc env [<]) in rewrite sym (appendNilRightNeutral vars) in args export @@ -441,7 +442,7 @@ mkConstant : {vars : _} -> mkConstant fc [] tm = tm -- mkConstant {vars = x :: _} fc (Let c val ty :: env) tm -- = mkConstant fc env (Bind fc x (Let c val ty) tm) -mkConstant {vars = x :: _} fc (b :: env) tm +mkConstant {vars = x :%: _} fc (b :: env) tm = let ty = binderType b in mkConstant fc env (Bind fc x (Lam fc (multiplicity b) Explicit ty) tm) @@ -460,7 +461,7 @@ newConstant {vars} fc rig env tm ty constrs = do let def = mkConstant fc env tm let defty = abstractFullEnvType fc env ty cn <- genName "postpone" - let guess = newDef fc cn rig [] defty (specified Public) + let guess = newDef fc cn rig SLNil defty (specified Public) (Guess def (length env) constrs) log "unify.constant" 5 $ "Adding new constant " ++ show (cn, fc, rig) logTerm "unify.constant" 10 ("New constant type " ++ show cn) defty @@ -469,7 +470,7 @@ newConstant {vars} fc rig env tm ty constrs pure (Meta fc cn idx envArgs) where envArgs : List (Term vars) - envArgs = let args = reverse (mkConstantAppArgs {done = []} True fc env [<]) in + envArgs = let args = reverse (mkConstantAppArgs {done = SLNil} True fc env [<]) in rewrite sym (appendNilRightNeutral vars) in args -- Create a new search with the given name and return type, @@ -483,7 +484,7 @@ newSearch : {vars : _} -> Env Term vars -> Name -> Term vars -> Core (Int, Term vars) newSearch {vars} fc rig depth def env n ty = do let hty = abstractEnvType fc env ty - let hole = newDef fc n rig [] hty (specified Public) (BySearch rig depth def) + let hole = newDef fc n rig SLNil hty (specified Public) (BySearch rig depth def) log "unify.search" 10 $ "Adding new search " ++ show fc ++ " " ++ show n logTermNF "unify.search" 10 "New search type" [] hty idx <- addDef n hole @@ -491,7 +492,7 @@ newSearch {vars} fc rig depth def env n ty pure (idx, Meta fc n idx envArgs) where envArgs : List (Term vars) - envArgs = let args = reverse (mkConstantAppArgs {done = []} False fc env [<]) in + envArgs = let args = reverse (mkConstantAppArgs {done = SLNil} False fc env [<]) in rewrite sym (appendNilRightNeutral vars) in args -- Add a hole which stands for a delayed elaborator @@ -504,14 +505,14 @@ newDelayed : {vars : _} -> (ty : Term vars) -> Core (Int, Term vars) newDelayed {vars} fc rig env n ty = do let hty = abstractEnvType fc env ty - let hole = newDef fc n rig [] hty (specified Public) Delayed + let hole = newDef fc n rig SLNil hty (specified Public) Delayed idx <- addDef n hole log "unify.delay" 10 $ "Added delayed elaborator " ++ show (n, idx) addHoleName fc n idx pure (idx, Meta fc n idx envArgs) where envArgs : List (Term vars) - envArgs = let args = reverse (mkConstantAppArgs {done = []} False fc env [<]) in + envArgs = let args = reverse (mkConstantAppArgs {done = SLNil} False fc env [<]) in rewrite sym (appendNilRightNeutral vars) in args export @@ -618,15 +619,15 @@ checkUserHolesAfter : {auto u : Ref UST UState} -> Int -> Bool -> Core () checkUserHolesAfter base now = do gs_map <- getGuesses - let gs = toList gs_map + let gs = toScopedList gs_map log "unify.unsolved" 10 $ "Unsolved guesses " ++ show gs - traverse_ (checkValidHole base) gs + Core.Core.ScopedList.traverse_ (checkValidHole base) gs hs_map <- getCurrentHoles - let hs = toList hs_map + let hs = toScopedList hs_map let hs' = if any isUserName (map (snd . snd) hs) - then [] else hs + then SLNil else hs when (now && not (isNil hs')) $ - throw (UnsolvedHoles (map snd (nubBy nameEq hs))) + throw (UnsolvedHoles (map snd (nubBy nameEq $ toList hs))) -- Note the hole names, to ensure they are resolved -- by the end of elaborating the current source file traverse_ addDelayedHoleName hs' diff --git a/src/Core/Value.idr b/src/Core/Value.idr index 640e3ad169..2225ba46ad 100644 --- a/src/Core/Value.idr +++ b/src/Core/Value.idr @@ -4,6 +4,7 @@ import Core.Context import Core.Core import Core.Env import Core.TT +import Core.Name.ScopedList %default covering @@ -57,46 +58,46 @@ cbv = { strategy := CBV } defaultOpts mutual public export - data LocalEnv : List Name -> List Name -> Type where - Nil : LocalEnv free [] - (::) : Closure free -> LocalEnv free vars -> LocalEnv free (x :: vars) + data LocalEnv : ScopedList Name -> ScopedList Name -> Type where + Nil : LocalEnv free SLNil + (::) : Closure free -> LocalEnv free vars -> LocalEnv free (x :%: vars) public export - data Closure : List Name -> Type where + data Closure : ScopedList Name -> Type where MkClosure : {vars : _} -> (opts : EvalOpts) -> LocalEnv free vars -> Env Term free -> - Term (vars ++ free) -> Closure free + Term (vars +%+ free) -> Closure free MkNFClosure : EvalOpts -> Env Term free -> NF free -> Closure free -- The head of a value: things you can apply arguments to public export - data NHead : List Name -> Type where + data NHead : ScopedList Name -> Type where NLocal : Maybe Bool -> (idx : Nat) -> (0 p : IsVar nm idx vars) -> NHead vars NRef : NameType -> Name -> NHead vars - NMeta : Name -> Int -> List (Closure vars) -> NHead vars + NMeta : Name -> Int -> ScopedList (Closure vars) -> NHead vars -- Values themselves. 'Closure' is an unevaluated thunk, which means -- we can wait until necessary to reduce constructor arguments public export - data NF : List Name -> Type where + data NF : ScopedList Name -> Type where NBind : FC -> (x : Name) -> Binder (Closure vars) -> (Defs -> Closure vars -> Core (NF vars)) -> NF vars -- Each closure is associated with the file context of the App node that -- had it as an argument. It's necessary so as to not lose file context -- information when creating the normal form. - NApp : FC -> NHead vars -> List (FC, Closure vars) -> NF vars + NApp : FC -> NHead vars -> ScopedList (FC, Closure vars) -> NF vars NDCon : FC -> Name -> (tag : Int) -> (arity : Nat) -> - List (FC, Closure vars) -> NF vars + ScopedList (FC, Closure vars) -> NF vars NTCon : FC -> Name -> (tag : Int) -> (arity : Nat) -> - List (FC, Closure vars) -> NF vars + ScopedList (FC, Closure vars) -> NF vars NAs : FC -> UseSide -> NF vars -> NF vars -> NF vars NDelayed : FC -> LazyReason -> NF vars -> NF vars NDelay : FC -> LazyReason -> Closure vars -> Closure vars -> NF vars - NForce : FC -> LazyReason -> NF vars -> List (FC, Closure vars) -> NF vars + NForce : FC -> LazyReason -> NF vars -> ScopedList (FC, Closure vars) -> NF vars NPrimVal : FC -> Constant -> NF vars NErased : FC -> WhyErased (NF vars) -> NF vars NType : FC -> Name -> NF vars @@ -107,13 +108,13 @@ mutual %name NF nf export -ntCon : FC -> Name -> Int -> Nat -> List (FC, Closure vars) -> NF vars +ntCon : FC -> Name -> Int -> Nat -> ScopedList (FC, Closure vars) -> NF vars -- Part of the machinery for matching on types - I believe this won't affect -- universe checking so put a dummy name. -ntCon fc (UN (Basic "Type")) tag Z [] = NType fc (MN "top" 0) -ntCon fc n tag Z [] = case isConstantType n of +ntCon fc (UN (Basic "Type")) tag Z SLNil = NType fc (MN "top" 0) +ntCon fc n tag Z SLNil = case isConstantType n of Just c => NPrimVal fc $ PrT c - Nothing => NTCon fc n tag Z [] + Nothing => NTCon fc n tag Z SLNil ntCon fc n tag arity args = NTCon fc n tag arity args export diff --git a/src/Idris/Doc/String.idr b/src/Idris/Doc/String.idr index 4b99d76297..6bb78c5b6b 100644 --- a/src/Idris/Doc/String.idr +++ b/src/Idris/Doc/String.idr @@ -96,7 +96,7 @@ prettyType syn ty = do ||| Look up implementations getImplDocs : {auto c : Ref Ctxt Defs} -> {auto s : Ref Syn SyntaxInfo} -> - (keep : Term [] -> Core Bool) -> + (keep : Term SLNil -> Core Bool) -> Core (List (Doc IdrisDocAnn)) getImplDocs keep = do defs <- get Ctxt diff --git a/src/Idris/Elab/Implementation.idr b/src/Idris/Elab/Implementation.idr index c55b50e788..b4e0c2950d 100644 --- a/src/Idris/Elab/Implementation.idr +++ b/src/Idris/Elab/Implementation.idr @@ -139,6 +139,7 @@ elabImplementation {vars} ifc vis opts_in pass env nest is cons iname ps named i -- alias for something syn <- get Syn defs <- get Ctxt + let varsList = toList vars inames <- lookupCtxtName iname (gamma defs) let [cndata] = concatMap (\n => lookupName n (ifaces syn)) (map fst inames) @@ -177,7 +178,7 @@ elabImplementation {vars} ifc vis opts_in pass env nest is cons iname ps named i let initTy = bindImpls is $ bindConstraints vfc AutoImplicit cons (apply (IVar vfc iname) ps) let paramBinds = if !isUnboundImplicits - then findBindableNames True vars [] initTy + then findBindableNames True varsList [] initTy else [] let impTy = doBind paramBinds initTy @@ -403,26 +404,27 @@ elabImplementation {vars} ifc vis opts_in pass env nest is cons iname ps named i = do -- Get the specialised type by applying the method to the -- parameters n <- inCurrentNS (methName meth.name) + let varsList = toList vars -- Avoid any name clashes between parameters and method types by -- renaming IBindVars in the method types which appear in the -- parameters let upds' = !(traverse (applyCon impName) allmeths) - let mty_in = substNames vars upds' meth.type + let mty_in = substNames varsList upds' meth.type let (upds, mty_in) = runState Prelude.Nil (renameIBinds impsp (findImplicits mty_in) mty_in) -- Finally update the method type so that implicits from the -- parameters are passed through to any earlier methods which -- appear in the type - let mty_in = substNames vars methupds mty_in + let mty_in = substNames varsList methupds mty_in -- Substitute _ in for the implicit parameters, then -- substitute in the explicit parameters. let mty_iparams - = substBindVars vars + = substBindVars varsList (map (\n => (n, Implicit vfc False)) imppnames) mty_in let mty_params - = substNames vars (zip pnames ps) mty_iparams + = substNames varsList (zip pnames ps) mty_iparams log "elab.implementation" 5 $ "Substitute implicits " ++ show imppnames ++ " parameters " ++ show (zip pnames ps) ++ " " ++ show mty_in ++ " is " ++ @@ -433,7 +435,7 @@ elabImplementation {vars} ifc vis opts_in pass env nest is cons iname ps named i mty_params let ibound = findImplicits mbase - mty <- bindTypeNamesUsed ifc ibound vars mbase + mty <- bindTypeNamesUsed ifc ibound varsList mbase log "elab.implementation" 3 $ "Method " ++ show meth.name ++ " ==> " ++ diff --git a/src/Idris/Elab/Interface.idr b/src/Idris/Elab/Interface.idr index d90ada1770..e7e6b3a398 100644 --- a/src/Idris/Elab/Interface.idr +++ b/src/Idris/Elab/Interface.idr @@ -117,8 +117,8 @@ mkIfaceData {vars} ifc def_vis env constraints n conName ps dets meths retty = apply (IVar vfc n) (map (IVar EmptyFC) pNames) conty = mkTy Implicit (map jname ps) $ mkTy AutoImplicit (map bhere constraints) (mkTy Explicit (map bname meths) retty) - con = MkImpTy vfc EmptyFC conName !(bindTypeNames ifc [] (pNames ++ map fst meths ++ vars) conty) - bound = pNames ++ map fst meths ++ vars in + con = MkImpTy vfc EmptyFC conName !(bindTypeNames ifc [] (pNames ++ map fst meths ++ toList vars) conty) + bound = pNames ++ map fst meths ++ toList vars in pure $ IData vfc def_vis Nothing {- ?? -} $ MkImpData vfc n @@ -155,7 +155,7 @@ getMethDecl : {vars : _} -> Core (nm, RigCount, RawImp) getMethDecl {vars} env nest params mnames (c, nm, ty) = do let paramNames = map fst params - ty_imp <- bindTypeNames EmptyFC [] (paramNames ++ mnames ++ vars) ty + ty_imp <- bindTypeNames EmptyFC [] (paramNames ++ mnames ++ toList vars) ty pure (nm, c, stripParams paramNames ty_imp) where -- We don't want the parameters to explicitly appear in the method @@ -189,8 +189,8 @@ getMethToplevel {vars} env vis iname cname constraints allmeths params sig -- Make the constraint application explicit for any method names -- which appear in other method types let ty_constr = - substNames vars (map applyCon allmeths) sig.type - ty_imp <- bindTypeNames EmptyFC [] vars (bindPs params $ bindIFace vfc ity ty_constr) + substNames (toList vars) (map applyCon allmeths) sig.type + ty_imp <- bindTypeNames EmptyFC [] (toList vars) (bindPs params $ bindIFace vfc ity ty_constr) cn <- inCurrentNS sig.name let tydecl = IClaim vfc sig.count vis (if sig.isData then [Inline, Invertible] else [Inline]) @@ -254,7 +254,7 @@ getConstraintHint : {vars : _} -> getConstraintHint {vars} fc env vis iname cname constraints meths params (cn, con) = do let ity = apply (IVar fc iname) (map (IVar fc) params) let fty = IPi fc top Explicit Nothing ity con - ty_imp <- bindTypeNames fc [] (meths ++ vars) fty + ty_imp <- bindTypeNames fc [] (meths ++ toList vars) fty let hintname = DN ("Constraint " ++ show con) (UN (Basic $ "__" ++ show iname ++ "_" ++ show con)) @@ -449,9 +449,9 @@ elabInterface {vars} ifc def_vis env nest constraints iname params dets mcon bod tydecls let dty = bindPs params -- bind parameters $ bindIFace vdfc ity -- bind interface (?!) - $ substNames vars methNameMap dty + $ substNames (toList vars) methNameMap dty - dty_imp <- bindTypeNames dfc [] (map name tydecls ++ vars) dty + dty_imp <- bindTypeNames dfc [] (map name tydecls ++ toList vars) dty log "elab.interface.default" 5 $ "Default method " ++ show dn ++ " : " ++ show dty_imp let dtydecl = IClaim vdfc rig (collapseDefault def_vis) [] (MkImpTy EmptyFC EmptyFC dn dty_imp) @@ -524,6 +524,6 @@ elabInterface {vars} ifc def_vis env nest constraints iname params dets mcon bod meth_names paramNames) nconstraints log "elab.interface" 5 $ "Constraint hints from " ++ show constraints ++ ": " ++ show chints - traverse_ (processDecl [] nest env) (concatMap snd chints) + Core.Core.traverse_ (processDecl [] nest env) (concatMap snd chints) traverse_ (\n => do mn <- inCurrentNS n setFlag vfc mn TCInline) (map fst chints) diff --git a/src/Idris/Error.idr b/src/Idris/Error.idr index ca30b4546e..be8b1c2476 100644 --- a/src/Idris/Error.idr +++ b/src/Idris/Error.idr @@ -392,7 +392,7 @@ perrorRaw (NotCovering fc n (MissingCases cs)) = pure $ errorDesc (code (pretty0 !(prettyName n)) <++> reflow "is not covering.") <+> line <+> !(ploc fc) <+> line <+> reflow "Missing cases" <+> colon <+> line - <+> indent 4 (vsep !(traverse (pshow []) cs)) <+> line + <+> indent 4 (vsep !(traverse (pshow []) (toList cs))) <+> line perrorRaw (NotCovering fc n (NonCoveringCall ns)) = pure $ errorDesc (pretty0 !(prettyName n) <++> reflow "is not covering.") <+> line <+> !(ploc fc) <+> line diff --git a/src/Idris/Parser.idr b/src/Idris/Parser.idr index fd941c347a..73cd5c0023 100644 --- a/src/Idris/Parser.idr +++ b/src/Idris/Parser.idr @@ -1125,7 +1125,7 @@ mutual let fc = boundToFC fname x in toLines xs [< StrLiteral fc (last strs)] $ acc :< (line <>> [StrLiteral fc str]) - <>< map (\str => [StrLiteral fc str]) (init strs) + <>< (the (List _) $ map (\str => [StrLiteral fc str]) (init strs)) fnDirectOpt : OriginDesc -> Rule PFnOpt fnDirectOpt fname diff --git a/src/Idris/REPL.idr b/src/Idris/REPL.idr index a343926d1b..65935be2da 100644 --- a/src/Idris/REPL.idr +++ b/src/Idris/REPL.idr @@ -386,7 +386,7 @@ findInTree p hint m match : (NonEmptyFC, Name) -> Bool match (_, n) = matches hint n && checkCandidate n -record TermWithType (vars : List Name) where +record TermWithType (vars : ScopedList Name) where constructor WithType termOf : Term vars typeOf : Term vars @@ -416,7 +416,7 @@ inferAndElab : Env Term vars -> Core (TermWithType vars) inferAndElab emode itm env - = do ttimp <- desugar AnyExpr vars itm + = do ttimp <- desugar AnyExpr (toList vars) itm let ttimpWithIt = ILocal replFC !getItDecls ttimp inidx <- resolveName (UN $ Basic "[input]") -- a TMP HACK to prioritise list syntax for List: hide @@ -628,7 +628,7 @@ processEdit (ExprSearch upd line name hints) if upd then updateFile (proofSearch name (show itm') (integerToNat (cast (line - 1)))) else pure $ DisplayEdit (prettyBy Syntax itm') - [(n, nidx, PMDef pi [] (STerm _ tm) _ _)] => + [(n, nidx, PMDef pi SLNil (STerm _ tm) _ _)] => case holeInfo pi of NotHole => pure $ EditError "Not a searchable hole" SolvedHole locs => @@ -854,7 +854,7 @@ inferAndNormalize : {auto c : Ref Ctxt Defs} -> {auto o : Ref ROpts REPLOpts} -> REPLEval -> PTerm -> - Core (TermWithType []) + Core (TermWithType SLNil) inferAndNormalize emode itm = do (tm `WithType` ty) <- inferAndElab (elabMode emode) itm [] logTerm "repl.eval" 10 "Elaborated input" tm @@ -902,8 +902,8 @@ process (Eval itm) let norm = replEval emode evalResultName <- DN "it" <$> genName "evalResult" ignore $ addDef evalResultName - $ newDef replFC evalResultName top [] ty defaulted - $ PMDef defaultPI [] (STerm 0 ntm) (STerm 0 ntm) [] + $ newDef replFC evalResultName top SLNil ty defaulted + $ PMDef defaultPI SLNil (STerm 0 ntm) (STerm 0 ntm) [] addToSave evalResultName put ROpts ({ evalResultName := Just evalResultName } opts) if showTypes opts diff --git a/src/TTImp/Elab/Ambiguity.idr b/src/TTImp/Elab/Ambiguity.idr index 41932f8eff..304a284694 100644 --- a/src/TTImp/Elab/Ambiguity.idr +++ b/src/TTImp/Elab/Ambiguity.idr @@ -190,14 +190,14 @@ Show TypeMatch where mutual mightMatchD : {auto c : Ref Ctxt Defs} -> {vars : _} -> - Defs -> NF vars -> NF [] -> Core TypeMatch + Defs -> NF vars -> NF SLNil -> Core TypeMatch mightMatchD defs l r = mightMatch defs (stripDelay l) (stripDelay r) mightMatchArg : {auto c : Ref Ctxt Defs} -> {vars : _} -> Defs -> - Closure vars -> Closure [] -> + Closure vars -> Closure SLNil -> Core Bool mightMatchArg defs l r = pure $ case !(mightMatchD defs !(evalClosure defs l) !(evalClosure defs r)) of @@ -207,10 +207,10 @@ mutual mightMatchArgs : {auto c : Ref Ctxt Defs} -> {vars : _} -> Defs -> - List (Closure vars) -> List (Closure []) -> + ScopedList (Closure vars) -> ScopedList (Closure SLNil) -> Core Bool - mightMatchArgs defs [] [] = pure True - mightMatchArgs defs (x :: xs) (y :: ys) + mightMatchArgs defs SLNil SLNil = pure True + mightMatchArgs defs (x :%: xs) (y :%: ys) = do amatch <- mightMatchArg defs x y if amatch then mightMatchArgs defs xs ys @@ -219,7 +219,7 @@ mutual mightMatch : {auto c : Ref Ctxt Defs} -> {vars : _} -> - Defs -> NF vars -> NF [] -> Core TypeMatch + Defs -> NF vars -> NF SLNil -> Core TypeMatch mightMatch defs target (NBind fc n (Pi _ _ _ _) sc) = mightMatchD defs target !(sc defs (toClosure defaultOpts [] (Erased fc Placeholder))) mightMatch defs (NBind _ _ _ _) (NBind _ _ _ _) = pure Poly -- lambdas might match diff --git a/src/TTImp/Elab/App.idr b/src/TTImp/Elab/App.idr index ede3102048..88da7b3e9f 100644 --- a/src/TTImp/Elab/App.idr +++ b/src/TTImp/Elab/App.idr @@ -854,7 +854,7 @@ checkApp rig elabinfo nest env fc (IVar fc' n) expargs autoargs namedargs exp = do tm <- Normalise.normalisePrims (`boundSafe` elabMode elabinfo) isIPrimVal (onLHS (elabMode elabinfo)) - prims n expargs (fst res) env + prims n (fromList expargs) (fst res) env pure (fromMaybe (fst res) tm, snd res) where diff --git a/src/TTImp/Elab/Binders.idr b/src/TTImp/Elab/Binders.idr index 9f191e4282..578a8bb777 100644 --- a/src/TTImp/Elab/Binders.idr +++ b/src/TTImp/Elab/Binders.idr @@ -70,7 +70,7 @@ checkPi rig elabinfo nest env fc rigf info n argTy retTy expTy (tyv, tyt) <- check pirig elabinfo nest env argTy (Just (gType fc tyu)) info' <- checkPiInfo rigf elabinfo nest env info (Just (gnf env tyv)) - let env' : Env Term (n :: _) = Pi fc rigf info' tyv :: env + let env' : Env Term (n :%: _) = Pi fc rigf info' tyv :: env let nest' = weaken (dropName n nest) scu <- uniVar fc (scopev, scopet) <- @@ -115,7 +115,7 @@ inferLambda rig elabinfo nest env fc rigl info n argTy scope expTy u <- uniVar fc (tyv, tyt) <- check erased elabinfo nest env argTy (Just (gType fc u)) info' <- checkPiInfo rigl elabinfo nest env info (Just (gnf env tyv)) - let env' : Env Term (n :: _) = Lam fc rigb info' tyv :: env + let env' : Env Term (n :%: _) = Lam fc rigb info' tyv :: env let nest' = weaken (dropName n nest) (scopev, scopet) <- inScope fc env' (\e' => check {e=e'} rig elabinfo @@ -172,7 +172,7 @@ checkLambda rig_in elabinfo nest env fc rigl info n argTy scope (Just expty_in) argTy (Just (gType fc u)) info' <- checkPiInfo rigl elabinfo nest env info (Just (gnf env tyv)) let rigb = rigl `glb` c - let env' : Env Term (n :: _) = Lam fc rigb info' tyv :: env + let env' : Env Term (n :%: _) = Lam fc rigb info' tyv :: env ignore $ convert fc elabinfo env (gnf env tyv) (gnf env pty) let nest' = weaken (dropName n nest) (scopev, scopet) <- @@ -199,8 +199,8 @@ checkLambda rig_in elabinfo nest env fc rigl info n argTy scope (Just expty_in) _ => inferLambda rig elabinfo nest env fc rigl info n argTy scope (Just expty_in) weakenExp : {x, vars : _} -> - Env Term (x :: vars) -> - Maybe (Glued vars) -> Core (Maybe (Glued (x :: vars))) + Env Term (x :%: vars) -> + Maybe (Glued vars) -> Core (Maybe (Glued (x :%: vars))) weakenExp env Nothing = pure Nothing weakenExp env (Just gtm) = do tm <- getTerm gtm @@ -248,7 +248,7 @@ checkLet rigc_in elabinfo nest env fc lhsFC rigl n nTy nVal scope expty {vars} elabinfo -- without preciseInf nest env nVal (Just (gnf env tyv)) pure (fst c, snd c, rigl |*| rigc)) - let env' : Env Term (n :: _) = Lam fc rigb Explicit tyv :: env + let env' : Env Term (n :%: _) = Lam fc rigb Explicit tyv :: env let nest' = weaken (dropName n nest) expScope <- weakenExp env' expty (scopev, gscopet) <- diff --git a/src/TTImp/Elab/Case.idr b/src/TTImp/Elab/Case.idr index 949df82749..54ab768010 100644 --- a/src/TTImp/Elab/Case.idr +++ b/src/TTImp/Elab/Case.idr @@ -53,9 +53,9 @@ changeVar old new (TForce fc r p) = TForce fc r (changeVar old new p) changeVar old new tm = tm -findLater : (x : Name) -> (newer : List Name) -> Var (newer ++ x :: older) -findLater x [] = MkVar First -findLater {older} x (_ :: xs) +findLater : (x : Name) -> (newer : ScopedList Name) -> Var (newer +%+ x :%: older) +findLater x SLNil = MkVar First +findLater {older} x (_ :%: xs) = let MkVar p = findLater {older} x xs in MkVar (Later p) @@ -101,10 +101,10 @@ findImpsIn fc env ns ty = when (not (isNil ns)) $ throw (TryWithImplicits fc env (reverse ns)) -merge : {vs : List Name} -> - List (Var vs) -> List (Var vs) -> List (Var vs) -merge [] xs = xs -merge (v :: vs) xs +merge : {vs : ScopedList Name} -> + ScopedList (Var vs) -> List (Var vs) -> List (Var vs) +merge SLNil xs = xs +merge (v :%: vs) xs = merge vs (v :: filter (v /=) xs) -- Extend the list of variables we need in the environment so far, removing @@ -121,7 +121,7 @@ extendNeeded b env needed findScrutinee : {vs : _} -> Env Term vs -> RawImp -> Maybe (Var vs) -findScrutinee {vs = n' :: _} (b :: bs) (IVar loc' n) +findScrutinee {vs = n' :%: _} (b :: bs) (IVar loc' n) = if n' == n && not (isLet b) then Just (MkVar First) else do MkVar p <- findScrutinee bs (IVar loc' n) @@ -233,7 +233,7 @@ caseBlock {vars} rigc elabinfo fc nest env opts scr scrtm scrty caseRig alts exp when (not (isNil fullImps)) $ findImpsIn fc [] [] casefnty cidx <- addDef casen ({ eraseArgs := erasedargs } (newDef fc casen (if isErased rigc then erased else top) - [] casefnty vis None)) + SLNil casefnty vis None)) traverse_ (processFnOpt fc False casen) opts @@ -306,7 +306,7 @@ caseBlock {vars} rigc elabinfo fc nest env opts scr scrtm scrty caseRig alts exp addEnv : {vs : _} -> Int -> Env Term vs -> List Name -> (List (Name, Name), List RawImp) addEnv idx [] used = ([], []) - addEnv idx {vs = v :: vs} (b :: bs) used + addEnv idx {vs = v :%: vs} (b :: bs) used = let n = getBindName idx v used (ns, rest) = addEnv (idx + 1) bs (snd n :: used) ns' = n :: ns in @@ -431,7 +431,7 @@ checkCase rig elabinfo nest env fc opts scr scrty_in alts exp = throw (GenericMsg fc "Can't infer type for case scrutinee") checkConcrete _ = pure () - applyTo : Defs -> RawImp -> NF [] -> Core RawImp + applyTo : Defs -> RawImp -> NF SLNil -> Core RawImp applyTo defs ty (NBind fc _ (Pi _ _ Explicit _) sc) = applyTo defs (IApp fc ty (Implicit fc False)) !(sc defs (toClosure defaultOpts [] (Erased fc Placeholder))) @@ -441,7 +441,7 @@ checkCase rig elabinfo nest env fc opts scr scrty_in alts exp applyTo defs ty _ = pure ty -- Get the name and type of the family the scrutinee is in - getRetTy : Defs -> NF [] -> Core (Maybe (Name, NF [])) + getRetTy : Defs -> NF SLNil -> Core (Maybe (Name, NF SLNil)) getRetTy defs (NBind fc _ (Pi _ _ _ _) sc) = getRetTy defs !(sc defs (toClosure defaultOpts [] (Erased fc Placeholder))) getRetTy defs (NTCon _ n _ arity _) diff --git a/src/TTImp/Elab/Check.idr b/src/TTImp/Elab/Check.idr index 45698c5fbe..4ceadfb533 100644 --- a/src/TTImp/Elab/Check.idr +++ b/src/TTImp/Elab/Check.idr @@ -62,7 +62,7 @@ Eq ElabOpt where -- Descriptions of implicit name bindings. They're either just the name, -- or a binding of an @-pattern which has an associated pattern. public export -data ImplBinding : List Name -> Type where +data ImplBinding : ScopedList Name -> Type where NameBinding : {vars : _} -> FC -> RigCount -> PiInfo (Term vars) -> (elabAs : Term vars) -> (expTy : Term vars) -> @@ -113,9 +113,9 @@ bindingPiInfo (AsBinding _ p _ _ _) = p -- Current elaboration state (preserved/updated throughout elaboration) public export -record EState (vars : List Name) where +record EState (vars : ScopedList Name) where constructor MkEState - {outer : List Name} + {outer : ScopedList Name} -- The function/constructor name we're currently working on (resolved id) defining : Int -- The outer environment in which we're running the elaborator. Things here should @@ -190,7 +190,7 @@ saveHole n = update EST { saveHoles $= insert n () } weakenedEState : {n, vars : _} -> {auto e : Ref EST (EState vars)} -> - Core (Ref EST (EState (n :: vars))) + Core (Ref EST (EState (n :%: vars))) weakenedEState {e} = do est <- get EST eref <- newRef EST $ @@ -203,7 +203,7 @@ weakenedEState {e} pure eref where wknTms : (Name, ImplBinding vs) -> - (Name, ImplBinding (n :: vs)) + (Name, ImplBinding (n :%: vs)) wknTms (f, NameBinding fc c p x y) = (f, NameBinding fc c (map weaken p) (weaken x) (weaken y)) wknTms (f, AsBinding c p x y z) @@ -211,8 +211,8 @@ weakenedEState {e} strengthenedEState : {n, vars : _} -> Ref Ctxt Defs -> - Ref EST (EState (n :: vars)) -> - FC -> Env Term (n :: vars) -> + Ref EST (EState (n :%: vars)) -> + FC -> Env Term (n :%: vars) -> Core (EState vars) strengthenedEState {n} {vars} c e fc env = do est <- get EST @@ -228,7 +228,7 @@ strengthenedEState {n} {vars} c e fc env } est where - dropSub : Thin xs (y :: ys) -> Core (Thin xs ys) + dropSub : Thin xs (y :%: ys) -> Core (Thin xs ys) dropSub (Drop sub) = pure sub dropSub _ = throw (InternalError "Badly formed weakened environment") @@ -239,7 +239,7 @@ strengthenedEState {n} {vars} c e fc env -- never actualy *use* that hole - this process is only to ensure that the -- unbound implicit doesn't depend on any variables it doesn't have -- in scope. - removeArgVars : List (Term (n :: vs)) -> Maybe (List (Term vs)) + removeArgVars : List (Term (n :%: vs)) -> Maybe (List (Term vs)) removeArgVars [] = pure [] removeArgVars (Local fc r (S k) p :: args) = do args' <- removeArgVars args @@ -251,7 +251,7 @@ strengthenedEState {n} {vars} c e fc env args' <- removeArgVars args pure (a' :: args') - removeArg : Term (n :: vs) -> Maybe (Term vs) + removeArg : Term (n :%: vs) -> Maybe (Term vs) removeArg tm = case getFnArgs tm of (f, args) => @@ -259,7 +259,7 @@ strengthenedEState {n} {vars} c e fc env f' <- shrink f (Drop Refl) pure (apply (getLoc f) f' args') - strTms : Defs -> (Name, ImplBinding (n :: vars)) -> + strTms : Defs -> (Name, ImplBinding (n :%: vars)) -> Core (Name, ImplBinding vars) strTms defs (f, NameBinding fc c p x y) = do xnf <- normaliseHoles defs env x @@ -282,7 +282,7 @@ strengthenedEState {n} {vars} c e fc env pure (f, AsBinding c p' x' y' z') _ => throw (BadUnboundImplicit fc env f y) - dropTop : (Var (n :: vs)) -> Maybe (Var vs) + dropTop : (Var (n :%: vs)) -> Maybe (Var vs) dropTop (MkVar First) = Nothing dropTop (MkVar (Later p)) = Just (MkVar p) @@ -290,8 +290,8 @@ export inScope : {n, vars : _} -> {auto c : Ref Ctxt Defs} -> {auto e : Ref EST (EState vars)} -> - FC -> Env Term (n :: vars) -> - (Ref EST (EState (n :: vars)) -> Core a) -> + FC -> Env Term (n :%: vars) -> + (Ref EST (EState (n :%: vars)) -> Core a) -> Core a inScope {c} {e} fc env elab = do e' <- weakenedEState @@ -413,7 +413,7 @@ uniVar : {auto c : Ref Ctxt Defs} -> FC -> Core Name uniVar fc = do n <- genName "u" - idx <- addDef n (newDef fc n erased [] (Erased fc Placeholder) (specified Public) None) + idx <- addDef n (newDef fc n erased SLNil (Erased fc Placeholder) (specified Public) None) pure (Resolved idx) export @@ -452,8 +452,8 @@ searchVar fc rig depth def env nest n ty else find x xs envHints : List Name -> Env Term vars -> - Core (vars' ** (Term (vars' ++ vars) -> Term vars, Env Term (vars' ++ vars))) - envHints [] env = pure ([] ** (id, env)) + Core (vars' ** (Term (vars' +%+ vars) -> Term vars, Env Term (vars' +%+ vars))) + envHints [] env = pure (SLNil ** (id, env)) envHints (n :: ns) env = do (vs ** (f, env')) <- envHints ns env let Just (nestn, argns, tmf) = find !(toFullNames n) (names nest) @@ -468,7 +468,7 @@ searchVar fc rig depth def env nest n ty let binder = Let fc top (weakenNs (mkSizeOf vs) app) (weakenNs (mkSizeOf vs) tyenv) varn <- toFullNames n' - pure ((varn :: vs) ** + pure ((varn :%: vs) ** (\t => f (Bind fc varn binder t), binder :: env')) diff --git a/src/TTImp/Elab/Delayed.idr b/src/TTImp/Elab/Delayed.idr index 9d1257472a..dd534e20c9 100644 --- a/src/TTImp/Elab/Delayed.idr +++ b/src/TTImp/Elab/Delayed.idr @@ -28,7 +28,7 @@ mkClosedElab : {vars : _} -> mkClosedElab fc [] elab = do (tm, _) <- elab pure tm -mkClosedElab {vars = x :: vars} fc (b :: env) elab +mkClosedElab {vars = x :%: vars} fc (b :: env) elab = mkClosedElab fc env (do (sc', _) <- elab let b' = newBinder b @@ -163,11 +163,11 @@ mutual mismatchNF defs (NTCon _ xn xt _ xargs) (NTCon _ yn yt _ yargs) = if xn /= yn then pure True - else anyM (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) + else anyMScoped (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) mismatchNF defs (NDCon _ _ xt _ xargs) (NDCon _ _ yt _ yargs) = if xt /= yt then pure True - else anyM (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) + else anyMScoped (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) mismatchNF defs (NPrimVal _ xc) (NPrimVal _ yc) = pure (xc /= yc) mismatchNF defs (NDelayed _ _ x) (NDelayed _ _ y) = mismatchNF defs x y mismatchNF defs (NDelay _ _ _ x) (NDelay _ _ _ y) @@ -187,11 +187,11 @@ contra : {auto c : Ref Ctxt Defs} -> contra defs (NTCon _ xn xt xa xargs) (NTCon _ yn yt ya yargs) = if xn /= yn then pure True - else anyM (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) + else anyMScoped (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) contra defs (NDCon _ _ xt _ xargs) (NDCon _ _ yt _ yargs) = if xt /= yt then pure True - else anyM (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) + else anyMScoped (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) contra defs (NPrimVal _ x) (NPrimVal _ y) = pure (x /= y) contra defs (NDCon _ _ _ _ _) (NPrimVal _ _) = pure True contra defs (NPrimVal _ _) (NDCon _ _ _ _ _) = pure True @@ -258,7 +258,7 @@ retryDelayed' errmode p acc (d@(_, i, hints, elab) :: ds) updateDef (Resolved i) (const (Just (PMDef (MkPMDefInfo NotHole True False) - [] (STerm 0 tm) (STerm 0 tm) []))) + SLNil (STerm 0 tm) (STerm 0 tm) []))) logTerm "elab.update" 5 ("Resolved delayed hole " ++ show i) tm logTermNF "elab.update" 5 ("Resolved delayed hole NF " ++ show i) [] tm removeHole i diff --git a/src/TTImp/Elab/ImplicitBind.idr b/src/TTImp/Elab/ImplicitBind.idr index 2f173b5674..9cfae362d6 100644 --- a/src/TTImp/Elab/ImplicitBind.idr +++ b/src/TTImp/Elab/ImplicitBind.idr @@ -100,7 +100,7 @@ mkPatternHole {vars'} loc rig n topenv imode (Just expty_in) Env Term vs -> Term vs -> Thin newvars vs -> Maybe (Term newvars) bindInner env ty Refl = Just ty - bindInner {vs = x :: _} (b :: env) ty (Drop p) + bindInner {vs = x :%: _} (b :: env) ty (Drop p) = bindInner env (Bind loc x b ty) p bindInner _ _ _ = Nothing @@ -169,34 +169,34 @@ bindUnsolved {vars} fc elabmode _ _ => inTerm) fc env tm bindtm -swapIsVarH : {idx : Nat} -> (0 p : IsVar nm idx (x :: y :: xs)) -> - Var (y :: x :: xs) +swapIsVarH : {idx : Nat} -> (0 p : IsVar nm idx (x :%: y :%: xs)) -> + Var (y :%: x :%: xs) swapIsVarH First = MkVar (Later First) swapIsVarH (Later p) = swapP p -- it'd be nice to do this all at the top -- level, but that will need an improvement -- in erasability checking where - swapP : forall name . {idx : _} -> (0 p : IsVar name idx (y :: xs)) -> - Var (y :: x :: xs) + swapP : forall name . {idx : _} -> (0 p : IsVar name idx (y :%: xs)) -> + Var (y :%: x :%: xs) swapP First = MkVar First swapP (Later x) = MkVar (Later (Later x)) -swapIsVar : (vs : List Name) -> - {idx : Nat} -> (0 p : IsVar nm idx (vs ++ x :: y :: xs)) -> - Var (vs ++ y :: x :: xs) -swapIsVar [] prf = swapIsVarH prf -swapIsVar (x :: xs) First = MkVar First -swapIsVar (x :: xs) (Later p) +swapIsVar : (vs : ScopedList Name) -> + {idx : Nat} -> (0 p : IsVar nm idx (vs +%+ x :%: y :%: xs)) -> + Var (vs +%+ y :%: x :%: xs) +swapIsVar SLNil prf = swapIsVarH prf +swapIsVar (x :%: xs) First = MkVar First +swapIsVar (x :%: xs) (Later p) = let MkVar p' = swapIsVar xs p in MkVar (Later p') -swapVars : {vs : List Name} -> - Term (vs ++ x :: y :: ys) -> Term (vs ++ y :: x :: ys) +swapVars : {vs : ScopedList Name} -> + Term (vs +%+ x :%: y :%: ys) -> Term (vs +%+ y :%: x :%: ys) swapVars (Local fc x idx p) = let MkVar p' = swapIsVar _ p in Local fc x _ p' swapVars (Ref fc x name) = Ref fc x name swapVars (Meta fc n i xs) = Meta fc n i (map swapVars xs) swapVars {vs} (Bind fc x b scope) - = Bind fc x (map swapVars b) (swapVars {vs = x :: vs} scope) + = Bind fc x (map swapVars b) (swapVars {vs = x :%: vs} scope) swapVars (App fc fn arg) = App fc (swapVars fn) (swapVars arg) swapVars (As fc s nm pat) = As fc s (swapVars nm) (swapVars pat) swapVars (TDelayed fc x tm) = TDelayed fc x (swapVars tm) @@ -212,13 +212,13 @@ swapVars (TType fc u) = TType fc u -- move it under implicit binders that don't depend on it, and stop -- when hitting any non-implicit binder push : {vs : _} -> - FC -> (n : Name) -> Binder (Term vs) -> Term (n :: vs) -> Term vs + FC -> (n : Name) -> Binder (Term vs) -> Term (n :%: vs) -> Term vs push ofc n b tm@(Bind fc (PV x i) (Pi fc' c Implicit ty) sc) -- only push past 'PV's = case shrink ty (Drop Refl) of Nothing => -- needs explicit pi, do nothing Bind ofc n b tm Just ty' => Bind fc (PV x i) (Pi fc' c Implicit ty') - (push ofc n (map weaken b) (swapVars {vs = []} sc)) + (push ofc n (map weaken b) (swapVars {vs = SLNil} sc)) push ofc n b tm = Bind ofc n b tm -- Move any implicit arguments as far to the left as possible - this helps @@ -259,7 +259,7 @@ bindImplVars {vars} fc mode gam env imps_in scope scty getBinds : (imps : List (Name, Name, ImplBinding vs)) -> Bounds new -> (tm : Term vs) -> (ty : Term vs) -> - (Term (new ++ vs), Term (new ++ vs)) + (Term (new +%+ vs), Term (new +%+ vs)) getBinds [] bs tm ty = (refsToLocals bs tm, refsToLocals bs ty) getBinds {new} ((n, metan, NameBinding loc c p _ bty) :: imps) bs tm ty = let (tm', ty') = getBinds imps (Add n metan bs) tm ty diff --git a/src/TTImp/Elab/Record.idr b/src/TTImp/Elab/Record.idr index 822f6610d3..9b3fec65d4 100644 --- a/src/TTImp/Elab/Record.idr +++ b/src/TTImp/Elab/Record.idr @@ -26,12 +26,12 @@ getRecordType : Env Term vars -> NF vars -> Maybe Name getRecordType env (NTCon _ n _ _ _) = Just n getRecordType env _ = Nothing -getNames : {auto c : Ref Ctxt Defs} -> Defs -> NF [] -> Core $ SortedSet Name +getNames : {auto c : Ref Ctxt Defs} -> Defs -> NF SLNil -> Core $ SortedSet Name getNames defs (NApp _ hd args) = do eargs <- traverse (evalClosure defs . snd) args pure $ nheadNames hd `union` concat !(traverse (getNames defs) eargs) where - nheadNames : NHead [] -> SortedSet Name + nheadNames : NHead SLNil -> SortedSet Name nheadNames (NRef Bound n) = singleton n nheadNames _ = empty getNames defs (NDCon _ _ _ _ args) @@ -92,7 +92,7 @@ findFieldsAndTypeArgs defs con where getExpNames : SortedSet Name -> List (String, Maybe Name, Maybe Name) -> - NF [] -> + NF SLNil -> Core (List (String, Maybe Name, Maybe Name), SortedSet Name) getExpNames names expNames (NBind fc x (Pi _ _ p ty) sc) = do let imp = case p of diff --git a/src/TTImp/Elab/Rewrite.idr b/src/TTImp/Elab/Rewrite.idr index 3deb47cff2..8ccd10b706 100644 --- a/src/TTImp/Elab/Rewrite.idr +++ b/src/TTImp/Elab/Rewrite.idr @@ -37,7 +37,7 @@ getRewriteTerms : {vars : _} -> getRewriteTerms loc defs (NTCon nfc eq t a args) err = if !(isEqualTy eq) then case reverse $ map snd args of - (rhs :: lhs :: rhsty :: lhsty :: _) => + (rhs :%: lhs :%: rhsty :%: lhsty :%: _) => pure (!(evalClosure defs lhs), !(evalClosure defs rhs), !(evalClosure defs lhsty)) @@ -145,7 +145,7 @@ checkRewrite {vars} rigc elabinfo nest env ifc rule tm (Just expected) (rwtm, grwty) <- inScope vfc (pbind :: env) $ \e' => inScope {e=e'} vfc env' $ \e'' => - let offset = mkSizeOf [rname, pname] in + let offset = mkSizeOf (rname :%: pname :%: SLNil) in check {e = e''} rigc elabinfo (weakenNs offset nest) env' (apply (IVar vfc lemma.name) [ IVar vfc pname diff --git a/src/TTImp/Elab/RunElab.idr b/src/TTImp/Elab/RunElab.idr index 457a9b962d..ea64948a97 100644 --- a/src/TTImp/Elab/RunElab.idr +++ b/src/TTImp/Elab/RunElab.idr @@ -116,7 +116,7 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp -- parses and resolves `Language.Reflection.LookupDir` lookupDir : Defs -> NF vars -> Core String - lookupDir defs (NDCon _ conName _ _ []) + lookupDir defs (NDCon _ conName _ _ SLNil) = do defs <- get Ctxt NS ns (UN (Basic n)) <- toFullNames conName | fnm => failWith defs $ "bad lookup dir fullnames " ++ show fnm @@ -156,25 +156,25 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp pathDoesNotEscape n ("." ::rest) = pathDoesNotEscape n rest pathDoesNotEscape n (_ ::rest) = pathDoesNotEscape (S n) rest - elabCon : Defs -> String -> List (Closure vars) -> Core (NF vars) - elabCon defs "Pure" [_,val] + elabCon : Defs -> String -> ScopedList (Closure vars) -> Core (NF vars) + elabCon defs "Pure" (_ :%: val :%: SLNil) = do empty <- clearDefs defs evalClosure empty val - elabCon defs "Map" [_,_,fm,act] + elabCon defs "Map" (_ :%: _ :%: fm :%: act :%: SLNil) -- fm : A -> B -- elab : A = do act <- elabScript rig fc nest env !(evalClosure defs act) exp act <- quote defs env act fm <- evalClosure defs fm applyToStack defs withHoles env fm [(getLoc act, toClosure withAll env act)] - elabCon defs "Ap" [_,_,actF,actX] + elabCon defs "Ap" (_ :%: _ :%: actF :%: actX :%: SLNil) -- actF : Elab (A -> B) -- actX : Elab A = do actF <- elabScript rig fc nest env !(evalClosure defs actF) exp actX <- elabScript rig fc nest env !(evalClosure defs actX) exp actX <- quote defs env actX applyToStack defs withHoles env actF [(getLoc actX, toClosure withAll env actX)] - elabCon defs "Bind" [_,_,act,k] + elabCon defs "Bind" (_ :%: _ :%: act :%: k :%: SLNil) -- act : Elab A -- k : A -> Elab B -- 1) Run elabScript on act stripping off Elab @@ -187,14 +187,14 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp k <- evalClosure defs k r <- applyToStack defs withAll env k [(getLoc act, toClosure withAll env act)] elabScript rig fc nest env r exp - elabCon defs "Fail" [_, mbfc, msg] + elabCon defs "Fail" (_ :%: mbfc :%: msg :%: SLNil) = do msg' <- evalClosure defs msg throw $ RunElabFail $ GenericMsg !(reifyFC defs mbfc) !(reify defs msg') - elabCon defs "Warn" [mbfc, msg] + elabCon defs "Warn" (mbfc :%: msg :%: SLNil) = do msg' <- evalClosure defs msg recordWarning $ GenericWarn !(reifyFC defs mbfc) !(reify defs msg') scriptRet () - elabCon defs "Try" [_, elab1, elab2] + elabCon defs "Try" (_ :%: elab1 :%: elab2 :%: SLNil) = tryUnify (do constart <- getNextEntry res <- elabScript rig fc nest env !(evalClosure defs elab1) exp -- We ensure that all of the constraints introduced during the elab script @@ -203,14 +203,14 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp solveConstraintsAfter constart inTerm LastChance pure res) (elabScript rig fc nest env !(evalClosure defs elab2) exp) - elabCon defs "LogMsg" [topic, verb, str] + elabCon defs "LogMsg" (topic :%: verb :%: str :%: SLNil) = do topic' <- evalClosure defs topic verb' <- evalClosure defs verb unverifiedLogC !(reify defs topic') !(reify defs verb') $ do str' <- evalClosure defs str reify defs str' scriptRet () - elabCon defs "LogTerm" [topic, verb, str, tm] + elabCon defs "LogTerm" (topic :%: verb :%: str :%: tm :%: SLNil) = do topic' <- evalClosure defs topic verb' <- evalClosure defs verb unverifiedLogC !(reify defs topic') !(reify defs verb') $ @@ -219,7 +219,7 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp pure $ !(reify defs str') ++ ": " ++ show (the RawImp !(reify defs tm')) scriptRet () - elabCon defs "LogSugaredTerm" [topic, verb, str, tm] + elabCon defs "LogSugaredTerm" (topic :%: verb :%: str :%: tm :%: SLNil) = do topic' <- evalClosure defs topic verb' <- evalClosure defs verb unverifiedLogC !(reify defs topic') !(reify defs verb') $ @@ -228,7 +228,7 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp ptm <- pterm (map defaultKindedName tm') pure $ !(reify defs str') ++ ": " ++ show ptm scriptRet () - elabCon defs "Check" [exp, ttimp] + elabCon defs "Check" (exp :%: ttimp :%: SLNil) = do exp' <- evalClosure defs exp ttimp' <- evalClosure defs ttimp tidx <- resolveName (UN $ Basic "[elaborator script]") @@ -238,12 +238,12 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp (Just (glueBack defs env exp')) empty <- clearDefs defs nf empty env checktm - elabCon defs "Quote" [exp, tm] + elabCon defs "Quote" (exp :%: tm :%: SLNil) = do tm' <- evalClosure defs tm defs <- get Ctxt empty <- clearDefs defs scriptRet $ map rawName !(unelabUniqueBinders env !(quote empty env tm')) - elabCon defs "Lambda" [x, _, scope] + elabCon defs "Lambda" (x :%: _ :%: scope :%: SLNil) = do empty <- clearDefs defs NBind bfc x (Lam fc' c p ty) sc <- evalClosure defs scope | _ => failWith defs "Not a lambda" @@ -264,23 +264,23 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp quotePi Implicit = pure Implicit quotePi AutoImplicit = pure AutoImplicit quotePi (DefImplicit t) = failWith defs "Can't add default lambda" - elabCon defs "Goal" [] + elabCon defs "Goal" SLNil = do let Just gty = exp | Nothing => nfOpts withAll defs env !(reflect fc defs False env (the (Maybe RawImp) Nothing)) ty <- getTerm gty scriptRet (Just $ map rawName $ !(unelabUniqueBinders env ty)) - elabCon defs "LocalVars" [] + elabCon defs "LocalVars" SLNil = scriptRet vars - elabCon defs "GenSym" [str] + elabCon defs "GenSym" (str :%: SLNil) = do str' <- evalClosure defs str n <- genVarName !(reify defs str') scriptRet n - elabCon defs "InCurrentNS" [n] + elabCon defs "InCurrentNS" (n :%: SLNil) = do n' <- evalClosure defs n nsn <- inCurrentNS !(reify defs n') scriptRet nsn - elabCon defs "GetType" [n] + elabCon defs "GetType" (n :%: SLNil) = do n' <- evalClosure defs n res <- lookupTyName !(reify defs n') (gamma defs) scriptRet !(traverse unelabType res) @@ -288,15 +288,15 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp unelabType : (Name, Int, ClosedTerm) -> Core (Name, RawImp) unelabType (n, _, ty) = pure (n, map rawName !(unelabUniqueBinders [] ty)) - elabCon defs "GetInfo" [n] + elabCon defs "GetInfo" (n :%: SLNil) = do n' <- evalClosure defs n res <- lookupNameInfo !(reify defs n') (gamma defs) scriptRet res - elabCon defs "GetVis" [n] + elabCon defs "GetVis" (n :%: SLNil) = do dn <- reify defs !(evalClosure defs n) ds <- lookupCtxtName dn (gamma defs) scriptRet $ map (\(n,_,d) => (n, collapseDefault $ visibility d)) ds - elabCon defs "GetLocalType" [n] + elabCon defs "GetLocalType" (n :%: SLNil) = do n' <- evalClosure defs n n <- reify defs n' case defined n env of @@ -305,28 +305,28 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp let bty = binderType binder scriptRet $ map rawName !(unelabUniqueBinders env bty) _ => failWith defs $ show n ++ " is not a local variable" - elabCon defs "GetCons" [n] + elabCon defs "GetCons" (n :%: SLNil) = do n' <- evalClosure defs n cn <- reify defs n' Just (TCon _ _ _ _ _ _ cons _) <- lookupDefExact cn (gamma defs) | _ => failWith defs $ show cn ++ " is not a type" scriptRet cons - elabCon defs "GetReferredFns" [n] + elabCon defs "GetReferredFns" (n :%: SLNil) = do dn <- reify defs !(evalClosure defs n) Just def <- lookupCtxtExact dn (gamma defs) | Nothing => failWith defs $ show dn ++ " is not a definition" ns <- deepRefersTo def scriptRet ns - elabCon defs "GetCurrentFn" [] + elabCon defs "GetCurrentFn" SLNil = do defs <- get Ctxt scriptRet defs.defsStack - elabCon defs "Declare" [d] + elabCon defs "Declare" (d :%: SLNil) = do d' <- evalClosure defs d decls <- reify defs d' - traverse_ (processDecl [] (MkNested []) []) decls + Core.Core.ScopedList.traverse_ (processDecl [] (MkNested []) []) decls scriptRet () - elabCon defs "ReadFile" [lk, pth] + elabCon defs "ReadFile" (lk :%: pth :%: SLNil) = do pathPrefix <- lookupDir defs !(evalClosure defs lk) path <- reify defs !(evalClosure defs pth) validatePath defs path @@ -335,7 +335,7 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp | False => scriptRet $ Nothing {ty=String} contents <- readFile fullPath scriptRet $ Just contents - elabCon defs "WriteFile" [lk, pth, contents] + elabCon defs "WriteFile" (lk :%: pth :%: contents :%: SLNil) = do pathPrefix <- lookupDir defs !(evalClosure defs lk) path <- reify defs !(evalClosure defs pth) validatePath defs path @@ -344,7 +344,7 @@ elabScript rig fc nest env script@(NDCon nfc nm t ar args) exp whenJust (parent fullPath) ensureDirectoryExists writeFile fullPath contents scriptRet () - elabCon defs "IdrisDir" [lk] + elabCon defs "IdrisDir" (lk :%: SLNil) = do evalClosure defs lk >>= lookupDir defs >>= scriptRet elabCon defs n args = failWith defs $ "unexpected Elab constructor " ++ n ++ ", or incorrect count of arguments: " ++ show (length args) diff --git a/src/TTImp/Elab/Utils.idr b/src/TTImp/Elab/Utils.idr index 39adffc147..de006ee275 100644 --- a/src/TTImp/Elab/Utils.idr +++ b/src/TTImp/Elab/Utils.idr @@ -14,16 +14,16 @@ import TTImp.TTImp %default covering detagSafe : {auto c : Ref Ctxt Defs} -> - Defs -> NF [] -> Core Bool + Defs -> NF SLNil -> Core Bool detagSafe defs (NTCon _ n _ _ args) = do Just (TCon _ _ _ _ _ _ _ (Just detags)) <- lookupDefExact n (gamma defs) | _ => pure False - args' <- traverse (evalClosure defs . snd) args + args' <- traverse (evalClosure defs . snd) (toList args) pure $ notErased 0 detags args' where -- if any argument positions are in the detaggable set, and unerased, then -- detagging is safe - notErased : Nat -> List Nat -> List (NF []) -> Bool + notErased : Nat -> List Nat -> List (NF SLNil) -> Bool notErased i [] _ = True -- Don't need an index available notErased i ns [] = False notErased i ns (NErased _ Impossible :: rest) @@ -33,7 +33,7 @@ detagSafe defs (NTCon _ n _ _ args) detagSafe defs _ = pure False findErasedFrom : {auto c : Ref Ctxt Defs} -> - Defs -> Nat -> NF [] -> Core (List Nat, List Nat) + Defs -> Nat -> NF SLNil -> Core (List Nat, List Nat) findErasedFrom defs pos (NBind fc x (Pi _ c _ aty) scf) = do -- In the scope, use 'Erased fc Impossible' to mean 'argument is erased'. -- It's handy here, because we can use it to tell if a detaggable @@ -95,7 +95,7 @@ bindNotReq fc i (b :: env) (Keep p) ns tm = let tmptm = subst (Ref fc Bound (MN "arg" i)) tm (ns', btm) = bindNotReq fc (1 + i) env p ns tmptm in (ns', refToLocal (MN "arg" i) _ btm) -bindNotReq {vs = n :: _} fc i (b :: env) (Drop p) ns tm +bindNotReq {vs = n :%: _} fc i (b :: env) (Drop p) ns tm = bindNotReq fc i env p ((plicit b, n) :: ns) (Bind fc _ (Pi (binderLoc b) (multiplicity b) Explicit (binderType b)) tm) @@ -107,11 +107,11 @@ bindReq : {vs : _} -> bindReq {vs} fc env Refl ns tm = pure (ns, notLets [] _ env, abstractEnvType fc env tm) where - notLets : List Name -> (vars : List Name) -> Env Term vars -> List Name - notLets acc [] _ = acc - notLets acc (v :: vs) (b :: env) = if isLet b then notLets acc vs env + notLets : List Name -> (vars : ScopedList Name) -> Env Term vars -> List Name + notLets acc SLNil _ = acc + notLets acc (v :%: vs) (b :: env) = if isLet b then notLets acc vs env else notLets (v :: acc) vs env -bindReq {vs = n :: _} fc (b :: env) (Keep p) ns tm +bindReq {vs = n :%: _} fc (b :: env) (Keep p) ns tm = do b' <- shrinkBinder b p bindReq fc env p ((plicit b, n) :: ns) (Bind fc _ (Pi (binderLoc b) (multiplicity b) Explicit (binderType b')) tm) @@ -126,18 +126,18 @@ data ArgUsed = Used1 -- been used | Used0 -- not used | LocalVar -- don't care if it's used -data Usage : List Name -> Type where - Nil : Usage [] - (::) : ArgUsed -> Usage xs -> Usage (x :: xs) +data Usage : ScopedList Name -> Type where + Nil : Usage SLNil + (::) : ArgUsed -> Usage xs -> Usage (x :%: xs) -initUsed : (xs : List Name) -> Usage xs -initUsed [] = [] -initUsed (x :: xs) = Used0 :: initUsed xs +initUsed : (xs : ScopedList Name) -> Usage xs +initUsed SLNil = [] +initUsed (x :%: xs) = Used0 :: initUsed xs -initUsedCase : (xs : List Name) -> Usage xs -initUsedCase [] = [] -initUsedCase [x] = [Used0] -initUsedCase (x :: xs) = LocalVar :: initUsedCase xs +initUsedCase : (xs : ScopedList Name) -> Usage xs +initUsedCase SLNil = [] +initUsedCase (x :%: SLNil) = [Used0] +initUsedCase (x :%: xs) = LocalVar :: initUsedCase xs setUsedVar : {idx : _} -> (0 _ : IsVar n idx xs) -> Usage xs -> Usage xs @@ -158,17 +158,17 @@ setUsed : {idx : _} -> (0 _ : IsVar n idx vars) -> Core () setUsed p = update Used $ setUsedVar p -extendUsed : ArgUsed -> (new : List Name) -> Usage vars -> Usage (new ++ vars) -extendUsed a [] x = x -extendUsed a (y :: xs) x = a :: extendUsed a xs x +extendUsed : ArgUsed -> (new : ScopedList Name) -> Usage vars -> Usage (new +%+ vars) +extendUsed a SLNil x = x +extendUsed a (y :%: xs) x = a :: extendUsed a xs x -dropUsed : (new : List Name) -> Usage (new ++ vars) -> Usage vars -dropUsed [] x = x -dropUsed (x :: xs) (u :: us) = dropUsed xs us +dropUsed : (new : ScopedList Name) -> Usage (new +%+ vars) -> Usage vars +dropUsed SLNil x = x +dropUsed (x :%: xs) (u :: us) = dropUsed xs us -inExtended : ArgUsed -> (new : List Name) -> +inExtended : ArgUsed -> (new : ScopedList Name) -> {auto u : Ref Used (Usage vars)} -> - (Ref Used (Usage (new ++ vars)) -> Core a) -> + (Ref Used (Usage (new +%+ vars)) -> Core a) -> Core a inExtended a new sc = do used <- get Used @@ -198,7 +198,7 @@ termInlineSafe (Meta fc x y xs) termInlineSafe (Bind fc x b scope) = do bok <- binderInlineSafe b if bok - then inExtended LocalVar [x] (\u' => termInlineSafe scope) + then inExtended LocalVar (x :%: SLNil) (\u' => termInlineSafe scope) else pure False where binderInlineSafe : Binder (Term vars) -> Core Bool @@ -246,7 +246,7 @@ mutual caseAltInlineSafe (ConCase x tag args sc) = inExtended Used0 args (\u' => caseInlineSafe sc) caseAltInlineSafe (DelayCase ty arg sc) - = inExtended Used0 [ty, arg] (\u' => caseInlineSafe sc) + = inExtended Used0 (ty :%: arg :%: SLNil) (\u' => caseInlineSafe sc) caseAltInlineSafe (ConstCase x sc) = caseInlineSafe sc caseAltInlineSafe (DefaultCase sc) = caseInlineSafe sc diff --git a/src/TTImp/Impossible.idr b/src/TTImp/Impossible.idr index a733747610..dce6c0808a 100644 --- a/src/TTImp/Impossible.idr +++ b/src/TTImp/Impossible.idr @@ -22,13 +22,13 @@ import Data.List -- they involve resoling interfaces - they'll just become unmatchable patterns. match : {auto c : Ref Ctxt Defs} -> - NF [] -> (Name, Int, ClosedTerm) -> Core Bool + NF SLNil -> (Name, Int, ClosedTerm) -> Core Bool match nty (n, i, rty) = do defs <- get Ctxt rtynf <- nf defs [] rty sameRet nty rtynf where - sameRet : NF [] -> NF [] -> Core Bool + sameRet : NF SLNil -> NF SLNil -> Core Bool sameRet _ (NApp _ _ _) = pure True sameRet _ (NErased _ _) = pure True sameRet (NApp _ _ _) _ = pure True @@ -43,7 +43,7 @@ match nty (n, i, rty) sameRet _ _ = pure False dropNoMatch : {auto c : Ref Ctxt Defs} -> - Maybe (NF []) -> List (Name, Int, GlobalDef) -> + Maybe (NF SLNil) -> List (Name, Int, GlobalDef) -> Core (List (Name, Int, GlobalDef)) dropNoMatch _ [t] = pure [t] dropNoMatch Nothing ts = pure ts @@ -52,13 +52,13 @@ dropNoMatch (Just nty) ts filterM (match nty . map (map type)) ts nextVar : {auto q : Ref QVar Int} -> - FC -> Core (Term []) + FC -> Core (Term SLNil) nextVar fc = do i <- get QVar put QVar (i + 1) pure (Ref fc Bound (MN "imp" i)) -badClause : Term [] -> List RawImp -> List RawImp -> List (Name, RawImp) -> Core a +badClause : Term SLNil -> List RawImp -> List RawImp -> List (Name, RawImp) -> Core a badClause fn exps autos named = throw (GenericMsg (getLoc fn) ("Badly formed impossible clause " @@ -67,7 +67,7 @@ badClause fn exps autos named mutual processArgs : {auto c : Ref Ctxt Defs} -> {auto q : Ref QVar Int} -> - Term [] -> NF [] -> + Term SLNil -> NF SLNil -> (expargs : List RawImp) -> (autoargs : List RawImp) -> (namedargs : List (Name, RawImp)) -> @@ -121,7 +121,7 @@ mutual buildApp : {auto c : Ref Ctxt Defs} -> {auto q : Ref QVar Int} -> - FC -> Name -> Maybe (Closure []) -> + FC -> Name -> Maybe (Closure SLNil) -> (expargs : List RawImp) -> (autoargs : List RawImp) -> (namedargs : List (Name, RawImp)) -> @@ -150,7 +150,7 @@ mutual mkTerm : {auto c : Ref Ctxt Defs} -> {auto q : Ref QVar Int} -> - RawImp -> Maybe (Closure []) -> + RawImp -> Maybe (Closure SLNil) -> (expargs : List RawImp) -> (autoargs : List RawImp) -> (namedargs : List (Name, RawImp)) -> diff --git a/src/TTImp/Interactive/CaseSplit.idr b/src/TTImp/Interactive/CaseSplit.idr index e741314556..0717da6532 100644 --- a/src/TTImp/Interactive/CaseSplit.idr +++ b/src/TTImp/Interactive/CaseSplit.idr @@ -92,7 +92,7 @@ getDefining tm -- For the name on the lhs, return the function name being defined, the -- type name, and the possible constructors. findCons : {auto c : Ref Ctxt Defs} -> - Name -> Term [] -> Core (SplitResult (Name, Name, List Name)) + Name -> Term SLNil -> Core (SplitResult (Name, Name, List Name)) findCons n lhs = case getDefining lhs of Nothing => pure (SplitFail @@ -123,7 +123,7 @@ findAllVars (Bind _ x (PLet _ _ _ _) sc) findAllVars t = toList (dropNS <$> getDefining t) export -explicitlyBound : Defs -> NF [] -> Core (List Name) +explicitlyBound : Defs -> NF SLNil -> Core (List Name) explicitlyBound defs (NBind fc x (Pi _ _ _ _) sc) = pure $ x :: !(explicitlyBound defs !(sc defs (toClosure defaultOpts [] (Erased fc Placeholder)))) @@ -131,7 +131,7 @@ explicitlyBound defs _ = pure [] export getEnvArgNames : {auto c : Ref Ctxt Defs} -> - Defs -> Nat -> NF [] -> Core (List String) + Defs -> Nat -> NF SLNil -> Core (List String) getEnvArgNames defs Z sc = getArgNames defs !(explicitlyBound defs sc) [] [] sc getEnvArgNames defs (S k) (NBind fc n _ sc) = getEnvArgNames defs k !(sc defs (toClosure defaultOpts [] (Erased fc Placeholder))) diff --git a/src/TTImp/Interactive/ExprSearch.idr b/src/TTImp/Interactive/ExprSearch.idr index 36fad55967..57ccf31a5b 100644 --- a/src/TTImp/Interactive/ExprSearch.idr +++ b/src/TTImp/Interactive/ExprSearch.idr @@ -43,7 +43,7 @@ import Libraries.Data.WithDefault -- of the LHS. Only recursive calls with a different structure are okay. record RecData where constructor MkRecData - {localVars : List Name} + {localVars : ScopedList Name} recname : Name -- resolved name lhsapp : Term localVars @@ -147,18 +147,18 @@ search : {auto c : Ref Ctxt Defs} -> getAllEnv : {vars : _} -> FC -> SizeOf done -> Env Term vars -> - List (Term (done ++ vars), Term (done ++ vars)) + List (Term (done +%+ vars), Term (done +%+ vars)) getAllEnv fc done [] = [] -getAllEnv {vars = v :: vs} {done} fc p (b :: env) +getAllEnv {vars = v :%: vs} {done} fc p (b :: env) = let rest = getAllEnv fc (sucR p) env 0 var = mkIsVar (hasLength p) usable = usableName v in if usable then (Local fc Nothing _ var, - rewrite appendAssociative done [v] vs in + rewrite appendAssociative done (v :%: SLNil) vs in weakenNs (sucR p) (binderType b)) :: - rewrite appendAssociative done [v] vs in rest - else rewrite appendAssociative done [v] vs in rest + rewrite appendAssociative done (v :%: SLNil) vs in rest + else rewrite appendAssociative done (v :%: SLNil) vs in rest where usableName : Name -> Bool usableName (UN _) = True @@ -334,7 +334,7 @@ getSuccessful {vars} fc rig opts mkHole env ty topty all let base = maybe "arg" (\r => nameRoot (recname r) ++ "_rhs") (recData opts) - hn <- uniqueBasicName defs (map nameRoot vars) base + hn <- uniqueBasicName defs (toList $ map nameRoot vars) base (idx, tm) <- newMeta fc rig env (UN $ Basic hn) ty (Hole (length env) (holeInit False)) False @@ -488,7 +488,7 @@ searchLocalWith {vars} fc nofn rig opts hints env ((p, pty) :: rest) ty topty findPos : Defs -> Term vars -> (Term vars -> Term vars) -> NF vars -> NF vars -> Core (Search (Term vars, ExprDefs)) - findPos defs prf f x@(NTCon pfc pn _ _ [(fc1, xty), (fc2, yty)]) target + findPos defs prf f x@(NTCon pfc pn _ _ ((fc1, xty) :%: (fc2, yty) :%: SLNil)) target = getSuccessful fc rig opts False env ty topty [findDirect defs prf f x target, (do fname <- maybe (throw (InternalError "No fst")) @@ -505,16 +505,16 @@ searchLocalWith {vars} fc nofn rig opts hints env ((p, pty) :: rest) ty topty [(do xtynf <- evalClosure defs xty findPos defs prf (\arg => applyStackWithFC (Ref fc Func fname) - [(fc1, xtytm), - (fc2, ytytm), - (fc, f arg)]) + ((fc1, xtytm) :%: + (fc2, ytytm) :%: + (fc, f arg) :%: SLNil)) xtynf target), (do ytynf <- evalClosure defs yty findPos defs prf (\arg => applyStackWithFC (Ref fc Func sname) - [(fc1, xtytm), - (fc2, ytytm), - (fc, f arg)]) + ((fc1, xtytm) :%: + (fc2, ytytm) :%: + (fc, f arg) :%: SLNil)) ytynf target)] else noResult)] findPos defs prf f nty target = findDirect defs prf f nty target @@ -683,7 +683,7 @@ tryIntermediateRec fc rig opts hints env ty topty (Just rd) recsearch <- tryRecursive fc rig opts' hints env letty topty rd makeHelper fc rig opts' env letty ty recsearch where - isSingleCon : Defs -> NF [] -> Core Bool + isSingleCon : Defs -> NF SLNil -> Core Bool isSingleCon defs (NBind fc x (Pi _ _ _ _) sc) = isSingleCon defs !(sc defs (toClosure defaultOpts [] (Erased fc Placeholder))) @@ -701,7 +701,7 @@ searchType : {vars : _} -> ClosedTerm -> Nat -> Term vars -> Core (Search (Term vars, ExprDefs)) searchType fc rig opts hints env topty (S k) (Bind bfc n b@(Pi fc' c info ty) sc) - = do let env' : Env Term (n :: _) = b :: env + = do let env' : Env Term (n :%: _) = b :: env log "interaction.search" 10 $ "Introduced lambda, search for " ++ show sc scVal <- searchType fc rig opts hints env' topty k sc pure (map (\ (sc, ds) => (Bind bfc n (Lam fc' c info ty) sc, ds)) scVal) @@ -710,8 +710,8 @@ searchType {vars} fc rig opts hints env topty Z (Bind bfc n b@(Pi fc' c info ty) getSuccessful fc rig opts False env ty topty [searchLocal fc rig opts hints env (Bind bfc n b sc) topty, (do defs <- get Ctxt - let n' = UN $ Basic !(getArgName defs n [] vars !(nf defs env ty)) - let env' : Env Term (n' :: _) = b :: env + let n' = UN $ Basic !(getArgName defs n [] (toList vars) !(nf defs env ty)) + let env' : Env Term (n' :%: _) = b :: env let sc' = compat sc log "interaction.search" 10 $ "Introduced lambda, search for " ++ show sc' scVal <- searchType fc rig opts hints env' topty Z sc' @@ -856,7 +856,7 @@ exprSearchOpts opts fc n_in hints -- the REPL does this step, but doing it here too because -- expression search might be invoked some other way let Hole _ _ = definition gdef - | PMDef pi [] (STerm _ tm) _ _ + | PMDef pi SLNil (STerm _ tm) _ _ => do raw <- unelab [] !(toFullNames !(normaliseHoles defs [] tm)) one (map rawName raw) | _ => throw (GenericMsg fc "Name is already defined") diff --git a/src/TTImp/Interactive/Intro.idr b/src/TTImp/Interactive/Intro.idr index 57973f0dcf..3f178d231e 100644 --- a/src/TTImp/Interactive/Intro.idr +++ b/src/TTImp/Interactive/Intro.idr @@ -59,7 +59,7 @@ parameters let pcons = papply replFC (PRef replFC cons) new_holes res <- catch (do -- We're desugaring it to the corresponding TTImp - icons <- desugar AnyExpr lhsCtxt pcons + icons <- desugar AnyExpr (toList lhsCtxt) pcons ccons <- checkTerm hidx {-is this correct?-} InExpr [] (MkNested []) env icons gty newdefs <- get Ctxt ncons <- normaliseHoles newdefs env ccons diff --git a/src/TTImp/Interactive/MakeLemma.idr b/src/TTImp/Interactive/MakeLemma.idr index 312a6fd5c5..66fc23a369 100644 --- a/src/TTImp/Interactive/MakeLemma.idr +++ b/src/TTImp/Interactive/MakeLemma.idr @@ -49,7 +49,7 @@ getArgs : {vars : _} -> getArgs {vars} env (S k) (Bind _ x b@(Pi _ c _ ty) sc) = do defs <- get Ctxt ty' <- map (map rawName) $ unelab env !(normalise defs env ty) - let x' = UN $ Basic !(uniqueBasicName defs (map nameRoot vars) (nameRoot x)) + let x' = UN $ Basic !(uniqueBasicName defs (toList $ map nameRoot vars) (nameRoot x)) (sc', ty) <- getArgs (b :: env) k (compat {n = x'} sc) -- Don't need to use the name if it's not used in the scope type let mn = if c == top diff --git a/src/TTImp/PartialEval.idr b/src/TTImp/PartialEval.idr index 75a1a26dd4..c2a68b2bd5 100644 --- a/src/TTImp/PartialEval.idr +++ b/src/TTImp/PartialEval.idr @@ -9,6 +9,7 @@ import Core.Metadata import Core.Normalise import Core.Value import Core.UnifyState +import Core.Name.ScopedList import Idris.REPL.Opts import Idris.Syntax @@ -42,12 +43,12 @@ Show a => Show (ArgMode' a) where show Dynamic = "Dynamic" -getStatic : ArgMode -> Maybe (Term []) +getStatic : ArgMode -> Maybe (Term SLNil) getStatic Dynamic = Nothing getStatic (Static t) = Just t specialiseTy : {vars : _} -> - Nat -> List (Nat, Term []) -> Term vars -> Term vars + Nat -> List (Nat, Term SLNil) -> Term vars -> Term vars specialiseTy i specs (Bind fc x (Pi fc' c p ty) sc) = case lookup i specs of Nothing => Bind fc x (Pi fc' c p ty) $ -- easier later if everything explicit @@ -75,7 +76,7 @@ substLocs : {vs : _} -> substLocs [] tm = tm substLocs ((i, tm') :: subs) tm = substLocs subs (substLoc i tm' tm) -mkSubsts : Nat -> List (Nat, Term []) -> +mkSubsts : Nat -> List (Nat, Term SLNil) -> List (Term vs) -> Term vs -> Maybe (List (Nat, Term vs)) mkSubsts i specs [] rhs = Just [] mkSubsts i specs (arg :: args) rhs @@ -93,7 +94,7 @@ mkSubsts i specs (arg :: args) rhs -- In the case where all the specialised positions are variables on the LHS, -- substitute the term in on the RHS -specPatByVar : List (Nat, Term []) -> +specPatByVar : List (Nat, Term SLNil) -> (vs ** (Env Term vs, Term vs, Term vs)) -> Maybe (vs ** (Env Term vs, Term vs, Term vs)) specPatByVar specs (vs ** (env, lhs, rhs)) @@ -102,7 +103,7 @@ specPatByVar specs (vs ** (env, lhs, rhs)) let lhs' = apply (getLoc fn) fn args pure (vs ** (env, substLocs psubs lhs', substLocs psubs rhs)) -specByVar : List (Nat, Term []) -> +specByVar : List (Nat, Term SLNil) -> List (vs ** (Env Term vs, Term vs, Term vs)) -> Maybe (List (vs ** (Env Term vs, Term vs, Term vs))) specByVar specs [] = pure [] @@ -111,19 +112,19 @@ specByVar specs (p :: ps) ps' <- specByVar specs ps pure (p' :: ps') -dropSpec : Nat -> List (Nat, Term []) -> List a -> List a -dropSpec i sargs [] = [] -dropSpec i sargs (x :: xs) +dropSpec : Nat -> List (Nat, Term SLNil) -> ScopedList a -> ScopedList a +dropSpec i sargs SLNil = SLNil +dropSpec i sargs (x :%: xs) = case lookup i sargs of - Nothing => x :: dropSpec (1 + i) sargs xs + Nothing => x :%: dropSpec (1 + i) sargs xs Just _ => dropSpec (1 + i) sargs xs getSpecPats : {auto c : Ref Ctxt Defs} -> FC -> Name -> - (fn : Name) -> (stk : List (FC, Term vars)) -> - NF [] -> -- Type of 'fn' + (fn : Name) -> (stk : ScopedList (FC, Term vars)) -> + NF SLNil -> -- Type of 'fn' List (Nat, ArgMode) -> -- All the arguments - List (Nat, Term []) -> -- Just the static ones + List (Nat, Term SLNil) -> -- Just the static ones List (vs ** (Env Term vs, Term vs, Term vs)) -> Core (Maybe (List ImpClause)) getSpecPats fc pename fn stk fnty args sargs pats @@ -150,7 +151,7 @@ getSpecPats fc pename fn stk fnty args sargs pats -- Build a RHS from the type of the function to be specialised, the -- dynamic argument names, and the list of given arguments. We assume -- the latter two correspond appropriately. - mkRHSargs : NF [] -> RawImp -> List String -> List (Nat, ArgMode) -> + mkRHSargs : NF SLNil -> RawImp -> List String -> List (Nat, ArgMode) -> Core RawImp mkRHSargs (NBind _ x (Pi _ _ Explicit _) sc) app (a :: as) ((_, Dynamic) :: ds) = do defs <- get Ctxt @@ -182,24 +183,24 @@ getSpecPats fc pename fn stk fnty args sargs pats mkRHSargs _ app _ _ = pure app - getRawArgs : List (Arg' Name) -> RawImp -> List (Arg' Name) - getRawArgs args (IApp fc f arg) = getRawArgs (Explicit fc arg :: args) f + getRawArgs : ScopedList (Arg' Name) -> RawImp -> ScopedList (Arg' Name) + getRawArgs args (IApp fc f arg) = getRawArgs (Explicit fc arg :%: args) f getRawArgs args (INamedApp fc f n arg) - = getRawArgs (Named fc n arg :: args) f + = getRawArgs (Named fc n arg :%: args) f getRawArgs args (IAutoApp fc f arg) - = getRawArgs (Auto fc arg :: args) f + = getRawArgs (Auto fc arg :%: args) f getRawArgs args tm = args - reapply : RawImp -> List (Arg' Name) -> RawImp - reapply f [] = f - reapply f (Explicit fc arg :: args) = reapply (IApp fc f arg) args - reapply f (Named fc n arg :: args) + reapply : RawImp -> ScopedList (Arg' Name) -> RawImp + reapply f SLNil = f + reapply f (Explicit fc arg :%: args) = reapply (IApp fc f arg) args + reapply f (Named fc n arg :%: args) = reapply (INamedApp fc f n arg) args - reapply f (Auto fc arg :: args) + reapply f (Auto fc arg :%: args) = reapply (IAutoApp fc f arg) args dropArgs : Name -> RawImp -> RawImp - dropArgs pename tm = reapply (IVar fc pename) (dropSpec 0 sargs (getRawArgs [] tm)) + dropArgs pename tm = reapply (IVar fc pename) (dropSpec 0 sargs (getRawArgs SLNil tm)) unelabPat : Name -> (vs ** (Env Term vs, Term vs, Term vs)) -> Core ImpClause @@ -246,7 +247,7 @@ mkSpecDef : {auto c : Ref Ctxt Defs} -> {auto s : Ref Syn SyntaxInfo} -> {auto o : Ref ROpts REPLOpts} -> FC -> GlobalDef -> - Name -> List (Nat, ArgMode) -> Name -> List (FC, Term vars) -> + Name -> List (Nat, ArgMode) -> Name -> ScopedList (FC, Term vars) -> Core (Term vars) mkSpecDef {vars} fc gdef pename sargs fn stk = handleUnify {unResolve = True} @@ -282,7 +283,7 @@ mkSpecDef {vars} fc gdef pename sargs fn stk log "specialise.flags" 20 "Defining \{show pename} with flags: \{show defflags}" peidx <- addDef pename $ the (GlobalDef -> GlobalDef) { flags := defflags } - $ newDef fc pename top [] sty (specified Public) None + $ newDef fc pename top SLNil sty (specified Public) None addToSave (Resolved peidx) -- Reduce the function to be specialised, and reduce any name in @@ -400,7 +401,7 @@ specialise : {vars : _} -> {auto s : Ref Syn SyntaxInfo} -> {auto o : Ref ROpts REPLOpts} -> FC -> Env Term vars -> GlobalDef -> - Name -> List (FC, Term vars) -> + Name -> ScopedList (FC, Term vars) -> Core (Maybe (Term vars)) specialise {vars} fc env gdef fn stk = case specArgs gdef of @@ -408,7 +409,7 @@ specialise {vars} fc env gdef fn stk specs => do fnfull <- toFullNames fn -- If all the arguments are concrete (meaning, no local variables - -- or holes in them, so they can be a Term []) we can specialise + -- or holes in them, so they can be a Term SLNil) we can specialise Just sargs <- getSpecArgs 0 specs stk | Nothing => pure Nothing defs <- get Ctxt @@ -424,13 +425,13 @@ specialise {vars} fc env gdef fn stk Just _ => pure Nothing where concrete : {vars : _} -> - Term vars -> Maybe (Term []) + Term vars -> Maybe (Term SLNil) concrete tm = shrink tm none - getSpecArgs : Nat -> List Nat -> List (FC, Term vars) -> + getSpecArgs : Nat -> List Nat -> ScopedList (FC, Term vars) -> Core (Maybe (List (Nat, ArgMode))) - getSpecArgs i specs [] = pure (Just []) - getSpecArgs i specs ((_, x) :: xs) + getSpecArgs i specs SLNil = pure (Just []) + getSpecArgs i specs ((_, x) :%: xs) = do Just xs' <- getSpecArgs (1 + i) specs xs | Nothing => pure Nothing if i `elem` specs @@ -448,7 +449,7 @@ findSpecs : {vars : _} -> {auto u : Ref UST UState} -> {auto s : Ref Syn SyntaxInfo} -> {auto o : Ref ROpts REPLOpts} -> - Env Term vars -> List (FC, Term vars) -> Term vars -> + Env Term vars -> ScopedList (FC, Term vars) -> Term vars -> Core (Term vars) findSpecs env stk (Ref fc Func fn) = do defs <- get Ctxt @@ -458,24 +459,24 @@ findSpecs env stk (Ref fc Func fn) | Nothing => pure (applyStackWithFC (Ref fc Func fn) stk) pure r findSpecs env stk (Meta fc n i args) - = do args' <- traverse (findSpecs env []) args + = do args' <- traverse (findSpecs env SLNil) args pure $ applyStackWithFC (Meta fc n i args') stk findSpecs env stk (Bind fc x b sc) - = do b' <- traverse (findSpecs env []) b - sc' <- findSpecs (b' :: env) [] sc + = do b' <- traverse (findSpecs env SLNil) b + sc' <- findSpecs (b' :: env) SLNil sc pure $ applyStackWithFC (Bind fc x b' sc') stk findSpecs env stk (App fc fn arg) - = do arg' <- findSpecs env [] arg - findSpecs env ((fc, arg') :: stk) fn + = do arg' <- findSpecs env SLNil arg + findSpecs env ((fc, arg') :%: stk) fn findSpecs env stk (TDelayed fc r tm) - = do tm' <- findSpecs env [] tm + = do tm' <- findSpecs env SLNil tm pure $ applyStackWithFC (TDelayed fc r tm') stk findSpecs env stk (TDelay fc r ty tm) - = do ty' <- findSpecs env [] ty - tm' <- findSpecs env [] tm + = do ty' <- findSpecs env SLNil ty + tm' <- findSpecs env SLNil tm pure $ applyStackWithFC (TDelay fc r ty' tm') stk findSpecs env stk (TForce fc r tm) - = do tm' <- findSpecs env [] tm + = do tm' <- findSpecs env SLNil tm pure $ applyStackWithFC (TForce fc r tm') stk findSpecs env stk tm = pure $ applyStackWithFC tm stk @@ -499,11 +500,11 @@ mutual {auto s : Ref Syn SyntaxInfo} -> {auto o : Ref ROpts REPLOpts} -> Ref QVar Int -> Defs -> Bounds bound -> - Env Term free -> List (Closure free) -> - Core (List (Term (bound ++ free))) - quoteArgs q defs bounds env [] = pure [] - quoteArgs q defs bounds env (a :: args) - = pure $ (!(quoteGenNF q defs bounds env !(evalClosure defs a)) :: + Env Term free -> ScopedList (Closure free) -> + Core (ScopedList (Term (bound +%+ free))) + quoteArgs q defs bounds env SLNil = pure SLNil + quoteArgs q defs bounds env (a :%: args) + = pure $ (!(quoteGenNF q defs bounds env !(evalClosure defs a)) :%: !(quoteArgs q defs bounds env args)) quoteArgsWithFC : {auto c : Ref Ctxt Defs} -> @@ -513,8 +514,8 @@ mutual {auto o : Ref ROpts REPLOpts} -> {bound, free : _} -> Ref QVar Int -> Defs -> Bounds bound -> - Env Term free -> List (FC, Closure free) -> - Core (List (FC, Term (bound ++ free))) + Env Term free -> ScopedList (FC, Closure free) -> + Core (ScopedList (FC, Term (bound +%+ free))) quoteArgsWithFC q defs bounds env terms = pure $ zip (map fst terms) !(quoteArgs q defs bounds env (map snd terms)) @@ -526,15 +527,15 @@ mutual {auto o : Ref ROpts REPLOpts} -> Ref QVar Int -> Defs -> FC -> Bounds bound -> Env Term free -> NHead free -> - Core (Term (bound ++ free)) + Core (Term (bound +%+ free)) quoteHead {bound} q defs fc bounds env (NLocal mrig _ prf) = let MkVar prf' = addLater bound prf in pure $ Local fc mrig _ prf' where - addLater : {idx : _} -> (ys : List Name) -> (0 p : IsVar n idx xs) -> - Var (ys ++ xs) - addLater [] isv = MkVar isv - addLater (x :: xs) isv + addLater : {idx : _} -> (ys : ScopedList Name) -> (0 p : IsVar n idx xs) -> + Var (ys +%+ xs) + addLater SLNil isv = MkVar isv + addLater (x :%: xs) isv = let MkVar isv' = addLater xs isv in MkVar (Later isv') quoteHead q defs fc bounds env (NRef Bound (MN n i)) @@ -556,7 +557,7 @@ mutual quoteHead q defs fc bounds env (NRef nt n) = pure $ Ref fc nt n quoteHead q defs fc bounds env (NMeta n i args) = do args' <- quoteArgs q defs bounds env args - pure $ Meta fc n i args' + pure $ Meta fc n i (toList args') quotePi : {bound, free : _} -> {auto c : Ref Ctxt Defs} -> @@ -566,7 +567,7 @@ mutual {auto o : Ref ROpts REPLOpts} -> Ref QVar Int -> Defs -> Bounds bound -> Env Term free -> PiInfo (Closure free) -> - Core (PiInfo (Term (bound ++ free))) + Core (PiInfo (Term (bound +%+ free))) quotePi q defs bounds env Explicit = pure Explicit quotePi q defs bounds env Implicit = pure Implicit quotePi q defs bounds env AutoImplicit = pure AutoImplicit @@ -582,7 +583,7 @@ mutual {auto o : Ref ROpts REPLOpts} -> Ref QVar Int -> Defs -> Bounds bound -> Env Term free -> Binder (Closure free) -> - Core (Binder (Term (bound ++ free))) + Core (Binder (Term (bound +%+ free))) quoteBinder q defs bounds env (Lam fc r p ty) = do ty' <- quoteGenNF q defs bounds env !(evalClosure defs ty) p' <- quotePi q defs bounds env p @@ -615,7 +616,7 @@ mutual {auto o : Ref ROpts REPLOpts} -> Ref QVar Int -> Defs -> Bounds bound -> - Env Term vars -> NF vars -> Core (Term (bound ++ vars)) + Env Term vars -> NF vars -> Core (Term (bound +%+ vars)) quoteGenNF q defs bound env (NBind fc n b sc) = do var <- bName "qv" sc' <- quoteGenNF q defs (Add n var bound) env @@ -642,7 +643,7 @@ mutual pure $ applyStackWithFC (Ref fc Func fn) args' pure r where - extendEnv : Bounds bs -> Env Term vs -> Env Term (bs ++ vs) + extendEnv : Bounds bs -> Env Term vs -> Env Term (bs +%+ vs) extendEnv None env = env extendEnv (Add x n bs) env -- We're just using this to evaluate holes in the right scope, so @@ -718,7 +719,7 @@ applySpecialise : {vars : _} -> Term vars -> -- initial RHS Core (Term vars) applySpecialise env Nothing tm - = findSpecs env [] tm -- not specialising, just search through RHS + = findSpecs env SLNil tm -- not specialising, just search through RHS applySpecialise env (Just ls) tmin -- specialising, evaluate RHS while looking -- for names to specialise = do defs <- get Ctxt diff --git a/src/TTImp/ProcessData.idr b/src/TTImp/ProcessData.idr index e3e6b001bb..653072dd02 100644 --- a/src/TTImp/ProcessData.idr +++ b/src/TTImp/ProcessData.idr @@ -127,22 +127,22 @@ checkCon {vars} opts nest env vis tn_in tn (MkImpTy fc _ cn_in ty_raw) -- Get the indices of the constructor type (with non-constructor parts erased) getIndexPats : {auto c : Ref Ctxt Defs} -> - ClosedTerm -> Core (List (NF [])) + ClosedTerm -> Core (List (NF SLNil)) getIndexPats tm = do defs <- get Ctxt tmnf <- nf defs [] tm ret <- getRetType defs tmnf getPats defs ret where - getRetType : Defs -> NF [] -> Core (NF []) + getRetType : Defs -> NF SLNil -> Core (NF SLNil) getRetType defs (NBind fc _ (Pi _ _ _ _) sc) = do sc' <- sc defs (toClosure defaultOpts [] (Erased fc Placeholder)) getRetType defs sc' getRetType defs t = pure t - getPats : Defs -> NF [] -> Core (List (NF [])) + getPats : Defs -> NF SLNil -> Core (List (NF SLNil)) getPats defs (NTCon fc _ _ _ args) - = traverse (evalClosure defs . snd) args + = traverse (evalClosure defs . snd) (toList args) getPats defs _ = pure [] -- Can't happen if we defined the type successfully! getDetags : {auto c : Ref Ctxt Defs} -> @@ -156,7 +156,7 @@ getDetags fc tys xs => pure $ Just xs where mutual - disjointArgs : List (NF []) -> List (NF []) -> Core Bool + disjointArgs : List (NF SLNil) -> List (NF SLNil) -> Core Bool disjointArgs [] _ = pure False disjointArgs _ [] = pure False disjointArgs (a :: args) (a' :: args') @@ -164,25 +164,25 @@ getDetags fc tys then pure True else disjointArgs args args' - disjoint : NF [] -> NF [] -> Core Bool + disjoint : NF SLNil -> NF SLNil -> Core Bool disjoint (NDCon _ _ t _ args) (NDCon _ _ t' _ args') = if t /= t' then pure True else do defs <- get Ctxt - argsnf <- traverse (evalClosure defs . snd) args - args'nf <- traverse (evalClosure defs . snd) args' + argsnf <- traverse (evalClosure defs . snd) (toList args) + args'nf <- traverse (evalClosure defs . snd) (toList args') disjointArgs argsnf args'nf disjoint (NTCon _ n _ _ args) (NDCon _ n' _ _ args') = if n /= n' then pure True else do defs <- get Ctxt - argsnf <- traverse (evalClosure defs . snd) args - args'nf <- traverse (evalClosure defs . snd) args' + argsnf <- traverse (evalClosure defs . snd) (toList args) + args'nf <- traverse (evalClosure defs . snd) (toList args') disjointArgs argsnf args'nf disjoint (NPrimVal _ c) (NPrimVal _ c') = pure (c /= c') disjoint _ _ = pure False - allDisjointWith : NF [] -> List (NF []) -> Core Bool + allDisjointWith : NF SLNil -> List (NF SLNil) -> Core Bool allDisjointWith val [] = pure True allDisjointWith (NErased _ _) _ = pure False allDisjointWith val (nf :: nfs) @@ -190,7 +190,7 @@ getDetags fc tys if ok then allDisjointWith val nfs else pure False - allDisjoint : List (NF []) -> Core Bool + allDisjoint : List (NF SLNil) -> Core Bool allDisjoint [] = pure True allDisjoint (NErased _ _ :: _) = pure False allDisjoint (nf :: nfs) @@ -199,7 +199,7 @@ getDetags fc tys else pure False -- Which argument positions have completely disjoint contructors - getDisjointPos : Nat -> List (List (NF [])) -> Core (List Nat) + getDisjointPos : Nat -> List (List (NF SLNil)) -> Core (List Nat) getDisjointPos i [] = pure [] getDisjointPos i (args :: argss) = do rest <- getDisjointPos (1 + i) argss @@ -209,7 +209,7 @@ getDetags fc tys -- If exactly one argument is unerased, return its position getRelevantArg : {auto c : Ref Ctxt Defs} -> - Defs -> Nat -> Maybe Nat -> Bool -> NF [] -> + Defs -> Nat -> Maybe Nat -> Bool -> NF SLNil -> Core (Maybe (Bool, Nat)) getRelevantArg defs i rel world (NBind fc _ (Pi _ rig _ val) sc) = branchZero (getRelevantArg defs (1 + i) rel world @@ -405,7 +405,7 @@ processData : {vars : _} -> ImpData -> Core () processData {vars} eopts nest env fc def_vis mbtot (MkImpLater dfc n_in ty_raw) = do n <- inCurrentNS n_in - ty_raw <- bindTypeNames fc [] vars ty_raw + ty_raw <- bindTypeNames fc [] (toList vars) ty_raw defs <- get Ctxt -- Check 'n' is undefined @@ -448,7 +448,7 @@ processData {vars} eopts nest env fc def_vis mbtot (MkImpData dfc n_in mty_raw o defs <- get Ctxt mmetasfullty <- flip traverseOpt mty_raw $ \ ty_raw => do - ty_raw <- bindTypeNames fc [] vars ty_raw + ty_raw <- bindTypeNames fc [] (toList vars) ty_raw u <- uniVar fc (ty, _) <- diff --git a/src/TTImp/ProcessDecls.idr b/src/TTImp/ProcessDecls.idr index 779e289a9f..97eaad64d3 100644 --- a/src/TTImp/ProcessDecls.idr +++ b/src/TTImp/ProcessDecls.idr @@ -172,22 +172,22 @@ processTTImpDecls {vars} nest env decls where bindConNames : ImpTy -> Core ImpTy bindConNames (MkImpTy fc nameFC n ty) - = do ty' <- bindTypeNames fc [] vars ty + = do ty' <- bindTypeNames fc [] (toList vars) ty pure (MkImpTy fc nameFC n ty') bindDataNames : ImpData -> Core ImpData bindDataNames (MkImpData fc n t opts cons) - = do t' <- traverseOpt (bindTypeNames fc [] vars) t + = do t' <- traverseOpt (bindTypeNames fc [] (toList vars)) t cons' <- traverse bindConNames cons pure (MkImpData fc n t' opts cons') bindDataNames (MkImpLater fc n t) - = do t' <- bindTypeNames fc [] vars t + = do t' <- bindTypeNames fc [] (toList vars) t pure (MkImpLater fc n t') -- bind implicits to make raw TTImp source a bit friendlier bindNames : ImpDecl -> Core ImpDecl bindNames (IClaim fc c vis opts (MkImpTy tfc nameFC n ty)) - = do ty' <- bindTypeNames fc [] vars ty + = do ty' <- bindTypeNames fc [] (toList vars) ty pure (IClaim fc c vis opts (MkImpTy tfc nameFC n ty')) bindNames (IData fc vis mbtot d) = do d' <- bindDataNames d diff --git a/src/TTImp/ProcessDef.idr b/src/TTImp/ProcessDef.idr index b10a44b68f..a0e07dbb27 100644 --- a/src/TTImp/ProcessDef.idr +++ b/src/TTImp/ProcessDef.idr @@ -51,11 +51,11 @@ mutual mismatchNF defs (NTCon _ xn xt _ xargs) (NTCon _ yn yt _ yargs) = if xn /= yn then pure True - else anyM (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) + else anyMScoped (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) mismatchNF defs (NDCon _ _ xt _ xargs) (NDCon _ _ yt _ yargs) = if xt /= yt then pure True - else anyM (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) + else anyMScoped (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) mismatchNF defs (NPrimVal _ xc) (NPrimVal _ yc) = pure (xc /= yc) mismatchNF defs (NDelayed _ _ x) (NDelayed _ _ y) = mismatchNF defs x y mismatchNF defs (NDelay _ _ _ x) (NDelay _ _ _ y) @@ -99,12 +99,12 @@ impossibleOK : {auto c : Ref Ctxt Defs} -> impossibleOK defs (NTCon _ xn xt xa xargs) (NTCon _ yn yt ya yargs) = if xn /= yn then pure True - else anyM (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) + else anyMScoped (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) -- If it's a data constructor, any mismatch will do impossibleOK defs (NDCon _ _ xt _ xargs) (NDCon _ _ yt _ yargs) = if xt /= yt then pure True - else anyM (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) + else anyMScoped (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs) impossibleOK defs (NPrimVal _ x) (NPrimVal _ y) = pure (x /= y) -- NPrimVal is apart from NDCon, NTCon, NBind, and NType @@ -161,7 +161,7 @@ recoverable : {auto c : Ref Ctxt Defs} -> recoverable defs (NTCon _ xn xt xa xargs) (NTCon _ yn yt ya yargs) = if xn /= yn then pure False - else pure $ not !(anyM (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs)) + else pure $ not !(anyMScoped (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs)) -- Type constructor vs. primitive type recoverable defs (NTCon _ _ _ _ _) (NPrimVal _ _) = pure False recoverable defs (NPrimVal _ _) (NTCon _ _ _ _ _) = pure False @@ -179,7 +179,7 @@ recoverable defs _ (NTCon _ _ _ _ _) = pure True recoverable defs (NDCon _ _ xt _ xargs) (NDCon _ _ yt _ yargs) = if xt /= yt then pure False - else pure $ not !(anyM (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs)) + else pure $ not !(anyMScoped (mismatch defs) (zipWith (curry $ mapHom snd) xargs yargs)) -- Data constructor vs. primitive constant recoverable defs (NDCon _ _ _ _ _) (NPrimVal _ _) = pure False recoverable defs (NPrimVal _ _) (NDCon _ _ _ _ _) = pure False @@ -284,7 +284,7 @@ findLinear top bound rig tm accessible _ r = r findLinArg : {vars : _} -> - RigCount -> NF [] -> List (Term vars) -> + RigCount -> NF SLNil -> List (Term vars) -> Core (List (Name, RigCount)) findLinArg rig ty@(NBind _ _ (Pi _ c _ _) _) (As fc u a p :: as) = if isLinear c @@ -387,7 +387,7 @@ checkLHS {vars} trans mult n opts nest env fc lhs_in lhs <- if trans then pure lhs_bound - else implicitsAs n defs vars lhs_bound + else implicitsAs n defs (toList vars) lhs_bound logC "declare.def.lhs" 5 $ do pure $ "Checking LHS of " ++ show !(getFullName (Resolved n)) -- todo: add Pretty RawImp instance @@ -633,24 +633,24 @@ checkClause {vars} mult vis totreq hashit n opts nest env bindWithArgs : (rig : RigCount) -> (wvalTy : Term xs) -> Maybe (Name, Term xs) -> (wvalEnv : Env Term xs) -> - Core (ext : List Name - ** ( Env Term (ext ++ xs) - , Term (ext ++ xs) - , (Term (ext ++ xs) -> Term xs) + Core (ext : ScopedList Name + ** ( Env Term (ext +%+ xs) + , Term (ext +%+ xs) + , (Term (ext +%+ xs) -> Term xs) )) bindWithArgs {xs} rig wvalTy Nothing wvalEnv = let wargn : Name wargn = MN "warg" 0 - wargs : List Name - wargs = [wargn] + wargs : ScopedList Name + wargs = (wargn :%: SLNil) - scenv : Env Term (wargs ++ xs) + scenv : Env Term (wargs +%+ xs) := Pi vfc top Explicit wvalTy :: wvalEnv - var : Term (wargs ++ xs) + var : Term (wargs +%+ xs) := Local vfc (Just False) Z First - binder : Term (wargs ++ xs) -> Term xs + binder : Term (wargs +%+ xs) -> Term xs := Bind vfc wargn (Pi vfc rig Explicit wvalTy) in pure (wargs ** (scenv, var, binder)) @@ -665,11 +665,11 @@ checkClause {vars} mult vis totreq hashit n opts nest env let wargn : Name wargn = MN "warg" 0 - wargs : List Name - wargs = [name, wargn] + wargs : ScopedList Name + wargs = (name :%: wargn :%: SLNil) wvalTy' := weaken wvalTy - eqTy : Term (MN "warg" 0 :: xs) + eqTy : Term (MN "warg" 0 :%: xs) := apply vfc eqTyCon [ wvalTy' , wvalTy' @@ -677,15 +677,15 @@ checkClause {vars} mult vis totreq hashit n opts nest env , Local vfc (Just False) Z First ] - scenv : Env Term (wargs ++ xs) + scenv : Env Term (wargs +%+ xs) := Pi vfc top Implicit eqTy :: Pi vfc top Explicit wvalTy :: wvalEnv - var : Term (wargs ++ xs) + var : Term (wargs +%+ xs) := Local vfc (Just False) (S Z) (Later First) - binder : Term (wargs ++ xs) -> Term xs + binder : Term (wargs +%+ xs) -> Term xs := \ t => Bind vfc wargn (Pi vfc rig Explicit wvalTy) $ Bind vfc name (Pi vfc rig Implicit eqTy) t @@ -696,7 +696,7 @@ checkClause {vars} mult vis totreq hashit n opts nest env -- function. Hence, turn it to Keep whatever keepOldEnv : {0 outer : _} -> {vs : _} -> (outprf : Thin outer vs) -> Thin vs' vs -> - (vs'' : List Name ** Thin vs'' vs) + (vs'' : ScopedList Name ** Thin vs'' vs) keepOldEnv {vs} Refl p = (vs ** Refl) keepOldEnv {vs} p Refl = (vs ** Refl) keepOldEnv (Drop p) (Drop p') @@ -736,13 +736,13 @@ checkClause {vars} mult vis totreq hashit n opts nest env pure (ImpossibleClause ploc newlhs) -- TODO: remove -nameListEq : (xs : List Name) -> (ys : List Name) -> Maybe (xs = ys) -nameListEq [] [] = Just Refl -nameListEq (x :: xs) (y :: ys) with (nameEq x y) - nameListEq (x :: xs) (x :: ys) | (Just Refl) with (nameListEq xs ys) - nameListEq (x :: xs) (x :: xs) | (Just Refl) | Just Refl= Just Refl - nameListEq (x :: xs) (x :: ys) | (Just Refl) | Nothing = Nothing - nameListEq (x :: xs) (y :: ys) | Nothing = Nothing +nameListEq : (xs : ScopedList Name) -> (ys : ScopedList Name) -> Maybe (xs = ys) +nameListEq SLNil SLNil = Just Refl +nameListEq (x :%: xs) (y :%: ys) with (nameEq x y) + nameListEq (x :%: xs) (x :%: ys) | (Just Refl) with (nameListEq xs ys) + nameListEq (x :%: xs) (x :%: xs) | (Just Refl) | Just Refl= Just Refl + nameListEq (x :%: xs) (x :%: ys) | (Just Refl) | Nothing = Nothing + nameListEq (x :%: xs) (y :%: ys) | Nothing = Nothing nameListEq _ _ = Nothing -- Calculate references for the given name, and recursively if they haven't @@ -1118,7 +1118,7 @@ processDef opts nest env fc n_in cs_in then pure Nothing else pure (Just tm)) where - closeEnv : Defs -> NF [] -> Core ClosedTerm + closeEnv : Defs -> NF SLNil -> Core ClosedTerm closeEnv defs (NBind _ x (PVar _ _ _ _) sc) = closeEnv defs !(sc defs (toClosure defaultOpts [] (Ref fc Bound x))) closeEnv defs nf = quote defs [] nf diff --git a/src/TTImp/ProcessFnOpt.idr b/src/TTImp/ProcessFnOpt.idr index a2db835821..ac6cf434f3 100644 --- a/src/TTImp/ProcessFnOpt.idr +++ b/src/TTImp/ProcessFnOpt.idr @@ -10,7 +10,7 @@ import TTImp.TTImp import Libraries.Data.NameMap -getRetTy : Defs -> NF [] -> Core Name +getRetTy : Defs -> NF SLNil -> Core Name getRetTy defs (NBind fc _ (Pi _ _ _ _) sc) = getRetTy defs !(sc defs (toClosure defaultOpts [] (Erased fc Placeholder))) getRetTy defs (NTCon _ n _ _ _) = pure n @@ -85,7 +85,7 @@ processFnOpt fc _ ndef (SpecArgs ns) else insertDeps (pos :: acc) ps ns -- Collect the argument names which the dynamic args depend on - collectDDeps : NF [] -> Core (List Name) + collectDDeps : NF SLNil -> Core (List Name) collectDDeps (NBind tfc x (Pi _ _ _ nty) sc) = do defs <- get Ctxt empty <- clearDefs defs @@ -101,14 +101,14 @@ processFnOpt fc _ ndef (SpecArgs ns) -- Return names the type depends on, and whether it's a parameter mutual - getDepsArgs : Bool -> List (NF []) -> NameMap Bool -> + getDepsArgs : Bool -> ScopedList (NF SLNil) -> NameMap Bool -> Core (NameMap Bool) - getDepsArgs inparam [] ns = pure ns - getDepsArgs inparam (a :: as) ns + getDepsArgs inparam SLNil ns = pure ns + getDepsArgs inparam (a :%: as) ns = do ns' <- getDeps inparam a ns getDepsArgs inparam as ns' - getDeps : Bool -> NF [] -> NameMap Bool -> + getDeps : Bool -> NF SLNil -> NameMap Bool -> Core (NameMap Bool) getDeps inparam (NBind _ x (Pi _ _ _ pty) sc) ns = do defs <- get Ctxt @@ -132,19 +132,19 @@ processFnOpt fc _ ndef (SpecArgs ns) params <- case !(lookupDefExact n (gamma defs)) of Just (TCon _ _ ps _ _ _ _ _) => pure ps _ => pure [] - let (ps, ds) = splitPs 0 params (map snd args) + let (ps, ds) = splitPs 0 params (map snd (toList args)) ns' <- getDepsArgs True !(traverse (evalClosure defs) ps) ns getDepsArgs False !(traverse (evalClosure defs) ds) ns' where -- Split into arguments in parameter position, and others - splitPs : Nat -> List Nat -> List (Closure []) -> - (List (Closure []), List (Closure [])) - splitPs n params [] = ([], []) + splitPs : Nat -> List Nat -> List (Closure SLNil) -> + (ScopedList (Closure SLNil), ScopedList (Closure SLNil)) + splitPs n params [] = (SLNil, SLNil) splitPs n params (x :: xs) = let (ps', ds') = splitPs (1 + n) params xs in if n `elem` params - then (x :: ps', ds') - else (ps', x :: ds') + then (x :%: ps', ds') + else (ps', x :%: ds') getDeps inparam (NDelayed _ _ t) ns = getDeps inparam t ns getDeps inparams nf ns = pure ns @@ -155,7 +155,7 @@ processFnOpt fc _ ndef (SpecArgs ns) List Name -> -- things depended on by dynamic args -- We're assuming it's a short list, so just use -- List and don't worry about duplicates. - List (Name, Nat) -> NF [] -> Core (List Nat) + List (Name, Nat) -> NF SLNil -> Core (List Nat) collectSpec acc ddeps ps (NBind tfc x (Pi _ _ _ nty) sc) = do defs <- get Ctxt empty <- clearDefs defs @@ -174,7 +174,7 @@ processFnOpt fc _ ndef (SpecArgs ns) else collectSpec acc ddeps ps sc' collectSpec acc ddeps ps _ = pure acc - getNamePos : Nat -> NF [] -> Core (List (Name, Nat)) + getNamePos : Nat -> NF SLNil -> Core (List (Name, Nat)) getNamePos i (NBind tfc x (Pi _ _ _ _) sc) = do defs <- get Ctxt ns' <- getNamePos (1 + i) !(sc defs (toClosure defaultOpts [] (Erased tfc Placeholder))) diff --git a/src/TTImp/ProcessParams.idr b/src/TTImp/ProcessParams.idr index 71a948c30d..40164b9c8e 100644 --- a/src/TTImp/ProcessParams.idr +++ b/src/TTImp/ProcessParams.idr @@ -43,7 +43,7 @@ processParams {vars} {c} {m} {u} nest env fc ps ds -- then read off the environment from the elaborated type. This way -- we'll get all the implicit names we need let pty_raw = mkParamTy ps - pty_imp <- bindTypeNames fc [] vars (IBindHere fc (PI erased) pty_raw) + pty_imp <- bindTypeNames fc [] (toList vars) (IBindHere fc (PI erased) pty_raw) log "declare.param" 10 $ "Checking " ++ show pty_imp u <- uniVar fc pty <- checkTerm (-1) InType [] diff --git a/src/TTImp/ProcessRecord.idr b/src/TTImp/ProcessRecord.idr index 5ea0037055..465d74a1bc 100644 --- a/src/TTImp/ProcessRecord.idr +++ b/src/TTImp/ProcessRecord.idr @@ -144,7 +144,7 @@ elabRecord {vars} eopts fc env nest newns def_vis mbtot tn_in params0 opts conNa Core (List ImpParameter) -- New telescope of parameters, including missing bindings preElabAsData tn = do let fc = virtualiseFC fc - let dataTy = IBindHere fc (PI erased) !(bindTypeNames fc [] vars (mkDataTy fc params0)) + let dataTy = IBindHere fc (PI erased) !(bindTypeNames fc [] (toList vars) (mkDataTy fc params0)) -- we don't use MkImpLater because users may have already declared the record ahead of time let dt = MkImpData fc tn (Just dataTy) opts [] log "declare.record" 10 $ "Pre-declare record data type: \{show dt}" @@ -165,9 +165,9 @@ elabRecord {vars} eopts fc env nest newns def_vis mbtot tn_in params0 opts conNa -- a LHS, or inside a `parameters` block) and so we need to start by dropping -- these local variables from the fully elaborated record's type -- We'll use the `env` thus obtained to unelab the remaining scope - dropLeadingPis : {vs : _} -> (vars : List Name) -> Term vs -> Env Term vs -> + dropLeadingPis : {vs : _} -> (vars : ScopedList Name) -> Term vs -> Env Term vs -> Core (vars' ** (Env Term vars', Term vars')) - dropLeadingPis [] ty env + dropLeadingPis SLNil ty env = do unless (null vars) $ logC "declare.record.parameters" 60 $ pure $ unlines [ "We elaborated \{show tn} in a non-empty local context." @@ -175,7 +175,7 @@ elabRecord {vars} eopts fc env nest newns def_vis mbtot tn_in params0 opts conNa , " Remaining type: \{show !(toFullNames ty)}" ] pure (_ ** (env, ty)) - dropLeadingPis (var :: vars) (Bind fc n b@(Pi _ _ _ _) ty) env + dropLeadingPis (var :%: vars) (Bind fc n b@(Pi _ _ _ _) ty) env = dropLeadingPis vars ty (b :: env) dropLeadingPis _ ty _ = throw (InternalError "Malformed record type \{show ty}") @@ -219,7 +219,7 @@ elabRecord {vars} eopts fc env nest newns def_vis mbtot tn_in params0 opts conNa = do let fc = virtualiseFC fc let conty = mkTy (paramTelescope params) $ mkTy (map farg fields) (recTy tn params) - let boundNames = paramNames params ++ map fname fields ++ vars + let boundNames = paramNames params ++ map fname fields ++ (toList vars) let con = MkImpTy (virtualiseFC fc) EmptyFC cname !(bindTypeNames fc [] boundNames conty) let dt = MkImpData fc tn Nothing opts [con] @@ -260,13 +260,13 @@ elabRecord {vars} eopts fc env nest newns def_vis mbtot tn_in params0 opts conNa unNameNS <- inCurrentNS (UN $ Basic fldNameStr) ty <- unelabNest !nestDrop tyenv ty_chk - let ty' = substNames vars upds $ map rawName ty + let ty' = substNames (toList vars) upds $ map rawName ty log "declare.record.field" 5 $ "Field type: " ++ show ty' let rname = MN "rec" 0 -- Claim the projection type projTy <- bindTypeNames fc [] - (map fst params ++ map fname fields ++ vars) $ + (map fst params ++ map fname fields ++ toList vars) $ mkTy (paramTelescope params) $ IPi bfc top Explicit (Just rname) (recTy tn params) ty' diff --git a/src/TTImp/ProcessType.idr b/src/TTImp/ProcessType.idr index 8937caa1e5..033a9877bf 100644 --- a/src/TTImp/ProcessType.idr +++ b/src/TTImp/ProcessType.idr @@ -91,32 +91,32 @@ initDef fc n env ty (_ :: opts) = initDef fc n env ty opts -- generalising partially evaluated definitions and (potentially) in interactive -- editing findInferrable : {auto c : Ref Ctxt Defs} -> - Defs -> NF [] -> Core (List Nat) + Defs -> NF SLNil -> Core (List Nat) findInferrable defs ty = fi 0 0 [] [] ty where mutual -- Add to the inferrable arguments from the given type. An argument is -- inferrable if it's guarded by a constructor, or on its own findInf : List Nat -> List (Name, Nat) -> - NF [] -> Core (List Nat) - findInf acc pos (NApp _ (NRef Bound n) []) + NF SLNil -> Core (List Nat) + findInf acc pos (NApp _ (NRef Bound n) SLNil) = case lookup n pos of Nothing => pure acc Just p => if p `elem` acc then pure acc else pure (p :: acc) findInf acc pos (NDCon _ _ _ _ args) - = do args' <- traverse (evalClosure defs . snd) args + = do args' <- traverse (evalClosure defs . snd) (toList args) findInfs acc pos args' findInf acc pos (NTCon _ _ _ _ args) - = do args' <- traverse (evalClosure defs . snd) args + = do args' <- traverse (evalClosure defs . snd) (toList args) findInfs acc pos args' findInf acc pos (NDelayed _ _ t) = findInf acc pos t findInf acc _ _ = pure acc - findInfs : List Nat -> List (Name, Nat) -> List (NF []) -> Core (List Nat) + findInfs : List Nat -> List (Name, Nat) -> List (NF SLNil) -> Core (List Nat) findInfs acc pos [] = pure acc findInfs acc pos (n :: ns) = findInf !(findInfs acc pos ns) pos n - fi : Nat -> Int -> List (Name, Nat) -> List Nat -> NF [] -> Core (List Nat) + fi : Nat -> Int -> List (Name, Nat) -> List Nat -> NF SLNil -> Core (List Nat) fi pos i args acc (NBind fc x (Pi _ _ _ aty) sc) = do let argn = MN "inf" i sc' <- sc defs (toClosure defaultOpts [] (Ref fc Bound argn)) diff --git a/src/TTImp/Reflect.idr b/src/TTImp/Reflect.idr index 618ce77c1e..e445099e86 100644 --- a/src/TTImp/Reflect.idr +++ b/src/TTImp/Reflect.idr @@ -16,7 +16,7 @@ export Reify BindMode where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of - (UN (Basic "PI"), [(_, c)]) + (UN (Basic "PI"), ((_, c) :%: SLNil)) => do c' <- reify defs !(evalClosure defs c) pure (PI c') (UN (Basic "PATTERN"), _) => pure PATTERN @@ -90,11 +90,11 @@ mutual Reify RawImp where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), map snd args) of - (UN (Basic "IVar"), [fc, n]) + (UN (Basic "IVar"), (fc :%: n :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) n' <- reify defs !(evalClosure defs n) pure (IVar fc' n') - (UN (Basic "IPi"), [fc, c, p, mn, aty, rty]) + (UN (Basic "IPi"), (fc :%: c :%: p :%: mn :%: aty :%: rty :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) c' <- reify defs !(evalClosure defs c) p' <- reify defs !(evalClosure defs p) @@ -102,7 +102,7 @@ mutual aty' <- reify defs !(evalClosure defs aty) rty' <- reify defs !(evalClosure defs rty) pure (IPi fc' c' p' mn' aty' rty') - (UN (Basic "ILam"), [fc, c, p, mn, aty, lty]) + (UN (Basic "ILam"), (fc :%: c :%: p :%: mn :%: aty :%: lty :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) c' <- reify defs !(evalClosure defs c) p' <- reify defs !(evalClosure defs p) @@ -110,7 +110,7 @@ mutual aty' <- reify defs !(evalClosure defs aty) lty' <- reify defs !(evalClosure defs lty) pure (ILam fc' c' p' mn' aty' lty') - (UN (Basic "ILet"), [fc, lhsFC, c, n, ty, val, sc]) + (UN (Basic "ILet"), (fc :%: lhsFC :%: c :%: n :%: ty :%: val :%: sc :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) lhsFC' <- reify defs !(evalClosure defs lhsFC) c' <- reify defs !(evalClosure defs c) @@ -119,124 +119,124 @@ mutual val' <- reify defs !(evalClosure defs val) sc' <- reify defs !(evalClosure defs sc) pure (ILet fc' lhsFC' c' n' ty' val' sc') - (UN (Basic "ICase"), [fc, opts, sc, ty, cs]) + (UN (Basic "ICase"), (fc :%: opts :%: sc :%: ty :%: cs :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) opts' <- reify defs !(evalClosure defs opts) sc' <- reify defs !(evalClosure defs sc) ty' <- reify defs !(evalClosure defs ty) cs' <- reify defs !(evalClosure defs cs) pure (ICase fc' opts' sc' ty' cs') - (UN (Basic "ILocal"), [fc, ds, sc]) + (UN (Basic "ILocal"), (fc :%: ds :%: sc :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) ds' <- reify defs !(evalClosure defs ds) sc' <- reify defs !(evalClosure defs sc) pure (ILocal fc' ds' sc') - (UN (Basic "IUpdate"), [fc, ds, sc]) + (UN (Basic "IUpdate"), (fc :%: ds :%: sc :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) ds' <- reify defs !(evalClosure defs ds) sc' <- reify defs !(evalClosure defs sc) pure (IUpdate fc' ds' sc') - (UN (Basic "IApp"), [fc, f, a]) + (UN (Basic "IApp"), (fc :%: f :%: a :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) f' <- reify defs !(evalClosure defs f) a' <- reify defs !(evalClosure defs a) pure (IApp fc' f' a') - (UN (Basic "INamedApp"), [fc, f, m, a]) + (UN (Basic "INamedApp"), (fc :%: f :%: m :%: a :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) f' <- reify defs !(evalClosure defs f) m' <- reify defs !(evalClosure defs m) a' <- reify defs !(evalClosure defs a) pure (INamedApp fc' f' m' a') - (UN (Basic "IAutoApp"), [fc, f, a]) + (UN (Basic "IAutoApp"), (fc :%: f :%: a :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) f' <- reify defs !(evalClosure defs f) a' <- reify defs !(evalClosure defs a) pure (IAutoApp fc' f' a') - (UN (Basic "IWithApp"), [fc, f, a]) + (UN (Basic "IWithApp"), (fc :%: f :%: a :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) f' <- reify defs !(evalClosure defs f) a' <- reify defs !(evalClosure defs a) pure (IWithApp fc' f' a') - (UN (Basic "ISearch"), [fc, d]) + (UN (Basic "ISearch"), (fc :%: d :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) d' <- reify defs !(evalClosure defs d) pure (ISearch fc' d') - (UN (Basic "IAlternative"), [fc, t, as]) + (UN (Basic "IAlternative"), (fc :%: t :%: as :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) t' <- reify defs !(evalClosure defs t) as' <- reify defs !(evalClosure defs as) pure (IAlternative fc' t' as') - (UN (Basic "IRewrite"), [fc, t, sc]) + (UN (Basic "IRewrite"), (fc :%: t :%: sc :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) t' <- reify defs !(evalClosure defs t) sc' <- reify defs !(evalClosure defs sc) pure (IRewrite fc' t' sc') - (UN (Basic "IBindHere"), [fc, t, sc]) + (UN (Basic "IBindHere"), (fc :%: t :%: sc :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) t' <- reify defs !(evalClosure defs t) sc' <- reify defs !(evalClosure defs sc) pure (IBindHere fc' t' sc') - (UN (Basic "IBindVar"), [fc, n]) + (UN (Basic "IBindVar"), (fc :%: n :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) n' <- reify defs !(evalClosure defs n) pure (IBindVar fc' n') - (UN (Basic "IAs"), [fc, nameFC, s, n, t]) + (UN (Basic "IAs"), (fc :%: nameFC :%: s :%: n :%: t :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) nameFC' <- reify defs !(evalClosure defs nameFC) s' <- reify defs !(evalClosure defs s) n' <- reify defs !(evalClosure defs n) t' <- reify defs !(evalClosure defs t) pure (IAs fc' nameFC' s' n' t') - (UN (Basic "IMustUnify"), [fc, r, t]) + (UN (Basic "IMustUnify"), (fc :%: r :%: t :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) r' <- reify defs !(evalClosure defs r) t' <- reify defs !(evalClosure defs t) pure (IMustUnify fc' r' t') - (UN (Basic "IDelayed"), [fc, r, t]) + (UN (Basic "IDelayed"), (fc :%: r :%: t :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) r' <- reify defs !(evalClosure defs r) t' <- reify defs !(evalClosure defs t) pure (IDelayed fc' r' t') - (UN (Basic "IDelay"), [fc, t]) + (UN (Basic "IDelay"), (fc :%: t :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) t' <- reify defs !(evalClosure defs t) pure (IDelay fc' t') - (UN (Basic "IForce"), [fc, t]) + (UN (Basic "IForce"), (fc :%: t :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) t' <- reify defs !(evalClosure defs t) pure (IForce fc' t') - (UN (Basic "IQuote"), [fc, t]) + (UN (Basic "IQuote"), (fc :%: t :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) t' <- reify defs !(evalClosure defs t) pure (IQuote fc' t') - (UN (Basic "IQuoteName"), [fc, t]) + (UN (Basic "IQuoteName"), (fc :%: t :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) t' <- reify defs !(evalClosure defs t) pure (IQuoteName fc' t') - (UN (Basic "IQuoteDecl"), [fc, t]) + (UN (Basic "IQuoteDecl"), (fc :%: t :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) t' <- reify defs !(evalClosure defs t) pure (IQuoteDecl fc' t') - (UN (Basic "IUnquote"), [fc, t]) + (UN (Basic "IUnquote"), (fc :%: t :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) t' <- reify defs !(evalClosure defs t) pure (IUnquote fc' t') - (UN (Basic "IPrimVal"), [fc, t]) + (UN (Basic "IPrimVal"), (fc :%: t :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) t' <- reify defs !(evalClosure defs t) pure (IPrimVal fc' t') - (UN (Basic "IType"), [fc]) + (UN (Basic "IType"), (fc :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) pure (IType fc') - (UN (Basic "IHole"), [fc, n]) + (UN (Basic "IHole"), (fc :%: n :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) n' <- reify defs !(evalClosure defs n) pure (IHole fc' n') - (UN (Basic "Implicit"), [fc, n]) + (UN (Basic "Implicit"), (fc :%: n :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) n' <- reify defs !(evalClosure defs n) pure (Implicit fc' n') - (UN (Basic "IWithUnambigNames"), [fc, ns, t]) + (UN (Basic "IWithUnambigNames"), (fc :%: ns :%: t :%: SLNil)) => do fc' <- reify defs !(evalClosure defs fc) ns' <- reify defs !(evalClosure defs ns) t' <- reify defs !(evalClosure defs t) @@ -248,11 +248,11 @@ mutual Reify IFieldUpdate where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of - (UN (Basic "ISetField"), [(_, x), (_, y)]) + (UN (Basic "ISetField"), ((_, x) :%: (_, y) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) pure (ISetField x' y') - (UN (Basic "ISetFieldApp"), [(_, x), (_, y)]) + (UN (Basic "ISetFieldApp"), ((_, x) :%: (_, y) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) pure (ISetFieldApp x' y') @@ -267,7 +267,7 @@ mutual => pure FirstSuccess (UN (Basic "Unique"), _) => pure Unique - (UN (Basic "UniqueDefault"), [(_, x)]) + (UN (Basic "UniqueDefault"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (UniqueDefault x') _ => cantReify val "AltType" @@ -282,25 +282,25 @@ mutual (UN (Basic "NoInline"), _) => pure NoInline (UN (Basic "Deprecate"), _) => pure Deprecate (UN (Basic "TCInline"), _) => pure TCInline - (UN (Basic "Hint"), [(_, x)]) + (UN (Basic "Hint"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (Hint x') - (UN (Basic "GlobalHint"), [(_, x)]) + (UN (Basic "GlobalHint"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (GlobalHint x') (UN (Basic "ExternFn"), _) => pure ExternFn - (UN (Basic "ForeignFn"), [(_, x)]) + (UN (Basic "ForeignFn"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (ForeignFn x') - (UN (Basic "ForeignExport"), [(_, x)]) + (UN (Basic "ForeignExport"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (ForeignExport x') (UN (Basic "Invertible"), _) => pure Invertible - (UN (Basic "Totality"), [(_, x)]) + (UN (Basic "Totality"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (Totality x') (UN (Basic "Macro"), _) => pure Macro - (UN (Basic "SpecArgs"), [(_, x)]) + (UN (Basic "SpecArgs"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (SpecArgs x') _ => cantReify val "FnOpt" @@ -310,7 +310,7 @@ mutual Reify ImpTy where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), map snd args) of - (UN (Basic "MkTy"), [w, x, y, z]) + (UN (Basic "MkTy"), (w :%: x :%: y :%: z :%: SLNil)) => do w' <- reify defs !(evalClosure defs w) x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) @@ -323,7 +323,7 @@ mutual Reify DataOpt where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), args) of - (UN (Basic "SearchBy"), [(_, x)]) + (UN (Basic "SearchBy"), ((_, x) :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (SearchBy x') (UN (Basic "NoHints"), _) => pure NoHints @@ -337,14 +337,14 @@ mutual Reify ImpData where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), map snd args) of - (UN (Basic "MkData"), [v,w,x,y,z]) + (UN (Basic "MkData"), (v :%: w :%: x :%: y :%: z :%: SLNil)) => do v' <- reify defs !(evalClosure defs v) w' <- reify defs !(evalClosure defs w) x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) z' <- reify defs !(evalClosure defs z) pure (MkImpData v' w' x' y' z') - (UN (Basic "MkLater"), [x,y,z]) + (UN (Basic "MkLater"), (x :%: y :%: z :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) z' <- reify defs !(evalClosure defs z) @@ -356,7 +356,7 @@ mutual Reify IField where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), map snd args) of - (UN (Basic "MkIField"), [v,w,x,y,z]) + (UN (Basic "MkIField"), (v :%: w :%: x :%: y :%: z :%: SLNil)) => do v' <- reify defs !(evalClosure defs v) w' <- reify defs !(evalClosure defs w) x' <- reify defs !(evalClosure defs x) @@ -370,7 +370,7 @@ mutual Reify ImpRecord where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), map snd args) of - (UN (Basic "MkRecord"), [v,w,x,y,z,a]) + (UN (Basic "MkRecord"), (v :%: w :%: x :%: y :%: z :%: a :%: SLNil)) => do v' <- reify defs !(evalClosure defs v) w' <- reify defs !(evalClosure defs w) x' <- reify defs !(evalClosure defs x) @@ -385,7 +385,7 @@ mutual Reify WithFlag where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), map snd args) of - (UN (Basic "Syntactic"), []) + (UN (Basic "Syntactic"), SLNil) => pure Syntactic _ => cantReify val "WithFlag" reify defs val = cantReify val "WithFlag" @@ -394,12 +394,12 @@ mutual Reify ImpClause where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), map snd args) of - (UN (Basic "PatClause"), [x,y,z]) + (UN (Basic "PatClause"), (x :%: y :%: z :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) z' <- reify defs !(evalClosure defs z) pure (PatClause x' y' z') - (UN (Basic "WithClause"), [u,v,w,x,y,z,a]) + (UN (Basic "WithClause"), (u :%: v :%: w :%: x :%: y :%: z :%: a :%: SLNil)) => do u' <- reify defs !(evalClosure defs u) v' <- reify defs !(evalClosure defs v) w' <- reify defs !(evalClosure defs w) @@ -408,7 +408,7 @@ mutual z' <- reify defs !(evalClosure defs z) a' <- reify defs !(evalClosure defs a) pure (WithClause u' v' w' x' y' z' a') - (UN (Basic "ImpossibleClause"), [x,y]) + (UN (Basic "ImpossibleClause"), (x :%: y :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) pure (ImpossibleClause x' y') @@ -419,53 +419,53 @@ mutual Reify ImpDecl where reify defs val@(NDCon _ n _ _ args) = case (dropAllNS !(full (gamma defs) n), map snd args) of - (UN (Basic "IClaim"), [v,w,x,y,z]) + (UN (Basic "IClaim"), (v :%: w :%: x :%: y :%: z :%: SLNil)) => do v' <- reify defs !(evalClosure defs v) w' <- reify defs !(evalClosure defs w) x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) z' <- reify defs !(evalClosure defs z) pure (IClaim v' w' x' y' z') - (UN (Basic "IData"), [x,y,z,w]) + (UN (Basic "IData"), (x :%: y :%: z :%: w :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) z' <- reify defs !(evalClosure defs z) w' <- reify defs !(evalClosure defs w) pure (IData x' y' z' w') - (UN (Basic "IDef"), [x,y,z]) + (UN (Basic "IDef"), (x :%: y :%: z :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) z' <- reify defs !(evalClosure defs z) pure (IDef x' y' z') - (UN (Basic "IParameters"), [x,y,z]) + (UN (Basic "IParameters"), (x :%: y :%: z :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) z' <- reify defs !(evalClosure defs z) pure (IParameters x' y' z') - (UN (Basic "IRecord"), [w,x,y,z,u]) + (UN (Basic "IRecord"), (w :%: x :%: y :%: z :%: u :%: SLNil)) => do w' <- reify defs !(evalClosure defs w) x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) z' <- reify defs !(evalClosure defs z) u' <- reify defs !(evalClosure defs u) pure (IRecord w' x' y' z' u') - (UN (Basic "IFail"), [w,x,y]) + (UN (Basic "IFail"), (w :%: x :%: y :%: SLNil)) => do w' <- reify defs !(evalClosure defs w) x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) pure (IFail w' x' y') - (UN (Basic "INamespace"), [w,x,y]) + (UN (Basic "INamespace"), (w :%: x :%: y :%: SLNil)) => do w' <- reify defs !(evalClosure defs w) x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) pure (INamespace w' x' y') - (UN (Basic "ITransform"), [w,x,y,z]) + (UN (Basic "ITransform"), (w :%: x :%: y :%: z :%: SLNil)) => do w' <- reify defs !(evalClosure defs w) x' <- reify defs !(evalClosure defs x) y' <- reify defs !(evalClosure defs y) z' <- reify defs !(evalClosure defs z) pure (ITransform w' x' y' z') - (UN (Basic "ILog"), [x]) + (UN (Basic "ILog"), (x :%: SLNil)) => do x' <- reify defs !(evalClosure defs x) pure (ILog x') _ => cantReify val "Decl" diff --git a/src/TTImp/TTImp.idr b/src/TTImp/TTImp.idr index c3087d4d7a..2b9e1719f1 100644 --- a/src/TTImp/TTImp.idr +++ b/src/TTImp/TTImp.idr @@ -20,7 +20,7 @@ import Libraries.Data.WithDefault -- Information about names in nested blocks public export -record NestedNames (vars : List Name) where +record NestedNames (vars : ScopedList Name) where constructor MkNested -- A map from names to the decorated version of the name, and the new name -- applied to its enclosing environment @@ -35,7 +35,7 @@ Weaken NestedNames where weakenNs {ns = wkns} s (MkNested ns) = MkNested (map wknName ns) where wknName : (Name, (Maybe Name, List (Var vars), FC -> NameType -> Term vars)) -> - (Name, (Maybe Name, List (Var (wkns ++ vars)), FC -> NameType -> Term (wkns ++ vars))) + (Name, (Maybe Name, List (Var (wkns +%+ vars)), FC -> NameType -> Term (wkns +%+ vars))) wknName (n, (mn, vars, rep)) = (n, (mn, map (weakenNs s) vars, \fc, nt => weakenNs s (rep fc nt))) @@ -722,7 +722,7 @@ implicitsAs n defs ns tm -- in the lhs: this is used to determine when to stop searching for further -- implicits to add. findImps : List (Maybe Name) -> List (Maybe Name) -> - List Name -> NF [] -> + List Name -> NF SLNil -> Core (List (Name, PiInfo RawImp)) -- #834 When we are in a local definition, we have an explicit telescope -- corresponding to the variables bound in the parent function. diff --git a/src/TTImp/Unelab.idr b/src/TTImp/Unelab.idr index f4b6361b29..dbacdcf1f3 100644 --- a/src/TTImp/Unelab.idr +++ b/src/TTImp/Unelab.idr @@ -115,7 +115,7 @@ mutual = TForce fc r (substVars xs y) substVars xs tm = tm - substArgs : SizeOf vs -> List (List (Var vs), Term vars) -> Term vs -> Term (vs ++ vars) + substArgs : SizeOf vs -> List (List (Var vs), Term vars) -> Term vs -> Term (vs +%+ vars) substArgs p substs tm = let substs' = map (bimap (map $ embed {outer = vars}) (weakenNs p)) substs @@ -155,7 +155,7 @@ mutual (argpos : Nat) -> List (Term vars) -> Core (Maybe IRawImp) mkCase pats argpos args = do unless (null args) $ log "unelab.case.clause" 20 $ - unwords $ "Ignoring" :: map show args + unwords $ "Ignoring" :: map show (toList $ args) let Just scrutinee = idxOrMaybe argpos args | _ => pure Nothing fc = getLoc scrutinee @@ -252,7 +252,7 @@ mutual next (NS ns n) = NS ns (next n) next n = MN (show n) 0 - uniqueLocal : List Name -> Name -> Name + uniqueLocal : ScopedList Name -> Name -> Name uniqueLocal vs n = if n `elem` vs then uniqueLocal vs (next n) @@ -329,8 +329,8 @@ mutual (umode : UnelabMode) -> (nest : List (Name, Nat)) -> FC -> Env Term vars -> (x : Name) -> - Binder (Term vars) -> Term (x :: vars) -> - IRawImp -> Term (x :: vars) -> + Binder (Term vars) -> Term (x :%: vars) -> + IRawImp -> Term (x :%: vars) -> Core (IRawImp, Glued vars) unelabBinder umode nest fc env x (Lam fc' rig p ty) sctm sc scty = do (ty', _) <- unelabTy umode nest env ty @@ -401,7 +401,7 @@ unelabNest : {vars : _} -> Env Term vars -> Term vars -> Core IRawImp unelabNest nest env (Meta fc n i args) - = do let mkn = nameRoot n ++ showScope args + = do let mkn = nameRoot n ++ (showScope $ toList args) pure (IHole fc mkn) where toName : Term vars -> Maybe Name diff --git a/src/TTImp/Utils.idr b/src/TTImp/Utils.idr index 9f7f562c1f..7c4e531af0 100644 --- a/src/TTImp/Utils.idr +++ b/src/TTImp/Utils.idr @@ -592,7 +592,7 @@ getArgName defs x bound allvars ty findNamesM : NF vars -> Core (Maybe (List String)) findNamesM (NBind _ x (Pi _ _ _ _) _) = pure (Just ["f", "g"]) - findNamesM (NTCon _ n _ d [(_, v)]) = do + findNamesM (NTCon _ n _ d ((_, v) :%: SLNil)) = do case dropNS !(full (gamma defs) n) of UN (Basic "List") => do nf <- evalClosure defs v