-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS -- [traits]: function definitions import Base import Traits.Types open Primitives namespace traits /- [traits::Bool::{0}::get_bool]: forward function -/ def Bool.get_bool (self : Bool) : Result Bool := Result.ret self /- Trait implementation: [traits::Bool::{0}] -/ def Bool.BoolTraitInst : BoolTrait Bool := { get_bool := Bool.get_bool } /- [traits::BoolTrait::ret_true]: forward function -/ def BoolTrait.ret_true {Self : Type} (self_clause : BoolTrait Self) (self : Self) : Result Bool := Result.ret true /- [traits::test_bool_trait_bool]: forward function -/ def test_bool_trait_bool (x : Bool) : Result Bool := do let b ← Bool.get_bool x if b then BoolTrait.ret_true Bool.BoolTraitInst x else Result.ret false /- [traits::Option::{1}::get_bool]: forward function -/ def Option.get_bool (T : Type) (self : Option T) : Result Bool := match self with | Option.none => Result.ret false | Option.some t => Result.ret true /- Trait implementation: [traits::Option::{1}] -/ def Option.BoolTraitInst (T : Type) : BoolTrait (Option T) := { get_bool := Option.get_bool T } /- [traits::test_bool_trait_option]: forward function -/ def test_bool_trait_option (T : Type) (x : Option T) : Result Bool := do let b ← Option.get_bool T x if b then BoolTrait.ret_true (Option.BoolTraitInst T) x else Result.ret false /- [traits::test_bool_trait]: forward function -/ def test_bool_trait (T : Type) (inst : BoolTrait T) (x : T) : Result Bool := inst.get_bool x /- [traits::u64::{2}::to_u64]: forward function -/ def u64.to_u64 (self : U64) : Result U64 := Result.ret self /- Trait implementation: [traits::u64::{2}] -/ def u64.ToU64Inst : ToU64 U64 := { to_u64 := u64.to_u64 } /- [traits::Tuple2::{3}::to_u64]: forward function -/ def Tuple2.to_u64 (A : Type) (inst : ToU64 A) (self : (A × A)) : Result U64 := do let (t, t0) := self let i ← inst.to_u64 t let i0 ← inst.to_u64 t0 i + i0 /- Trait implementation: [traits::Tuple2::{3}] -/ def Tuple2.ToU64Inst (A : Type) (inst : ToU64 A) : ToU64 (A × A) := { to_u64 := Tuple2.to_u64 A inst } /- [traits::f]: forward function -/ def f (T : Type) (inst : ToU64 T) (x : (T × T)) : Result U64 := Tuple2.to_u64 T inst x /- [traits::g]: forward function -/ def g (T : Type) (inst : ToU64 (T × T)) (x : (T × T)) : Result U64 := inst.to_u64 x /- [traits::h0]: forward function -/ def h0 (x : U64) : Result U64 := u64.to_u64 x /- [traits::Wrapper::{4}::to_u64]: forward function -/ def Wrapper.to_u64 (T : Type) (inst : ToU64 T) (self : Wrapper T) : Result U64 := inst.to_u64 self.x /- Trait implementation: [traits::Wrapper::{4}] -/ def Wrapper.ToU64Inst (T : Type) (inst : ToU64 T) : ToU64 (Wrapper T) := { to_u64 := Wrapper.to_u64 T inst } /- [traits::h1]: forward function -/ def h1 (x : Wrapper U64) : Result U64 := Wrapper.to_u64 U64 u64.ToU64Inst x /- [traits::h2]: forward function -/ def h2 (T : Type) (inst : ToU64 T) (x : Wrapper T) : Result U64 := Wrapper.to_u64 T inst x /- [traits::u64::{5}::to_type]: forward function -/ def u64.to_type (self : U64) : Result Bool := Result.ret (self > (U64.ofInt 0)) /- Trait implementation: [traits::u64::{5}] -/ def u64.ToTypeInst : ToType U64 Bool := { to_type := u64.to_type } /- [traits::h3]: forward function -/ def h3 (T1 T2 : Type) (inst : OfType T1) (inst0 : ToType T2 T1) (y : T2) : Result T1 := inst.of_type T2 inst0 y /- [traits::h4]: forward function -/ def h4 (T1 T2 : Type) (inst : OfTypeBis T1 T2) (inst0 : ToType T2 T1) (y : T2) : Result T1 := inst.of_type y /- [traits::TestType::{6}::test::TestType1::{0}::test]: forward function -/ def TestType.test.TestType1.test (self : TestType.test.TestType1) : Result Bool := Result.ret (self._0 > (U64.ofInt 1)) /- Trait implementation: [traits::TestType::{6}::test::TestType1::{0}] -/ def TestType.test.TestType1.TestTypetestTestTraitInst : TestType.test.TestTrait TestType.test.TestType1 := { test := TestType.test.TestType1.test } /- [traits::TestType::{6}::test]: forward function -/ def TestType.test (T : Type) (inst : ToU64 T) (self : TestType T) (x : T) : Result Bool := do let x0 ← inst.to_u64 x if x0 > (U64.ofInt 0) then TestType.test.TestType1.test { _0 := (U64.ofInt 0) } else Result.ret false /- [traits::BoolWrapper::{7}::to_type]: forward function -/ def BoolWrapper.to_type (T : Type) (inst : ToType Bool T) (self : BoolWrapper) : Result T := inst.to_type self._0 /- Trait implementation: [traits::BoolWrapper::{7}] -/ def BoolWrapper.ToTypeInst (T : Type) (inst : ToType Bool T) : ToType BoolWrapper T := { to_type := BoolWrapper.to_type T inst } /- [traits::WithConstTy::LEN2] -/ def with_const_ty_len2_body : Result Usize := Result.ret (Usize.ofInt 32) def with_const_ty_len2_c : Usize := eval_global with_const_ty_len2_body (by simp) /- [traits::Bool::{8}::LEN1] -/ def bool_len1_body : Result Usize := Result.ret (Usize.ofInt 12) def bool_len1_c : Usize := eval_global bool_len1_body (by simp) /- [traits::Bool::{8}::f]: merged forward/backward function (there is a single backward function, and the forward function returns ()) -/ def Bool.f (i : U64) (a : Array U8 (Usize.ofInt 32)) : Result U64 := Result.ret i /- Trait implementation: [traits::Bool::{8}] -/ def Bool.WithConstTyInst : WithConstTy Bool (Usize.ofInt 32) := { LEN1 := bool_len1_c LEN2 := with_const_ty_len2_c V := U8 W := U64 W_clause_0 := u64.ToU64Inst f := Bool.f } /- [traits::use_with_const_ty1]: forward function -/ def use_with_const_ty1 (H : Type) (LEN : Usize) (inst : WithConstTy H LEN) : Result Usize := let i := inst.LEN1 Result.ret i /- [traits::use_with_const_ty2]: forward function -/ def use_with_const_ty2 (H : Type) (LEN : Usize) (inst : WithConstTy H LEN) (w : inst.W) : Result Unit := Result.ret () /- [traits::use_with_const_ty3]: forward function -/ def use_with_const_ty3 (H : Type) (LEN : Usize) (inst : WithConstTy H LEN) (x : inst.W) : Result U64 := inst.W_clause_0.to_u64 x /- [traits::test_where1]: forward function -/ def test_where1 (T : Type) (_x : T) : Result Unit := Result.ret () /- [traits::test_where2]: forward function -/ def test_where2 (T : Type) (inst : WithConstTy T (Usize.ofInt 32)) (_x : U32) : Result Unit := Result.ret () /- [traits::test_child_trait1]: forward function -/ def test_child_trait1 (T : Type) (inst : ChildTrait T) (x : T) : Result alloc.string.String := inst.parent_clause_0.get_name x /- [traits::test_child_trait2]: forward function -/ def test_child_trait2 (T : Type) (inst : ChildTrait T) (x : T) : Result inst.parent_clause_0.W := inst.parent_clause_0.get_w x /- [traits::order1]: forward function -/ def order1 (T U : Type) (inst : ParentTrait0 T) (inst0 : ParentTrait0 U) : Result Unit := Result.ret () end traits