From e684179ff9c582495b720c3b1310e0648ca24a58 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Fri, 6 Oct 2023 12:13:38 +0200 Subject: Generate the Traits test files for Lean --- tests/lean/Traits.lean | 1 + 1 file changed, 1 insertion(+) create mode 100644 tests/lean/Traits.lean (limited to 'tests/lean/Traits.lean') diff --git a/tests/lean/Traits.lean b/tests/lean/Traits.lean new file mode 100644 index 00000000..5e812e95 --- /dev/null +++ b/tests/lean/Traits.lean @@ -0,0 +1 @@ +import Traits.Funs -- cgit v1.2.3 From c57dec640d4e12c3dc66969d626bbbca2eb733fd Mon Sep 17 00:00:00 2001 From: Son Ho Date: Thu, 9 Nov 2023 11:43:47 +0100 Subject: Modify some options and update the Makefile --- tests/lean/Traits.lean | 411 ++++++++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 410 insertions(+), 1 deletion(-) (limited to 'tests/lean/Traits.lean') diff --git a/tests/lean/Traits.lean b/tests/lean/Traits.lean index 5e812e95..94ae0bb0 100644 --- a/tests/lean/Traits.lean +++ b/tests/lean/Traits.lean @@ -1 +1,410 @@ -import Traits.Funs +-- THIS FILE WAS AUTOMATICALLY GENERATED BY AENEAS +-- [traits] +import Base +open Primitives + +namespace traits + +/- Trait declaration: [traits::BoolTrait] -/ +structure BoolTrait (Self : Type) where + get_bool : Self → Result Bool + +/- [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 + | none => Result.ret false + | 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 + +/- Trait declaration: [traits::ToU64] -/ +structure ToU64 (Self : Type) where + to_u64 : Self → Result U64 + +/- [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] -/ +structure Wrapper (T : Type) where + x : T + +/- [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 + +/- Trait declaration: [traits::ToType] -/ +structure ToType (Self T : Type) where + to_type : Self → Result T + +/- [traits::u64::{5}::to_type]: forward function -/ +def u64.to_type (self : U64) : Result Bool := + Result.ret (self > 0#u64) + +/- Trait implementation: [traits::u64::{5}] -/ +def u64.ToTypeInst : ToType U64 Bool := { + to_type := u64.to_type +} + +/- Trait declaration: [traits::OfType] -/ +structure OfType (Self : Type) where + of_type : forall (T : Type) (inst : ToType T Self), T → Result Self + +/- [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 + +/- Trait declaration: [traits::OfTypeBis] -/ +structure OfTypeBis (Self T : Type) where + parent_clause_0 : ToType T Self + of_type : T → Result Self + +/- [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] -/ +structure TestType (T : Type) where + _0 : T + +/- [traits::TestType::{6}::test::TestType1] -/ +structure TestType.test.TestType1 where + _0 : U64 + +/- Trait declaration: [traits::TestType::{6}::test::TestTrait] -/ +structure TestType.test.TestTrait (Self : Type) where + test : Self → Result Bool + +/- [traits::TestType::{6}::test::TestType1::{0}::test]: forward function -/ +def TestType.test.TestType1.test + (self : TestType.test.TestType1) : Result Bool := + Result.ret (self._0 > 1#u64) + +/- 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 > 0#u64 + then TestType.test.TestType1.test { _0 := 0#u64 } + else Result.ret false + +/- [traits::BoolWrapper] -/ +structure BoolWrapper where + _0 : Bool + +/- [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 32#usize +def with_const_ty_len2_c : Usize := + eval_global with_const_ty_len2_body (by simp) + +/- Trait declaration: [traits::WithConstTy] -/ +structure WithConstTy (Self : Type) (LEN : Usize) where + LEN1 : Usize + LEN2 : Usize + V : Type + W : Type + W_clause_0 : ToU64 W + f : W → Array U8 LEN → Result W + +/- [traits::Bool::{8}::LEN1] -/ +def bool_len1_body : Result Usize := Result.ret 12#usize +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 32#usize) : Result U64 := + Result.ret i + +/- Trait implementation: [traits::Bool::{8}] -/ +def Bool.WithConstTyInst : WithConstTy Bool 32#usize := { + 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 32#usize) (_x : U32) : Result Unit := + Result.ret () + +/- [alloc::string::String] -/ +axiom alloc.string.String : Type + +/- Trait declaration: [traits::ParentTrait0] -/ +structure ParentTrait0 (Self : Type) where + W : Type + get_name : Self → Result alloc.string.String + get_w : Self → Result W + +/- Trait declaration: [traits::ParentTrait1] -/ +structure ParentTrait1 (Self : Type) where + +/- Trait declaration: [traits::ChildTrait] -/ +structure ChildTrait (Self : Type) where + parent_clause_0 : ParentTrait0 Self + parent_clause_1 : ParentTrait1 Self + +/- [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 () + +/- Trait declaration: [traits::ChildTrait1] -/ +structure ChildTrait1 (Self : Type) where + parent_clause_0 : ParentTrait1 Self + +/- Trait implementation: [traits::usize::{9}] -/ +def usize.ParentTrait1Inst : ParentTrait1 Usize := { +} + +/- Trait implementation: [traits::usize::{10}] -/ +def usize.ChildTrait1Inst : ChildTrait1 Usize := { + parent_clause_0 := usize.ParentTrait1Inst +} + +/- Trait declaration: [traits::Iterator] -/ +structure Iterator (Self : Type) where + Item : Type + +/- Trait declaration: [traits::IntoIterator] -/ +structure IntoIterator (Self : Type) where + Item : Type + IntoIter : Type + IntoIter_clause_0 : Iterator IntoIter + into_iter : Self → Result IntoIter + +/- Trait declaration: [traits::FromResidual] -/ +structure FromResidual (Self T : Type) where + +/- Trait declaration: [traits::Try] -/ +structure Try (Self : Type) where + parent_clause_0 : FromResidual Self Residual + Residual : Type + +/- Trait declaration: [traits::CFnOnce] -/ +structure CFnOnce (Self Args : Type) where + Output : Type + call_once : Self → Args → Result Output + +/- Trait declaration: [traits::CFnMut] -/ +structure CFnMut (Self Args : Type) where + parent_clause_0 : CFnOnce Self Args + call_mut : Self → Args → Result parent_clause_0.Output + call_mut_back : Self → Args → parent_clause_0.Output → Result Self + +/- Trait declaration: [traits::CFn] -/ +structure CFn (Self Args : Type) where + parent_clause_0 : CFnMut Self Args + call_mut : Self → Args → Result parent_clause_0.parent_clause_0.Output + +/- Trait declaration: [core::ops::function::FnOnce] -/ +structure core.ops.function.FnOnce (Self Args : Type) where + Output : Type + call_once : Self → Args → Result Output + +/- Trait declaration: [core::ops::function::FnMut] -/ +structure core.ops.function.FnMut (Self Args : Type) where + parent_clause_0 : core.ops.function.FnOnce Self Args + call_mut : Self → Args → Result parent_clause_0.Output + call_mut_back : Self → Args → parent_clause_0.Output → Result Self + +/- Trait declaration: [core::ops::function::Fn] -/ +structure core.ops.function.Fn (Self Args : Type) where + parent_clause_0 : core.ops.function.FnMut Self Args + call : Self → Args → Result parent_clause_0.parent_clause_0.Output + +/- [traits::map_option]: forward function -/ +def map_option + (T F : Type) (inst : core.ops.function.Fn F T) (x : Option T) (f0 : F) : + Result (Option T) + := + match x with + | none => Result.ret none + | some x0 => do + let t ← inst.call f0 x0 + Result.ret (some t) + +/- Trait declaration: [traits::WithTarget] -/ +structure WithTarget (Self : Type) where + Target : Type + +/- Trait declaration: [traits::ParentTrait2] -/ +structure ParentTrait2 (Self : Type) where + U : Type + U_clause_0 : WithTarget U + +/- Trait declaration: [traits::ChildTrait2] -/ +structure ChildTrait2 (Self : Type) where + parent_clause_0 : ParentTrait2 Self + convert : parent_clause_0.U → Result parent_clause_0.U_clause_0.Target + +/- Trait implementation: [traits::u32::{11}] -/ +def u32.WithTargetInst : WithTarget U32 := { + Target := U32 +} + +/- Trait implementation: [traits::u32::{12}] -/ +def u32.ParentTrait2Inst : ParentTrait2 U32 := { + U := U32 + U_clause_0 := u32.WithTargetInst +} + +/- [traits::u32::{13}::convert]: forward function -/ +def u32.convert (x : U32) : Result U32 := + Result.ret x + +/- Trait implementation: [traits::u32::{13}] -/ +def u32.ChildTrait2Inst : ChildTrait2 U32 := { + parent_clause_0 := u32.ParentTrait2Inst + convert := u32.convert +} + +/- [traits::incr_u32]: forward function -/ +def incr_u32 (x : U32) : Result U32 := + x + 1#u32 + +end traits -- cgit v1.2.3 From 38a01551719d0cefdd4a1c21a050674154b0b087 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Thu, 9 Nov 2023 13:44:03 +0100 Subject: Regenerate the tests --- tests/lean/Traits.lean | 61 ++++++++++++++------------------------------------ 1 file changed, 17 insertions(+), 44 deletions(-) (limited to 'tests/lean/Traits.lean') diff --git a/tests/lean/Traits.lean b/tests/lean/Traits.lean index 94ae0bb0..12e7eafa 100644 --- a/tests/lean/Traits.lean +++ b/tests/lean/Traits.lean @@ -322,51 +322,8 @@ structure FromResidual (Self T : Type) where /- Trait declaration: [traits::Try] -/ structure Try (Self : Type) where - parent_clause_0 : FromResidual Self Residual Residual : Type - -/- Trait declaration: [traits::CFnOnce] -/ -structure CFnOnce (Self Args : Type) where - Output : Type - call_once : Self → Args → Result Output - -/- Trait declaration: [traits::CFnMut] -/ -structure CFnMut (Self Args : Type) where - parent_clause_0 : CFnOnce Self Args - call_mut : Self → Args → Result parent_clause_0.Output - call_mut_back : Self → Args → parent_clause_0.Output → Result Self - -/- Trait declaration: [traits::CFn] -/ -structure CFn (Self Args : Type) where - parent_clause_0 : CFnMut Self Args - call_mut : Self → Args → Result parent_clause_0.parent_clause_0.Output - -/- Trait declaration: [core::ops::function::FnOnce] -/ -structure core.ops.function.FnOnce (Self Args : Type) where - Output : Type - call_once : Self → Args → Result Output - -/- Trait declaration: [core::ops::function::FnMut] -/ -structure core.ops.function.FnMut (Self Args : Type) where - parent_clause_0 : core.ops.function.FnOnce Self Args - call_mut : Self → Args → Result parent_clause_0.Output - call_mut_back : Self → Args → parent_clause_0.Output → Result Self - -/- Trait declaration: [core::ops::function::Fn] -/ -structure core.ops.function.Fn (Self Args : Type) where - parent_clause_0 : core.ops.function.FnMut Self Args - call : Self → Args → Result parent_clause_0.parent_clause_0.Output - -/- [traits::map_option]: forward function -/ -def map_option - (T F : Type) (inst : core.ops.function.Fn F T) (x : Option T) (f0 : F) : - Result (Option T) - := - match x with - | none => Result.ret none - | some x0 => do - let t ← inst.call f0 x0 - Result.ret (some t) + parent_clause_0 : FromResidual Self Residual /- Trait declaration: [traits::WithTarget] -/ structure WithTarget (Self : Type) where @@ -407,4 +364,20 @@ def u32.ChildTrait2Inst : ChildTrait2 U32 := { def incr_u32 (x : U32) : Result U32 := x + 1#u32 +/- Trait declaration: [traits::CFnOnce] -/ +structure CFnOnce (Self Args : Type) where + Output : Type + call_once : Self → Args → Result Output + +/- Trait declaration: [traits::CFnMut] -/ +structure CFnMut (Self Args : Type) where + parent_clause_0 : CFnOnce Self Args + call_mut : Self → Args → Result parent_clause_0.Output + call_mut_back : Self → Args → parent_clause_0.Output → Result Self + +/- Trait declaration: [traits::CFn] -/ +structure CFn (Self Args : Type) where + parent_clause_0 : CFnMut Self Args + call_mut : Self → Args → Result parent_clause_0.parent_clause_0.Output + end traits -- cgit v1.2.3 From 46ab0fc047c69d14e52415bd8b76be167771cd58 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 21 Nov 2023 11:46:10 +0100 Subject: Regenerate the Traits files --- tests/lean/Traits.lean | 132 +++++++++++++++++++++++++------------------------ 1 file changed, 67 insertions(+), 65 deletions(-) (limited to 'tests/lean/Traits.lean') diff --git a/tests/lean/Traits.lean b/tests/lean/Traits.lean index 12e7eafa..51a05581 100644 --- a/tests/lean/Traits.lean +++ b/tests/lean/Traits.lean @@ -9,12 +9,12 @@ namespace traits structure BoolTrait (Self : Type) where get_bool : Self → Result Bool -/- [traits::Bool::{0}::get_bool]: forward function -/ +/- [traits::{bool}::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 := { +/- Trait implementation: [traits::{bool}] -/ +def traits.BoolTraitBoolInst : BoolTrait Bool := { get_bool := Bool.get_bool } @@ -28,17 +28,18 @@ 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 + then BoolTrait.ret_true traits.BoolTraitBoolInst x else Result.ret false -/- [traits::Option::{1}::get_bool]: forward function -/ +/- [traits::{core::option::Option#1}::get_bool]: forward function -/ def Option.get_bool (T : Type) (self : Option T) : Result Bool := match self with | none => Result.ret false | some t => Result.ret true -/- Trait implementation: [traits::Option::{1}] -/ -def Option.BoolTraitInst (T : Type) : BoolTrait (Option T) := { +/- Trait implementation: [traits::{core::option::Option#1}] -/ +def traits.BoolTraitcoreoptionOptionTInst (T : Type) : BoolTrait (Option T) + := { get_bool := Option.get_bool T } @@ -47,7 +48,7 @@ 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 + then BoolTrait.ret_true (traits.BoolTraitcoreoptionOptionTInst T) x else Result.ret false /- [traits::test_bool_trait]: forward function -/ @@ -58,31 +59,31 @@ def test_bool_trait (T : Type) (inst : BoolTrait T) (x : T) : Result Bool := structure ToU64 (Self : Type) where to_u64 : Self → Result U64 -/- [traits::u64::{2}::to_u64]: forward function -/ -def u64.to_u64 (self : U64) : Result U64 := +/- [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 +/- Trait implementation: [traits::{u64#2}] -/ +def traits.ToU64U64Inst : 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 := +/- [traits::{(A, A)#3}::to_u64]: forward function -/ +def Pair.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 +/- Trait implementation: [traits::{(A, A)#3}] -/ +def traits.ToU64TupleAAInst (A : Type) (inst : ToU64 A) : ToU64 (A × A) := { + to_u64 := Pair.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 + Pair.to_u64 T inst x /- [traits::g]: forward function -/ def g (T : Type) (inst : ToU64 (T × T)) (x : (T × T)) : Result U64 := @@ -90,25 +91,26 @@ def g (T : Type) (inst : ToU64 (T × T)) (x : (T × T)) : Result U64 := /- [traits::h0]: forward function -/ def h0 (x : U64) : Result U64 := - u64.to_u64 x + U64.to_u64 x /- [traits::Wrapper] -/ structure Wrapper (T : Type) where x : T -/- [traits::Wrapper::{4}::to_u64]: forward function -/ +/- [traits::{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) := { +/- Trait implementation: [traits::{traits::Wrapper#4}] -/ +def traits.ToU64traitsWrapperTInst (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 + Wrapper.to_u64 U64 traits.ToU64U64Inst x /- [traits::h2]: forward function -/ def h2 (T : Type) (inst : ToU64 T) (x : Wrapper T) : Result U64 := @@ -118,13 +120,13 @@ def h2 (T : Type) (inst : ToU64 T) (x : Wrapper T) : Result U64 := structure ToType (Self T : Type) where to_type : Self → Result T -/- [traits::u64::{5}::to_type]: forward function -/ -def u64.to_type (self : U64) : Result Bool := +/- [traits::{u64#5}::to_type]: forward function -/ +def U64.to_type (self : U64) : Result Bool := Result.ret (self > 0#u64) -/- Trait implementation: [traits::u64::{5}] -/ -def u64.ToTypeInst : ToType U64 Bool := { - to_type := u64.to_type +/- Trait implementation: [traits::{u64#5}] -/ +def traits.ToTypeU64BoolInst : ToType U64 Bool := { + to_type := U64.to_type } /- Trait declaration: [traits::OfType] -/ @@ -154,26 +156,26 @@ def h4 structure TestType (T : Type) where _0 : T -/- [traits::TestType::{6}::test::TestType1] -/ +/- [traits::{traits::TestType#6}::test::TestType1] -/ structure TestType.test.TestType1 where _0 : U64 -/- Trait declaration: [traits::TestType::{6}::test::TestTrait] -/ +/- Trait declaration: [traits::{traits::TestType#6}::test::TestTrait] -/ structure TestType.test.TestTrait (Self : Type) where test : Self → Result Bool -/- [traits::TestType::{6}::test::TestType1::{0}::test]: forward function -/ +/- [traits::{traits::TestType#6}::test::{traits::{traits::TestType#6}::test::TestType1}::test]: forward function -/ def TestType.test.TestType1.test (self : TestType.test.TestType1) : Result Bool := Result.ret (self._0 > 1#u64) -/- Trait implementation: [traits::TestType::{6}::test::TestType1::{0}] -/ -def TestType.test.TestType1.TestTypetestTestTraitInst : TestType.test.TestTrait - TestType.test.TestType1 := { +/- Trait implementation: [traits::{traits::TestType#6}::test::{traits::{traits::TestType#6}::test::TestType1}] -/ +def traits.TestType.test.TestTraittraitstraitsTestTypeTtestTestType1Inst : + TestType.test.TestTrait TestType.test.TestType1 := { test := TestType.test.TestType1.test } -/- [traits::TestType::{6}::test]: forward function -/ +/- [traits::{traits::TestType#6}::test]: forward function -/ def TestType.test (T : Type) (inst : ToU64 T) (self : TestType T) (x : T) : Result Bool := do @@ -186,14 +188,14 @@ def TestType.test structure BoolWrapper where _0 : Bool -/- [traits::BoolWrapper::{7}::to_type]: forward function -/ +/- [traits::{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 := { +/- Trait implementation: [traits::{traits::BoolWrapper#7}] -/ +def traits.ToTypetraitsBoolWrapperTInst (T : Type) (inst : ToType Bool T) : + ToType BoolWrapper T := { to_type := BoolWrapper.to_type T inst } @@ -211,22 +213,22 @@ structure WithConstTy (Self : Type) (LEN : Usize) where W_clause_0 : ToU64 W f : W → Array U8 LEN → Result W -/- [traits::Bool::{8}::LEN1] -/ +/- [traits::{bool#8}::LEN1] -/ def bool_len1_body : Result Usize := Result.ret 12#usize def bool_len1_c : Usize := eval_global bool_len1_body (by simp) -/- [traits::Bool::{8}::f]: merged forward/backward function +/- [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 32#usize) : Result U64 := Result.ret i -/- Trait implementation: [traits::Bool::{8}] -/ -def Bool.WithConstTyInst : WithConstTy Bool 32#usize := { +/- Trait implementation: [traits::{bool#8}] -/ +def traits.WithConstTyBool32Inst : WithConstTy Bool 32#usize := { LEN1 := bool_len1_c LEN2 := with_const_ty_len2_c V := U8 W := U64 - W_clause_0 := u64.ToU64Inst + W_clause_0 := traits.ToU64U64Inst f := Bool.f } @@ -297,13 +299,13 @@ def order1 structure ChildTrait1 (Self : Type) where parent_clause_0 : ParentTrait1 Self -/- Trait implementation: [traits::usize::{9}] -/ -def usize.ParentTrait1Inst : ParentTrait1 Usize := { +/- Trait implementation: [traits::{usize#9}] -/ +def traits.ParentTrait1UsizeInst : ParentTrait1 Usize := { } -/- Trait implementation: [traits::usize::{10}] -/ -def usize.ChildTrait1Inst : ChildTrait1 Usize := { - parent_clause_0 := usize.ParentTrait1Inst +/- Trait implementation: [traits::{usize#10}] -/ +def traits.ChildTrait1UsizeInst : ChildTrait1 Usize := { + parent_clause_0 := traits.ParentTrait1UsizeInst } /- Trait declaration: [traits::Iterator] -/ @@ -339,31 +341,27 @@ structure ChildTrait2 (Self : Type) where parent_clause_0 : ParentTrait2 Self convert : parent_clause_0.U → Result parent_clause_0.U_clause_0.Target -/- Trait implementation: [traits::u32::{11}] -/ -def u32.WithTargetInst : WithTarget U32 := { +/- Trait implementation: [traits::{u32#11}] -/ +def traits.WithTargetU32Inst : WithTarget U32 := { Target := U32 } -/- Trait implementation: [traits::u32::{12}] -/ -def u32.ParentTrait2Inst : ParentTrait2 U32 := { +/- Trait implementation: [traits::{u32#12}] -/ +def traits.ParentTrait2U32Inst : ParentTrait2 U32 := { U := U32 - U_clause_0 := u32.WithTargetInst + U_clause_0 := traits.WithTargetU32Inst } -/- [traits::u32::{13}::convert]: forward function -/ -def u32.convert (x : U32) : Result U32 := +/- [traits::{u32#13}::convert]: forward function -/ +def U32.convert (x : U32) : Result U32 := Result.ret x -/- Trait implementation: [traits::u32::{13}] -/ -def u32.ChildTrait2Inst : ChildTrait2 U32 := { - parent_clause_0 := u32.ParentTrait2Inst - convert := u32.convert +/- Trait implementation: [traits::{u32#13}] -/ +def traits.ChildTrait2U32Inst : ChildTrait2 U32 := { + parent_clause_0 := traits.ParentTrait2U32Inst + convert := U32.convert } -/- [traits::incr_u32]: forward function -/ -def incr_u32 (x : U32) : Result U32 := - x + 1#u32 - /- Trait declaration: [traits::CFnOnce] -/ structure CFnOnce (Self Args : Type) where Output : Type @@ -380,4 +378,8 @@ structure CFn (Self Args : Type) where parent_clause_0 : CFnMut Self Args call_mut : Self → Args → Result parent_clause_0.parent_clause_0.Output +/- [traits::incr_u32]: forward function -/ +def incr_u32 (x : U32) : Result U32 := + x + 1#u32 + end traits -- cgit v1.2.3 From 137cc7335e64fcb70c254e7fd2a6fa353fb43e61 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 21 Nov 2023 14:57:38 +0100 Subject: Regenerate the files --- tests/lean/Traits.lean | 219 ++++++++++++++++++++++++++++++++----------------- 1 file changed, 146 insertions(+), 73 deletions(-) (limited to 'tests/lean/Traits.lean') diff --git a/tests/lean/Traits.lean b/tests/lean/Traits.lean index 51a05581..f854da76 100644 --- a/tests/lean/Traits.lean +++ b/tests/lean/Traits.lean @@ -5,25 +5,30 @@ open Primitives namespace traits -/- Trait declaration: [traits::BoolTrait] -/ +/- Trait declaration: [traits::BoolTrait] + Source: 'src/traits.rs', lines 1:0-1:19 -/ structure BoolTrait (Self : Type) where get_bool : Self → Result Bool -/- [traits::{bool}::get_bool]: forward function -/ +/- [traits::{bool}::get_bool]: forward function + Source: 'src/traits.rs', lines 12:4-12:30 -/ def Bool.get_bool (self : Bool) : Result Bool := Result.ret self -/- Trait implementation: [traits::{bool}] -/ +/- Trait implementation: [traits::{bool}] + Source: 'src/traits.rs', lines 11:0-11:23 -/ def traits.BoolTraitBoolInst : BoolTrait Bool := { get_bool := Bool.get_bool } -/- [traits::BoolTrait::ret_true]: forward function -/ +/- [traits::BoolTrait::ret_true]: forward function + Source: 'src/traits.rs', lines 6:4-6:30 -/ def BoolTrait.ret_true {Self : Type} (self_clause : BoolTrait Self) (self : Self) : Result Bool := Result.ret true -/- [traits::test_bool_trait_bool]: forward function -/ +/- [traits::test_bool_trait_bool]: forward function + Source: 'src/traits.rs', lines 17:0-17:44 -/ def test_bool_trait_bool (x : Bool) : Result Bool := do let b ← Bool.get_bool x @@ -31,19 +36,22 @@ def test_bool_trait_bool (x : Bool) : Result Bool := then BoolTrait.ret_true traits.BoolTraitBoolInst x else Result.ret false -/- [traits::{core::option::Option#1}::get_bool]: forward function -/ +/- [traits::{core::option::Option#1}::get_bool]: forward function + Source: 'src/traits.rs', lines 23:4-23:30 -/ def Option.get_bool (T : Type) (self : Option T) : Result Bool := match self with | none => Result.ret false | some t => Result.ret true -/- Trait implementation: [traits::{core::option::Option#1}] -/ +/- Trait implementation: [traits::{core::option::Option#1}] + Source: 'src/traits.rs', lines 22:0-22:31 -/ def traits.BoolTraitcoreoptionOptionTInst (T : Type) : BoolTrait (Option T) := { get_bool := Option.get_bool T } -/- [traits::test_bool_trait_option]: forward function -/ +/- [traits::test_bool_trait_option]: forward function + Source: 'src/traits.rs', lines 31:0-31:54 -/ def test_bool_trait_option (T : Type) (x : Option T) : Result Bool := do let b ← Option.get_bool T x @@ -51,24 +59,29 @@ def test_bool_trait_option (T : Type) (x : Option T) : Result Bool := then BoolTrait.ret_true (traits.BoolTraitcoreoptionOptionTInst T) x else Result.ret false -/- [traits::test_bool_trait]: forward function -/ +/- [traits::test_bool_trait]: forward function + Source: 'src/traits.rs', lines 35:0-35:50 -/ def test_bool_trait (T : Type) (inst : BoolTrait T) (x : T) : Result Bool := inst.get_bool x -/- Trait declaration: [traits::ToU64] -/ +/- Trait declaration: [traits::ToU64] + Source: 'src/traits.rs', lines 39:0-39:15 -/ structure ToU64 (Self : Type) where to_u64 : Self → Result U64 -/- [traits::{u64#2}::to_u64]: forward function -/ +/- [traits::{u64#2}::to_u64]: forward function + Source: 'src/traits.rs', lines 44:4-44:26 -/ def U64.to_u64 (self : U64) : Result U64 := Result.ret self -/- Trait implementation: [traits::{u64#2}] -/ +/- Trait implementation: [traits::{u64#2}] + Source: 'src/traits.rs', lines 43:0-43:18 -/ def traits.ToU64U64Inst : ToU64 U64 := { to_u64 := U64.to_u64 } -/- [traits::{(A, A)#3}::to_u64]: forward function -/ +/- [traits::{(A, A)#3}::to_u64]: forward function + Source: 'src/traits.rs', lines 50:4-50:26 -/ def Pair.to_u64 (A : Type) (inst : ToU64 A) (self : (A × A)) : Result U64 := do let (t, t0) := self @@ -76,106 +89,128 @@ def Pair.to_u64 (A : Type) (inst : ToU64 A) (self : (A × A)) : Result U64 := let i0 ← inst.to_u64 t0 i + i0 -/- Trait implementation: [traits::{(A, A)#3}] -/ +/- Trait implementation: [traits::{(A, A)#3}] + Source: 'src/traits.rs', lines 49:0-49:31 -/ def traits.ToU64TupleAAInst (A : Type) (inst : ToU64 A) : ToU64 (A × A) := { to_u64 := Pair.to_u64 A inst } -/- [traits::f]: forward function -/ +/- [traits::f]: forward function + Source: 'src/traits.rs', lines 55:0-55:36 -/ def f (T : Type) (inst : ToU64 T) (x : (T × T)) : Result U64 := Pair.to_u64 T inst x -/- [traits::g]: forward function -/ +/- [traits::g]: forward function + Source: 'src/traits.rs', lines 59:0-61:18 -/ def g (T : Type) (inst : ToU64 (T × T)) (x : (T × T)) : Result U64 := inst.to_u64 x -/- [traits::h0]: forward function -/ +/- [traits::h0]: forward function + Source: 'src/traits.rs', lines 66:0-66:24 -/ def h0 (x : U64) : Result U64 := U64.to_u64 x -/- [traits::Wrapper] -/ +/- [traits::Wrapper] + Source: 'src/traits.rs', lines 70:0-70:21 -/ structure Wrapper (T : Type) where x : T -/- [traits::{traits::Wrapper#4}::to_u64]: forward function -/ +/- [traits::{traits::Wrapper#4}::to_u64]: forward function + Source: 'src/traits.rs', lines 75:4-75:26 -/ def Wrapper.to_u64 (T : Type) (inst : ToU64 T) (self : Wrapper T) : Result U64 := inst.to_u64 self.x -/- Trait implementation: [traits::{traits::Wrapper#4}] -/ +/- Trait implementation: [traits::{traits::Wrapper#4}] + Source: 'src/traits.rs', lines 74:0-74:35 -/ def traits.ToU64traitsWrapperTInst (T : Type) (inst : ToU64 T) : ToU64 (Wrapper T) := { to_u64 := Wrapper.to_u64 T inst } -/- [traits::h1]: forward function -/ +/- [traits::h1]: forward function + Source: 'src/traits.rs', lines 80:0-80:33 -/ def h1 (x : Wrapper U64) : Result U64 := Wrapper.to_u64 U64 traits.ToU64U64Inst x -/- [traits::h2]: forward function -/ +/- [traits::h2]: forward function + Source: 'src/traits.rs', lines 84:0-84:41 -/ def h2 (T : Type) (inst : ToU64 T) (x : Wrapper T) : Result U64 := Wrapper.to_u64 T inst x -/- Trait declaration: [traits::ToType] -/ +/- Trait declaration: [traits::ToType] + Source: 'src/traits.rs', lines 88:0-88:19 -/ structure ToType (Self T : Type) where to_type : Self → Result T -/- [traits::{u64#5}::to_type]: forward function -/ +/- [traits::{u64#5}::to_type]: forward function + Source: 'src/traits.rs', lines 93:4-93:28 -/ def U64.to_type (self : U64) : Result Bool := Result.ret (self > 0#u64) -/- Trait implementation: [traits::{u64#5}] -/ +/- Trait implementation: [traits::{u64#5}] + Source: 'src/traits.rs', lines 92:0-92:25 -/ def traits.ToTypeU64BoolInst : ToType U64 Bool := { to_type := U64.to_type } -/- Trait declaration: [traits::OfType] -/ +/- Trait declaration: [traits::OfType] + Source: 'src/traits.rs', lines 98:0-98:16 -/ structure OfType (Self : Type) where of_type : forall (T : Type) (inst : ToType T Self), T → Result Self -/- [traits::h3]: forward function -/ +/- [traits::h3]: forward function + Source: 'src/traits.rs', lines 104:0-104:50 -/ def h3 (T1 T2 : Type) (inst : OfType T1) (inst0 : ToType T2 T1) (y : T2) : Result T1 := inst.of_type T2 inst0 y -/- Trait declaration: [traits::OfTypeBis] -/ +/- Trait declaration: [traits::OfTypeBis] + Source: 'src/traits.rs', lines 109:0-109:36 -/ structure OfTypeBis (Self T : Type) where parent_clause_0 : ToType T Self of_type : T → Result Self -/- [traits::h4]: forward function -/ +/- [traits::h4]: forward function + Source: 'src/traits.rs', lines 118:0-118:57 -/ def h4 (T1 T2 : Type) (inst : OfTypeBis T1 T2) (inst0 : ToType T2 T1) (y : T2) : Result T1 := inst.of_type y -/- [traits::TestType] -/ +/- [traits::TestType] + Source: 'src/traits.rs', lines 122:0-122:22 -/ structure TestType (T : Type) where _0 : T -/- [traits::{traits::TestType#6}::test::TestType1] -/ +/- [traits::{traits::TestType#6}::test::TestType1] + Source: 'src/traits.rs', lines 127:8-127:24 -/ structure TestType.test.TestType1 where _0 : U64 -/- Trait declaration: [traits::{traits::TestType#6}::test::TestTrait] -/ +/- Trait declaration: [traits::{traits::TestType#6}::test::TestTrait] + Source: 'src/traits.rs', lines 128:8-128:23 -/ structure TestType.test.TestTrait (Self : Type) where test : Self → Result Bool -/- [traits::{traits::TestType#6}::test::{traits::{traits::TestType#6}::test::TestType1}::test]: forward function -/ +/- [traits::{traits::TestType#6}::test::{traits::{traits::TestType#6}::test::TestType1}::test]: forward function + Source: 'src/traits.rs', lines 139:12-139:34 -/ def TestType.test.TestType1.test (self : TestType.test.TestType1) : Result Bool := Result.ret (self._0 > 1#u64) -/- Trait implementation: [traits::{traits::TestType#6}::test::{traits::{traits::TestType#6}::test::TestType1}] -/ +/- Trait implementation: [traits::{traits::TestType#6}::test::{traits::{traits::TestType#6}::test::TestType1}] + Source: 'src/traits.rs', lines 138:8-138:36 -/ def traits.TestType.test.TestTraittraitstraitsTestTypeTtestTestType1Inst : TestType.test.TestTrait TestType.test.TestType1 := { test := TestType.test.TestType1.test } -/- [traits::{traits::TestType#6}::test]: forward function -/ +/- [traits::{traits::TestType#6}::test]: forward function + Source: 'src/traits.rs', lines 126:4-126:36 -/ def TestType.test (T : Type) (inst : ToU64 T) (self : TestType T) (x : T) : Result Bool := do @@ -184,27 +219,32 @@ def TestType.test then TestType.test.TestType1.test { _0 := 0#u64 } else Result.ret false -/- [traits::BoolWrapper] -/ +/- [traits::BoolWrapper] + Source: 'src/traits.rs', lines 150:0-150:22 -/ structure BoolWrapper where _0 : Bool -/- [traits::{traits::BoolWrapper#7}::to_type]: forward function -/ +/- [traits::{traits::BoolWrapper#7}::to_type]: forward function + Source: 'src/traits.rs', lines 156:4-156:25 -/ def BoolWrapper.to_type (T : Type) (inst : ToType Bool T) (self : BoolWrapper) : Result T := inst.to_type self._0 -/- Trait implementation: [traits::{traits::BoolWrapper#7}] -/ +/- Trait implementation: [traits::{traits::BoolWrapper#7}] + Source: 'src/traits.rs', lines 152:0-152:33 -/ def traits.ToTypetraitsBoolWrapperTInst (T : Type) (inst : ToType Bool T) : ToType BoolWrapper T := { to_type := BoolWrapper.to_type T inst } -/- [traits::WithConstTy::LEN2] -/ +/- [traits::WithConstTy::LEN2] + Source: 'src/traits.rs', lines 164:4-164:21 -/ def with_const_ty_len2_body : Result Usize := Result.ret 32#usize def with_const_ty_len2_c : Usize := eval_global with_const_ty_len2_body (by simp) -/- Trait declaration: [traits::WithConstTy] -/ +/- Trait declaration: [traits::WithConstTy] + Source: 'src/traits.rs', lines 161:0-161:39 -/ structure WithConstTy (Self : Type) (LEN : Usize) where LEN1 : Usize LEN2 : Usize @@ -213,16 +253,19 @@ structure WithConstTy (Self : Type) (LEN : Usize) where W_clause_0 : ToU64 W f : W → Array U8 LEN → Result W -/- [traits::{bool#8}::LEN1] -/ +/- [traits::{bool#8}::LEN1] + Source: 'src/traits.rs', lines 175:4-175:21 -/ def bool_len1_body : Result Usize := Result.ret 12#usize 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 ()) -/ + (there is a single backward function, and the forward function returns ()) + Source: 'src/traits.rs', lines 180:4-180:39 -/ def Bool.f (i : U64) (a : Array U8 32#usize) : Result U64 := Result.ret i -/- Trait implementation: [traits::{bool#8}] -/ +/- Trait implementation: [traits::{bool#8}] + Source: 'src/traits.rs', lines 174:0-174:29 -/ def traits.WithConstTyBool32Inst : WithConstTy Bool 32#usize := { LEN1 := bool_len1_c LEN2 := with_const_ty_len2_c @@ -232,153 +275,183 @@ def traits.WithConstTyBool32Inst : WithConstTy Bool 32#usize := { f := Bool.f } -/- [traits::use_with_const_ty1]: forward function -/ +/- [traits::use_with_const_ty1]: forward function + Source: 'src/traits.rs', lines 183:0-183:75 -/ 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 -/ +/- [traits::use_with_const_ty2]: forward function + Source: 'src/traits.rs', lines 187:0-187:73 -/ 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 -/ +/- [traits::use_with_const_ty3]: forward function + Source: 'src/traits.rs', lines 189:0-189:80 -/ 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 -/ +/- [traits::test_where1]: forward function + Source: 'src/traits.rs', lines 193:0-193:40 -/ def test_where1 (T : Type) (_x : T) : Result Unit := Result.ret () -/- [traits::test_where2]: forward function -/ +/- [traits::test_where2]: forward function + Source: 'src/traits.rs', lines 194:0-194:57 -/ def test_where2 (T : Type) (inst : WithConstTy T 32#usize) (_x : U32) : Result Unit := Result.ret () -/- [alloc::string::String] -/ +/- [alloc::string::String] + Source: '/rustc/d59363ad0b6391b7fc5bbb02c9ccf9300eef3753/library/alloc/src/string.rs', lines 365:0-365:17 -/ axiom alloc.string.String : Type -/- Trait declaration: [traits::ParentTrait0] -/ +/- Trait declaration: [traits::ParentTrait0] + Source: 'src/traits.rs', lines 200:0-200:22 -/ structure ParentTrait0 (Self : Type) where W : Type get_name : Self → Result alloc.string.String get_w : Self → Result W -/- Trait declaration: [traits::ParentTrait1] -/ +/- Trait declaration: [traits::ParentTrait1] + Source: 'src/traits.rs', lines 205:0-205:22 -/ structure ParentTrait1 (Self : Type) where -/- Trait declaration: [traits::ChildTrait] -/ +/- Trait declaration: [traits::ChildTrait] + Source: 'src/traits.rs', lines 206:0-206:49 -/ structure ChildTrait (Self : Type) where parent_clause_0 : ParentTrait0 Self parent_clause_1 : ParentTrait1 Self -/- [traits::test_child_trait1]: forward function -/ +/- [traits::test_child_trait1]: forward function + Source: 'src/traits.rs', lines 209:0-209:56 -/ 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 -/ +/- [traits::test_child_trait2]: forward function + Source: 'src/traits.rs', lines 213:0-213:54 -/ 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 -/ +/- [traits::order1]: forward function + Source: 'src/traits.rs', lines 219:0-219:59 -/ def order1 (T U : Type) (inst : ParentTrait0 T) (inst0 : ParentTrait0 U) : Result Unit := Result.ret () -/- Trait declaration: [traits::ChildTrait1] -/ +/- Trait declaration: [traits::ChildTrait1] + Source: 'src/traits.rs', lines 222:0-222:35 -/ structure ChildTrait1 (Self : Type) where parent_clause_0 : ParentTrait1 Self -/- Trait implementation: [traits::{usize#9}] -/ +/- Trait implementation: [traits::{usize#9}] + Source: 'src/traits.rs', lines 224:0-224:27 -/ def traits.ParentTrait1UsizeInst : ParentTrait1 Usize := { } -/- Trait implementation: [traits::{usize#10}] -/ +/- Trait implementation: [traits::{usize#10}] + Source: 'src/traits.rs', lines 225:0-225:26 -/ def traits.ChildTrait1UsizeInst : ChildTrait1 Usize := { parent_clause_0 := traits.ParentTrait1UsizeInst } -/- Trait declaration: [traits::Iterator] -/ +/- Trait declaration: [traits::Iterator] + Source: 'src/traits.rs', lines 229:0-229:18 -/ structure Iterator (Self : Type) where Item : Type -/- Trait declaration: [traits::IntoIterator] -/ +/- Trait declaration: [traits::IntoIterator] + Source: 'src/traits.rs', lines 233:0-233:22 -/ structure IntoIterator (Self : Type) where Item : Type IntoIter : Type IntoIter_clause_0 : Iterator IntoIter into_iter : Self → Result IntoIter -/- Trait declaration: [traits::FromResidual] -/ +/- Trait declaration: [traits::FromResidual] + Source: 'src/traits.rs', lines 250:0-250:21 -/ structure FromResidual (Self T : Type) where -/- Trait declaration: [traits::Try] -/ +/- Trait declaration: [traits::Try] + Source: 'src/traits.rs', lines 246:0-246:48 -/ structure Try (Self : Type) where Residual : Type parent_clause_0 : FromResidual Self Residual -/- Trait declaration: [traits::WithTarget] -/ +/- Trait declaration: [traits::WithTarget] + Source: 'src/traits.rs', lines 252:0-252:20 -/ structure WithTarget (Self : Type) where Target : Type -/- Trait declaration: [traits::ParentTrait2] -/ +/- Trait declaration: [traits::ParentTrait2] + Source: 'src/traits.rs', lines 256:0-256:22 -/ structure ParentTrait2 (Self : Type) where U : Type U_clause_0 : WithTarget U -/- Trait declaration: [traits::ChildTrait2] -/ +/- Trait declaration: [traits::ChildTrait2] + Source: 'src/traits.rs', lines 260:0-260:35 -/ structure ChildTrait2 (Self : Type) where parent_clause_0 : ParentTrait2 Self convert : parent_clause_0.U → Result parent_clause_0.U_clause_0.Target -/- Trait implementation: [traits::{u32#11}] -/ +/- Trait implementation: [traits::{u32#11}] + Source: 'src/traits.rs', lines 264:0-264:23 -/ def traits.WithTargetU32Inst : WithTarget U32 := { Target := U32 } -/- Trait implementation: [traits::{u32#12}] -/ +/- Trait implementation: [traits::{u32#12}] + Source: 'src/traits.rs', lines 268:0-268:25 -/ def traits.ParentTrait2U32Inst : ParentTrait2 U32 := { U := U32 U_clause_0 := traits.WithTargetU32Inst } -/- [traits::{u32#13}::convert]: forward function -/ +/- [traits::{u32#13}::convert]: forward function + Source: 'src/traits.rs', lines 273:4-273:29 -/ def U32.convert (x : U32) : Result U32 := Result.ret x -/- Trait implementation: [traits::{u32#13}] -/ +/- Trait implementation: [traits::{u32#13}] + Source: 'src/traits.rs', lines 272:0-272:24 -/ def traits.ChildTrait2U32Inst : ChildTrait2 U32 := { parent_clause_0 := traits.ParentTrait2U32Inst convert := U32.convert } -/- Trait declaration: [traits::CFnOnce] -/ +/- Trait declaration: [traits::CFnOnce] + Source: 'src/traits.rs', lines 286:0-286:23 -/ structure CFnOnce (Self Args : Type) where Output : Type call_once : Self → Args → Result Output -/- Trait declaration: [traits::CFnMut] -/ +/- Trait declaration: [traits::CFnMut] + Source: 'src/traits.rs', lines 292:0-292:37 -/ structure CFnMut (Self Args : Type) where parent_clause_0 : CFnOnce Self Args call_mut : Self → Args → Result parent_clause_0.Output call_mut_back : Self → Args → parent_clause_0.Output → Result Self -/- Trait declaration: [traits::CFn] -/ +/- Trait declaration: [traits::CFn] + Source: 'src/traits.rs', lines 296:0-296:33 -/ structure CFn (Self Args : Type) where parent_clause_0 : CFnMut Self Args call_mut : Self → Args → Result parent_clause_0.parent_clause_0.Output -/- [traits::incr_u32]: forward function -/ +/- [traits::incr_u32]: forward function + Source: 'src/traits.rs', lines 300:0-300:30 -/ def incr_u32 (x : U32) : Result U32 := x + 1#u32 -- cgit v1.2.3 From 184e27bce209f7a852c2adc7e0598ed75ac8452d Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 21 Nov 2023 18:58:04 +0100 Subject: Regenerate the files --- tests/lean/Traits.lean | 38 +++++++++++++++++++++----------------- 1 file changed, 21 insertions(+), 17 deletions(-) (limited to 'tests/lean/Traits.lean') diff --git a/tests/lean/Traits.lean b/tests/lean/Traits.lean index f854da76..0aa68c5e 100644 --- a/tests/lean/Traits.lean +++ b/tests/lean/Traits.lean @@ -170,7 +170,7 @@ def h3 /- Trait declaration: [traits::OfTypeBis] Source: 'src/traits.rs', lines 109:0-109:36 -/ structure OfTypeBis (Self T : Type) where - parent_clause_0 : ToType T Self + ToTypeTSelfInst : ToType T Self of_type : T → Result Self /- [traits::h4]: forward function @@ -327,20 +327,22 @@ structure ParentTrait1 (Self : Type) where /- Trait declaration: [traits::ChildTrait] Source: 'src/traits.rs', lines 206:0-206:49 -/ structure ChildTrait (Self : Type) where - parent_clause_0 : ParentTrait0 Self - parent_clause_1 : ParentTrait1 Self + ParentTrait0SelfInst : ParentTrait0 Self + ParentTrait1SelfInst : ParentTrait1 Self /- [traits::test_child_trait1]: forward function Source: 'src/traits.rs', lines 209:0-209:56 -/ def test_child_trait1 (T : Type) (inst : ChildTrait T) (x : T) : Result alloc.string.String := - inst.parent_clause_0.get_name x + inst.ParentTrait0SelfInst.get_name x /- [traits::test_child_trait2]: forward function Source: 'src/traits.rs', lines 213:0-213:54 -/ def test_child_trait2 - (T : Type) (inst : ChildTrait T) (x : T) : Result inst.parent_clause_0.W := - inst.parent_clause_0.get_w x + (T : Type) (inst : ChildTrait T) (x : T) : + Result inst.ParentTrait0SelfInst.W + := + inst.ParentTrait0SelfInst.get_w x /- [traits::order1]: forward function Source: 'src/traits.rs', lines 219:0-219:59 -/ @@ -353,7 +355,7 @@ def order1 /- Trait declaration: [traits::ChildTrait1] Source: 'src/traits.rs', lines 222:0-222:35 -/ structure ChildTrait1 (Self : Type) where - parent_clause_0 : ParentTrait1 Self + ParentTrait1SelfInst : ParentTrait1 Self /- Trait implementation: [traits::{usize#9}] Source: 'src/traits.rs', lines 224:0-224:27 -/ @@ -363,7 +365,7 @@ def traits.ParentTrait1UsizeInst : ParentTrait1 Usize := { /- Trait implementation: [traits::{usize#10}] Source: 'src/traits.rs', lines 225:0-225:26 -/ def traits.ChildTrait1UsizeInst : ChildTrait1 Usize := { - parent_clause_0 := traits.ParentTrait1UsizeInst + ParentTrait1SelfInst := traits.ParentTrait1UsizeInst } /- Trait declaration: [traits::Iterator] @@ -387,7 +389,7 @@ structure FromResidual (Self T : Type) where Source: 'src/traits.rs', lines 246:0-246:48 -/ structure Try (Self : Type) where Residual : Type - parent_clause_0 : FromResidual Self Residual + FromResidualSelftraitsTrySelfResidualInst : FromResidual Self Residual /- Trait declaration: [traits::WithTarget] Source: 'src/traits.rs', lines 252:0-252:20 -/ @@ -403,8 +405,9 @@ structure ParentTrait2 (Self : Type) where /- Trait declaration: [traits::ChildTrait2] Source: 'src/traits.rs', lines 260:0-260:35 -/ structure ChildTrait2 (Self : Type) where - parent_clause_0 : ParentTrait2 Self - convert : parent_clause_0.U → Result parent_clause_0.U_clause_0.Target + ParentTrait2SelfInst : ParentTrait2 Self + convert : ParentTrait2SelfInst.U → Result + ParentTrait2SelfInst.U_clause_0.Target /- Trait implementation: [traits::{u32#11}] Source: 'src/traits.rs', lines 264:0-264:23 -/ @@ -427,7 +430,7 @@ def U32.convert (x : U32) : Result U32 := /- Trait implementation: [traits::{u32#13}] Source: 'src/traits.rs', lines 272:0-272:24 -/ def traits.ChildTrait2U32Inst : ChildTrait2 U32 := { - parent_clause_0 := traits.ParentTrait2U32Inst + ParentTrait2SelfInst := traits.ParentTrait2U32Inst convert := U32.convert } @@ -440,15 +443,16 @@ structure CFnOnce (Self Args : Type) where /- Trait declaration: [traits::CFnMut] Source: 'src/traits.rs', lines 292:0-292:37 -/ structure CFnMut (Self Args : Type) where - parent_clause_0 : CFnOnce Self Args - call_mut : Self → Args → Result parent_clause_0.Output - call_mut_back : Self → Args → parent_clause_0.Output → Result Self + CFnOnceSelfArgsInst : CFnOnce Self Args + call_mut : Self → Args → Result CFnOnceSelfArgsInst.Output + call_mut_back : Self → Args → CFnOnceSelfArgsInst.Output → Result Self /- Trait declaration: [traits::CFn] Source: 'src/traits.rs', lines 296:0-296:33 -/ structure CFn (Self Args : Type) where - parent_clause_0 : CFnMut Self Args - call_mut : Self → Args → Result parent_clause_0.parent_clause_0.Output + CFnMutSelfArgsInst : CFnMut Self Args + call_mut : Self → Args → Result + CFnMutSelfArgsInst.CFnOnceSelfArgsInst.Output /- [traits::incr_u32]: forward function Source: 'src/traits.rs', lines 300:0-300:30 -/ -- cgit v1.2.3 From 84a505ed9f193885175308ecc837922a41176b5b Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 22 Nov 2023 09:11:06 +0100 Subject: Regenerate the test files --- tests/lean/Traits.lean | 104 +++++++++++++++++++++++++++++-------------------- 1 file changed, 62 insertions(+), 42 deletions(-) (limited to 'tests/lean/Traits.lean') diff --git a/tests/lean/Traits.lean b/tests/lean/Traits.lean index 0aa68c5e..9ac4736c 100644 --- a/tests/lean/Traits.lean +++ b/tests/lean/Traits.lean @@ -61,8 +61,9 @@ def test_bool_trait_option (T : Type) (x : Option T) : Result Bool := /- [traits::test_bool_trait]: forward function Source: 'src/traits.rs', lines 35:0-35:50 -/ -def test_bool_trait (T : Type) (inst : BoolTrait T) (x : T) : Result Bool := - inst.get_bool x +def test_bool_trait + (T : Type) (BoolTraitTInst : BoolTrait T) (x : T) : Result Bool := + BoolTraitTInst.get_bool x /- Trait declaration: [traits::ToU64] Source: 'src/traits.rs', lines 39:0-39:15 -/ @@ -82,28 +83,31 @@ def traits.ToU64U64Inst : ToU64 U64 := { /- [traits::{(A, A)#3}::to_u64]: forward function Source: 'src/traits.rs', lines 50:4-50:26 -/ -def Pair.to_u64 (A : Type) (inst : ToU64 A) (self : (A × A)) : Result U64 := +def Pair.to_u64 + (A : Type) (ToU64AInst : ToU64 A) (self : (A × A)) : Result U64 := do let (t, t0) := self - let i ← inst.to_u64 t - let i0 ← inst.to_u64 t0 + let i ← ToU64AInst.to_u64 t + let i0 ← ToU64AInst.to_u64 t0 i + i0 /- Trait implementation: [traits::{(A, A)#3}] Source: 'src/traits.rs', lines 49:0-49:31 -/ -def traits.ToU64TupleAAInst (A : Type) (inst : ToU64 A) : ToU64 (A × A) := { - to_u64 := Pair.to_u64 A inst +def traits.ToU64TupleAAInst (A : Type) (ToU64AInst : ToU64 A) : ToU64 (A × A) + := { + to_u64 := Pair.to_u64 A ToU64AInst } /- [traits::f]: forward function Source: 'src/traits.rs', lines 55:0-55:36 -/ -def f (T : Type) (inst : ToU64 T) (x : (T × T)) : Result U64 := - Pair.to_u64 T inst x +def f (T : Type) (ToU64TInst : ToU64 T) (x : (T × T)) : Result U64 := + Pair.to_u64 T ToU64TInst x /- [traits::g]: forward function Source: 'src/traits.rs', lines 59:0-61:18 -/ -def g (T : Type) (inst : ToU64 (T × T)) (x : (T × T)) : Result U64 := - inst.to_u64 x +def g + (T : Type) (ToU64TupleTTInst : ToU64 (T × T)) (x : (T × T)) : Result U64 := + ToU64TupleTTInst.to_u64 x /- [traits::h0]: forward function Source: 'src/traits.rs', lines 66:0-66:24 -/ @@ -118,14 +122,14 @@ structure Wrapper (T : Type) where /- [traits::{traits::Wrapper#4}::to_u64]: forward function Source: 'src/traits.rs', lines 75:4-75:26 -/ def Wrapper.to_u64 - (T : Type) (inst : ToU64 T) (self : Wrapper T) : Result U64 := - inst.to_u64 self.x + (T : Type) (ToU64TInst : ToU64 T) (self : Wrapper T) : Result U64 := + ToU64TInst.to_u64 self.x /- Trait implementation: [traits::{traits::Wrapper#4}] Source: 'src/traits.rs', lines 74:0-74:35 -/ -def traits.ToU64traitsWrapperTInst (T : Type) (inst : ToU64 T) : ToU64 (Wrapper - T) := { - to_u64 := Wrapper.to_u64 T inst +def traits.ToU64traitsWrapperTInst (T : Type) (ToU64TInst : ToU64 T) : ToU64 + (Wrapper T) := { + to_u64 := Wrapper.to_u64 T ToU64TInst } /- [traits::h1]: forward function @@ -135,8 +139,8 @@ def h1 (x : Wrapper U64) : Result U64 := /- [traits::h2]: forward function Source: 'src/traits.rs', lines 84:0-84:41 -/ -def h2 (T : Type) (inst : ToU64 T) (x : Wrapper T) : Result U64 := - Wrapper.to_u64 T inst x +def h2 (T : Type) (ToU64TInst : ToU64 T) (x : Wrapper T) : Result U64 := + Wrapper.to_u64 T ToU64TInst x /- Trait declaration: [traits::ToType] Source: 'src/traits.rs', lines 88:0-88:19 -/ @@ -157,15 +161,17 @@ def traits.ToTypeU64BoolInst : ToType U64 Bool := { /- Trait declaration: [traits::OfType] Source: 'src/traits.rs', lines 98:0-98:16 -/ structure OfType (Self : Type) where - of_type : forall (T : Type) (inst : ToType T Self), T → Result Self + of_type : forall (T : Type) (ToTypeTSelfInst : ToType T Self), T → Result + Self /- [traits::h3]: forward function Source: 'src/traits.rs', lines 104:0-104:50 -/ def h3 - (T1 T2 : Type) (inst : OfType T1) (inst0 : ToType T2 T1) (y : T2) : + (T1 T2 : Type) (OfTypeT1Inst : OfType T1) (ToTypeT2T1Inst : ToType T2 T1) + (y : T2) : Result T1 := - inst.of_type T2 inst0 y + OfTypeT1Inst.of_type T2 ToTypeT2T1Inst y /- Trait declaration: [traits::OfTypeBis] Source: 'src/traits.rs', lines 109:0-109:36 -/ @@ -176,10 +182,11 @@ structure OfTypeBis (Self T : Type) where /- [traits::h4]: forward function Source: 'src/traits.rs', lines 118:0-118:57 -/ def h4 - (T1 T2 : Type) (inst : OfTypeBis T1 T2) (inst0 : ToType T2 T1) (y : T2) : + (T1 T2 : Type) (OfTypeBisT1T2Inst : OfTypeBis T1 T2) (ToTypeT2T1Inst : ToType + T2 T1) (y : T2) : Result T1 := - inst.of_type y + OfTypeBisT1T2Inst.of_type y /- [traits::TestType] Source: 'src/traits.rs', lines 122:0-122:22 -/ @@ -212,9 +219,11 @@ def traits.TestType.test.TestTraittraitstraitsTestTypeTtestTestType1Inst : /- [traits::{traits::TestType#6}::test]: forward function Source: 'src/traits.rs', lines 126:4-126:36 -/ def TestType.test - (T : Type) (inst : ToU64 T) (self : TestType T) (x : T) : Result Bool := + (T : Type) (ToU64TInst : ToU64 T) (self : TestType T) (x : T) : + Result Bool + := do - let x0 ← inst.to_u64 x + let x0 ← ToU64TInst.to_u64 x if x0 > 0#u64 then TestType.test.TestType1.test { _0 := 0#u64 } else Result.ret false @@ -227,14 +236,16 @@ structure BoolWrapper where /- [traits::{traits::BoolWrapper#7}::to_type]: forward function Source: 'src/traits.rs', lines 156:4-156:25 -/ def BoolWrapper.to_type - (T : Type) (inst : ToType Bool T) (self : BoolWrapper) : Result T := - inst.to_type self._0 + (T : Type) (ToTypeBoolTInst : ToType Bool T) (self : BoolWrapper) : + Result T + := + ToTypeBoolTInst.to_type self._0 /- Trait implementation: [traits::{traits::BoolWrapper#7}] Source: 'src/traits.rs', lines 152:0-152:33 -/ -def traits.ToTypetraitsBoolWrapperTInst (T : Type) (inst : ToType Bool T) : - ToType BoolWrapper T := { - to_type := BoolWrapper.to_type T inst +def traits.ToTypetraitsBoolWrapperTInst (T : Type) (ToTypeBoolTInst : ToType + Bool T) : ToType BoolWrapper T := { + to_type := BoolWrapper.to_type T ToTypeBoolTInst } /- [traits::WithConstTy::LEN2] @@ -278,14 +289,17 @@ def traits.WithConstTyBool32Inst : WithConstTy Bool 32#usize := { /- [traits::use_with_const_ty1]: forward function Source: 'src/traits.rs', lines 183:0-183:75 -/ def use_with_const_ty1 - (H : Type) (LEN : Usize) (inst : WithConstTy H LEN) : Result Usize := - let i := inst.LEN1 + (H : Type) (LEN : Usize) (WithConstTyHLENInst : WithConstTy H LEN) : + Result Usize + := + let i := WithConstTyHLENInst.LEN1 Result.ret i /- [traits::use_with_const_ty2]: forward function Source: 'src/traits.rs', lines 187:0-187:73 -/ def use_with_const_ty2 - (H : Type) (LEN : Usize) (inst : WithConstTy H LEN) (w : inst.W) : + (H : Type) (LEN : Usize) (WithConstTyHLENInst : WithConstTy H LEN) + (w : WithConstTyHLENInst.W) : Result Unit := Result.ret () @@ -293,10 +307,11 @@ def use_with_const_ty2 /- [traits::use_with_const_ty3]: forward function Source: 'src/traits.rs', lines 189:0-189:80 -/ def use_with_const_ty3 - (H : Type) (LEN : Usize) (inst : WithConstTy H LEN) (x : inst.W) : + (H : Type) (LEN : Usize) (WithConstTyHLENInst : WithConstTy H LEN) + (x : WithConstTyHLENInst.W) : Result U64 := - inst.W_clause_0.to_u64 x + WithConstTyHLENInst.W_clause_0.to_u64 x /- [traits::test_where1]: forward function Source: 'src/traits.rs', lines 193:0-193:40 -/ @@ -306,7 +321,9 @@ def test_where1 (T : Type) (_x : T) : Result Unit := /- [traits::test_where2]: forward function Source: 'src/traits.rs', lines 194:0-194:57 -/ def test_where2 - (T : Type) (inst : WithConstTy T 32#usize) (_x : U32) : Result Unit := + (T : Type) (WithConstTyT32Inst : WithConstTy T 32#usize) (_x : U32) : + Result Unit + := Result.ret () /- [alloc::string::String] @@ -333,21 +350,24 @@ structure ChildTrait (Self : Type) where /- [traits::test_child_trait1]: forward function Source: 'src/traits.rs', lines 209:0-209:56 -/ def test_child_trait1 - (T : Type) (inst : ChildTrait T) (x : T) : Result alloc.string.String := - inst.ParentTrait0SelfInst.get_name x + (T : Type) (ChildTraitTInst : ChildTrait T) (x : T) : + Result alloc.string.String + := + ChildTraitTInst.ParentTrait0SelfInst.get_name x /- [traits::test_child_trait2]: forward function Source: 'src/traits.rs', lines 213:0-213:54 -/ def test_child_trait2 - (T : Type) (inst : ChildTrait T) (x : T) : - Result inst.ParentTrait0SelfInst.W + (T : Type) (ChildTraitTInst : ChildTrait T) (x : T) : + Result ChildTraitTInst.ParentTrait0SelfInst.W := - inst.ParentTrait0SelfInst.get_w x + ChildTraitTInst.ParentTrait0SelfInst.get_w x /- [traits::order1]: forward function Source: 'src/traits.rs', lines 219:0-219:59 -/ def order1 - (T U : Type) (inst : ParentTrait0 T) (inst0 : ParentTrait0 U) : + (T U : Type) (ParentTrait0TInst : ParentTrait0 T) (ParentTrait0UInst : + ParentTrait0 U) : Result Unit := Result.ret () -- cgit v1.2.3 From d84040e000333d6d2a212fb849a38fb73a65eb48 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Fri, 24 Nov 2023 17:41:42 +0100 Subject: Regenerate the files --- tests/lean/Traits.lean | 10 ++-------- 1 file changed, 2 insertions(+), 8 deletions(-) (limited to 'tests/lean/Traits.lean') diff --git a/tests/lean/Traits.lean b/tests/lean/Traits.lean index 9ac4736c..e7795d9c 100644 --- a/tests/lean/Traits.lean +++ b/tests/lean/Traits.lean @@ -326,15 +326,11 @@ def test_where2 := Result.ret () -/- [alloc::string::String] - Source: '/rustc/d59363ad0b6391b7fc5bbb02c9ccf9300eef3753/library/alloc/src/string.rs', lines 365:0-365:17 -/ -axiom alloc.string.String : Type - /- Trait declaration: [traits::ParentTrait0] Source: 'src/traits.rs', lines 200:0-200:22 -/ structure ParentTrait0 (Self : Type) where W : Type - get_name : Self → Result alloc.string.String + get_name : Self → Result String get_w : Self → Result W /- Trait declaration: [traits::ParentTrait1] @@ -350,9 +346,7 @@ structure ChildTrait (Self : Type) where /- [traits::test_child_trait1]: forward function Source: 'src/traits.rs', lines 209:0-209:56 -/ def test_child_trait1 - (T : Type) (ChildTraitTInst : ChildTrait T) (x : T) : - Result alloc.string.String - := + (T : Type) (ChildTraitTInst : ChildTrait T) (x : T) : Result String := ChildTraitTInst.ParentTrait0SelfInst.get_name x /- [traits::test_child_trait2]: forward function -- cgit v1.2.3 From 59cb4312df866ec00a1fddec62a1e660b358052d Mon Sep 17 00:00:00 2001 From: Son Ho Date: Wed, 29 Nov 2023 11:46:35 +0100 Subject: Regenerate the test files --- tests/lean/Traits.lean | 5 ----- 1 file changed, 5 deletions(-) (limited to 'tests/lean/Traits.lean') diff --git a/tests/lean/Traits.lean b/tests/lean/Traits.lean index e7795d9c..653384d6 100644 --- a/tests/lean/Traits.lean +++ b/tests/lean/Traits.lean @@ -468,9 +468,4 @@ structure CFn (Self Args : Type) where call_mut : Self → Args → Result CFnMutSelfArgsInst.CFnOnceSelfArgsInst.Output -/- [traits::incr_u32]: forward function - Source: 'src/traits.rs', lines 300:0-300:30 -/ -def incr_u32 (x : U32) : Result U32 := - x + 1#u32 - end traits -- cgit v1.2.3