(* *********************************************************************) (* *) (* The Compcert verified compiler *) (* *) (* Xavier Leroy, INRIA Paris-Rocquencourt *) (* *) (* Copyright Institut National de Recherche en Informatique et en *) (* Automatique. All rights reserved. This file is distributed *) (* under the terms of the GNU General Public License as published by *) (* the Free Software Foundation, either version 2 of the License, or *) (* (at your option) any later version. This file is also distributed *) (* under the terms of the INRIA Non-Commercial License Agreement. *) (* *) (* *********************************************************************) include "datatypes/pairs.ma". include "Plogic/equality.ma". include "Plogic/connectives.ma". (* * Error reporting and the error monad. *) (* (** * Representation of error messages. *) (** Compile-time errors produce an error message, represented in Coq as a list of either substrings or positive numbers encoding a source-level identifier (see module AST). *) Inductive errcode: Type := | MSG: string -> errcode | CTX: positive -> errcode. Definition errmsg: Type := list errcode. Definition msg (s: string) : errmsg := MSG s :: nil. *) (* * * The error monad *) (* * Compilation functions that can fail have return type [res A]. The return value is either [OK res] to indicate success, or [Error msg] to indicate failure. *) ninductive res (A: Type) : Type ≝ | OK: A → res A | Error: (* FIXME errmsg →*) res A. (*Implicit Arguments Error [A].*) (* * To automate the propagation of errors, we use a monadic style with the following [bind] operation. *) ndefinition bind ≝ λA,B:Type. λf: res A. λg: A → res B. match f with [ OK x ⇒ g x | Error (*msg*) ⇒ Error ? (*msg*) ]. ndefinition bind2 ≝ λA,B,C: Type. λf: res (A × B). λg: A → B → res C. match f with [ OK v ⇒ match v with [ mk_pair x y ⇒ g x y ] | Error (*msg*) => Error ? (*msg*) ]. (* Not sure what level to use *) notation > "vbox('do' ident v ← e; break e')" with precedence 40 for @{'bind ${e} (λ${ident v}.${e'})}. notation > "vbox('do' ident v : ty ← e; break e')" with precedence 40 for @{'bind ${e} (λ${ident v} : ${ty}.${e'})}. notation < "vbox('do' \nbsp ident v ← e; break e')" with precedence 40 for @{'bind ${e} (λ${ident v}.${e'})}. notation < "vbox('do' \nbsp ident v : ty ← e; break e')" with precedence 40 for @{'bind ${e} (λ${ident v} : ${ty}.${e'})}. interpretation "error monad bind" 'bind e f = (bind ?? e f). notation > "vbox('do' 〈ident v1, ident v2〉 ← e; break e')" with precedence 40 for @{'bind2 ${e} (λ${ident v1}.λ${ident v2}.${e'})}. notation > "vbox('do' 〈ident v1 : ty1, ident v2 : ty2〉 ← e; break e')" with precedence 40 for @{'bind2 ${e} (λ${ident v1} : ${ty1}.λ${ident v2} : ${ty2}.${e'})}. notation < "vbox('do' \nbsp 〈ident v1, ident v2〉 ← e; break e')" with precedence 40 for @{'bind2 ${e} (λ${ident v1}.λ${ident v2}.${e'})}. notation < "vbox('do' \nbsp 〈ident v1 : ty1, ident v2 : ty2〉 ← e; break e')" with precedence 40 for @{'bind2 ${e} (λ${ident v1} : ${ty1}.λ${ident v2} : ${ty2}.${e'})}. interpretation "error monad pair bind" 'bind2 e f = (bind2 ??? e f). (*interpretation "error monad ret" 'ret e = (ret ? e). notation "'ret' e" non associative with precedence 45 for @{'ret ${e}}.*) (* (** The [do] notation, inspired by Haskell's, keeps the code readable. *) Notation "'do' X <- A ; B" := (bind A (fun X => B)) (at level 200, X ident, A at level 100, B at level 200) : error_monad_scope. Notation "'do' ( X , Y ) <- A ; B" := (bind2 A (fun X Y => B)) (at level 200, X ident, Y ident, A at level 100, B at level 200) : error_monad_scope. *) nremark bind_inversion: ∀A,B: Type. ∀f: res A. ∀g: A → res B. ∀y: B. bind ?? f g = OK ? y → ∃x. f = OK ? x ∧ g x = OK ? y. #A B f g y; ncases f; ##[ #a e; @a; nwhd in e:(??%?); /2/; ##| #H; nwhd in H:(??%?); ndestruct (H); ##] nqed. nremark bind2_inversion: ∀A,B,C: Type. ∀f: res (A×B). ∀g: A → B → res C. ∀z: C. bind2 ??? f g = OK ? z → ∃x. ∃y. f = OK ? 〈x, y〉 ∧ g x y = OK ? z. #A B C f g z; ncases f; ##[ #ab; ncases ab; #a b e; @a; @b; nwhd in e:(??%?); /2/; ##| #H; nwhd in H:(??%?); ndestruct ##] nqed. (* Open Local Scope error_monad_scope. (** This is the familiar monadic map iterator. *) Fixpoint mmap (A B: Type) (f: A -> res B) (l: list A) {struct l} : res (list B) := match l with | nil => OK nil | hd :: tl => do hd' <- f hd; do tl' <- mmap f tl; OK (hd' :: tl') end. Remark mmap_inversion: forall (A B: Type) (f: A -> res B) (l: list A) (l': list B), mmap f l = OK l' -> list_forall2 (fun x y => f x = OK y) l l'. Proof. induction l; simpl; intros. inversion_clear H. constructor. destruct (bind_inversion _ _ H) as [hd' [P Q]]. destruct (bind_inversion _ _ Q) as [tl' [R S]]. inversion_clear S. constructor. auto. auto. Qed. (** * Reasoning over monadic computations *) (** The [monadInv H] tactic below simplifies hypotheses of the form << H: (do x <- a; b) = OK res >> By definition of the bind operation, both computations [a] and [b] must succeed for their composition to succeed. The tactic therefore generates the following hypotheses: x: ... H1: a = OK x H2: b x = OK res *) Ltac monadInv1 H := match type of H with | (OK _ = OK _) => inversion H; clear H; try subst | (Error _ = OK _) => discriminate | (bind ?F ?G = OK ?X) => let x := fresh "x" in ( let EQ1 := fresh "EQ" in ( let EQ2 := fresh "EQ" in ( destruct (bind_inversion F G H) as [x [EQ1 EQ2]]; clear H; try (monadInv1 EQ2)))) | (bind2 ?F ?G = OK ?X) => let x1 := fresh "x" in ( let x2 := fresh "x" in ( let EQ1 := fresh "EQ" in ( let EQ2 := fresh "EQ" in ( destruct (bind2_inversion F G H) as [x1 [x2 [EQ1 EQ2]]]; clear H; try (monadInv1 EQ2))))) | (mmap ?F ?L = OK ?M) => generalize (mmap_inversion F L H); intro end. Ltac monadInv H := match type of H with | (OK _ = OK _) => monadInv1 H | (Error _ = OK _) => monadInv1 H | (bind ?F ?G = OK ?X) => monadInv1 H | (bind2 ?F ?G = OK ?X) => monadInv1 H | (?F _ _ _ _ _ _ _ _ = OK _) => ((progress simpl in H) || unfold F in H); monadInv1 H | (?F _ _ _ _ _ _ _ = OK _) => ((progress simpl in H) || unfold F in H); monadInv1 H | (?F _ _ _ _ _ _ = OK _) => ((progress simpl in H) || unfold F in H); monadInv1 H | (?F _ _ _ _ _ = OK _) => ((progress simpl in H) || unfold F in H); monadInv1 H | (?F _ _ _ _ = OK _) => ((progress simpl in H) || unfold F in H); monadInv1 H | (?F _ _ _ = OK _) => ((progress simpl in H) || unfold F in H); monadInv1 H | (?F _ _ = OK _) => ((progress simpl in H) || unfold F in H); monadInv1 H | (?F _ = OK _) => ((progress simpl in H) || unfold F in H); monadInv1 H end. *)