-
Notifications
You must be signed in to change notification settings - Fork 432
Commit
This commit does not belong to any branch on this repository, and may belong to a fork outside of the repository.
fix: reference implementation ByteArray.copySlice (#2967)
Fixes reference implementation of `ByteArray.copySlice`, as reported #2966. Adds tests. --------- Co-authored-by: Joachim Breitner <mail@joachim-breitner.de>
- Loading branch information
Showing
3 changed files
with
82 additions
and
11 deletions.
There are no files selected for viewing
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
Original file line number | Diff line number | Diff line change |
---|---|---|
@@ -0,0 +1,81 @@ | ||
import Lean | ||
|
||
def testCopySlice (src : Array UInt8) (srcOff : Nat) (dest : Array UInt8) (destOff len : Nat) : Bool := | ||
(ByteArray.copySlice ⟨src⟩ srcOff ⟨dest⟩ destOff len |>.toList) == | ||
(ByteArray.copySlice ⟨src⟩ srcOff ⟨dest⟩ destOff len |>.toList) | ||
|
||
-- We verify that the `@[extern]` implementation of `ByteArray.copySlice` matches the formal definition, | ||
-- by equating two copies using `==`, unfolding the definition of one, and then calling `native_decide` | ||
macro "testCopySliceTactic" : tactic => | ||
`(tactic| | ||
(dsimp [testCopySlice] | ||
conv => | ||
congr | ||
congr | ||
dsimp [ByteArray.copySlice] | ||
native_decide)) | ||
|
||
-- These used to fail, as reported in #2966 | ||
|
||
example : testCopySlice #[1,2,3] 1 #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] 0 6 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 1 #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] 0 20 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 1 #[4, 5, 6] 0 6 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 1 #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] 1 6 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 1 #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] 1 20 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 0 #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] 0 6 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 0 #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] 0 20 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 0 #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] 1 6 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 0 #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] 1 20 := by | ||
testCopySliceTactic | ||
|
||
|
||
-- These worked prior to #2966; extra text cases can't hurt! | ||
|
||
example : testCopySlice #[1,2,3] 1 #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] 0 2 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 1 #[4, 5, 6] 0 2 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 1 #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] 1 2 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 1 #[4, 5, 6] 1 2 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 1 #[4, 5, 6] 1 6 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 1 #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] 0 2 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 0 #[4, 5, 6] 0 2 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 0 #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] 1 2 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 0 #[4, 5, 6] 1 2 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 0 #[4, 5, 6] 1 6 := by | ||
testCopySliceTactic | ||
|
||
example : testCopySlice #[1,2,3] 0 #[4, 5, 6] 0 6 := by | ||
testCopySliceTactic |
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
Original file line number | Diff line number | Diff line change |
---|---|---|
@@ -1,11 +1 @@ | ||
test_extern.lean:7:0-7:17: error: test_extern: HAdd.hAdd does not have an @[extern] attribute | ||
test_extern.lean:9:0-9:86: error: native implementation did not agree with reference implementation! | ||
Compare the outputs of: | ||
#eval ByteArray.copySlice { data := #[1, 2, 3] } 1 { data := #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] } 0 6 | ||
and | ||
#eval { | ||
data := | ||
Array.extract { data := #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] }.data 0 0 ++ | ||
Array.extract { data := #[1, 2, 3] }.data 1 (1 + 6) ++ | ||
Array.extract { data := #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] }.data (0 + 6) | ||
(Array.size { data := #[4, 5, 6, 7, 8, 9, 10, 11, 12, 13] }.data) } |