X-Git-Url: http://git.megacz.com/?a=blobdiff_plain;f=src%2FGeneral.v;h=3c8252849faa7cbca0b3e22963216901e25adaef;hb=8efffc7368b5e54c42461f45a9708ff2828409a4;hp=f3e4379049e4a5a5664dea49f325ba7410e9b9eb;hpb=8c26722a1ee110077968a8a166eb7130266b2035;p=coq-hetmet.git diff --git a/src/General.v b/src/General.v index f3e4379..3c82528 100644 --- a/src/General.v +++ b/src/General.v @@ -615,9 +615,24 @@ Lemma extensionality_composes : forall t1 t2 t3 (f f':t1->t2) (g g':t2->t3), + + +CoInductive Fresh A T := +{ next : forall a:A, (T a * Fresh A T) +; split : Fresh A T * Fresh A T +}. + + + + + Definition map2 {A}{B}(f:A->B)(t:A*A) : (B*B) := ((f (fst t)), (f (snd t))). +(* string stuff *) +Variable eol : string. +Extract Constant eol => "'\n':[]". + (* the Error monad *) Inductive OrError (T:Type) := @@ -634,6 +649,18 @@ Definition orErrorBind {T:Type} (oe:OrError T) {Q:Type} (f:T -> OrError Q) := end. Notation "a >>= b" := (@orErrorBind _ a _ b) (at level 20). +Open Scope string_scope. +Definition orErrorBindWithMessage {T:Type} (oe:OrError T) {Q:Type} (f:T -> OrError Q) err_msg := + match oe with + | Error s => Error (err_msg +++ eol +++ " " +++ s) + | OK t => f t + end. + +Notation "a >>=[ S ] b" := (@orErrorBindWithMessage _ a _ b S) (at level 20). + +Definition addErrorMessage s {T} (x:OrError T) := + x >>=[ s ] (fun y => OK y). + Inductive Indexed {T:Type}(f:T -> Type) : ???T -> Type := | Indexed_Error : forall error_message:string, Indexed f (Error error_message) | Indexed_OK : forall t, f t -> Indexed f (OK t)