Idris2/tests/idris2/interactive017/RLE.idr
Edwin Brady 3a6614b227 Look at intermediate results in program search
This has involved quite a bit of reorganisation and some improvements in
resugaring so that the results look nice. In summary:

* Expression search now gives back a RawImp rather than a checked term,
  which allows it to include case expressions
* Case with one pattern is resugared to a destructuring let
* Some name generation issues address in function generation

We look at intermediate results for local variables which are functions
that return a concrete type, or recursive calls that return a single
constructor type. In these cases, we:

* let bind the local variable/recursive call
* generate a new definition for the scope, as a 'case' function

When we recursively generate the definition, it's a bit more restricted
so as not to explode the search space. We only take the first result, we
only look one constructor deep, and we go right to left on variable
splitting so only deconstruct the name we've just added.
2020-08-04 12:51:57 +01:00

96 lines
1.7 KiB
Idris

module RLE
import Decidable.Equality
rep : Nat -> a -> List a
data RunLength : List a -> Type where
Empty : RunLength []
Run : (n : Nat) -> (x : a) -> (rest : RunLength more) ->
RunLength (rep n x ++ more)
data Singleton : a -> Type where
Val : (x : a) -> Singleton x
uncompress : RunLength xs -> Singleton xs
{-
rle : DecEq a => (xs : List a) -> RLE xs
rle [] = Empty
rle (x :: xs) with (rle xs)
rle (x :: []) | Empty = Run 1 x Empty
rle (x :: (rep n y ++ more)) | (Run n y comp) with (decEq x y)
rle (y :: (rep n y ++ more)) | (Run n y comp) | (Yes Refl)
= Run (S n) y comp
rle (x :: (rep n y ++ more)) | (Run n y comp) | (No contra)
= Run 1 x $ Run n y comp
data Singleton : a -> Type where
Val : (x : a) -> Singleton x
uncompress : RLE xs -> Singleton xs
-- uncompress Empty = Val []
-- uncompress (Run n x comp)
-- = let Val rec = uncompress comp in
-- Val (rep n x ++ rec)
uncompressHelp : (x : a) -> RLE more -> (n : Nat) -> Singleton more -> Singleton (rep n x ++ more)
uncompressHelp x comp n (Val _) = Val (rep n x ++ more)
uncompress' : RLE xs -> Singleton xs
uncompress' Empty = Val []
uncompress' (Run n x comp)
= let rec = uncompress' comp in
uncompressHelp x comp n rec
{-
rle [] = Empty
rle (x :: xs) with (rle xs)
rle (x :: []) | Empty = Run 1 x Empty
rle (x :: (rep n y ++ more)) | (Run n y comp) with (decEq x y)
rle (y :: (rep n y ++ more)) | (Run n y comp) | (Yes Refl)
= Run (S n) y comp
rle (x :: (rep n y ++ more)) | (Run n y comp) | (No contra)
= Run 1 x $ Run n y comp
uncompress Empty = Val []
uncompress (Run n x comp)
= let Val uncomp = uncompress comp in
Val (rep n x ++ uncomp)
-}
-}