Skip to content

Commit

Permalink
[ test ] Add failing tests for quantities of proofs
Browse files Browse the repository at this point in the history
  • Loading branch information
spcfox committed Nov 18, 2024
1 parent ae9a9a0 commit ba57b5b
Show file tree
Hide file tree
Showing 4 changed files with 51 additions and 0 deletions.
28 changes: 28 additions & 0 deletions tests/idris2/with/with012/WithProof0Fail.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,28 @@
module WithProof0Fail

%default total

%hide List.filter

data So : Bool -> Type where
Oh : So True

eqToSo : b = True -> So b
eqToSo Refl = Oh

data All : (0 p : a -> Type) -> List a -> Type where
Nil : All p Nil
(::) : {0 xs : List a} -> p x -> All p xs -> All p (x :: xs)

filter : (p : a -> Bool) -> (xs : List a) -> List a
filter p [] = []
filter p (x :: xs) with (p x)
filter p (x :: xs) | False = filter p xs
filter p (x :: xs) | True = x :: filter p xs

failing "While processing right hand side of with block in allFilter. prf is not accessible in this context"
allFilter : (p : a -> Bool) -> (xs : List a) -> All (So . p) (filter p xs)
allFilter p [] = []
allFilter p (x :: xs) with (p x) proof 0 prf
allFilter p (x :: xs) | False = allFilter p xs
allFilter p (x :: xs) | True = eqToSo prf :: allFilter p xs
19 changes: 19 additions & 0 deletions tests/idris2/with/with012/WithProof1Fail.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,19 @@
module WithProof1Fail

%default total

%hide List.filter

filter : (p : a -> Bool) -> (xs : List a) -> List a
filter p [] = []
filter p (x :: xs) with (p x)
filter p (x :: xs) | False = filter p xs
filter p (x :: xs) | True = x :: filter p xs

failing "While processing right hand side of with block in filterSquared. There are 0 uses of linear name eq"
filterSquared : (p : a -> Bool) -> (xs : List a) ->
filter p (filter p xs) === filter p xs
filterSquared p [] = Refl
filterSquared p (x :: xs) with (p x) proof 1 eq
filterSquared p (x :: xs) | False = filterSquared p xs
filterSquared p (x :: xs) | True = filterSquared p xs
2 changes: 2 additions & 0 deletions tests/idris2/with/with012/expected
Original file line number Diff line number Diff line change
@@ -1,2 +1,4 @@
1/1: Building WithProof0 (WithProof0.idr)
1/1: Building WithProof1 (WithProof1.idr)
1/1: Building WithProof0Fail (WithProof0Fail.idr)
1/1: Building WithProof1Fail (WithProof1Fail.idr)
2 changes: 2 additions & 0 deletions tests/idris2/with/with012/run
Original file line number Diff line number Diff line change
Expand Up @@ -2,3 +2,5 @@

check WithProof0.idr
check WithProof1.idr
check WithProof0Fail.idr
check WithProof1Fail.idr

0 comments on commit ba57b5b

Please sign in to comment.