summaryrefslogtreecommitdiff
path: root/tests/coq/misc
diff options
context:
space:
mode:
authorSon HO2023-11-22 15:06:43 +0100
committerGitHub2023-11-22 15:06:43 +0100
commitbacf3f5f6f5f6a9aa650d5ae8d12a132fd747039 (patch)
tree9953d7af1fe406cdc750030a43a5e4d6245cd763 /tests/coq/misc
parent587f1ebc0178acb19029d3fc9a729c197082aba7 (diff)
parent01cfd899119174ef7c5941c99dd251711f4ee701 (diff)
Merge pull request #45 from AeneasVerif/son_merge_types
Big cleanup
Diffstat (limited to '')
-rw-r--r--tests/coq/misc/Constants.v84
-rw-r--r--tests/coq/misc/External_Funs.v27
-rw-r--r--tests/coq/misc/External_Opaque.v15
-rw-r--r--tests/coq/misc/External_Types.v3
-rw-r--r--tests/coq/misc/Loops.v201
-rw-r--r--tests/coq/misc/NoNestedBorrows.v212
-rw-r--r--tests/coq/misc/Paper.v33
-rw-r--r--tests/coq/misc/PoloniusList.v9
-rw-r--r--tests/coq/misc/Primitives.v88
9 files changed, 433 insertions, 239 deletions
diff --git a/tests/coq/misc/Constants.v b/tests/coq/misc/Constants.v
index 03653f69..20edb2b1 100644
--- a/tests/coq/misc/Constants.v
+++ b/tests/coq/misc/Constants.v
@@ -8,124 +8,152 @@ Import ListNotations.
Local Open Scope Primitives_scope.
Module Constants.
-(** [constants::X0] *)
+(** [constants::X0]
+ Source: 'src/constants.rs', lines 5:0-5:17 *)
Definition x0_body : result u32 := Return 0%u32.
Definition x0_c : u32 := x0_body%global.
-(** [constants::X1] *)
+(** [constants::X1]
+ Source: 'src/constants.rs', lines 7:0-7:17 *)
Definition x1_body : result u32 := Return core_u32_max.
Definition x1_c : u32 := x1_body%global.
-(** [constants::X2] *)
+(** [constants::X2]
+ Source: 'src/constants.rs', lines 10:0-10:17 *)
Definition x2_body : result u32 := Return 3%u32.
Definition x2_c : u32 := x2_body%global.
-(** [constants::incr]: forward function *)
+(** [constants::incr]: forward function
+ Source: 'src/constants.rs', lines 17:0-17:32 *)
Definition incr (n : u32) : result u32 :=
u32_add n 1%u32.
-(** [constants::X3] *)
+(** [constants::X3]
+ Source: 'src/constants.rs', lines 15:0-15:17 *)
Definition x3_body : result u32 := incr 32%u32.
Definition x3_c : u32 := x3_body%global.
-(** [constants::mk_pair0]: forward function *)
+(** [constants::mk_pair0]: forward function
+ Source: 'src/constants.rs', lines 23:0-23:51 *)
Definition mk_pair0 (x : u32) (y : u32) : result (u32 * u32) :=
Return (x, y).
-(** [constants::Pair] *)
+(** [constants::Pair]
+ Source: 'src/constants.rs', lines 36:0-36:23 *)
Record Pair_t (T1 T2 : Type) := mkPair_t { pair_x : T1; pair_y : T2; }.
Arguments mkPair_t { _ _ }.
Arguments pair_x { _ _ }.
Arguments pair_y { _ _ }.
-(** [constants::mk_pair1]: forward function *)
+(** [constants::mk_pair1]: forward function
+ Source: 'src/constants.rs', lines 27:0-27:55 *)
Definition mk_pair1 (x : u32) (y : u32) : result (Pair_t u32 u32) :=
Return {| pair_x := x; pair_y := y |}
.
-(** [constants::P0] *)
+(** [constants::P0]
+ Source: 'src/constants.rs', lines 31:0-31:24 *)
Definition p0_body : result (u32 * u32) := mk_pair0 0%u32 1%u32.
Definition p0_c : (u32 * u32) := p0_body%global.
-(** [constants::P1] *)
+(** [constants::P1]
+ Source: 'src/constants.rs', lines 32:0-32:28 *)
Definition p1_body : result (Pair_t u32 u32) := mk_pair1 0%u32 1%u32.
Definition p1_c : Pair_t u32 u32 := p1_body%global.
-(** [constants::P2] *)
+(** [constants::P2]
+ Source: 'src/constants.rs', lines 33:0-33:24 *)
Definition p2_body : result (u32 * u32) := Return (0%u32, 1%u32).
Definition p2_c : (u32 * u32) := p2_body%global.
-(** [constants::P3] *)
+(** [constants::P3]
+ Source: 'src/constants.rs', lines 34:0-34:28 *)
Definition p3_body : result (Pair_t u32 u32) :=
Return {| pair_x := 0%u32; pair_y := 1%u32 |}
.
Definition p3_c : Pair_t u32 u32 := p3_body%global.
-(** [constants::Wrap] *)
+(** [constants::Wrap]
+ Source: 'src/constants.rs', lines 49:0-49:18 *)
Record Wrap_t (T : Type) := mkWrap_t { wrap_value : T; }.
Arguments mkWrap_t { _ }.
Arguments wrap_value { _ }.
-(** [constants::Wrap::{0}::new]: forward function *)
+(** [constants::{constants::Wrap<T>}::new]: forward function
+ Source: 'src/constants.rs', lines 54:4-54:41 *)
Definition wrap_new (T : Type) (value : T) : result (Wrap_t T) :=
Return {| wrap_value := value |}
.
-(** [constants::Y] *)
+(** [constants::Y]
+ Source: 'src/constants.rs', lines 41:0-41:22 *)
Definition y_body : result (Wrap_t i32) := wrap_new i32 2%i32.
Definition y_c : Wrap_t i32 := y_body%global.
-(** [constants::unwrap_y]: forward function *)
+(** [constants::unwrap_y]: forward function
+ Source: 'src/constants.rs', lines 43:0-43:30 *)
Definition unwrap_y : result i32 :=
Return y_c.(wrap_value).
-(** [constants::YVAL] *)
+(** [constants::YVAL]
+ Source: 'src/constants.rs', lines 47:0-47:19 *)
Definition yval_body : result i32 := unwrap_y.
Definition yval_c : i32 := yval_body%global.
-(** [constants::get_z1::Z1] *)
+(** [constants::get_z1::Z1]
+ Source: 'src/constants.rs', lines 62:4-62:17 *)
Definition get_z1_z1_body : result i32 := Return 3%i32.
Definition get_z1_z1_c : i32 := get_z1_z1_body%global.
-(** [constants::get_z1]: forward function *)
+(** [constants::get_z1]: forward function
+ Source: 'src/constants.rs', lines 61:0-61:28 *)
Definition get_z1 : result i32 :=
Return get_z1_z1_c.
-(** [constants::add]: forward function *)
+(** [constants::add]: forward function
+ Source: 'src/constants.rs', lines 66:0-66:39 *)
Definition add (a : i32) (b : i32) : result i32 :=
i32_add a b.
-(** [constants::Q1] *)
+(** [constants::Q1]
+ Source: 'src/constants.rs', lines 74:0-74:17 *)
Definition q1_body : result i32 := Return 5%i32.
Definition q1_c : i32 := q1_body%global.
-(** [constants::Q2] *)
+(** [constants::Q2]
+ Source: 'src/constants.rs', lines 75:0-75:17 *)
Definition q2_body : result i32 := Return q1_c.
Definition q2_c : i32 := q2_body%global.
-(** [constants::Q3] *)
+(** [constants::Q3]
+ Source: 'src/constants.rs', lines 76:0-76:17 *)
Definition q3_body : result i32 := add q2_c 3%i32.
Definition q3_c : i32 := q3_body%global.
-(** [constants::get_z2]: forward function *)
+(** [constants::get_z2]: forward function
+ Source: 'src/constants.rs', lines 70:0-70:28 *)
Definition get_z2 : result i32 :=
i <- get_z1; i0 <- add i q3_c; add q1_c i0.
-(** [constants::S1] *)
+(** [constants::S1]
+ Source: 'src/constants.rs', lines 80:0-80:18 *)
Definition s1_body : result u32 := Return 6%u32.
Definition s1_c : u32 := s1_body%global.
-(** [constants::S2] *)
+(** [constants::S2]
+ Source: 'src/constants.rs', lines 81:0-81:18 *)
Definition s2_body : result u32 := incr s1_c.
Definition s2_c : u32 := s2_body%global.
-(** [constants::S3] *)
+(** [constants::S3]
+ Source: 'src/constants.rs', lines 82:0-82:29 *)
Definition s3_body : result (Pair_t u32 u32) := Return p3_c.
Definition s3_c : Pair_t u32 u32 := s3_body%global.
-(** [constants::S4] *)
+(** [constants::S4]
+ Source: 'src/constants.rs', lines 83:0-83:29 *)
Definition s4_body : result (Pair_t u32 u32) := mk_pair1 7%u32 8%u32.
Definition s4_c : Pair_t u32 u32 := s4_body%global.
diff --git a/tests/coq/misc/External_Funs.v b/tests/coq/misc/External_Funs.v
index 018ce13c..0a14c7d1 100644
--- a/tests/coq/misc/External_Funs.v
+++ b/tests/coq/misc/External_Funs.v
@@ -12,7 +12,8 @@ Require Export External_Opaque.
Import External_Opaque.
Module External_Funs.
-(** [external::swap]: forward function *)
+(** [external::swap]: forward function
+ Source: 'src/external.rs', lines 6:0-6:46 *)
Definition swap
(T : Type) (x : T) (y : T) (st : state) : result (state * unit) :=
p <- core_mem_swap T x y st;
@@ -24,7 +25,8 @@ Definition swap
Return (st2, tt)
.
-(** [external::swap]: backward function 0 *)
+(** [external::swap]: backward function 0
+ Source: 'src/external.rs', lines 6:0-6:46 *)
Definition swap_back
(T : Type) (x : T) (y : T) (st : state) (st0 : state) :
result (state * (T * T))
@@ -38,7 +40,8 @@ Definition swap_back
Return (st0, (x0, y0))
.
-(** [external::test_new_non_zero_u32]: forward function *)
+(** [external::test_new_non_zero_u32]: forward function
+ Source: 'src/external.rs', lines 11:0-11:60 *)
Definition test_new_non_zero_u32
(x : u32) (st : state) : result (state * core_num_nonzero_NonZeroU32_t) :=
p <- core_num_nonzero_NonZeroU32_new x st;
@@ -46,7 +49,8 @@ Definition test_new_non_zero_u32
core_option_Option_unwrap core_num_nonzero_NonZeroU32_t o st0
.
-(** [external::test_vec]: forward function *)
+(** [external::test_vec]: forward function
+ Source: 'src/external.rs', lines 17:0-17:17 *)
Definition test_vec : result unit :=
let v := alloc_vec_Vec_new u32 in
_ <- alloc_vec_Vec_push u32 v 0%u32;
@@ -56,7 +60,8 @@ Definition test_vec : result unit :=
(** Unit test for [external::test_vec] *)
Check (test_vec )%return.
-(** [external::custom_swap]: forward function *)
+(** [external::custom_swap]: forward function
+ Source: 'src/external.rs', lines 24:0-24:66 *)
Definition custom_swap
(T : Type) (x : T) (y : T) (st : state) : result (state * T) :=
p <- core_mem_swap T x y st;
@@ -68,7 +73,8 @@ Definition custom_swap
Return (st2, x0)
.
-(** [external::custom_swap]: backward function 0 *)
+(** [external::custom_swap]: backward function 0
+ Source: 'src/external.rs', lines 24:0-24:66 *)
Definition custom_swap_back
(T : Type) (x : T) (y : T) (st : state) (ret : T) (st0 : state) :
result (state * (T * T))
@@ -82,13 +88,15 @@ Definition custom_swap_back
Return (st0, (ret, y0))
.
-(** [external::test_custom_swap]: forward function *)
+(** [external::test_custom_swap]: forward function
+ Source: 'src/external.rs', lines 29:0-29:59 *)
Definition test_custom_swap
(x : u32) (y : u32) (st : state) : result (state * unit) :=
p <- custom_swap u32 x y st; let (st0, _) := p in Return (st0, tt)
.
-(** [external::test_custom_swap]: backward function 0 *)
+(** [external::test_custom_swap]: backward function 0
+ Source: 'src/external.rs', lines 29:0-29:59 *)
Definition test_custom_swap_back
(x : u32) (y : u32) (st : state) (st0 : state) :
result (state * (u32 * u32))
@@ -96,7 +104,8 @@ Definition test_custom_swap_back
custom_swap_back u32 x y st 1%u32 st0
.
-(** [external::test_swap_non_zero]: forward function *)
+(** [external::test_swap_non_zero]: forward function
+ Source: 'src/external.rs', lines 35:0-35:44 *)
Definition test_swap_non_zero (x : u32) (st : state) : result (state * u32) :=
p <- swap u32 x 0%u32 st;
let (st0, _) := p in
diff --git a/tests/coq/misc/External_Opaque.v b/tests/coq/misc/External_Opaque.v
index 80be37e7..b482431f 100644
--- a/tests/coq/misc/External_Opaque.v
+++ b/tests/coq/misc/External_Opaque.v
@@ -10,27 +10,32 @@ Require Export External_Types.
Import External_Types.
Module External_Opaque.
-(** [core::mem::swap]: forward function *)
+(** [core::mem::swap]: forward function
+ Source: '/rustc/d59363ad0b6391b7fc5bbb02c9ccf9300eef3753/library/core/src/mem/mod.rs', lines 726:0-726:42 *)
Axiom core_mem_swap :
forall(T : Type), T -> T -> state -> result (state * unit)
.
-(** [core::mem::swap]: backward function 0 *)
+(** [core::mem::swap]: backward function 0
+ Source: '/rustc/d59363ad0b6391b7fc5bbb02c9ccf9300eef3753/library/core/src/mem/mod.rs', lines 726:0-726:42 *)
Axiom core_mem_swap_back0 :
forall(T : Type), T -> T -> state -> state -> result (state * T)
.
-(** [core::mem::swap]: backward function 1 *)
+(** [core::mem::swap]: backward function 1
+ Source: '/rustc/d59363ad0b6391b7fc5bbb02c9ccf9300eef3753/library/core/src/mem/mod.rs', lines 726:0-726:42 *)
Axiom core_mem_swap_back1 :
forall(T : Type), T -> T -> state -> state -> result (state * T)
.
-(** [core::num::nonzero::NonZeroU32::{14}::new]: forward function *)
+(** [core::num::nonzero::{core::num::nonzero::NonZeroU32#14}::new]: forward function
+ Source: '/rustc/d59363ad0b6391b7fc5bbb02c9ccf9300eef3753/library/core/src/num/nonzero.rs', lines 79:16-79:57 *)
Axiom core_num_nonzero_NonZeroU32_new
: u32 -> state -> result (state * (option core_num_nonzero_NonZeroU32_t))
.
-(** [core::option::Option::{0}::unwrap]: forward function *)
+(** [core::option::{core::option::Option<T>}::unwrap]: forward function
+ Source: '/rustc/d59363ad0b6391b7fc5bbb02c9ccf9300eef3753/library/core/src/option.rs', lines 932:4-932:34 *)
Axiom core_option_Option_unwrap :
forall(T : Type), option T -> state -> result (state * T)
.
diff --git a/tests/coq/misc/External_Types.v b/tests/coq/misc/External_Types.v
index 9e49ca41..c638670c 100644
--- a/tests/coq/misc/External_Types.v
+++ b/tests/coq/misc/External_Types.v
@@ -8,7 +8,8 @@ Import ListNotations.
Local Open Scope Primitives_scope.
Module External_Types.
-(** [core::num::nonzero::NonZeroU32] *)
+(** [core::num::nonzero::NonZeroU32]
+ Source: '/rustc/d59363ad0b6391b7fc5bbb02c9ccf9300eef3753/library/core/src/num/nonzero.rs', lines 50:12-50:33 *)
Axiom core_num_nonzero_NonZeroU32_t : Type.
(** The state type used in the state-error monad *)
diff --git a/tests/coq/misc/Loops.v b/tests/coq/misc/Loops.v
index 1c0eab17..4929ddd0 100644
--- a/tests/coq/misc/Loops.v
+++ b/tests/coq/misc/Loops.v
@@ -8,7 +8,8 @@ Import ListNotations.
Local Open Scope Primitives_scope.
Module Loops.
-(** [loops::sum]: loop 0: forward function *)
+(** [loops::sum]: loop 0: forward function
+ Source: 'src/loops.rs', lines 4:0-14:1 *)
Fixpoint sum_loop (n : nat) (max : u32) (i : u32) (s : u32) : result u32 :=
match n with
| O => Fail_ OutOfFuel
@@ -19,12 +20,14 @@ Fixpoint sum_loop (n : nat) (max : u32) (i : u32) (s : u32) : result u32 :=
end
.
-(** [loops::sum]: forward function *)
+(** [loops::sum]: forward function
+ Source: 'src/loops.rs', lines 4:0-4:27 *)
Definition sum (n : nat) (max : u32) : result u32 :=
sum_loop n max 0%u32 0%u32
.
-(** [loops::sum_with_mut_borrows]: loop 0: forward function *)
+(** [loops::sum_with_mut_borrows]: loop 0: forward function
+ Source: 'src/loops.rs', lines 19:0-31:1 *)
Fixpoint sum_with_mut_borrows_loop
(n : nat) (max : u32) (mi : u32) (ms : u32) : result u32 :=
match n with
@@ -39,12 +42,14 @@ Fixpoint sum_with_mut_borrows_loop
end
.
-(** [loops::sum_with_mut_borrows]: forward function *)
+(** [loops::sum_with_mut_borrows]: forward function
+ Source: 'src/loops.rs', lines 19:0-19:44 *)
Definition sum_with_mut_borrows (n : nat) (max : u32) : result u32 :=
sum_with_mut_borrows_loop n max 0%u32 0%u32
.
-(** [loops::sum_with_shared_borrows]: loop 0: forward function *)
+(** [loops::sum_with_shared_borrows]: loop 0: forward function
+ Source: 'src/loops.rs', lines 34:0-48:1 *)
Fixpoint sum_with_shared_borrows_loop
(n : nat) (max : u32) (i : u32) (s : u32) : result u32 :=
match n with
@@ -59,13 +64,15 @@ Fixpoint sum_with_shared_borrows_loop
end
.
-(** [loops::sum_with_shared_borrows]: forward function *)
+(** [loops::sum_with_shared_borrows]: forward function
+ Source: 'src/loops.rs', lines 34:0-34:47 *)
Definition sum_with_shared_borrows (n : nat) (max : u32) : result u32 :=
sum_with_shared_borrows_loop n max 0%u32 0%u32
.
(** [loops::clear]: loop 0: 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/loops.rs', lines 52:0-58:1 *)
Fixpoint clear_loop
(n : nat) (v : alloc_vec_Vec u32) (i : usize) : result (alloc_vec_Vec u32) :=
match n with
@@ -77,20 +84,22 @@ Fixpoint clear_loop
i1 <- usize_add i 1%usize;
v0 <-
alloc_vec_Vec_index_mut_back u32 usize
- (core_slice_index_usize_coresliceindexSliceIndexInst u32) v i 0%u32;
+ (core_slice_index_SliceIndexUsizeSliceTInst u32) v i 0%u32;
clear_loop n0 v0 i1)
else Return v
end
.
(** [loops::clear]: 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/loops.rs', lines 52:0-52:30 *)
Definition clear
(n : nat) (v : alloc_vec_Vec u32) : result (alloc_vec_Vec u32) :=
clear_loop n v 0%usize
.
-(** [loops::List] *)
+(** [loops::List]
+ Source: 'src/loops.rs', lines 60:0-60:16 *)
Inductive List_t (T : Type) :=
| List_Cons : T -> List_t T -> List_t T
| List_Nil : List_t T
@@ -99,7 +108,8 @@ Inductive List_t (T : Type) :=
Arguments List_Cons { _ }.
Arguments List_Nil { _ }.
-(** [loops::list_mem]: loop 0: forward function *)
+(** [loops::list_mem]: loop 0: forward function
+ Source: 'src/loops.rs', lines 66:0-75:1 *)
Fixpoint list_mem_loop (n : nat) (x : u32) (ls : List_t u32) : result bool :=
match n with
| O => Fail_ OutOfFuel
@@ -111,12 +121,14 @@ Fixpoint list_mem_loop (n : nat) (x : u32) (ls : List_t u32) : result bool :=
end
.
-(** [loops::list_mem]: forward function *)
+(** [loops::list_mem]: forward function
+ Source: 'src/loops.rs', lines 66:0-66:52 *)
Definition list_mem (n : nat) (x : u32) (ls : List_t u32) : result bool :=
list_mem_loop n x ls
.
-(** [loops::list_nth_mut_loop]: loop 0: forward function *)
+(** [loops::list_nth_mut_loop]: loop 0: forward function
+ Source: 'src/loops.rs', lines 78:0-88:1 *)
Fixpoint list_nth_mut_loop_loop
(T : Type) (n : nat) (ls : List_t T) (i : u32) : result T :=
match n with
@@ -132,13 +144,15 @@ Fixpoint list_nth_mut_loop_loop
end
.
-(** [loops::list_nth_mut_loop]: forward function *)
+(** [loops::list_nth_mut_loop]: forward function
+ Source: 'src/loops.rs', lines 78:0-78:71 *)
Definition list_nth_mut_loop
(T : Type) (n : nat) (ls : List_t T) (i : u32) : result T :=
list_nth_mut_loop_loop T n ls i
.
-(** [loops::list_nth_mut_loop]: loop 0: backward function 0 *)
+(** [loops::list_nth_mut_loop]: loop 0: backward function 0
+ Source: 'src/loops.rs', lines 78:0-88:1 *)
Fixpoint list_nth_mut_loop_loop_back
(T : Type) (n : nat) (ls : List_t T) (i : u32) (ret : T) :
result (List_t T)
@@ -159,7 +173,8 @@ Fixpoint list_nth_mut_loop_loop_back
end
.
-(** [loops::list_nth_mut_loop]: backward function 0 *)
+(** [loops::list_nth_mut_loop]: backward function 0
+ Source: 'src/loops.rs', lines 78:0-78:71 *)
Definition list_nth_mut_loop_back
(T : Type) (n : nat) (ls : List_t T) (i : u32) (ret : T) :
result (List_t T)
@@ -167,7 +182,8 @@ Definition list_nth_mut_loop_back
list_nth_mut_loop_loop_back T n ls i ret
.
-(** [loops::list_nth_shared_loop]: loop 0: forward function *)
+(** [loops::list_nth_shared_loop]: loop 0: forward function
+ Source: 'src/loops.rs', lines 91:0-101:1 *)
Fixpoint list_nth_shared_loop_loop
(T : Type) (n : nat) (ls : List_t T) (i : u32) : result T :=
match n with
@@ -183,13 +199,15 @@ Fixpoint list_nth_shared_loop_loop
end
.
-(** [loops::list_nth_shared_loop]: forward function *)
+(** [loops::list_nth_shared_loop]: forward function
+ Source: 'src/loops.rs', lines 91:0-91:66 *)
Definition list_nth_shared_loop
(T : Type) (n : nat) (ls : List_t T) (i : u32) : result T :=
list_nth_shared_loop_loop T n ls i
.
-(** [loops::get_elem_mut]: loop 0: forward function *)
+(** [loops::get_elem_mut]: loop 0: forward function
+ Source: 'src/loops.rs', lines 103:0-117:1 *)
Fixpoint get_elem_mut_loop
(n : nat) (x : usize) (ls : List_t usize) : result usize :=
match n with
@@ -202,19 +220,20 @@ Fixpoint get_elem_mut_loop
end
.
-(** [loops::get_elem_mut]: forward function *)
+(** [loops::get_elem_mut]: forward function
+ Source: 'src/loops.rs', lines 103:0-103:73 *)
Definition get_elem_mut
(n : nat) (slots : alloc_vec_Vec (List_t usize)) (x : usize) :
result usize
:=
l <-
alloc_vec_Vec_index_mut (List_t usize) usize
- (core_slice_index_usize_coresliceindexSliceIndexInst (List_t usize))
- slots 0%usize;
+ (core_slice_index_SliceIndexUsizeSliceTInst (List_t usize)) slots 0%usize;
get_elem_mut_loop n x l
.
-(** [loops::get_elem_mut]: loop 0: backward function 0 *)
+(** [loops::get_elem_mut]: loop 0: backward function 0
+ Source: 'src/loops.rs', lines 103:0-117:1 *)
Fixpoint get_elem_mut_loop_back
(n : nat) (x : usize) (ls : List_t usize) (ret : usize) :
result (List_t usize)
@@ -233,22 +252,23 @@ Fixpoint get_elem_mut_loop_back
end
.
-(** [loops::get_elem_mut]: backward function 0 *)
+(** [loops::get_elem_mut]: backward function 0
+ Source: 'src/loops.rs', lines 103:0-103:73 *)
Definition get_elem_mut_back
(n : nat) (slots : alloc_vec_Vec (List_t usize)) (x : usize) (ret : usize) :
result (alloc_vec_Vec (List_t usize))
:=
l <-
alloc_vec_Vec_index_mut (List_t usize) usize
- (core_slice_index_usize_coresliceindexSliceIndexInst (List_t usize))
- slots 0%usize;
+ (core_slice_index_SliceIndexUsizeSliceTInst (List_t usize)) slots 0%usize;
l0 <- get_elem_mut_loop_back n x l ret;
alloc_vec_Vec_index_mut_back (List_t usize) usize
- (core_slice_index_usize_coresliceindexSliceIndexInst (List_t usize)) slots
- 0%usize l0
+ (core_slice_index_SliceIndexUsizeSliceTInst (List_t usize)) slots 0%usize
+ l0
.
-(** [loops::get_elem_shared]: loop 0: forward function *)
+(** [loops::get_elem_shared]: loop 0: forward function
+ Source: 'src/loops.rs', lines 119:0-133:1 *)
Fixpoint get_elem_shared_loop
(n : nat) (x : usize) (ls : List_t usize) : result usize :=
match n with
@@ -262,34 +282,38 @@ Fixpoint get_elem_shared_loop
end
.
-(** [loops::get_elem_shared]: forward function *)
+(** [loops::get_elem_shared]: forward function
+ Source: 'src/loops.rs', lines 119:0-119:68 *)
Definition get_elem_shared
(n : nat) (slots : alloc_vec_Vec (List_t usize)) (x : usize) :
result usize
:=
l <-
alloc_vec_Vec_index (List_t usize) usize
- (core_slice_index_usize_coresliceindexSliceIndexInst (List_t usize))
- slots 0%usize;
+ (core_slice_index_SliceIndexUsizeSliceTInst (List_t usize)) slots 0%usize;
get_elem_shared_loop n x l
.
-(** [loops::id_mut]: forward function *)
+(** [loops::id_mut]: forward function
+ Source: 'src/loops.rs', lines 135:0-135:50 *)
Definition id_mut (T : Type) (ls : List_t T) : result (List_t T) :=
Return ls.
-(** [loops::id_mut]: backward function 0 *)
+(** [loops::id_mut]: backward function 0
+ Source: 'src/loops.rs', lines 135:0-135:50 *)
Definition id_mut_back
(T : Type) (ls : List_t T) (ret : List_t T) : result (List_t T) :=
Return ret
.
-(** [loops::id_shared]: forward function *)
+(** [loops::id_shared]: forward function
+ Source: 'src/loops.rs', lines 139:0-139:45 *)
Definition id_shared (T : Type) (ls : List_t T) : result (List_t T) :=
Return ls
.
-(** [loops::list_nth_mut_loop_with_id]: loop 0: forward function *)
+(** [loops::list_nth_mut_loop_with_id]: loop 0: forward function
+ Source: 'src/loops.rs', lines 144:0-155:1 *)
Fixpoint list_nth_mut_loop_with_id_loop
(T : Type) (n : nat) (i : u32) (ls : List_t T) : result T :=
match n with
@@ -305,13 +329,15 @@ Fixpoint list_nth_mut_loop_with_id_loop
end
.
-(** [loops::list_nth_mut_loop_with_id]: forward function *)
+(** [loops::list_nth_mut_loop_with_id]: forward function
+ Source: 'src/loops.rs', lines 144:0-144:75 *)
Definition list_nth_mut_loop_with_id
(T : Type) (n : nat) (ls : List_t T) (i : u32) : result T :=
ls0 <- id_mut T ls; list_nth_mut_loop_with_id_loop T n i ls0
.
-(** [loops::list_nth_mut_loop_with_id]: loop 0: backward function 0 *)
+(** [loops::list_nth_mut_loop_with_id]: loop 0: backward function 0
+ Source: 'src/loops.rs', lines 144:0-155:1 *)
Fixpoint list_nth_mut_loop_with_id_loop_back
(T : Type) (n : nat) (i : u32) (ls : List_t T) (ret : T) :
result (List_t T)
@@ -332,7 +358,8 @@ Fixpoint list_nth_mut_loop_with_id_loop_back
end
.
-(** [loops::list_nth_mut_loop_with_id]: backward function 0 *)
+(** [loops::list_nth_mut_loop_with_id]: backward function 0
+ Source: 'src/loops.rs', lines 144:0-144:75 *)
Definition list_nth_mut_loop_with_id_back
(T : Type) (n : nat) (ls : List_t T) (i : u32) (ret : T) :
result (List_t T)
@@ -342,7 +369,8 @@ Definition list_nth_mut_loop_with_id_back
id_mut_back T ls l
.
-(** [loops::list_nth_shared_loop_with_id]: loop 0: forward function *)
+(** [loops::list_nth_shared_loop_with_id]: loop 0: forward function
+ Source: 'src/loops.rs', lines 158:0-169:1 *)
Fixpoint list_nth_shared_loop_with_id_loop
(T : Type) (n : nat) (i : u32) (ls : List_t T) : result T :=
match n with
@@ -359,13 +387,15 @@ Fixpoint list_nth_shared_loop_with_id_loop
end
.
-(** [loops::list_nth_shared_loop_with_id]: forward function *)
+(** [loops::list_nth_shared_loop_with_id]: forward function
+ Source: 'src/loops.rs', lines 158:0-158:70 *)
Definition list_nth_shared_loop_with_id
(T : Type) (n : nat) (ls : List_t T) (i : u32) : result T :=
ls0 <- id_shared T ls; list_nth_shared_loop_with_id_loop T n i ls0
.
-(** [loops::list_nth_mut_loop_pair]: loop 0: forward function *)
+(** [loops::list_nth_mut_loop_pair]: loop 0: forward function
+ Source: 'src/loops.rs', lines 174:0-195:1 *)
Fixpoint list_nth_mut_loop_pair_loop
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -388,7 +418,8 @@ Fixpoint list_nth_mut_loop_pair_loop
end
.
-(** [loops::list_nth_mut_loop_pair]: forward function *)
+(** [loops::list_nth_mut_loop_pair]: forward function
+ Source: 'src/loops.rs', lines 174:0-178:27 *)
Definition list_nth_mut_loop_pair
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -396,7 +427,8 @@ Definition list_nth_mut_loop_pair
list_nth_mut_loop_pair_loop T n ls0 ls1 i
.
-(** [loops::list_nth_mut_loop_pair]: loop 0: backward function 0 *)
+(** [loops::list_nth_mut_loop_pair]: loop 0: backward function 0
+ Source: 'src/loops.rs', lines 174:0-195:1 *)
Fixpoint list_nth_mut_loop_pair_loop_back'a
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) (ret : T) :
result (List_t T)
@@ -421,7 +453,8 @@ Fixpoint list_nth_mut_loop_pair_loop_back'a
end
.
-(** [loops::list_nth_mut_loop_pair]: backward function 0 *)
+(** [loops::list_nth_mut_loop_pair]: backward function 0
+ Source: 'src/loops.rs', lines 174:0-178:27 *)
Definition list_nth_mut_loop_pair_back'a
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) (ret : T) :
result (List_t T)
@@ -429,7 +462,8 @@ Definition list_nth_mut_loop_pair_back'a
list_nth_mut_loop_pair_loop_back'a T n ls0 ls1 i ret
.
-(** [loops::list_nth_mut_loop_pair]: loop 0: backward function 1 *)
+(** [loops::list_nth_mut_loop_pair]: loop 0: backward function 1
+ Source: 'src/loops.rs', lines 174:0-195:1 *)
Fixpoint list_nth_mut_loop_pair_loop_back'b
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) (ret : T) :
result (List_t T)
@@ -454,7 +488,8 @@ Fixpoint list_nth_mut_loop_pair_loop_back'b
end
.
-(** [loops::list_nth_mut_loop_pair]: backward function 1 *)
+(** [loops::list_nth_mut_loop_pair]: backward function 1
+ Source: 'src/loops.rs', lines 174:0-178:27 *)
Definition list_nth_mut_loop_pair_back'b
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) (ret : T) :
result (List_t T)
@@ -462,7 +497,8 @@ Definition list_nth_mut_loop_pair_back'b
list_nth_mut_loop_pair_loop_back'b T n ls0 ls1 i ret
.
-(** [loops::list_nth_shared_loop_pair]: loop 0: forward function *)
+(** [loops::list_nth_shared_loop_pair]: loop 0: forward function
+ Source: 'src/loops.rs', lines 198:0-219:1 *)
Fixpoint list_nth_shared_loop_pair_loop
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -485,7 +521,8 @@ Fixpoint list_nth_shared_loop_pair_loop
end
.
-(** [loops::list_nth_shared_loop_pair]: forward function *)
+(** [loops::list_nth_shared_loop_pair]: forward function
+ Source: 'src/loops.rs', lines 198:0-202:19 *)
Definition list_nth_shared_loop_pair
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -493,7 +530,8 @@ Definition list_nth_shared_loop_pair
list_nth_shared_loop_pair_loop T n ls0 ls1 i
.
-(** [loops::list_nth_mut_loop_pair_merge]: loop 0: forward function *)
+(** [loops::list_nth_mut_loop_pair_merge]: loop 0: forward function
+ Source: 'src/loops.rs', lines 223:0-238:1 *)
Fixpoint list_nth_mut_loop_pair_merge_loop
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -517,7 +555,8 @@ Fixpoint list_nth_mut_loop_pair_merge_loop
end
.
-(** [loops::list_nth_mut_loop_pair_merge]: forward function *)
+(** [loops::list_nth_mut_loop_pair_merge]: forward function
+ Source: 'src/loops.rs', lines 223:0-227:27 *)
Definition list_nth_mut_loop_pair_merge
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -525,7 +564,8 @@ Definition list_nth_mut_loop_pair_merge
list_nth_mut_loop_pair_merge_loop T n ls0 ls1 i
.
-(** [loops::list_nth_mut_loop_pair_merge]: loop 0: backward function 0 *)
+(** [loops::list_nth_mut_loop_pair_merge]: loop 0: backward function 0
+ Source: 'src/loops.rs', lines 223:0-238:1 *)
Fixpoint list_nth_mut_loop_pair_merge_loop_back
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32)
(ret : (T * T)) :
@@ -552,7 +592,8 @@ Fixpoint list_nth_mut_loop_pair_merge_loop_back
end
.
-(** [loops::list_nth_mut_loop_pair_merge]: backward function 0 *)
+(** [loops::list_nth_mut_loop_pair_merge]: backward function 0
+ Source: 'src/loops.rs', lines 223:0-227:27 *)
Definition list_nth_mut_loop_pair_merge_back
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32)
(ret : (T * T)) :
@@ -561,7 +602,8 @@ Definition list_nth_mut_loop_pair_merge_back
list_nth_mut_loop_pair_merge_loop_back T n ls0 ls1 i ret
.
-(** [loops::list_nth_shared_loop_pair_merge]: loop 0: forward function *)
+(** [loops::list_nth_shared_loop_pair_merge]: loop 0: forward function
+ Source: 'src/loops.rs', lines 241:0-256:1 *)
Fixpoint list_nth_shared_loop_pair_merge_loop
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -585,7 +627,8 @@ Fixpoint list_nth_shared_loop_pair_merge_loop
end
.
-(** [loops::list_nth_shared_loop_pair_merge]: forward function *)
+(** [loops::list_nth_shared_loop_pair_merge]: forward function
+ Source: 'src/loops.rs', lines 241:0-245:19 *)
Definition list_nth_shared_loop_pair_merge
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -593,7 +636,8 @@ Definition list_nth_shared_loop_pair_merge
list_nth_shared_loop_pair_merge_loop T n ls0 ls1 i
.
-(** [loops::list_nth_mut_shared_loop_pair]: loop 0: forward function *)
+(** [loops::list_nth_mut_shared_loop_pair]: loop 0: forward function
+ Source: 'src/loops.rs', lines 259:0-274:1 *)
Fixpoint list_nth_mut_shared_loop_pair_loop
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -617,7 +661,8 @@ Fixpoint list_nth_mut_shared_loop_pair_loop
end
.
-(** [loops::list_nth_mut_shared_loop_pair]: forward function *)
+(** [loops::list_nth_mut_shared_loop_pair]: forward function
+ Source: 'src/loops.rs', lines 259:0-263:23 *)
Definition list_nth_mut_shared_loop_pair
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -625,7 +670,8 @@ Definition list_nth_mut_shared_loop_pair
list_nth_mut_shared_loop_pair_loop T n ls0 ls1 i
.
-(** [loops::list_nth_mut_shared_loop_pair]: loop 0: backward function 0 *)
+(** [loops::list_nth_mut_shared_loop_pair]: loop 0: backward function 0
+ Source: 'src/loops.rs', lines 259:0-274:1 *)
Fixpoint list_nth_mut_shared_loop_pair_loop_back
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) (ret : T) :
result (List_t T)
@@ -650,7 +696,8 @@ Fixpoint list_nth_mut_shared_loop_pair_loop_back
end
.
-(** [loops::list_nth_mut_shared_loop_pair]: backward function 0 *)
+(** [loops::list_nth_mut_shared_loop_pair]: backward function 0
+ Source: 'src/loops.rs', lines 259:0-263:23 *)
Definition list_nth_mut_shared_loop_pair_back
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) (ret : T) :
result (List_t T)
@@ -658,7 +705,8 @@ Definition list_nth_mut_shared_loop_pair_back
list_nth_mut_shared_loop_pair_loop_back T n ls0 ls1 i ret
.
-(** [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: forward function *)
+(** [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: forward function
+ Source: 'src/loops.rs', lines 278:0-293:1 *)
Fixpoint list_nth_mut_shared_loop_pair_merge_loop
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -682,7 +730,8 @@ Fixpoint list_nth_mut_shared_loop_pair_merge_loop
end
.
-(** [loops::list_nth_mut_shared_loop_pair_merge]: forward function *)
+(** [loops::list_nth_mut_shared_loop_pair_merge]: forward function
+ Source: 'src/loops.rs', lines 278:0-282:23 *)
Definition list_nth_mut_shared_loop_pair_merge
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -690,7 +739,8 @@ Definition list_nth_mut_shared_loop_pair_merge
list_nth_mut_shared_loop_pair_merge_loop T n ls0 ls1 i
.
-(** [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: backward function 0 *)
+(** [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: backward function 0
+ Source: 'src/loops.rs', lines 278:0-293:1 *)
Fixpoint list_nth_mut_shared_loop_pair_merge_loop_back
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) (ret : T) :
result (List_t T)
@@ -716,7 +766,8 @@ Fixpoint list_nth_mut_shared_loop_pair_merge_loop_back
end
.
-(** [loops::list_nth_mut_shared_loop_pair_merge]: backward function 0 *)
+(** [loops::list_nth_mut_shared_loop_pair_merge]: backward function 0
+ Source: 'src/loops.rs', lines 278:0-282:23 *)
Definition list_nth_mut_shared_loop_pair_merge_back
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) (ret : T) :
result (List_t T)
@@ -724,7 +775,8 @@ Definition list_nth_mut_shared_loop_pair_merge_back
list_nth_mut_shared_loop_pair_merge_loop_back T n ls0 ls1 i ret
.
-(** [loops::list_nth_shared_mut_loop_pair]: loop 0: forward function *)
+(** [loops::list_nth_shared_mut_loop_pair]: loop 0: forward function
+ Source: 'src/loops.rs', lines 297:0-312:1 *)
Fixpoint list_nth_shared_mut_loop_pair_loop
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -748,7 +800,8 @@ Fixpoint list_nth_shared_mut_loop_pair_loop
end
.
-(** [loops::list_nth_shared_mut_loop_pair]: forward function *)
+(** [loops::list_nth_shared_mut_loop_pair]: forward function
+ Source: 'src/loops.rs', lines 297:0-301:23 *)
Definition list_nth_shared_mut_loop_pair
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -756,7 +809,8 @@ Definition list_nth_shared_mut_loop_pair
list_nth_shared_mut_loop_pair_loop T n ls0 ls1 i
.
-(** [loops::list_nth_shared_mut_loop_pair]: loop 0: backward function 1 *)
+(** [loops::list_nth_shared_mut_loop_pair]: loop 0: backward function 1
+ Source: 'src/loops.rs', lines 297:0-312:1 *)
Fixpoint list_nth_shared_mut_loop_pair_loop_back
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) (ret : T) :
result (List_t T)
@@ -781,7 +835,8 @@ Fixpoint list_nth_shared_mut_loop_pair_loop_back
end
.
-(** [loops::list_nth_shared_mut_loop_pair]: backward function 1 *)
+(** [loops::list_nth_shared_mut_loop_pair]: backward function 1
+ Source: 'src/loops.rs', lines 297:0-301:23 *)
Definition list_nth_shared_mut_loop_pair_back
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) (ret : T) :
result (List_t T)
@@ -789,7 +844,8 @@ Definition list_nth_shared_mut_loop_pair_back
list_nth_shared_mut_loop_pair_loop_back T n ls0 ls1 i ret
.
-(** [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: forward function *)
+(** [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: forward function
+ Source: 'src/loops.rs', lines 316:0-331:1 *)
Fixpoint list_nth_shared_mut_loop_pair_merge_loop
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -813,7 +869,8 @@ Fixpoint list_nth_shared_mut_loop_pair_merge_loop
end
.
-(** [loops::list_nth_shared_mut_loop_pair_merge]: forward function *)
+(** [loops::list_nth_shared_mut_loop_pair_merge]: forward function
+ Source: 'src/loops.rs', lines 316:0-320:23 *)
Definition list_nth_shared_mut_loop_pair_merge
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
@@ -821,7 +878,8 @@ Definition list_nth_shared_mut_loop_pair_merge
list_nth_shared_mut_loop_pair_merge_loop T n ls0 ls1 i
.
-(** [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: backward function 0 *)
+(** [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: backward function 0
+ Source: 'src/loops.rs', lines 316:0-331:1 *)
Fixpoint list_nth_shared_mut_loop_pair_merge_loop_back
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) (ret : T) :
result (List_t T)
@@ -847,7 +905,8 @@ Fixpoint list_nth_shared_mut_loop_pair_merge_loop_back
end
.
-(** [loops::list_nth_shared_mut_loop_pair_merge]: backward function 0 *)
+(** [loops::list_nth_shared_mut_loop_pair_merge]: backward function 0
+ Source: 'src/loops.rs', lines 316:0-320:23 *)
Definition list_nth_shared_mut_loop_pair_merge_back
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) (ret : T) :
result (List_t T)
diff --git a/tests/coq/misc/NoNestedBorrows.v b/tests/coq/misc/NoNestedBorrows.v
index c7af496f..b044d24f 100644
--- a/tests/coq/misc/NoNestedBorrows.v
+++ b/tests/coq/misc/NoNestedBorrows.v
@@ -8,14 +8,16 @@ Import ListNotations.
Local Open Scope Primitives_scope.
Module NoNestedBorrows.
-(** [no_nested_borrows::Pair] *)
+(** [no_nested_borrows::Pair]
+ Source: 'src/no_nested_borrows.rs', lines 4:0-4:23 *)
Record Pair_t (T1 T2 : Type) := mkPair_t { pair_x : T1; pair_y : T2; }.
Arguments mkPair_t { _ _ }.
Arguments pair_x { _ _ }.
Arguments pair_y { _ _ }.
-(** [no_nested_borrows::List] *)
+(** [no_nested_borrows::List]
+ Source: 'src/no_nested_borrows.rs', lines 9:0-9:16 *)
Inductive List_t (T : Type) :=
| List_Cons : T -> List_t T -> List_t T
| List_Nil : List_t T
@@ -24,21 +26,26 @@ Inductive List_t (T : Type) :=
Arguments List_Cons { _ }.
Arguments List_Nil { _ }.
-(** [no_nested_borrows::One] *)
+(** [no_nested_borrows::One]
+ Source: 'src/no_nested_borrows.rs', lines 20:0-20:16 *)
Inductive One_t (T1 : Type) := | One_One : T1 -> One_t T1.
Arguments One_One { _ }.
-(** [no_nested_borrows::EmptyEnum] *)
+(** [no_nested_borrows::EmptyEnum]
+ Source: 'src/no_nested_borrows.rs', lines 26:0-26:18 *)
Inductive EmptyEnum_t := | EmptyEnum_Empty : EmptyEnum_t.
-(** [no_nested_borrows::Enum] *)
+(** [no_nested_borrows::Enum]
+ Source: 'src/no_nested_borrows.rs', lines 32:0-32:13 *)
Inductive Enum_t := | Enum_Variant1 : Enum_t | Enum_Variant2 : Enum_t.
-(** [no_nested_borrows::EmptyStruct] *)
+(** [no_nested_borrows::EmptyStruct]
+ Source: 'src/no_nested_borrows.rs', lines 39:0-39:22 *)
Record EmptyStruct_t := mkEmptyStruct_t { }.
-(** [no_nested_borrows::Sum] *)
+(** [no_nested_borrows::Sum]
+ Source: 'src/no_nested_borrows.rs', lines 41:0-41:20 *)
Inductive Sum_t (T1 T2 : Type) :=
| Sum_Left : T1 -> Sum_t T1 T2
| Sum_Right : T2 -> Sum_t T1 T2
@@ -47,59 +54,72 @@ Inductive Sum_t (T1 T2 : Type) :=
Arguments Sum_Left { _ _ }.
Arguments Sum_Right { _ _ }.
-(** [no_nested_borrows::neg_test]: forward function *)
+(** [no_nested_borrows::neg_test]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 48:0-48:30 *)
Definition neg_test (x : i32) : result i32 :=
i32_neg x.
-(** [no_nested_borrows::add_test]: forward function *)
+(** [no_nested_borrows::add_test]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 54:0-54:38 *)
Definition add_test (x : u32) (y : u32) : result u32 :=
u32_add x y.
-(** [no_nested_borrows::subs_test]: forward function *)
+(** [no_nested_borrows::subs_test]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 60:0-60:39 *)
Definition subs_test (x : u32) (y : u32) : result u32 :=
u32_sub x y.
-(** [no_nested_borrows::div_test]: forward function *)
+(** [no_nested_borrows::div_test]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 66:0-66:38 *)
Definition div_test (x : u32) (y : u32) : result u32 :=
u32_div x y.
-(** [no_nested_borrows::div_test1]: forward function *)
+(** [no_nested_borrows::div_test1]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 73:0-73:31 *)
Definition div_test1 (x : u32) : result u32 :=
u32_div x 2%u32.
-(** [no_nested_borrows::rem_test]: forward function *)
+(** [no_nested_borrows::rem_test]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 78:0-78:38 *)
Definition rem_test (x : u32) (y : u32) : result u32 :=
u32_rem x y.
-(** [no_nested_borrows::mul_test]: forward function *)
+(** [no_nested_borrows::mul_test]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 82:0-82:38 *)
Definition mul_test (x : u32) (y : u32) : result u32 :=
u32_mul x y.
-(** [no_nested_borrows::CONST0] *)
+(** [no_nested_borrows::CONST0]
+ Source: 'src/no_nested_borrows.rs', lines 91:0-91:23 *)
Definition const0_body : result usize := usize_add 1%usize 1%usize.
Definition const0_c : usize := const0_body%global.
-(** [no_nested_borrows::CONST1] *)
+(** [no_nested_borrows::CONST1]
+ Source: 'src/no_nested_borrows.rs', lines 92:0-92:23 *)
Definition const1_body : result usize := usize_mul 2%usize 2%usize.
Definition const1_c : usize := const1_body%global.
-(** [no_nested_borrows::cast_test]: forward function *)
+(** [no_nested_borrows::cast_test]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 94:0-94:31 *)
Definition cast_test (x : u32) : result i32 :=
scalar_cast U32 I32 x.
-(** [no_nested_borrows::test2]: forward function *)
+(** [no_nested_borrows::test2]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 99:0-99:14 *)
Definition test2 : result unit :=
_ <- u32_add 23%u32 44%u32; Return tt.
(** Unit test for [no_nested_borrows::test2] *)
Check (test2 )%return.
-(** [no_nested_borrows::get_max]: forward function *)
+(** [no_nested_borrows::get_max]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 111:0-111:37 *)
Definition get_max (x : u32) (y : u32) : result u32 :=
if x s>= y then Return x else Return y
.
-(** [no_nested_borrows::test3]: forward function *)
+(** [no_nested_borrows::test3]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 119:0-119:14 *)
Definition test3 : result unit :=
x <- get_max 4%u32 3%u32;
y <- get_max 10%u32 11%u32;
@@ -110,7 +130,8 @@ Definition test3 : result unit :=
(** Unit test for [no_nested_borrows::test3] *)
Check (test3 )%return.
-(** [no_nested_borrows::test_neg1]: forward function *)
+(** [no_nested_borrows::test_neg1]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 126:0-126:18 *)
Definition test_neg1 : result unit :=
y <- i32_neg 3%i32; if negb (y s= (-3)%i32) then Fail_ Failure else Return tt
.
@@ -118,7 +139,8 @@ Definition test_neg1 : result unit :=
(** Unit test for [no_nested_borrows::test_neg1] *)
Check (test_neg1 )%return.
-(** [no_nested_borrows::refs_test1]: forward function *)
+(** [no_nested_borrows::refs_test1]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 133:0-133:19 *)
Definition refs_test1 : result unit :=
if negb (1%i32 s= 1%i32) then Fail_ Failure else Return tt
.
@@ -126,7 +148,8 @@ Definition refs_test1 : result unit :=
(** Unit test for [no_nested_borrows::refs_test1] *)
Check (refs_test1 )%return.
-(** [no_nested_borrows::refs_test2]: forward function *)
+(** [no_nested_borrows::refs_test2]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 144:0-144:19 *)
Definition refs_test2 : result unit :=
if negb (2%i32 s= 2%i32)
then Fail_ Failure
@@ -142,38 +165,45 @@ Definition refs_test2 : result unit :=
(** Unit test for [no_nested_borrows::refs_test2] *)
Check (refs_test2 )%return.
-(** [no_nested_borrows::test_list1]: forward function *)
+(** [no_nested_borrows::test_list1]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 160:0-160:19 *)
Definition test_list1 : result unit :=
Return tt.
(** Unit test for [no_nested_borrows::test_list1] *)
Check (test_list1 )%return.
-(** [no_nested_borrows::test_box1]: forward function *)
+(** [no_nested_borrows::test_box1]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 165:0-165:18 *)
Definition test_box1 : result unit :=
- let b := 1%i32 in
- let x := b in
+ let b := 0%i32 in
+ b0 <- alloc_boxed_Box_deref_mut_back i32 b 1%i32;
+ x <- alloc_boxed_Box_deref i32 b0;
if negb (x s= 1%i32) then Fail_ Failure else Return tt
.
(** Unit test for [no_nested_borrows::test_box1] *)
Check (test_box1 )%return.
-(** [no_nested_borrows::copy_int]: forward function *)
+(** [no_nested_borrows::copy_int]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 175:0-175:30 *)
Definition copy_int (x : i32) : result i32 :=
Return x.
-(** [no_nested_borrows::test_unreachable]: forward function *)
+(** [no_nested_borrows::test_unreachable]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 181:0-181:32 *)
Definition test_unreachable (b : bool) : result unit :=
if b then Fail_ Failure else Return tt
.
-(** [no_nested_borrows::test_panic]: forward function *)
+(** [no_nested_borrows::test_panic]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 189:0-189:26 *)
Definition test_panic (b : bool) : result unit :=
if b then Fail_ Failure else Return tt
.
-(** [no_nested_borrows::test_copy_int]: forward function *)
+(** [no_nested_borrows::test_copy_int]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 196:0-196:22 *)
Definition test_copy_int : result unit :=
y <- copy_int 0%i32; if negb (0%i32 s= y) then Fail_ Failure else Return tt
.
@@ -181,12 +211,14 @@ Definition test_copy_int : result unit :=
(** Unit test for [no_nested_borrows::test_copy_int] *)
Check (test_copy_int )%return.
-(** [no_nested_borrows::is_cons]: forward function *)
+(** [no_nested_borrows::is_cons]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 203:0-203:38 *)
Definition is_cons (T : Type) (l : List_t T) : result bool :=
match l with | List_Cons t l0 => Return true | List_Nil => Return false end
.
-(** [no_nested_borrows::test_is_cons]: forward function *)
+(** [no_nested_borrows::test_is_cons]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 210:0-210:21 *)
Definition test_is_cons : result unit :=
let l := List_Nil in
b <- is_cons i32 (List_Cons 0%i32 l);
@@ -196,7 +228,8 @@ Definition test_is_cons : result unit :=
(** Unit test for [no_nested_borrows::test_is_cons] *)
Check (test_is_cons )%return.
-(** [no_nested_borrows::split_list]: forward function *)
+(** [no_nested_borrows::split_list]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 216:0-216:48 *)
Definition split_list (T : Type) (l : List_t T) : result (T * (List_t T)) :=
match l with
| List_Cons hd tl => Return (hd, tl)
@@ -204,7 +237,8 @@ Definition split_list (T : Type) (l : List_t T) : result (T * (List_t T)) :=
end
.
-(** [no_nested_borrows::test_split_list]: forward function *)
+(** [no_nested_borrows::test_split_list]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 224:0-224:24 *)
Definition test_split_list : result unit :=
let l := List_Nil in
p <- split_list i32 (List_Cons 0%i32 l);
@@ -215,18 +249,21 @@ Definition test_split_list : result unit :=
(** Unit test for [no_nested_borrows::test_split_list] *)
Check (test_split_list )%return.
-(** [no_nested_borrows::choose]: forward function *)
+(** [no_nested_borrows::choose]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 231:0-231:70 *)
Definition choose (T : Type) (b : bool) (x : T) (y : T) : result T :=
if b then Return x else Return y
.
-(** [no_nested_borrows::choose]: backward function 0 *)
+(** [no_nested_borrows::choose]: backward function 0
+ Source: 'src/no_nested_borrows.rs', lines 231:0-231:70 *)
Definition choose_back
(T : Type) (b : bool) (x : T) (y : T) (ret : T) : result (T * T) :=
if b then Return (ret, y) else Return (x, ret)
.
-(** [no_nested_borrows::choose_test]: forward function *)
+(** [no_nested_borrows::choose_test]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 239:0-239:20 *)
Definition choose_test : result unit :=
z <- choose i32 true 0%i32 0%i32;
z0 <- i32_add z 1%i32;
@@ -243,16 +280,19 @@ Definition choose_test : result unit :=
(** Unit test for [no_nested_borrows::choose_test] *)
Check (choose_test )%return.
-(** [no_nested_borrows::test_char]: forward function *)
+(** [no_nested_borrows::test_char]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 251:0-251:26 *)
Definition test_char : result char :=
Return (char_of_byte Coq.Init.Byte.x61).
-(** [no_nested_borrows::Tree] *)
+(** [no_nested_borrows::Tree]
+ Source: 'src/no_nested_borrows.rs', lines 256:0-256:16 *)
Inductive Tree_t (T : Type) :=
| Tree_Leaf : T -> Tree_t T
| Tree_Node : T -> NodeElem_t T -> Tree_t T -> Tree_t T
-(** [no_nested_borrows::NodeElem] *)
+(** [no_nested_borrows::NodeElem]
+ Source: 'src/no_nested_borrows.rs', lines 261:0-261:20 *)
with NodeElem_t (T : Type) :=
| NodeElem_Cons : Tree_t T -> NodeElem_t T -> NodeElem_t T
| NodeElem_Nil : NodeElem_t T
@@ -264,7 +304,8 @@ Arguments Tree_Node { _ }.
Arguments NodeElem_Cons { _ }.
Arguments NodeElem_Nil { _ }.
-(** [no_nested_borrows::list_length]: forward function *)
+(** [no_nested_borrows::list_length]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 296:0-296:48 *)
Fixpoint list_length (T : Type) (l : List_t T) : result u32 :=
match l with
| List_Cons t l1 => i <- list_length T l1; u32_add 1%u32 i
@@ -272,7 +313,8 @@ Fixpoint list_length (T : Type) (l : List_t T) : result u32 :=
end
.
-(** [no_nested_borrows::list_nth_shared]: forward function *)
+(** [no_nested_borrows::list_nth_shared]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 304:0-304:62 *)
Fixpoint list_nth_shared (T : Type) (l : List_t T) (i : u32) : result T :=
match l with
| List_Cons x tl =>
@@ -283,7 +325,8 @@ Fixpoint list_nth_shared (T : Type) (l : List_t T) (i : u32) : result T :=
end
.
-(** [no_nested_borrows::list_nth_mut]: forward function *)
+(** [no_nested_borrows::list_nth_mut]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 320:0-320:67 *)
Fixpoint list_nth_mut (T : Type) (l : List_t T) (i : u32) : result T :=
match l with
| List_Cons x tl =>
@@ -294,7 +337,8 @@ Fixpoint list_nth_mut (T : Type) (l : List_t T) (i : u32) : result T :=
end
.
-(** [no_nested_borrows::list_nth_mut]: backward function 0 *)
+(** [no_nested_borrows::list_nth_mut]: backward function 0
+ Source: 'src/no_nested_borrows.rs', lines 320:0-320:67 *)
Fixpoint list_nth_mut_back
(T : Type) (l : List_t T) (i : u32) (ret : T) : result (List_t T) :=
match l with
@@ -309,7 +353,8 @@ Fixpoint list_nth_mut_back
end
.
-(** [no_nested_borrows::list_rev_aux]: forward function *)
+(** [no_nested_borrows::list_rev_aux]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 336:0-336:63 *)
Fixpoint list_rev_aux
(T : Type) (li : List_t T) (lo : List_t T) : result (List_t T) :=
match li with
@@ -319,13 +364,15 @@ Fixpoint list_rev_aux
.
(** [no_nested_borrows::list_rev]: 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/no_nested_borrows.rs', lines 350:0-350:42 *)
Definition list_rev (T : Type) (l : List_t T) : result (List_t T) :=
let li := core_mem_replace (List_t T) l List_Nil in
list_rev_aux T li List_Nil
.
-(** [no_nested_borrows::test_list_functions]: forward function *)
+(** [no_nested_borrows::test_list_functions]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 355:0-355:28 *)
Definition test_list_functions : result unit :=
let l := List_Nil in
let l0 := List_Cons 2%i32 l in
@@ -362,63 +409,74 @@ Definition test_list_functions : result unit :=
(** Unit test for [no_nested_borrows::test_list_functions] *)
Check (test_list_functions )%return.
-(** [no_nested_borrows::id_mut_pair1]: forward function *)
+(** [no_nested_borrows::id_mut_pair1]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 371:0-371:89 *)
Definition id_mut_pair1 (T1 T2 : Type) (x : T1) (y : T2) : result (T1 * T2) :=
Return (x, y)
.
-(** [no_nested_borrows::id_mut_pair1]: backward function 0 *)
+(** [no_nested_borrows::id_mut_pair1]: backward function 0
+ Source: 'src/no_nested_borrows.rs', lines 371:0-371:89 *)
Definition id_mut_pair1_back
(T1 T2 : Type) (x : T1) (y : T2) (ret : (T1 * T2)) : result (T1 * T2) :=
let (t, t0) := ret in Return (t, t0)
.
-(** [no_nested_borrows::id_mut_pair2]: forward function *)
+(** [no_nested_borrows::id_mut_pair2]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 375:0-375:88 *)
Definition id_mut_pair2 (T1 T2 : Type) (p : (T1 * T2)) : result (T1 * T2) :=
let (t, t0) := p in Return (t, t0)
.
-(** [no_nested_borrows::id_mut_pair2]: backward function 0 *)
+(** [no_nested_borrows::id_mut_pair2]: backward function 0
+ Source: 'src/no_nested_borrows.rs', lines 375:0-375:88 *)
Definition id_mut_pair2_back
(T1 T2 : Type) (p : (T1 * T2)) (ret : (T1 * T2)) : result (T1 * T2) :=
let (t, t0) := ret in Return (t, t0)
.
-(** [no_nested_borrows::id_mut_pair3]: forward function *)
+(** [no_nested_borrows::id_mut_pair3]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 379:0-379:93 *)
Definition id_mut_pair3 (T1 T2 : Type) (x : T1) (y : T2) : result (T1 * T2) :=
Return (x, y)
.
-(** [no_nested_borrows::id_mut_pair3]: backward function 0 *)
+(** [no_nested_borrows::id_mut_pair3]: backward function 0
+ Source: 'src/no_nested_borrows.rs', lines 379:0-379:93 *)
Definition id_mut_pair3_back'a
(T1 T2 : Type) (x : T1) (y : T2) (ret : T1) : result T1 :=
Return ret
.
-(** [no_nested_borrows::id_mut_pair3]: backward function 1 *)
+(** [no_nested_borrows::id_mut_pair3]: backward function 1
+ Source: 'src/no_nested_borrows.rs', lines 379:0-379:93 *)
Definition id_mut_pair3_back'b
(T1 T2 : Type) (x : T1) (y : T2) (ret : T2) : result T2 :=
Return ret
.
-(** [no_nested_borrows::id_mut_pair4]: forward function *)
+(** [no_nested_borrows::id_mut_pair4]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 383:0-383:92 *)
Definition id_mut_pair4 (T1 T2 : Type) (p : (T1 * T2)) : result (T1 * T2) :=
let (t, t0) := p in Return (t, t0)
.
-(** [no_nested_borrows::id_mut_pair4]: backward function 0 *)
+(** [no_nested_borrows::id_mut_pair4]: backward function 0
+ Source: 'src/no_nested_borrows.rs', lines 383:0-383:92 *)
Definition id_mut_pair4_back'a
(T1 T2 : Type) (p : (T1 * T2)) (ret : T1) : result T1 :=
Return ret
.
-(** [no_nested_borrows::id_mut_pair4]: backward function 1 *)
+(** [no_nested_borrows::id_mut_pair4]: backward function 1
+ Source: 'src/no_nested_borrows.rs', lines 383:0-383:92 *)
Definition id_mut_pair4_back'b
(T1 T2 : Type) (p : (T1 * T2)) (ret : T2) : result T2 :=
Return ret
.
-(** [no_nested_borrows::StructWithTuple] *)
+(** [no_nested_borrows::StructWithTuple]
+ Source: 'src/no_nested_borrows.rs', lines 390:0-390:34 *)
Record StructWithTuple_t (T1 T2 : Type) :=
mkStructWithTuple_t {
structWithTuple_p : (T1 * T2);
@@ -428,22 +486,26 @@ mkStructWithTuple_t {
Arguments mkStructWithTuple_t { _ _ }.
Arguments structWithTuple_p { _ _ }.
-(** [no_nested_borrows::new_tuple1]: forward function *)
+(** [no_nested_borrows::new_tuple1]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 394:0-394:48 *)
Definition new_tuple1 : result (StructWithTuple_t u32 u32) :=
Return {| structWithTuple_p := (1%u32, 2%u32) |}
.
-(** [no_nested_borrows::new_tuple2]: forward function *)
+(** [no_nested_borrows::new_tuple2]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 398:0-398:48 *)
Definition new_tuple2 : result (StructWithTuple_t i16 i16) :=
Return {| structWithTuple_p := (1%i16, 2%i16) |}
.
-(** [no_nested_borrows::new_tuple3]: forward function *)
+(** [no_nested_borrows::new_tuple3]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 402:0-402:48 *)
Definition new_tuple3 : result (StructWithTuple_t u64 i64) :=
Return {| structWithTuple_p := (1%u64, 2%i64) |}
.
-(** [no_nested_borrows::StructWithPair] *)
+(** [no_nested_borrows::StructWithPair]
+ Source: 'src/no_nested_borrows.rs', lines 407:0-407:33 *)
Record StructWithPair_t (T1 T2 : Type) :=
mkStructWithPair_t {
structWithPair_p : Pair_t T1 T2;
@@ -453,12 +515,14 @@ mkStructWithPair_t {
Arguments mkStructWithPair_t { _ _ }.
Arguments structWithPair_p { _ _ }.
-(** [no_nested_borrows::new_pair1]: forward function *)
+(** [no_nested_borrows::new_pair1]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 411:0-411:46 *)
Definition new_pair1 : result (StructWithPair_t u32 u32) :=
Return {| structWithPair_p := {| pair_x := 1%u32; pair_y := 2%u32 |} |}
.
-(** [no_nested_borrows::test_constants]: forward function *)
+(** [no_nested_borrows::test_constants]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 419:0-419:23 *)
Definition test_constants : result unit :=
swt <- new_tuple1;
let (i, _) := swt.(structWithTuple_p) in
@@ -484,7 +548,8 @@ Definition test_constants : result unit :=
(** Unit test for [no_nested_borrows::test_constants] *)
Check (test_constants )%return.
-(** [no_nested_borrows::test_weird_borrows1]: forward function *)
+(** [no_nested_borrows::test_weird_borrows1]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 428:0-428:28 *)
Definition test_weird_borrows1 : result unit :=
Return tt.
@@ -492,27 +557,32 @@ Definition test_weird_borrows1 : result unit :=
Check (test_weird_borrows1 )%return.
(** [no_nested_borrows::test_mem_replace]: 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/no_nested_borrows.rs', lines 438:0-438:37 *)
Definition test_mem_replace (px : u32) : result u32 :=
let y := core_mem_replace u32 px 1%u32 in
if negb (y s= 0%u32) then Fail_ Failure else Return 2%u32
.
-(** [no_nested_borrows::test_shared_borrow_bool1]: forward function *)
+(** [no_nested_borrows::test_shared_borrow_bool1]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 445:0-445:47 *)
Definition test_shared_borrow_bool1 (b : bool) : result u32 :=
if b then Return 0%u32 else Return 1%u32
.
-(** [no_nested_borrows::test_shared_borrow_bool2]: forward function *)
+(** [no_nested_borrows::test_shared_borrow_bool2]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 458:0-458:40 *)
Definition test_shared_borrow_bool2 : result u32 :=
Return 0%u32.
-(** [no_nested_borrows::test_shared_borrow_enum1]: forward function *)
+(** [no_nested_borrows::test_shared_borrow_enum1]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 473:0-473:52 *)
Definition test_shared_borrow_enum1 (l : List_t u32) : result u32 :=
match l with | List_Cons i l0 => Return 1%u32 | List_Nil => Return 0%u32 end
.
-(** [no_nested_borrows::test_shared_borrow_enum2]: forward function *)
+(** [no_nested_borrows::test_shared_borrow_enum2]: forward function
+ Source: 'src/no_nested_borrows.rs', lines 485:0-485:40 *)
Definition test_shared_borrow_enum2 : result u32 :=
Return 0%u32.
diff --git a/tests/coq/misc/Paper.v b/tests/coq/misc/Paper.v
index d3852e6b..4a49096f 100644
--- a/tests/coq/misc/Paper.v
+++ b/tests/coq/misc/Paper.v
@@ -9,11 +9,13 @@ Local Open Scope Primitives_scope.
Module 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 *)
Definition ref_incr (x : i32) : result i32 :=
i32_add x 1%i32.
-(** [paper::test_incr]: forward function *)
+(** [paper::test_incr]: forward function
+ Source: 'src/paper.rs', lines 8:0-8:18 *)
Definition test_incr : result unit :=
x <- ref_incr 0%i32; if negb (x s= 1%i32) then Fail_ Failure else Return tt
.
@@ -21,18 +23,21 @@ Definition test_incr : result unit :=
(** Unit test for [paper::test_incr] *)
Check (test_incr )%return.
-(** [paper::choose]: forward function *)
+(** [paper::choose]: forward function
+ Source: 'src/paper.rs', lines 15:0-15:70 *)
Definition choose (T : Type) (b : bool) (x : T) (y : T) : result T :=
if b then Return x else Return y
.
-(** [paper::choose]: backward function 0 *)
+(** [paper::choose]: backward function 0
+ Source: 'src/paper.rs', lines 15:0-15:70 *)
Definition choose_back
(T : Type) (b : bool) (x : T) (y : T) (ret : T) : result (T * T) :=
if b then Return (ret, y) else Return (x, ret)
.
-(** [paper::test_choose]: forward function *)
+(** [paper::test_choose]: forward function
+ Source: 'src/paper.rs', lines 23:0-23:20 *)
Definition test_choose : result unit :=
z <- choose i32 true 0%i32 0%i32;
z0 <- i32_add z 1%i32;
@@ -49,7 +54,8 @@ Definition test_choose : result unit :=
(** Unit test for [paper::test_choose] *)
Check (test_choose )%return.
-(** [paper::List] *)
+(** [paper::List]
+ Source: 'src/paper.rs', lines 35:0-35:16 *)
Inductive List_t (T : Type) :=
| List_Cons : T -> List_t T -> List_t T
| List_Nil : List_t T
@@ -58,7 +64,8 @@ Inductive List_t (T : Type) :=
Arguments List_Cons { _ }.
Arguments List_Nil { _ }.
-(** [paper::list_nth_mut]: forward function *)
+(** [paper::list_nth_mut]: forward function
+ Source: 'src/paper.rs', lines 42:0-42:67 *)
Fixpoint list_nth_mut (T : Type) (l : List_t T) (i : u32) : result T :=
match l with
| List_Cons x tl =>
@@ -69,7 +76,8 @@ Fixpoint list_nth_mut (T : Type) (l : List_t T) (i : u32) : result T :=
end
.
-(** [paper::list_nth_mut]: backward function 0 *)
+(** [paper::list_nth_mut]: backward function 0
+ Source: 'src/paper.rs', lines 42:0-42:67 *)
Fixpoint list_nth_mut_back
(T : Type) (l : List_t T) (i : u32) (ret : T) : result (List_t T) :=
match l with
@@ -84,7 +92,8 @@ Fixpoint list_nth_mut_back
end
.
-(** [paper::sum]: forward function *)
+(** [paper::sum]: forward function
+ Source: 'src/paper.rs', lines 57:0-57:32 *)
Fixpoint sum (l : List_t i32) : result i32 :=
match l with
| List_Cons x tl => i <- sum tl; i32_add x i
@@ -92,7 +101,8 @@ Fixpoint sum (l : List_t i32) : result i32 :=
end
.
-(** [paper::test_nth]: forward function *)
+(** [paper::test_nth]: forward function
+ Source: 'src/paper.rs', lines 68:0-68:17 *)
Definition test_nth : result unit :=
let l := List_Nil in
let l0 := List_Cons 3%i32 l in
@@ -107,7 +117,8 @@ Definition test_nth : result unit :=
(** Unit test for [paper::test_nth] *)
Check (test_nth )%return.
-(** [paper::call_choose]: forward function *)
+(** [paper::call_choose]: forward function
+ Source: 'src/paper.rs', lines 76:0-76:44 *)
Definition call_choose (p : (u32 * u32)) : result u32 :=
let (px, py) := p in
pz <- choose u32 true px py;
diff --git a/tests/coq/misc/PoloniusList.v b/tests/coq/misc/PoloniusList.v
index 4848444f..a0820e40 100644
--- a/tests/coq/misc/PoloniusList.v
+++ b/tests/coq/misc/PoloniusList.v
@@ -8,7 +8,8 @@ Import ListNotations.
Local Open Scope Primitives_scope.
Module PoloniusList.
-(** [polonius_list::List] *)
+(** [polonius_list::List]
+ Source: 'src/polonius_list.rs', lines 3:0-3:16 *)
Inductive List_t (T : Type) :=
| List_Cons : T -> List_t T -> List_t T
| List_Nil : List_t T
@@ -17,7 +18,8 @@ Inductive List_t (T : Type) :=
Arguments List_Cons { _ }.
Arguments List_Nil { _ }.
-(** [polonius_list::get_list_at_x]: forward function *)
+(** [polonius_list::get_list_at_x]: forward function
+ Source: 'src/polonius_list.rs', lines 13:0-13:76 *)
Fixpoint get_list_at_x (ls : List_t u32) (x : u32) : result (List_t u32) :=
match ls with
| List_Cons hd tl =>
@@ -26,7 +28,8 @@ Fixpoint get_list_at_x (ls : List_t u32) (x : u32) : result (List_t u32) :=
end
.
-(** [polonius_list::get_list_at_x]: backward function 0 *)
+(** [polonius_list::get_list_at_x]: backward function 0
+ Source: 'src/polonius_list.rs', lines 13:0-13:76 *)
Fixpoint get_list_at_x_back
(ls : List_t u32) (x : u32) (ret : List_t u32) : result (List_t u32) :=
match ls with
diff --git a/tests/coq/misc/Primitives.v b/tests/coq/misc/Primitives.v
index 85e38f01..83f860b6 100644
--- a/tests/coq/misc/Primitives.v
+++ b/tests/coq/misc/Primitives.v
@@ -467,14 +467,14 @@ Definition alloc_boxed_Box_deref_mut (T : Type) (x : T) : result T := Return x.
Definition alloc_boxed_Box_deref_mut_back (T : Type) (_ : T) (x : T) : result T := Return x.
(* Trait instance *)
-Definition alloc_boxed_Box_coreOpsDerefInst (Self : Type) : core_ops_deref_Deref Self := {|
+Definition alloc_boxed_Box_coreopsDerefInst (Self : Type) : core_ops_deref_Deref Self := {|
core_ops_deref_Deref_target := Self;
core_ops_deref_Deref_deref := alloc_boxed_Box_deref Self;
|}.
(* Trait instance *)
-Definition alloc_boxed_Box_coreOpsDerefMutInst (Self : Type) : core_ops_deref_DerefMut Self := {|
- core_ops_deref_DerefMut_derefInst := alloc_boxed_Box_coreOpsDerefInst Self;
+Definition alloc_boxed_Box_coreopsDerefMutInst (Self : Type) : core_ops_deref_DerefMut Self := {|
+ core_ops_deref_DerefMut_derefInst := alloc_boxed_Box_coreopsDerefInst Self;
core_ops_deref_DerefMut_deref_mut := alloc_boxed_Box_deref_mut Self;
core_ops_deref_DerefMut_deref_mut_back := alloc_boxed_Box_deref_mut_back Self;
|}.
@@ -576,7 +576,7 @@ Definition alloc_vec_Vec_insert (T: Type) (v: alloc_vec_Vec T) (i: usize) (x: T)
else Fail_ Failure).
(* Helper *)
-Axiom alloc_vec_Vec_index_usize : forall {T : Type} (v : alloc_vec_Vec T) (i : usize) (x : T), result T.
+Axiom alloc_vec_Vec_index_usize : forall {T : Type} (v : alloc_vec_Vec T) (i : usize), result T.
(* Helper *)
Axiom alloc_vec_Vec_update_usize : forall {T : Type} (v : alloc_vec_Vec T) (i : usize) (x : T), result (alloc_vec_Vec T).
@@ -620,18 +620,18 @@ Definition core_slice_index_Slice_index
end.
(* [core::slice::index::Range:::get]: forward function *)
-Axiom core_slice_index_Range_get : forall (T : Type) (i : core_ops_range_Range usize) (s : slice T), result (option (slice T)).
+Axiom core_slice_index_RangeUsize_get : forall (T : Type) (i : core_ops_range_Range usize) (s : slice T), result (option (slice T)).
(* [core::slice::index::Range::get_mut]: forward function *)
-Axiom core_slice_index_Range_get_mut :
+Axiom core_slice_index_RangeUsize_get_mut :
forall (T : Type), core_ops_range_Range usize -> slice T -> result (option (slice T)).
(* [core::slice::index::Range::get_mut]: backward function 0 *)
-Axiom core_slice_index_Range_get_mut_back :
+Axiom core_slice_index_RangeUsize_get_mut_back :
forall (T : Type), core_ops_range_Range usize -> slice T -> option (slice T) -> result (slice T).
(* [core::slice::index::Range::get_unchecked]: forward function *)
-Definition core_slice_index_Range_get_unchecked
+Definition core_slice_index_RangeUsize_get_unchecked
(T : Type) :
core_ops_range_Range usize -> const_raw_ptr (slice T) -> result (const_raw_ptr (slice T)) :=
(* Don't know what the model should be - for now we always fail to make
@@ -639,7 +639,7 @@ Definition core_slice_index_Range_get_unchecked
fun _ _ => Fail_ Failure.
(* [core::slice::index::Range::get_unchecked_mut]: forward function *)
-Definition core_slice_index_Range_get_unchecked_mut
+Definition core_slice_index_RangeUsize_get_unchecked_mut
(T : Type) :
core_ops_range_Range usize -> mut_raw_ptr (slice T) -> result (mut_raw_ptr (slice T)) :=
(* Don't know what the model should be - for now we always fail to make
@@ -647,15 +647,15 @@ Definition core_slice_index_Range_get_unchecked_mut
fun _ _ => Fail_ Failure.
(* [core::slice::index::Range::index]: forward function *)
-Axiom core_slice_index_Range_index :
+Axiom core_slice_index_RangeUsize_index :
forall (T : Type), core_ops_range_Range usize -> slice T -> result (slice T).
(* [core::slice::index::Range::index_mut]: forward function *)
-Axiom core_slice_index_Range_index_mut :
+Axiom core_slice_index_RangeUsize_index_mut :
forall (T : Type), core_ops_range_Range usize -> slice T -> result (slice T).
(* [core::slice::index::Range::index_mut]: backward function 0 *)
-Axiom core_slice_index_Range_index_mut_back :
+Axiom core_slice_index_RangeUsize_index_mut_back :
forall (T : Type), core_ops_range_Range usize -> slice T -> slice T -> result (slice T).
(* [core::slice::index::[T]::index_mut]: forward function *)
@@ -683,44 +683,44 @@ Axiom core_array_Array_index_mut_back :
forall (T Idx : Type) (N : usize) (inst : core_ops_index_IndexMut (slice T) Idx)
(a : array T N) (i : Idx) (x : inst.(core_ops_index_IndexMut_indexInst).(core_ops_index_Index_Output)), result (array T N).
-(* Trait implementation: [core::slice::index::[T]] *)
-Definition core_slice_index_Slice_coreopsindexIndexInst (T Idx : Type)
- (inst : core_slice_index_SliceIndex Idx (slice T)) :
- core_ops_index_Index (slice T) Idx := {|
- core_ops_index_Index_Output := inst.(core_slice_index_SliceIndex_Output);
- core_ops_index_Index_index := core_slice_index_Slice_index T Idx inst;
-|}.
-
(* Trait implementation: [core::slice::index::private_slice_index::Range] *)
-Definition core_slice_index_private_slice_index_Range_coresliceindexprivate_slice_indexSealedInst
+Definition core_slice_index_private_slice_index_SealedRangeUsizeInst
: core_slice_index_private_slice_index_Sealed (core_ops_range_Range usize) := tt.
(* Trait implementation: [core::slice::index::Range] *)
-Definition core_slice_index_Range_coresliceindexSliceIndexInst (T : Type) :
+Definition core_slice_index_SliceIndexRangeUsizeSliceTInst (T : Type) :
core_slice_index_SliceIndex (core_ops_range_Range usize) (slice T) := {|
- core_slice_index_SliceIndex_sealedInst := core_slice_index_private_slice_index_Range_coresliceindexprivate_slice_indexSealedInst;
+ core_slice_index_SliceIndex_sealedInst := core_slice_index_private_slice_index_SealedRangeUsizeInst;
core_slice_index_SliceIndex_Output := slice T;
- core_slice_index_SliceIndex_get := core_slice_index_Range_get T;
- core_slice_index_SliceIndex_get_mut := core_slice_index_Range_get_mut T;
- core_slice_index_SliceIndex_get_mut_back := core_slice_index_Range_get_mut_back T;
- core_slice_index_SliceIndex_get_unchecked := core_slice_index_Range_get_unchecked T;
- core_slice_index_SliceIndex_get_unchecked_mut := core_slice_index_Range_get_unchecked_mut T;
- core_slice_index_SliceIndex_index := core_slice_index_Range_index T;
- core_slice_index_SliceIndex_index_mut := core_slice_index_Range_index_mut T;
- core_slice_index_SliceIndex_index_mut_back := core_slice_index_Range_index_mut_back T;
+ core_slice_index_SliceIndex_get := core_slice_index_RangeUsize_get T;
+ core_slice_index_SliceIndex_get_mut := core_slice_index_RangeUsize_get_mut T;
+ core_slice_index_SliceIndex_get_mut_back := core_slice_index_RangeUsize_get_mut_back T;
+ core_slice_index_SliceIndex_get_unchecked := core_slice_index_RangeUsize_get_unchecked T;
+ core_slice_index_SliceIndex_get_unchecked_mut := core_slice_index_RangeUsize_get_unchecked_mut T;
+ core_slice_index_SliceIndex_index := core_slice_index_RangeUsize_index T;
+ core_slice_index_SliceIndex_index_mut := core_slice_index_RangeUsize_index_mut T;
+ core_slice_index_SliceIndex_index_mut_back := core_slice_index_RangeUsize_index_mut_back T;
+|}.
+
+(* Trait implementation: [core::slice::index::[T]] *)
+Definition core_ops_index_IndexSliceTIInst (T Idx : Type)
+ (inst : core_slice_index_SliceIndex Idx (slice T)) :
+ core_ops_index_Index (slice T) Idx := {|
+ core_ops_index_Index_Output := inst.(core_slice_index_SliceIndex_Output);
+ core_ops_index_Index_index := core_slice_index_Slice_index T Idx inst;
|}.
(* Trait implementation: [core::slice::index::[T]] *)
-Definition core_slice_index_Slice_coreopsindexIndexMutInst (T Idx : Type)
+Definition core_ops_index_IndexMutSliceTIInst (T Idx : Type)
(inst : core_slice_index_SliceIndex Idx (slice T)) :
core_ops_index_IndexMut (slice T) Idx := {|
- core_ops_index_IndexMut_indexInst := core_slice_index_Slice_coreopsindexIndexInst T Idx inst;
+ core_ops_index_IndexMut_indexInst := core_ops_index_IndexSliceTIInst T Idx inst;
core_ops_index_IndexMut_index_mut := core_slice_index_Slice_index_mut T Idx inst;
core_ops_index_IndexMut_index_mut_back := core_slice_index_Slice_index_mut_back T Idx inst;
|}.
(* Trait implementation: [core::array::[T; N]] *)
-Definition core_array_Array_coreopsindexIndexInst (T Idx : Type) (N : usize)
+Definition core_ops_index_IndexArrayInst (T Idx : Type) (N : usize)
(inst : core_ops_index_Index (slice T) Idx) :
core_ops_index_Index (array T N) Idx := {|
core_ops_index_Index_Output := inst.(core_ops_index_Index_Output);
@@ -728,10 +728,10 @@ Definition core_array_Array_coreopsindexIndexInst (T Idx : Type) (N : usize)
|}.
(* Trait implementation: [core::array::[T; N]] *)
-Definition core_array_Array_coreopsindexIndexMutInst (T Idx : Type) (N : usize)
+Definition core_ops_index_IndexMutArrayInst (T Idx : Type) (N : usize)
(inst : core_ops_index_IndexMut (slice T) Idx) :
core_ops_index_IndexMut (array T N) Idx := {|
- core_ops_index_IndexMut_indexInst := core_array_Array_coreopsindexIndexInst T Idx N inst.(core_ops_index_IndexMut_indexInst);
+ core_ops_index_IndexMut_indexInst := core_ops_index_IndexArrayInst T Idx N inst.(core_ops_index_IndexMut_indexInst);
core_ops_index_IndexMut_index_mut := core_array_Array_index_mut T Idx N inst;
core_ops_index_IndexMut_index_mut_back := core_array_Array_index_mut_back T Idx N inst;
|}.
@@ -765,13 +765,13 @@ Axiom core_slice_index_usize_index_mut_back :
forall (T : Type), usize -> slice T -> T -> result (slice T).
(* Trait implementation: [core::slice::index::private_slice_index::usize] *)
-Definition core_slice_index_private_slice_index_usize_coresliceindexprivate_slice_indexSealedInst
+Definition core_slice_index_private_slice_index_SealedUsizeInst
: core_slice_index_private_slice_index_Sealed usize := tt.
(* Trait implementation: [core::slice::index::usize] *)
-Definition core_slice_index_usize_coresliceindexSliceIndexInst (T : Type) :
+Definition core_slice_index_SliceIndexUsizeSliceTInst (T : Type) :
core_slice_index_SliceIndex usize (slice T) := {|
- core_slice_index_SliceIndex_sealedInst := core_slice_index_private_slice_index_usize_coresliceindexprivate_slice_indexSealedInst;
+ core_slice_index_SliceIndex_sealedInst := core_slice_index_private_slice_index_SealedUsizeInst;
core_slice_index_SliceIndex_Output := T;
core_slice_index_SliceIndex_get := core_slice_index_usize_get T;
core_slice_index_SliceIndex_get_mut := core_slice_index_usize_get_mut T;
@@ -815,8 +815,16 @@ Definition alloc_vec_Vec_coreopsindexIndexMutInst (T Idx : Type)
(*** Theorems *)
+Axiom alloc_vec_Vec_index_eq : forall {a : Type} (v : alloc_vec_Vec a) (i : usize) (x : a),
+ alloc_vec_Vec_index a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i =
+ alloc_vec_Vec_index_usize v i.
+
+Axiom alloc_vec_Vec_index_mut_eq : forall {a : Type} (v : alloc_vec_Vec a) (i : usize) (x : a),
+ alloc_vec_Vec_index_mut a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i =
+ alloc_vec_Vec_index_usize v i.
+
Axiom alloc_vec_Vec_index_mut_back_eq : forall {a : Type} (v : alloc_vec_Vec a) (i : usize) (x : a),
- alloc_vec_Vec_index_mut_back a usize (core_slice_index_usize_coresliceindexSliceIndexInst a) v i x =
+ alloc_vec_Vec_index_mut_back a usize (core_slice_index_SliceIndexUsizeSliceTInst a) v i x =
alloc_vec_Vec_update_usize v i x.
End Primitives.