(** 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 Export Daisy.Infra.Abbrevs Daisy.Infra.RealSimps Daisy.Infra.NatSet. (** 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 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: 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 |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) := (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 (eps:R) (E:env) :(exp R) -> R -> Prop := | Var_load x v: E x = Some v -> eval_exp eps E (Var R x) v | Const_dist n delta: Rle (Rabs delta) eps -> eval_exp eps E (Const n) (perturb n delta) | Unop_neg f1 v1: eval_exp eps E f1 v1 -> eval_exp eps E (Unop Neg f1) (evalUnop Neg v1) | Unop_inv f1 v1 delta: Rle (Rabs delta) eps -> eval_exp eps E f1 v1 -> eval_exp eps E (Unop Inv f1) (perturb (evalUnop Inv v1) delta) | Binop_dist op f1 f2 v1 v2 delta: Rle (Rabs delta) eps -> eval_exp eps E f1 v1 -> eval_exp eps E f2 v2 -> eval_exp eps E (Binop op f1 f2) (perturb (evalBinop op v1 v2) delta). (** 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 => freeVars e1 | Binop b e1 e2 => NatSet.union (freeVars e1) (freeVars e2) | _ => 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. (** Evaluation with 0 as machine epsilon is deterministic **) Lemma meps_0_deterministic (f:exp R) (E:env): forall v1 v2, eval_exp 0 E f v1 -> eval_exp 0 E 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 H3 in H0; inversion H0; 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. **) Lemma binary_unfolding b f1 f2 eps E vF: eval_exp eps E (Binop b f1 f2) vF -> exists vF1 vF2, eval_exp eps E f1 vF1 /\ eval_exp eps E f2 vF2 /\ eval_exp eps (updEnv 2 vF2 (updEnv 1 vF1 emptyEnv)) (Binop b (Var R 1) (Var R 2)) vF. Proof. intros eval_float. inversion eval_float; subst. exists v1 ; exists v2; repeat split; try auto. constructor; try auto. - constructor. unfold updEnv; cbv; auto. - constructor. unfold updEnv; cbv; 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 H2; 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 (eps:R) (E:env): (bexp R) -> Prop -> Prop := leq_eval (f1:exp R) (f2:exp R) (v1:R) (v2:R): eval_exp eps E f1 v1 -> eval_exp eps E f2 v2 -> bval eps E (leq f1 f2) (Rle v1 v2) |less_eval (f1:exp R) (f2:exp R) (v1:R) (v2:R): eval_exp eps E f1 v1 -> eval_exp eps E f2 v2 -> bval eps 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.