include "common/IOMonad.ma". include "common/Events.ma". (* Checking types of values given to / received from an external function call. *) definition eventval_type : ∀ty:typ. Type[0] ≝ λty. match ty with [ ASTint ⇒ int | ASTptr _ ⇒ False | ASTfloat ⇒ float ]. definition mk_eventval: ∀ty:typ. eventval_type ty → eventval ≝ λty:typ. match ty return λty'.eventval_type ty' → eventval with [ ASTint ⇒ λv.EVint v | ASTptr r ⇒ ? | ASTfloat ⇒ λv.EVfloat v ]. *; qed. definition mk_val: ∀ty:typ. eventval_type ty → val ≝ λty:typ. match ty return λty'.eventval_type ty' → val with [ ASTint ⇒ λv.Vint v | ASTptr r ⇒ ? | ASTfloat ⇒ λv.Vfloat v ]. *; qed. lemma mk_val_correct: ∀ty:typ. ∀v:eventval_type ty. eventval_match (mk_eventval ty v) ty (mk_val ty v). #ty cases ty; normalize; // * *; qed. definition convert_eventval : ∀ev:eventval. ∀ty:typ. res val ≝ λev,ty. match ty with [ ASTint ⇒ match ev with [ EVint i ⇒ OK ? (Vint i) | _ ⇒ Error ? ] | ASTfloat ⇒ match ev with [ EVfloat f ⇒ OK ? (Vfloat f) | _ ⇒ Error ? ] | _ ⇒ Error ? ]. definition check_eventval' : ∀v:val. ∀ty:typ. res eventval ≝ λv,ty. match ty with [ ASTint ⇒ match v with [ Vint i ⇒ OK ? (EVint i) | _ ⇒ Error ? ] | ASTfloat ⇒ match v with [ Vfloat f ⇒ OK ? (EVfloat f) | _ ⇒ Error ? ] | _ ⇒ Error ? ]. let rec check_eventval_list (vs: list val) (tys: list typ) : res (list eventval) ≝ match vs with [ nil ⇒ match tys with [ nil ⇒ OK ? (nil ?) | _ ⇒ Error ? ] | cons v vt ⇒ match tys with [ nil ⇒ Error ? | cons ty tyt ⇒ do ev ← check_eventval' v ty; do evt ← check_eventval_list vt tyt; OK ? (ev::evt) ] ]. (* IO monad *) (* Interactions are function calls that return a value and do not change the rest of the Clight program's state. *) record io_out : Type[0] ≝ { io_function: ident ; io_args: list eventval ; io_in_typ: typ }. definition io_in ≝ λo. eventval_type (io_in_typ o). definition do_io : ident → list eventval → ∀t:typ. IO io_out io_in (eventval_type t) ≝ λfn,args,t. Interact ??? (mk_io_out fn args t) (λres. Value ??? res). definition ret: ∀T. T → (IO io_out io_in T) ≝ λT,x.(Value ?? T x).