(* *********************************************************************) (* *) (* 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 "basics/types.ma". include "basics/logic.ma". include "basics/list.ma". include "common/PreIdentifiers.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[0] := | MSG: String → errcode | CTX: ∀tag:String. identifier tag → errcode. definition errmsg: Type[0] ≝ list errcode. definition msg : String → errmsg ≝ λs. [MSG s]. (* * * 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. *) inductive res (A: Type[0]) : Type[0] ≝ | OK: A → res A | Error: errmsg → res A. (*Implicit Arguments Error [A].*) (* * To automate the propagation of errors, we use a monadic style with the following [bind] operation. *) definition bind ≝ λA,B:Type[0]. λf: res A. λg: A → res B. match f with [ OK x ⇒ g x | Error msg ⇒ Error ? msg ]. definition bind2 ≝ λA,B,C: Type[0]. λf: res (A × B). λg: A → B → res C. match f with [ OK v ⇒ match v with [ pair x y ⇒ g x y ] | Error msg => Error ? msg ]. definition bind3 ≝ λA,B,C,D: Type[0]. λf: res (A × B × C). λg: A → B → C → res D. match f with [ OK v ⇒ match v with [ pair xy z ⇒ match xy with [ pair x y ⇒ g x y z ] ] | 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 Prod bind" 'bind2 e f = (bind2 ??? e f). notation > "vbox('do' 〈ident v1, ident v2, ident v3〉 ← e; break e')" with precedence 40 for @{'bind3 ${e} (λ${ident v1}.λ${ident v2}.λ${ident v3}.${e'})}. notation > "vbox('do' 〈ident v1 : ty1, ident v2 : ty2, ident v3 : ty3〉 ← e; break e')" with precedence 40 for @{'bind3 ${e} (λ${ident v1} : ${ty1}.λ${ident v2} : ${ty2}.λ${ident v3} : ${ty3}.${e'})}. notation < "vbox('do' \nbsp 〈ident v1, ident v2, ident v3〉 ← e; break e')" with precedence 40 for @{'bind3 ${e} (λ${ident v1}.λ${ident v2}.λ${ident v3}.${e'})}. notation < "vbox('do' \nbsp 〈ident v1 : ty1, ident v2 : ty2, ident v3 : ty3〉 ← e; break e')" with precedence 40 for @{'bind3 ${e} (λ${ident v1} : ${ty1}.λ${ident v2} : ${ty2}.λ${ident v3} : ${ty3}.${e'})}. interpretation "error monad triple bind" 'bind3 e f = (bind3 ???? 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. *) lemma bind_inversion: ∀A,B: Type[0]. ∀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 cases f; [ #a #e %{a} whd in e:(??%?); /2/; | #m #H whd in H:(??%?); destruct (H); ] qed. lemma bind2_inversion: ∀A,B,C: Type[0]. ∀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 cases f; [ #ab cases ab; #a #b #e %{a} %{b} whd in e:(??%?); /2/; | #m #H whd in H:(??%?); destruct ] qed. (* Open Local Scope error_monad_scope. (** This is the familiar monadic map iterator. *) *) let rec mmap (A, B: Type[0]) (f: A → res B) (l: list A) on l : res (list B) ≝ match l with [ nil ⇒ OK ? [] | cons hd tl ⇒ do hd' ← f hd; do tl' ← mmap ?? f tl; OK ? (hd'::tl') ]. (* lemma mmap_inversion: ∀A, B: Type[0]. ∀f: A -> res B. ∀l: list A. ∀l': list B. mmap A B 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. *) (* A monadic fold_left2 *) axiom WrongLength: String. let rec mfold_left2 (A: Type[0]) (B: Type[0]) (C: Type[0]) (f: A → B → C → res A) (accu: res A) (left: list B) (right: list C) on left: res A ≝ match left with [ nil ⇒ match right with [ nil ⇒ accu | cons hd tl ⇒ Error ? (msg WrongLength) ] | cons hd tl ⇒ match right with [ nil ⇒ Error ? (msg WrongLength) | cons hd' tl' ⇒ do accu ← accu; mfold_left2 … f (f accu hd hd') tl tl' ] ]. (* (** * 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. *) definition opt_to_res ≝ λA.λmsg.λv:option A. match v with [ None ⇒ Error A msg | Some v ⇒ OK A v ]. lemma opt_OK: ∀A,m,P,e. (∀v. e = Some ? v → P v) → match opt_to_res A m e with [ Error _ ⇒ True | OK v ⇒ P v ]. #A #m #P #e elim e; /2/; qed. definition res_to_opt : ∀A:Type[0]. res A → option A ≝ λA.λv. match v with [ Error _ ⇒ None ? | OK v ⇒ Some … v].