(** Formalization of the base expression language for the daisy framework **) From Coq Require Import Reals.Reals micromega.Psatz QArith.QArith QArith.Qreals Structures.Orders. 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_sym u1 u2: unopEq u1 u2 = unopEq u2 u1. Proof. destruct u1,u2; compute; 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 . Definition evalFma (v1:R) (v2:R) (v3:R):= evalBinop Plus v1 (evalBinop Mult v2 v3). (** 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 | Fma: 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) | Fma e11 e12 e13, Fma e21 e22 e23 => (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. - 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; auto. + auto. - 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; cbn in *; try rewrite Nat.eqb_eq in *; Daisy_compute; try congruence; type_conv; subst; try auto. - rewrite mTypeEq_refl; simpl. rewrite Qeq_bool_iff in *; lra. - rewrite unopEq_compat_eq in *; subst. rewrite unopEq_refl; simpl. eapply IHe; eauto. - rewrite binopEq_compat_eq in *; subst. rewrite binopEq_refl; simpl. apply andb_true_iff. split; [eapply IHe1; eauto | eapply IHe2; eauto]. - rewrite andb_true_iff. rewrite andb_true_iff. split; [split; [eapply IHe1; eauto | eapply IHe2; eauto] | eapply IHe3; eauto]. - rewrite mTypeEq_refl; simpl. eapply IHe; eauto. Qed. Module Type OrderType := Coq.Structures.Orders.OrderedType. Module ExpOrderedType (V_ordered:OrderType) <: OrderType. Module V_orderedFacts := OrdersFacts.OrderedTypeFacts (V_ordered). Definition V := V_ordered.t. Definition t := exp V. Fixpoint expCompare (e1:exp V) (e2:exp V) := match e1, e2 with |Var _ n1, Var _ n2 => Nat.compare n1 n2 |Var _ n1, _ => Lt | Const m1 v1, Const m2 v2 => if mTypeEq m1 m2 then V_ordered.compare v1 v2 else (if morePrecise m1 m2 then Lt else Gt) | Const _ _, Var _ _ => Gt | Const _ _, _ => Lt | Unop u1 e1, Unop u2 e2 => if unopEq u1 u2 then expCompare e1 e2 else (if unopEq u1 Neg then Lt else Gt) | Unop _ _, Binop _ _ _ => Lt | Unop _ _, Downcast _ _ => Lt | Unop _ _, _ => Gt | Downcast m1 e1, Downcast m2 e2 => if mTypeEq m1 m2 then expCompare e1 e2 else (if morePrecise m1 m2 then Lt else Gt) | Downcast _ _, Binop _ _ _ => Lt | Downcast _ _, _ => Gt | Binop b1 e11 e12, Binop b2 e21 e22 => let res := match b1, b2 with | Plus, Plus => Eq | Plus, _ => Lt | Sub, Sub => Eq | Sub, Plus => Gt | Sub, _ => Lt | Mult, Mult => Eq | Mult, Div => Lt | Mult, _ => Gt | Div, Div => Eq | Div, _ => Gt end in match res with | Eq => match expCompare e11 e21 with | Eq => expCompare e12 e22 | Lt => Lt | Gt => Gt end | _ => res end |_ , _ => Gt end. Lemma expCompare_refl e: expCompare e e = Eq. Proof. induction e; simpl. - apply Nat.compare_refl. - rewrite mTypeEq_refl. apply V_orderedFacts.compare_refl. - rewrite unopEq_refl; auto. - rewrite IHe1, IHe2. destruct b; auto. - rewrite mTypeEq_refl; auto. Qed. (* Lemma expCompare_eq_compat_eq e1: *) (* forall e2, *) (* expCompare e1 e2 = Eq <-> *) (* eq e1 e2. *) (* Proof. *) (* induction e1; destruct e2; split; intros * cmp_res; simpl in *; *) (* subst; try congruence. *) (* - rewrite Nat.compare_eq_iff in cmp_res; subst; auto. *) (* - inversion cmp_res; subst; simpl. apply Nat.compare_refl. *) (* - destruct (mTypeEq m m0 *) (* apply V_orderedFacts.compare_eq in cmp_res. *) (* f_equal. *) (* hnf in cmp_res. *) Lemma expCompare_eq_trans e1 : forall e2 e3, expCompare e1 e2 = Eq -> expCompare e2 e3 = Eq -> expCompare e1 e3 = Eq. Proof. induction e1; intros * eq12 eq23; destruct e2; destruct e3; simpl in *; try congruence. - rewrite Nat.compare_eq_iff in *; subst; auto. - destruct (mTypeEq m m0) eqn:?; [ destruct (mTypeEq m0 m1) eqn:? | destruct (morePrecise m m0) eqn:?; congruence]; [ | destruct (morePrecise m0 m1) eqn:?; congruence]. type_conv. rewrite mTypeEq_refl. rewrite V_orderedFacts.compare_eq_iff in *; eapply V_orderedFacts.eq_trans; eauto. - destruct (unopEq u u0) eqn:?; destruct (unopEq u0 u1) eqn:?; try rewrite unopEq_compat_eq in *; subst; try rewrite unopEq_refl; try congruence. + eapply IHe1; eauto. + destruct (unopEq u0 Neg); congruence. + destruct (unopEq u Neg); congruence. + destruct (unopEq u Neg); congruence. - destruct b; destruct b0; try congruence; destruct b1; try congruence; destruct (expCompare e1_1 e2_1) eqn:?; destruct (expCompare e2_1 e3_1) eqn:?; try congruence; try erewrite IHe1_1; eauto. - destruct (mTypeEq m m0) eqn:?; destruct (mTypeEq m0 m1) eqn:?; type_conv; try rewrite mTypeEq_refl. + eapply IHe1; eauto. + destruct (morePrecise m0 m1); congruence. + destruct (morePrecise m m1); congruence. + destruct (morePrecise m m0); congruence. Qed. Lemma expCompare_antisym e1: forall e2, expCompare e1 e2 = CompOpp (expCompare e2 e1). Proof. induction e1; destruct e2; simpl; try auto. - apply Nat.compare_antisym. - rewrite mTypeEq_sym. destruct (mTypeEq m0 m) eqn:?; type_conv; try congruence; try rewrite mTypeEq_refl. + apply V_orderedFacts.compare_antisym. + destruct (morePrecise m m0) eqn:?; destruct (morePrecise m0 m) eqn:?; try (split; auto; fail). * pose proof (morePrecise_antisym _ _ Heqb0 Heqb1); type_conv; congruence. * destruct m, m0; unfold morePrecise in *; cbv; congruence. - rewrite unopEq_sym. destruct (unopEq u0 u) eqn:?; try rewrite unopEq_compat_eq in *; subst; try rewrite unopEq_refl, IHe1; try (apply IHe1). destruct (unopEq u Neg) eqn:?; try rewrite unopEq_compat_eq in *; destruct (unopEq u0 Neg) eqn:?; try rewrite unopEq_compat_eq in *; subst; simpl in *; try congruence. destruct u, u0; simpl in *; congruence. - destruct b, b0; simpl; try (split; auto; fail); destruct (expCompare e1_1 e2_1) eqn:first_comp; rewrite IHe1_1 in *; simpl in *; rewrite CompOpp_iff in first_comp; rewrite first_comp; simpl; try auto. - rewrite mTypeEq_sym. destruct (mTypeEq m0 m) eqn:?; type_conv; try auto. + destruct (morePrecise m m0) eqn:?; destruct (morePrecise m0 m) eqn:?; try (split; auto; fail). * pose proof (morePrecise_antisym _ _ Heqb0 Heqb1); type_conv; congruence. * destruct m, m0; unfold morePrecise in *; cbv; congruence. Qed. Lemma expCompare_lt_eq_is_lt e1: forall e2 e3, expCompare e1 e2 = Lt -> expCompare e2 e3 = Eq -> expCompare e1 e3 = Lt. Proof. induction e1; intros * compare_lt compare_eq; destruct e2; simpl in *; destruct e3; try congruence. - rewrite Nat.compare_eq_iff in compare_eq; subst; auto. - destruct (mTypeEq m m0) eqn:?; destruct (mTypeEq m0 m1) eqn:?. + pose proof (V_orderedFacts.compare_compat). unfold Proper in H. apply V_orderedFacts.compare_eq_iff in compare_eq. specialize (H v v (V_orderedFacts.eq_refl v) v0 v1 compare_eq). type_conv; rewrite mTypeEq_refl, <- H; auto. + rewrite mTypeEq_compat_eq in Heqb; subst. rewrite Heqb0. destruct (morePrecise m0 m1) eqn:?; congruence. + rewrite mTypeEq_compat_eq in Heqb0; subst. rewrite Heqb; destruct (morePrecise m m1) eqn:?; congruence. + destruct (morePrecise m0 m1); congruence. - destruct (unopEq u u0) eqn:?; destruct (unopEq u0 u1) eqn:?; try rewrite unopEq_compat_eq in *; subst; try rewrite unopEq_refl; try auto; try congruence. + eapply IHe1; eauto. + destruct (unopEq u0 Neg); congruence. + destruct (unopEq u Neg); try congruence. destruct (unopEq u u1); congruence. + destruct (unopEq u0 Neg); congruence. - destruct b; destruct b0; try congruence; destruct b1; try congruence; destruct (expCompare e1_1 e2_1) eqn:?; destruct (expCompare e2_1 e3_1) eqn:?; try congruence; try (erewrite IHe1_1; eauto; fail ""); try erewrite expCompare_eq_trans; eauto. - destruct (mTypeEq m m0) eqn:?; destruct (mTypeEq m0 m1) eqn:?. + type_conv; subst. rewrite mTypeEq_refl. eapply IHe1; eauto. + destruct (morePrecise m0 m1); congruence. + rewrite mTypeEq_compat_eq in Heqb0; subst. rewrite Heqb. destruct (morePrecise m m1) eqn:?; congruence. + destruct (morePrecise m0 m1); congruence. Qed. Lemma expCompare_eq_lt_is_lt e1: forall e2 e3, expCompare e1 e2 = Eq -> expCompare e2 e3 = Lt -> expCompare e1 e3 = Lt. Proof. induction e1; intros * compare_eq compare_lt; destruct e2; simpl in *; destruct e3; try congruence. - rewrite Nat.compare_eq_iff in compare_eq; subst; auto. - destruct (mTypeEq m m0) eqn:?; destruct (mTypeEq m0 m1) eqn:?. + pose proof (V_orderedFacts.compare_compat). unfold Proper in H. apply V_orderedFacts.compare_eq_iff in compare_eq. specialize (H v v0 compare_eq v1 v1 (V_orderedFacts.eq_refl v1)). type_conv; rewrite mTypeEq_refl, H; auto. + rewrite mTypeEq_compat_eq in Heqb; subst. rewrite Heqb0. destruct (morePrecise m0 m1) eqn:?; congruence. + rewrite mTypeEq_compat_eq in Heqb0; subst. rewrite Heqb; destruct (morePrecise m m1) eqn:?; congruence. + destruct (morePrecise m m0); congruence. - destruct (unopEq u u0) eqn:?; destruct (unopEq u0 u1) eqn:?; try rewrite unopEq_compat_eq in *; subst; try rewrite unopEq_refl; try auto; try congruence. + eapply IHe1; eauto. + rewrite Heqb0. destruct (unopEq u0 Neg); congruence. + destruct (unopEq u Neg); congruence. + destruct (unopEq u Neg); congruence. - destruct b; destruct b0; destruct b1; try congruence; destruct (expCompare e1_1 e2_1) eqn:?; destruct (expCompare e2_1 e3_1) eqn:?; try congruence; try (erewrite IHe1_1; eauto; fail ""); try erewrite expCompare_eq_trans; eauto. - destruct (mTypeEq m m0) eqn:?; destruct (mTypeEq m0 m1) eqn:?. + type_conv; subst. rewrite mTypeEq_refl. eapply IHe1; eauto. + rewrite mTypeEq_compat_eq in Heqb; subst. rewrite Heqb0. destruct (morePrecise m0 m1); congruence. + rewrite mTypeEq_compat_eq in Heqb0; subst. rewrite Heqb. destruct (morePrecise m m1) eqn:?; congruence. + destruct (morePrecise m m0); congruence. Qed. Definition eq e1 e2 := expCompare e1 e2 = Eq. Definition lt (e1:exp V) (e2: exp V):= expCompare e1 e2 = Lt. Instance lt_strorder : StrictOrder lt. Proof. split. - unfold Irreflexive. unfold Reflexive. intros x; unfold complement. intros lt_x. induction x; unfold lt in *; simpl in lt_x. + rewrite PeanoNat.Nat.compare_refl in lt_x. congruence. + rewrite mTypeEq_refl, V_orderedFacts.compare_refl in *; congruence. + rewrite unopEq_refl in *; simpl in *. apply IHx; auto. + destruct b; destruct (expCompare x1 x1) eqn:?; try congruence. + rewrite mTypeEq_refl in lt_x. apply IHx; auto. - unfold Transitive. intros e1. unfold lt. induction e1; intros * lt_e1_e2 lt_e2_e3; simpl in *; destruct y; destruct z; simpl in *; try auto; try congruence. + rewrite <- nat_compare_lt in *. omega. + destruct (mTypeEq m m0) eqn:?; destruct (mTypeEq m0 m1) eqn:?. * type_conv; rewrite mTypeEq_refl, V_orderedFacts.compare_lt_iff in *; eapply V_orderedFacts.lt_trans; eauto. * rewrite mTypeEq_compat_eq in Heqb; subst. rewrite Heqb0. assumption. * rewrite mTypeEq_compat_eq in Heqb0; subst. rewrite Heqb; assumption. * destruct (mTypeEq m m1) eqn:?. { rewrite mTypeEq_compat_eq in Heqb1; subst. destruct (morePrecise m0 m1) eqn:?; destruct (morePrecise m1 m0) eqn:?; try congruence. pose proof (morePrecise_antisym _ _ Heqb1 Heqb2). type_conv; congruence. } { destruct (morePrecise m m0) eqn:?; destruct (morePrecise m0 m1) eqn:?; try congruence. erewrite morePrecise_trans; eauto. } + destruct (unopEq u u0) eqn:?; destruct (unopEq u0 u1) eqn:?; try rewrite unopEq_compat_eq in *; subst; [ destruct (expCompare e1 y) eqn:?; try congruence; rewrite unopEq_refl; eapply IHe1; eauto | destruct (unopEq u0 Neg) eqn:?; try congruence; rewrite unopEq_compat_eq in *; subst | |]. * rewrite Heqb0; auto. * destruct (unopEq u Neg) eqn:?; try congruence; rewrite unopEq_compat_eq in *; subst. rewrite Heqb; auto. * destruct (unopEq u u1) eqn:?; try congruence. rewrite unopEq_compat_eq in Heqb1; subst. destruct (unopEq u1 Neg) eqn:?; try congruence; destruct (unopEq u0 Neg) eqn:?; try congruence; rewrite unopEq_compat_eq in *; subst. simpl in *; congruence. + destruct b; destruct b0; destruct b1; try congruence; destruct (expCompare e1_1 y1) eqn:?; try congruence; destruct (expCompare y1 z1) eqn:?; try congruence; try (erewrite expCompare_eq_trans; eauto; fail); try (erewrite expCompare_eq_lt_is_lt; eauto; fail); try (erewrite expCompare_lt_eq_is_lt; eauto; fail); try (erewrite IHe1_1; eauto). + destruct (mTypeEq m m0) eqn:?; destruct (mTypeEq m0 m1) eqn:?; [type_conv; subst; rewrite mTypeEq_refl | | | ]. * eapply IHe1; eauto. * rewrite mTypeEq_compat_eq in Heqb; subst. rewrite Heqb0; destruct (morePrecise m0 m1); congruence. * rewrite mTypeEq_compat_eq in Heqb0; subst. rewrite Heqb. destruct (morePrecise m m1); congruence. * destruct (mTypeEq m m1) eqn:?. { rewrite mTypeEq_compat_eq in Heqb1; subst. destruct (morePrecise m1 m0) eqn:?; try congruence. destruct (morePrecise m0 m1) eqn:?; try congruence. pose proof (morePrecise_antisym _ _ Heqb1 Heqb2). type_conv; subst. congruence. } { destruct (morePrecise m m0) eqn:?; try congruence. destruct (morePrecise m0 m1) eqn:?; try congruence. pose proof (morePrecise_trans _ _ _ Heqb2 Heqb3). rewrite H; auto. } Defined. Lemma eq_compat: Proper (eq ==> eq ==> iff) eq. Proof. unfold Proper; hnf. intros e1; induction e1; intros e2 e1_eq_e2; hnf; intros e3 e4 e3_eq_e4; unfold lt, eq in *; destruct e2,e3,e4; simpl in *; try congruence; try (split; auto; fail). - repeat rewrite Nat.compare_eq_iff in *; subst. split; try auto. - destruct (mTypeEq m m0) eqn:?; destruct (mTypeEq m1 m2) eqn:?; [type_conv | | |]. + rewrite V_orderedFacts.compare_eq_iff in *. rewrite (V_orderedFacts.compare_compat e1_eq_e2 e3_eq_e4). split; auto. + destruct (morePrecise m1 m2); congruence. + destruct (morePrecise m m0); congruence. + destruct (morePrecise m m0); congruence. - destruct (unopEq u u0) eqn:?; destruct (unopEq u1 u2) eqn:?; try rewrite unopEq_compat_eq in *; subst; try (destruct (unopEq u Neg); congruence); try (destruct (unopEq u1 Neg); congruence). specialize (IHe1 e2 e1_eq_e2 e3 e4 e3_eq_e4). simpl in *. destruct (unopEq u0 u2); try rewrite IHe1; split; auto. - destruct b; destruct b0; destruct b1; destruct b2; try congruence; try (split; auto; fail); destruct (expCompare e1_1 e2_1) eqn:?; destruct (expCompare e3_1 e4_1) eqn:?; try congruence; destruct (expCompare e1_1 e3_1) eqn:?; destruct (expCompare e2_1 e4_1) eqn:?; try (split; congruence); try (specialize (IHe1_2 _ e1_eq_e2 _ _ e3_eq_e4); simpl in *; rewrite IHe1_2 in *; split; auto; fail); try (split; try congruence; intros); try (specialize (IHe1_1 _ Heqc _ _ Heqc0); simpl in *; rewrite IHe1_1 in *; congruence); try (specialize (IHe1_1 _ Heqc _ _ Heqc0); simpl in *; rewrite <- IHe1_1 in *; congruence). - destruct (mTypeEq m m0) eqn:?; destruct (mTypeEq m1 m2) eqn:?; [type_conv | | |]. + specialize (IHe1 _ e1_eq_e2 _ _ e3_eq_e4); simpl in *. destruct (mTypeEq m0 m2); try congruence. split; auto. + destruct (morePrecise m1 m2); congruence. + destruct (morePrecise m m0); congruence. + destruct (morePrecise m m0); congruence. Qed. Instance lt_compat: Proper (eq ==> eq ==> iff) lt. Proof. unfold Proper; hnf. intros e1; induction e1; intros e2 e1_eq_e2; hnf; intros e3 e4 e3_eq_e4; unfold lt, eq in *; destruct e2,e3,e4; simpl in *; try congruence; try (split; auto; fail). - rewrite Nat.compare_eq_iff in *; subst. split; try auto. - destruct (mTypeEq m m0) eqn:?; destruct (mTypeEq m1 m2) eqn:?; [type_conv | | |]. + rewrite V_orderedFacts.compare_eq_iff in *. rewrite (V_orderedFacts.compare_compat e1_eq_e2 e3_eq_e4). split; auto. + destruct (morePrecise m1 m2); congruence. + destruct (morePrecise m m0); congruence. + destruct (morePrecise m m0); congruence. - destruct (unopEq u u0) eqn:?; destruct (unopEq u1 u2) eqn:?; try rewrite unopEq_compat_eq in *; subst; try (destruct (unopEq u Neg); congruence); try (destruct (unopEq u1 Neg); congruence). specialize (IHe1 e2 e1_eq_e2 e3 e4 e3_eq_e4). simpl in *. destruct (unopEq u0 u2); try rewrite IHe1; split; auto. - pose proof eq_compat as eq_comp. unfold Proper, eq in eq_comp. destruct b, b0, b1, b2; try congruence; try (split; auto; fail); destruct (expCompare e1_1 e2_1) eqn:?; destruct (expCompare e3_1 e4_1) eqn:?; try congruence; destruct (expCompare e1_1 e3_1) eqn:?; destruct (expCompare e2_1 e4_1) eqn:?; try (split; congruence); try (specialize (IHe1_2 _ e1_eq_e2 _ _ e3_eq_e4); simpl in *; rewrite IHe1_2 in *; split; auto; fail); try (split; try congruence; intros); try (specialize (IHe1_1 _ Heqc _ _ Heqc0); simpl in *; rewrite IHe1_1 in *; congruence); try (specialize (IHe1_1 _ Heqc _ _ Heqc0); simpl in *; rewrite <- IHe1_1 in *; congruence); try (rewrite (eq_comp _ _ Heqc _ _ Heqc0) in *; congruence); try (rewrite <- (eq_comp _ _ Heqc _ _ Heqc0) in *; congruence). - destruct (mTypeEq m m0) eqn:?; destruct (mTypeEq m1 m2) eqn:?; [type_conv | | |]. + specialize (IHe1 _ e1_eq_e2 _ _ e3_eq_e4); simpl in *. destruct (mTypeEq m0 m2); try congruence. split; auto. + destruct (morePrecise m1 m2); congruence. + destruct (morePrecise m m0); congruence. + destruct (morePrecise m m0); congruence. Defined. Lemma compare_spec : forall x y, CompSpec eq lt x y (expCompare x y). Proof. intros e1 e2. destruct (expCompare e1 e2) eqn:?. - apply CompEq. unfold eq; auto. - apply CompLt. unfold lt; auto. - apply CompGt. unfold lt. rewrite expCompare_antisym in Heqc. rewrite CompOpp_iff in Heqc. simpl in *; auto. Qed. Instance eq_equiv: Equivalence eq. Proof. split; unfold Reflexive, Symmetric, Transitive, eq. - apply expCompare_refl. - intros. rewrite expCompare_antisym in * |-. rewrite CompOpp_iff in * |- . auto. - apply expCompare_eq_trans. Defined. Parameter eq_dec : forall x y, { eq x y } + { ~ eq x y }. Definition eq_refl : forall x, eq x x. Proof. apply expCompare_refl. Defined. Definition eq_sym : forall x y, eq x y -> eq y x. Proof. unfold eq; intros. rewrite expCompare_antisym in * |-. rewrite CompOpp_iff in * |-. auto. Defined. Definition eq_trans : forall x y z, eq x y -> eq y z -> eq x z. Proof. apply expCompare_eq_trans. Defined. Definition lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z. Proof. pose proof lt_strorder as [_ Trans]. apply Trans. Defined. Definition lt_not_eq : forall x y : t, lt x y -> ~ eq x y. Proof. intros. unfold lt,eq in *. hnf; intros; congruence. Defined. Definition compare e1 e2:= expCompare e1 e2. (* Definition compare (e1 e2:t) :Compare lt eq e1 e2. *) (* Proof. *) (* destruct (expCompare e1 e2) eqn:?. *) (* - eapply EQ. unfold eq; auto. *) (* - eapply LT; auto. *) (* - eapply GT. rewrite expCompare_antisym in * |-. *) (* rewrite CompOpp_iff in *. *) (* auto. *) (* Defined. *) End ExpOrderedType. 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) | Fma e1 e2 e3 => Fma (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) | Fma e1 e2 e3 => Fma (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) | Fma_dist m1 m2 m3 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 (Fma f1 f2 f3) (perturb (evalFma 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 Fma_dist' m1 m2 m3 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 (evalFma v1 v2 v3) delta -> m' = join3 m1 m2 m3 -> eval_exp E Gamma (Fma f1 f2 f3) v m'. Proof. intros; subst; auto. Qed. Hint Resolve Fma_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) | Fma 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 lemmas. 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 fma_unfolding f1 f2 f3 E v1 v2 v3 m1 m2 m3 Gamma delta: (Rabs delta <= Q2R (mTypeToQ (join3 m1 m2 m3)))%R -> 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 (Fma f1 f2 f3) (perturb (evalFma v1 v2 v3) delta) (join3 m1 m2 m3) -> eval_exp (updEnv 3 v3 (updEnv 2 v2 (updEnv 1 v1 emptyEnv))) (updDefVars 3 m3 (updDefVars 2 m2 (updDefVars 1 m1 Gamma))) (Fma (Var R 1) (Var R 2) (Var R 3)) (perturb (evalFma v1 v2 v3) delta) (join3 m1 m2 m3). Proof. 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. *)