From 50d1542f830b7ceb73efd34573b6b56b4971a114 Mon Sep 17 00:00:00 2001 From: Son Ho Date: Tue, 9 May 2023 10:40:19 +0200 Subject: Regenerate the translated files for Lean --- tests/lean/misc-loops/Loops/Funs.lean | 281 +++++++++++++++------------------- 1 file changed, 123 insertions(+), 158 deletions(-) (limited to 'tests/lean/misc-loops/Loops/Funs.lean') diff --git a/tests/lean/misc-loops/Loops/Funs.lean b/tests/lean/misc-loops/Loops/Funs.lean index f79a27a9..fd8d62d7 100644 --- a/tests/lean/misc-loops/Loops/Funs.lean +++ b/tests/lean/misc-loops/Loops/Funs.lean @@ -5,79 +5,78 @@ import Loops.Types import Loops.Clauses.Clauses /- [loops::sum] -/ -def sum_loop_fwd (max : UInt32) (i : UInt32) (s : UInt32) : (Result UInt32) := +def sum_loop_fwd (max : U32) (i : U32) (s : U32) : (Result U32) := if h: i < max then do - let s0 ← UInt32.checked_add s i - let i0 ← UInt32.checked_add i (UInt32.ofNatCore 1 (by intlit)) + let s0 ← s + i + let i0 ← i + (U32.ofInt 1 (by intlit)) sum_loop_fwd max i0 s0 - else UInt32.checked_mul s (UInt32.ofNatCore 2 (by intlit)) + else s * (U32.ofInt 2 (by intlit)) termination_by sum_loop_fwd max i s => sum_loop_terminates max i s decreasing_by sum_loop_decreases max i s /- [loops::sum] -/ -def sum_fwd (max : UInt32) : Result UInt32 := - sum_loop_fwd max (UInt32.ofNatCore 0 (by intlit)) - (UInt32.ofNatCore 0 (by intlit)) +def sum_fwd (max : U32) : Result U32 := + sum_loop_fwd max (U32.ofInt 0 (by intlit)) (U32.ofInt 0 (by intlit)) /- [loops::sum_with_mut_borrows] -/ def sum_with_mut_borrows_loop_fwd - (max : UInt32) (mi : UInt32) (ms : UInt32) : (Result UInt32) := + (max : U32) (mi : U32) (ms : U32) : (Result U32) := if h: mi < max then do - let ms0 ← UInt32.checked_add ms mi - let mi0 ← UInt32.checked_add mi (UInt32.ofNatCore 1 (by intlit)) + let ms0 ← ms + mi + let mi0 ← mi + (U32.ofInt 1 (by intlit)) sum_with_mut_borrows_loop_fwd max mi0 ms0 - else UInt32.checked_mul ms (UInt32.ofNatCore 2 (by intlit)) + else ms * (U32.ofInt 2 (by intlit)) termination_by sum_with_mut_borrows_loop_fwd max mi ms => sum_with_mut_borrows_loop_terminates max mi ms decreasing_by sum_with_mut_borrows_loop_decreases max mi ms /- [loops::sum_with_mut_borrows] -/ -def sum_with_mut_borrows_fwd (max : UInt32) : Result UInt32 := - sum_with_mut_borrows_loop_fwd max (UInt32.ofNatCore 0 (by intlit)) - (UInt32.ofNatCore 0 (by intlit)) +def sum_with_mut_borrows_fwd (max : U32) : Result U32 := + sum_with_mut_borrows_loop_fwd max (U32.ofInt 0 (by intlit)) + (U32.ofInt 0 (by intlit)) /- [loops::sum_with_shared_borrows] -/ def sum_with_shared_borrows_loop_fwd - (max : UInt32) (i : UInt32) (s : UInt32) : (Result UInt32) := + (max : U32) (i : U32) (s : U32) : (Result U32) := if h: i < max then do - let i0 ← UInt32.checked_add i (UInt32.ofNatCore 1 (by intlit)) - let s0 ← UInt32.checked_add s i0 + let i0 ← i + (U32.ofInt 1 (by intlit)) + let s0 ← s + i0 sum_with_shared_borrows_loop_fwd max i0 s0 - else UInt32.checked_mul s (UInt32.ofNatCore 2 (by intlit)) + else s * (U32.ofInt 2 (by intlit)) termination_by sum_with_shared_borrows_loop_fwd max i s => sum_with_shared_borrows_loop_terminates max i s decreasing_by sum_with_shared_borrows_loop_decreases max i s /- [loops::sum_with_shared_borrows] -/ -def sum_with_shared_borrows_fwd (max : UInt32) : Result UInt32 := - sum_with_shared_borrows_loop_fwd max (UInt32.ofNatCore 0 (by intlit)) - (UInt32.ofNatCore 0 (by intlit)) +def sum_with_shared_borrows_fwd (max : U32) : Result U32 := + sum_with_shared_borrows_loop_fwd max (U32.ofInt 0 (by intlit)) + (U32.ofInt 0 (by intlit)) /- [loops::clear] -/ -def clear_loop_fwd_back (v : Vec UInt32) (i : USize) : (Result (Vec UInt32)) := - let i0 := vec_len UInt32 v +def clear_loop_fwd_back (v : Vec U32) (i : Usize) : (Result (Vec U32)) := + let i0 := vec_len U32 v if h: i < i0 then do - let i1 ← USize.checked_add i (USize.ofNatCore 1 (by intlit)) - let v0 ← vec_index_mut_back UInt32 v i (UInt32.ofNatCore 0 (by intlit)) + let i1 ← i + (Usize.ofInt 1 (by intlit)) + let v0 ← vec_index_mut_back U32 v i (U32.ofInt 0 (by intlit)) clear_loop_fwd_back v0 i1 else Result.ret v termination_by clear_loop_fwd_back v i => clear_loop_terminates v i decreasing_by clear_loop_decreases v i /- [loops::clear] -/ -def clear_fwd_back (v : Vec UInt32) : Result (Vec UInt32) := - clear_loop_fwd_back v (USize.ofNatCore 0 (by intlit)) +def clear_fwd_back (v : Vec U32) : Result (Vec U32) := + clear_loop_fwd_back v (Usize.ofInt 0 (by intlit)) /- [loops::list_mem] -/ -def list_mem_loop_fwd (x : UInt32) (ls : list_t UInt32) : (Result Bool) := +def list_mem_loop_fwd (x : U32) (ls : list_t U32) : (Result Bool) := match h: ls with | list_t.Cons y tl => if h: y = x @@ -88,19 +87,19 @@ termination_by list_mem_loop_fwd x ls => list_mem_loop_terminates x ls decreasing_by list_mem_loop_decreases x ls /- [loops::list_mem] -/ -def list_mem_fwd (x : UInt32) (ls : list_t UInt32) : Result Bool := +def list_mem_fwd (x : U32) (ls : list_t U32) : Result Bool := list_mem_loop_fwd x ls /- [loops::list_nth_mut_loop] -/ def list_nth_mut_loop_loop_fwd - (T : Type) (ls : list_t T) (i : UInt32) : (Result T) := + (T : Type) (ls : list_t T) (i : U32) : (Result T) := match h: ls with | list_t.Cons x tl => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret x else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) list_nth_mut_loop_loop_fwd T tl i0 | list_t.Nil => Result.fail Error.panic termination_by list_nth_mut_loop_loop_fwd ls i => @@ -108,19 +107,19 @@ termination_by list_nth_mut_loop_loop_fwd ls i => decreasing_by list_nth_mut_loop_loop_decreases ls i /- [loops::list_nth_mut_loop] -/ -def list_nth_mut_loop_fwd (T : Type) (ls : list_t T) (i : UInt32) : Result T := +def list_nth_mut_loop_fwd (T : Type) (ls : list_t T) (i : U32) : Result T := list_nth_mut_loop_loop_fwd T ls i /- [loops::list_nth_mut_loop] -/ def list_nth_mut_loop_loop_back - (T : Type) (ls : list_t T) (i : UInt32) (ret0 : T) : (Result (list_t T)) := + (T : Type) (ls : list_t T) (i : U32) (ret0 : T) : (Result (list_t T)) := match h: ls with | list_t.Cons x tl => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (list_t.Cons ret0 tl) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) let tl0 ← list_nth_mut_loop_loop_back T tl i0 ret0 Result.ret (list_t.Cons x tl0) | list_t.Nil => Result.fail Error.panic @@ -130,19 +129,19 @@ decreasing_by list_nth_mut_loop_loop_decreases ls i /- [loops::list_nth_mut_loop] -/ def list_nth_mut_loop_back - (T : Type) (ls : list_t T) (i : UInt32) (ret0 : T) : Result (list_t T) := + (T : Type) (ls : list_t T) (i : U32) (ret0 : T) : Result (list_t T) := list_nth_mut_loop_loop_back T ls i ret0 /- [loops::list_nth_shared_loop] -/ def list_nth_shared_loop_loop_fwd - (T : Type) (ls : list_t T) (i : UInt32) : (Result T) := + (T : Type) (ls : list_t T) (i : U32) : (Result T) := match h: ls with | list_t.Cons x tl => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret x else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) list_nth_shared_loop_loop_fwd T tl i0 | list_t.Nil => Result.fail Error.panic termination_by list_nth_shared_loop_loop_fwd ls i => @@ -150,12 +149,11 @@ termination_by list_nth_shared_loop_loop_fwd ls i => decreasing_by list_nth_shared_loop_loop_decreases ls i /- [loops::list_nth_shared_loop] -/ -def list_nth_shared_loop_fwd - (T : Type) (ls : list_t T) (i : UInt32) : Result T := +def list_nth_shared_loop_fwd (T : Type) (ls : list_t T) (i : U32) : Result T := list_nth_shared_loop_loop_fwd T ls i /- [loops::get_elem_mut] -/ -def get_elem_mut_loop_fwd (x : USize) (ls : list_t USize) : (Result USize) := +def get_elem_mut_loop_fwd (x : Usize) (ls : list_t Usize) : (Result Usize) := match h: ls with | list_t.Cons y tl => if h: y = x @@ -166,15 +164,15 @@ termination_by get_elem_mut_loop_fwd x ls => get_elem_mut_loop_terminates x ls decreasing_by get_elem_mut_loop_decreases x ls /- [loops::get_elem_mut] -/ -def get_elem_mut_fwd (slots : Vec (list_t USize)) (x : USize) : Result USize := +def get_elem_mut_fwd (slots : Vec (list_t Usize)) (x : Usize) : Result Usize := do let l ← - vec_index_mut_fwd (list_t USize) slots (USize.ofNatCore 0 (by intlit)) + vec_index_mut_fwd (list_t Usize) slots (Usize.ofInt 0 (by intlit)) get_elem_mut_loop_fwd x l /- [loops::get_elem_mut] -/ def get_elem_mut_loop_back - (x : USize) (ls : list_t USize) (ret0 : USize) : (Result (list_t USize)) := + (x : Usize) (ls : list_t Usize) (ret0 : Usize) : (Result (list_t Usize)) := match h: ls with | list_t.Cons y tl => if h: y = x @@ -190,18 +188,18 @@ decreasing_by get_elem_mut_loop_decreases x ls /- [loops::get_elem_mut] -/ def get_elem_mut_back - (slots : Vec (list_t USize)) (x : USize) (ret0 : USize) : - Result (Vec (list_t USize)) + (slots : Vec (list_t Usize)) (x : Usize) (ret0 : Usize) : + Result (Vec (list_t Usize)) := do let l ← - vec_index_mut_fwd (list_t USize) slots (USize.ofNatCore 0 (by intlit)) + vec_index_mut_fwd (list_t Usize) slots (Usize.ofInt 0 (by intlit)) let l0 ← get_elem_mut_loop_back x l ret0 - vec_index_mut_back (list_t USize) slots (USize.ofNatCore 0 (by intlit)) l0 + vec_index_mut_back (list_t Usize) slots (Usize.ofInt 0 (by intlit)) l0 /- [loops::get_elem_shared] -/ def get_elem_shared_loop_fwd - (x : USize) (ls : list_t USize) : (Result USize) := + (x : Usize) (ls : list_t Usize) : (Result Usize) := match h: ls with | list_t.Cons y tl => if h: y = x @@ -214,10 +212,9 @@ decreasing_by get_elem_shared_loop_decreases x ls /- [loops::get_elem_shared] -/ def get_elem_shared_fwd - (slots : Vec (list_t USize)) (x : USize) : Result USize := + (slots : Vec (list_t Usize)) (x : Usize) : Result Usize := do - let l ← - vec_index_fwd (list_t USize) slots (USize.ofNatCore 0 (by intlit)) + let l ← vec_index_fwd (list_t Usize) slots (Usize.ofInt 0 (by intlit)) get_elem_shared_loop_fwd x l /- [loops::id_mut] -/ @@ -235,14 +232,14 @@ def id_shared_fwd (T : Type) (ls : list_t T) : Result (list_t T) := /- [loops::list_nth_mut_loop_with_id] -/ def list_nth_mut_loop_with_id_loop_fwd - (T : Type) (i : UInt32) (ls : list_t T) : (Result T) := + (T : Type) (i : U32) (ls : list_t T) : (Result T) := match h: ls with | list_t.Cons x tl => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret x else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) list_nth_mut_loop_with_id_loop_fwd T i0 tl | list_t.Nil => Result.fail Error.panic termination_by list_nth_mut_loop_with_id_loop_fwd i ls => @@ -251,21 +248,21 @@ decreasing_by list_nth_mut_loop_with_id_loop_decreases i ls /- [loops::list_nth_mut_loop_with_id] -/ def list_nth_mut_loop_with_id_fwd - (T : Type) (ls : list_t T) (i : UInt32) : Result T := + (T : Type) (ls : list_t T) (i : U32) : Result T := do let ls0 ← id_mut_fwd T ls list_nth_mut_loop_with_id_loop_fwd T i ls0 /- [loops::list_nth_mut_loop_with_id] -/ def list_nth_mut_loop_with_id_loop_back - (T : Type) (i : UInt32) (ls : list_t T) (ret0 : T) : (Result (list_t T)) := + (T : Type) (i : U32) (ls : list_t T) (ret0 : T) : (Result (list_t T)) := match h: ls with | list_t.Cons x tl => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (list_t.Cons ret0 tl) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) let tl0 ← list_nth_mut_loop_with_id_loop_back T i0 tl ret0 Result.ret (list_t.Cons x tl0) | list_t.Nil => Result.fail Error.panic @@ -275,7 +272,7 @@ decreasing_by list_nth_mut_loop_with_id_loop_decreases i ls /- [loops::list_nth_mut_loop_with_id] -/ def list_nth_mut_loop_with_id_back - (T : Type) (ls : list_t T) (i : UInt32) (ret0 : T) : Result (list_t T) := + (T : Type) (ls : list_t T) (i : U32) (ret0 : T) : Result (list_t T) := do let ls0 ← id_mut_fwd T ls let l ← list_nth_mut_loop_with_id_loop_back T i ls0 ret0 @@ -283,14 +280,14 @@ def list_nth_mut_loop_with_id_back /- [loops::list_nth_shared_loop_with_id] -/ def list_nth_shared_loop_with_id_loop_fwd - (T : Type) (i : UInt32) (ls : list_t T) : (Result T) := + (T : Type) (i : U32) (ls : list_t T) : (Result T) := match h: ls with | list_t.Cons x tl => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret x else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) list_nth_shared_loop_with_id_loop_fwd T i0 tl | list_t.Nil => Result.fail Error.panic termination_by list_nth_shared_loop_with_id_loop_fwd i ls => @@ -299,25 +296,23 @@ decreasing_by list_nth_shared_loop_with_id_loop_decreases i ls /- [loops::list_nth_shared_loop_with_id] -/ def list_nth_shared_loop_with_id_fwd - (T : Type) (ls : list_t T) (i : UInt32) : Result T := + (T : Type) (ls : list_t T) (i : U32) : Result T := do let ls0 ← id_shared_fwd T ls list_nth_shared_loop_with_id_loop_fwd T i ls0 /- [loops::list_nth_mut_loop_pair] -/ def list_nth_mut_loop_pair_loop_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - (Result (T × T)) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : (Result (T × T)) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (x0, x1) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) list_nth_mut_loop_pair_loop_fwd T tl0 tl1 i0 | list_t.Nil => Result.fail Error.panic | list_t.Nil => Result.fail Error.panic @@ -327,25 +322,23 @@ decreasing_by list_nth_mut_loop_pair_loop_decreases ls0 ls1 i /- [loops::list_nth_mut_loop_pair] -/ def list_nth_mut_loop_pair_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - Result (T × T) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : Result (T × T) := list_nth_mut_loop_pair_loop_fwd T ls0 ls1 i /- [loops::list_nth_mut_loop_pair] -/ def list_nth_mut_loop_pair_loop_back'a - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) (ret0 : T) : + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) (ret0 : T) : (Result (list_t T)) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (list_t.Cons ret0 tl0) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) let tl00 ← list_nth_mut_loop_pair_loop_back'a T tl0 tl1 i0 ret0 Result.ret (list_t.Cons x0 tl00) | list_t.Nil => Result.fail Error.panic @@ -356,25 +349,25 @@ decreasing_by list_nth_mut_loop_pair_loop_decreases ls0 ls1 i /- [loops::list_nth_mut_loop_pair] -/ def list_nth_mut_loop_pair_back'a - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) (ret0 : T) : + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) (ret0 : T) : Result (list_t T) := list_nth_mut_loop_pair_loop_back'a T ls0 ls1 i ret0 /- [loops::list_nth_mut_loop_pair] -/ def list_nth_mut_loop_pair_loop_back'b - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) (ret0 : T) : + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) (ret0 : T) : (Result (list_t T)) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (list_t.Cons ret0 tl1) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) let tl10 ← list_nth_mut_loop_pair_loop_back'b T tl0 tl1 i0 ret0 Result.ret (list_t.Cons x1 tl10) | list_t.Nil => Result.fail Error.panic @@ -385,25 +378,23 @@ decreasing_by list_nth_mut_loop_pair_loop_decreases ls0 ls1 i /- [loops::list_nth_mut_loop_pair] -/ def list_nth_mut_loop_pair_back'b - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) (ret0 : T) : + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) (ret0 : T) : Result (list_t T) := list_nth_mut_loop_pair_loop_back'b T ls0 ls1 i ret0 /- [loops::list_nth_shared_loop_pair] -/ def list_nth_shared_loop_pair_loop_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - (Result (T × T)) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : (Result (T × T)) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (x0, x1) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) list_nth_shared_loop_pair_loop_fwd T tl0 tl1 i0 | list_t.Nil => Result.fail Error.panic | list_t.Nil => Result.fail Error.panic @@ -413,25 +404,21 @@ decreasing_by list_nth_shared_loop_pair_loop_decreases ls0 ls1 i /- [loops::list_nth_shared_loop_pair] -/ def list_nth_shared_loop_pair_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - Result (T × T) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : Result (T × T) := list_nth_shared_loop_pair_loop_fwd T ls0 ls1 i /- [loops::list_nth_mut_loop_pair_merge] -/ def list_nth_mut_loop_pair_merge_loop_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - (Result (T × T)) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : (Result (T × T)) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (x0, x1) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) list_nth_mut_loop_pair_merge_loop_fwd T tl0 tl1 i0 | list_t.Nil => Result.fail Error.panic | list_t.Nil => Result.fail Error.panic @@ -441,27 +428,25 @@ decreasing_by list_nth_mut_loop_pair_merge_loop_decreases ls0 ls1 i /- [loops::list_nth_mut_loop_pair_merge] -/ def list_nth_mut_loop_pair_merge_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - Result (T × T) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : Result (T × T) := list_nth_mut_loop_pair_merge_loop_fwd T ls0 ls1 i /- [loops::list_nth_mut_loop_pair_merge] -/ def list_nth_mut_loop_pair_merge_loop_back - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) (ret0 : (T × T)) : + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) (ret0 : (T × T)) : (Result ((list_t T) × (list_t T))) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then let (t, t0) := ret0 Result.ret (list_t.Cons t tl0, list_t.Cons t0 tl1) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) let (tl00, tl10) ← list_nth_mut_loop_pair_merge_loop_back T tl0 tl1 i0 ret0 Result.ret (list_t.Cons x0 tl00, list_t.Cons x1 tl10) @@ -473,25 +458,23 @@ decreasing_by list_nth_mut_loop_pair_merge_loop_decreases ls0 ls1 i /- [loops::list_nth_mut_loop_pair_merge] -/ def list_nth_mut_loop_pair_merge_back - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) (ret0 : (T × T)) : + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) (ret0 : (T × T)) : Result ((list_t T) × (list_t T)) := list_nth_mut_loop_pair_merge_loop_back T ls0 ls1 i ret0 /- [loops::list_nth_shared_loop_pair_merge] -/ def list_nth_shared_loop_pair_merge_loop_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - (Result (T × T)) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : (Result (T × T)) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (x0, x1) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) list_nth_shared_loop_pair_merge_loop_fwd T tl0 tl1 i0 | list_t.Nil => Result.fail Error.panic | list_t.Nil => Result.fail Error.panic @@ -501,25 +484,21 @@ decreasing_by list_nth_shared_loop_pair_merge_loop_decreases ls0 ls1 i /- [loops::list_nth_shared_loop_pair_merge] -/ def list_nth_shared_loop_pair_merge_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - Result (T × T) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : Result (T × T) := list_nth_shared_loop_pair_merge_loop_fwd T ls0 ls1 i /- [loops::list_nth_mut_shared_loop_pair] -/ def list_nth_mut_shared_loop_pair_loop_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - (Result (T × T)) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : (Result (T × T)) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (x0, x1) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) list_nth_mut_shared_loop_pair_loop_fwd T tl0 tl1 i0 | list_t.Nil => Result.fail Error.panic | list_t.Nil => Result.fail Error.panic @@ -529,25 +508,23 @@ decreasing_by list_nth_mut_shared_loop_pair_loop_decreases ls0 ls1 i /- [loops::list_nth_mut_shared_loop_pair] -/ def list_nth_mut_shared_loop_pair_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - Result (T × T) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : Result (T × T) := list_nth_mut_shared_loop_pair_loop_fwd T ls0 ls1 i /- [loops::list_nth_mut_shared_loop_pair] -/ def list_nth_mut_shared_loop_pair_loop_back - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) (ret0 : T) : + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) (ret0 : T) : (Result (list_t T)) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (list_t.Cons ret0 tl0) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) let tl00 ← list_nth_mut_shared_loop_pair_loop_back T tl0 tl1 i0 ret0 Result.ret (list_t.Cons x0 tl00) @@ -559,25 +536,23 @@ decreasing_by list_nth_mut_shared_loop_pair_loop_decreases ls0 ls1 i /- [loops::list_nth_mut_shared_loop_pair] -/ def list_nth_mut_shared_loop_pair_back - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) (ret0 : T) : + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) (ret0 : T) : Result (list_t T) := list_nth_mut_shared_loop_pair_loop_back T ls0 ls1 i ret0 /- [loops::list_nth_mut_shared_loop_pair_merge] -/ def list_nth_mut_shared_loop_pair_merge_loop_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - (Result (T × T)) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : (Result (T × T)) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (x0, x1) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) list_nth_mut_shared_loop_pair_merge_loop_fwd T tl0 tl1 i0 | list_t.Nil => Result.fail Error.panic | list_t.Nil => Result.fail Error.panic @@ -587,25 +562,23 @@ decreasing_by list_nth_mut_shared_loop_pair_merge_loop_decreases ls0 ls1 i /- [loops::list_nth_mut_shared_loop_pair_merge] -/ def list_nth_mut_shared_loop_pair_merge_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - Result (T × T) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : Result (T × T) := list_nth_mut_shared_loop_pair_merge_loop_fwd T ls0 ls1 i /- [loops::list_nth_mut_shared_loop_pair_merge] -/ def list_nth_mut_shared_loop_pair_merge_loop_back - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) (ret0 : T) : + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) (ret0 : T) : (Result (list_t T)) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (list_t.Cons ret0 tl0) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) let tl00 ← list_nth_mut_shared_loop_pair_merge_loop_back T tl0 tl1 i0 ret0 Result.ret (list_t.Cons x0 tl00) @@ -617,25 +590,23 @@ decreasing_by list_nth_mut_shared_loop_pair_merge_loop_decreases ls0 ls1 i /- [loops::list_nth_mut_shared_loop_pair_merge] -/ def list_nth_mut_shared_loop_pair_merge_back - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) (ret0 : T) : + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) (ret0 : T) : Result (list_t T) := list_nth_mut_shared_loop_pair_merge_loop_back T ls0 ls1 i ret0 /- [loops::list_nth_shared_mut_loop_pair] -/ def list_nth_shared_mut_loop_pair_loop_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - (Result (T × T)) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : (Result (T × T)) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (x0, x1) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) list_nth_shared_mut_loop_pair_loop_fwd T tl0 tl1 i0 | list_t.Nil => Result.fail Error.panic | list_t.Nil => Result.fail Error.panic @@ -645,25 +616,23 @@ decreasing_by list_nth_shared_mut_loop_pair_loop_decreases ls0 ls1 i /- [loops::list_nth_shared_mut_loop_pair] -/ def list_nth_shared_mut_loop_pair_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - Result (T × T) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : Result (T × T) := list_nth_shared_mut_loop_pair_loop_fwd T ls0 ls1 i /- [loops::list_nth_shared_mut_loop_pair] -/ def list_nth_shared_mut_loop_pair_loop_back - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) (ret0 : T) : + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) (ret0 : T) : (Result (list_t T)) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (list_t.Cons ret0 tl1) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) let tl10 ← list_nth_shared_mut_loop_pair_loop_back T tl0 tl1 i0 ret0 Result.ret (list_t.Cons x1 tl10) @@ -675,25 +644,23 @@ decreasing_by list_nth_shared_mut_loop_pair_loop_decreases ls0 ls1 i /- [loops::list_nth_shared_mut_loop_pair] -/ def list_nth_shared_mut_loop_pair_back - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) (ret0 : T) : + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) (ret0 : T) : Result (list_t T) := list_nth_shared_mut_loop_pair_loop_back T ls0 ls1 i ret0 /- [loops::list_nth_shared_mut_loop_pair_merge] -/ def list_nth_shared_mut_loop_pair_merge_loop_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - (Result (T × T)) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : (Result (T × T)) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (x0, x1) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) list_nth_shared_mut_loop_pair_merge_loop_fwd T tl0 tl1 i0 | list_t.Nil => Result.fail Error.panic | list_t.Nil => Result.fail Error.panic @@ -703,25 +670,23 @@ decreasing_by list_nth_shared_mut_loop_pair_merge_loop_decreases ls0 ls1 i /- [loops::list_nth_shared_mut_loop_pair_merge] -/ def list_nth_shared_mut_loop_pair_merge_fwd - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) : - Result (T × T) - := + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) : Result (T × T) := list_nth_shared_mut_loop_pair_merge_loop_fwd T ls0 ls1 i /- [loops::list_nth_shared_mut_loop_pair_merge] -/ def list_nth_shared_mut_loop_pair_merge_loop_back - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) (ret0 : T) : + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) (ret0 : T) : (Result (list_t T)) := match h: ls0 with | list_t.Cons x0 tl0 => match h: ls1 with | list_t.Cons x1 tl1 => - if h: i = (UInt32.ofNatCore 0 (by intlit)) + if h: i = (U32.ofInt 0 (by intlit)) then Result.ret (list_t.Cons ret0 tl1) else do - let i0 ← UInt32.checked_sub i (UInt32.ofNatCore 1 (by intlit)) + let i0 ← i - (U32.ofInt 1 (by intlit)) let tl10 ← list_nth_shared_mut_loop_pair_merge_loop_back T tl0 tl1 i0 ret0 Result.ret (list_t.Cons x1 tl10) @@ -733,7 +698,7 @@ decreasing_by list_nth_shared_mut_loop_pair_merge_loop_decreases ls0 ls1 i /- [loops::list_nth_shared_mut_loop_pair_merge] -/ def list_nth_shared_mut_loop_pair_merge_back - (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : UInt32) (ret0 : T) : + (T : Type) (ls0 : list_t T) (ls1 : list_t T) (i : U32) (ret0 : T) : Result (list_t T) := list_nth_shared_mut_loop_pair_merge_loop_back T ls0 ls1 i ret0 -- cgit v1.2.3