(** THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS *) (** [traits] *) Require Import Primitives. Import Primitives. Require Import Coq.ZArith.ZArith. Require Import List. Import ListNotations. Local Open Scope Primitives_scope. Module Traits. (** Trait declaration: [traits::BoolTrait] *) Record BoolTrait_t (Self : Type) := mkBoolTrait_t { BoolTrait_t_get_bool : Self -> result bool; }. Arguments mkBoolTrait_t { _ }. Arguments BoolTrait_t_get_bool { _ }. (** [traits::Bool::{0}::get_bool]: forward function *) Definition bool_get_bool (self : bool) : result bool := Return self. (** Trait implementation: [traits::Bool::{0}] *) Definition Bool_BoolTraitInst : BoolTrait_t bool := {| BoolTrait_t_get_bool := bool_get_bool; |}. (** [traits::BoolTrait::ret_true]: forward function *) Definition boolTrait_ret_true {Self : Type} (self_clause : BoolTrait_t Self) (self : Self) : result bool := Return true . (** [traits::test_bool_trait_bool]: forward function *) Definition test_bool_trait_bool (x : bool) : result bool := b <- bool_get_bool x; if b then boolTrait_ret_true Bool_BoolTraitInst x else Return false . (** [traits::Option::{1}::get_bool]: forward function *) Definition option_get_bool (T : Type) (self : option T) : result bool := match self with | None => Return false | Some t => Return true end . (** Trait implementation: [traits::Option::{1}] *) Definition Option_BoolTraitInst (T : Type) : BoolTrait_t (option T) := {| BoolTrait_t_get_bool := option_get_bool T; |}. (** [traits::test_bool_trait_option]: forward function *) Definition test_bool_trait_option (T : Type) (x : option T) : result bool := b <- option_get_bool T x; if b then boolTrait_ret_true (Option_BoolTraitInst T) x else Return false . (** [traits::test_bool_trait]: forward function *) Definition test_bool_trait (T : Type) (inst : BoolTrait_t T) (x : T) : result bool := inst.(BoolTrait_t_get_bool) x . (** Trait declaration: [traits::ToU64] *) Record ToU64_t (Self : Type) := mkToU64_t { ToU64_t_to_u64 : Self -> result u64; }. Arguments mkToU64_t { _ }. Arguments ToU64_t_to_u64 { _ }. (** [traits::u64::{2}::to_u64]: forward function *) Definition u64_to_u64 (self : u64) : result u64 := Return self. (** Trait implementation: [traits::u64::{2}] *) Definition u64_ToU64Inst : ToU64_t u64 := {| ToU64_t_to_u64 := u64_to_u64; |}. (** [traits::Tuple2::{3}::to_u64]: forward function *) Definition tuple2_to_u64 (A : Type) (inst : ToU64_t A) (self : (A * A)) : result u64 := let (t, t0) := self in i <- inst.(ToU64_t_to_u64) t; i0 <- inst.(ToU64_t_to_u64) t0; u64_add i i0 . (** Trait implementation: [traits::Tuple2::{3}] *) Definition Tuple2_ToU64Inst (A : Type) (inst : ToU64_t A) : ToU64_t (A * A) := {| ToU64_t_to_u64 := tuple2_to_u64 A inst; |}. (** [traits::f]: forward function *) Definition f (T : Type) (inst : ToU64_t T) (x : (T * T)) : result u64 := tuple2_to_u64 T inst x . (** [traits::g]: forward function *) Definition g (T : Type) (inst : ToU64_t (T * T)) (x : (T * T)) : result u64 := inst.(ToU64_t_to_u64) x . (** [traits::h0]: forward function *) Definition h0 (x : u64) : result u64 := u64_to_u64 x. (** [traits::Wrapper] *) Record Wrapper_t (T : Type) := mkWrapper_t { wrapper_x : T; }. Arguments mkWrapper_t { _ }. Arguments wrapper_x { _ }. (** [traits::Wrapper::{4}::to_u64]: forward function *) Definition wrapper_to_u64 (T : Type) (inst : ToU64_t T) (self : Wrapper_t T) : result u64 := inst.(ToU64_t_to_u64) self.(wrapper_x) . (** Trait implementation: [traits::Wrapper::{4}] *) Definition Wrapper_ToU64Inst (T : Type) (inst : ToU64_t T) : ToU64_t (Wrapper_t T) := {| ToU64_t_to_u64 := wrapper_to_u64 T inst; |}. (** [traits::h1]: forward function *) Definition h1 (x : Wrapper_t u64) : result u64 := wrapper_to_u64 u64 u64_ToU64Inst x . (** [traits::h2]: forward function *) Definition h2 (T : Type) (inst : ToU64_t T) (x : Wrapper_t T) : result u64 := wrapper_to_u64 T inst x . (** Trait declaration: [traits::ToType] *) Record ToType_t (Self T : Type) := mkToType_t { ToType_t_to_type : Self -> result T; }. Arguments mkToType_t { _ _ }. Arguments ToType_t_to_type { _ _ }. (** [traits::u64::{5}::to_type]: forward function *) Definition u64_to_type (self : u64) : result bool := Return (self s> 0%u64). (** Trait implementation: [traits::u64::{5}] *) Definition u64_ToTypeInst : ToType_t u64 bool := {| ToType_t_to_type := u64_to_type; |}. (** Trait declaration: [traits::OfType] *) Record OfType_t (Self : Type) := mkOfType_t { OfType_t_of_type : forall (T : Type) (inst : ToType_t T Self), T -> result Self; }. Arguments mkOfType_t { _ }. Arguments OfType_t_of_type { _ }. (** [traits::h3]: forward function *) Definition h3 (T1 T2 : Type) (inst : OfType_t T1) (inst0 : ToType_t T2 T1) (y : T2) : result T1 := inst.(OfType_t_of_type) T2 inst0 y . (** Trait declaration: [traits::OfTypeBis] *) Record OfTypeBis_t (Self T : Type) := mkOfTypeBis_t { OfTypeBis_tOfTypeBis_t_parent_clause_0 : ToType_t T Self; OfTypeBis_t_of_type : T -> result Self; }. Arguments mkOfTypeBis_t { _ _ }. Arguments OfTypeBis_tOfTypeBis_t_parent_clause_0 { _ _ }. Arguments OfTypeBis_t_of_type { _ _ }. (** [traits::h4]: forward function *) Definition h4 (T1 T2 : Type) (inst : OfTypeBis_t T1 T2) (inst0 : ToType_t T2 T1) (y : T2) : result T1 := inst.(OfTypeBis_t_of_type) y . (** [traits::TestType] *) Record TestType_t (T : Type) := mkTestType_t { testType_0 : T; }. Arguments mkTestType_t { _ }. Arguments testType_0 { _ }. (** [traits::TestType::{6}::test::TestType1] *) Record TestType_test_TestType1_t := mkTestType_test_TestType1_t { testType_test_TestType1_0 : u64; } . (** Trait declaration: [traits::TestType::{6}::test::TestTrait] *) Record TestType_test_TestTrait_t (Self : Type) := mkTestType_test_TestTrait_t { TestType_test_TestTrait_t_test : Self -> result bool; }. Arguments mkTestType_test_TestTrait_t { _ }. Arguments TestType_test_TestTrait_t_test { _ }. (** [traits::TestType::{6}::test::TestType1::{0}::test]: forward function *) Definition testType_test_TestType1_test (self : TestType_test_TestType1_t) : result bool := Return (self.(testType_test_TestType1_0) s> 1%u64) . (** Trait implementation: [traits::TestType::{6}::test::TestType1::{0}] *) Definition TestType_test_TestType1_TestType_test_TestTraitInst : TestType_test_TestTrait_t TestType_test_TestType1_t := {| TestType_test_TestTrait_t_test := testType_test_TestType1_test; |}. (** [traits::TestType::{6}::test]: forward function *) Definition testType_test (T : Type) (inst : ToU64_t T) (self : TestType_t T) (x : T) : result bool := x0 <- inst.(ToU64_t_to_u64) x; if x0 s> 0%u64 then testType_test_TestType1_test {| testType_test_TestType1_0 := 0%u64 |} else Return false . (** [traits::BoolWrapper] *) Record BoolWrapper_t := mkBoolWrapper_t { boolWrapper_0 : bool; }. (** [traits::BoolWrapper::{7}::to_type]: forward function *) Definition boolWrapper_to_type (T : Type) (inst : ToType_t bool T) (self : BoolWrapper_t) : result T := inst.(ToType_t_to_type) self.(boolWrapper_0) . (** Trait implementation: [traits::BoolWrapper::{7}] *) Definition BoolWrapper_ToTypeInst (T : Type) (inst : ToType_t bool T) : ToType_t BoolWrapper_t T := {| ToType_t_to_type := boolWrapper_to_type T inst; |}. (** [traits::WithConstTy::LEN2] *) Definition with_const_ty_len2_body : result usize := Return 32%usize. Definition with_const_ty_len2_c : usize := with_const_ty_len2_body%global. (** Trait declaration: [traits::WithConstTy] *) Record WithConstTy_t (Self : Type) (LEN : usize) := mkWithConstTy_t { WithConstTy_tWithConstTy_t_LEN1 : usize; WithConstTy_tWithConstTy_t_LEN2 : usize; WithConstTy_tWithConstTy_t_V : Type; WithConstTy_tWithConstTy_t_W : Type; WithConstTy_tWithConstTy_t_W_clause_0 : ToU64_t WithConstTy_tWithConstTy_t_W; WithConstTy_t_f : WithConstTy_tWithConstTy_t_W -> array u8 LEN -> result WithConstTy_tWithConstTy_t_W; }. Arguments mkWithConstTy_t { _ _ }. Arguments WithConstTy_tWithConstTy_t_LEN1 { _ _ }. Arguments WithConstTy_tWithConstTy_t_LEN2 { _ _ }. Arguments WithConstTy_tWithConstTy_t_V { _ _ }. Arguments WithConstTy_tWithConstTy_t_W { _ _ }. Arguments WithConstTy_tWithConstTy_t_W_clause_0 { _ _ }. Arguments WithConstTy_t_f { _ _ }. (** [traits::Bool::{8}::LEN1] *) Definition bool_len1_body : result usize := Return 12%usize. Definition bool_len1_c : usize := bool_len1_body%global. (** [traits::Bool::{8}::f]: merged forward/backward function (there is a single backward function, and the forward function returns ()) *) Definition bool_f (i : u64) (a : array u8 32%usize) : result u64 := Return i. (** Trait implementation: [traits::Bool::{8}] *) Definition Bool_WithConstTyInst : WithConstTy_t bool 32%usize := {| WithConstTy_tWithConstTy_t_LEN1 := bool_len1_c; WithConstTy_tWithConstTy_t_LEN2 := with_const_ty_len2_c; WithConstTy_tWithConstTy_t_V := u8; WithConstTy_tWithConstTy_t_W := u64; WithConstTy_tWithConstTy_t_W_clause_0 := u64_ToU64Inst; WithConstTy_t_f := bool_f; |}. (** [traits::use_with_const_ty1]: forward function *) Definition use_with_const_ty1 (H : Type) (LEN : usize) (inst : WithConstTy_t H LEN) : result usize := let i := inst.(WithConstTy_tWithConstTy_t_LEN1) in Return i . (** [traits::use_with_const_ty2]: forward function *) Definition use_with_const_ty2 (H : Type) (LEN : usize) (inst : WithConstTy_t H LEN) (w : inst.(WithConstTy_tWithConstTy_t_W)) : result unit := Return tt . (** [traits::use_with_const_ty3]: forward function *) Definition use_with_const_ty3 (H : Type) (LEN : usize) (inst : WithConstTy_t H LEN) (x : inst.(WithConstTy_tWithConstTy_t_W)) : result u64 := inst.(WithConstTy_tWithConstTy_t_W_clause_0).(ToU64_t_to_u64) x . (** [traits::test_where1]: forward function *) Definition test_where1 (T : Type) (_x : T) : result unit := Return tt. (** [traits::test_where2]: forward function *) Definition test_where2 (T : Type) (inst : WithConstTy_t T 32%usize) (_x : u32) : result unit := Return tt . (** [alloc::string::String] *) Axiom alloc_string_String_t : Type. (** Trait declaration: [traits::ParentTrait0] *) Record ParentTrait0_t (Self : Type) := mkParentTrait0_t { ParentTrait0_tParentTrait0_t_W : Type; ParentTrait0_t_get_name : Self -> result alloc_string_String_t; ParentTrait0_t_get_w : Self -> result ParentTrait0_tParentTrait0_t_W; }. Arguments mkParentTrait0_t { _ }. Arguments ParentTrait0_tParentTrait0_t_W { _ }. Arguments ParentTrait0_t_get_name { _ }. Arguments ParentTrait0_t_get_w { _ }. (** Trait declaration: [traits::ParentTrait1] *) Record ParentTrait1_t (Self : Type) := mkParentTrait1_t{}. Arguments mkParentTrait1_t { _ }. (** Trait declaration: [traits::ChildTrait] *) Record ChildTrait_t (Self : Type) := mkChildTrait_t { ChildTrait_tChildTrait_t_parent_clause_0 : ParentTrait0_t Self; ChildTrait_tChildTrait_t_parent_clause_1 : ParentTrait1_t Self; }. Arguments mkChildTrait_t { _ }. Arguments ChildTrait_tChildTrait_t_parent_clause_0 { _ }. Arguments ChildTrait_tChildTrait_t_parent_clause_1 { _ }. (** [traits::test_child_trait1]: forward function *) Definition test_child_trait1 (T : Type) (inst : ChildTrait_t T) (x : T) : result alloc_string_String_t := inst.(ChildTrait_tChildTrait_t_parent_clause_0).(ParentTrait0_t_get_name) x . (** [traits::test_child_trait2]: forward function *) Definition test_child_trait2 (T : Type) (inst : ChildTrait_t T) (x : T) : result inst.(ChildTrait_tChildTrait_t_parent_clause_0).(ParentTrait0_tParentTrait0_t_W) := inst.(ChildTrait_tChildTrait_t_parent_clause_0).(ParentTrait0_t_get_w) x . (** [traits::order1]: forward function *) Definition order1 (T U : Type) (inst : ParentTrait0_t T) (inst0 : ParentTrait0_t U) : result unit := Return tt . (** Trait declaration: [traits::ChildTrait1] *) Record ChildTrait1_t (Self : Type) := mkChildTrait1_t { ChildTrait1_tChildTrait1_t_parent_clause_0 : ParentTrait1_t Self; }. Arguments mkChildTrait1_t { _ }. Arguments ChildTrait1_tChildTrait1_t_parent_clause_0 { _ }. (** Trait implementation: [traits::usize::{9}] *) Definition usize_ParentTrait1Inst : ParentTrait1_t usize := mkParentTrait1_t. (** Trait implementation: [traits::usize::{10}] *) Definition usize_ChildTrait1Inst : ChildTrait1_t usize := {| ChildTrait1_tChildTrait1_t_parent_clause_0 := usize_ParentTrait1Inst; |}. (** Trait declaration: [traits::Iterator] *) Record Iterator_t (Self : Type) := mkIterator_t { Iterator_tIterator_t_Item : Type; }. Arguments mkIterator_t { _ }. Arguments Iterator_tIterator_t_Item { _ }. (** Trait declaration: [traits::IntoIterator] *) Record IntoIterator_t (Self : Type) := mkIntoIterator_t { IntoIterator_tIntoIterator_t_Item : Type; IntoIterator_tIntoIterator_t_IntoIter : Type; IntoIterator_tIntoIterator_t_IntoIter_clause_0 : Iterator_t IntoIterator_tIntoIterator_t_IntoIter; IntoIterator_t_into_iter : Self -> result IntoIterator_tIntoIterator_t_IntoIter; }. Arguments mkIntoIterator_t { _ }. Arguments IntoIterator_tIntoIterator_t_Item { _ }. Arguments IntoIterator_tIntoIterator_t_IntoIter { _ }. Arguments IntoIterator_tIntoIterator_t_IntoIter_clause_0 { _ }. Arguments IntoIterator_t_into_iter { _ }. (** Trait declaration: [traits::FromResidual] *) Record FromResidual_t (Self T : Type) := mkFromResidual_t{}. Arguments mkFromResidual_t { _ _ }. (** Trait declaration: [traits::Try] *) Record Try_t (Self : Type) := mkTry_t { Try_tTry_t_Residual : Type; Try_tTry_t_parent_clause_0 : FromResidual_t Self Try_tTry_t_Residual; }. Arguments mkTry_t { _ }. Arguments Try_tTry_t_Residual { _ }. Arguments Try_tTry_t_parent_clause_0 { _ }. (** Trait declaration: [traits::WithTarget] *) Record WithTarget_t (Self : Type) := mkWithTarget_t { WithTarget_tWithTarget_t_Target : Type; }. Arguments mkWithTarget_t { _ }. Arguments WithTarget_tWithTarget_t_Target { _ }. (** Trait declaration: [traits::ParentTrait2] *) Record ParentTrait2_t (Self : Type) := mkParentTrait2_t { ParentTrait2_tParentTrait2_t_U : Type; ParentTrait2_tParentTrait2_t_U_clause_0 : WithTarget_t ParentTrait2_tParentTrait2_t_U; }. Arguments mkParentTrait2_t { _ }. Arguments ParentTrait2_tParentTrait2_t_U { _ }. Arguments ParentTrait2_tParentTrait2_t_U_clause_0 { _ }. (** Trait declaration: [traits::ChildTrait2] *) Record ChildTrait2_t (Self : Type) := mkChildTrait2_t { ChildTrait2_tChildTrait2_t_parent_clause_0 : ParentTrait2_t Self; ChildTrait2_t_convert : (ChildTrait2_tChildTrait2_t_parent_clause_0).(ParentTrait2_tParentTrait2_t_U) -> result (ChildTrait2_tChildTrait2_t_parent_clause_0).(ParentTrait2_tParentTrait2_t_U_clause_0).(WithTarget_tWithTarget_t_Target); }. Arguments mkChildTrait2_t { _ }. Arguments ChildTrait2_tChildTrait2_t_parent_clause_0 { _ }. Arguments ChildTrait2_t_convert { _ }. (** Trait implementation: [traits::u32::{11}] *) Definition u32_WithTargetInst : WithTarget_t u32 := {| WithTarget_tWithTarget_t_Target := u32; |}. (** Trait implementation: [traits::u32::{12}] *) Definition u32_ParentTrait2Inst : ParentTrait2_t u32 := {| ParentTrait2_tParentTrait2_t_U := u32; ParentTrait2_tParentTrait2_t_U_clause_0 := u32_WithTargetInst; |}. (** [traits::u32::{13}::convert]: forward function *) Definition u32_convert (x : u32) : result u32 := Return x. (** Trait implementation: [traits::u32::{13}] *) Definition u32_ChildTrait2Inst : ChildTrait2_t u32 := {| ChildTrait2_tChildTrait2_t_parent_clause_0 := u32_ParentTrait2Inst; ChildTrait2_t_convert := u32_convert; |}. (** [traits::incr_u32]: forward function *) Definition incr_u32 (x : u32) : result u32 := u32_add x 1%u32. (** Trait declaration: [traits::CFnOnce] *) Record CFnOnce_t (Self Args : Type) := mkCFnOnce_t { CFnOnce_tCFnOnce_t_Output : Type; CFnOnce_t_call_once : Self -> Args -> result CFnOnce_tCFnOnce_t_Output; }. Arguments mkCFnOnce_t { _ _ }. Arguments CFnOnce_tCFnOnce_t_Output { _ _ }. Arguments CFnOnce_t_call_once { _ _ }. (** Trait declaration: [traits::CFnMut] *) Record CFnMut_t (Self Args : Type) := mkCFnMut_t { CFnMut_tCFnMut_t_parent_clause_0 : CFnOnce_t Self Args; CFnMut_t_call_mut : Self -> Args -> result (CFnMut_tCFnMut_t_parent_clause_0).(CFnOnce_tCFnOnce_t_Output); CFnMut_t_call_mut_back : Self -> Args -> (CFnMut_tCFnMut_t_parent_clause_0).(CFnOnce_tCFnOnce_t_Output) -> result Self; }. Arguments mkCFnMut_t { _ _ }. Arguments CFnMut_tCFnMut_t_parent_clause_0 { _ _ }. Arguments CFnMut_t_call_mut { _ _ }. Arguments CFnMut_t_call_mut_back { _ _ }. (** Trait declaration: [traits::CFn] *) Record CFn_t (Self Args : Type) := mkCFn_t { CFn_tCFn_t_parent_clause_0 : CFnMut_t Self Args; CFn_t_call_mut : Self -> Args -> result (CFn_tCFn_t_parent_clause_0).(CFnMut_tCFnMut_t_parent_clause_0).(CFnOnce_tCFnOnce_t_Output); }. Arguments mkCFn_t { _ _ }. Arguments CFn_tCFn_t_parent_clause_0 { _ _ }. Arguments CFn_t_call_mut { _ _ }. End Traits .