From 137cc7335e64fcb70c254e7fd2a6fa353fb43e61 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 21 Nov 2023 14:57:38 +0100 Subject: Regenerate the files --- tests/lean/Paper.lean | 33 ++++++++++++++++++++++----------- 1 file changed, 22 insertions(+), 11 deletions(-) (limited to 'tests/lean/Paper.lean') diff --git a/tests/lean/Paper.lean b/tests/lean/Paper.lean index ae4dd243..37e0e70e 100644 --- a/tests/lean/Paper.lean +++ b/tests/lean/Paper.lean @@ -6,11 +6,13 @@ open Primitives namespace paper /- [paper::ref_incr]: merged forward/backward function - (there is a single backward function, and the forward function returns ()) -/ + (there is a single backward function, and the forward function returns ()) + Source: 'src/paper.rs', lines 4:0-4:28 -/ def ref_incr (x : I32) : Result I32 := x + 1#i32 -/- [paper::test_incr]: forward function -/ +/- [paper::test_incr]: forward function + Source: 'src/paper.rs', lines 8:0-8:18 -/ def test_incr : Result Unit := do let x ← ref_incr 0#i32 @@ -21,20 +23,23 @@ def test_incr : Result Unit := /- Unit test for [paper::test_incr] -/ #assert (test_incr == .ret ()) -/- [paper::choose]: forward function -/ +/- [paper::choose]: forward function + Source: 'src/paper.rs', lines 15:0-15:70 -/ def choose (T : Type) (b : Bool) (x : T) (y : T) : Result T := if b then Result.ret x else Result.ret y -/- [paper::choose]: backward function 0 -/ +/- [paper::choose]: backward function 0 + Source: 'src/paper.rs', lines 15:0-15:70 -/ def choose_back (T : Type) (b : Bool) (x : T) (y : T) (ret0 : T) : Result (T × T) := if b then Result.ret (ret0, y) else Result.ret (x, ret0) -/- [paper::test_choose]: forward function -/ +/- [paper::test_choose]: forward function + Source: 'src/paper.rs', lines 23:0-23:20 -/ def test_choose : Result Unit := do let z ← choose I32 true 0#i32 0#i32 @@ -53,12 +58,14 @@ def test_choose : Result Unit := /- Unit test for [paper::test_choose] -/ #assert (test_choose == .ret ()) -/- [paper::List] -/ +/- [paper::List] + Source: 'src/paper.rs', lines 35:0-35:16 -/ inductive List (T : Type) := | Cons : T → List T → List T | Nil : List T -/- [paper::list_nth_mut]: forward function -/ +/- [paper::list_nth_mut]: forward function + Source: 'src/paper.rs', lines 42:0-42:67 -/ divergent def list_nth_mut (T : Type) (l : List T) (i : U32) : Result T := match l with | List.Cons x tl => @@ -69,7 +76,8 @@ divergent def list_nth_mut (T : Type) (l : List T) (i : U32) : Result T := list_nth_mut T tl i0 | List.Nil => Result.fail Error.panic -/- [paper::list_nth_mut]: backward function 0 -/ +/- [paper::list_nth_mut]: backward function 0 + Source: 'src/paper.rs', lines 42:0-42:67 -/ divergent def list_nth_mut_back (T : Type) (l : List T) (i : U32) (ret0 : T) : Result (List T) := match l with @@ -83,7 +91,8 @@ divergent def list_nth_mut_back Result.ret (List.Cons x tl0) | List.Nil => Result.fail Error.panic -/- [paper::sum]: forward function -/ +/- [paper::sum]: forward function + Source: 'src/paper.rs', lines 57:0-57:32 -/ divergent def sum (l : List I32) : Result I32 := match l with | List.Cons x tl => do @@ -91,7 +100,8 @@ divergent def sum (l : List I32) : Result I32 := x + i | List.Nil => Result.ret 0#i32 -/- [paper::test_nth]: forward function -/ +/- [paper::test_nth]: forward function + Source: 'src/paper.rs', lines 68:0-68:17 -/ def test_nth : Result Unit := do let l := List.Nil @@ -108,7 +118,8 @@ def test_nth : Result Unit := /- Unit test for [paper::test_nth] -/ #assert (test_nth == .ret ()) -/- [paper::call_choose]: forward function -/ +/- [paper::call_choose]: forward function + Source: 'src/paper.rs', lines 76:0-76:44 -/ def call_choose (p : (U32 × U32)) : Result U32 := do let (px, py) := p -- cgit v1.2.3