(** Formalization of the base expression language for the daisy framework Required in all files, since we will always reason about expressions. **) Require Import Coq.Reals.Reals Coq.micromega.Psatz Coq.QArith.QArith. Require Import Daisy.Infra.RealSimps Daisy.Infra.Abbrevs. Set Implicit Arguments. (** Expressions will use binary operators. Define them first **) Inductive binop : Type := Plus | Sub | Mult | Div. 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. (** 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 on 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. Note that we differentiate between wether we use a variable from the environment or a parameter. Parameters do not have error bounds in the invariants, so they must be perturbed, but variables from the program will be perturbed upon binding, so we do not need to perturb them. **) Inductive exp (V:Type): Type := Var: nat -> exp V | Param: nat -> exp V | Const: V -> exp V | Unop: unop -> exp V -> exp V | Binop: binop -> exp V -> 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 _ v1 => match e2 with |Var _ v2 => v1 =? v2 | _=> false end |Param _ v1 => match e2 with |Param _ v2 => v1 =? v2 | _=> false end |Const n1 => match e2 with |Const n2 => 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 end. (** Define a perturbation function to ease writing of basic definitions **) Definition perturb (r:R) (e:R) := Rmult r (Rplus 1 e). (** Define expression evaluation relation parametric by an "error" epsilon. This value will be used later to express float computations using a perturbation of the real valued computation by (1 + delta), where |delta| <= machine epsilon. It is important that variables are not perturbed when loading from an environment. This is the case, since loading a float value should not increase an additional error. Unary negation is special! We do not have a new error here since IEE 754 gives us a sign bit **) Inductive eval_exp (eps:R) (env:env_ty) : (exp R) -> R -> Prop := Var_load x: eval_exp eps env (Var R x) (env x) | Param_acc x delta: ((Rabs delta) <= eps)%R -> eval_exp eps env (Param R x) (perturb (env x) delta) | Const_dist n delta: Rle (Rabs delta) eps -> eval_exp eps env (Const n) (perturb n delta) | Unop_neg f1 v1: eval_exp eps env f1 v1 -> eval_exp eps env (Unop Neg f1) (evalUnop Neg v1) | Unop_inv f1 v1 delta: Rle (Rabs delta) eps -> eval_exp eps env f1 v1 -> eval_exp eps env (Unop Inv f1) (perturb (evalUnop Inv v1) delta) | Binop_dist op f1 f2 v1 v2 delta: Rle (Rabs delta) eps -> eval_exp eps env f1 v1 -> eval_exp eps env f2 v2 -> eval_exp eps env (Binop op f1 f2) (perturb (evalBinop op v1 v2) delta). (** 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) (env:env_ty): forall v1 v2, eval_exp R0 env f v1 -> eval_exp R0 env f v2 -> v1 = v2. Proof. induction f; intros v1 v2 eval_v1 eval_v2; inversion eval_v1; inversion eval_v2; repeat try rewrite delta_0_deterministic; subst; auto. - rewrite (IHf v0 v3); auto. - inversion H3. - inversion H4. - rewrite (IHf v0 v3); auto. - rewrite (IHf1 v0 v4); auto. rewrite (IHf2 v3 v5); 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. This relies on the property that variables are not perturbed as opposed to parameters **) Lemma binary_unfolding (b:binop) (f1:exp R) (f2:exp R) (eps:R) (cenv:env_ty) (v:R): (eval_exp eps cenv (Binop b f1 f2) v <-> exists v1 v2, eval_exp eps cenv f1 v1 /\ eval_exp eps cenv f2 v2 /\ eval_exp eps (updEnv 2 v2 (updEnv 1 v1 cenv)) (Binop b (Var R 1) (Var R 2)) v). Proof. split. - intros eval_bin. inversion eval_bin; subst. exists v1, v2. repeat split; try auto. constructor; try auto; constructor; auto. - intros exists_val. destruct exists_val as [v1 [v2 [eval_f1 [eval_f2 eval_e_env]]]]. inversion eval_e_env; subst. inversion H4; inversion H5; subst. unfold updEnv in *; simpl in *. constructor; auto. Qed. (** Analogous lemma for unary expressions. **) Lemma unary_unfolding (e:exp R) (eps:R) (cenv:env_ty) (v:R): (eval_exp eps cenv (Unop Inv e) v <-> exists v1, eval_exp eps cenv e v1 /\ eval_exp eps (updEnv 1 v1 cenv) (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_env]]. inversion eval_e_env; subst. inversion H1; subst. unfold updEnv in *; simpl in *. constructor; 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 booleans for reals **) Inductive bval (eps:R) (env:env_ty) : (bexp R) -> Prop -> Prop := leq_eval (f1:exp R) (f2:exp R) (v1:R) (v2:R): eval_exp eps env f1 v1 -> eval_exp eps env f2 v2 -> bval eps env (leq f1 f2) (Rle v1 v2) |less_eval (f1:exp R) (f2:exp R) (v1:R) (v2:R): eval_exp eps env f1 v1 -> eval_exp eps env f2 v2 -> bval eps env (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.