X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FGeneral.v;h=5d19e9c60bf7b75bdd3cf32962dda8dcefbf07b5;hp=f97989d2677f6c1cbafc6783ea70b1423901a9fc;hb=cacf56c9e223e864884317718b09c33bd6a37635;hpb=794719aadd55d760c0514f45c23b9cb450d92d9f diff --git a/src/General.v b/src/General.v index f97989d..5d19e9c 100644 --- a/src/General.v +++ b/src/General.v @@ -478,6 +478,7 @@ Fixpoint listToString {T:Type}{tst:ToString T}(l:list T) : string := Instance ListToString {T:Type}{tst:ToString T} : ToString (list T) := { toString := @listToString _ _ }. + (*******************************************************************************) (* Tree Flags *) @@ -952,6 +953,16 @@ Lemma list2vecOrFail {T}(l:list T)(n:nat)(error_message:nat->nat->string) : ???( apply (Error (error_message (length l) n)). Defined. +(* this makes a type function application, ensuring not to oversaturate it (though if it was undersaturated we can't fix that) *) +Fixpoint split_list {T}(l:list T)(n:nat) : ???(list T * list T) := + match n with + | O => OK (nil , l) + | S n' => match l with + | nil => Error "take_list failed" + | h::t => split_list t n' >>= fun t' => let (t1,t2) := t' in OK ((h::t1),t2) + end + end. + (* Uniques *) Variable UniqSupply : Type. Extract Inlined Constant UniqSupply => "UniqSupply.UniqSupply". Variable Unique : Type. Extract Inlined Constant Unique => "Unique.Unique".