From a0c58326c43a7a8026b3d4c158017bf126180e90 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Fri, 22 Dec 2023 23:23:30 +0100 Subject: Regenerate the test files and add the fstar-split tests --- tests/lean/Traits.lean | 161 +++++++++++++++++++++++++++---------------------- 1 file changed, 89 insertions(+), 72 deletions(-) (limited to 'tests/lean/Traits.lean') diff --git a/tests/lean/Traits.lean b/tests/lean/Traits.lean index 17118203..63d07d85 100644 --- a/tests/lean/Traits.lean +++ b/tests/lean/Traits.lean @@ -10,7 +10,7 @@ namespace traits structure BoolTrait (Self : Type) where get_bool : Self → Result Bool -/- [traits::{bool}::get_bool]: forward function +/- [traits::{bool}::get_bool]: Source: 'src/traits.rs', lines 12:4-12:30 -/ def Bool.get_bool (self : Bool) : Result Bool := Result.ret self @@ -21,27 +21,27 @@ def traits.BoolTraitBoolInst : BoolTrait Bool := { get_bool := Bool.get_bool } -/- [traits::BoolTrait::ret_true]: forward function +/- [traits::BoolTrait::ret_true]: 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]: 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 - if b - then BoolTrait.ret_true traits.BoolTraitBoolInst x - else Result.ret false + let b ← Bool.get_bool x + if b + 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]: 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 + | some _ => Result.ret true /- Trait implementation: [traits::{core::option::Option#1}] Source: 'src/traits.rs', lines 22:0-22:31 -/ @@ -50,16 +50,16 @@ 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]: 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 - if b - then BoolTrait.ret_true (traits.BoolTraitcoreoptionOptionTInst T) x - else Result.ret false + let b ← Option.get_bool T x + if b + then BoolTrait.ret_true (traits.BoolTraitcoreoptionOptionTInst T) x + else Result.ret false -/- [traits::test_bool_trait]: forward function +/- [traits::test_bool_trait]: Source: 'src/traits.rs', lines 35:0-35:50 -/ def test_bool_trait (T : Type) (BoolTraitTInst : BoolTrait T) (x : T) : Result Bool := @@ -70,7 +70,7 @@ def test_bool_trait structure ToU64 (Self : Type) where to_u64 : Self → Result U64 -/- [traits::{u64#2}::to_u64]: forward function +/- [traits::{u64#2}::to_u64]: Source: 'src/traits.rs', lines 44:4-44:26 -/ def U64.to_u64 (self : U64) : Result U64 := Result.ret self @@ -81,15 +81,15 @@ def traits.ToU64U64Inst : ToU64 U64 := { to_u64 := U64.to_u64 } -/- [traits::{(A, A)#3}::to_u64]: forward function +/- [traits::{(A, A)#3}::to_u64]: Source: 'src/traits.rs', lines 50:4-50:26 -/ def Pair.to_u64 (A : Type) (ToU64AInst : ToU64 A) (self : (A × A)) : Result U64 := do - let (t, t0) := self - let i ← ToU64AInst.to_u64 t - let i0 ← ToU64AInst.to_u64 t0 - i + i0 + let (t, t1) := self + let i ← ToU64AInst.to_u64 t + let i1 ← ToU64AInst.to_u64 t1 + i + i1 /- Trait implementation: [traits::{(A, A)#3}] Source: 'src/traits.rs', lines 49:0-49:31 -/ @@ -98,18 +98,18 @@ def traits.ToU64TupleAAInst (A : Type) (ToU64AInst : ToU64 A) : ToU64 (A × A) to_u64 := Pair.to_u64 A ToU64AInst } -/- [traits::f]: forward function +/- [traits::f]: Source: 'src/traits.rs', lines 55:0-55:36 -/ def f (T : Type) (ToU64TInst : ToU64 T) (x : (T × T)) : Result U64 := Pair.to_u64 T ToU64TInst x -/- [traits::g]: forward function +/- [traits::g]: Source: 'src/traits.rs', lines 59:0-61:18 -/ def g (T : Type) (ToU64TupleTTInst : ToU64 (T × T)) (x : (T × T)) : Result U64 := ToU64TupleTTInst.to_u64 x -/- [traits::h0]: forward function +/- [traits::h0]: Source: 'src/traits.rs', lines 66:0-66:24 -/ def h0 (x : U64) : Result U64 := U64.to_u64 x @@ -119,7 +119,7 @@ def h0 (x : U64) : Result U64 := structure Wrapper (T : Type) where x : T -/- [traits::{traits::Wrapper#4}::to_u64]: forward function +/- [traits::{traits::Wrapper#4}::to_u64]: Source: 'src/traits.rs', lines 75:4-75:26 -/ def Wrapper.to_u64 (T : Type) (ToU64TInst : ToU64 T) (self : Wrapper T) : Result U64 := @@ -132,12 +132,12 @@ def traits.ToU64traitsWrapperTInst (T : Type) (ToU64TInst : ToU64 T) : ToU64 to_u64 := Wrapper.to_u64 T ToU64TInst } -/- [traits::h1]: forward function +/- [traits::h1]: 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]: Source: 'src/traits.rs', lines 84:0-84:41 -/ def h2 (T : Type) (ToU64TInst : ToU64 T) (x : Wrapper T) : Result U64 := Wrapper.to_u64 T ToU64TInst x @@ -147,7 +147,7 @@ def h2 (T : Type) (ToU64TInst : 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 +/- [traits::{u64#5}::to_type]: Source: 'src/traits.rs', lines 93:4-93:28 -/ def U64.to_type (self : U64) : Result Bool := Result.ret (self > 0#u64) @@ -164,7 +164,7 @@ structure OfType (Self : Type) where of_type : forall (T : Type) (ToTypeTSelfInst : ToType T Self), T → Result Self -/- [traits::h3]: forward function +/- [traits::h3]: Source: 'src/traits.rs', lines 104:0-104:50 -/ def h3 (T1 T2 : Type) (OfTypeT1Inst : OfType T1) (ToTypeT2T1Inst : ToType T2 T1) @@ -179,7 +179,7 @@ structure OfTypeBis (Self T : Type) where ToTypeTSelfInst : ToType T Self of_type : T → Result Self -/- [traits::h4]: forward function +/- [traits::h4]: Source: 'src/traits.rs', lines 118:0-118:57 -/ def h4 (T1 T2 : Type) (OfTypeBisT1T2Inst : OfTypeBis T1 T2) (ToTypeT2T1Inst : ToType @@ -201,7 +201,7 @@ def h4 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]: Source: 'src/traits.rs', lines 139:12-139:34 -/ def TestType.test.TestType1.test (self : TestType.test.TestType1) : Result Bool := @@ -214,23 +214,23 @@ def traits.TestType.test.TestTraittraitstraitsTestTypeTtestTestType1Inst : test := TestType.test.TestType1.test } -/- [traits::{traits::TestType#6}::test]: forward function +/- [traits::{traits::TestType#6}::test]: Source: 'src/traits.rs', lines 126:4-126:36 -/ def TestType.test (T : Type) (ToU64TInst : ToU64 T) (self : TestType T) (x : T) : Result Bool := do - let x0 ← ToU64TInst.to_u64 x - if x0 > 0#u64 - then TestType.test.TestType1.test 0#u64 - else Result.ret false + let x1 ← ToU64TInst.to_u64 x + if x1 > 0#u64 + then TestType.test.TestType1.test 0#u64 + else Result.ret false /- [traits::BoolWrapper] Source: 'src/traits.rs', lines 150:0-150:22 -/ @[reducible] def BoolWrapper := Bool -/- [traits::{traits::BoolWrapper#7}::to_type]: forward function +/- [traits::{traits::BoolWrapper#7}::to_type]: Source: 'src/traits.rs', lines 156:4-156:25 -/ def BoolWrapper.to_type (T : Type) (ToTypeBoolTInst : ToType Bool T) (self : BoolWrapper) : @@ -266,8 +266,7 @@ structure WithConstTy (Self : Type) (LEN : Usize) where 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 ()) +/- [traits::{bool#8}::f]: Source: 'src/traits.rs', lines 180:4-180:39 -/ def Bool.f (i : U64) (a : Array U8 32#usize) : Result U64 := Result.ret i @@ -283,7 +282,7 @@ def traits.WithConstTyBool32Inst : WithConstTy Bool 32#usize := { f := Bool.f } -/- [traits::use_with_const_ty1]: forward function +/- [traits::use_with_const_ty1]: Source: 'src/traits.rs', lines 183:0-183:75 -/ def use_with_const_ty1 (H : Type) (LEN : Usize) (WithConstTyHLENInst : WithConstTy H LEN) : @@ -292,7 +291,7 @@ def use_with_const_ty1 let i := WithConstTyHLENInst.LEN1 Result.ret i -/- [traits::use_with_const_ty2]: forward function +/- [traits::use_with_const_ty2]: Source: 'src/traits.rs', lines 187:0-187:73 -/ def use_with_const_ty2 (H : Type) (LEN : Usize) (WithConstTyHLENInst : WithConstTy H LEN) @@ -301,7 +300,7 @@ def use_with_const_ty2 := Result.ret () -/- [traits::use_with_const_ty3]: forward function +/- [traits::use_with_const_ty3]: Source: 'src/traits.rs', lines 189:0-189:80 -/ def use_with_const_ty3 (H : Type) (LEN : Usize) (WithConstTyHLENInst : WithConstTy H LEN) @@ -310,12 +309,12 @@ def use_with_const_ty3 := WithConstTyHLENInst.W_clause_0.to_u64 x -/- [traits::test_where1]: forward function +/- [traits::test_where1]: 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]: Source: 'src/traits.rs', lines 194:0-194:57 -/ def test_where2 (T : Type) (WithConstTyT32Inst : WithConstTy T 32#usize) (_x : U32) : @@ -340,22 +339,30 @@ structure ChildTrait (Self : Type) where ParentTrait0SelfInst : ParentTrait0 Self ParentTrait1SelfInst : ParentTrait1 Self -/- [traits::test_child_trait1]: forward function - Source: 'src/traits.rs', lines 209:0-209:56 -/ +/- [traits::test_parent_trait0]: + Source: 'src/traits.rs', lines 208:0-208:57 -/ +def test_parent_trait0 + (T : Type) (ParentTrait0TInst : ParentTrait0 T) (x : T) : + Result ParentTrait0TInst.W + := + ParentTrait0TInst.get_w x + +/- [traits::test_child_trait1]: + Source: 'src/traits.rs', lines 213:0-213:56 -/ def test_child_trait1 (T : Type) (ChildTraitTInst : ChildTrait T) (x : T) : Result String := ChildTraitTInst.ParentTrait0SelfInst.get_name x -/- [traits::test_child_trait2]: forward function - Source: 'src/traits.rs', lines 213:0-213:54 -/ +/- [traits::test_child_trait2]: + Source: 'src/traits.rs', lines 217:0-217:54 -/ def test_child_trait2 (T : Type) (ChildTraitTInst : ChildTrait T) (x : T) : Result ChildTraitTInst.ParentTrait0SelfInst.W := ChildTraitTInst.ParentTrait0SelfInst.get_w x -/- [traits::order1]: forward function - Source: 'src/traits.rs', lines 219:0-219:59 -/ +/- [traits::order1]: + Source: 'src/traits.rs', lines 223:0-223:59 -/ def order1 (T U : Type) (ParentTrait0TInst : ParentTrait0 T) (ParentTrait0UInst : ParentTrait0 U) : @@ -364,28 +371,28 @@ def order1 Result.ret () /- Trait declaration: [traits::ChildTrait1] - Source: 'src/traits.rs', lines 222:0-222:35 -/ + Source: 'src/traits.rs', lines 226:0-226:35 -/ structure ChildTrait1 (Self : Type) where ParentTrait1SelfInst : ParentTrait1 Self /- Trait implementation: [traits::{usize#9}] - Source: 'src/traits.rs', lines 224:0-224:27 -/ + Source: 'src/traits.rs', lines 228:0-228:27 -/ def traits.ParentTrait1UsizeInst : ParentTrait1 Usize := { } /- Trait implementation: [traits::{usize#10}] - Source: 'src/traits.rs', lines 225:0-225:26 -/ + Source: 'src/traits.rs', lines 229:0-229:26 -/ def traits.ChildTrait1UsizeInst : ChildTrait1 Usize := { ParentTrait1SelfInst := traits.ParentTrait1UsizeInst } /- Trait declaration: [traits::Iterator] - Source: 'src/traits.rs', lines 229:0-229:18 -/ + Source: 'src/traits.rs', lines 233:0-233:18 -/ structure Iterator (Self : Type) where Item : Type /- Trait declaration: [traits::IntoIterator] - Source: 'src/traits.rs', lines 233:0-233:22 -/ + Source: 'src/traits.rs', lines 237:0-237:22 -/ structure IntoIterator (Self : Type) where Item : Type IntoIter : Type @@ -393,76 +400,86 @@ structure IntoIterator (Self : Type) where into_iter : Self → Result IntoIter /- Trait declaration: [traits::FromResidual] - Source: 'src/traits.rs', lines 250:0-250:21 -/ + Source: 'src/traits.rs', lines 254:0-254:21 -/ structure FromResidual (Self T : Type) where /- Trait declaration: [traits::Try] - Source: 'src/traits.rs', lines 246:0-246:48 -/ + Source: 'src/traits.rs', lines 250:0-250:48 -/ structure Try (Self : Type) where Residual : Type FromResidualSelftraitsTrySelfResidualInst : FromResidual Self Residual /- Trait declaration: [traits::WithTarget] - Source: 'src/traits.rs', lines 252:0-252:20 -/ + Source: 'src/traits.rs', lines 256:0-256:20 -/ structure WithTarget (Self : Type) where Target : Type /- Trait declaration: [traits::ParentTrait2] - Source: 'src/traits.rs', lines 256:0-256:22 -/ + Source: 'src/traits.rs', lines 260:0-260:22 -/ structure ParentTrait2 (Self : Type) where U : Type U_clause_0 : WithTarget U /- Trait declaration: [traits::ChildTrait2] - Source: 'src/traits.rs', lines 260:0-260:35 -/ + Source: 'src/traits.rs', lines 264:0-264:35 -/ structure ChildTrait2 (Self : Type) where 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 -/ + Source: 'src/traits.rs', lines 268:0-268:23 -/ def traits.WithTargetU32Inst : WithTarget U32 := { Target := U32 } /- Trait implementation: [traits::{u32#12}] - Source: 'src/traits.rs', lines 268:0-268:25 -/ + Source: 'src/traits.rs', lines 272:0-272:25 -/ def traits.ParentTrait2U32Inst : ParentTrait2 U32 := { U := U32 U_clause_0 := traits.WithTargetU32Inst } -/- [traits::{u32#13}::convert]: forward function - Source: 'src/traits.rs', lines 273:4-273:29 -/ +/- [traits::{u32#13}::convert]: + Source: 'src/traits.rs', lines 277:4-277:29 -/ def U32.convert (x : U32) : Result U32 := Result.ret x /- Trait implementation: [traits::{u32#13}] - Source: 'src/traits.rs', lines 272:0-272:24 -/ + Source: 'src/traits.rs', lines 276:0-276:24 -/ def traits.ChildTrait2U32Inst : ChildTrait2 U32 := { ParentTrait2SelfInst := traits.ParentTrait2U32Inst convert := U32.convert } /- Trait declaration: [traits::CFnOnce] - Source: 'src/traits.rs', lines 286:0-286:23 -/ + Source: 'src/traits.rs', lines 290:0-290:23 -/ structure CFnOnce (Self Args : Type) where Output : Type call_once : Self → Args → Result Output /- Trait declaration: [traits::CFnMut] - Source: 'src/traits.rs', lines 292:0-292:37 -/ + Source: 'src/traits.rs', lines 296:0-296:37 -/ structure CFnMut (Self Args : Type) where CFnOnceSelfArgsInst : CFnOnce Self Args - call_mut : Self → Args → Result CFnOnceSelfArgsInst.Output - call_mut_back : Self → Args → CFnOnceSelfArgsInst.Output → Result Self + call_mut : Self → Args → Result (CFnOnceSelfArgsInst.Output × Self) /- Trait declaration: [traits::CFn] - Source: 'src/traits.rs', lines 296:0-296:33 -/ + Source: 'src/traits.rs', lines 300:0-300:33 -/ structure CFn (Self Args : Type) where CFnMutSelfArgsInst : CFnMut Self Args - call_mut : Self → Args → Result - CFnMutSelfArgsInst.CFnOnceSelfArgsInst.Output + call : Self → Args → Result CFnMutSelfArgsInst.CFnOnceSelfArgsInst.Output + +/- Trait declaration: [traits::GetTrait] + Source: 'src/traits.rs', lines 304:0-304:18 -/ +structure GetTrait (Self : Type) where + W : Type + get_w : Self → Result W + +/- [traits::test_get_trait]: + Source: 'src/traits.rs', lines 309:0-309:49 -/ +def test_get_trait + (T : Type) (GetTraitTInst : GetTrait T) (x : T) : Result GetTraitTInst.W := + GetTraitTInst.get_w x end traits -- cgit v1.2.3 From a4decc7654bc6f3301c0174124d21fdbc2dbc708 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Sat, 23 Dec 2023 00:59:55 +0100 Subject: Regenerate the files --- tests/lean/Traits.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) (limited to 'tests/lean/Traits.lean') diff --git a/tests/lean/Traits.lean b/tests/lean/Traits.lean index 63d07d85..35f9e5bf 100644 --- a/tests/lean/Traits.lean +++ b/tests/lean/Traits.lean @@ -288,8 +288,7 @@ def use_with_const_ty1 (H : Type) (LEN : Usize) (WithConstTyHLENInst : WithConstTy H LEN) : Result Usize := - let i := WithConstTyHLENInst.LEN1 - Result.ret i + Result.ret WithConstTyHLENInst.LEN1 /- [traits::use_with_const_ty2]: Source: 'src/traits.rs', lines 187:0-187:73 -/ -- cgit v1.2.3