summaryrefslogtreecommitdiff
path: root/tests/coq/misc
diff options
context:
space:
mode:
authorSon Ho2023-10-27 13:34:03 +0200
committerSon Ho2023-10-27 13:34:03 +0200
commit49117ba254679f98938223711810191c3f7d788f (patch)
tree2a29b2be5becbb90104c552b77b4f9d3b3fa2546 /tests/coq/misc
parent1c55079dac0c51e45f7d77c67bfdbdb8c52ed192 (diff)
Regenerate the Coq test files
Diffstat (limited to 'tests/coq/misc')
-rw-r--r--tests/coq/misc/Constants.v56
-rw-r--r--tests/coq/misc/External_Funs.v39
-rw-r--r--tests/coq/misc/External_Opaque.v8
-rw-r--r--tests/coq/misc/External_Types.v2
-rw-r--r--tests/coq/misc/Loops.v412
-rw-r--r--tests/coq/misc/NoNestedBorrows.v305
-rw-r--r--tests/coq/misc/Paper.v67
-rw-r--r--tests/coq/misc/PoloniusList.v22
-rw-r--r--tests/coq/misc/Primitives.v405
9 files changed, 802 insertions, 514 deletions
diff --git a/tests/coq/misc/Constants.v b/tests/coq/misc/Constants.v
index 5dd78a09..7cb1a642 100644
--- a/tests/coq/misc/Constants.v
+++ b/tests/coq/misc/Constants.v
@@ -21,36 +21,35 @@ Definition x2_body : result u32 := Return 3%u32.
Definition x2_c : u32 := x2_body%global.
(** [constants::incr]: forward function *)
-Definition incr_fwd (n : u32) : result u32 :=
+Definition incr (n : u32) : result u32 :=
u32_add n 1%u32.
(** [constants::X3] *)
-Definition x3_body : result u32 := incr_fwd 32%u32.
+Definition x3_body : result u32 := incr 32%u32.
Definition x3_c : u32 := x3_body%global.
(** [constants::mk_pair0]: forward function *)
-Definition mk_pair0_fwd (x : u32) (y : u32) : result (u32 * u32) :=
- Return (x, y)
-.
+Definition mk_pair0 (x : u32) (y : u32) : result (u32 * u32) :=
+ Return (x, y).
(** [constants::Pair] *)
-Record Pair_t (T1 T2 : Type) := mkPair_t { Pair_x : T1; Pair_y : T2; }.
+Record Pair_t (T1 T2 : Type) := mkPair_t { pair_x : T1; pair_y : T2; }.
Arguments mkPair_t {T1} {T2} _ _.
-Arguments Pair_x {T1} {T2}.
-Arguments Pair_y {T1} {T2}.
+Arguments pair_x {T1} {T2}.
+Arguments pair_y {T1} {T2}.
(** [constants::mk_pair1]: forward function *)
-Definition mk_pair1_fwd (x : u32) (y : u32) : result (Pair_t u32 u32) :=
- Return {| Pair_x := x; Pair_y := y |}
+Definition mk_pair1 (x : u32) (y : u32) : result (Pair_t u32 u32) :=
+ Return {| pair_x := x; pair_y := y |}
.
(** [constants::P0] *)
-Definition p0_body : result (u32 * u32) := mk_pair0_fwd 0%u32 1%u32.
+Definition p0_body : result (u32 * u32) := mk_pair0 0%u32 1%u32.
Definition p0_c : (u32 * u32) := p0_body%global.
(** [constants::P1] *)
-Definition p1_body : result (Pair_t u32 u32) := mk_pair1_fwd 0%u32 1%u32.
+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] *)
@@ -59,31 +58,31 @@ Definition p2_c : (u32 * u32) := p2_body%global.
(** [constants::P3] *)
Definition p3_body : result (Pair_t u32 u32) :=
- Return {| Pair_x := 0%u32; Pair_y := 1%u32 |}
+ Return {| pair_x := 0%u32; pair_y := 1%u32 |}
.
Definition p3_c : Pair_t u32 u32 := p3_body%global.
(** [constants::Wrap] *)
-Record Wrap_t (T : Type) := mkWrap_t { Wrap_val : T; }.
+Record Wrap_t (T : Type) := mkWrap_t { wrap_value : T; }.
Arguments mkWrap_t {T} _.
-Arguments Wrap_val {T}.
+Arguments wrap_value {T}.
(** [constants::Wrap::{0}::new]: forward function *)
-Definition wrap_new_fwd (T : Type) (val : T) : result (Wrap_t T) :=
- Return {| Wrap_val := val |}
+Definition wrap_new (T : Type) (value : T) : result (Wrap_t T) :=
+ Return {| wrap_value := value |}
.
(** [constants::Y] *)
-Definition y_body : result (Wrap_t i32) := wrap_new_fwd i32 2%i32.
+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 *)
-Definition unwrap_y_fwd : result i32 :=
- Return y_c.(Wrap_val).
+Definition unwrap_y : result i32 :=
+ Return y_c.(wrap_value).
(** [constants::YVAL] *)
-Definition yval_body : result i32 := unwrap_y_fwd.
+Definition yval_body : result i32 := unwrap_y.
Definition yval_c : i32 := yval_body%global.
(** [constants::get_z1::Z1] *)
@@ -91,11 +90,11 @@ 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 *)
-Definition get_z1_fwd : result i32 :=
+Definition get_z1 : result i32 :=
Return get_z1_z1_c.
(** [constants::add]: forward function *)
-Definition add_fwd (a : i32) (b : i32) : result i32 :=
+Definition add (a : i32) (b : i32) : result i32 :=
i32_add a b.
(** [constants::Q1] *)
@@ -107,20 +106,19 @@ Definition q2_body : result i32 := Return q1_c.
Definition q2_c : i32 := q2_body%global.
(** [constants::Q3] *)
-Definition q3_body : result i32 := add_fwd q2_c 3%i32.
+Definition q3_body : result i32 := add q2_c 3%i32.
Definition q3_c : i32 := q3_body%global.
(** [constants::get_z2]: forward function *)
-Definition get_z2_fwd : result i32 :=
- i <- get_z1_fwd; i0 <- add_fwd i q3_c; add_fwd q1_c i0
-.
+Definition get_z2 : result i32 :=
+ i <- get_z1; i0 <- add i q3_c; add q1_c i0.
(** [constants::S1] *)
Definition s1_body : result u32 := Return 6%u32.
Definition s1_c : u32 := s1_body%global.
(** [constants::S2] *)
-Definition s2_body : result u32 := incr_fwd s1_c.
+Definition s2_body : result u32 := incr s1_c.
Definition s2_c : u32 := s2_body%global.
(** [constants::S3] *)
@@ -128,7 +126,7 @@ Definition s3_body : result (Pair_t u32 u32) := Return p3_c.
Definition s3_c : Pair_t u32 u32 := s3_body%global.
(** [constants::S4] *)
-Definition s4_body : result (Pair_t u32 u32) := mk_pair1_fwd 7%u32 8%u32.
+Definition s4_body : result (Pair_t u32 u32) := mk_pair1 7%u32 8%u32.
Definition s4_c : Pair_t u32 u32 := s4_body%global.
End Constants .
diff --git a/tests/coq/misc/External_Funs.v b/tests/coq/misc/External_Funs.v
index 28370b2b..018ce13c 100644
--- a/tests/coq/misc/External_Funs.v
+++ b/tests/coq/misc/External_Funs.v
@@ -13,9 +13,9 @@ Import External_Opaque.
Module External_Funs.
(** [external::swap]: forward function *)
-Definition swap_fwd
+Definition swap
(T : Type) (x : T) (y : T) (st : state) : result (state * unit) :=
- p <- core_mem_swap_fwd T x y st;
+ p <- core_mem_swap T x y st;
let (st0, _) := p in
p0 <- core_mem_swap_back0 T x y st st0;
let (st1, _) := p0 in
@@ -29,7 +29,7 @@ Definition swap_back
(T : Type) (x : T) (y : T) (st : state) (st0 : state) :
result (state * (T * T))
:=
- p <- core_mem_swap_fwd T x y st;
+ p <- core_mem_swap T x y st;
let (st1, _) := p in
p0 <- core_mem_swap_back0 T x y st st1;
let (st2, x0) := p0 in
@@ -39,25 +39,27 @@ Definition swap_back
.
(** [external::test_new_non_zero_u32]: forward function *)
-Definition test_new_non_zero_u32_fwd
- (x : u32) (st : state) : result (state * Core_num_nonzero_non_zero_u32_t) :=
- p <- core_num_nonzero_non_zero_u32_new_fwd x st;
- let (st0, opt) := p in
- core_option_option_unwrap_fwd Core_num_nonzero_non_zero_u32_t opt st0
+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;
+ let (st0, o) := p in
+ core_option_Option_unwrap core_num_nonzero_NonZeroU32_t o st0
.
(** [external::test_vec]: forward function *)
-Definition test_vec_fwd : result unit :=
- let v := vec_new u32 in _ <- vec_push_back u32 v 0%u32; Return tt
+Definition test_vec : result unit :=
+ let v := alloc_vec_Vec_new u32 in
+ _ <- alloc_vec_Vec_push u32 v 0%u32;
+ Return tt
.
(** Unit test for [external::test_vec] *)
-Check (test_vec_fwd )%return.
+Check (test_vec )%return.
(** [external::custom_swap]: forward function *)
-Definition custom_swap_fwd
+Definition custom_swap
(T : Type) (x : T) (y : T) (st : state) : result (state * T) :=
- p <- core_mem_swap_fwd T x y st;
+ p <- core_mem_swap T x y st;
let (st0, _) := p in
p0 <- core_mem_swap_back0 T x y st st0;
let (st1, x0) := p0 in
@@ -71,7 +73,7 @@ Definition custom_swap_back
(T : Type) (x : T) (y : T) (st : state) (ret : T) (st0 : state) :
result (state * (T * T))
:=
- p <- core_mem_swap_fwd T x y st;
+ p <- core_mem_swap T x y st;
let (st1, _) := p in
p0 <- core_mem_swap_back0 T x y st st1;
let (st2, _) := p0 in
@@ -81,9 +83,9 @@ Definition custom_swap_back
.
(** [external::test_custom_swap]: forward function *)
-Definition test_custom_swap_fwd
+Definition test_custom_swap
(x : u32) (y : u32) (st : state) : result (state * unit) :=
- p <- custom_swap_fwd u32 x y st; let (st0, _) := p in Return (st0, tt)
+ p <- custom_swap u32 x y st; let (st0, _) := p in Return (st0, tt)
.
(** [external::test_custom_swap]: backward function 0 *)
@@ -95,9 +97,8 @@ Definition test_custom_swap_back
.
(** [external::test_swap_non_zero]: forward function *)
-Definition test_swap_non_zero_fwd
- (x : u32) (st : state) : result (state * u32) :=
- p <- swap_fwd u32 x 0%u32 st;
+Definition test_swap_non_zero (x : u32) (st : state) : result (state * u32) :=
+ p <- swap u32 x 0%u32 st;
let (st0, _) := p in
p0 <- swap_back u32 x 0%u32 st st0;
let (st1, p1) := p0 in
diff --git a/tests/coq/misc/External_Opaque.v b/tests/coq/misc/External_Opaque.v
index d2ee42d4..80be37e7 100644
--- a/tests/coq/misc/External_Opaque.v
+++ b/tests/coq/misc/External_Opaque.v
@@ -11,7 +11,7 @@ Import External_Types.
Module External_Opaque.
(** [core::mem::swap]: forward function *)
-Axiom core_mem_swap_fwd :
+Axiom core_mem_swap :
forall(T : Type), T -> T -> state -> result (state * unit)
.
@@ -26,12 +26,12 @@ Axiom core_mem_swap_back1 :
.
(** [core::num::nonzero::NonZeroU32::{14}::new]: forward function *)
-Axiom core_num_nonzero_non_zero_u32_new_fwd
- : u32 -> state -> result (state * (option Core_num_nonzero_non_zero_u32_t))
+Axiom core_num_nonzero_NonZeroU32_new
+ : u32 -> state -> result (state * (option core_num_nonzero_NonZeroU32_t))
.
(** [core::option::Option::{0}::unwrap]: forward function *)
-Axiom core_option_option_unwrap_fwd :
+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 1883fa6c..9e49ca41 100644
--- a/tests/coq/misc/External_Types.v
+++ b/tests/coq/misc/External_Types.v
@@ -9,7 +9,7 @@ Local Open Scope Primitives_scope.
Module External_Types.
(** [core::num::nonzero::NonZeroU32] *)
-Axiom Core_num_nonzero_non_zero_u32_t : Type.
+Axiom core_num_nonzero_NonZeroU32_t : Type.
(** The state type used in the state-error monad *)
Axiom state : Type.
diff --git a/tests/coq/misc/Loops.v b/tests/coq/misc/Loops.v
index 82e57576..180a1d68 100644
--- a/tests/coq/misc/Loops.v
+++ b/tests/coq/misc/Loops.v
@@ -9,23 +9,23 @@ Local Open Scope Primitives_scope.
Module Loops.
(** [loops::sum]: loop 0: forward function *)
-Fixpoint sum_loop_fwd (n : nat) (max : u32) (i : u32) (s : u32) : result u32 :=
+Fixpoint sum_loop (n : nat) (max : u32) (i : u32) (s : u32) : result u32 :=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
if i s< max
- then (s0 <- u32_add s i; i0 <- u32_add i 1%u32; sum_loop_fwd n0 max i0 s0)
+ then (s0 <- u32_add s i; i0 <- u32_add i 1%u32; sum_loop n0 max i0 s0)
else u32_mul s 2%u32
end
.
(** [loops::sum]: forward function *)
-Definition sum_fwd (n : nat) (max : u32) : result u32 :=
- sum_loop_fwd n max 0%u32 0%u32
+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 *)
-Fixpoint sum_with_mut_borrows_loop_fwd
+Fixpoint sum_with_mut_borrows_loop
(n : nat) (max : u32) (mi : u32) (ms : u32) : result u32 :=
match n with
| O => Fail_ OutOfFuel
@@ -34,18 +34,18 @@ Fixpoint sum_with_mut_borrows_loop_fwd
then (
ms0 <- u32_add ms mi;
mi0 <- u32_add mi 1%u32;
- sum_with_mut_borrows_loop_fwd n0 max mi0 ms0)
+ sum_with_mut_borrows_loop n0 max mi0 ms0)
else u32_mul ms 2%u32
end
.
(** [loops::sum_with_mut_borrows]: forward function *)
-Definition sum_with_mut_borrows_fwd (n : nat) (max : u32) : result u32 :=
- sum_with_mut_borrows_loop_fwd n max 0%u32 0%u32
+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 *)
-Fixpoint sum_with_shared_borrows_loop_fwd
+Fixpoint sum_with_shared_borrows_loop
(n : nat) (max : u32) (i : u32) (s : u32) : result u32 :=
match n with
| O => Fail_ OutOfFuel
@@ -54,87 +54,88 @@ Fixpoint sum_with_shared_borrows_loop_fwd
then (
i0 <- u32_add i 1%u32;
s0 <- u32_add s i0;
- sum_with_shared_borrows_loop_fwd n0 max i0 s0)
+ sum_with_shared_borrows_loop n0 max i0 s0)
else u32_mul s 2%u32
end
.
(** [loops::sum_with_shared_borrows]: forward function *)
-Definition sum_with_shared_borrows_fwd (n : nat) (max : u32) : result u32 :=
- sum_with_shared_borrows_loop_fwd n max 0%u32 0%u32
+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 ()) *)
-Fixpoint clear_loop_fwd_back
- (n : nat) (v : vec u32) (i : usize) : result (vec u32) :=
+Fixpoint clear_loop
+ (n : nat) (v : alloc_vec_Vec u32) (i : usize) : result (alloc_vec_Vec u32) :=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
- let i0 := vec_len u32 v in
+ let i0 := alloc_vec_Vec_len u32 v in
if i s< i0
then (
i1 <- usize_add i 1%usize;
- v0 <- vec_index_mut_back u32 v i 0%u32;
- clear_loop_fwd_back n0 v0 i1)
+ v0 <-
+ alloc_vec_Vec_index_mut_back u32 usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst 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 ()) *)
-Definition clear_fwd_back (n : nat) (v : vec u32) : result (vec u32) :=
- clear_loop_fwd_back n v 0%usize
+Definition clear
+ (n : nat) (v : alloc_vec_Vec u32) : result (alloc_vec_Vec u32) :=
+ clear_loop n v 0%usize
.
(** [loops::List] *)
Inductive List_t (T : Type) :=
-| ListCons : T -> List_t T -> List_t T
-| ListNil : List_t T
+| List_Cons : T -> List_t T -> List_t T
+| List_Nil : List_t T
.
-Arguments ListCons {T} _ _.
-Arguments ListNil {T}.
+Arguments List_Cons {T} _ _.
+Arguments List_Nil {T}.
(** [loops::list_mem]: loop 0: forward function *)
-Fixpoint list_mem_loop_fwd
- (n : nat) (x : u32) (ls : List_t u32) : result bool :=
+Fixpoint list_mem_loop (n : nat) (x : u32) (ls : List_t u32) : result bool :=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | ListCons y tl =>
- if y s= x then Return true else list_mem_loop_fwd n0 x tl
- | ListNil => Return false
+ | List_Cons y tl => if y s= x then Return true else list_mem_loop n0 x tl
+ | List_Nil => Return false
end
end
.
(** [loops::list_mem]: forward function *)
-Definition list_mem_fwd (n : nat) (x : u32) (ls : List_t u32) : result bool :=
- list_mem_loop_fwd n x ls
+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 *)
-Fixpoint list_nth_mut_loop_loop_fwd
+Fixpoint list_nth_mut_loop_loop
(T : Type) (n : nat) (ls : List_t T) (i : u32) : result T :=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | ListCons x tl =>
+ | List_Cons x tl =>
if i s= 0%u32
then Return x
- else (i0 <- u32_sub i 1%u32; list_nth_mut_loop_loop_fwd T n0 tl i0)
- | ListNil => Fail_ Failure
+ else (i0 <- u32_sub i 1%u32; list_nth_mut_loop_loop T n0 tl i0)
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::list_nth_mut_loop]: forward function *)
-Definition list_nth_mut_loop_fwd
+Definition list_nth_mut_loop
(T : Type) (n : nat) (ls : List_t T) (i : u32) : result T :=
- list_nth_mut_loop_loop_fwd T n ls i
+ list_nth_mut_loop_loop T n ls i
.
(** [loops::list_nth_mut_loop]: loop 0: backward function 0 *)
@@ -146,14 +147,14 @@ Fixpoint list_nth_mut_loop_loop_back
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | ListCons x tl =>
+ | List_Cons x tl =>
if i s= 0%u32
- then Return (ListCons ret tl)
+ then Return (List_Cons ret tl)
else (
i0 <- u32_sub i 1%u32;
tl0 <- list_nth_mut_loop_loop_back T n0 tl i0 ret;
- Return (ListCons x tl0))
- | ListNil => Fail_ Failure
+ Return (List_Cons x tl0))
+ | List_Nil => Fail_ Failure
end
end
.
@@ -167,46 +168,50 @@ Definition list_nth_mut_loop_back
.
(** [loops::list_nth_shared_loop]: loop 0: forward function *)
-Fixpoint list_nth_shared_loop_loop_fwd
+Fixpoint list_nth_shared_loop_loop
(T : Type) (n : nat) (ls : List_t T) (i : u32) : result T :=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | ListCons x tl =>
+ | List_Cons x tl =>
if i s= 0%u32
then Return x
- else (i0 <- u32_sub i 1%u32; list_nth_shared_loop_loop_fwd T n0 tl i0)
- | ListNil => Fail_ Failure
+ else (i0 <- u32_sub i 1%u32; list_nth_shared_loop_loop T n0 tl i0)
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::list_nth_shared_loop]: forward function *)
-Definition list_nth_shared_loop_fwd
+Definition list_nth_shared_loop
(T : Type) (n : nat) (ls : List_t T) (i : u32) : result T :=
- list_nth_shared_loop_loop_fwd T n ls i
+ list_nth_shared_loop_loop T n ls i
.
(** [loops::get_elem_mut]: loop 0: forward function *)
-Fixpoint get_elem_mut_loop_fwd
+Fixpoint get_elem_mut_loop
(n : nat) (x : usize) (ls : List_t usize) : result usize :=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | ListCons y tl =>
- if y s= x then Return y else get_elem_mut_loop_fwd n0 x tl
- | ListNil => Fail_ Failure
+ | List_Cons y tl => if y s= x then Return y else get_elem_mut_loop n0 x tl
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::get_elem_mut]: forward function *)
-Definition get_elem_mut_fwd
- (n : nat) (slots : vec (List_t usize)) (x : usize) : result usize :=
- l <- vec_index_mut_fwd (List_t usize) slots 0%usize;
- get_elem_mut_loop_fwd n x l
+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;
+ get_elem_mut_loop n x l
.
(** [loops::get_elem_mut]: loop 0: backward function 0 *)
@@ -218,50 +223,60 @@ Fixpoint get_elem_mut_loop_back
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | ListCons y tl =>
+ | List_Cons y tl =>
if y s= x
- then Return (ListCons ret tl)
- else (tl0 <- get_elem_mut_loop_back n0 x tl ret; Return (ListCons y tl0))
- | ListNil => Fail_ Failure
+ then Return (List_Cons ret tl)
+ else (
+ tl0 <- get_elem_mut_loop_back n0 x tl ret; Return (List_Cons y tl0))
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::get_elem_mut]: backward function 0 *)
Definition get_elem_mut_back
- (n : nat) (slots : vec (List_t usize)) (x : usize) (ret : usize) :
- result (vec (List_t usize))
+ (n : nat) (slots : alloc_vec_Vec (List_t usize)) (x : usize) (ret : usize) :
+ result (alloc_vec_Vec (List_t usize))
:=
- l <- vec_index_mut_fwd (List_t usize) slots 0%usize;
+ l <-
+ alloc_vec_Vec_index_mut (List_t usize) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (List_t usize))
+ slots 0%usize;
l0 <- get_elem_mut_loop_back n x l ret;
- vec_index_mut_back (List_t usize) slots 0%usize l0
+ alloc_vec_Vec_index_mut_back (List_t usize) usize
+ (core_slice_index_usize_coresliceindexSliceIndexInst (List_t usize)) slots
+ 0%usize l0
.
(** [loops::get_elem_shared]: loop 0: forward function *)
-Fixpoint get_elem_shared_loop_fwd
+Fixpoint get_elem_shared_loop
(n : nat) (x : usize) (ls : List_t usize) : result usize :=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | ListCons y tl =>
- if y s= x then Return y else get_elem_shared_loop_fwd n0 x tl
- | ListNil => Fail_ Failure
+ | List_Cons y tl =>
+ if y s= x then Return y else get_elem_shared_loop n0 x tl
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::get_elem_shared]: forward function *)
-Definition get_elem_shared_fwd
- (n : nat) (slots : vec (List_t usize)) (x : usize) : result usize :=
- l <- vec_index_fwd (List_t usize) slots 0%usize;
- get_elem_shared_loop_fwd n x l
+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;
+ get_elem_shared_loop n x l
.
(** [loops::id_mut]: forward function *)
-Definition id_mut_fwd (T : Type) (ls : List_t T) : result (List_t T) :=
- Return ls
-.
+Definition id_mut (T : Type) (ls : List_t T) : result (List_t T) :=
+ Return ls.
(** [loops::id_mut]: backward function 0 *)
Definition id_mut_back
@@ -270,31 +285,30 @@ Definition id_mut_back
.
(** [loops::id_shared]: forward function *)
-Definition id_shared_fwd (T : Type) (ls : List_t T) : result (List_t T) :=
+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 *)
-Fixpoint list_nth_mut_loop_with_id_loop_fwd
+Fixpoint list_nth_mut_loop_with_id_loop
(T : Type) (n : nat) (i : u32) (ls : List_t T) : result T :=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | ListCons x tl =>
+ | List_Cons x tl =>
if i s= 0%u32
then Return x
- else (
- i0 <- u32_sub i 1%u32; list_nth_mut_loop_with_id_loop_fwd T n0 i0 tl)
- | ListNil => Fail_ Failure
+ else (i0 <- u32_sub i 1%u32; list_nth_mut_loop_with_id_loop T n0 i0 tl)
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::list_nth_mut_loop_with_id]: forward function *)
-Definition list_nth_mut_loop_with_id_fwd
+Definition list_nth_mut_loop_with_id
(T : Type) (n : nat) (ls : List_t T) (i : u32) : result T :=
- ls0 <- id_mut_fwd T ls; list_nth_mut_loop_with_id_loop_fwd T n i ls0
+ 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 *)
@@ -306,14 +320,14 @@ Fixpoint list_nth_mut_loop_with_id_loop_back
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | ListCons x tl =>
+ | List_Cons x tl =>
if i s= 0%u32
- then Return (ListCons ret tl)
+ then Return (List_Cons ret tl)
else (
i0 <- u32_sub i 1%u32;
tl0 <- list_nth_mut_loop_with_id_loop_back T n0 i0 tl ret;
- Return (ListCons x tl0))
- | ListNil => Fail_ Failure
+ Return (List_Cons x tl0))
+ | List_Nil => Fail_ Failure
end
end
.
@@ -323,36 +337,36 @@ Definition list_nth_mut_loop_with_id_back
(T : Type) (n : nat) (ls : List_t T) (i : u32) (ret : T) :
result (List_t T)
:=
- ls0 <- id_mut_fwd T ls;
+ ls0 <- id_mut T ls;
l <- list_nth_mut_loop_with_id_loop_back T n i ls0 ret;
id_mut_back T ls l
.
(** [loops::list_nth_shared_loop_with_id]: loop 0: forward function *)
-Fixpoint list_nth_shared_loop_with_id_loop_fwd
+Fixpoint list_nth_shared_loop_with_id_loop
(T : Type) (n : nat) (i : u32) (ls : List_t T) : result T :=
match n with
| O => Fail_ OutOfFuel
| S n0 =>
match ls with
- | ListCons x tl =>
+ | List_Cons x tl =>
if i s= 0%u32
then Return x
else (
- i0 <- u32_sub i 1%u32; list_nth_shared_loop_with_id_loop_fwd T n0 i0 tl)
- | ListNil => Fail_ Failure
+ i0 <- u32_sub i 1%u32; list_nth_shared_loop_with_id_loop T n0 i0 tl)
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::list_nth_shared_loop_with_id]: forward function *)
-Definition list_nth_shared_loop_with_id_fwd
+Definition list_nth_shared_loop_with_id
(T : Type) (n : nat) (ls : List_t T) (i : u32) : result T :=
- ls0 <- id_shared_fwd T ls; list_nth_shared_loop_with_id_loop_fwd T n i ls0
+ 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 *)
-Fixpoint list_nth_mut_loop_pair_loop_fwd
+Fixpoint list_nth_mut_loop_pair_loop
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
:=
@@ -360,27 +374,26 @@ Fixpoint list_nth_mut_loop_pair_loop_fwd
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
then Return (x0, x1)
else (
- i0 <- u32_sub i 1%u32;
- list_nth_mut_loop_pair_loop_fwd T n0 tl0 tl1 i0)
- | ListNil => Fail_ Failure
+ i0 <- u32_sub i 1%u32; list_nth_mut_loop_pair_loop T n0 tl0 tl1 i0)
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::list_nth_mut_loop_pair]: forward function *)
-Definition list_nth_mut_loop_pair_fwd
+Definition list_nth_mut_loop_pair
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
:=
- list_nth_mut_loop_pair_loop_fwd T n ls0 ls1 i
+ list_nth_mut_loop_pair_loop T n ls0 ls1 i
.
(** [loops::list_nth_mut_loop_pair]: loop 0: backward function 0 *)
@@ -392,18 +405,18 @@ Fixpoint list_nth_mut_loop_pair_loop_back'a
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
- then Return (ListCons ret tl0)
+ then Return (List_Cons ret tl0)
else (
i0 <- u32_sub i 1%u32;
tl00 <- list_nth_mut_loop_pair_loop_back'a T n0 tl0 tl1 i0 ret;
- Return (ListCons x0 tl00))
- | ListNil => Fail_ Failure
+ Return (List_Cons x0 tl00))
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
@@ -425,18 +438,18 @@ Fixpoint list_nth_mut_loop_pair_loop_back'b
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
- then Return (ListCons ret tl1)
+ then Return (List_Cons ret tl1)
else (
i0 <- u32_sub i 1%u32;
tl10 <- list_nth_mut_loop_pair_loop_back'b T n0 tl0 tl1 i0 ret;
- Return (ListCons x1 tl10))
- | ListNil => Fail_ Failure
+ Return (List_Cons x1 tl10))
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
@@ -450,7 +463,7 @@ Definition list_nth_mut_loop_pair_back'b
.
(** [loops::list_nth_shared_loop_pair]: loop 0: forward function *)
-Fixpoint list_nth_shared_loop_pair_loop_fwd
+Fixpoint list_nth_shared_loop_pair_loop
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
:=
@@ -458,31 +471,30 @@ Fixpoint list_nth_shared_loop_pair_loop_fwd
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
then Return (x0, x1)
else (
- i0 <- u32_sub i 1%u32;
- list_nth_shared_loop_pair_loop_fwd T n0 tl0 tl1 i0)
- | ListNil => Fail_ Failure
+ i0 <- u32_sub i 1%u32; list_nth_shared_loop_pair_loop T n0 tl0 tl1 i0)
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::list_nth_shared_loop_pair]: forward function *)
-Definition list_nth_shared_loop_pair_fwd
+Definition list_nth_shared_loop_pair
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
:=
- list_nth_shared_loop_pair_loop_fwd T n ls0 ls1 i
+ list_nth_shared_loop_pair_loop T n ls0 ls1 i
.
(** [loops::list_nth_mut_loop_pair_merge]: loop 0: forward function *)
-Fixpoint list_nth_mut_loop_pair_merge_loop_fwd
+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)
:=
@@ -490,27 +502,27 @@ Fixpoint list_nth_mut_loop_pair_merge_loop_fwd
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
then Return (x0, x1)
else (
i0 <- u32_sub i 1%u32;
- list_nth_mut_loop_pair_merge_loop_fwd T n0 tl0 tl1 i0)
- | ListNil => Fail_ Failure
+ list_nth_mut_loop_pair_merge_loop T n0 tl0 tl1 i0)
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::list_nth_mut_loop_pair_merge]: forward function *)
-Definition list_nth_mut_loop_pair_merge_fwd
+Definition list_nth_mut_loop_pair_merge
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
:=
- list_nth_mut_loop_pair_merge_loop_fwd T n ls0 ls1 i
+ list_nth_mut_loop_pair_merge_loop T n ls0 ls1 i
.
(** [loops::list_nth_mut_loop_pair_merge]: loop 0: backward function 0 *)
@@ -523,19 +535,19 @@ Fixpoint list_nth_mut_loop_pair_merge_loop_back
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
- then let (t, t0) := ret in Return (ListCons t tl0, ListCons t0 tl1)
+ then let (t, t0) := ret in Return (List_Cons t tl0, List_Cons t0 tl1)
else (
i0 <- u32_sub i 1%u32;
p <- list_nth_mut_loop_pair_merge_loop_back T n0 tl0 tl1 i0 ret;
let (tl00, tl10) := p in
- Return (ListCons x0 tl00, ListCons x1 tl10))
- | ListNil => Fail_ Failure
+ Return (List_Cons x0 tl00, List_Cons x1 tl10))
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
@@ -550,7 +562,7 @@ Definition list_nth_mut_loop_pair_merge_back
.
(** [loops::list_nth_shared_loop_pair_merge]: loop 0: forward function *)
-Fixpoint list_nth_shared_loop_pair_merge_loop_fwd
+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)
:=
@@ -558,31 +570,31 @@ Fixpoint list_nth_shared_loop_pair_merge_loop_fwd
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
then Return (x0, x1)
else (
i0 <- u32_sub i 1%u32;
- list_nth_shared_loop_pair_merge_loop_fwd T n0 tl0 tl1 i0)
- | ListNil => Fail_ Failure
+ list_nth_shared_loop_pair_merge_loop T n0 tl0 tl1 i0)
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::list_nth_shared_loop_pair_merge]: forward function *)
-Definition list_nth_shared_loop_pair_merge_fwd
+Definition list_nth_shared_loop_pair_merge
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
:=
- list_nth_shared_loop_pair_merge_loop_fwd T n ls0 ls1 i
+ list_nth_shared_loop_pair_merge_loop T n ls0 ls1 i
.
(** [loops::list_nth_mut_shared_loop_pair]: loop 0: forward function *)
-Fixpoint list_nth_mut_shared_loop_pair_loop_fwd
+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)
:=
@@ -590,27 +602,27 @@ Fixpoint list_nth_mut_shared_loop_pair_loop_fwd
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
then Return (x0, x1)
else (
i0 <- u32_sub i 1%u32;
- list_nth_mut_shared_loop_pair_loop_fwd T n0 tl0 tl1 i0)
- | ListNil => Fail_ Failure
+ list_nth_mut_shared_loop_pair_loop T n0 tl0 tl1 i0)
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::list_nth_mut_shared_loop_pair]: forward function *)
-Definition list_nth_mut_shared_loop_pair_fwd
+Definition list_nth_mut_shared_loop_pair
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
:=
- list_nth_mut_shared_loop_pair_loop_fwd T n ls0 ls1 i
+ list_nth_mut_shared_loop_pair_loop T n ls0 ls1 i
.
(** [loops::list_nth_mut_shared_loop_pair]: loop 0: backward function 0 *)
@@ -622,18 +634,18 @@ Fixpoint list_nth_mut_shared_loop_pair_loop_back
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
- then Return (ListCons ret tl0)
+ then Return (List_Cons ret tl0)
else (
i0 <- u32_sub i 1%u32;
tl00 <- list_nth_mut_shared_loop_pair_loop_back T n0 tl0 tl1 i0 ret;
- Return (ListCons x0 tl00))
- | ListNil => Fail_ Failure
+ Return (List_Cons x0 tl00))
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
@@ -647,7 +659,7 @@ Definition list_nth_mut_shared_loop_pair_back
.
(** [loops::list_nth_mut_shared_loop_pair_merge]: loop 0: forward function *)
-Fixpoint list_nth_mut_shared_loop_pair_merge_loop_fwd
+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)
:=
@@ -655,27 +667,27 @@ Fixpoint list_nth_mut_shared_loop_pair_merge_loop_fwd
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
then Return (x0, x1)
else (
i0 <- u32_sub i 1%u32;
- list_nth_mut_shared_loop_pair_merge_loop_fwd T n0 tl0 tl1 i0)
- | ListNil => Fail_ Failure
+ list_nth_mut_shared_loop_pair_merge_loop T n0 tl0 tl1 i0)
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::list_nth_mut_shared_loop_pair_merge]: forward function *)
-Definition list_nth_mut_shared_loop_pair_merge_fwd
+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)
:=
- list_nth_mut_shared_loop_pair_merge_loop_fwd T n ls0 ls1 i
+ 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 *)
@@ -687,19 +699,19 @@ Fixpoint list_nth_mut_shared_loop_pair_merge_loop_back
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
- then Return (ListCons ret tl0)
+ then Return (List_Cons ret tl0)
else (
i0 <- u32_sub i 1%u32;
tl00 <-
list_nth_mut_shared_loop_pair_merge_loop_back T n0 tl0 tl1 i0 ret;
- Return (ListCons x0 tl00))
- | ListNil => Fail_ Failure
+ Return (List_Cons x0 tl00))
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
@@ -713,7 +725,7 @@ Definition list_nth_mut_shared_loop_pair_merge_back
.
(** [loops::list_nth_shared_mut_loop_pair]: loop 0: forward function *)
-Fixpoint list_nth_shared_mut_loop_pair_loop_fwd
+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)
:=
@@ -721,27 +733,27 @@ Fixpoint list_nth_shared_mut_loop_pair_loop_fwd
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
then Return (x0, x1)
else (
i0 <- u32_sub i 1%u32;
- list_nth_shared_mut_loop_pair_loop_fwd T n0 tl0 tl1 i0)
- | ListNil => Fail_ Failure
+ list_nth_shared_mut_loop_pair_loop T n0 tl0 tl1 i0)
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::list_nth_shared_mut_loop_pair]: forward function *)
-Definition list_nth_shared_mut_loop_pair_fwd
+Definition list_nth_shared_mut_loop_pair
(T : Type) (n : nat) (ls0 : List_t T) (ls1 : List_t T) (i : u32) :
result (T * T)
:=
- list_nth_shared_mut_loop_pair_loop_fwd T n ls0 ls1 i
+ list_nth_shared_mut_loop_pair_loop T n ls0 ls1 i
.
(** [loops::list_nth_shared_mut_loop_pair]: loop 0: backward function 1 *)
@@ -753,18 +765,18 @@ Fixpoint list_nth_shared_mut_loop_pair_loop_back
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
- then Return (ListCons ret tl1)
+ then Return (List_Cons ret tl1)
else (
i0 <- u32_sub i 1%u32;
tl10 <- list_nth_shared_mut_loop_pair_loop_back T n0 tl0 tl1 i0 ret;
- Return (ListCons x1 tl10))
- | ListNil => Fail_ Failure
+ Return (List_Cons x1 tl10))
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
@@ -778,7 +790,7 @@ Definition list_nth_shared_mut_loop_pair_back
.
(** [loops::list_nth_shared_mut_loop_pair_merge]: loop 0: forward function *)
-Fixpoint list_nth_shared_mut_loop_pair_merge_loop_fwd
+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)
:=
@@ -786,27 +798,27 @@ Fixpoint list_nth_shared_mut_loop_pair_merge_loop_fwd
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
then Return (x0, x1)
else (
i0 <- u32_sub i 1%u32;
- list_nth_shared_mut_loop_pair_merge_loop_fwd T n0 tl0 tl1 i0)
- | ListNil => Fail_ Failure
+ list_nth_shared_mut_loop_pair_merge_loop T n0 tl0 tl1 i0)
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
(** [loops::list_nth_shared_mut_loop_pair_merge]: forward function *)
-Definition list_nth_shared_mut_loop_pair_merge_fwd
+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)
:=
- list_nth_shared_mut_loop_pair_merge_loop_fwd T n ls0 ls1 i
+ 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 *)
@@ -818,19 +830,19 @@ Fixpoint list_nth_shared_mut_loop_pair_merge_loop_back
| O => Fail_ OutOfFuel
| S n0 =>
match ls0 with
- | ListCons x0 tl0 =>
+ | List_Cons x0 tl0 =>
match ls1 with
- | ListCons x1 tl1 =>
+ | List_Cons x1 tl1 =>
if i s= 0%u32
- then Return (ListCons ret tl1)
+ then Return (List_Cons ret tl1)
else (
i0 <- u32_sub i 1%u32;
tl10 <-
list_nth_shared_mut_loop_pair_merge_loop_back T n0 tl0 tl1 i0 ret;
- Return (ListCons x1 tl10))
- | ListNil => Fail_ Failure
+ Return (List_Cons x1 tl10))
+ | List_Nil => Fail_ Failure
end
- | ListNil => Fail_ Failure
+ | List_Nil => Fail_ Failure
end
end
.
diff --git a/tests/coq/misc/NoNestedBorrows.v b/tests/coq/misc/NoNestedBorrows.v
index c1c24e00..47cdc3af 100644
--- a/tests/coq/misc/NoNestedBorrows.v
+++ b/tests/coq/misc/NoNestedBorrows.v
@@ -9,113 +9,113 @@ Local Open Scope Primitives_scope.
Module NoNestedBorrows.
(** [no_nested_borrows::Pair] *)
-Record Pair_t (T1 T2 : Type) := mkPair_t { Pair_x : T1; Pair_y : T2; }.
+Record Pair_t (T1 T2 : Type) := mkPair_t { pair_x : T1; pair_y : T2; }.
Arguments mkPair_t {T1} {T2} _ _.
-Arguments Pair_x {T1} {T2}.
-Arguments Pair_y {T1} {T2}.
+Arguments pair_x {T1} {T2}.
+Arguments pair_y {T1} {T2}.
(** [no_nested_borrows::List] *)
Inductive List_t (T : Type) :=
-| ListCons : T -> List_t T -> List_t T
-| ListNil : List_t T
+| List_Cons : T -> List_t T -> List_t T
+| List_Nil : List_t T
.
-Arguments ListCons {T} _ _.
-Arguments ListNil {T}.
+Arguments List_Cons {T} _ _.
+Arguments List_Nil {T}.
(** [no_nested_borrows::One] *)
-Inductive One_t (T1 : Type) := | OneOne : T1 -> One_t T1.
+Inductive One_t (T1 : Type) := | One_One : T1 -> One_t T1.
-Arguments OneOne {T1} _.
+Arguments One_One {T1} _.
(** [no_nested_borrows::EmptyEnum] *)
-Inductive Empty_enum_t := | EmptyEnumEmpty : Empty_enum_t.
+Inductive EmptyEnum_t := | EmptyEnum_Empty : EmptyEnum_t.
(** [no_nested_borrows::Enum] *)
-Inductive Enum_t := | EnumVariant1 : Enum_t | EnumVariant2 : Enum_t.
+Inductive Enum_t := | Enum_Variant1 : Enum_t | Enum_Variant2 : Enum_t.
(** [no_nested_borrows::EmptyStruct] *)
-Record Empty_struct_t := mkEmpty_struct_t { }.
+Record EmptyStruct_t := mkEmptyStruct_t { }.
(** [no_nested_borrows::Sum] *)
Inductive Sum_t (T1 T2 : Type) :=
-| SumLeft : T1 -> Sum_t T1 T2
-| SumRight : T2 -> Sum_t T1 T2
+| Sum_Left : T1 -> Sum_t T1 T2
+| Sum_Right : T2 -> Sum_t T1 T2
.
-Arguments SumLeft {T1} {T2} _.
-Arguments SumRight {T1} {T2} _.
+Arguments Sum_Left {T1} {T2} _.
+Arguments Sum_Right {T1} {T2} _.
(** [no_nested_borrows::neg_test]: forward function *)
-Definition neg_test_fwd (x : i32) : result i32 :=
+Definition neg_test (x : i32) : result i32 :=
i32_neg x.
(** [no_nested_borrows::add_test]: forward function *)
-Definition add_test_fwd (x : u32) (y : u32) : result u32 :=
+Definition add_test (x : u32) (y : u32) : result u32 :=
u32_add x y.
(** [no_nested_borrows::subs_test]: forward function *)
-Definition subs_test_fwd (x : u32) (y : u32) : result u32 :=
+Definition subs_test (x : u32) (y : u32) : result u32 :=
u32_sub x y.
(** [no_nested_borrows::div_test]: forward function *)
-Definition div_test_fwd (x : u32) (y : u32) : result u32 :=
+Definition div_test (x : u32) (y : u32) : result u32 :=
u32_div x y.
(** [no_nested_borrows::div_test1]: forward function *)
-Definition div_test1_fwd (x : u32) : result u32 :=
+Definition div_test1 (x : u32) : result u32 :=
u32_div x 2%u32.
(** [no_nested_borrows::rem_test]: forward function *)
-Definition rem_test_fwd (x : u32) (y : u32) : result u32 :=
+Definition rem_test (x : u32) (y : u32) : result u32 :=
u32_rem x y.
(** [no_nested_borrows::cast_test]: forward function *)
-Definition cast_test_fwd (x : u32) : result i32 :=
+Definition cast_test (x : u32) : result i32 :=
scalar_cast U32 I32 x.
(** [no_nested_borrows::test2]: forward function *)
-Definition test2_fwd : result unit :=
+Definition test2 : result unit :=
_ <- u32_add 23%u32 44%u32; Return tt.
(** Unit test for [no_nested_borrows::test2] *)
-Check (test2_fwd )%return.
+Check (test2 )%return.
(** [no_nested_borrows::get_max]: forward function *)
-Definition get_max_fwd (x : u32) (y : u32) : result u32 :=
+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 *)
-Definition test3_fwd : result unit :=
- x <- get_max_fwd 4%u32 3%u32;
- y <- get_max_fwd 10%u32 11%u32;
+Definition test3 : result unit :=
+ x <- get_max 4%u32 3%u32;
+ y <- get_max 10%u32 11%u32;
z <- u32_add x y;
if negb (z s= 15%u32) then Fail_ Failure else Return tt
.
(** Unit test for [no_nested_borrows::test3] *)
-Check (test3_fwd )%return.
+Check (test3 )%return.
(** [no_nested_borrows::test_neg1]: forward function *)
-Definition test_neg1_fwd : result unit :=
+Definition test_neg1 : result unit :=
y <- i32_neg 3%i32; if negb (y s= (-3)%i32) then Fail_ Failure else Return tt
.
(** Unit test for [no_nested_borrows::test_neg1] *)
-Check (test_neg1_fwd )%return.
+Check (test_neg1 )%return.
(** [no_nested_borrows::refs_test1]: forward function *)
-Definition refs_test1_fwd : result unit :=
+Definition refs_test1 : result unit :=
if negb (1%i32 s= 1%i32) then Fail_ Failure else Return tt
.
(** Unit test for [no_nested_borrows::refs_test1] *)
-Check (refs_test1_fwd )%return.
+Check (refs_test1 )%return.
(** [no_nested_borrows::refs_test2]: forward function *)
-Definition refs_test2_fwd : result unit :=
+Definition refs_test2 : result unit :=
if negb (2%i32 s= 2%i32)
then Fail_ Failure
else
@@ -128,85 +128,83 @@ Definition refs_test2_fwd : result unit :=
.
(** Unit test for [no_nested_borrows::refs_test2] *)
-Check (refs_test2_fwd )%return.
+Check (refs_test2 )%return.
(** [no_nested_borrows::test_list1]: forward function *)
-Definition test_list1_fwd : result unit :=
+Definition test_list1 : result unit :=
Return tt.
(** Unit test for [no_nested_borrows::test_list1] *)
-Check (test_list1_fwd )%return.
+Check (test_list1 )%return.
(** [no_nested_borrows::test_box1]: forward function *)
-Definition test_box1_fwd : result unit :=
+Definition test_box1 : result unit :=
let b := 1%i32 in
let x := b in
if negb (x s= 1%i32) then Fail_ Failure else Return tt
.
(** Unit test for [no_nested_borrows::test_box1] *)
-Check (test_box1_fwd )%return.
+Check (test_box1 )%return.
(** [no_nested_borrows::copy_int]: forward function *)
-Definition copy_int_fwd (x : i32) : result i32 :=
+Definition copy_int (x : i32) : result i32 :=
Return x.
(** [no_nested_borrows::test_unreachable]: forward function *)
-Definition test_unreachable_fwd (b : bool) : result unit :=
+Definition test_unreachable (b : bool) : result unit :=
if b then Fail_ Failure else Return tt
.
(** [no_nested_borrows::test_panic]: forward function *)
-Definition test_panic_fwd (b : bool) : result unit :=
+Definition test_panic (b : bool) : result unit :=
if b then Fail_ Failure else Return tt
.
(** [no_nested_borrows::test_copy_int]: forward function *)
-Definition test_copy_int_fwd : result unit :=
- y <- copy_int_fwd 0%i32;
- if negb (0%i32 s= y) then Fail_ Failure else Return tt
+Definition test_copy_int : result unit :=
+ y <- copy_int 0%i32; if negb (0%i32 s= y) then Fail_ Failure else Return tt
.
(** Unit test for [no_nested_borrows::test_copy_int] *)
-Check (test_copy_int_fwd )%return.
+Check (test_copy_int )%return.
(** [no_nested_borrows::is_cons]: forward function *)
-Definition is_cons_fwd (T : Type) (l : List_t T) : result bool :=
- match l with | ListCons t l0 => Return true | ListNil => Return false end
+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 *)
-Definition test_is_cons_fwd : result unit :=
- let l := ListNil in
- b <- is_cons_fwd i32 (ListCons 0%i32 l);
+Definition test_is_cons : result unit :=
+ let l := List_Nil in
+ b <- is_cons i32 (List_Cons 0%i32 l);
if negb b then Fail_ Failure else Return tt
.
(** Unit test for [no_nested_borrows::test_is_cons] *)
-Check (test_is_cons_fwd )%return.
+Check (test_is_cons )%return.
(** [no_nested_borrows::split_list]: forward function *)
-Definition split_list_fwd
- (T : Type) (l : List_t T) : result (T * (List_t T)) :=
+Definition split_list (T : Type) (l : List_t T) : result (T * (List_t T)) :=
match l with
- | ListCons hd tl => Return (hd, tl)
- | ListNil => Fail_ Failure
+ | List_Cons hd tl => Return (hd, tl)
+ | List_Nil => Fail_ Failure
end
.
(** [no_nested_borrows::test_split_list]: forward function *)
-Definition test_split_list_fwd : result unit :=
- let l := ListNil in
- p <- split_list_fwd i32 (ListCons 0%i32 l);
+Definition test_split_list : result unit :=
+ let l := List_Nil in
+ p <- split_list i32 (List_Cons 0%i32 l);
let (hd, _) := p in
if negb (hd s= 0%i32) then Fail_ Failure else Return tt
.
(** Unit test for [no_nested_borrows::test_split_list] *)
-Check (test_split_list_fwd )%return.
+Check (test_split_list )%return.
(** [no_nested_borrows::choose]: forward function *)
-Definition choose_fwd (T : Type) (b : bool) (x : T) (y : T) : result T :=
+Definition choose (T : Type) (b : bool) (x : T) (y : T) : result T :=
if b then Return x else Return y
.
@@ -217,8 +215,8 @@ Definition choose_back
.
(** [no_nested_borrows::choose_test]: forward function *)
-Definition choose_test_fwd : result unit :=
- z <- choose_fwd i32 true 0%i32 0%i32;
+Definition choose_test : result unit :=
+ z <- choose i32 true 0%i32 0%i32;
z0 <- i32_add z 1%i32;
if negb (z0 s= 1%i32)
then Fail_ Failure
@@ -231,57 +229,56 @@ Definition choose_test_fwd : result unit :=
.
(** Unit test for [no_nested_borrows::choose_test] *)
-Check (choose_test_fwd )%return.
+Check (choose_test )%return.
(** [no_nested_borrows::test_char]: forward function *)
-Definition test_char_fwd : result char :=
- Return (char_of_byte Coq.Init.Byte.x61)
-.
+Definition test_char : result char :=
+ Return (char_of_byte Coq.Init.Byte.x61).
(** [no_nested_borrows::Tree] *)
Inductive Tree_t (T : Type) :=
-| TreeLeaf : T -> Tree_t T
-| TreeNode : T -> Node_elem_t T -> Tree_t T -> Tree_t T
+| Tree_Leaf : T -> Tree_t T
+| Tree_Node : T -> NodeElem_t T -> Tree_t T -> Tree_t T
(** [no_nested_borrows::NodeElem] *)
-with Node_elem_t (T : Type) :=
-| NodeElemCons : Tree_t T -> Node_elem_t T -> Node_elem_t T
-| NodeElemNil : Node_elem_t T
+with NodeElem_t (T : Type) :=
+| NodeElem_Cons : Tree_t T -> NodeElem_t T -> NodeElem_t T
+| NodeElem_Nil : NodeElem_t T
.
-Arguments TreeLeaf {T} _.
-Arguments TreeNode {T} _ _ _.
+Arguments Tree_Leaf {T} _.
+Arguments Tree_Node {T} _ _ _.
-Arguments NodeElemCons {T} _ _.
-Arguments NodeElemNil {T}.
+Arguments NodeElem_Cons {T} _ _.
+Arguments NodeElem_Nil {T}.
(** [no_nested_borrows::list_length]: forward function *)
-Fixpoint list_length_fwd (T : Type) (l : List_t T) : result u32 :=
+Fixpoint list_length (T : Type) (l : List_t T) : result u32 :=
match l with
- | ListCons t l1 => i <- list_length_fwd T l1; u32_add 1%u32 i
- | ListNil => Return 0%u32
+ | List_Cons t l1 => i <- list_length T l1; u32_add 1%u32 i
+ | List_Nil => Return 0%u32
end
.
(** [no_nested_borrows::list_nth_shared]: forward function *)
-Fixpoint list_nth_shared_fwd (T : Type) (l : List_t T) (i : u32) : result T :=
+Fixpoint list_nth_shared (T : Type) (l : List_t T) (i : u32) : result T :=
match l with
- | ListCons x tl =>
+ | List_Cons x tl =>
if i s= 0%u32
then Return x
- else (i0 <- u32_sub i 1%u32; list_nth_shared_fwd T tl i0)
- | ListNil => Fail_ Failure
+ else (i0 <- u32_sub i 1%u32; list_nth_shared T tl i0)
+ | List_Nil => Fail_ Failure
end
.
(** [no_nested_borrows::list_nth_mut]: forward function *)
-Fixpoint list_nth_mut_fwd (T : Type) (l : List_t T) (i : u32) : result T :=
+Fixpoint list_nth_mut (T : Type) (l : List_t T) (i : u32) : result T :=
match l with
- | ListCons x tl =>
+ | List_Cons x tl =>
if i s= 0%u32
then Return x
- else (i0 <- u32_sub i 1%u32; list_nth_mut_fwd T tl i0)
- | ListNil => Fail_ Failure
+ else (i0 <- u32_sub i 1%u32; list_nth_mut T tl i0)
+ | List_Nil => Fail_ Failure
end
.
@@ -289,73 +286,72 @@ Fixpoint list_nth_mut_fwd (T : Type) (l : List_t T) (i : u32) : result T :=
Fixpoint list_nth_mut_back
(T : Type) (l : List_t T) (i : u32) (ret : T) : result (List_t T) :=
match l with
- | ListCons x tl =>
+ | List_Cons x tl =>
if i s= 0%u32
- then Return (ListCons ret tl)
+ then Return (List_Cons ret tl)
else (
i0 <- u32_sub i 1%u32;
tl0 <- list_nth_mut_back T tl i0 ret;
- Return (ListCons x tl0))
- | ListNil => Fail_ Failure
+ Return (List_Cons x tl0))
+ | List_Nil => Fail_ Failure
end
.
(** [no_nested_borrows::list_rev_aux]: forward function *)
-Fixpoint list_rev_aux_fwd
+Fixpoint list_rev_aux
(T : Type) (li : List_t T) (lo : List_t T) : result (List_t T) :=
match li with
- | ListCons hd tl => list_rev_aux_fwd T tl (ListCons hd lo)
- | ListNil => Return lo
+ | List_Cons hd tl => list_rev_aux T tl (List_Cons hd lo)
+ | List_Nil => Return lo
end
.
(** [no_nested_borrows::list_rev]: merged forward/backward function
(there is a single backward function, and the forward function returns ()) *)
-Definition list_rev_fwd_back (T : Type) (l : List_t T) : result (List_t T) :=
- let li := mem_replace_fwd (List_t T) l ListNil in
- list_rev_aux_fwd T li ListNil
+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 *)
-Definition test_list_functions_fwd : result unit :=
- let l := ListNil in
- let l0 := ListCons 2%i32 l in
- let l1 := ListCons 1%i32 l0 in
- i <- list_length_fwd i32 (ListCons 0%i32 l1);
+Definition test_list_functions : result unit :=
+ let l := List_Nil in
+ let l0 := List_Cons 2%i32 l in
+ let l1 := List_Cons 1%i32 l0 in
+ i <- list_length i32 (List_Cons 0%i32 l1);
if negb (i s= 3%u32)
then Fail_ Failure
else (
- i0 <- list_nth_shared_fwd i32 (ListCons 0%i32 l1) 0%u32;
+ i0 <- list_nth_shared i32 (List_Cons 0%i32 l1) 0%u32;
if negb (i0 s= 0%i32)
then Fail_ Failure
else (
- i1 <- list_nth_shared_fwd i32 (ListCons 0%i32 l1) 1%u32;
+ i1 <- list_nth_shared i32 (List_Cons 0%i32 l1) 1%u32;
if negb (i1 s= 1%i32)
then Fail_ Failure
else (
- i2 <- list_nth_shared_fwd i32 (ListCons 0%i32 l1) 2%u32;
+ i2 <- list_nth_shared i32 (List_Cons 0%i32 l1) 2%u32;
if negb (i2 s= 2%i32)
then Fail_ Failure
else (
- ls <- list_nth_mut_back i32 (ListCons 0%i32 l1) 1%u32 3%i32;
- i3 <- list_nth_shared_fwd i32 ls 0%u32;
+ ls <- list_nth_mut_back i32 (List_Cons 0%i32 l1) 1%u32 3%i32;
+ i3 <- list_nth_shared i32 ls 0%u32;
if negb (i3 s= 0%i32)
then Fail_ Failure
else (
- i4 <- list_nth_shared_fwd i32 ls 1%u32;
+ i4 <- list_nth_shared i32 ls 1%u32;
if negb (i4 s= 3%i32)
then Fail_ Failure
else (
- i5 <- list_nth_shared_fwd i32 ls 2%u32;
+ i5 <- list_nth_shared i32 ls 2%u32;
if negb (i5 s= 2%i32) then Fail_ Failure else Return tt))))))
.
(** Unit test for [no_nested_borrows::test_list_functions] *)
-Check (test_list_functions_fwd )%return.
+Check (test_list_functions )%return.
(** [no_nested_borrows::id_mut_pair1]: forward function *)
-Definition id_mut_pair1_fwd
- (T1 T2 : Type) (x : T1) (y : T2) : result (T1 * T2) :=
+Definition id_mut_pair1 (T1 T2 : Type) (x : T1) (y : T2) : result (T1 * T2) :=
Return (x, y)
.
@@ -366,8 +362,7 @@ Definition id_mut_pair1_back
.
(** [no_nested_borrows::id_mut_pair2]: forward function *)
-Definition id_mut_pair2_fwd
- (T1 T2 : Type) (p : (T1 * T2)) : result (T1 * T2) :=
+Definition id_mut_pair2 (T1 T2 : Type) (p : (T1 * T2)) : result (T1 * T2) :=
let (t, t0) := p in Return (t, t0)
.
@@ -378,8 +373,7 @@ Definition id_mut_pair2_back
.
(** [no_nested_borrows::id_mut_pair3]: forward function *)
-Definition id_mut_pair3_fwd
- (T1 T2 : Type) (x : T1) (y : T2) : result (T1 * T2) :=
+Definition id_mut_pair3 (T1 T2 : Type) (x : T1) (y : T2) : result (T1 * T2) :=
Return (x, y)
.
@@ -396,8 +390,7 @@ Definition id_mut_pair3_back'b
.
(** [no_nested_borrows::id_mut_pair4]: forward function *)
-Definition id_mut_pair4_fwd
- (T1 T2 : Type) (p : (T1 * T2)) : result (T1 * T2) :=
+Definition id_mut_pair4 (T1 T2 : Type) (p : (T1 * T2)) : result (T1 * T2) :=
let (t, t0) := p in Return (t, t0)
.
@@ -414,101 +407,101 @@ Definition id_mut_pair4_back'b
.
(** [no_nested_borrows::StructWithTuple] *)
-Record Struct_with_tuple_t (T1 T2 : Type) :=
-mkStruct_with_tuple_t {
- Struct_with_tuple_p : (T1 * T2);
+Record StructWithTuple_t (T1 T2 : Type) :=
+mkStructWithTuple_t {
+ structWithTuple_p : (T1 * T2);
}
.
-Arguments mkStruct_with_tuple_t {T1} {T2} _.
-Arguments Struct_with_tuple_p {T1} {T2}.
+Arguments mkStructWithTuple_t {T1} {T2} _.
+Arguments structWithTuple_p {T1} {T2}.
(** [no_nested_borrows::new_tuple1]: forward function *)
-Definition new_tuple1_fwd : result (Struct_with_tuple_t u32 u32) :=
- Return {| Struct_with_tuple_p := (1%u32, 2%u32) |}
+Definition new_tuple1 : result (StructWithTuple_t u32 u32) :=
+ Return {| structWithTuple_p := (1%u32, 2%u32) |}
.
(** [no_nested_borrows::new_tuple2]: forward function *)
-Definition new_tuple2_fwd : result (Struct_with_tuple_t i16 i16) :=
- Return {| Struct_with_tuple_p := (1%i16, 2%i16) |}
+Definition new_tuple2 : result (StructWithTuple_t i16 i16) :=
+ Return {| structWithTuple_p := (1%i16, 2%i16) |}
.
(** [no_nested_borrows::new_tuple3]: forward function *)
-Definition new_tuple3_fwd : result (Struct_with_tuple_t u64 i64) :=
- Return {| Struct_with_tuple_p := (1%u64, 2%i64) |}
+Definition new_tuple3 : result (StructWithTuple_t u64 i64) :=
+ Return {| structWithTuple_p := (1%u64, 2%i64) |}
.
(** [no_nested_borrows::StructWithPair] *)
-Record Struct_with_pair_t (T1 T2 : Type) :=
-mkStruct_with_pair_t {
- Struct_with_pair_p : Pair_t T1 T2;
+Record StructWithPair_t (T1 T2 : Type) :=
+mkStructWithPair_t {
+ structWithPair_p : Pair_t T1 T2;
}
.
-Arguments mkStruct_with_pair_t {T1} {T2} _.
-Arguments Struct_with_pair_p {T1} {T2}.
+Arguments mkStructWithPair_t {T1} {T2} _.
+Arguments structWithPair_p {T1} {T2}.
(** [no_nested_borrows::new_pair1]: forward function *)
-Definition new_pair1_fwd : result (Struct_with_pair_t u32 u32) :=
- Return {| Struct_with_pair_p := {| Pair_x := 1%u32; Pair_y := 2%u32 |} |}
+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 *)
-Definition test_constants_fwd : result unit :=
- swt <- new_tuple1_fwd;
- let (i, _) := swt.(Struct_with_tuple_p) in
+Definition test_constants : result unit :=
+ swt <- new_tuple1;
+ let (i, _) := swt.(structWithTuple_p) in
if negb (i s= 1%u32)
then Fail_ Failure
else (
- swt0 <- new_tuple2_fwd;
- let (i0, _) := swt0.(Struct_with_tuple_p) in
+ swt0 <- new_tuple2;
+ let (i0, _) := swt0.(structWithTuple_p) in
if negb (i0 s= 1%i16)
then Fail_ Failure
else (
- swt1 <- new_tuple3_fwd;
- let (i1, _) := swt1.(Struct_with_tuple_p) in
+ swt1 <- new_tuple3;
+ let (i1, _) := swt1.(structWithTuple_p) in
if negb (i1 s= 1%u64)
then Fail_ Failure
else (
- swp <- new_pair1_fwd;
- if negb (swp.(Struct_with_pair_p).(Pair_x) s= 1%u32)
+ swp <- new_pair1;
+ if negb (swp.(structWithPair_p).(pair_x) s= 1%u32)
then Fail_ Failure
else Return tt)))
.
(** Unit test for [no_nested_borrows::test_constants] *)
-Check (test_constants_fwd )%return.
+Check (test_constants )%return.
(** [no_nested_borrows::test_weird_borrows1]: forward function *)
-Definition test_weird_borrows1_fwd : result unit :=
+Definition test_weird_borrows1 : result unit :=
Return tt.
(** Unit test for [no_nested_borrows::test_weird_borrows1] *)
-Check (test_weird_borrows1_fwd )%return.
+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 ()) *)
-Definition test_mem_replace_fwd_back (px : u32) : result u32 :=
- let y := mem_replace_fwd u32 px 1%u32 in
+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 *)
-Definition test_shared_borrow_bool1_fwd (b : bool) : result u32 :=
+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 *)
-Definition test_shared_borrow_bool2_fwd : result u32 :=
+Definition test_shared_borrow_bool2 : result u32 :=
Return 0%u32.
(** [no_nested_borrows::test_shared_borrow_enum1]: forward function *)
-Definition test_shared_borrow_enum1_fwd (l : List_t u32) : result u32 :=
- match l with | ListCons i l0 => Return 1%u32 | ListNil => Return 0%u32 end
+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 *)
-Definition test_shared_borrow_enum2_fwd : result u32 :=
+Definition test_shared_borrow_enum2 : result u32 :=
Return 0%u32.
End NoNestedBorrows .
diff --git a/tests/coq/misc/Paper.v b/tests/coq/misc/Paper.v
index 175a523d..d397995b 100644
--- a/tests/coq/misc/Paper.v
+++ b/tests/coq/misc/Paper.v
@@ -10,20 +10,19 @@ Module Paper.
(** [paper::ref_incr]: merged forward/backward function
(there is a single backward function, and the forward function returns ()) *)
-Definition ref_incr_fwd_back (x : i32) : result i32 :=
+Definition ref_incr (x : i32) : result i32 :=
i32_add x 1%i32.
(** [paper::test_incr]: forward function *)
-Definition test_incr_fwd : result unit :=
- x <- ref_incr_fwd_back 0%i32;
- if negb (x s= 1%i32) then Fail_ Failure else Return tt
+Definition test_incr : result unit :=
+ x <- ref_incr 0%i32; if negb (x s= 1%i32) then Fail_ Failure else Return tt
.
(** Unit test for [paper::test_incr] *)
-Check (test_incr_fwd )%return.
+Check (test_incr )%return.
(** [paper::choose]: forward function *)
-Definition choose_fwd (T : Type) (b : bool) (x : T) (y : T) : result T :=
+Definition choose (T : Type) (b : bool) (x : T) (y : T) : result T :=
if b then Return x else Return y
.
@@ -34,8 +33,8 @@ Definition choose_back
.
(** [paper::test_choose]: forward function *)
-Definition test_choose_fwd : result unit :=
- z <- choose_fwd i32 true 0%i32 0%i32;
+Definition test_choose : result unit :=
+ z <- choose i32 true 0%i32 0%i32;
z0 <- i32_add z 1%i32;
if negb (z0 s= 1%i32)
then Fail_ Failure
@@ -48,25 +47,25 @@ Definition test_choose_fwd : result unit :=
.
(** Unit test for [paper::test_choose] *)
-Check (test_choose_fwd )%return.
+Check (test_choose )%return.
(** [paper::List] *)
Inductive List_t (T : Type) :=
-| ListCons : T -> List_t T -> List_t T
-| ListNil : List_t T
+| List_Cons : T -> List_t T -> List_t T
+| List_Nil : List_t T
.
-Arguments ListCons {T} _ _.
-Arguments ListNil {T}.
+Arguments List_Cons {T} _ _.
+Arguments List_Nil {T}.
(** [paper::list_nth_mut]: forward function *)
-Fixpoint list_nth_mut_fwd (T : Type) (l : List_t T) (i : u32) : result T :=
+Fixpoint list_nth_mut (T : Type) (l : List_t T) (i : u32) : result T :=
match l with
- | ListCons x tl =>
+ | List_Cons x tl =>
if i s= 0%u32
then Return x
- else (i0 <- u32_sub i 1%u32; list_nth_mut_fwd T tl i0)
- | ListNil => Fail_ Failure
+ else (i0 <- u32_sub i 1%u32; list_nth_mut T tl i0)
+ | List_Nil => Fail_ Failure
end
.
@@ -74,44 +73,44 @@ Fixpoint list_nth_mut_fwd (T : Type) (l : List_t T) (i : u32) : result T :=
Fixpoint list_nth_mut_back
(T : Type) (l : List_t T) (i : u32) (ret : T) : result (List_t T) :=
match l with
- | ListCons x tl =>
+ | List_Cons x tl =>
if i s= 0%u32
- then Return (ListCons ret tl)
+ then Return (List_Cons ret tl)
else (
i0 <- u32_sub i 1%u32;
tl0 <- list_nth_mut_back T tl i0 ret;
- Return (ListCons x tl0))
- | ListNil => Fail_ Failure
+ Return (List_Cons x tl0))
+ | List_Nil => Fail_ Failure
end
.
(** [paper::sum]: forward function *)
-Fixpoint sum_fwd (l : List_t i32) : result i32 :=
+Fixpoint sum (l : List_t i32) : result i32 :=
match l with
- | ListCons x tl => i <- sum_fwd tl; i32_add x i
- | ListNil => Return 0%i32
+ | List_Cons x tl => i <- sum tl; i32_add x i
+ | List_Nil => Return 0%i32
end
.
(** [paper::test_nth]: forward function *)
-Definition test_nth_fwd : result unit :=
- let l := ListNil in
- let l0 := ListCons 3%i32 l in
- let l1 := ListCons 2%i32 l0 in
- x <- list_nth_mut_fwd i32 (ListCons 1%i32 l1) 2%u32;
+Definition test_nth : result unit :=
+ let l := List_Nil in
+ let l0 := List_Cons 3%i32 l in
+ let l1 := List_Cons 2%i32 l0 in
+ x <- list_nth_mut i32 (List_Cons 1%i32 l1) 2%u32;
x0 <- i32_add x 1%i32;
- l2 <- list_nth_mut_back i32 (ListCons 1%i32 l1) 2%u32 x0;
- i <- sum_fwd l2;
+ l2 <- list_nth_mut_back i32 (List_Cons 1%i32 l1) 2%u32 x0;
+ i <- sum l2;
if negb (i s= 7%i32) then Fail_ Failure else Return tt
.
(** Unit test for [paper::test_nth] *)
-Check (test_nth_fwd )%return.
+Check (test_nth )%return.
(** [paper::call_choose]: forward function *)
-Definition call_choose_fwd (p : (u32 * u32)) : result u32 :=
+Definition call_choose (p : (u32 * u32)) : result u32 :=
let (px, py) := p in
- pz <- choose_fwd u32 true px py;
+ pz <- choose u32 true px py;
pz0 <- u32_add pz 1%u32;
p0 <- choose_back u32 true px py pz0;
let (px0, _) := p0 in
diff --git a/tests/coq/misc/PoloniusList.v b/tests/coq/misc/PoloniusList.v
index 54021bdf..4f804b55 100644
--- a/tests/coq/misc/PoloniusList.v
+++ b/tests/coq/misc/PoloniusList.v
@@ -10,19 +10,19 @@ Module PoloniusList.
(** [polonius_list::List] *)
Inductive List_t (T : Type) :=
-| ListCons : T -> List_t T -> List_t T
-| ListNil : List_t T
+| List_Cons : T -> List_t T -> List_t T
+| List_Nil : List_t T
.
-Arguments ListCons {T} _ _.
-Arguments ListNil {T}.
+Arguments List_Cons {T} _ _.
+Arguments List_Nil {T}.
(** [polonius_list::get_list_at_x]: forward function *)
-Fixpoint get_list_at_x_fwd (ls : List_t u32) (x : u32) : result (List_t u32) :=
+Fixpoint get_list_at_x (ls : List_t u32) (x : u32) : result (List_t u32) :=
match ls with
- | ListCons hd tl =>
- if hd s= x then Return (ListCons hd tl) else get_list_at_x_fwd tl x
- | ListNil => Return ListNil
+ | List_Cons hd tl =>
+ if hd s= x then Return (List_Cons hd tl) else get_list_at_x tl x
+ | List_Nil => Return List_Nil
end
.
@@ -30,11 +30,11 @@ Fixpoint get_list_at_x_fwd (ls : List_t u32) (x : u32) : result (List_t u32) :=
Fixpoint get_list_at_x_back
(ls : List_t u32) (x : u32) (ret : List_t u32) : result (List_t u32) :=
match ls with
- | ListCons hd tl =>
+ | List_Cons hd tl =>
if hd s= x
then Return ret
- else (tl0 <- get_list_at_x_back tl x ret; Return (ListCons hd tl0))
- | ListNil => Return ret
+ else (tl0 <- get_list_at_x_back tl x ret; Return (List_Cons hd tl0))
+ | List_Nil => Return ret
end
.
diff --git a/tests/coq/misc/Primitives.v b/tests/coq/misc/Primitives.v
index 8d6c9c8d..85e38f01 100644
--- a/tests/coq/misc/Primitives.v
+++ b/tests/coq/misc/Primitives.v
@@ -63,13 +63,15 @@ Check (if true then Return (1 + 2) else Fail_ Failure)%global = 3.
(*** Misc *)
-
Definition string := Coq.Strings.String.string.
Definition char := Coq.Strings.Ascii.ascii.
Definition char_of_byte := Coq.Strings.Ascii.ascii_of_byte.
-Definition mem_replace_fwd (a : Type) (x : a) (y : a) : a := x .
-Definition mem_replace_back (a : Type) (x : a) (y : a) : a := y .
+Definition core_mem_replace (a : Type) (x : a) (y : a) : a := x .
+Definition core_mem_replace_back (a : Type) (x : a) (y : a) : a := y .
+
+Record mut_raw_ptr (T : Type) := { mut_raw_ptr_v : T }.
+Record const_raw_ptr (T : Type) := { const_raw_ptr_v : T }.
(*** Scalars *)
@@ -408,12 +410,75 @@ Definition core_i64_max := i64_max %i64.
Definition core_i128_max := i64_max %i128.
Axiom core_isize_max : isize. (** TODO *)
-(*** Range *)
-Record range (T : Type) := mk_range {
- start: T;
- end_: T;
+(*** core::ops *)
+
+(* Trait declaration: [core::ops::index::Index] *)
+Record core_ops_index_Index (Self Idx : Type) := mk_core_ops_index_Index {
+ core_ops_index_Index_Output : Type;
+ core_ops_index_Index_index : Self -> Idx -> result core_ops_index_Index_Output;
+}.
+Arguments mk_core_ops_index_Index {_ _}.
+Arguments core_ops_index_Index_Output {_ _}.
+Arguments core_ops_index_Index_index {_ _}.
+
+(* Trait declaration: [core::ops::index::IndexMut] *)
+Record core_ops_index_IndexMut (Self Idx : Type) := mk_core_ops_index_IndexMut {
+ core_ops_index_IndexMut_indexInst : core_ops_index_Index Self Idx;
+ core_ops_index_IndexMut_index_mut : Self -> Idx -> result core_ops_index_IndexMut_indexInst.(core_ops_index_Index_Output);
+ core_ops_index_IndexMut_index_mut_back : Self -> Idx -> core_ops_index_IndexMut_indexInst.(core_ops_index_Index_Output) -> result Self;
+}.
+Arguments mk_core_ops_index_IndexMut {_ _}.
+Arguments core_ops_index_IndexMut_indexInst {_ _}.
+Arguments core_ops_index_IndexMut_index_mut {_ _}.
+Arguments core_ops_index_IndexMut_index_mut_back {_ _}.
+
+(* Trait declaration [core::ops::deref::Deref] *)
+Record core_ops_deref_Deref (Self : Type) := mk_core_ops_deref_Deref {
+ core_ops_deref_Deref_target : Type;
+ core_ops_deref_Deref_deref : Self -> result core_ops_deref_Deref_target;
+}.
+Arguments mk_core_ops_deref_Deref {_}.
+Arguments core_ops_deref_Deref_target {_}.
+Arguments core_ops_deref_Deref_deref {_}.
+
+(* Trait declaration [core::ops::deref::DerefMut] *)
+Record core_ops_deref_DerefMut (Self : Type) := mk_core_ops_deref_DerefMut {
+ core_ops_deref_DerefMut_derefInst : core_ops_deref_Deref Self;
+ core_ops_deref_DerefMut_deref_mut : Self -> result core_ops_deref_DerefMut_derefInst.(core_ops_deref_Deref_target);
+ core_ops_deref_DerefMut_deref_mut_back : Self -> core_ops_deref_DerefMut_derefInst.(core_ops_deref_Deref_target) -> result Self;
}.
-Arguments mk_range {_}.
+Arguments mk_core_ops_deref_DerefMut {_}.
+Arguments core_ops_deref_DerefMut_derefInst {_}.
+Arguments core_ops_deref_DerefMut_deref_mut {_}.
+Arguments core_ops_deref_DerefMut_deref_mut_back {_}.
+
+Record core_ops_range_Range (T : Type) := mk_core_ops_range_Range {
+ core_ops_range_Range_start : T;
+ core_ops_range_Range_end_ : T;
+}.
+Arguments mk_core_ops_range_Range {_}.
+Arguments core_ops_range_Range_start {_}.
+Arguments core_ops_range_Range_end_ {_}.
+
+(*** [alloc] *)
+
+Definition alloc_boxed_Box_deref (T : Type) (x : T) : result T := Return x.
+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 := {|
+ 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;
+ 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;
+|}.
+
(*** Arrays *)
Definition array T (n : usize) := { l: list T | Z.of_nat (length l) = to_Z n}.
@@ -433,51 +498,50 @@ Qed.
(* TODO: finish the definitions *)
Axiom mk_array : forall (T : Type) (n : usize) (l : list T), array T n.
-Axiom array_index_shared : forall (T : Type) (n : usize) (x : array T n) (i : usize), result T.
-Axiom array_index_mut_fwd : forall (T : Type) (n : usize) (x : array T n) (i : usize), result T.
-Axiom array_index_mut_back : forall (T : Type) (n : usize) (x : array T n) (i : usize) (nx : T), result (array T n).
+(* For initialization *)
+Axiom array_repeat : forall (T : Type) (n : usize) (x : T), array T n.
+
+Axiom array_index_usize : forall (T : Type) (n : usize) (x : array T n) (i : usize), result T.
+Axiom array_update_usize : forall (T : Type) (n : usize) (x : array T n) (i : usize) (nx : T), result (array T n).
(*** Slice *)
Definition slice T := { l: list T | Z.of_nat (length l) <= usize_max}.
Axiom slice_len : forall (T : Type) (s : slice T), usize.
-Axiom slice_index_shared : forall (T : Type) (x : slice T) (i : usize), result T.
-Axiom slice_index_mut_fwd : forall (T : Type) (x : slice T) (i : usize), result T.
-Axiom slice_index_mut_back : forall (T : Type) (x : slice T) (i : usize) (nx : T), result (slice T).
+Axiom slice_index_usize : forall (T : Type) (x : slice T) (i : usize), result T.
+Axiom slice_update_usize : forall (T : Type) (x : slice T) (i : usize) (nx : T), result (slice T).
(*** Subslices *)
-Axiom array_to_slice_shared : forall (T : Type) (n : usize) (x : array T n), result (slice T).
-Axiom array_to_slice_mut_fwd : forall (T : Type) (n : usize) (x : array T n), result (slice T).
-Axiom array_to_slice_mut_back : forall (T : Type) (n : usize) (x : array T n) (s : slice T), result (array T n).
+Axiom array_to_slice : forall (T : Type) (n : usize) (x : array T n), result (slice T).
+Axiom array_from_slice : forall (T : Type) (n : usize) (x : array T n) (s : slice T), result (array T n).
+
+Axiom array_subslice: forall (T : Type) (n : usize) (x : array T n) (r : core_ops_range_Range usize), result (slice T).
+Axiom array_update_subslice: forall (T : Type) (n : usize) (x : array T n) (r : core_ops_range_Range usize) (ns : slice T), result (array T n).
-Axiom array_subslice_shared: forall (T : Type) (n : usize) (x : array T n) (r : range usize), result (slice T).
-Axiom array_subslice_mut_fwd: forall (T : Type) (n : usize) (x : array T n) (r : range usize), result (slice T).
-Axiom array_subslice_mut_back: forall (T : Type) (n : usize) (x : array T n) (r : range usize) (ns : slice T), result (array T n).
-Axiom slice_subslice_shared: forall (T : Type) (x : slice T) (r : range usize), result (slice T).
-Axiom slice_subslice_mut_fwd: forall (T : Type) (x : slice T) (r : range usize), result (slice T).
-Axiom slice_subslice_mut_back: forall (T : Type) (x : slice T) (r : range usize) (ns : slice T), result (slice T).
+Axiom slice_subslice: forall (T : Type) (x : slice T) (r : core_ops_range_Range usize), result (slice T).
+Axiom slice_update_subslice: forall (T : Type) (x : slice T) (r : core_ops_range_Range usize) (ns : slice T), result (slice T).
(*** Vectors *)
-Definition vec T := { l: list T | Z.of_nat (length l) <= usize_max }.
+Definition alloc_vec_Vec T := { l: list T | Z.of_nat (length l) <= usize_max }.
-Definition vec_to_list {T: Type} (v: vec T) : list T := proj1_sig v.
+Definition alloc_vec_Vec_to_list {T: Type} (v: alloc_vec_Vec T) : list T := proj1_sig v.
-Definition vec_length {T: Type} (v: vec T) : Z := Z.of_nat (length (vec_to_list v)).
+Definition alloc_vec_Vec_length {T: Type} (v: alloc_vec_Vec T) : Z := Z.of_nat (length (alloc_vec_Vec_to_list v)).
-Definition vec_new (T: Type) : vec T := (exist _ [] le_0_usize_max).
+Definition alloc_vec_Vec_new (T: Type) : alloc_vec_Vec T := (exist _ [] le_0_usize_max).
-Lemma vec_len_in_usize {T} (v: vec T) : usize_min <= vec_length v <= usize_max.
+Lemma alloc_vec_Vec_len_in_usize {T} (v: alloc_vec_Vec T) : usize_min <= alloc_vec_Vec_length v <= usize_max.
Proof.
- unfold vec_length, usize_min.
+ unfold alloc_vec_Vec_length, usize_min.
split.
- lia.
- apply (proj2_sig v).
Qed.
-Definition vec_len (T: Type) (v: vec T) : usize :=
- exist _ (vec_length v) (vec_len_in_usize v).
+Definition alloc_vec_Vec_len (T: Type) (v: alloc_vec_Vec T) : usize :=
+ exist _ (alloc_vec_Vec_length v) (alloc_vec_Vec_len_in_usize v).
Fixpoint list_update {A} (l: list A) (n: nat) (a: A)
: list A :=
@@ -488,50 +552,271 @@ Fixpoint list_update {A} (l: list A) (n: nat) (a: A)
| S m => x :: (list_update t m a)
end end.
-Definition vec_bind {A B} (v: vec A) (f: list A -> result (list B)) : result (vec B) :=
- l <- f (vec_to_list v) ;
+Definition alloc_vec_Vec_bind {A B} (v: alloc_vec_Vec A) (f: list A -> result (list B)) : result (alloc_vec_Vec B) :=
+ l <- f (alloc_vec_Vec_to_list v) ;
match sumbool_of_bool (scalar_le_max Usize (Z.of_nat (length l))) with
| left H => Return (exist _ l (scalar_le_max_valid _ _ H))
| right _ => Fail_ Failure
end.
(* The **forward** function shouldn't be used *)
-Definition vec_push_fwd (T: Type) (v: vec T) (x: T) : unit := tt.
+Definition alloc_vec_Vec_push_fwd (T: Type) (v: alloc_vec_Vec T) (x: T) : unit := tt.
-Definition vec_push_back (T: Type) (v: vec T) (x: T) : result (vec T) :=
- vec_bind v (fun l => Return (l ++ [x])).
+Definition alloc_vec_Vec_push (T: Type) (v: alloc_vec_Vec T) (x: T) : result (alloc_vec_Vec T) :=
+ alloc_vec_Vec_bind v (fun l => Return (l ++ [x])).
(* The **forward** function shouldn't be used *)
-Definition vec_insert_fwd (T: Type) (v: vec T) (i: usize) (x: T) : result unit :=
- if to_Z i <? vec_length v then Return tt else Fail_ Failure.
+Definition alloc_vec_Vec_insert_fwd (T: Type) (v: alloc_vec_Vec T) (i: usize) (x: T) : result unit :=
+ if to_Z i <? alloc_vec_Vec_length v then Return tt else Fail_ Failure.
-Definition vec_insert_back (T: Type) (v: vec T) (i: usize) (x: T) : result (vec T) :=
- vec_bind v (fun l =>
+Definition alloc_vec_Vec_insert (T: Type) (v: alloc_vec_Vec T) (i: usize) (x: T) : result (alloc_vec_Vec T) :=
+ alloc_vec_Vec_bind v (fun l =>
if to_Z i <? Z.of_nat (length l)
then Return (list_update l (usize_to_nat i) x)
else Fail_ Failure).
-(* The **backward** function shouldn't be used *)
-Definition vec_index_fwd (T: Type) (v: vec T) (i: usize) : result T :=
- match nth_error (vec_to_list v) (usize_to_nat i) with
- | Some n => Return n
- | None => Fail_ Failure
- end.
-
-Definition vec_index_back (T: Type) (v: vec T) (i: usize) (x: T) : result unit :=
- if to_Z i <? vec_length v then Return tt else Fail_ Failure.
-
-(* The **backward** function shouldn't be used *)
-Definition vec_index_mut_fwd (T: Type) (v: vec T) (i: usize) : result T :=
- match nth_error (vec_to_list v) (usize_to_nat i) with
- | Some n => Return n
- | None => Fail_ Failure
+(* Helper *)
+Axiom alloc_vec_Vec_index_usize : forall {T : Type} (v : alloc_vec_Vec T) (i : usize) (x : T), 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).
+
+(* Trait declaration: [core::slice::index::private_slice_index::Sealed] *)
+Definition core_slice_index_private_slice_index_Sealed (self : Type) := unit.
+
+(* Trait declaration: [core::slice::index::SliceIndex] *)
+Record core_slice_index_SliceIndex (Self T : Type) := mk_core_slice_index_SliceIndex {
+ core_slice_index_SliceIndex_sealedInst : core_slice_index_private_slice_index_Sealed Self;
+ core_slice_index_SliceIndex_Output : Type;
+ core_slice_index_SliceIndex_get : Self -> T -> result (option core_slice_index_SliceIndex_Output);
+ core_slice_index_SliceIndex_get_mut : Self -> T -> result (option core_slice_index_SliceIndex_Output);
+ core_slice_index_SliceIndex_get_mut_back : Self -> T -> option core_slice_index_SliceIndex_Output -> result T;
+ core_slice_index_SliceIndex_get_unchecked : Self -> const_raw_ptr T -> result (const_raw_ptr core_slice_index_SliceIndex_Output);
+ core_slice_index_SliceIndex_get_unchecked_mut : Self -> mut_raw_ptr T -> result (mut_raw_ptr core_slice_index_SliceIndex_Output);
+ core_slice_index_SliceIndex_index : Self -> T -> result core_slice_index_SliceIndex_Output;
+ core_slice_index_SliceIndex_index_mut : Self -> T -> result core_slice_index_SliceIndex_Output;
+ core_slice_index_SliceIndex_index_mut_back : Self -> T -> core_slice_index_SliceIndex_Output -> result T;
+}.
+Arguments mk_core_slice_index_SliceIndex {_ _}.
+Arguments core_slice_index_SliceIndex_sealedInst {_ _}.
+Arguments core_slice_index_SliceIndex_Output {_ _}.
+Arguments core_slice_index_SliceIndex_get {_ _}.
+Arguments core_slice_index_SliceIndex_get_mut {_ _}.
+Arguments core_slice_index_SliceIndex_get_mut_back {_ _}.
+Arguments core_slice_index_SliceIndex_get_unchecked {_ _}.
+Arguments core_slice_index_SliceIndex_get_unchecked_mut {_ _}.
+Arguments core_slice_index_SliceIndex_index {_ _}.
+Arguments core_slice_index_SliceIndex_index_mut {_ _}.
+Arguments core_slice_index_SliceIndex_index_mut_back {_ _}.
+
+(* [core::slice::index::[T]::index]: forward function *)
+Definition core_slice_index_Slice_index
+ (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T))
+ (s : slice T) (i : Idx) : result inst.(core_slice_index_SliceIndex_Output) :=
+ x <- inst.(core_slice_index_SliceIndex_get) i s;
+ match x with
+ | None => Fail_ Failure
+ | Some x => Return x
end.
-Definition vec_index_mut_back (T: Type) (v: vec T) (i: usize) (x: T) : result (vec T) :=
- vec_bind v (fun l =>
- if to_Z i <? Z.of_nat (length l)
- then Return (list_update l (usize_to_nat i) x)
- else Fail_ Failure).
+(* [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)).
+
+(* [core::slice::index::Range::get_mut]: forward function *)
+Axiom core_slice_index_Range_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 :
+ 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
+ (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
+ sure code which uses it fails *)
+ fun _ _ => Fail_ Failure.
+
+(* [core::slice::index::Range::get_unchecked_mut]: forward function *)
+Definition core_slice_index_Range_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
+ sure code which uses it fails *)
+ fun _ _ => Fail_ Failure.
+
+(* [core::slice::index::Range::index]: forward function *)
+Axiom core_slice_index_Range_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 :
+ 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 :
+ forall (T : Type), core_ops_range_Range usize -> slice T -> slice T -> result (slice T).
+
+(* [core::slice::index::[T]::index_mut]: forward function *)
+Axiom core_slice_index_Slice_index_mut :
+ forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)),
+ slice T -> Idx -> result inst.(core_slice_index_SliceIndex_Output).
+
+(* [core::slice::index::[T]::index_mut]: backward function 0 *)
+Axiom core_slice_index_Slice_index_mut_back :
+ forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T)),
+ slice T -> Idx -> inst.(core_slice_index_SliceIndex_Output) -> result (slice T).
+
+(* [core::array::[T; N]::index]: forward function *)
+Axiom core_array_Array_index :
+ forall (T Idx : Type) (N : usize) (inst : core_ops_index_Index (slice T) Idx)
+ (a : array T N) (i : Idx), result inst.(core_ops_index_Index_Output).
+
+(* [core::array::[T; N]::index_mut]: forward function *)
+Axiom core_array_Array_index_mut :
+ forall (T Idx : Type) (N : usize) (inst : core_ops_index_IndexMut (slice T) Idx)
+ (a : array T N) (i : Idx), result inst.(core_ops_index_IndexMut_indexInst).(core_ops_index_Index_Output).
+
+(* [core::array::[T; N]::index_mut]: backward function 0 *)
+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
+ : 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) :
+ 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_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;
+|}.
+
+(* Trait implementation: [core::slice::index::[T]] *)
+Definition core_slice_index_Slice_coreopsindexIndexMutInst (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_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)
+ (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);
+ core_ops_index_Index_index := core_array_Array_index T Idx N inst;
+|}.
+
+(* Trait implementation: [core::array::[T; N]] *)
+Definition core_array_Array_coreopsindexIndexMutInst (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_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;
+|}.
+
+(* [core::slice::index::usize::get]: forward function *)
+Axiom core_slice_index_usize_get : forall (T : Type), usize -> slice T -> result (option T).
+
+(* [core::slice::index::usize::get_mut]: forward function *)
+Axiom core_slice_index_usize_get_mut : forall (T : Type), usize -> slice T -> result (option T).
+
+(* [core::slice::index::usize::get_mut]: backward function 0 *)
+Axiom core_slice_index_usize_get_mut_back :
+ forall (T : Type), usize -> slice T -> option T -> result (slice T).
+
+(* [core::slice::index::usize::get_unchecked]: forward function *)
+Axiom core_slice_index_usize_get_unchecked :
+ forall (T : Type), usize -> const_raw_ptr (slice T) -> result (const_raw_ptr T).
+
+(* [core::slice::index::usize::get_unchecked_mut]: forward function *)
+Axiom core_slice_index_usize_get_unchecked_mut :
+ forall (T : Type), usize -> mut_raw_ptr (slice T) -> result (mut_raw_ptr T).
+
+(* [core::slice::index::usize::index]: forward function *)
+Axiom core_slice_index_usize_index : forall (T : Type), usize -> slice T -> result T.
+
+(* [core::slice::index::usize::index_mut]: forward function *)
+Axiom core_slice_index_usize_index_mut : forall (T : Type), usize -> slice T -> result T.
+
+(* [core::slice::index::usize::index_mut]: backward function 0 *)
+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
+ : core_slice_index_private_slice_index_Sealed usize := tt.
+
+(* Trait implementation: [core::slice::index::usize] *)
+Definition core_slice_index_usize_coresliceindexSliceIndexInst (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_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;
+ core_slice_index_SliceIndex_get_mut_back := core_slice_index_usize_get_mut_back T;
+ core_slice_index_SliceIndex_get_unchecked := core_slice_index_usize_get_unchecked T;
+ core_slice_index_SliceIndex_get_unchecked_mut := core_slice_index_usize_get_unchecked_mut T;
+ core_slice_index_SliceIndex_index := core_slice_index_usize_index T;
+ core_slice_index_SliceIndex_index_mut := core_slice_index_usize_index_mut T;
+ core_slice_index_SliceIndex_index_mut_back := core_slice_index_usize_index_mut_back T;
+|}.
+
+(* [alloc::vec::Vec::index]: forward function *)
+Axiom alloc_vec_Vec_index : forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T))
+ (Self : alloc_vec_Vec T) (i : Idx), result inst.(core_slice_index_SliceIndex_Output).
+
+(* [alloc::vec::Vec::index_mut]: forward function *)
+Axiom alloc_vec_Vec_index_mut : forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T))
+ (Self : alloc_vec_Vec T) (i : Idx), result inst.(core_slice_index_SliceIndex_Output).
+
+(* [alloc::vec::Vec::index_mut]: backward function 0 *)
+Axiom alloc_vec_Vec_index_mut_back :
+ forall (T Idx : Type) (inst : core_slice_index_SliceIndex Idx (slice T))
+ (Self : alloc_vec_Vec T) (i : Idx) (x : inst.(core_slice_index_SliceIndex_Output)), result (alloc_vec_Vec T).
+
+(* Trait implementation: [alloc::vec::Vec] *)
+Definition alloc_vec_Vec_coreopsindexIndexInst (T Idx : Type)
+ (inst : core_slice_index_SliceIndex Idx (slice T)) :
+ core_ops_index_Index (alloc_vec_Vec T) Idx := {|
+ core_ops_index_Index_Output := inst.(core_slice_index_SliceIndex_Output);
+ core_ops_index_Index_index := alloc_vec_Vec_index T Idx inst;
+|}.
+
+(* Trait implementation: [alloc::vec::Vec] *)
+Definition alloc_vec_Vec_coreopsindexIndexMutInst (T Idx : Type)
+ (inst : core_slice_index_SliceIndex Idx (slice T)) :
+ core_ops_index_IndexMut (alloc_vec_Vec T) Idx := {|
+ core_ops_index_IndexMut_indexInst := alloc_vec_Vec_coreopsindexIndexInst T Idx inst;
+ core_ops_index_IndexMut_index_mut := alloc_vec_Vec_index_mut T Idx inst;
+ core_ops_index_IndexMut_index_mut_back := alloc_vec_Vec_index_mut_back T Idx inst;
+|}.
+
+(*** Theorems *)
+
+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_update_usize v i x.
End Primitives.