@@ -19,6 +19,8 @@ import Data.String
19
19
import Data.Vect
20
20
import Libraries.Data.List.LengthMatch
21
21
import Libraries.Data.SortedSet
22
+ import Libraries.Data.SnocList.SizeOf
23
+ import Libraries.Data.SnocList.LengthMatch
22
24
23
25
import Decidable.Equality
24
26
@@ -476,8 +478,8 @@ nextNames {vars} fc root (pats :< p) fty
476
478
newPats : (pargs : SnocList Pat) -> LengthMatch pargs ns ->
477
479
NamedPats vars (todo ++ ns) ->
478
480
NamedPats vars ns
479
- newPats [< ] NilMatch rest = []
480
- newPats (xs : < newpat) (ConsMatch w) (pi :: rest)
481
+ newPats [< ] LinMatch rest = []
482
+ newPats (xs : < newpat) (SnocMatch w) (pi :: rest)
481
483
= { pat : = newpat} pi :: newPats xs w rest
482
484
483
485
updateNames : SnocList (Name, Pat) -> SnocList (Name, Name)
@@ -600,7 +602,7 @@ groupCons fc fn pvars cs
600
602
((DelayGroup {tyarg} {valarg} ((MkPatClause pvars ps tid tm) :: rest)) :: gs)
601
603
| (DelayMatch {tyarg} {valarg})
602
604
= do let l = mkSizeOf [< valarg, tyarg]
603
- let newps = newPats [< parg, pty] (ConsMatch ( ConsMatch NilMatch )) ps
605
+ let newps = newPats [< parg, pty] (SnocMatch ( SnocMatch LinMatch )) ps
604
606
let pats' = updatePatNames (updateNames [< (valarg, parg), (tyarg, pty)])
605
607
(weakenNs l pats)
606
608
let newclause : PatClause (vars' :< valarg :< tyarg)
@@ -685,10 +687,10 @@ data ScoredPats : SnocList Name -> SnocList Name -> Type where
685
687
zeroedScore : {ps : _} -> List (NamedPats ns (ps :< p)) -> ScoredPats ns (ps :< p)
686
688
zeroedScore nps = Scored nps (replicate (S $ length ps) 0 )
687
689
688
- ||| Proof that a value `v` inserted in the middle of a list with
689
- ||| prefix `ps` and suffix `qs` can equivalently be snoced with
690
+ ||| Proof that a value `v` inserted in the middle of a snoc list with
691
+ ||| prefix `ps` and suffix `qs` can equivalently be consed with
690
692
||| `ps` or consed with `qs` before appending `qs` to `ps`.
691
- elemInsertedMiddle : (v : a) -> (ps,qs : SnocList a) -> ((qs :< v) ++ ps) = (qs ++ (ps `snoc` v ))
693
+ elemInsertedMiddle : (v : a) -> (ps,qs : SnocList a) -> ((qs :< v) ++ ps) = (qs ++ (v `cons` ps ))
692
694
elemInsertedMiddle v [< ] qs = Refl
693
695
elemInsertedMiddle v (xs : < x) qs = rewrite elemInsertedMiddle v xs qs in Refl
694
696
@@ -704,7 +706,7 @@ highScore : {prev : SnocList Name} ->
704
706
highScore [< ] [] high idx True = Nothing
705
707
highScore [< ] [] high idx False = Just idx
706
708
highScore (xs : < x) (y :: ys) high idx duped =
707
- let next = highScore {prev = prev `snoc` x } xs ys
709
+ let next = highScore {prev = x `cons` prev } xs ys
708
710
prf = elemInsertedMiddle x prev xs
709
711
in rewrite prf in
710
712
case compare y high of
@@ -1097,7 +1099,7 @@ mkPat args orig (Ref fc (DataCon t a) n) = pure $ PCon fc n t a args
1097
1099
mkPat args orig (Ref fc (TyCon t a) n) = pure $ PTyCon fc n a args
1098
1100
mkPat args orig (Ref fc Func n)
1099
1101
= do prims <- getPrimitiveNames
1100
- mtm <- normalisePrims (const True ) isPConst True prims n args orig []
1102
+ mtm <- normalisePrims (const True ) isPConst True prims n args orig [< ]
1101
1103
case mtm of
1102
1104
Just tm => if tm /= orig -- check we made progress; if there's an
1103
1105
-- unresolved interface, we might be stuck
@@ -1146,34 +1148,34 @@ mkPatClause fc fn args ty pid (ps, rhs)
1146
1148
= maybe (throw (CaseCompile fc fn DifferingArgNumbers ))
1147
1149
(\ eq =>
1148
1150
do defs <- get Ctxt
1149
- nty <- nf defs [] ty
1151
+ nty <- nf defs [< ] ty
1150
1152
ns <- mkNames args ps eq (Just nty)
1151
1153
log " compile.casetree" 20 $
1152
1154
" Make pat clause for names " ++ show ns
1153
1155
++ " in LHS " ++ show ps
1154
1156
pure (MkPatClause [] ns pid
1155
- (rewrite sym (appendNilRightNeutral args) in
1157
+ (rewrite sym (appendLinLeftNeutral args) in
1156
1158
(weakenNs (mkSizeOf args) rhs))))
1157
1159
(checkLengthMatch args ps)
1158
1160
where
1159
1161
mkNames : (vars : SnocList Name) -> (ps : SnocList Pat) ->
1160
1162
LengthMatch vars ps -> Maybe (NF [<]) ->
1161
1163
Core (NamedPats vars vars)
1162
- mkNames [< ] [< ] NilMatch fty = pure []
1163
- mkNames (args : < arg) (ps :< p) (ConsMatch eq) fty
1164
+ mkNames [< ] [< ] LinMatch fty = pure []
1165
+ mkNames (args : < arg) (ps :< p) (SnocMatch eq) fty
1164
1166
= do defs <- get Ctxt
1165
1167
empty <- clearDefs defs
1166
1168
fa_tys <- the (Core (Maybe _ , ArgType _ )) $
1167
1169
case fty of
1168
1170
Nothing => pure (Nothing , CaseBuilder . Unknown )
1169
1171
Just (NBind pfc _ (Pi _ c _ farg) fsc) =>
1170
- pure (Just ! (fsc defs (toClosure defaultOpts [] (Ref pfc Bound arg))),
1172
+ pure (Just ! (fsc defs (toClosure defaultOpts [< ] (Ref pfc Bound arg))),
1171
1173
Known c (embed {outer = args : < arg}
1172
- ! (quote empty [] farg)))
1174
+ ! (quote empty [< ] farg)))
1173
1175
Just t =>
1174
1176
pure (Nothing ,
1175
1177
Stuck (embed {outer = args : < arg}
1176
- ! (quote empty [] t)))
1178
+ ! (quote empty [< ] t)))
1177
1179
pure (MkInfo p First (Builtin . snd fa_tys)
1178
1180
:: weaken ! (mkNames args ps eq (Builtin . fst fa_tys)))
1179
1181
@@ -1199,7 +1201,7 @@ patCompile fc fn phase ty (p :: ps) def
1199
1201
log " compile.casetree" 10 $ show pats
1200
1202
i <- newRef PName (the Int 0 )
1201
1203
cases <- match fc fn phase pats
1202
- (rewrite sym (appendNilRightNeutral ns) in
1204
+ (rewrite sym (appendLinLeftNeutral ns) in
1203
1205
map (weakenNs n) def)
1204
1206
pure (_ ** cases)
1205
1207
where
@@ -1216,7 +1218,7 @@ patCompile fc fn phase ty (p :: ps) def
1216
1218
getNames i [< ] = ([< ] ** zero)
1217
1219
getNames i (xs : < x) =
1218
1220
let (ns ** n) = getNames (i + 1 ) xs
1219
- in (MN " arg" ns : < i ** suc n)
1221
+ in (ns : < MN "arg" i ** suc n)
1220
1222
1221
1223
toPatClause : {auto c : Ref Ctxt Defs} ->
1222
1224
FC -> Name -> (ClosedTerm, ClosedTerm) ->
@@ -1346,12 +1348,12 @@ getPMDef : {auto c : Ref Ctxt Defs} ->
1346
1348
getPMDef fc phase fn ty []
1347
1349
= do log " compile.casetree.getpmdef" 20 " getPMDef: No clauses!"
1348
1350
defs <- get Ctxt
1349
- pure (! (getArgs 0 ! (nf defs [] ty)) ** (Unmatched " No clauses" , []))
1351
+ pure (! (getArgs 0 ! (nf defs [< ] ty)) ** (Unmatched " No clauses" , []))
1350
1352
where
1351
1353
getArgs : Int -> NF [<] -> Core (SnocList Name)
1352
1354
getArgs i (NBind fc x (Pi _ _ _ _ ) sc)
1353
1355
= do defs <- get Ctxt
1354
- sc' <- sc defs (toClosure defaultOpts [] (Erased fc Placeholder ))
1356
+ sc' <- sc defs (toClosure defaultOpts [< ] (Erased fc Placeholder ))
1355
1357
pure (! (getArgs i sc') :< MN " arg" i)
1356
1358
getArgs i _ = pure [< ]
1357
1359
getPMDef fc phase fn ty clauses
0 commit comments