Skip to content

Commit 8230ceb

Browse files
spcfoxGulinSS
authored andcommitted
ScopedSnocList: Cleanup and refactor
1 parent 0a3ca35 commit 8230ceb

File tree

17 files changed

+112
-111
lines changed

17 files changed

+112
-111
lines changed

libs/base/Data/SnocList/Operations.idr

Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -96,6 +96,17 @@ lengthHomomorphism sx (sy :< x) = Calc $
9696
~~ 1 + (length sx + length sy) ...(cong (1+) $ lengthHomomorphism _ _)
9797
~~ length sx + (1 + length sy) ...(plusSuccRightSucc _ _)
9898

99+
export
100+
lengthDistributesOverFish : (sx : SnocList a) -> (ys : List a) ->
101+
length (sx <>< ys) === length sx + length ys
102+
lengthDistributesOverFish sx [] = sym $ plusZeroRightNeutral _
103+
lengthDistributesOverFish sx (y :: ys) = Calc $
104+
|~ length ((sx :< y) <>< ys)
105+
~~ length (sx :< y) + length ys ...( lengthDistributesOverFish (sx :< y) ys)
106+
~~ S (length sx) + length ys ...( Refl )
107+
~~ length sx + S (length ys) ...( plusSuccRightSucc _ _ )
108+
~~ length sx + length (y :: ys) ...( Refl )
109+
99110
-- cons-list operations on snoc-lists
100111

101112
||| Take `n` first elements from `sx`, returning the whole list if

src/Compiler/CompileExpr.idr

Lines changed: 0 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -23,12 +23,6 @@ import Libraries.Data.SnocList.Extra
2323

2424
%default covering
2525

26-
-- For ease of type level reasoning!
27-
public export
28-
rev : SnocList a -> SnocList a
29-
rev [<] = [<]
30-
rev (xs :< x) = [<x] ++ rev xs
31-
3226
data Args
3327
= NewTypeBy Nat Nat
3428
| EraseArgs Nat (List Nat)

src/Compiler/Inline.idr

Lines changed: 1 addition & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -18,14 +18,11 @@ import Data.Maybe
1818
import Data.List
1919
import Data.SnocList
2020
import Data.Vect
21-
import Libraries.Data.List.LengthMatch
21+
2222
import Libraries.Data.NameMap
2323
import Libraries.Data.WithDefault
24-
2524
import Libraries.Data.List.SizeOf
26-
import Libraries.Data.SnocList.LengthMatch
2725
import Libraries.Data.SnocList.SizeOf
28-
import Libraries.Data.SnocList.HasLength
2926
import Libraries.Data.SnocList.Extra
3027

3128
%default covering

src/Compiler/LambdaLift.idr

Lines changed: 8 additions & 19 deletions
Original file line numberDiff line numberDiff line change
@@ -18,6 +18,7 @@ import Core.TT
1818

1919
import Data.List
2020
import Data.SnocList
21+
import Data.SnocList.Operations
2122
import Data.Vect
2223

2324
import Libraries.Data.SnocList.Extra
@@ -360,31 +361,19 @@ record Used (vars : SnocList Name) where
360361
initUsed : {vars : _} -> Used vars
361362
initUsed {vars} = MkUsed (replicate (length vars) False)
362363

363-
lengthDistributesOverAppend
364-
: (xs, ys : SnocList a)
365-
-> length (ys ++ xs) = length xs + length ys
366-
lengthDistributesOverAppend [<] ys = Refl
367-
lengthDistributesOverAppend (xs :< x) ys =
368-
cong S $ lengthDistributesOverAppend xs ys
369-
370364
weakenUsed : {outer : _} -> Used vars -> Used (vars ++ outer)
371365
weakenUsed {outer} (MkUsed xs) =
372-
MkUsed (rewrite lengthDistributesOverAppend outer vars in
373-
(replicate (length outer) False ++ xs))
374-
375-
-- TODO
376-
-- lengthDistributesOverAppendFish
377-
-- : (xs : List a)
378-
-- -> (ys : SnocList a)
379-
-- -> length (ys <>< xs) = length xs + length ys
366+
MkUsed (rewrite lengthHomomorphism vars outer in
367+
rewrite plusCommutative (length vars) (length outer) in
368+
replicate (length outer) False ++ xs)
380369

381370
weakenUsedFish : {outer : _} -> Used vars -> Used (vars <>< outer)
382371
weakenUsedFish {outer} (MkUsed xs) =
383372
do rewrite fishAsSnocAppend vars outer
384-
MkUsed $ do
385-
rewrite lengthDistributesOverAppend (cast outer) vars
386-
-- See lengthDistributesOverAppendFish
387-
(believe_me $ replicate (length outer) False ++ xs)
373+
MkUsed $ do rewrite lengthHomomorphism vars (cast outer)
374+
rewrite Extra.lengthDistributesOverFish [<] outer
375+
rewrite plusCommutative (length vars) (length outer)
376+
replicate (length outer) False ++ xs
388377

389378
contractUsed : (Used (vars :< x)) -> Used vars
390379
contractUsed (MkUsed xs) = MkUsed (tail xs)

src/Core/Case/CaseBuilder.idr

Lines changed: 4 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -222,9 +222,9 @@ covering
222222
where
223223
showAll : {vs, ts : _} -> NamedPats vs ts -> String
224224
showAll [] = ""
225-
showAll {ts = _ :< t } [x]
225+
showAll {ts = _ :< t} [x]
226226
= show t ++ " " ++ show (pat x) ++ " [" ++ show (argType x) ++ "]"
227-
showAll {ts = _ :< t } (x :: xs)
227+
showAll {ts = _ :< t} (x :: xs)
228228
= show t ++ " " ++ show (pat x) ++ " [" ++ show (argType x) ++ "]"
229229
++ ", " ++ showAll xs
230230

@@ -233,7 +233,7 @@ covering
233233
where
234234
prettyAll : {vs, ts : _} -> NamedPats vs ts -> List (Doc IdrisSyntax)
235235
prettyAll [] = []
236-
prettyAll {ts = _ :< t } (x :: xs)
236+
prettyAll {ts = _ :< t} (x :: xs)
237237
= parens (pretty0 t <++> equals <++> pretty (pat x))
238238
:: prettyAll xs
239239

@@ -597,10 +597,6 @@ updatePatNames _ [] = []
597597
updatePatNames ns (pi :: ps)
598598
= { pat $= update } pi :: updatePatNames ns ps
599599
where
600-
lookup : Name -> SnocList (Name, Name) -> Maybe Name
601-
lookup n [<] = Nothing
602-
lookup n (ns :< (x, n')) = if x == n then Just n' else lookup n ns
603-
604600
update : Pat -> Pat
605601
update (PAs fc n p)
606602
= case lookup n ns of
@@ -1226,7 +1222,7 @@ mkPat args orig (Ref fc (DataCon t a) n) = pure $ PCon fc n t a (rev args)
12261222
mkPat args orig (Ref fc (TyCon t a) n) = pure $ PTyCon fc n a (rev args)
12271223
mkPat args orig (Ref fc Func n)
12281224
= do prims <- getPrimitiveNames
1229-
mtm <- normalisePrims (const True) isPConst True prims n args orig [<]
1225+
mtm <- normalisePrims (const True) isPConst True prims n (cast args) orig [<]
12301226
case mtm of
12311227
Just tm => if tm /= orig -- check we made progress; if there's an
12321228
-- unresolved interface, we might be stuck

src/Core/Context/Context.idr

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -195,7 +195,7 @@ export
195195
covering
196196
Show Clause where
197197
show (MkClause {vars} env lhs rhs)
198-
= show (toList $ reverse vars) ++ ": " ++ show lhs ++ " = " ++ show rhs
198+
= show (asList vars) ++ ": " ++ show lhs ++ " = " ++ show rhs
199199

200200
public export
201201
data DefFlag

src/Core/Env.idr

Lines changed: 0 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -78,15 +78,6 @@ bindEnv loc (env :< b) tm
7878
Explicit
7979
(binderType b)) tm)
8080

81-
public export
82-
revNs : (vs, ns : SnocList a) -> reverse vs ++ reverse ns = reverse (ns ++ vs)
83-
revNs [<] ns = rewrite appendLinLeftNeutral (reverse ns) in Refl
84-
revNs (vs :< v) ns
85-
= rewrite Extra.revOnto [<v] vs in
86-
rewrite Extra.revOnto [<v] (ns ++ vs) in
87-
rewrite sym $ revNs vs ns in
88-
rewrite appendAssociative [<v] (reverse vs) (reverse ns) in Refl
89-
9081
-- Weaken by all the names at once at the end, to save multiple traversals
9182
-- in big environments
9283
-- Also reversing the names at the end saves significant time over concatenating

src/Core/Normalise.idr

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -362,7 +362,7 @@ normalisePrims : {auto c : Ref Ctxt Defs} -> {vs : _} ->
362362
List Name ->
363363
-- view of the potential redex
364364
(n : Name) -> -- function name
365-
(args : SnocList arg) -> -- arguments from inside out (arg1, ..., argk)
365+
(args : List arg) -> -- arguments from inside out (arg1, ..., argk)
366366
-- actual term to evaluate if needed
367367
(tm : Term vs) -> -- original term (n arg1 ... argk)
368368
Env Term vs -> -- evaluation environment
@@ -371,7 +371,7 @@ normalisePrims : {auto c : Ref Ctxt Defs} -> {vs : _} ->
371371
normalisePrims boundSafe viewConstant all prims n args tm env
372372
= do let True = isPrimName prims !(getFullName n) -- is a primitive
373373
| _ => pure Nothing
374-
let (_ :< mc) = reverse args -- with at least one argument
374+
let (mc :: _) = args -- with at least one argument
375375
| _ => pure Nothing
376376
let (Just c) = viewConstant mc -- that is a constant
377377
| _ => pure Nothing

src/Core/Normalise/Convert.idr

Lines changed: 13 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -14,6 +14,8 @@ import Data.List
1414
import Data.SnocList
1515

1616
import Libraries.Data.List.SizeOf
17+
import Libraries.Data.SnocList.HasLength
18+
import Libraries.Data.SnocList.SizeOf
1719

1820
%default covering
1921

@@ -369,10 +371,12 @@ mutual
369371

370372
convGen q inf defs env (NApp fc val args) (NApp _ val' args')
371373
= if !(chkConvHead q inf defs env val val')
372-
then do i <- getInfPos val
373-
allConv q inf defs env
374-
(cast {from = List (FC, Closure vars)} $ dropInf 0 i $ cast args) -- TODO: UGH!
375-
(cast {from = List (FC, Closure vars)} $ dropInf 0 i $ cast args')
374+
then case !(getInfPos val) of
375+
[] => allConv q inf defs env args args'
376+
i => allConv q inf defs env
377+
(dropInf i args $ mkSizeOf args)
378+
(dropInf i args' $ mkSizeOf args')
379+
376380
else chkConvCaseBlock fc q inf defs env val args1 val' args2
377381
where
378382
getInfPos : NHead vars -> Core (List Nat)
@@ -384,13 +388,12 @@ mutual
384388
else pure []
385389
getInfPos _ = pure []
386390

387-
dropInf : Nat -> List Nat -> List a -> List a
388-
dropInf _ [] xs = xs
389-
dropInf _ _ [] = []
390-
dropInf i ds (x :: xs)
391+
dropInf : List Nat -> (sx : SnocList a) -> SizeOf sx -> SnocList a
392+
dropInf _ [<] _ = [<]
393+
dropInf ds (sx :< x) (MkSizeOf (S i) (S p))
391394
= if i `elem` ds
392-
then dropInf (S i) ds xs
393-
else x :: dropInf (S i) ds xs
395+
then dropInf ds sx (MkSizeOf i p)
396+
else dropInf ds sx (MkSizeOf i p) :< x
394397

395398
-- Discard file context information irrelevant for conversion checking
396399
args1 : SnocList (Closure vars)

src/Core/SchemeEval/Compile.idr

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -97,7 +97,7 @@ reverseOnto acc [<] = acc
9797
reverseOnto acc (sx :< x) = reverseOnto (acc :< x) sx
9898

9999
reverse : SchVars vars -> SchVars (reverse vars)
100-
reverse sx = reverseOnto [<] sx
100+
reverse = reverseOnto [<]
101101

102102
getSchVar : {idx : _} -> (0 _ : IsVar n idx vars) -> SchVars vars -> String
103103
getSchVar First (xs :< Bound x) = x

0 commit comments

Comments
 (0)