(** 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 Daisy.Infra.RationalSimps. 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. Definition binopEqBool (b1:binop) (b2:binop) := match b1, b2 with | Plus, Plus => true | Sub, Sub => true | Mult, Mult => true | Div, Div => true | _,_ => false 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, o2 with | Neg, Neg => true | Inv, Inv => true | _ , _ => false end. Lemma unopEqBool_refl b: unopEqBool b b = true. Proof. case b; auto. Qed. (** 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: 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, e2 with | Var _ v1, Var _ v2 => (v1 =? v2) | Const m1 n1, Const m2 n2 => andb (mTypeEqBool m1 m2) (Qeq_bool n1 n2) | Unop o1 e11, Unop o2 e22 => andb (unopEqBool o1 o2) (expEqBool e11 e22) | Binop o1 e11 e12, Binop o2 e21 e22 => andb (binopEqBool o1 o2) (andb (expEqBool e11 e21) (expEqBool e12 e22)) | Downcast m1 f1, Downcast m2 f2 => andb (mTypeEqBool m1 m2) (expEqBool f1 f2) | _, _ => false end. Lemma expEqBool_refl e: expEqBool e e = true. Proof. induction e; try (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. Lemma beq_nat_sym a b: beq_nat a b = beq_nat b a. Proof. case_eq (a =? b); intros. - apply beq_nat_true in H. rewrite H. apply beq_nat_refl. - apply beq_nat_false in H. case_eq (b =? a); intros. + apply beq_nat_true in H0. rewrite H0 in H. auto. + auto. Qed. Lemma expEqBool_sym e e': expEqBool e e' = expEqBool e' e. Proof. revert e'. induction e; intros e'; destruct e'; simpl; try auto. - apply beq_nat_sym. - f_equal. + apply mTypeEqBool_sym; auto. + apply Qeq_bool_sym. - f_equal. + destruct u; auto. + apply IHe. - f_equal. + destruct b; auto. + f_equal. * apply IHe1. * apply IHe2. - f_equal. + apply mTypeEqBool_sym; auto. + apply IHe. Qed. Lemma expEqBool_trans e f g: expEqBool e f = true -> expEqBool f g = true -> expEqBool e g = true. Proof. revert e f g; induction e; destruct f; intros; simpl in H; inversion H; rewrite H; clear H; destruct g; simpl in H0; inversion H0; rewrite H0; clear H0; try (apply andb_true_iff in H1; destruct H1; apply andb_true_iff in H2; destruct H2; simpl). - apply beq_nat_true in H2. apply beq_nat_true in H1. subst. unfold expEqBool. rewrite <- beq_nat_refl. auto. - apply EquivEqBoolEq in H1. apply EquivEqBoolEq in H. subst. rewrite mTypeEqBool_refl; simpl. apply Qeq_bool_iff in H2. apply Qeq_bool_iff in H0. apply Qeq_bool_iff. lra. - assert (u = u0) by (destruct u; destruct u0; inversion H1; auto). assert (u0 = u1) by (destruct u0; destruct u1; inversion H; auto). subst. assert (unopEqBool u1 u1 = true) by (destruct u1; auto). apply andb_true_iff; split; try auto. eapply IHe; eauto. - apply andb_true_iff; split. + destruct b; destruct b0; destruct b1; auto. + apply andb_true_iff in H2; destruct H2. apply andb_true_iff in H0; destruct H0. apply andb_true_iff; split. eapply IHe1; eauto. eapply IHe2; eauto. - apply EquivEqBoolEq in H1. apply EquivEqBoolEq in H. subst. rewrite mTypeEqBool_refl; simpl. eapply IHe; eauto. Qed. Fixpoint toRExp (e:exp Q) := match e with |Var _ v => Var R 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 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. (* TODO: put to REValVars here? *) (** 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) (defVars: nat -> option mType) :(exp R) -> R -> mType -> Prop := | Var_load m x v: defVars x = Some m -> E x = Some v -> eval_exp E defVars (Var R x) v m | Const_dist m n delta: Rle (Rabs delta) (Q2R (meps m)) -> eval_exp E defVars (Const m n) (perturb n delta) m | Unop_neg m f1 v1: eval_exp E defVars f1 v1 m -> eval_exp E defVars (Unop Neg f1) (evalUnop Neg v1) m | Unop_inv m f1 v1 delta: Rle (Rabs delta) (Q2R (meps m)) -> eval_exp E defVars f1 v1 m -> eval_exp E defVars (Unop Inv f1) (perturb (evalUnop Inv v1) delta) m | Binop_dist m1 m2 op f1 f2 v1 v2 delta: Rle (Rabs delta) (Q2R (meps (computeJoin m1 m2))) -> eval_exp E defVars f1 v1 m1 -> eval_exp E defVars f2 v2 m2 -> ((op = Div) -> (~ v2 = 0)%R) -> eval_exp E defVars (Binop op f1 f2) (perturb (evalBinop op v1 v2) delta) (computeJoin m1 m2) | 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 defVars f1 v1 m1 -> eval_exp E defVars (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. (* TODO: need of `general` case? *) Lemma general_meps_0_deterministic (f:exp R) (E:env) defVars: forall v1 v2 m1, m1 = M0 -> eval_exp E defVars (toREval f) v1 m1 -> eval_exp E defVars (toREval f) v2 M0 -> v1 = v2. Proof. induction f; intros * 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 H6 in H1; inversion H1; 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. destruct m0; destruct m2; inversion H5. destruct m3; destruct m4; inversion H11. simpl in *. rewrite (IHf1 v0 v4 M0); auto. rewrite (IHf2 v5 v3 M0); auto. rewrite Q2R0_is_0 in H2,H12. rewrite delta_0_deterministic; auto. rewrite delta_0_deterministic; auto. - simpl toREval in eval_v1. simpl toREval in eval_v2. apply (IHf v1 v2 m1); auto. Qed. (* Lemma rnd_0_deterministic f E m v: *) (* eval_exp E (toREval (Downcast m f)) v M0 <-> *) (* eval_exp E (toREval f) v M0. *) (* Proof. *) (* split; intros. *) (* - simpl in H. auto. *) (* - simpl; auto. *) (* Qed. *) (** Evaluation with 0 as machine epsilon is deterministic **) Lemma meps_0_deterministic (f:exp R) (E:env) defVars: forall v1 v2, eval_exp E defVars (toREval f) v1 M0 -> eval_exp E defVars (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. Fixpoint toREvalVars (d:nat -> option mType) (n:nat) := match d n with | Some m => Some M0 | None => None end. (** 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. **) (* TODO: write updDefVars function as well *) Lemma binary_unfolding b f1 f2 m E vF defVars: eval_exp E defVars (Binop b f1 f2) vF m -> exists vF1 vF2 m1 m2, m = computeJoin m1 m2 /\ eval_exp E defVars f1 vF1 m1 /\ eval_exp E defVars f2 vF2 m2 /\ eval_exp (updEnv 2 vF2 (updEnv 1 vF1 emptyEnv)) (fun n => if (n =? 2) then Some m2 else if (n =? 1) then Some m1 else defVars n) (Binop b (Var R 1) (Var R 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). eapply Var_load; eauto. Qed. (* (* Analogous lemma for unary expressions. *) *) (* Lemma unary_unfolding (e:exp R) (m:mType) (E:env) (v:R) defVars: *) (* (eval_exp E defVars (Unop Inv e) v m -> *) (* exists v1 m1, *) (* eval_exp E defVars e v1 m1 /\ *) (* eval_exp (updEnv 1 v1 E) (fun n => if (n =? 1) then Some m1 else defVars n) (Unop Inv (Var R 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. *) (* 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. *)