(** Formalization of the base expression language for the daisy framework **) Require Import Coq.Reals.Reals Coq.micromega.Psatz Coq.QArith.QArith Coq.QArith.Qreals. Require Import Daisy.Infra.RealRationalProps. Require Export Daisy.Infra.Abbrevs Daisy.Infra.RealSimps Daisy.Infra.NatSet Daisy.IntervalArithQ Daisy.IntervalArith Daisy.Infra.MachineType. (** Expressions will use binary operators. Define them first **) Inductive binop : Type := Plus | Sub | Mult | Div. (** TODO: simplify pattern matching **) Definition binopEqBool (b1:binop) (b2:binop) := match b1 with Plus => match b2 with Plus => true |_ => false end | Sub => match b2 with Sub => true |_ => false end | Mult => match b2 with Mult => true |_ => false end | Div => match b2 with Div => true |_ => false end end. (** Next define an evaluation function for binary operators on reals. Errors are added on the expression evaluation level later. **) Definition evalBinop (o:binop) (v1:R) (v2:R) := match o with | Plus => Rplus v1 v2 | Sub => Rminus v1 v2 | Mult => Rmult v1 v2 | Div => Rdiv v1 v2 end. Lemma binopEqBool_refl b: binopEqBool b b = true. Proof. case b; auto. Qed. (** Expressions will use unary operators. Define them first **) Inductive unop: Type := Neg | Inv. Definition unopEqBool (o1:unop) (o2:unop) := match o1 with |Neg => match o2 with |Neg => true |_=> false end |Inv => match o2 with |Inv => true |_ => false end end. (** Define evaluation for unary operators on reals. Errors are added in the expression evaluation level later. **) Definition evalUnop (o:unop) (v:R):= match o with |Neg => (- v)%R |Inv => (/ v)%R end . (** Define expressions parametric over some value type V. Will ease reasoning about different instantiations later. **) Inductive exp (V:Type): Type := Var: mType -> nat -> exp V | Const: mType -> V -> exp V | Unop: unop -> exp V -> exp V | Binop: binop -> exp V -> exp V -> exp V | Downcast: mType -> exp V -> exp V. (** Boolean equality function on expressions. Used in certificates to define the analysis result as function **) Fixpoint expEqBool (e1:exp Q) (e2:exp Q) := match e1 with |Var _ m1 v1 => match e2 with |Var _ m2 v2 => andb (mTypeEqBool m1 m2) (v1 =? v2) | _=> false end |Const m1 n1 => match e2 with |Const m2 n2 => andb (mTypeEqBool m1 m2) (Qeq_bool n1 n2) | _=> false end |Unop o1 e11 => match e2 with |Unop o2 e22 => andb (unopEqBool o1 o2) (expEqBool e11 e22) |_ => false end |Binop o1 e11 e12 => match e2 with |Binop o2 e21 e22 => andb (binopEqBool o1 o2) (andb (expEqBool e11 e21) (expEqBool e12 e22)) |_ => false end |Downcast m1 f1 => match e2 with |Downcast m2 f2 => andb (mTypeEqBool m1 m2) (expEqBool f1 f2) |_ => false end end. Lemma expEqBool_refl e: expEqBool e e = true. Proof. induction e; apply andb_true_iff; split; simpl in *; auto; try (apply EquivEqBoolEq; auto). - symmetry; apply beq_nat_refl. - apply Qeq_bool_iff; lra. - case u; auto. - case b; auto. - apply andb_true_iff; split. apply IHe1. apply IHe2. Qed. Fixpoint toRExp (e:exp Q) := match e with |Var _ m v => Var R m v |Const m n => Const m (Q2R n) |Unop o e1 => Unop o (toRExp e1) |Binop o e1 e2 => Binop o (toRExp e1) (toRExp e2) |Downcast m e1 => Downcast m (toRExp e1) end. Fixpoint toREval (e:exp R) := match e with | Var _ _ v => Var R M0 v | Const _ n => Const M0 n | Unop o e1 => Unop o (toREval e1) | Binop o e1 e2 => Binop o (toREval e1) (toREval e2) | Downcast _ e1 => (toREval e1) end. Definition toREvalEnv (E:env) : env := fun (n:nat) => let s := (E n) in match s with | None => None | Some (r, _) => Some (r, M0) end. (** Define a perturbation function to ease writing of basic definitions **) Definition perturb (r:R) (e:R) := (r * (1 + e))%R. (** Define expression evaluation relation parametric by an "error" epsilon. The result value expresses float computations according to the IEEE standard, using a perturbation of the real valued computation by (1 + delta), where |delta| <= machine epsilon. **) Inductive eval_exp (E:env) :(exp R) -> R -> mType -> Prop := | Var_load m (*m1*) x v: (* isMorePrecise m m1 = true ->*) (**mTypeEqBool m m1 = true ->*) E x = Some (v, m) -> eval_exp E (Var R m x) v m | Const_dist m n delta: Rle (Rabs delta) (Q2R (meps m)) -> eval_exp E (Const m n) (perturb n delta) m | Unop_neg m f1 v1: eval_exp E f1 v1 m -> eval_exp E (Unop Neg f1) (evalUnop Neg v1) m | Unop_inv m f1 v1 delta: Rle (Rabs delta) (Q2R (meps m)) -> eval_exp E f1 v1 m -> eval_exp E (Unop Inv f1) (perturb (evalUnop Inv v1) delta) m | Binop_dist m m1 m2 op f1 f2 v1 v2 delta: isJoinOf m m1 m2 = true -> Rle (Rabs delta) (Q2R (meps m)) -> eval_exp E f1 v1 m1 -> eval_exp E f2 v2 m2 -> eval_exp E (Binop op f1 f2) (perturb (evalBinop op v1 v2) delta) m | Downcast_dist m m1 f1 v1 delta: (* Downcast expression f1 (evaluating to machine type m1), to a machine type m, less precise than m1.*) isMorePrecise m1 m = true -> Rle (Rabs delta) (Q2R (meps m)) -> eval_exp E f1 v1 m1 -> eval_exp E (Downcast m f1) (perturb v1 delta) m. (** Define the set of "used" variables of an expression to be the set of variables occuring in it **) Fixpoint usedVars (V:Type) (e:exp V) :NatSet.t := match e with | Var _ _ x => NatSet.singleton x | Unop u e1 => usedVars e1 | Binop b e1 e2 => NatSet.union (usedVars e1) (usedVars e2) | Downcast _ e1 => usedVars e1 | _ => NatSet.empty end. (** If |delta| <= 0 then perturb v delta is exactly v. **) Lemma delta_0_deterministic (v:R) (delta:R): (Rabs delta <= 0)%R -> perturb v delta = v. Proof. intros abs_0; apply Rabs_0_impl_eq in abs_0; subst. unfold perturb. lra. Qed. Lemma general_meps_0_deterministic (f:exp R) (E:env): forall v1 v2 m1, m1 = M0 -> eval_exp E (toREval f) v1 m1 -> eval_exp E (toREval f) v2 M0 -> v1 = v2. Proof. induction f; intros v1 v2 m1 m10_eq eval_v1 eval_v2. - inversion eval_v1; inversion eval_v2; subst; auto; try repeat (repeat rewrite delta_0_deterministic; simpl in *; rewrite Q2R0_is_0 in *; subst; auto); simpl. rewrite H7 in H3; inversion H3; subst; auto. - inversion eval_v1; inversion eval_v2; subst; auto; try repeat (repeat rewrite delta_0_deterministic; simpl in *; rewrite Q2R0_is_0 in *; subst; auto); simpl. - inversion eval_v1; inversion eval_v2; subst; auto; try repeat (repeat rewrite delta_0_deterministic; simpl in *; rewrite Q2R0_is_0 in *; subst; auto); simpl. + apply Ropp_eq_compat. apply (IHf v0 v3 M0); auto. + inversion H4. + inversion H5. + rewrite (IHf v0 v3 M0); auto. - inversion eval_v1; inversion eval_v2; subst; auto; try repeat (repeat rewrite delta_0_deterministic; simpl in *; rewrite Q2R0_is_0 in *; subst; auto); simpl. assert (M0 = M0) as M00 by auto. pose proof (ifM0isJoin_l M0 m0 m2 M00 H2); auto. pose proof (ifM0isJoin_r M0 m0 m2 M00 H2); auto. pose proof (ifM0isJoin_l M0 m4 m5 M00 H11); auto. pose proof (ifM0isJoin_r M0 m4 m5 M00 H11); auto. subst. rewrite (IHf1 v0 v4 M0); auto. rewrite (IHf2 v5 v3 M0); auto. - simpl toREval in eval_v1. simpl toREval in eval_v2. apply (IHf v1 v2 m1); auto. Qed. (** Evaluation with 0 as machine epsilon is deterministic **) Lemma meps_0_deterministic (f:exp R) (E:env): forall v1 v2, eval_exp E (toREval f) v1 M0 -> eval_exp E (toREval f) v2 M0 -> v1 = v2. Proof. intros v1 v2 ev1 ev2. assert (M0 = M0) by auto. apply (general_meps_0_deterministic f H ev1 ev2). Qed. (** Helping lemma. Needed in soundness proof. For each evaluation of using an arbitrary epsilon, we can replace it by evaluating the subexpressions and then binding the result values to different variables in the Environment. **) Lemma binary_unfolding b f1 f2 m E vF: eval_exp E (Binop b f1 f2) vF m -> exists vF1 vF2 m1 m2, eval_exp E f1 vF1 m1 /\ eval_exp E f2 vF2 m2 /\ eval_exp (updEnv 2 m2 vF2 (updEnv 1 m1 vF1 emptyEnv)) (Binop b (Var R m1 1) (Var R m2 2)) vF m. Proof. intros eval_float. inversion eval_float; subst. exists v1 ; exists v2; exists m1; exists m2; repeat split; try auto. eapply Binop_dist; eauto. pose proof (isMorePrecise_refl m1). eapply Var_load; eauto. pose proof (isMorePrecise_refl m2). (* unfold mTypeEqBool; apply Qeq_bool_iff; apply Qeq_refl. *) eapply Var_load; eauto. (* unfold mTypeEqBool; apply Qeq_bool_iff; apply Qeq_refl. *) Qed. (* (** *) (* Analogous lemma for unary expressions. *) (* **) *) Lemma unary_unfolding (e:exp R) (m:mType) (E:env) (v:R): (eval_exp E (Unop Inv e) v m -> exists v1 m1, eval_exp E e v1 m1 /\ eval_exp (updEnv 1 m1 v1 E) (Unop Inv (Var R m1 1)) v m). Proof. intros eval_un. inversion eval_un; subst. exists v1; exists m. repeat split; try auto. econstructor; try auto. pose proof (isMorePrecise_refl m). econstructor; eauto. (* - intros exists_val. *) (* destruct exists_val as [v1 [m1 [eval_f1 eval_e_E]]]. *) (* inversion eval_e_E; subst. *) (* inversion H1; subst. *) (* econstructor; eauto. *) (* unfold updEnv in H6. *) (* simpl in H6. *) (* inversion H6. *) (* rewrite <- H2. *) (* rewrite <- H1. *) (* auto. *) Qed. (** Using the parametric expressions, define boolean expressions for conditionals **) Inductive bexp (V:Type) : Type := leq: exp V -> exp V -> bexp V | less: exp V -> exp V -> bexp V. (** Define evaluation of boolean expressions **) (* Inductive bval (E:env): (bexp R) -> Prop -> Prop := *) (* leq_eval (f1:exp R) (f2:exp R) (v1:R) (v2:R): *) (* eval_exp E f1 v1 -> *) (* eval_exp E f2 v2 -> *) (* bval E (leq f1 f2) (Rle v1 v2) *) (* |less_eval (f1:exp R) (f2:exp R) (v1:R) (v2:R): *) (* eval_exp E f1 v1 -> *) (* eval_exp E f2 v2 -> *) (* bval E (less f1 f2) (Rlt v1 v2). *) (* (** *) (* Simplify arithmetic later by making > >= only abbreviations *) (* **) *) (* Definition gr := fun (V:Type) (f1: exp V) (f2: exp V) => less f2 f1. *) (* Definition greq := fun (V:Type) (f1:exp V) (f2: exp V) => leq f2 f1. *)