diff --git a/CHANGELOG.md b/CHANGELOG.md index ef6af5660..0684c4dc3 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -102,6 +102,9 @@ Other changes: * Added an environment variable `IDRIS2_PACKAGE_PATH` for extending where to look for packages. * Support for auto-completion in bash-like shells was added. +* Fixed case-splitting to respect any indentation there may be in the term being + case-split and the surrounding symbols, instead of filtering out the + whitespace and putting it back as indentation. Changes since Idris 2 v0.2.1 ---------------------------- diff --git a/libs/base/Data/List.idr b/libs/base/Data/List.idr index 15687f80b..2c6a19019 100644 --- a/libs/base/Data/List.idr +++ b/libs/base/Data/List.idr @@ -196,6 +196,9 @@ public export union : Eq a => List a -> List a -> List a union = unionBy (==) +||| Like @span@ but using a predicate that might convert a to b, i.e. given a +||| predicate from a to Maybe b and a list of as, returns a tuple consisting of +||| the longest prefix of the list where a -> Just b, and the rest of the list. public export spanBy : (a -> Maybe b) -> List a -> (List b, List a) spanBy p [] = ([], []) @@ -203,6 +206,9 @@ spanBy p (x :: xs) = case p x of Nothing => ([], x :: xs) Just y => let (ys, zs) = spanBy p xs in (y :: ys, zs) +||| Given a predicate and a list, returns a tuple consisting of the longest +||| prefix of the list whose elements satisfy the predicate, and the rest of the +||| list. public export span : (a -> Bool) -> List a -> (List a, List a) span p [] = ([], []) diff --git a/src/Idris/IDEMode/CaseSplit.idr b/src/Idris/IDEMode/CaseSplit.idr index e095ed064..35923802c 100644 --- a/src/Idris/IDEMode/CaseSplit.idr +++ b/src/Idris/IDEMode/CaseSplit.idr @@ -60,37 +60,62 @@ toStrUpdate _ = pure [] -- can't replace non user names data UPD : Type where +||| Given a list of definitions, a list of mappings from `RawName` to `String`, +||| and a list of tokens to update, work out the updates to do, apply them, and +||| return the result. doUpdates : {auto u : Ref UPD (List String)} -> Defs -> Updates -> List SourcePart -> Core (List SourcePart) -doUpdates defs ups [] = pure [] +doUpdates defs ups [] = pure [] -- no more tokens to update, so we are done +-- if we have an LBrace (i.e. `{`), handle its contents doUpdates defs ups (LBrace :: xs) - = let (ws, nws) = spanSpace xs in map (LBrace :: ws ++) $ - case nws of - Name n :: RBrace :: rest => - pure (Name n :: - Whitespace " " :: Equal :: Whitespace " " :: - !(doUpdates defs ups (Name n :: RBrace :: rest))) - Name n :: Equal :: rest => - pure (Name n :: - Whitespace " " :: Equal :: Whitespace " " :: - !(doUpdates defs ups rest)) - _ => doUpdates defs ups xs + -- the cases we care about are easy to detect w/o whitespace, so separate it + = let (ws, nws) = span isWhitespace xs in + case nws of + -- handle potential whitespace in the other parts + Name n :: rest => + let (ws', nws') = span isWhitespace rest in + case nws' of + -- brace is immediately closed, so generate a new + -- pattern-match on the values the name can have, e.g. + -- { x} where x : Nat would become { x = Z} + -- (and later { x = (S k)}) + RBrace :: rest' => + pure (LBrace :: ws ++ + Name n :: Whitespace " " :: Equal :: Whitespace " " :: + !(doUpdates defs ups (Name n :: ws' ++ RBrace :: rest')) + ) + -- preserve whitespace before (and after) the Equal + Equal :: rest' => + let (ws'', nws'') = span isWhitespace rest' in + pure (LBrace :: ws ++ + Name n :: ws' ++ Equal :: ws'' ++ + !(doUpdates defs ups nws'') + ) + -- handle everything else as usual, preserving whitespace + _ => pure (LBrace :: ws ++ Name n :: ws' ++ + !(doUpdates defs ups rest) + ) + -- not a special case: proceed as normal + _ => pure (LBrace :: [] ++ !(doUpdates defs ups xs)) where - spanSpace : List SourcePart -> (List SourcePart, List SourcePart) - spanSpace [] = ([], []) - spanSpace (RBrace :: xs) = ([], RBrace :: xs) - spanSpace (w@(Whitespace _) :: xs) = mapFst (w ::) (spanSpace xs) - spanSpace (x :: xs) = map (x ::) (spanSpace xs) + isWhitespace : SourcePart -> Bool + isWhitespace (Whitespace _) = True + isWhitespace _ = False +-- if we have a name, look up if it's a name we're updating. If it isn't, keep +-- the old name, otherwise update the name, i.e. replace with the new name doUpdates defs ups (Name n :: xs) = case lookup n ups of Nothing => pure (Name n :: !(doUpdates defs ups xs)) Just up => pure (Other up :: !(doUpdates defs ups xs)) +-- if we have a hole, get the used names, generate+register a new unique name, +-- and change the hole's name to the new one doUpdates defs ups (HoleName n :: xs) = do used <- get UPD n' <- uniqueName defs used n put UPD (n' :: used) pure $ HoleName n' :: !(doUpdates defs ups xs) +-- if it's not a thing we update, leave it and continue working on the rest doUpdates defs ups (x :: xs) = pure $ x :: !(doUpdates defs ups xs) diff --git a/tests/idris2/interactive013/Spacing.idr b/tests/idris2/interactive013/Spacing.idr index 52f369bd5..378fab4f8 100644 --- a/tests/idris2/interactive013/Spacing.idr +++ b/tests/idris2/interactive013/Spacing.idr @@ -1,4 +1,61 @@ module Spacing -id : {n : Nat} -> Nat -id { n} = ?a +no : {n : Nat} -> Nat +no {n} = ?no_rhs + +spaced : {n : Nat} -> Nat +spaced { n } = ?spaced_rhs + +s1 : {n : Nat} -> Nat +s1 { n} = ?s1_rhs + +s2 : {n : Nat} -> Nat +s2 { n} = ?s2_rhs + +s3 : {n : Nat} -> Nat +s3 { n} = ?s3_rhs + + +noSEq : {n : Nat} -> Nat +noSEq {n = a} = ?noSEq_rhs + +spacedEq : {n : Nat} -> Nat +spacedEq { n = a } = ?spacedEq_rhs + +s1Eq : {n : Nat} -> Nat +s1Eq { n = a} = ?s1Eq_rhs + +s2Eq : {n : Nat} -> Nat +s2Eq { n = a} = ?s2Eq_rhs + +s3Eq : {n : Nat} -> Nat +s3Eq { n = a} = ?s3Eq_rhs + + +weirdNo : {n : Nat} -> Nat +weirdNo {n } = ?weirdNo_rhs + +weird0a : {n : Nat} -> Nat +weird0a {n= a} = ?weird0a_rhs + +weird0b : {n : Nat} -> Nat +weird0b {n =b} = ?weird0b_rhs + +weird1a : {n : Nat} -> Nat +weird1a { n= a} = ?weird1a_rhs + +weird1b : {n : Nat} -> Nat +weird1b { n =b} = ?weird1b_rhs + +weird2a : {n : Nat} -> Nat +weird2a { n= a} = ?weird2a_rhs + +weird2b : {n : Nat} -> Nat +weird2b { n =b} = ?weird2b_rhs + +weirdSpacedA : {n : Nat} -> Nat +weirdSpacedA { n= a } = ?weirdSpacedA_rhs + +weirdSpacedB : {n : Nat} -> Nat +weirdSpacedB { n =b } = ?weirdSpacedB_rhs + diff --git a/tests/idris2/interactive013/expected b/tests/idris2/interactive013/expected index 3145378c8..5f0141a14 100644 --- a/tests/idris2/interactive013/expected +++ b/tests/idris2/interactive013/expected @@ -1,5 +1,40 @@ 1/1: Building Spacing (Spacing.idr) -Spacing> id { n = 0} = ?a_1 -id { n = (S k)} = ?a_2 -Spacing> -Bye for now! +Spacing> no {n = 0} = ?no_rhs_1 +no {n = (S k)} = ?no_rhs_2 +Spacing> spaced { n = 0 } = ?spaced_rhs_1 +spaced { n = (S k) } = ?spaced_rhs_2 +Spacing> s1 { n = 0} = ?s1_rhs_1 +s1 { n = (S k)} = ?s1_rhs_2 +Spacing> s2 { n = 0} = ?s2_rhs_1 +s2 { n = (S k)} = ?s2_rhs_2 +Spacing> s3 { n = 0} = ?s3_rhs_1 +s3 { n = (S k)} = ?s3_rhs_2 +Spacing> noSEq {n = 0} = ?noSEq_rhs_1 +noSEq {n = (S k)} = ?noSEq_rhs_2 +Spacing> spacedEq { n = 0 } = ?spacedEq_rhs_1 +spacedEq { n = (S k) } = ?spacedEq_rhs_2 +Spacing> s1Eq { n = 0} = ?s1Eq_rhs_1 +s1Eq { n = (S k)} = ?s1Eq_rhs_2 +Spacing> s2Eq { n = 0} = ?s2Eq_rhs_1 +s2Eq { n = (S k)} = ?s2Eq_rhs_2 +Spacing> s3Eq { n = 0} = ?s3Eq_rhs_1 +s3Eq { n = (S k)} = ?s3Eq_rhs_2 +Spacing> weirdNo {n = 0 } = ?weirdNo_rhs_1 +weirdNo {n = (S k) } = ?weirdNo_rhs_2 +Spacing> weird0a {n= 0} = ?weird0a_rhs_1 +weird0a {n= (S k)} = ?weird0a_rhs_2 +Spacing> weird0b {n =0} = ?weird0b_rhs_1 +weird0b {n =(S k)} = ?weird0b_rhs_2 +Spacing> weird1a { n= 0} = ?weird1a_rhs_1 +weird1a { n= (S k)} = ?weird1a_rhs_2 +Spacing> weird1b { n =0} = ?weird1b_rhs_1 +weird1b { n =(S k)} = ?weird1b_rhs_2 +Spacing> weird2a { n= 0} = ?weird2a_rhs_1 +weird2a { n= (S k)} = ?weird2a_rhs_2 +Spacing> weird2b { n =0} = ?weird2b_rhs_1 +weird2b { n =(S k)} = ?weird2b_rhs_2 +Spacing> weirdSpacedA { n= 0 } = ?weirdSpacedA_rhs_1 +weirdSpacedA { n= (S k) } = ?weirdSpacedA_rhs_2 +Spacing> weirdSpacedB { n =0 } = ?weirdSpacedB_rhs_1 +weirdSpacedB { n =(S k) } = ?weirdSpacedB_rhs_2 +Spacing> Bye for now! diff --git a/tests/idris2/interactive013/input b/tests/idris2/interactive013/input index 6cc8ff48c..cd5ef39dc 100644 --- a/tests/idris2/interactive013/input +++ b/tests/idris2/interactive013/input @@ -1,2 +1,20 @@ -:cs 4 6 n -:q \ No newline at end of file +:cs 4 5 n +:cs 7 10 n +:cs 10 6 n +:cs 13 7 n +:cs 16 8 n +:cs 20 12 a +:cs 23 16 a +:cs 26 12 a +:cs 29 13 a +:cs 32 14 a +:cs 36 10 n +:cs 39 13 a +:cs 42 13 b +:cs 45 14 a +:cs 48 14 b +:cs 51 16 a +:cs 54 16 b +:cs 57 20 a +:cs 60 20 b +:q