Skip to content

Commit

Permalink
[ elab ] Change quantity of the search function's argument to 0
Browse files Browse the repository at this point in the history
  • Loading branch information
buzden authored and gallais committed Jun 28, 2024
1 parent 715a304 commit 57f455d
Show file tree
Hide file tree
Showing 5 changed files with 33 additions and 1 deletion.
3 changes: 3 additions & 0 deletions CHANGELOG_NEXT.md
Original file line number Diff line number Diff line change
Expand Up @@ -193,6 +193,9 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO

* `Functor` is implemented for `PiInfo` from `Language.Reflection.TT`.

* Quantity of the argument for the type being searched in the `search` function
from `Language.Reflection` was changed to be zero.

#### Contrib

* `Data.List.Lazy` was moved from `contrib` to `base`.
Expand Down
2 changes: 1 addition & 1 deletion libs/base/Language/Reflection.idr
Original file line number Diff line number Diff line change
Expand Up @@ -287,5 +287,5 @@ catch elab = try (Just <$> elab) (pure Nothing)
||| Run proof search to attempt to find a value of the input type.
||| Useful to check whether a give interface constraint is satisfied.
export
search : Elaboration m => (ty : Type) -> m (Maybe ty)
search : Elaboration m => (0 ty : Type) -> m (Maybe ty)
search ty = catch $ check {expected = ty} `(%search)
25 changes: 25 additions & 0 deletions tests/idris2/reflection/reflection030/SearchZero.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,25 @@
import Language.Reflection

%default total

scr : ty -> ty -> Elab ty
scr l r = do
Just x <- search $ Semigroup ty
| Nothing => pure l
pure $ l <+> r

%language ElabReflection

data N = A | B | C | D

X : N
X = %runElab scr B C

xCorr : X = B
xCorr = Refl

Y : List Nat
Y = %runElab scr [1, 2, 3] [4, 5]

yCorr : Y = [1, 2, 3, 4, 5]
yCorr = Refl
1 change: 1 addition & 0 deletions tests/idris2/reflection/reflection030/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
1/1: Building SearchZero (SearchZero.idr)
3 changes: 3 additions & 0 deletions tests/idris2/reflection/reflection030/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
. ../../../testutils.sh

check SearchZero.idr

0 comments on commit 57f455d

Please sign in to comment.