(** 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 Daisy.Infra.Ltacs. 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 binopEq (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 binopEq_refl b: binopEq b b = true. Proof. case b; auto. Qed. Lemma binopEq_compat_eq b1 b2: binopEq b1 b2 = true <-> b1 = b2. Proof. split; case b1; case b2; intros; simpl in *; congruence. Qed. Lemma binopEq_compat_eq_false b1 b2: binopEq b1 b2 = false <-> ~ (b1 = b2). Proof. split; intros neq. - hnf; intros; subst. rewrite binopEq_refl in neq. congruence. - destruct b1; destruct b2; cbv; congruence. Qed. (** Expressions will use unary operators. Define them first **) Inductive unop: Type := Neg | Inv. Definition unopEq (o1:unop) (o2:unop) := match o1, o2 with | Neg, Neg => true | Inv, Inv => true | _ , _ => false end. Lemma unopEq_refl b: unopEq b b = true. Proof. case b; auto. Qed. Lemma unopEq_compat_eq b1 b2: unopEq b1 b2 = true <-> b1 = b2. Proof. split; case b1; case b2; intros; simpl in *; congruence. 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 . (** Expressions will use ternary operators Define them first **) Inductive terop: Type := Fma. Definition teropEq (o1: terop) (o2: terop) := match o1, o2 with | Fma, Fma => true end. Lemma teropEq_refl t: teropEq t t = true. Proof. now destruct t. Qed. Lemma teropEq_compat_eq t1 t2: teropEq t1 t2 = true <-> t1 = t2. Proof. now destruct t1, t2. Qed. Definition evalTerop (o: terop) (v1: R) (v2: R) (v3: R) := match o with | Fma => Rplus v1 (Rmult v2 v3) 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 | Terop: terop -> exp V -> 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 expEq (e1:exp Q) (e2:exp Q) := match e1, e2 with | Var _ v1, Var _ v2 => (v1 =? v2) | Const m1 n1, Const m2 n2 => (mTypeEq m1 m2) && (Qeq_bool n1 n2) | Unop o1 e11, Unop o2 e22 => (unopEq o1 o2) && (expEq e11 e22) | Binop o1 e11 e12, Binop o2 e21 e22 => (binopEq o1 o2) && (expEq e11 e21) && (expEq e12 e22) | Terop o1 e11 e12 e13, Terop o2 e21 e22 e23 => (teropEq o1 o2) && (expEq e11 e21) && (expEq e12 e22) && (expEq e13 e23) | Downcast m1 f1, Downcast m2 f2 => (mTypeEq m1 m2) && (expEq f1 f2) | _, _ => false end. Lemma expEq_refl e: expEq e e = true. Proof. induction e; try (apply andb_true_iff; split); simpl in *; auto . - symmetry; apply beq_nat_refl. - apply mTypeEq_refl. - apply Qeq_bool_iff; lra. - case u; auto. - case b; auto. - case t. firstorder. - apply mTypeEq_refl. Qed. Lemma expEq_sym e e': expEq e e' = expEq e' e. Proof. revert e'. induction e; intros e'; destruct e'; simpl; try auto. - apply Nat.eqb_sym. - f_equal. + apply mTypeEq_sym; auto. + apply Qeq_bool_sym. - f_equal. + destruct u; auto. + apply IHe. - f_equal. + f_equal. * destruct b; auto. * apply IHe1. + apply IHe2. - f_equal. + f_equal. * destruct t, t0. simpl. apply IHe1. * apply IHe2. + apply IHe3. - f_equal. + apply mTypeEq_sym; auto. + apply IHe. Qed. Lemma expEq_trans e f g: expEq e f = true -> expEq f g = true -> expEq e g = true. Proof. revert e f g; induction e; destruct f; intros g eq1 eq2; destruct g; simpl in *; try congruence; try rewrite Nat.eqb_eq in *; subst; try auto. - andb_to_prop eq1; andb_to_prop eq2. rewrite mTypeEq_compat_eq in L, L0; subst. rewrite mTypeEq_refl; simpl. rewrite Qeq_bool_iff in *; lra. - andb_to_prop eq1; andb_to_prop eq2. rewrite unopEq_compat_eq in *; subst. rewrite unopEq_refl; simpl. eapply IHe; eauto. - andb_to_prop eq1; andb_to_prop eq2. rewrite binopEq_compat_eq in *; subst. rewrite binopEq_refl; simpl. apply andb_true_iff. split; [eapply IHe1; eauto | eapply IHe2; eauto]. - andb_to_prop eq1; andb_to_prop eq2. rewrite teropEq_compat_eq in *; subst. rewrite teropEq_refl; simpl. rewrite andb_true_iff. rewrite andb_true_iff. split; [split; [eapply IHe1; eauto | eapply IHe2; eauto] | eapply IHe3; eauto]. - andb_to_prop eq1; andb_to_prop eq2. rewrite mTypeEq_compat_eq in *; subst. rewrite mTypeEq_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) | Terop o e1 e2 e3 => Terop o (toRExp e1) (toRExp e2) (toRExp e3) | 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) | Terop o e1 e2 e3 => Terop o (toREval e1) (toREval e2) (toREval e3) | Downcast _ e1 => Downcast M0 (toREval e1) end. Definition toRMap (d:nat -> option mType) (n:nat) := match d n with | Some m => Some M0 | None => None end. Arguments toRMap _ _/. (** Define a perturbation function to ease writing of basic definitions **) Definition perturb (r:R) (e:R) := (r * (1 + e))%R. Hint Unfold perturb. (** 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) (Gamma: nat -> option mType) :(exp R) -> R -> mType -> Prop := | Var_load m x v: Gamma x = Some m -> E x = Some v -> eval_exp E Gamma (Var R x) v m | Const_dist m n delta: Rle (Rabs delta) (Q2R (mTypeToQ m)) -> eval_exp E Gamma (Const m n) (perturb n delta) m | Unop_neg m f1 v1: eval_exp E Gamma f1 v1 m -> eval_exp E Gamma (Unop Neg f1) (evalUnop Neg v1) m | Unop_inv m f1 v1 delta: Rle (Rabs delta) (Q2R (mTypeToQ m)) -> eval_exp E Gamma f1 v1 m -> (~ v1 = 0)%R -> eval_exp E Gamma (Unop Inv f1) (perturb (evalUnop Inv v1) 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 (mTypeToQ m)) -> eval_exp E Gamma f1 v1 m1 -> eval_exp E Gamma (Downcast m f1) (perturb v1 delta) m | Binop_dist m1 m2 op f1 f2 v1 v2 delta: Rle (Rabs delta) (Q2R (mTypeToQ (join m1 m2))) -> eval_exp E Gamma f1 v1 m1 -> eval_exp E Gamma f2 v2 m2 -> ((op = Div) -> (~ v2 = 0)%R) -> eval_exp E Gamma (Binop op f1 f2) (perturb (evalBinop op v1 v2) delta) (join m1 m2) | Terop_dist m1 m2 m3 op f1 f2 f3 v1 v2 v3 delta: Rle (Rabs delta) (Q2R (mTypeToQ (join3 m1 m2 m3))) -> eval_exp E Gamma f1 v1 m1 -> eval_exp E Gamma f2 v2 m2 -> eval_exp E Gamma f3 v3 m3 -> eval_exp E Gamma (Terop op f1 f2 f3) (perturb (evalTerop op v1 v2 v3) delta) (join3 m1 m2 m3). Hint Constructors eval_exp. (** Show some simpler (more general) rule lemmata **) Lemma Const_dist' m n delta v m' E Gamma: Rle (Rabs delta) (Q2R (mTypeToQ m)) -> v = perturb n delta -> m' = m -> eval_exp E Gamma (Const m n) v m'. Proof. intros; subst; auto. Qed. Hint Resolve Const_dist'. Lemma Unop_neg' m f1 v1 v m' E Gamma: eval_exp E Gamma f1 v1 m -> v = evalUnop Neg v1 -> m' = m -> eval_exp E Gamma (Unop Neg f1) v m'. Proof. intros; subst; auto. Qed. Hint Resolve Unop_neg'. Lemma Unop_inv' m f1 v1 delta v m' E Gamma: Rle (Rabs delta) (Q2R (mTypeToQ m)) -> eval_exp E Gamma f1 v1 m -> (~ v1 = 0)%R -> v = perturb (evalUnop Inv v1) delta -> m' = m -> eval_exp E Gamma (Unop Inv f1) v m'. Proof. intros; subst; auto. Qed. Hint Resolve Unop_inv'. Lemma Downcast_dist' m m1 f1 v1 delta v m' E Gamma: isMorePrecise m1 m = true -> Rle (Rabs delta) (Q2R (mTypeToQ m)) -> eval_exp E Gamma f1 v1 m1 -> v = (perturb v1 delta) -> m' = m -> eval_exp E Gamma (Downcast m f1) v m'. Proof. intros; subst; eauto. Qed. Hint Resolve Downcast_dist'. Lemma Binop_dist' m1 m2 op f1 f2 v1 v2 delta v m' E Gamma: Rle (Rabs delta) (Q2R (mTypeToQ m')) -> eval_exp E Gamma f1 v1 m1 -> eval_exp E Gamma f2 v2 m2 -> ((op = Div) -> (~ v2 = 0)%R) -> v = perturb (evalBinop op v1 v2) delta -> m' = join m1 m2 -> eval_exp E Gamma (Binop op f1 f2) v m'. Proof. intros; subst; auto. Qed. Hint Resolve Binop_dist'. Lemma Terop_dist' m1 m2 m3 op f1 f2 f3 v1 v2 v3 delta v m' E Gamma: Rle (Rabs delta) (Q2R (mTypeToQ m')) -> eval_exp E Gamma f1 v1 m1 -> eval_exp E Gamma f2 v2 m2 -> eval_exp E Gamma f3 v3 m3 -> v = perturb (evalTerop op v1 v2 v3) delta -> m' = join3 m1 m2 m3 -> eval_exp E Gamma (Terop op f1 f2 f3) v m'. Proof. intros; subst; auto. Qed. Hint Resolve Terop_dist'. (** 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) | Terop t e1 e2 e3 => NatSet.union (usedVars e1) (NatSet.union (usedVars e2) (usedVars e3)) | Downcast _ e1 => usedVars e1 | _ => NatSet.empty end. Lemma toRMap_eval_M0 f v E Gamma m: eval_exp E (toRMap Gamma) (toREval f) v m -> m = M0. Proof. revert v E Gamma m. induction f; intros * eval_f; inversion eval_f; subst; repeat match goal with | H: context[toRMap _ _] |- _ => unfold toRMap in H | H: context[match ?Gamma ?v with | _ => _ end ] |- _ => destruct (Gamma v) eqn:? | H: Some ?m1 = Some ?m2 |- _ => inversion H; try auto | H: None = Some ?m |- _ => inversion H end; try auto. - eapply IHf; eauto. - eapply IHf; eauto. - assert (m1 = M0) by (eapply IHf1; eauto). assert (m2 = M0) by (eapply IHf2; eauto); subst; auto. - assert (m1 = M0) by (eapply IHf1; eauto). assert (m2 = M0) by (eapply IHf2; eauto). assert (m3 = M0) by (eapply IHf3; eauto); subst; auto. Qed. (** 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. (** Evaluation with 0 as machine epsilon is deterministic **) Lemma meps_0_deterministic (f:exp R) (E:env) Gamma: forall v1 v2, eval_exp E (toRMap Gamma) (toREval f) v1 M0 -> eval_exp E (toRMap Gamma) (toREval f) v2 M0 -> v1 = v2. Proof. induction f; intros v1 v2 ev1 ev2. - inversion ev1; inversion ev2; subst. rewrite H1 in H6. inversion H6; auto. - inversion ev1; inversion ev2; subst. simpl in *. rewrite Q2R0_is_0 in *; repeat (rewrite delta_0_deterministic; try auto). - inversion ev1; inversion ev2; subst; try congruence. + rewrite (IHf v0 v3); eauto. + rewrite (IHf v0 v3); eauto. simpl in *. rewrite Q2R0_is_0 in *; repeat (rewrite delta_0_deterministic; try auto). - inversion ev1; inversion ev2; subst. assert (m0 = M0) by (eapply toRMap_eval_M0; eauto). assert (m3 = M0) by (eapply toRMap_eval_M0; eauto). assert (m1 = M0) by (eapply toRMap_eval_M0; eauto). assert (m2 = M0) by (eapply toRMap_eval_M0; eauto). subst. rewrite (IHf1 v0 v4); try auto. rewrite (IHf2 v3 v5); try auto. simpl in *. rewrite Q2R0_is_0 in *. repeat (rewrite delta_0_deterministic; try auto). - inversion ev1; inversion ev2; subst. assert (m0 = M0) by (eapply toRMap_eval_M0; eauto). assert (m1 = M0) by (eapply toRMap_eval_M0; eauto). assert (m2 = M0) by (eapply toRMap_eval_M0; eauto). assert (m3 = M0) by (eapply toRMap_eval_M0; eauto). assert (m4 = M0) by (eapply toRMap_eval_M0; eauto). assert (m5 = M0) by (eapply toRMap_eval_M0; eauto). subst. rewrite (IHf1 v0 v5); try auto. rewrite (IHf2 v3 v6); try auto. rewrite (IHf3 v4 v7); try auto. simpl in *. rewrite Q2R0_is_0 in *. repeat (rewrite delta_0_deterministic; try auto). - inversion ev1; inversion ev2; subst. apply M0_least_precision in H1; apply M0_least_precision in H7; subst. rewrite (IHf v0 v3); try auto. simpl in *. rewrite Q2R0_is_0 in *. repeat (rewrite delta_0_deterministic; try auto). 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 E v1 v2 m1 m2 Gamma delta: (b = Div -> ~(v2 = 0 )%R) -> (Rabs delta <= Q2R (mTypeToQ (join m1 m2)))%R -> eval_exp E Gamma f1 v1 m1 -> eval_exp E Gamma f2 v2 m2 -> eval_exp E Gamma (Binop b f1 f2) (perturb (evalBinop b v1 v2) delta) (join m1 m2) -> eval_exp (updEnv 2 v2 (updEnv 1 v1 emptyEnv)) (updDefVars 2 m2 (updDefVars 1 m1 Gamma)) (Binop b (Var R 1) (Var R 2)) (perturb (evalBinop b v1 v2) delta) (join m1 m2). Proof. intros no_div_zero eval_f1 eval_f2 eval_float. econstructor; try auto. Qed. Lemma eval_eq_env e: forall E1 E2 Gamma v m, (forall x, E1 x = E2 x) -> eval_exp E1 Gamma e v m -> eval_exp E2 Gamma e v m. Proof. induction e; intros; (match_pat (eval_exp _ _ _ _ _) (fun H => inversion H; subst; simpl in *)); try eauto. eapply Var_load; auto. rewrite <- (H n); auto. Qed. (* (** Analogous lemma for unary expressions. **) Lemma unary_unfolding (e:exp R) (eps:R) (E:env) (v:R): (eval_exp eps E (Unop Inv e) v <-> exists v1, eval_exp eps E e v1 /\ eval_exp eps (updEnv 1 v1 E) (Unop Inv (Var R 1)) v). Proof. split. - intros eval_un. inversion eval_un; subst. exists v1. repeat split; try auto. constructor; try auto. constructor; auto. - intros exists_val. destruct exists_val as [v1 [eval_f1 eval_e_E]]. inversion eval_e_E; subst. inversion H1; subst. unfold updEnv in *; simpl in *. constructor; auto. inversion H3; subst; 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. *)