diff options
Diffstat (limited to '')
-rw-r--r-- | tests/coq/hashmap/Hashmap__Funs.v | 541 | ||||
-rw-r--r-- | tests/coq/hashmap/Hashmap__Types.v | 39 | ||||
-rw-r--r-- | tests/coq/hashmap/Makefile | 22 | ||||
-rw-r--r-- | tests/coq/hashmap/Primitives.v | 482 | ||||
-rw-r--r-- | tests/coq/hashmap/_CoqProject | 8 | ||||
-rw-r--r-- | tests/coq/hashmap_on_disk/HashmapMain__Funs.v | 596 | ||||
-rw-r--r-- | tests/coq/hashmap_on_disk/HashmapMain__Opaque.v | 21 | ||||
-rw-r--r-- | tests/coq/hashmap_on_disk/HashmapMain__Types.v | 42 | ||||
-rw-r--r-- | tests/coq/hashmap_on_disk/Makefile | 22 | ||||
-rw-r--r-- | tests/coq/hashmap_on_disk/Primitives.v | 482 | ||||
-rw-r--r-- | tests/coq/hashmap_on_disk/_CoqProject | 9 |
11 files changed, 2264 insertions, 0 deletions
diff --git a/tests/coq/hashmap/Hashmap__Funs.v b/tests/coq/hashmap/Hashmap__Funs.v new file mode 100644 index 00000000..93aa389b --- /dev/null +++ b/tests/coq/hashmap/Hashmap__Funs.v @@ -0,0 +1,541 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap]: function definitions *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Local Open Scope Primitives_scope. +Require Export Hashmap__Types . +Import Hashmap__Types . +Module Hashmap__Funs . + +(** [hashmap::hash_key] *) +Definition hash_key_fwd (k : usize) : result usize := Return k . + +(** [hashmap::HashMap::{0}::allocate_slots] *) +Fixpoint hash_map_allocate_slots_fwd + (T : Type) (n : nat) (slots : vec (List_t T)) (n0 : usize) : + result (vec (List_t T)) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + if n0 s= 0 %usize + then Return slots + else ( + slots0 <- vec_push_back (List_t T) slots ListNil; + i <- usize_sub n0 1 %usize; + v <- hash_map_allocate_slots_fwd T n1 slots0 i; + Return v) + end + . + +(** [hashmap::HashMap::{0}::new_with_capacity] *) +Definition hash_map_new_with_capacity_fwd + (T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize) + (max_load_divisor : usize) : + result (Hash_map_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + let v := vec_new (List_t T) in + slots <- hash_map_allocate_slots_fwd T n0 v capacity; + i <- usize_mul capacity max_load_dividend; + i0 <- usize_div i max_load_divisor; + Return (mkHash_map_t (0 %usize) (max_load_dividend, max_load_divisor) i0 + slots) + end + . + +(** [hashmap::HashMap::{0}::new] *) +Definition hash_map_new_fwd (T : Type) (n : nat) : result (Hash_map_t T) := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hm <- + hash_map_new_with_capacity_fwd T n0 (32 %usize) (4 %usize) (5 %usize); + Return hm + end + . + +(** [hashmap::HashMap::{0}::clear_slots] *) +Fixpoint hash_map_clear_slots_fwd_back + (T : Type) (n : nat) (slots : vec (List_t T)) (i : usize) : + result (vec (List_t T)) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + let i0 := vec_len (List_t T) slots in + if i s< i0 + then ( + slots0 <- vec_index_mut_back (List_t T) slots i ListNil; + i1 <- usize_add i 1 %usize; + slots1 <- hash_map_clear_slots_fwd_back T n0 slots0 i1; + Return slots1) + else Return slots + end + . + +(** [hashmap::HashMap::{0}::clear] *) +Definition hash_map_clear_fwd_back + (T : Type) (n : nat) (self : Hash_map_t T) : result (Hash_map_t T) := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match self with + | mkHash_map_t i p i0 v => + v0 <- hash_map_clear_slots_fwd_back T n0 v (0 %usize); + Return (mkHash_map_t (0 %usize) p i0 v0) + end + end + . + +(** [hashmap::HashMap::{0}::len] *) +Definition hash_map_len_fwd (T : Type) (self : Hash_map_t T) : result usize := + match self with | mkHash_map_t i p i0 v => Return i end . + +(** [hashmap::HashMap::{0}::insert_in_list] *) +Fixpoint hash_map_insert_in_list_fwd + (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : + result bool + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | ListCons ckey cvalue ls0 => + if ckey s= key + then Return false + else (b <- hash_map_insert_in_list_fwd T n0 key value ls0; Return b) + | ListNil => Return true + end + end + . + +(** [hashmap::HashMap::{0}::insert_in_list] *) +Fixpoint hash_map_insert_in_list_back + (T : Type) (n : nat) (key : usize) (value : T) (ls : List_t T) : + result (List_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | ListCons ckey cvalue ls0 => + if ckey s= key + then Return (ListCons ckey value ls0) + else ( + ls1 <- hash_map_insert_in_list_back T n0 key value ls0; + Return (ListCons ckey cvalue ls1)) + | ListNil => let l := ListNil in Return (ListCons key value l) + end + end + . + +(** [hashmap::HashMap::{0}::insert_no_resize] *) +Definition hash_map_insert_no_resize_fwd_back + (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) (value : T) : + result (Hash_map_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hash <- hash_key_fwd key; + match self with + | mkHash_map_t i p i0 v => + let i1 := vec_len (List_t T) v in + hash_mod <- usize_rem hash i1; + l <- vec_index_mut_fwd (List_t T) v hash_mod; + inserted <- hash_map_insert_in_list_fwd T n0 key value l; + if inserted + then ( + i2 <- usize_add i 1 %usize; + l0 <- hash_map_insert_in_list_back T n0 key value l; + v0 <- vec_index_mut_back (List_t T) v hash_mod l0; + Return (mkHash_map_t i2 p i0 v0)) + else ( + l0 <- hash_map_insert_in_list_back T n0 key value l; + v0 <- vec_index_mut_back (List_t T) v hash_mod l0; + Return (mkHash_map_t i p i0 v0)) + end + end + . + +(** [core::num::u32::{9}::MAX] *) +Definition core_num_u32_max_body : result u32 := Return (4294967295 %u32) . +Definition core_num_u32_max_c : u32 := core_num_u32_max_body%global . + +(** [hashmap::HashMap::{0}::move_elements_from_list] *) +Fixpoint hash_map_move_elements_from_list_fwd_back + (T : Type) (n : nat) (ntable : Hash_map_t T) (ls : List_t T) : + result (Hash_map_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | ListCons k v tl => + ntable0 <- hash_map_insert_no_resize_fwd_back T n0 ntable k v; + ntable1 <- hash_map_move_elements_from_list_fwd_back T n0 ntable0 tl; + Return ntable1 + | ListNil => Return ntable + end + end + . + +(** [hashmap::HashMap::{0}::move_elements] *) +Fixpoint hash_map_move_elements_fwd_back + (T : Type) (n : nat) (ntable : Hash_map_t T) (slots : vec (List_t T)) + (i : usize) : + result ((Hash_map_t T) * (vec (List_t T))) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + let i0 := vec_len (List_t T) slots in + if i s< i0 + then ( + l <- vec_index_mut_fwd (List_t T) slots i; + let ls := mem_replace_fwd (List_t T) l ListNil in + ntable0 <- hash_map_move_elements_from_list_fwd_back T n0 ntable ls; + let l0 := mem_replace_back (List_t T) l ListNil in + slots0 <- vec_index_mut_back (List_t T) slots i l0; + i1 <- usize_add i 1 %usize; + p <- hash_map_move_elements_fwd_back T n0 ntable0 slots0 i1; + let (ntable1, slots1) := p in + Return (ntable1, slots1)) + else Return (ntable, slots) + end + . + +(** [hashmap::HashMap::{0}::try_resize] *) +Definition hash_map_try_resize_fwd_back + (T : Type) (n : nat) (self : Hash_map_t T) : result (Hash_map_t T) := + match n with + | O => Fail_ OutOfFuel + | S n0 => + max_usize <- scalar_cast U32 Usize core_num_u32_max_c; + match self with + | mkHash_map_t i p i0 v => + let capacity := vec_len (List_t T) v in + n1 <- usize_div max_usize 2 %usize; + let (i1, i2) := p in + i3 <- usize_div n1 i1; + if capacity s<= i3 + then ( + i4 <- usize_mul capacity 2 %usize; + ntable <- hash_map_new_with_capacity_fwd T n0 i4 i1 i2; + p0 <- hash_map_move_elements_fwd_back T n0 ntable v (0 %usize); + let (ntable0, _) := p0 in + match ntable0 with + | mkHash_map_t i5 p1 i6 v0 => Return (mkHash_map_t i (i1, i2) i6 v0) + end) + else Return (mkHash_map_t i (i1, i2) i0 v) + end + end + . + +(** [hashmap::HashMap::{0}::insert] *) +Definition hash_map_insert_fwd_back + (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) (value : T) : + result (Hash_map_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + self0 <- hash_map_insert_no_resize_fwd_back T n0 self key value; + i <- hash_map_len_fwd T self0; + match self0 with + | mkHash_map_t i0 p i1 v => + if i s> i1 + then ( + self1 <- hash_map_try_resize_fwd_back T n0 (mkHash_map_t i0 p i1 v); + Return self1) + else Return (mkHash_map_t i0 p i1 v) + end + end + . + +(** [hashmap::HashMap::{0}::contains_key_in_list] *) +Fixpoint hash_map_contains_key_in_list_fwd + (T : Type) (n : nat) (key : usize) (ls : List_t T) : result bool := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | ListCons ckey t ls0 => + if ckey s= key + then Return true + else (b <- hash_map_contains_key_in_list_fwd T n0 key ls0; Return b) + | ListNil => Return false + end + end + . + +(** [hashmap::HashMap::{0}::contains_key] *) +Definition hash_map_contains_key_fwd + (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) : result bool := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hash <- hash_key_fwd key; + match self with + | mkHash_map_t i p i0 v => + let i1 := vec_len (List_t T) v in + hash_mod <- usize_rem hash i1; + l <- vec_index_fwd (List_t T) v hash_mod; + b <- hash_map_contains_key_in_list_fwd T n0 key l; + Return b + end + end + . + +(** [hashmap::HashMap::{0}::get_in_list] *) +Fixpoint hash_map_get_in_list_fwd + (T : Type) (n : nat) (key : usize) (ls : List_t T) : result T := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | ListCons ckey cvalue ls0 => + if ckey s= key + then Return cvalue + else (t <- hash_map_get_in_list_fwd T n0 key ls0; Return t) + | ListNil => Fail_ Failure + end + end + . + +(** [hashmap::HashMap::{0}::get] *) +Definition hash_map_get_fwd + (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) : result T := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hash <- hash_key_fwd key; + match self with + | mkHash_map_t i p i0 v => + let i1 := vec_len (List_t T) v in + hash_mod <- usize_rem hash i1; + l <- vec_index_fwd (List_t T) v hash_mod; + t <- hash_map_get_in_list_fwd T n0 key l; + Return t + end + end + . + +(** [hashmap::HashMap::{0}::get_mut_in_list] *) +Fixpoint hash_map_get_mut_in_list_fwd + (T : Type) (n : nat) (key : usize) (ls : List_t T) : result T := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | ListCons ckey cvalue ls0 => + if ckey s= key + then Return cvalue + else (t <- hash_map_get_mut_in_list_fwd T n0 key ls0; Return t) + | ListNil => Fail_ Failure + end + end + . + +(** [hashmap::HashMap::{0}::get_mut_in_list] *) +Fixpoint hash_map_get_mut_in_list_back + (T : Type) (n : nat) (key : usize) (ls : List_t T) (ret : T) : + result (List_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | ListCons ckey cvalue ls0 => + if ckey s= key + then Return (ListCons ckey ret ls0) + else ( + ls1 <- hash_map_get_mut_in_list_back T n0 key ls0 ret; + Return (ListCons ckey cvalue ls1)) + | ListNil => Fail_ Failure + end + end + . + +(** [hashmap::HashMap::{0}::get_mut] *) +Definition hash_map_get_mut_fwd + (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) : result T := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hash <- hash_key_fwd key; + match self with + | mkHash_map_t i p i0 v => + let i1 := vec_len (List_t T) v in + hash_mod <- usize_rem hash i1; + l <- vec_index_mut_fwd (List_t T) v hash_mod; + t <- hash_map_get_mut_in_list_fwd T n0 key l; + Return t + end + end + . + +(** [hashmap::HashMap::{0}::get_mut] *) +Definition hash_map_get_mut_back + (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) (ret : T) : + result (Hash_map_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hash <- hash_key_fwd key; + match self with + | mkHash_map_t i p i0 v => + let i1 := vec_len (List_t T) v in + hash_mod <- usize_rem hash i1; + l <- vec_index_mut_fwd (List_t T) v hash_mod; + l0 <- hash_map_get_mut_in_list_back T n0 key l ret; + v0 <- vec_index_mut_back (List_t T) v hash_mod l0; + Return (mkHash_map_t i p i0 v0) + end + end + . + +(** [hashmap::HashMap::{0}::remove_from_list] *) +Fixpoint hash_map_remove_from_list_fwd + (T : Type) (n : nat) (key : usize) (ls : List_t T) : result (option T) := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | ListCons ckey t tl => + if ckey s= key + then + let mv_ls := mem_replace_fwd (List_t T) (ListCons ckey t tl) ListNil in + match mv_ls with + | ListCons i cvalue tl0 => Return (Some cvalue) + | ListNil => Fail_ Failure + end + else (opt <- hash_map_remove_from_list_fwd T n0 key tl; Return opt) + | ListNil => Return None + end + end + . + +(** [hashmap::HashMap::{0}::remove_from_list] *) +Fixpoint hash_map_remove_from_list_back + (T : Type) (n : nat) (key : usize) (ls : List_t T) : result (List_t T) := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | ListCons ckey t tl => + if ckey s= key + then + let mv_ls := mem_replace_fwd (List_t T) (ListCons ckey t tl) ListNil in + match mv_ls with + | ListCons i cvalue tl0 => Return tl0 + | ListNil => Fail_ Failure + end + else ( + tl0 <- hash_map_remove_from_list_back T n0 key tl; + Return (ListCons ckey t tl0)) + | ListNil => Return ListNil + end + end + . + +(** [hashmap::HashMap::{0}::remove] *) +Definition hash_map_remove_fwd + (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) : + result (option T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hash <- hash_key_fwd key; + match self with + | mkHash_map_t i p i0 v => + let i1 := vec_len (List_t T) v in + hash_mod <- usize_rem hash i1; + l <- vec_index_mut_fwd (List_t T) v hash_mod; + x <- hash_map_remove_from_list_fwd T n0 key l; + match x with + | None => Return None + | Some x0 => i2 <- usize_sub i 1 %usize; let _ := i2 in Return (Some x0) + end + end + end + . + +(** [hashmap::HashMap::{0}::remove] *) +Definition hash_map_remove_back + (T : Type) (n : nat) (self : Hash_map_t T) (key : usize) : + result (Hash_map_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hash <- hash_key_fwd key; + match self with + | mkHash_map_t i p i0 v => + let i1 := vec_len (List_t T) v in + hash_mod <- usize_rem hash i1; + l <- vec_index_mut_fwd (List_t T) v hash_mod; + x <- hash_map_remove_from_list_fwd T n0 key l; + match x with + | None => + l0 <- hash_map_remove_from_list_back T n0 key l; + v0 <- vec_index_mut_back (List_t T) v hash_mod l0; + Return (mkHash_map_t i p i0 v0) + | Some x0 => + i2 <- usize_sub i 1 %usize; + l0 <- hash_map_remove_from_list_back T n0 key l; + v0 <- vec_index_mut_back (List_t T) v hash_mod l0; + Return (mkHash_map_t i2 p i0 v0) + end + end + end + . + +(** [hashmap::test1] *) +Definition test1_fwd (n : nat) : result unit := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hm <- hash_map_new_fwd u64 n0; + hm0 <- hash_map_insert_fwd_back u64 n0 hm (0 %usize) (42 %u64); + hm1 <- hash_map_insert_fwd_back u64 n0 hm0 (128 %usize) (18 %u64); + hm2 <- hash_map_insert_fwd_back u64 n0 hm1 (1024 %usize) (138 %u64); + hm3 <- hash_map_insert_fwd_back u64 n0 hm2 (1056 %usize) (256 %u64); + i <- hash_map_get_fwd u64 n0 hm3 (128 %usize); + if negb (i s= 18 %u64) + then Fail_ Failure + else ( + hm4 <- hash_map_get_mut_back u64 n0 hm3 (1024 %usize) (56 %u64); + i0 <- hash_map_get_fwd u64 n0 hm4 (1024 %usize); + if negb (i0 s= 56 %u64) + then Fail_ Failure + else ( + x <- hash_map_remove_fwd u64 n0 hm4 (1024 %usize); + match x with + | None => Fail_ Failure + | Some x0 => + if negb (x0 s= 56 %u64) + then Fail_ Failure + else ( + hm5 <- hash_map_remove_back u64 n0 hm4 (1024 %usize); + i1 <- hash_map_get_fwd u64 n0 hm5 (0 %usize); + if negb (i1 s= 42 %u64) + then Fail_ Failure + else ( + i2 <- hash_map_get_fwd u64 n0 hm5 (128 %usize); + if negb (i2 s= 18 %u64) + then Fail_ Failure + else ( + i3 <- hash_map_get_fwd u64 n0 hm5 (1056 %usize); + if negb (i3 s= 256 %u64) then Fail_ Failure else Return tt))) + end)) + end + . + +End Hashmap__Funs . diff --git a/tests/coq/hashmap/Hashmap__Types.v b/tests/coq/hashmap/Hashmap__Types.v new file mode 100644 index 00000000..b665179e --- /dev/null +++ b/tests/coq/hashmap/Hashmap__Types.v @@ -0,0 +1,39 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap]: type definitions *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Local Open Scope Primitives_scope. +Module Hashmap__Types . + +(** [hashmap::List] *) +Inductive List_t (T : Type) := +| ListCons : usize -> T -> List_t T -> List_t T +| ListNil : List_t T +. + +Arguments ListCons {T} _ _ _ . +Arguments ListNil {T} . + +(** [hashmap::HashMap] *) +Record Hash_map_t (T : Type) := +mkHash_map_t +{ + Hash_map_num_entries : usize; + Hash_map_max_load_factor : (usize * usize); + Hash_map_max_load : usize; + Hash_map_slots : vec (List_t T); +} +. + +Arguments mkHash_map_t {T} _ _ _ _ . +Arguments Hash_map_num_entries {T} . +Arguments Hash_map_max_load_factor {T} . +Arguments Hash_map_max_load {T} . +Arguments Hash_map_slots {T} . + +(** [core::num::u32::{9}::MAX] *) +Definition core_num_u32_max_body : result u32 := Return (4294967295 %u32) . +Definition core_num_u32_max_c : u32 := core_num_u32_max_body%global . + +End Hashmap__Types . diff --git a/tests/coq/hashmap/Makefile b/tests/coq/hashmap/Makefile new file mode 100644 index 00000000..ff1ccd39 --- /dev/null +++ b/tests/coq/hashmap/Makefile @@ -0,0 +1,22 @@ +# Makefile originally taken from coq-club + +%: Makefile.coq phony + +make -f Makefile.coq $@ + +all: Makefile.coq + +make -f Makefile.coq all + +clean: Makefile.coq + +make -f Makefile.coq clean + rm -f Makefile.coq + +Makefile.coq: _CoqProject Makefile + coq_makefile -f _CoqProject | sed 's/$$(COQCHK) $$(COQCHKFLAGS) $$(COQLIBS)/$$(COQCHK) $$(COQCHKFLAGS) $$(subst -Q,-R,$$(COQLIBS))/' > Makefile.coq + +_CoqProject: ; + +Makefile: ; + +phony: ; + +.PHONY: all clean phony diff --git a/tests/coq/hashmap/Primitives.v b/tests/coq/hashmap/Primitives.v new file mode 100644 index 00000000..9a97d6c7 --- /dev/null +++ b/tests/coq/hashmap/Primitives.v @@ -0,0 +1,482 @@ +Require Import Lia. +Require Coq.Strings.Ascii. +Require Coq.Strings.String. +Require Import Coq.Program.Equality. +Require Import Coq.ZArith.ZArith. +Require Import Coq.ZArith.Znat. +Require Import List. +Import ListNotations. + +Module Primitives. + + (* TODO: use more *) +Declare Scope Primitives_scope. + +(*** Result *) + +Inductive error := + | Failure + | OutOfFuel. + +Inductive result A := + | Return : A -> result A + | Fail_ : error -> result A. + +Arguments Return {_} a. +Arguments Fail_ {_}. + +Definition bind {A B} (m: result A) (f: A -> result B) : result B := + match m with + | Fail_ e => Fail_ e + | Return x => f x + end. + +Definition return_ {A: Type} (x: A) : result A := Return x. +Definition fail_ {A: Type} (e: error) : result A := Fail_ e. + +Notation "x <- c1 ; c2" := (bind c1 (fun x => c2)) + (at level 61, c1 at next level, right associativity). + +(** Monadic assert *) +Definition massert (b: bool) : result unit := + if b then Return tt else Fail_ Failure. + +(** Normalize and unwrap a successful result (used for globals) *) +Definition eval_result_refl {A} {x} (a: result A) (p: a = Return x) : A := + match a as r return (r = Return x -> A) with + | Return a' => fun _ => a' + | Fail_ e => fun p' => + False_rect _ (eq_ind (Fail_ e) + (fun e : result A => + match e with + | Return _ => False + | Fail_ e => True + end) + I (Return x) p') + end p. + +Notation "x %global" := (eval_result_refl x eq_refl) (at level 40). +Notation "x %return" := (eval_result_refl x eq_refl) (at level 40). + +(* Sanity check *) +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 . + +(*** Scalars *) + +Definition i8_min : Z := -128%Z. +Definition i8_max : Z := 127%Z. +Definition i16_min : Z := -32768%Z. +Definition i16_max : Z := 32767%Z. +Definition i32_min : Z := -2147483648%Z. +Definition i32_max : Z := 2147483647%Z. +Definition i64_min : Z := -9223372036854775808%Z. +Definition i64_max : Z := 9223372036854775807%Z. +Definition i128_min : Z := -170141183460469231731687303715884105728%Z. +Definition i128_max : Z := 170141183460469231731687303715884105727%Z. +Definition u8_min : Z := 0%Z. +Definition u8_max : Z := 255%Z. +Definition u16_min : Z := 0%Z. +Definition u16_max : Z := 65535%Z. +Definition u32_min : Z := 0%Z. +Definition u32_max : Z := 4294967295%Z. +Definition u64_min : Z := 0%Z. +Definition u64_max : Z := 18446744073709551615%Z. +Definition u128_min : Z := 0%Z. +Definition u128_max : Z := 340282366920938463463374607431768211455%Z. + +(** The bounds of [isize] and [usize] vary with the architecture. *) +Axiom isize_min : Z. +Axiom isize_max : Z. +Definition usize_min : Z := 0%Z. +Axiom usize_max : Z. + +Open Scope Z_scope. + +(** We provide those lemmas to reason about the bounds of [isize] and [usize] *) +Axiom isize_min_bound : isize_min <= i32_min. +Axiom isize_max_bound : i32_max <= isize_max. +Axiom usize_max_bound : u32_max <= usize_max. + +Inductive scalar_ty := + | Isize + | I8 + | I16 + | I32 + | I64 + | I128 + | Usize + | U8 + | U16 + | U32 + | U64 + | U128 +. + +Definition scalar_min (ty: scalar_ty) : Z := + match ty with + | Isize => isize_min + | I8 => i8_min + | I16 => i16_min + | I32 => i32_min + | I64 => i64_min + | I128 => i128_min + | Usize => usize_min + | U8 => u8_min + | U16 => u16_min + | U32 => u32_min + | U64 => u64_min + | U128 => u128_min +end. + +Definition scalar_max (ty: scalar_ty) : Z := + match ty with + | Isize => isize_max + | I8 => i8_max + | I16 => i16_max + | I32 => i32_max + | I64 => i64_max + | I128 => i128_max + | Usize => usize_max + | U8 => u8_max + | U16 => u16_max + | U32 => u32_max + | U64 => u64_max + | U128 => u128_max +end. + +(** We use the following conservative bounds to make sure we can compute bound + checks in most situations *) +Definition scalar_min_cons (ty: scalar_ty) : Z := + match ty with + | Isize => i32_min + | Usize => u32_min + | _ => scalar_min ty +end. + +Definition scalar_max_cons (ty: scalar_ty) : Z := + match ty with + | Isize => i32_max + | Usize => u32_max + | _ => scalar_max ty +end. + +Lemma scalar_min_cons_valid : forall ty, scalar_min ty <= scalar_min_cons ty . +Proof. + destruct ty; unfold scalar_min_cons, scalar_min; try lia. + - pose isize_min_bound; lia. + - apply Z.le_refl. +Qed. + +Lemma scalar_max_cons_valid : forall ty, scalar_max ty >= scalar_max_cons ty . +Proof. + destruct ty; unfold scalar_max_cons, scalar_max; try lia. + - pose isize_max_bound; lia. + - pose usize_max_bound. lia. +Qed. + +Definition scalar (ty: scalar_ty) : Type := + { x: Z | scalar_min ty <= x <= scalar_max ty }. + +Definition to_Z {ty} (x: scalar ty) : Z := proj1_sig x. + +(** Bounds checks: we start by using the conservative bounds, to make sure we + can compute in most situations, then we use the real bounds (for [isize] + and [usize]). *) +Definition scalar_ge_min (ty: scalar_ty) (x: Z) : bool := + Z.leb (scalar_min_cons ty) x || Z.leb (scalar_min ty) x. + +Definition scalar_le_max (ty: scalar_ty) (x: Z) : bool := + Z.leb x (scalar_max_cons ty) || Z.leb x (scalar_max ty). + +Lemma scalar_ge_min_valid (ty: scalar_ty) (x: Z) : + scalar_ge_min ty x = true -> scalar_min ty <= x . +Proof. + unfold scalar_ge_min. + pose (scalar_min_cons_valid ty). + lia. +Qed. + +Lemma scalar_le_max_valid (ty: scalar_ty) (x: Z) : + scalar_le_max ty x = true -> x <= scalar_max ty . +Proof. + unfold scalar_le_max. + pose (scalar_max_cons_valid ty). + lia. +Qed. + +Definition scalar_in_bounds (ty: scalar_ty) (x: Z) : bool := + scalar_ge_min ty x && scalar_le_max ty x . + +Lemma scalar_in_bounds_valid (ty: scalar_ty) (x: Z) : + scalar_in_bounds ty x = true -> scalar_min ty <= x <= scalar_max ty . +Proof. + unfold scalar_in_bounds. + intros H. + destruct (scalar_ge_min ty x) eqn:Hmin. + - destruct (scalar_le_max ty x) eqn:Hmax. + + pose (scalar_ge_min_valid ty x Hmin). + pose (scalar_le_max_valid ty x Hmax). + lia. + + inversion H. + - inversion H. +Qed. + +Import Sumbool. + +Definition mk_scalar (ty: scalar_ty) (x: Z) : result (scalar ty) := + match sumbool_of_bool (scalar_in_bounds ty x) with + | left H => Return (exist _ x (scalar_in_bounds_valid _ _ H)) + | right _ => Fail_ Failure + end. + +Definition scalar_add {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x + to_Z y). + +Definition scalar_sub {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x - to_Z y). + +Definition scalar_mul {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x * to_Z y). + +Definition scalar_div {ty} (x y: scalar ty) : result (scalar ty) := + if to_Z y =? 0 then Fail_ Failure else + mk_scalar ty (to_Z x / to_Z y). + +Definition scalar_rem {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (Z.rem (to_Z x) (to_Z y)). + +Definition scalar_neg {ty} (x: scalar ty) : result (scalar ty) := mk_scalar ty (-(to_Z x)). + +(** Cast an integer from a [src_ty] to a [tgt_ty] *) +(* TODO: check the semantics of casts in Rust *) +Definition scalar_cast (src_ty tgt_ty : scalar_ty) (x : scalar src_ty) : result (scalar tgt_ty) := + mk_scalar tgt_ty (to_Z x). + +(** Comparisons *) +Print Z.leb . + +Definition scalar_leb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.leb (to_Z x) (to_Z y) . + +Definition scalar_ltb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.ltb (to_Z x) (to_Z y) . + +Definition scalar_geb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.geb (to_Z x) (to_Z y) . + +Definition scalar_gtb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.gtb (to_Z x) (to_Z y) . + +Definition scalar_eqb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.eqb (to_Z x) (to_Z y) . + +Definition scalar_neqb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + negb (Z.eqb (to_Z x) (to_Z y)) . + + +(** The scalar types *) +Definition isize := scalar Isize. +Definition i8 := scalar I8. +Definition i16 := scalar I16. +Definition i32 := scalar I32. +Definition i64 := scalar I64. +Definition i128 := scalar I128. +Definition usize := scalar Usize. +Definition u8 := scalar U8. +Definition u16 := scalar U16. +Definition u32 := scalar U32. +Definition u64 := scalar U64. +Definition u128 := scalar U128. + +(** Negaion *) +Definition isize_neg := @scalar_neg Isize. +Definition i8_neg := @scalar_neg I8. +Definition i16_neg := @scalar_neg I16. +Definition i32_neg := @scalar_neg I32. +Definition i64_neg := @scalar_neg I64. +Definition i128_neg := @scalar_neg I128. + +(** Division *) +Definition isize_div := @scalar_div Isize. +Definition i8_div := @scalar_div I8. +Definition i16_div := @scalar_div I16. +Definition i32_div := @scalar_div I32. +Definition i64_div := @scalar_div I64. +Definition i128_div := @scalar_div I128. +Definition usize_div := @scalar_div Usize. +Definition u8_div := @scalar_div U8. +Definition u16_div := @scalar_div U16. +Definition u32_div := @scalar_div U32. +Definition u64_div := @scalar_div U64. +Definition u128_div := @scalar_div U128. + +(** Remainder *) +Definition isize_rem := @scalar_rem Isize. +Definition i8_rem := @scalar_rem I8. +Definition i16_rem := @scalar_rem I16. +Definition i32_rem := @scalar_rem I32. +Definition i64_rem := @scalar_rem I64. +Definition i128_rem := @scalar_rem I128. +Definition usize_rem := @scalar_rem Usize. +Definition u8_rem := @scalar_rem U8. +Definition u16_rem := @scalar_rem U16. +Definition u32_rem := @scalar_rem U32. +Definition u64_rem := @scalar_rem U64. +Definition u128_rem := @scalar_rem U128. + +(** Addition *) +Definition isize_add := @scalar_add Isize. +Definition i8_add := @scalar_add I8. +Definition i16_add := @scalar_add I16. +Definition i32_add := @scalar_add I32. +Definition i64_add := @scalar_add I64. +Definition i128_add := @scalar_add I128. +Definition usize_add := @scalar_add Usize. +Definition u8_add := @scalar_add U8. +Definition u16_add := @scalar_add U16. +Definition u32_add := @scalar_add U32. +Definition u64_add := @scalar_add U64. +Definition u128_add := @scalar_add U128. + +(** Substraction *) +Definition isize_sub := @scalar_sub Isize. +Definition i8_sub := @scalar_sub I8. +Definition i16_sub := @scalar_sub I16. +Definition i32_sub := @scalar_sub I32. +Definition i64_sub := @scalar_sub I64. +Definition i128_sub := @scalar_sub I128. +Definition usize_sub := @scalar_sub Usize. +Definition u8_sub := @scalar_sub U8. +Definition u16_sub := @scalar_sub U16. +Definition u32_sub := @scalar_sub U32. +Definition u64_sub := @scalar_sub U64. +Definition u128_sub := @scalar_sub U128. + +(** Multiplication *) +Definition isize_mul := @scalar_mul Isize. +Definition i8_mul := @scalar_mul I8. +Definition i16_mul := @scalar_mul I16. +Definition i32_mul := @scalar_mul I32. +Definition i64_mul := @scalar_mul I64. +Definition i128_mul := @scalar_mul I128. +Definition usize_mul := @scalar_mul Usize. +Definition u8_mul := @scalar_mul U8. +Definition u16_mul := @scalar_mul U16. +Definition u32_mul := @scalar_mul U32. +Definition u64_mul := @scalar_mul U64. +Definition u128_mul := @scalar_mul U128. + +(** Small utility *) +Definition usize_to_nat (x: usize) : nat := Z.to_nat (to_Z x). + +(** Notations *) +Notation "x %isize" := ((mk_scalar Isize x)%return) (at level 9). +Notation "x %i8" := ((mk_scalar I8 x)%return) (at level 9). +Notation "x %i16" := ((mk_scalar I16 x)%return) (at level 9). +Notation "x %i32" := ((mk_scalar I32 x)%return) (at level 9). +Notation "x %i64" := ((mk_scalar I64 x)%return) (at level 9). +Notation "x %i128" := ((mk_scalar I128 x)%return) (at level 9). +Notation "x %usize" := ((mk_scalar Usize x)%return) (at level 9). +Notation "x %u8" := ((mk_scalar U8 x)%return) (at level 9). +Notation "x %u16" := ((mk_scalar U16 x)%return) (at level 9). +Notation "x %u32" := ((mk_scalar U32 x)%return) (at level 9). +Notation "x %u64" := ((mk_scalar U64 x)%return) (at level 9). +Notation "x %u128" := ((mk_scalar U128 x)%return) (at level 9). + +Notation "x s= y" := (scalar_eqb x y) (at level 80) : Primitives_scope. +Notation "x s<> y" := (scalar_neqb x y) (at level 80) : Primitives_scope. +Notation "x s<= y" := (scalar_leb x y) (at level 80) : Primitives_scope. +Notation "x s< y" := (scalar_ltb x y) (at level 80) : Primitives_scope. +Notation "x s>= y" := (scalar_geb x y) (at level 80) : Primitives_scope. +Notation "x s> y" := (scalar_gtb x y) (at level 80) : Primitives_scope. + +(*** Vectors *) + +Definition 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 vec_length {T: Type} (v: vec T) : Z := Z.of_nat (length (vec_to_list v)). + +Lemma le_0_usize_max : 0 <= usize_max. +Proof. + pose (H := usize_max_bound). + unfold u32_max in H. + lia. +Qed. + +Definition vec_new (T: Type) : vec T := (exist _ [] le_0_usize_max). + +Lemma vec_len_in_usize {T} (v: vec T) : usize_min <= vec_length v <= usize_max. +Proof. + unfold 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). + +Fixpoint list_update {A} (l: list A) (n: nat) (a: A) + : list A := + match l with + | [] => [] + | x :: t => match n with + | 0%nat => a :: t + | 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) ; + 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 vec_push_back (T: Type) (v: vec T) (x: T) : result (vec T) := + 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 vec_insert_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). + +(* 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 + 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). + +End Primitives. diff --git a/tests/coq/hashmap/_CoqProject b/tests/coq/hashmap/_CoqProject new file mode 100644 index 00000000..2aa6b46c --- /dev/null +++ b/tests/coq/hashmap/_CoqProject @@ -0,0 +1,8 @@ +-R . Lib +-arg -w +-arg all + +Primitives.v + +Hashmap__Types.v +Hashmap__Funs.v
\ No newline at end of file diff --git a/tests/coq/hashmap_on_disk/HashmapMain__Funs.v b/tests/coq/hashmap_on_disk/HashmapMain__Funs.v new file mode 100644 index 00000000..c390ac1d --- /dev/null +++ b/tests/coq/hashmap_on_disk/HashmapMain__Funs.v @@ -0,0 +1,596 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: function definitions *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Local Open Scope Primitives_scope. +Require Export HashmapMain__Types . +Import HashmapMain__Types . +Require Export HashmapMain__Opaque . +Import HashmapMain__Opaque . +Module HashmapMain__Funs . + +(** [hashmap_main::hashmap::hash_key] *) +Definition hashmap_hash_key_fwd (k : usize) : result usize := Return k . + +(** [hashmap_main::hashmap::HashMap::{0}::allocate_slots] *) +Fixpoint hashmap_hash_map_allocate_slots_fwd + (T : Type) (n : nat) (slots : vec (Hashmap_list_t T)) (n0 : usize) : + result (vec (Hashmap_list_t T)) + := + match n with + | O => Fail_ OutOfFuel + | S n1 => + if n0 s= 0 %usize + then Return slots + else ( + slots0 <- vec_push_back (Hashmap_list_t T) slots HashmapListNil; + i <- usize_sub n0 1 %usize; + v <- hashmap_hash_map_allocate_slots_fwd T n1 slots0 i; + Return v) + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::new_with_capacity] *) +Definition hashmap_hash_map_new_with_capacity_fwd + (T : Type) (n : nat) (capacity : usize) (max_load_dividend : usize) + (max_load_divisor : usize) : + result (Hashmap_hash_map_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + let v := vec_new (Hashmap_list_t T) in + slots <- hashmap_hash_map_allocate_slots_fwd T n0 v capacity; + i <- usize_mul capacity max_load_dividend; + i0 <- usize_div i max_load_divisor; + Return (mkHashmap_hash_map_t (0 %usize) (max_load_dividend, + max_load_divisor) i0 slots) + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::new] *) +Definition hashmap_hash_map_new_fwd + (T : Type) (n : nat) : result (Hashmap_hash_map_t T) := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hm <- + hashmap_hash_map_new_with_capacity_fwd T n0 (32 %usize) (4 %usize) (5 + %usize); + Return hm + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::clear_slots] *) +Fixpoint hashmap_hash_map_clear_slots_fwd_back + (T : Type) (n : nat) (slots : vec (Hashmap_list_t T)) (i : usize) : + result (vec (Hashmap_list_t T)) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + let i0 := vec_len (Hashmap_list_t T) slots in + if i s< i0 + then ( + slots0 <- vec_index_mut_back (Hashmap_list_t T) slots i HashmapListNil; + i1 <- usize_add i 1 %usize; + slots1 <- hashmap_hash_map_clear_slots_fwd_back T n0 slots0 i1; + Return slots1) + else Return slots + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::clear] *) +Definition hashmap_hash_map_clear_fwd_back + (T : Type) (n : nat) (self : Hashmap_hash_map_t T) : + result (Hashmap_hash_map_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match self with + | mkHashmap_hash_map_t i p i0 v => + v0 <- hashmap_hash_map_clear_slots_fwd_back T n0 v (0 %usize); + Return (mkHashmap_hash_map_t (0 %usize) p i0 v0) + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::len] *) +Definition hashmap_hash_map_len_fwd + (T : Type) (self : Hashmap_hash_map_t T) : result usize := + match self with | mkHashmap_hash_map_t i p i0 v => Return i end . + +(** [hashmap_main::hashmap::HashMap::{0}::insert_in_list] *) +Fixpoint hashmap_hash_map_insert_in_list_fwd + (T : Type) (n : nat) (key : usize) (value : T) (ls : Hashmap_list_t T) : + result bool + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | HashmapListCons ckey cvalue ls0 => + if ckey s= key + then Return false + else ( + b <- hashmap_hash_map_insert_in_list_fwd T n0 key value ls0; Return b) + | HashmapListNil => Return true + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::insert_in_list] *) +Fixpoint hashmap_hash_map_insert_in_list_back + (T : Type) (n : nat) (key : usize) (value : T) (ls : Hashmap_list_t T) : + result (Hashmap_list_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | HashmapListCons ckey cvalue ls0 => + if ckey s= key + then Return (HashmapListCons ckey value ls0) + else ( + ls1 <- hashmap_hash_map_insert_in_list_back T n0 key value ls0; + Return (HashmapListCons ckey cvalue ls1)) + | HashmapListNil => + let l := HashmapListNil in Return (HashmapListCons key value l) + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::insert_no_resize] *) +Definition hashmap_hash_map_insert_no_resize_fwd_back + (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) (value : T) : + result (Hashmap_hash_map_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hash <- hashmap_hash_key_fwd key; + match self with + | mkHashmap_hash_map_t i p i0 v => + let i1 := vec_len (Hashmap_list_t T) v in + hash_mod <- usize_rem hash i1; + l <- vec_index_mut_fwd (Hashmap_list_t T) v hash_mod; + inserted <- hashmap_hash_map_insert_in_list_fwd T n0 key value l; + if inserted + then ( + i2 <- usize_add i 1 %usize; + l0 <- hashmap_hash_map_insert_in_list_back T n0 key value l; + v0 <- vec_index_mut_back (Hashmap_list_t T) v hash_mod l0; + Return (mkHashmap_hash_map_t i2 p i0 v0)) + else ( + l0 <- hashmap_hash_map_insert_in_list_back T n0 key value l; + v0 <- vec_index_mut_back (Hashmap_list_t T) v hash_mod l0; + Return (mkHashmap_hash_map_t i p i0 v0)) + end + end + . + +(** [core::num::u32::{9}::MAX] *) +Definition core_num_u32_max_body : result u32 := Return (4294967295 %u32) . +Definition core_num_u32_max_c : u32 := core_num_u32_max_body%global . + +(** [hashmap_main::hashmap::HashMap::{0}::move_elements_from_list] *) +Fixpoint hashmap_hash_map_move_elements_from_list_fwd_back + (T : Type) (n : nat) (ntable : Hashmap_hash_map_t T) (ls : Hashmap_list_t T) : + result (Hashmap_hash_map_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | HashmapListCons k v tl => + ntable0 <- hashmap_hash_map_insert_no_resize_fwd_back T n0 ntable k v; + ntable1 <- + hashmap_hash_map_move_elements_from_list_fwd_back T n0 ntable0 tl; + Return ntable1 + | HashmapListNil => Return ntable + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::move_elements] *) +Fixpoint hashmap_hash_map_move_elements_fwd_back + (T : Type) (n : nat) (ntable : Hashmap_hash_map_t T) + (slots : vec (Hashmap_list_t T)) (i : usize) : + result ((Hashmap_hash_map_t T) * (vec (Hashmap_list_t T))) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + let i0 := vec_len (Hashmap_list_t T) slots in + if i s< i0 + then ( + l <- vec_index_mut_fwd (Hashmap_list_t T) slots i; + let ls := mem_replace_fwd (Hashmap_list_t T) l HashmapListNil in + ntable0 <- + hashmap_hash_map_move_elements_from_list_fwd_back T n0 ntable ls; + let l0 := mem_replace_back (Hashmap_list_t T) l HashmapListNil in + slots0 <- vec_index_mut_back (Hashmap_list_t T) slots i l0; + i1 <- usize_add i 1 %usize; + p <- hashmap_hash_map_move_elements_fwd_back T n0 ntable0 slots0 i1; + let (ntable1, slots1) := p in + Return (ntable1, slots1)) + else Return (ntable, slots) + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::try_resize] *) +Definition hashmap_hash_map_try_resize_fwd_back + (T : Type) (n : nat) (self : Hashmap_hash_map_t T) : + result (Hashmap_hash_map_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + max_usize <- scalar_cast U32 Usize core_num_u32_max_c; + match self with + | mkHashmap_hash_map_t i p i0 v => + let capacity := vec_len (Hashmap_list_t T) v in + n1 <- usize_div max_usize 2 %usize; + let (i1, i2) := p in + i3 <- usize_div n1 i1; + if capacity s<= i3 + then ( + i4 <- usize_mul capacity 2 %usize; + ntable <- hashmap_hash_map_new_with_capacity_fwd T n0 i4 i1 i2; + p0 <- hashmap_hash_map_move_elements_fwd_back T n0 ntable v (0 %usize); + let (ntable0, _) := p0 in + match ntable0 with + | mkHashmap_hash_map_t i5 p1 i6 v0 => + Return (mkHashmap_hash_map_t i (i1, i2) i6 v0) + end) + else Return (mkHashmap_hash_map_t i (i1, i2) i0 v) + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::insert] *) +Definition hashmap_hash_map_insert_fwd_back + (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) (value : T) : + result (Hashmap_hash_map_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + self0 <- hashmap_hash_map_insert_no_resize_fwd_back T n0 self key value; + i <- hashmap_hash_map_len_fwd T self0; + match self0 with + | mkHashmap_hash_map_t i0 p i1 v => + if i s> i1 + then ( + self1 <- + hashmap_hash_map_try_resize_fwd_back T n0 (mkHashmap_hash_map_t i0 p + i1 v); + Return self1) + else Return (mkHashmap_hash_map_t i0 p i1 v) + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::contains_key_in_list] *) +Fixpoint hashmap_hash_map_contains_key_in_list_fwd + (T : Type) (n : nat) (key : usize) (ls : Hashmap_list_t T) : result bool := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | HashmapListCons ckey t ls0 => + if ckey s= key + then Return true + else ( + b <- hashmap_hash_map_contains_key_in_list_fwd T n0 key ls0; Return b) + | HashmapListNil => Return false + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::contains_key] *) +Definition hashmap_hash_map_contains_key_fwd + (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) : + result bool + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hash <- hashmap_hash_key_fwd key; + match self with + | mkHashmap_hash_map_t i p i0 v => + let i1 := vec_len (Hashmap_list_t T) v in + hash_mod <- usize_rem hash i1; + l <- vec_index_fwd (Hashmap_list_t T) v hash_mod; + b <- hashmap_hash_map_contains_key_in_list_fwd T n0 key l; + Return b + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::get_in_list] *) +Fixpoint hashmap_hash_map_get_in_list_fwd + (T : Type) (n : nat) (key : usize) (ls : Hashmap_list_t T) : result T := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | HashmapListCons ckey cvalue ls0 => + if ckey s= key + then Return cvalue + else (t <- hashmap_hash_map_get_in_list_fwd T n0 key ls0; Return t) + | HashmapListNil => Fail_ Failure + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::get] *) +Definition hashmap_hash_map_get_fwd + (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) : + result T + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hash <- hashmap_hash_key_fwd key; + match self with + | mkHashmap_hash_map_t i p i0 v => + let i1 := vec_len (Hashmap_list_t T) v in + hash_mod <- usize_rem hash i1; + l <- vec_index_fwd (Hashmap_list_t T) v hash_mod; + t <- hashmap_hash_map_get_in_list_fwd T n0 key l; + Return t + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list] *) +Fixpoint hashmap_hash_map_get_mut_in_list_fwd + (T : Type) (n : nat) (key : usize) (ls : Hashmap_list_t T) : result T := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | HashmapListCons ckey cvalue ls0 => + if ckey s= key + then Return cvalue + else (t <- hashmap_hash_map_get_mut_in_list_fwd T n0 key ls0; Return t) + | HashmapListNil => Fail_ Failure + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::get_mut_in_list] *) +Fixpoint hashmap_hash_map_get_mut_in_list_back + (T : Type) (n : nat) (key : usize) (ls : Hashmap_list_t T) (ret : T) : + result (Hashmap_list_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | HashmapListCons ckey cvalue ls0 => + if ckey s= key + then Return (HashmapListCons ckey ret ls0) + else ( + ls1 <- hashmap_hash_map_get_mut_in_list_back T n0 key ls0 ret; + Return (HashmapListCons ckey cvalue ls1)) + | HashmapListNil => Fail_ Failure + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::get_mut] *) +Definition hashmap_hash_map_get_mut_fwd + (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) : + result T + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hash <- hashmap_hash_key_fwd key; + match self with + | mkHashmap_hash_map_t i p i0 v => + let i1 := vec_len (Hashmap_list_t T) v in + hash_mod <- usize_rem hash i1; + l <- vec_index_mut_fwd (Hashmap_list_t T) v hash_mod; + t <- hashmap_hash_map_get_mut_in_list_fwd T n0 key l; + Return t + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::get_mut] *) +Definition hashmap_hash_map_get_mut_back + (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) (ret : T) : + result (Hashmap_hash_map_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hash <- hashmap_hash_key_fwd key; + match self with + | mkHashmap_hash_map_t i p i0 v => + let i1 := vec_len (Hashmap_list_t T) v in + hash_mod <- usize_rem hash i1; + l <- vec_index_mut_fwd (Hashmap_list_t T) v hash_mod; + l0 <- hashmap_hash_map_get_mut_in_list_back T n0 key l ret; + v0 <- vec_index_mut_back (Hashmap_list_t T) v hash_mod l0; + Return (mkHashmap_hash_map_t i p i0 v0) + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::remove_from_list] *) +Fixpoint hashmap_hash_map_remove_from_list_fwd + (T : Type) (n : nat) (key : usize) (ls : Hashmap_list_t T) : + result (option T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | HashmapListCons ckey t tl => + if ckey s= key + then + let mv_ls := + mem_replace_fwd (Hashmap_list_t T) (HashmapListCons ckey t tl) + HashmapListNil in + match mv_ls with + | HashmapListCons i cvalue tl0 => Return (Some cvalue) + | HashmapListNil => Fail_ Failure + end + else ( + opt <- hashmap_hash_map_remove_from_list_fwd T n0 key tl; Return opt) + | HashmapListNil => Return None + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::remove_from_list] *) +Fixpoint hashmap_hash_map_remove_from_list_back + (T : Type) (n : nat) (key : usize) (ls : Hashmap_list_t T) : + result (Hashmap_list_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + match ls with + | HashmapListCons ckey t tl => + if ckey s= key + then + let mv_ls := + mem_replace_fwd (Hashmap_list_t T) (HashmapListCons ckey t tl) + HashmapListNil in + match mv_ls with + | HashmapListCons i cvalue tl0 => Return tl0 + | HashmapListNil => Fail_ Failure + end + else ( + tl0 <- hashmap_hash_map_remove_from_list_back T n0 key tl; + Return (HashmapListCons ckey t tl0)) + | HashmapListNil => Return HashmapListNil + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::remove] *) +Definition hashmap_hash_map_remove_fwd + (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) : + result (option T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hash <- hashmap_hash_key_fwd key; + match self with + | mkHashmap_hash_map_t i p i0 v => + let i1 := vec_len (Hashmap_list_t T) v in + hash_mod <- usize_rem hash i1; + l <- vec_index_mut_fwd (Hashmap_list_t T) v hash_mod; + x <- hashmap_hash_map_remove_from_list_fwd T n0 key l; + match x with + | None => Return None + | Some x0 => i2 <- usize_sub i 1 %usize; let _ := i2 in Return (Some x0) + end + end + end + . + +(** [hashmap_main::hashmap::HashMap::{0}::remove] *) +Definition hashmap_hash_map_remove_back + (T : Type) (n : nat) (self : Hashmap_hash_map_t T) (key : usize) : + result (Hashmap_hash_map_t T) + := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hash <- hashmap_hash_key_fwd key; + match self with + | mkHashmap_hash_map_t i p i0 v => + let i1 := vec_len (Hashmap_list_t T) v in + hash_mod <- usize_rem hash i1; + l <- vec_index_mut_fwd (Hashmap_list_t T) v hash_mod; + x <- hashmap_hash_map_remove_from_list_fwd T n0 key l; + match x with + | None => + l0 <- hashmap_hash_map_remove_from_list_back T n0 key l; + v0 <- vec_index_mut_back (Hashmap_list_t T) v hash_mod l0; + Return (mkHashmap_hash_map_t i p i0 v0) + | Some x0 => + i2 <- usize_sub i 1 %usize; + l0 <- hashmap_hash_map_remove_from_list_back T n0 key l; + v0 <- vec_index_mut_back (Hashmap_list_t T) v hash_mod l0; + Return (mkHashmap_hash_map_t i2 p i0 v0) + end + end + end + . + +(** [hashmap_main::hashmap::test1] *) +Definition hashmap_test1_fwd (n : nat) : result unit := + match n with + | O => Fail_ OutOfFuel + | S n0 => + hm <- hashmap_hash_map_new_fwd u64 n0; + hm0 <- hashmap_hash_map_insert_fwd_back u64 n0 hm (0 %usize) (42 %u64); + hm1 <- hashmap_hash_map_insert_fwd_back u64 n0 hm0 (128 %usize) (18 %u64); + hm2 <- + hashmap_hash_map_insert_fwd_back u64 n0 hm1 (1024 %usize) (138 %u64); + hm3 <- + hashmap_hash_map_insert_fwd_back u64 n0 hm2 (1056 %usize) (256 %u64); + i <- hashmap_hash_map_get_fwd u64 n0 hm3 (128 %usize); + if negb (i s= 18 %u64) + then Fail_ Failure + else ( + hm4 <- hashmap_hash_map_get_mut_back u64 n0 hm3 (1024 %usize) (56 %u64); + i0 <- hashmap_hash_map_get_fwd u64 n0 hm4 (1024 %usize); + if negb (i0 s= 56 %u64) + then Fail_ Failure + else ( + x <- hashmap_hash_map_remove_fwd u64 n0 hm4 (1024 %usize); + match x with + | None => Fail_ Failure + | Some x0 => + if negb (x0 s= 56 %u64) + then Fail_ Failure + else ( + hm5 <- hashmap_hash_map_remove_back u64 n0 hm4 (1024 %usize); + i1 <- hashmap_hash_map_get_fwd u64 n0 hm5 (0 %usize); + if negb (i1 s= 42 %u64) + then Fail_ Failure + else ( + i2 <- hashmap_hash_map_get_fwd u64 n0 hm5 (128 %usize); + if negb (i2 s= 18 %u64) + then Fail_ Failure + else ( + i3 <- hashmap_hash_map_get_fwd u64 n0 hm5 (1056 %usize); + if negb (i3 s= 256 %u64) then Fail_ Failure else Return tt))) + end)) + end + . + +(** [hashmap_main::insert_on_disk] *) +Definition insert_on_disk_fwd + (n : nat) (key : usize) (value : u64) (st : state) : result (state * unit) := + match n with + | O => Fail_ OutOfFuel + | S n0 => + p <- hashmap_utils_deserialize_fwd st; + let (st0, hm) := p in + hm0 <- hashmap_hash_map_insert_fwd_back u64 n0 hm key value; + p0 <- hashmap_utils_serialize_fwd hm0 st0; + let (st1, _) := p0 in + Return (st1, tt) + end + . + +(** [hashmap_main::main] *) +Definition main_fwd : result unit := Return tt . + +(** Unit test for [hashmap_main::main] *) +Check (main_fwd )%return. + +End HashmapMain__Funs . diff --git a/tests/coq/hashmap_on_disk/HashmapMain__Opaque.v b/tests/coq/hashmap_on_disk/HashmapMain__Opaque.v new file mode 100644 index 00000000..f41c3ddf --- /dev/null +++ b/tests/coq/hashmap_on_disk/HashmapMain__Opaque.v @@ -0,0 +1,21 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: opaque function definitions *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Local Open Scope Primitives_scope. +Require Export HashmapMain__Types . +Import HashmapMain__Types . +Module HashmapMain__Opaque . + +(** [hashmap_main::hashmap_utils::deserialize] *) +Axiom hashmap_utils_deserialize_fwd + : state -> result (state * (Hashmap_hash_map_t u64)) + . + +(** [hashmap_main::hashmap_utils::serialize] *) +Axiom hashmap_utils_serialize_fwd + : Hashmap_hash_map_t u64 -> state -> result (state * unit) + . + +End HashmapMain__Opaque . diff --git a/tests/coq/hashmap_on_disk/HashmapMain__Types.v b/tests/coq/hashmap_on_disk/HashmapMain__Types.v new file mode 100644 index 00000000..8bdbd0bd --- /dev/null +++ b/tests/coq/hashmap_on_disk/HashmapMain__Types.v @@ -0,0 +1,42 @@ +(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) +(** [hashmap_main]: type definitions *) +Require Import Primitives. +Import Primitives. +Require Import Coq.ZArith.ZArith. +Local Open Scope Primitives_scope. +Module HashmapMain__Types . + +(** [hashmap_main::hashmap::List] *) +Inductive Hashmap_list_t (T : Type) := +| HashmapListCons : usize -> T -> Hashmap_list_t T -> Hashmap_list_t T +| HashmapListNil : Hashmap_list_t T +. + +Arguments HashmapListCons {T} _ _ _ . +Arguments HashmapListNil {T} . + +(** [hashmap_main::hashmap::HashMap] *) +Record Hashmap_hash_map_t (T : Type) := +mkHashmap_hash_map_t +{ + Hashmap_hash_map_num_entries : usize; + Hashmap_hash_map_max_load_factor : (usize * usize); + Hashmap_hash_map_max_load : usize; + Hashmap_hash_map_slots : vec (Hashmap_list_t T); +} +. + +Arguments mkHashmap_hash_map_t {T} _ _ _ _ . +Arguments Hashmap_hash_map_num_entries {T} . +Arguments Hashmap_hash_map_max_load_factor {T} . +Arguments Hashmap_hash_map_max_load {T} . +Arguments Hashmap_hash_map_slots {T} . + +(** [core::num::u32::{9}::MAX] *) +Definition core_num_u32_max_body : result u32 := Return (4294967295 %u32) . +Definition core_num_u32_max_c : u32 := core_num_u32_max_body%global . + +(** The state type used in the state-error monad *) +Axiom state : Type. + +End HashmapMain__Types . diff --git a/tests/coq/hashmap_on_disk/Makefile b/tests/coq/hashmap_on_disk/Makefile new file mode 100644 index 00000000..ff1ccd39 --- /dev/null +++ b/tests/coq/hashmap_on_disk/Makefile @@ -0,0 +1,22 @@ +# Makefile originally taken from coq-club + +%: Makefile.coq phony + +make -f Makefile.coq $@ + +all: Makefile.coq + +make -f Makefile.coq all + +clean: Makefile.coq + +make -f Makefile.coq clean + rm -f Makefile.coq + +Makefile.coq: _CoqProject Makefile + coq_makefile -f _CoqProject | sed 's/$$(COQCHK) $$(COQCHKFLAGS) $$(COQLIBS)/$$(COQCHK) $$(COQCHKFLAGS) $$(subst -Q,-R,$$(COQLIBS))/' > Makefile.coq + +_CoqProject: ; + +Makefile: ; + +phony: ; + +.PHONY: all clean phony diff --git a/tests/coq/hashmap_on_disk/Primitives.v b/tests/coq/hashmap_on_disk/Primitives.v new file mode 100644 index 00000000..9a97d6c7 --- /dev/null +++ b/tests/coq/hashmap_on_disk/Primitives.v @@ -0,0 +1,482 @@ +Require Import Lia. +Require Coq.Strings.Ascii. +Require Coq.Strings.String. +Require Import Coq.Program.Equality. +Require Import Coq.ZArith.ZArith. +Require Import Coq.ZArith.Znat. +Require Import List. +Import ListNotations. + +Module Primitives. + + (* TODO: use more *) +Declare Scope Primitives_scope. + +(*** Result *) + +Inductive error := + | Failure + | OutOfFuel. + +Inductive result A := + | Return : A -> result A + | Fail_ : error -> result A. + +Arguments Return {_} a. +Arguments Fail_ {_}. + +Definition bind {A B} (m: result A) (f: A -> result B) : result B := + match m with + | Fail_ e => Fail_ e + | Return x => f x + end. + +Definition return_ {A: Type} (x: A) : result A := Return x. +Definition fail_ {A: Type} (e: error) : result A := Fail_ e. + +Notation "x <- c1 ; c2" := (bind c1 (fun x => c2)) + (at level 61, c1 at next level, right associativity). + +(** Monadic assert *) +Definition massert (b: bool) : result unit := + if b then Return tt else Fail_ Failure. + +(** Normalize and unwrap a successful result (used for globals) *) +Definition eval_result_refl {A} {x} (a: result A) (p: a = Return x) : A := + match a as r return (r = Return x -> A) with + | Return a' => fun _ => a' + | Fail_ e => fun p' => + False_rect _ (eq_ind (Fail_ e) + (fun e : result A => + match e with + | Return _ => False + | Fail_ e => True + end) + I (Return x) p') + end p. + +Notation "x %global" := (eval_result_refl x eq_refl) (at level 40). +Notation "x %return" := (eval_result_refl x eq_refl) (at level 40). + +(* Sanity check *) +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 . + +(*** Scalars *) + +Definition i8_min : Z := -128%Z. +Definition i8_max : Z := 127%Z. +Definition i16_min : Z := -32768%Z. +Definition i16_max : Z := 32767%Z. +Definition i32_min : Z := -2147483648%Z. +Definition i32_max : Z := 2147483647%Z. +Definition i64_min : Z := -9223372036854775808%Z. +Definition i64_max : Z := 9223372036854775807%Z. +Definition i128_min : Z := -170141183460469231731687303715884105728%Z. +Definition i128_max : Z := 170141183460469231731687303715884105727%Z. +Definition u8_min : Z := 0%Z. +Definition u8_max : Z := 255%Z. +Definition u16_min : Z := 0%Z. +Definition u16_max : Z := 65535%Z. +Definition u32_min : Z := 0%Z. +Definition u32_max : Z := 4294967295%Z. +Definition u64_min : Z := 0%Z. +Definition u64_max : Z := 18446744073709551615%Z. +Definition u128_min : Z := 0%Z. +Definition u128_max : Z := 340282366920938463463374607431768211455%Z. + +(** The bounds of [isize] and [usize] vary with the architecture. *) +Axiom isize_min : Z. +Axiom isize_max : Z. +Definition usize_min : Z := 0%Z. +Axiom usize_max : Z. + +Open Scope Z_scope. + +(** We provide those lemmas to reason about the bounds of [isize] and [usize] *) +Axiom isize_min_bound : isize_min <= i32_min. +Axiom isize_max_bound : i32_max <= isize_max. +Axiom usize_max_bound : u32_max <= usize_max. + +Inductive scalar_ty := + | Isize + | I8 + | I16 + | I32 + | I64 + | I128 + | Usize + | U8 + | U16 + | U32 + | U64 + | U128 +. + +Definition scalar_min (ty: scalar_ty) : Z := + match ty with + | Isize => isize_min + | I8 => i8_min + | I16 => i16_min + | I32 => i32_min + | I64 => i64_min + | I128 => i128_min + | Usize => usize_min + | U8 => u8_min + | U16 => u16_min + | U32 => u32_min + | U64 => u64_min + | U128 => u128_min +end. + +Definition scalar_max (ty: scalar_ty) : Z := + match ty with + | Isize => isize_max + | I8 => i8_max + | I16 => i16_max + | I32 => i32_max + | I64 => i64_max + | I128 => i128_max + | Usize => usize_max + | U8 => u8_max + | U16 => u16_max + | U32 => u32_max + | U64 => u64_max + | U128 => u128_max +end. + +(** We use the following conservative bounds to make sure we can compute bound + checks in most situations *) +Definition scalar_min_cons (ty: scalar_ty) : Z := + match ty with + | Isize => i32_min + | Usize => u32_min + | _ => scalar_min ty +end. + +Definition scalar_max_cons (ty: scalar_ty) : Z := + match ty with + | Isize => i32_max + | Usize => u32_max + | _ => scalar_max ty +end. + +Lemma scalar_min_cons_valid : forall ty, scalar_min ty <= scalar_min_cons ty . +Proof. + destruct ty; unfold scalar_min_cons, scalar_min; try lia. + - pose isize_min_bound; lia. + - apply Z.le_refl. +Qed. + +Lemma scalar_max_cons_valid : forall ty, scalar_max ty >= scalar_max_cons ty . +Proof. + destruct ty; unfold scalar_max_cons, scalar_max; try lia. + - pose isize_max_bound; lia. + - pose usize_max_bound. lia. +Qed. + +Definition scalar (ty: scalar_ty) : Type := + { x: Z | scalar_min ty <= x <= scalar_max ty }. + +Definition to_Z {ty} (x: scalar ty) : Z := proj1_sig x. + +(** Bounds checks: we start by using the conservative bounds, to make sure we + can compute in most situations, then we use the real bounds (for [isize] + and [usize]). *) +Definition scalar_ge_min (ty: scalar_ty) (x: Z) : bool := + Z.leb (scalar_min_cons ty) x || Z.leb (scalar_min ty) x. + +Definition scalar_le_max (ty: scalar_ty) (x: Z) : bool := + Z.leb x (scalar_max_cons ty) || Z.leb x (scalar_max ty). + +Lemma scalar_ge_min_valid (ty: scalar_ty) (x: Z) : + scalar_ge_min ty x = true -> scalar_min ty <= x . +Proof. + unfold scalar_ge_min. + pose (scalar_min_cons_valid ty). + lia. +Qed. + +Lemma scalar_le_max_valid (ty: scalar_ty) (x: Z) : + scalar_le_max ty x = true -> x <= scalar_max ty . +Proof. + unfold scalar_le_max. + pose (scalar_max_cons_valid ty). + lia. +Qed. + +Definition scalar_in_bounds (ty: scalar_ty) (x: Z) : bool := + scalar_ge_min ty x && scalar_le_max ty x . + +Lemma scalar_in_bounds_valid (ty: scalar_ty) (x: Z) : + scalar_in_bounds ty x = true -> scalar_min ty <= x <= scalar_max ty . +Proof. + unfold scalar_in_bounds. + intros H. + destruct (scalar_ge_min ty x) eqn:Hmin. + - destruct (scalar_le_max ty x) eqn:Hmax. + + pose (scalar_ge_min_valid ty x Hmin). + pose (scalar_le_max_valid ty x Hmax). + lia. + + inversion H. + - inversion H. +Qed. + +Import Sumbool. + +Definition mk_scalar (ty: scalar_ty) (x: Z) : result (scalar ty) := + match sumbool_of_bool (scalar_in_bounds ty x) with + | left H => Return (exist _ x (scalar_in_bounds_valid _ _ H)) + | right _ => Fail_ Failure + end. + +Definition scalar_add {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x + to_Z y). + +Definition scalar_sub {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x - to_Z y). + +Definition scalar_mul {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (to_Z x * to_Z y). + +Definition scalar_div {ty} (x y: scalar ty) : result (scalar ty) := + if to_Z y =? 0 then Fail_ Failure else + mk_scalar ty (to_Z x / to_Z y). + +Definition scalar_rem {ty} (x y: scalar ty) : result (scalar ty) := mk_scalar ty (Z.rem (to_Z x) (to_Z y)). + +Definition scalar_neg {ty} (x: scalar ty) : result (scalar ty) := mk_scalar ty (-(to_Z x)). + +(** Cast an integer from a [src_ty] to a [tgt_ty] *) +(* TODO: check the semantics of casts in Rust *) +Definition scalar_cast (src_ty tgt_ty : scalar_ty) (x : scalar src_ty) : result (scalar tgt_ty) := + mk_scalar tgt_ty (to_Z x). + +(** Comparisons *) +Print Z.leb . + +Definition scalar_leb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.leb (to_Z x) (to_Z y) . + +Definition scalar_ltb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.ltb (to_Z x) (to_Z y) . + +Definition scalar_geb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.geb (to_Z x) (to_Z y) . + +Definition scalar_gtb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.gtb (to_Z x) (to_Z y) . + +Definition scalar_eqb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + Z.eqb (to_Z x) (to_Z y) . + +Definition scalar_neqb {ty : scalar_ty} (x : scalar ty) (y : scalar ty) : bool := + negb (Z.eqb (to_Z x) (to_Z y)) . + + +(** The scalar types *) +Definition isize := scalar Isize. +Definition i8 := scalar I8. +Definition i16 := scalar I16. +Definition i32 := scalar I32. +Definition i64 := scalar I64. +Definition i128 := scalar I128. +Definition usize := scalar Usize. +Definition u8 := scalar U8. +Definition u16 := scalar U16. +Definition u32 := scalar U32. +Definition u64 := scalar U64. +Definition u128 := scalar U128. + +(** Negaion *) +Definition isize_neg := @scalar_neg Isize. +Definition i8_neg := @scalar_neg I8. +Definition i16_neg := @scalar_neg I16. +Definition i32_neg := @scalar_neg I32. +Definition i64_neg := @scalar_neg I64. +Definition i128_neg := @scalar_neg I128. + +(** Division *) +Definition isize_div := @scalar_div Isize. +Definition i8_div := @scalar_div I8. +Definition i16_div := @scalar_div I16. +Definition i32_div := @scalar_div I32. +Definition i64_div := @scalar_div I64. +Definition i128_div := @scalar_div I128. +Definition usize_div := @scalar_div Usize. +Definition u8_div := @scalar_div U8. +Definition u16_div := @scalar_div U16. +Definition u32_div := @scalar_div U32. +Definition u64_div := @scalar_div U64. +Definition u128_div := @scalar_div U128. + +(** Remainder *) +Definition isize_rem := @scalar_rem Isize. +Definition i8_rem := @scalar_rem I8. +Definition i16_rem := @scalar_rem I16. +Definition i32_rem := @scalar_rem I32. +Definition i64_rem := @scalar_rem I64. +Definition i128_rem := @scalar_rem I128. +Definition usize_rem := @scalar_rem Usize. +Definition u8_rem := @scalar_rem U8. +Definition u16_rem := @scalar_rem U16. +Definition u32_rem := @scalar_rem U32. +Definition u64_rem := @scalar_rem U64. +Definition u128_rem := @scalar_rem U128. + +(** Addition *) +Definition isize_add := @scalar_add Isize. +Definition i8_add := @scalar_add I8. +Definition i16_add := @scalar_add I16. +Definition i32_add := @scalar_add I32. +Definition i64_add := @scalar_add I64. +Definition i128_add := @scalar_add I128. +Definition usize_add := @scalar_add Usize. +Definition u8_add := @scalar_add U8. +Definition u16_add := @scalar_add U16. +Definition u32_add := @scalar_add U32. +Definition u64_add := @scalar_add U64. +Definition u128_add := @scalar_add U128. + +(** Substraction *) +Definition isize_sub := @scalar_sub Isize. +Definition i8_sub := @scalar_sub I8. +Definition i16_sub := @scalar_sub I16. +Definition i32_sub := @scalar_sub I32. +Definition i64_sub := @scalar_sub I64. +Definition i128_sub := @scalar_sub I128. +Definition usize_sub := @scalar_sub Usize. +Definition u8_sub := @scalar_sub U8. +Definition u16_sub := @scalar_sub U16. +Definition u32_sub := @scalar_sub U32. +Definition u64_sub := @scalar_sub U64. +Definition u128_sub := @scalar_sub U128. + +(** Multiplication *) +Definition isize_mul := @scalar_mul Isize. +Definition i8_mul := @scalar_mul I8. +Definition i16_mul := @scalar_mul I16. +Definition i32_mul := @scalar_mul I32. +Definition i64_mul := @scalar_mul I64. +Definition i128_mul := @scalar_mul I128. +Definition usize_mul := @scalar_mul Usize. +Definition u8_mul := @scalar_mul U8. +Definition u16_mul := @scalar_mul U16. +Definition u32_mul := @scalar_mul U32. +Definition u64_mul := @scalar_mul U64. +Definition u128_mul := @scalar_mul U128. + +(** Small utility *) +Definition usize_to_nat (x: usize) : nat := Z.to_nat (to_Z x). + +(** Notations *) +Notation "x %isize" := ((mk_scalar Isize x)%return) (at level 9). +Notation "x %i8" := ((mk_scalar I8 x)%return) (at level 9). +Notation "x %i16" := ((mk_scalar I16 x)%return) (at level 9). +Notation "x %i32" := ((mk_scalar I32 x)%return) (at level 9). +Notation "x %i64" := ((mk_scalar I64 x)%return) (at level 9). +Notation "x %i128" := ((mk_scalar I128 x)%return) (at level 9). +Notation "x %usize" := ((mk_scalar Usize x)%return) (at level 9). +Notation "x %u8" := ((mk_scalar U8 x)%return) (at level 9). +Notation "x %u16" := ((mk_scalar U16 x)%return) (at level 9). +Notation "x %u32" := ((mk_scalar U32 x)%return) (at level 9). +Notation "x %u64" := ((mk_scalar U64 x)%return) (at level 9). +Notation "x %u128" := ((mk_scalar U128 x)%return) (at level 9). + +Notation "x s= y" := (scalar_eqb x y) (at level 80) : Primitives_scope. +Notation "x s<> y" := (scalar_neqb x y) (at level 80) : Primitives_scope. +Notation "x s<= y" := (scalar_leb x y) (at level 80) : Primitives_scope. +Notation "x s< y" := (scalar_ltb x y) (at level 80) : Primitives_scope. +Notation "x s>= y" := (scalar_geb x y) (at level 80) : Primitives_scope. +Notation "x s> y" := (scalar_gtb x y) (at level 80) : Primitives_scope. + +(*** Vectors *) + +Definition 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 vec_length {T: Type} (v: vec T) : Z := Z.of_nat (length (vec_to_list v)). + +Lemma le_0_usize_max : 0 <= usize_max. +Proof. + pose (H := usize_max_bound). + unfold u32_max in H. + lia. +Qed. + +Definition vec_new (T: Type) : vec T := (exist _ [] le_0_usize_max). + +Lemma vec_len_in_usize {T} (v: vec T) : usize_min <= vec_length v <= usize_max. +Proof. + unfold 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). + +Fixpoint list_update {A} (l: list A) (n: nat) (a: A) + : list A := + match l with + | [] => [] + | x :: t => match n with + | 0%nat => a :: t + | 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) ; + 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 vec_push_back (T: Type) (v: vec T) (x: T) : result (vec T) := + 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 vec_insert_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). + +(* 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 + 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). + +End Primitives. diff --git a/tests/coq/hashmap_on_disk/_CoqProject b/tests/coq/hashmap_on_disk/_CoqProject new file mode 100644 index 00000000..10e247ae --- /dev/null +++ b/tests/coq/hashmap_on_disk/_CoqProject @@ -0,0 +1,9 @@ +-R . Lib +-arg -w +-arg all + +Primitives.v + +HashmapMain__Funs.v +HashmapMain__Opaque.v +HashmapMain__Types.v
\ No newline at end of file |