From ea38ead03eca866924d8ae5b1d2f20675bf66670 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Fri, 22 Dec 2023 14:21:55 -0500 Subject: [PATCH 1/3] Move AbstractInterpretation to Bottomify and Fancy In preparation for making a fancier but slower bounds analysis pass. Note that this commit does not build on its own, and is here purely to make `git` notice renaming/copying of files. --- .../{ => Bottomify}/AbstractInterpretation.v | 0 src/AbstractInterpretation/Bottomify/Proofs.v | 1559 +++++++++++++++++ src/AbstractInterpretation/Bottomify/Wf.v | 1188 +++++++++++++ .../Bottomify/WfExtra.v | 35 + .../Fancy/AbstractInterpretation.v | 593 +++++++ .../{ => Fancy}/Proofs.v | 8 +- src/AbstractInterpretation/{ => Fancy}/Wf.v | 2 +- .../{ => Fancy}/WfExtra.v | 12 +- 8 files changed, 3386 insertions(+), 11 deletions(-) rename src/AbstractInterpretation/{ => Bottomify}/AbstractInterpretation.v (100%) create mode 100644 src/AbstractInterpretation/Bottomify/Proofs.v create mode 100644 src/AbstractInterpretation/Bottomify/Wf.v create mode 100644 src/AbstractInterpretation/Bottomify/WfExtra.v create mode 100644 src/AbstractInterpretation/Fancy/AbstractInterpretation.v rename src/AbstractInterpretation/{ => Fancy}/Proofs.v (99%) rename src/AbstractInterpretation/{ => Fancy}/Wf.v (99%) rename src/AbstractInterpretation/{ => Fancy}/WfExtra.v (72%) diff --git a/src/AbstractInterpretation/AbstractInterpretation.v b/src/AbstractInterpretation/Bottomify/AbstractInterpretation.v similarity index 100% rename from src/AbstractInterpretation/AbstractInterpretation.v rename to src/AbstractInterpretation/Bottomify/AbstractInterpretation.v diff --git a/src/AbstractInterpretation/Bottomify/Proofs.v b/src/AbstractInterpretation/Bottomify/Proofs.v new file mode 100644 index 0000000000..2021039c19 --- /dev/null +++ b/src/AbstractInterpretation/Bottomify/Proofs.v @@ -0,0 +1,1559 @@ +Require Import Coq.micromega.Lia. +Require Import Coq.ZArith.ZArith. +Require Import Coq.Classes.Morphisms. +Require Import Coq.Classes.RelationPairs. +Require Import Coq.Relations.Relations. +Require Import Crypto.Util.ZRange. +Require Import Crypto.Util.ZRange.Operations. +Require Import Crypto.Util.ZRange.BasicLemmas. +Require Import Crypto.Util.ZRange.SplitBounds. +Require Import Crypto.Util.Sum. +Require Import Crypto.Util.LetIn. +Require Import Crypto.Util.Prod. +Require Import Crypto.Util.Sigma. +Require Import Crypto.Util.Option. +Require Import Crypto.Util.ListUtil. +Require Import Crypto.Util.NatUtil. +Require Import Crypto.Util.ZUtil.AddGetCarry. +Require Import Crypto.Util.ZUtil.AddModulo. +Require Import Crypto.Util.ZUtil.CC. +Require Import Crypto.Util.ZUtil.MulSplit. +Require Import Crypto.Util.ZUtil.Rshi. +Require Import Crypto.Util.ZUtil.Zselect. +Require Import Crypto.Util.ZUtil.Tactics.LtbToLt. +Require Import Crypto.Util.ZUtil.Tactics.SplitMinMax. +Require Import Crypto.Util.ZUtil.Tactics.ReplaceNegWithPos. +Require Import Crypto.Util.HProp. +Require Import Crypto.Util.PER. +Require Import Crypto.Util.Tactics.BreakMatch. +Require Import Crypto.Util.Tactics.DestructHead. +Require Import Crypto.Util.Tactics.SplitInContext. +Require Import Crypto.Util.Tactics.UniquePose. +Require Import Crypto.Util.Tactics.SpecializeBy. +Require Import Crypto.Util.Tactics.SpecializeAllWays. +Require Import Crypto.Util.Tactics.Head. +Require Import Crypto.Util.Tactics.DoWithHyp. +Require Import Rewriter.Language.Language. +Require Import Rewriter.Language.Inversion. +Require Import Rewriter.Language.Wf. +Require Import Crypto.Language.API. +Require Import Crypto.CastLemmas. +Require Import Rewriter.Language.UnderLetsProofs. +Require Import Crypto.AbstractInterpretation.Bottomify.AbstractInterpretation. +Require Import Crypto.AbstractInterpretation.Bottomify.Wf. +Require Import Crypto.AbstractInterpretation.ZRangeProofs. + +Module Compilers. + Import Language.Compilers. + Import UnderLets.Compilers. + Import AbstractInterpretation.Compilers. + Import Language.Inversion.Compilers. + Import Language.Wf.Compilers. + Import Language.API.Compilers. + Import UnderLetsProofs.Compilers. + Import AbstractInterpretation.Bottomify.Wf.Compilers. + Import AbstractInterpretation.ZRangeProofs.Compilers. + Import AbstractInterpretation.Bottomify.Wf.Compilers.partial. + Import invert_expr. + + Local Notation related_bounded' b v1 v2 + := (ZRange.type.base.option.is_bounded_by b v1 = true + /\ ZRange.type.base.option.is_bounded_by b v2 = true + /\ v1 = v2) (only parsing). + Local Notation related_bounded + := (@type.related_hetero3 _ _ _ _ (fun t b v1 v2 => related_bounded' b v1 v2)). + + Module ZRange. + Module type. + Global Instance is_bounded_by_Proper_related_eq {t} + : Proper (eq ==> type.eqv ==> eq) (@ZRange.type.is_bounded_by t). + Proof. + cbv [Proper respectful]; destruct t; [ cbn | reflexivity ]; + intros; subst; reflexivity. + Qed. + + Module option. + Global Instance is_bounded_by_Proper_related_eq {t} + : Proper (eq ==> type.eqv ==> eq) (@ZRange.type.option.is_bounded_by t). + Proof. + cbv [Proper respectful]; destruct t; [ cbn | reflexivity ]; + intros; subst; reflexivity. + Qed. + End option. + End type. + End ZRange. + + Module Import partial. + Import AbstractInterpretation.Compilers.partial. + Import UnderLets.Compilers.UnderLets. + Section with_type. + Context {base_type : Type}. + Local Notation type := (type base_type). + Let type_base (x : base_type) : type := type.base x. + Local Coercion type_base : base_type >-> type. + Context {ident : type -> Type}. + Local Notation Expr := (@expr.Expr base_type ident). + Context (abstract_domain' base_interp : base_type -> Type) + (ident_interp : forall t, ident t -> type.interp base_interp t) + (abstraction_relation' : forall t, abstract_domain' t -> base_interp t -> Prop) + (bottom' : forall A, abstract_domain' A) + (bottom'_related : forall t v, abstraction_relation' t (bottom' t) v) + (skip_annotations_under : forall t, ident t -> bool) + (abstract_interp_ident : forall t, ident t -> type.interp abstract_domain' t) + (ident_interp_Proper : forall t (idc : ident t), type.related_hetero abstraction_relation' (abstract_interp_ident t idc) (ident_interp t idc)) + (ident_interp_Proper' : forall t, Proper (eq ==> type.eqv) (ident_interp t)) + (abstract_domain'_R : forall t, abstract_domain' t -> abstract_domain' t -> Prop) + (abstraction_relation'_Proper : forall t, Proper (abstract_domain'_R t ==> eq ==> Basics.impl) (abstraction_relation' t)) + {abstract_domain'_R_transitive : forall t, Transitive (@abstract_domain'_R t)} + {abstract_domain'_R_symmetric : forall t, Symmetric (@abstract_domain'_R t)} + {bottom'_Proper : forall t, Proper (abstract_domain'_R t) (bottom' t)} + (abstract_domain'_R_of_related : forall t st v, @abstraction_relation' t st v -> @abstract_domain'_R t st st). + Local Notation abstract_domain := (@abstract_domain base_type abstract_domain'). + Definition abstraction_relation {t} : abstract_domain t -> type.interp base_interp t -> Prop + := type.related_hetero (@abstraction_relation'). + Local Notation bottom := (@bottom base_type abstract_domain' (@bottom')). + Local Notation bottom_for_each_lhs_of_arrow := (@bottom_for_each_lhs_of_arrow base_type abstract_domain' (@bottom')). + Local Notation abstract_domain_R := (@abstract_domain_R base_type abstract_domain' abstract_domain'_R). + Local Notation var := (type.interp base_interp). + Local Notation expr := (@expr.expr base_type ident). + Local Notation UnderLets := (@UnderLets base_type ident). + Local Notation value := (@value base_type ident var abstract_domain'). + Local Notation value_with_lets := (@value_with_lets base_type ident var abstract_domain'). + Local Notation state_of_value := (@state_of_value base_type ident var abstract_domain' bottom'). + Context (annotate : forall (is_let_bound : bool) t, abstract_domain' t -> @expr var t -> @UnderLets var (@expr var t)) + (interp_ident : bool -> forall t, ident t -> value_with_lets t) + (ident_extract : forall t, ident t -> abstract_domain t) + (interp_annotate + : forall is_let_bound t st e + (He : abstraction_relation' t st (expr.interp (t:=type.base t) (@ident_interp) e)), + expr.interp (@ident_interp) (UnderLets.interp (@ident_interp) (@annotate is_let_bound t st e)) + = expr.interp (@ident_interp) e) + (ident_extract_Proper : forall t, Proper (eq ==> abstract_domain_R) (ident_extract t)). + Local Notation eta_expand_with_bound' := (@eta_expand_with_bound' base_type ident _ abstract_domain' annotate bottom'). + Local Notation eval_with_bound' := (@partial.eval_with_bound' base_type ident _ abstract_domain' annotate bottom' skip_annotations_under interp_ident). + Local Notation extract' := (@extract' base_type ident abstract_domain' bottom' ident_extract). + Local Notation extract_gen := (@extract_gen base_type ident abstract_domain' bottom' ident_extract). + Local Notation reify := (@reify base_type ident _ abstract_domain' annotate bottom'). + Local Notation reflect := (@reflect base_type ident _ abstract_domain' annotate bottom'). + Local Notation interp := (@interp base_type ident var abstract_domain' annotate bottom' skip_annotations_under interp_ident). + Local Notation bottomify := (@bottomify base_type ident _ abstract_domain' bottom'). + + Lemma bottom_related t v : @abstraction_relation t bottom v. + Proof using bottom'_related. cbv [abstraction_relation]; induction t; cbn; cbv [respectful_hetero]; eauto. Qed. + + Local Hint Resolve bottom_related : core typeclass_instances. + + Lemma bottom_for_each_lhs_of_arrow_related t v : type.and_for_each_lhs_of_arrow (@abstraction_relation) (@bottom_for_each_lhs_of_arrow t) v. + Proof using bottom'_related. induction t; cbn; eauto using bottom_related. Qed. + + Local Notation bottom_Proper := (@bottom_Proper base_type abstract_domain' bottom' abstract_domain'_R bottom'_Proper). + Local Notation bottom_for_each_lhs_of_arrow_Proper := (@bottom_for_each_lhs_of_arrow_Proper base_type abstract_domain' bottom' abstract_domain'_R bottom'_Proper). + + Local Hint Resolve (@bottom_Proper) (@bottom_for_each_lhs_of_arrow_Proper) : core typeclass_instances. + + Fixpoint related_bounded_value {t} : abstract_domain t -> value t -> type.interp base_interp t -> Prop + := match t return abstract_domain t -> value t -> type.interp base_interp t -> Prop with + | type.base t + => fun st '(e_st, e) v + => abstract_domain'_R t st e_st + /\ expr.interp ident_interp e = v + /\ abstraction_relation' t st v + | type.arrow s d + => fun st e v + => Proper type.eqv v + /\ forall st_s e_s v_s, + let st_s := match s with + | type.base _ => st_s + | type.arrow _ _ => bottom + end in + @related_bounded_value s st_s e_s v_s + -> @related_bounded_value d (st st_s) (UnderLets.interp ident_interp (e e_s)) (v v_s) + end. + Definition related_bounded_value_with_lets {t} : abstract_domain t -> value_with_lets t -> type.interp base_interp t -> Prop + := fun st e v => related_bounded_value st (UnderLets.interp ident_interp e) v. + + Definition related_of_related_bounded_value {t} st e v + : @related_bounded_value t st e v -> v == v. + Proof using Type. destruct t; [ reflexivity | intros [? ?]; assumption ]. Qed. + + Lemma abstract_domain'_R_refl_of_rel_l t x y (H : @abstract_domain'_R t x y) + : @abstract_domain'_R t x x. + Proof using abstract_domain'_R_symmetric abstract_domain'_R_transitive. eapply PER_valid_l; eassumption. Qed. + + Lemma abstract_domain'_R_refl_of_rel_r t x y (H : @abstract_domain'_R t x y) + : @abstract_domain'_R t y y. + Proof using abstract_domain'_R_symmetric abstract_domain'_R_transitive. eapply PER_valid_r; eassumption. Qed. + + Local Hint Immediate abstract_domain'_R_refl_of_rel_l abstract_domain'_R_refl_of_rel_r : core. + + Local Instance abstract_domain_R_Symmetric {t} : Symmetric (@abstract_domain_R t) := _ : Symmetric (type.related _). + Local Instance abstract_domain_R_Transitive {t} : Transitive (@abstract_domain_R t) := _ : Transitive (type.related _). + + Lemma abstract_domain_R_refl_of_rel_l t x y (H : @abstract_domain_R t x y) + : @abstract_domain_R t x x. + Proof using abstract_domain'_R_symmetric abstract_domain'_R_transitive. eapply PER_valid_l; eassumption. Qed. + + Lemma abstract_domain_R_refl_of_rel_r t x y (H : @abstract_domain_R t x y) + : @abstract_domain_R t y y. + Proof using abstract_domain'_R_symmetric abstract_domain'_R_transitive. eapply PER_valid_r; eassumption. Qed. + + Local Hint Immediate abstract_domain_R_refl_of_rel_l abstract_domain_R_refl_of_rel_r : core. + + Lemma related_bottom_for_each_lhs_of_arrow {t} v + : type.and_for_each_lhs_of_arrow (@abstraction_relation) (@bottom_for_each_lhs_of_arrow t) v. + Proof using bottom'_related. induction t; cbn; eauto. Qed. + + Local Hint Immediate related_bottom_for_each_lhs_of_arrow : core. + + Fixpoint fill_in_bottom_for_arrows {t} : abstract_domain t -> abstract_domain t + := match t with + | type.base t => fun x => x + | type.arrow s d + => fun f x => let x := match s with + | type.base _ => x + | type.arrow _ _ => bottom + end in + @fill_in_bottom_for_arrows d (f x) + end. + + Lemma abstract_domain_R_bottom_fill_arrows {t} + : abstract_domain_R (@bottom t) (fill_in_bottom_for_arrows (@bottom t)). + Proof using bottom'_Proper. + cbv [abstract_domain_R]; induction t as [t|s IHs d IHd]; cbn [fill_in_bottom_for_arrows bottom type.related]; + cbv [respectful Proper] in *; auto. + Qed. + + Lemma fill_in_bottom_for_arrows_bottom_related {t} v + : abstraction_relation (fill_in_bottom_for_arrows (@bottom t)) v. + Proof using bottom'_related. + cbv [abstraction_relation]; induction t; cbn; cbv [respectful_hetero]; eauto. + Qed. + + Hint Resolve fill_in_bottom_for_arrows_bottom_related : core. + + Local Instance fill_in_bottom_for_arrows_Proper {t} : Proper (abstract_domain_R ==> abstract_domain_R) (@fill_in_bottom_for_arrows t). + Proof using bottom'_Proper. + pose proof (@bottom_Proper). + cbv [Proper respectful abstract_domain_R] in *; induction t; cbn in *; cbv [respectful] in *; + intros; break_innermost_match; eauto. + Qed. + + Local Instance bottom_eqv_Proper_refl {t} : Proper type.eqv (@bottom t). + Proof using Type. cbv [Proper]; induction t; cbn in *; cbv [respectful] in *; eauto. Qed. + + Lemma bottom_eqv_refl {t} : @bottom t == @bottom t. + Proof using Type. apply bottom_eqv_Proper_refl. Qed. + Local Hint Resolve bottom_eqv_refl : core. + + Local Instance fill_in_bottom_for_arrows_Proper_eqv {t} : Proper (type.eqv ==> type.eqv) (@fill_in_bottom_for_arrows t). + Proof using Type. + cbv [Proper respectful] in *; induction t; cbn in *; cbv [respectful] in *; + intros; break_innermost_match; cbn in *; cbv [respectful] in *; eauto. + Qed. + + Lemma state_of_value_related_fill {t} v (HP : Proper abstract_domain_R (@state_of_value t v)) + : abstract_domain_R (@state_of_value t v) (fill_in_bottom_for_arrows (@state_of_value t v)). + Proof using bottom'_Proper. destruct t; [ assumption | apply abstract_domain_R_bottom_fill_arrows ]. Qed. + + Lemma eqv_bottom_fill_bottom {t} + : @bottom t == fill_in_bottom_for_arrows bottom. + Proof using Type. induction t; cbn; [ reflexivity | ]; cbv [respectful]; auto. Qed. + + Lemma eqv_fill_bottom_idempotent {t} v1 v2 + : v1 == v2 -> fill_in_bottom_for_arrows (fill_in_bottom_for_arrows v1) == @fill_in_bottom_for_arrows t v2. + Proof using Type. induction t; cbn; cbv [respectful]; break_innermost_match; auto. Qed. + + Lemma abstract_domain_R_fill_bottom_idempotent {t} v1 v2 + : abstract_domain_R v1 v2 + -> abstract_domain_R (fill_in_bottom_for_arrows (fill_in_bottom_for_arrows v1)) + (@fill_in_bottom_for_arrows t v2). + Proof using bottom'_Proper. + pose proof (@bottom_Proper) as Hb. + induction t as [|s IHs d IHd]; cbn; cbv [respectful Proper abstract_domain_R] in *; break_innermost_match; auto. + Qed. + + Lemma app_curried_state_of_value_fill {t} v x y + (H : type.and_for_each_lhs_of_arrow (@type.eqv) x y) + : type.app_curried (@state_of_value t v) x = type.app_curried (fill_in_bottom_for_arrows (@state_of_value t v)) y. + Proof using Type. + destruct t; [ reflexivity | cbv [state_of_value] ]. + apply type.app_curried_Proper; [ apply eqv_bottom_fill_bottom | assumption ]. + Qed. + + Lemma first_order_app_curried_fill_in_bottom_for_arrows_eq {t} f xs + (Ht : type.is_not_higher_order t = true) + : type.app_curried (t:=t) f xs = type.app_curried (fill_in_bottom_for_arrows f) xs. + Proof using Type. + clear -Ht. + induction t as [| [|s' d'] IHs d IHd]; cbn in *; try discriminate; auto. + Qed. + + Lemma first_order_abstraction_relation_fill_in_bottom_for_arrows_iff + {t} f v + (Ht : type.is_not_higher_order t = true) + : @abstraction_relation t f v + <-> @abstraction_relation t (fill_in_bottom_for_arrows f) v. + Proof using Type. + clear -Ht; cbv [abstraction_relation]. + split; induction t as [| [|s' d'] IHs d IHd]; + cbn in *; cbv [respectful_hetero]; try discriminate; auto. + Qed. + + Lemma related_state_of_value_of_related_bounded_value {t} st e v + : @related_bounded_value t st e v -> abstract_domain_R match t with + | type.base _ => st + | type.arrow _ _ => bottom + end (state_of_value e). + Proof using bottom'_Proper. intro H; destruct t; cbn in *; [ destruct e; apply H | repeat intro; refine bottom_Proper ]. Qed. + + Lemma related_state_of_value_of_related_bounded_value2 {t} st e v (st' := match t with + | type.base _ => st + | type.arrow _ _ => bottom + end) + : @related_bounded_value t st' e v -> abstract_domain_R st' (state_of_value e). + Proof using bottom'_Proper. intro H; destruct t; cbn in *; [ destruct e; apply H | repeat intro; refine bottom_Proper ]. Qed. + + Lemma related_bounded_value_Proper {t} st1 st2 (Hst : abstract_domain_R (fill_in_bottom_for_arrows st1) (fill_in_bottom_for_arrows st2)) + a a1 a2 + (Ha' : type.eqv a1 a2) + : @related_bounded_value t st1 a a1 -> @related_bounded_value t st2 a a2. + Proof using abstraction_relation'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_Proper. + induction t as [t|s IHs d IHd]; cbn [related_bounded_value type.related] in *; cbv [respectful abstract_domain_R] in *. + all: cbn [type.andb_each_lhs_of_arrow] in *. + all: rewrite ?Bool.andb_true_iff in *. + all: destruct_head'_and. + { intros; break_innermost_match; subst; + destruct_head'_and; repeat apply conj; auto. + { etransitivity; (idtac + symmetry); eassumption. } + { eapply abstraction_relation'_Proper; (eassumption + reflexivity). } } + { intros [? Hrel]. + split; [ repeat intro; etransitivity; (idtac + symmetry); eapply Ha'; (eassumption + (etransitivity; (idtac + symmetry); eassumption)) | ]. + pose proof (@bottom_Proper s) as Hsbot. + intros ?? v_s; destruct s; intros Hx; cbn [type.related] in *; + cbn [fill_in_bottom_for_arrows] in *; cbv [respectful] in *. + { specialize_by_assumption; cbn in *. + eapply IHd; [ cbn in Hst |- *; eapply Hst | apply Ha'; reflexivity | eapply Hrel, Hx ]; cbv [respectful]. + cbn [related_bounded_value] in *. + break_innermost_match_hyps; destruct_head'_and. + eauto. } + { eapply IHd; [ eapply Hst | apply Ha' | eapply Hrel, Hx ]; + [ eexact Hsbot | refine (@related_of_related_bounded_value _ _ _ v_s _); eassumption | refine bottom ]. } } + Qed. + + Lemma related_bounded_value_fill_bottom_iff {t} st1 st2 (Hst : abstract_domain_R st1 st2) + a a1 a2 + (Ha' : type.eqv a1 a2) + : @related_bounded_value t st1 a a1 <-> @related_bounded_value t (fill_in_bottom_for_arrows st2) a a2. + Proof using abstraction_relation'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_Proper. + split; eapply related_bounded_value_Proper; try solve [ (idtac + symmetry); assumption ]. + all: (idtac + symmetry); apply abstract_domain_R_fill_bottom_idempotent. + all: (idtac + symmetry); assumption. + Qed. + + Lemma related_bounded_value_Proper1 {t} + : Proper (abstract_domain_R ==> eq ==> eq ==> Basics.impl) (@related_bounded_value t). + Proof using abstraction_relation'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_Proper. + repeat intro; subst; eapply related_bounded_value_Proper. + { eapply fill_in_bottom_for_arrows_Proper; eassumption. } + { eapply related_of_related_bounded_value; eassumption. } + { assumption. } + Qed. + + Lemma related_bounded_value_Proper_eq {t} + : Proper (eq ==> eq ==> eq ==> Basics.impl) (@related_bounded_value t). + Proof using Type. + repeat intro; subst; assumption. + Qed. + + Lemma related_bounded_value_Proper_interp_eq_base {t} + : Proper (eq ==> RelProd eq (fun x y => expr.interp ident_interp x = expr.interp ident_interp y) ==> eq ==> Basics.impl) (@related_bounded_value (type.base t)). + Proof using Type. + repeat intro; subst. + cbv [value RelProd relation_conjunction predicate_intersection pointwise_extension RelCompFun] in *. + destruct_head'_prod; destruct_head'_and; cbn [fst snd] in *; subst. + cbv [related_bounded_value] in *; destruct_head'_and; repeat apply conj; subst; (idtac + symmetry); assumption. + Qed. + + Fixpoint interp_reify + annotate_with_state is_let_bound {t} st e v b_in + (Hb : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) b_in) + (H : related_bounded_value st e v) {struct t} + : (forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1), + type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify annotate_with_state is_let_bound t e b_in))) arg1 + = type.app_curried v arg2) + /\ (forall arg1 + (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1) + (Harg11 : Proper (type.and_for_each_lhs_of_arrow (@type.eqv)) arg1), + abstraction_relation' + _ + (type.app_curried (fill_in_bottom_for_arrows st) b_in) + (type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify annotate_with_state is_let_bound t e b_in))) arg1)) + with interp_reflect + annotate_with_state {t} st e v + (Hst_Proper : Proper abstract_domain_R st) + (H_val : expr.interp ident_interp e == v) + (Hst1 : abstraction_relation (fill_in_bottom_for_arrows st) (expr.interp ident_interp e)) + {struct t} + : related_bounded_value + st + (@reflect annotate_with_state t e st) + v. + Proof using interp_annotate abstraction_relation'_Proper bottom'_related bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric. + all: destruct t as [t|s d]; + [ clear interp_reify interp_reflect + | pose proof (fun is_let_bound => interp_reify annotate_with_state is_let_bound s) as interp_reify_s; + pose proof (fun is_let_bound => interp_reify annotate_with_state is_let_bound d) as interp_reify_d; + pose proof (interp_reflect annotate_with_state s) as interp_reflect_s; + pose proof (interp_reflect annotate_with_state d) as interp_reflect_d; + clear interp_reify interp_reflect; + pose proof (@abstract_domain_R_bottom_fill_arrows s); + pose proof (@abstract_domain_R_bottom_fill_arrows d) ]. + all: cbn [reify reflect] in *; fold (@reify) (@reflect) in *. + all: cbn [related_bounded_value type.related type.app_curried] in *. + all: cbn [UnderLets.interp expr.interp type.final_codomain type.andb_each_lhs_of_arrow type.is_base fst snd fill_in_bottom_for_arrows type.map_for_each_lhs_of_arrow type.for_each_lhs_of_arrow type.and_for_each_lhs_of_arrow partial.bottom_for_each_lhs_of_arrow partial.bottom] in *. + all: destruct annotate_with_state; try destruct is_let_bound. + all: repeat first [ reflexivity + | progress eta_expand + | progress cbv [type.is_not_higher_order] in * + | progress cbn [UnderLets.interp expr.interp type.final_codomain fst snd] in * + | progress subst + | progress destruct_head'_and + | progress destruct_head'_prod + | progress destruct_head_hnf' and + | progress destruct_head_hnf' prod + | progress destruct_head_hnf' unit + | progress split_and + | progress subst + | discriminate + | rewrite UnderLets.interp_splice + | rewrite UnderLets.interp_to_expr + | rewrite interp_annotate + | match goal with + | [ H : context[andb _ _ = true] |- _ ] => rewrite !Bool.andb_true_iff in H + | [ |- context[andb _ _ = true] ] => rewrite !Bool.andb_true_iff + end + | match goal with + | [ H : fst ?x = _ |- _ ] => is_var x; destruct x + | [ H : Proper _ ?st |- ?R (?st _) (?st _) ] => apply H + | [ |- ?R (state_of_value _) (state_of_value _) ] => cbv [state_of_value] in * + end + | solve [ repeat intro; apply bottom_Proper + | auto; cbv [Proper respectful Basics.impl] in *; eauto ] + | progress (repeat apply conj; intros * ) + | progress intros + | progress destruct_head'_or + | do_with_exactly_one_hyp ltac:(fun H => eapply H; clear H); + try assumption; auto; [] + | match goal with + | [ |- Proper ?R _ ] => (eapply PER_valid_l + eapply PER_valid_r); eassumption + | [ |- @related_bounded_value ?t ?st1 (reflect _ _ ?st2) _ ] + => (tryif first [ constr_eq st1 st2 | has_evar st1 | has_evar st2 ] + then fail + else (eapply (@related_bounded_value_Proper1 t st2 st1); + try reflexivity)) + | [ H : ?R ?x ?y |- ?R ?y ?x ] => symmetry; assumption + end + | break_innermost_match_step + | do_with_exactly_one_hyp ltac:(fun H => eapply H; clear H); + try assumption; auto + | match goal with + | [ |- abstraction_relation (fill_in_bottom_for_arrows (?f (state_of_value ?e))) _ ] + => replace (state_of_value e) with (match s with + | type.base _ => state_of_value e + | type.arrow _ _ => bottom + end) by (destruct s; reflexivity) + end + | progress fold (@reify) (@reflect) (@type.interp) (@type.related) (@type.related_hetero) in * + | match goal with + | [ |- type.related _ (expr.interp _ (UnderLets.interp _ (reify _ _ _ _))) _ ] + => rewrite type.related_iff_app_curried + | [ |- type.related_hetero _ (@state_of_value ?t _) _ ] + => is_var t; destruct t; cbv [state_of_value]; [ cbn | apply bottom_related ] + end ]. + Qed. + + Lemma interp_reify_first_order + annotate_with_state is_let_bound {t} st e v b_in + (Ht : type.is_not_higher_order t = true) + (Hb : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) b_in) + (H : related_bounded_value st e v) + : (forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1), + type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify annotate_with_state is_let_bound t e b_in))) arg1 + = type.app_curried v arg2) + /\ (forall arg1 + (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1) + (Harg11 : Proper (type.and_for_each_lhs_of_arrow (@type.eqv)) arg1), + abstraction_relation' + _ + (type.app_curried st b_in) + (type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify annotate_with_state is_let_bound t e b_in))) arg1)). + Proof using interp_annotate abstraction_relation'_Proper bottom'_related bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric. + rewrite first_order_app_curried_fill_in_bottom_for_arrows_eq by assumption. + apply interp_reify; assumption. + Qed. + + Lemma interp_reflect_first_order + annotate_with_state {t} st e v + (Ht : type.is_not_higher_order t = true) + (Hst_Proper : Proper abstract_domain_R st) + (H_val : expr.interp ident_interp e == v) + (Hst : abstraction_relation st (expr.interp ident_interp e)) + : related_bounded_value + st + (@reflect annotate_with_state t e st) + v. + Proof using interp_annotate abstraction_relation'_Proper bottom'_related bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric. + rewrite first_order_abstraction_relation_fill_in_bottom_for_arrows_iff in Hst by assumption. + apply interp_reflect; assumption. + Qed. + + Lemma related_bounded_value_annotate_base {t} + v_st st v + : @related_bounded_value (type.base t) v_st st v + -> @related_bounded_value (type.base t) v_st (fst st, UnderLets.interp ident_interp (annotate true t (fst st) (snd st))) v. + Proof using interp_annotate abstraction_relation'_Proper. + clear -interp_annotate abstraction_relation'_Proper. + cbv [Proper respectful Basics.impl] in *. + cbn; break_innermost_match; cbn; intros. + destruct_head'_and; subst; repeat apply conj; auto. + rewrite interp_annotate by eauto; reflexivity. + Qed. + + Lemma related_bounded_value_bottomify {t} v_st st v + : @related_bounded_value t v_st st v + -> @related_bounded_value t bottom (UnderLets.interp ident_interp (bottomify st)) v. + Proof using bottom'_Proper bottom'_related. + induction t; cbn in *; + repeat first [ progress subst + | progress cbv [respectful] in * + | progress cbn [UnderLets.interp] in * + | progress destruct_head'_and + | break_innermost_match_step + | progress intros + | apply conj + | reflexivity + | apply bottom'_Proper + | apply bottom'_related + | solve [ eauto ] + | rewrite UnderLets.interp_splice ]. + Qed. + + Context (interp_ident_Proper + : forall annotate_with_state t idc, + related_bounded_value (ident_extract t idc) (UnderLets.interp ident_interp (interp_ident annotate_with_state t idc)) (ident_interp t idc)). + + Lemma interp_interp + annotate_with_state G G' {t} (e_st e1 e2 e3 : expr t) + (HG : forall t v1 v2 v3, List.In (existT _ t (v1, v2, v3)) G + -> related_bounded_value_with_lets v1 v2 v3) + (HG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G' -> v1 == v2) + (Hwf : expr.wf3 G e_st e1 e2) + (Hwf' : expr.wf G' e2 e3) + : related_bounded_value_with_lets + (extract' e_st) + (interp annotate_with_state e1) + (expr.interp (@ident_interp) e2). + Proof using interp_ident_Proper interp_annotate abstraction_relation'_Proper ident_interp_Proper' abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_Proper bottom'_related. + clear -ident_interp_Proper' interp_ident_Proper interp_annotate abstraction_relation'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_Proper bottom'_related HG HG' Hwf Hwf'. + cbv [related_bounded_value_with_lets] in *; + revert dependent annotate_with_state; revert dependent G'; induction Hwf; intros; + cbn [extract' interp expr.interp UnderLets.interp List.In related_bounded_value reify reflect] in *; cbv [Let_In] in *. + all: destruct annotate_with_state eqn:?. + all: repeat first [ progress intros + | progress subst + | progress inversion_sigma + | progress inversion_prod + | progress destruct_head'_and + | progress destruct_head'_or + | progress destruct_head'_sig + | progress destruct_head'_sigT + | progress destruct_head'_prod + | progress eta_expand + | exfalso; assumption + | progress cbn [UnderLets.interp List.In eq_rect fst snd projT1 projT2] in * + | rewrite UnderLets.interp_splice + | rewrite interp_annotate + | solve [ cbv [Proper respectful Basics.impl] in *; unshelve eauto using related_of_related_bounded_value, related_bounded_value_bottomify ] + | progress specialize_by_assumption + | progress cbv [Let_In] in * + | progress cbn [state_of_value extract'] in * + | progress expr.invert_subst + | match goal with + | [ |- abstract_domain ?t ] => exact (@bottom t) + | [ H : expr.wf _ _ _ |- Proper type.eqv _ ] + => apply expr.wf_interp_Proper_gen1 in H; + [ cbv [Proper]; etransitivity; (idtac + symmetry); exact H | auto ] + | [ H : _ |- _ ] + => (tryif first [ constr_eq H HG | constr_eq H HG' ] + then fail + else (apply H; clear H)) + | [ |- related_bounded_value _ (fst _, UnderLets.interp _ (annotate _ _ _ _)) _ ] + => apply related_bounded_value_annotate_base + | [ H : context[match ?v with None => _ | _ => _ end] |- _ ] => destruct v eqn:? + | [ H : context[@related_bounded_value (type.base ?t) ?x _ ?y] + |- @related_bounded_value (type.base ?t) ?x _ ?y ] + => eapply related_bounded_value_Proper_interp_eq_base; [ reflexivity | split; hnf | reflexivity | eapply H ]; + cbn [fst snd expr.interp]; + [ reflexivity | reflexivity | .. ] + end + | apply conj + | match goal with + | [ H : _ = _ |- _ ] => rewrite H + end + | break_innermost_match_step + | progress expr.inversion_wf_one_constr + | match goal with + | [ H : _ |- _ ] + => (tryif first [ constr_eq H HG | constr_eq H HG' ] + then fail + else (eapply H; clear H; + [ lazymatch goal with + | [ |- expr.wf _ _ _ ] + => solve [ eassumption + | match goal with + | [ H : forall v1 v2, expr.wf _ _ _ |- expr.wf _ (?f ?x) _ ] + => apply (H x x) + end ] + | [ |- bool ] (* unused [annotate_with_state] argument to a [Proper ... /\ ...] lemma which is used for its [Proper] part *) + => constructor + | _ => idtac + end .. ]; + match goal with + [ |- ?G ] => assert_fails has_evar G + end)) + end + | reflexivity ]. + Qed. + + Lemma interp_eval_with_bound' + annotate_with_state {t} (e_st e1 e2 : expr t) + (Hwf : expr.wf3 nil e_st e1 e2) + (Hwf' : expr.wf nil e2 e2) + (Ht : type.is_not_higher_order t = true) + (st : type.for_each_lhs_of_arrow abstract_domain t) + (Hst : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) st) + : (forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) st arg1), + type.app_curried (expr.interp ident_interp (eval_with_bound' annotate_with_state e1 st)) arg1 + = type.app_curried (expr.interp ident_interp e2) arg2) + /\ (forall arg1 + (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) st arg1) + (Harg11 : Proper (type.and_for_each_lhs_of_arrow (@type.eqv)) arg1), + abstraction_relation' + _ + (extract_gen e_st st) + (type.app_curried (expr.interp ident_interp (eval_with_bound' annotate_with_state e1 st)) arg1)). + Proof using interp_annotate abstraction_relation'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_related interp_ident_Proper bottom'_Proper ident_interp_Proper'. + cbv [extract_gen extract' eval_with_bound']. + split; intros; rewrite UnderLets.interp_to_expr, UnderLets.interp_splice. + all: eapply interp_reify_first_order; eauto. + all: eapply interp_interp; eauto; wf_t. + Qed. + + Lemma interp_eta_expand_with_bound' + {t} (e1 e2 : expr t) + (Hwf : expr.wf nil e1 e2) + (b_in : type.for_each_lhs_of_arrow abstract_domain t) + (Hb_in : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) b_in) + : forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1), + type.app_curried (expr.interp ident_interp (eta_expand_with_bound' e1 b_in)) arg1 = type.app_curried (expr.interp ident_interp e2) arg2. + Proof using interp_annotate ident_interp_Proper' bottom'_related abstraction_relation'_Proper bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric. + cbv [eta_expand_with_bound']. + intros; rewrite UnderLets.interp_to_expr. + eapply interp_reify; eauto. + eapply interp_reflect; eauto using bottom_related. + eapply @expr.wf_interp_Proper_gen; auto using Hwf. + Qed. + + Lemma interp_extract'_from_wf_gen G + (HG : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> abstract_domain_R v1 v2) + {t} (e1 e2 : expr t) + (Hwf : expr.wf G e1 e2) + : abstract_domain_R (extract' e1) (extract' e2). + Proof using ident_extract_Proper bottom'_Proper. + cbv [abstract_domain_R] in *; induction Hwf; cbn [extract']; break_innermost_match. + all: repeat first [ progress subst + | progress inversion_sigma + | progress inversion_prod + | solve [ cbv [Proper respectful] in *; eauto ] + | progress cbv [respectful Let_In] in * + | progress cbn [type.related List.In eq_rect partial.bottom] in * + | progress intros + | progress destruct_head'_or + | apply bottom_Proper + | match goal with H : _ |- type.related _ _ _ => apply H; clear H end ]. + Qed. + + Lemma interp_extract'_from_wf {t} (e1 e2 : expr t) + (Hwf : expr.wf nil e1 e2) + : abstract_domain_R (extract' e1) (extract' e2). + Proof using ident_extract_Proper bottom'_Proper. + eapply interp_extract'_from_wf_gen; revgoals; wf_t. + Qed. + End with_type. + + Module ident. + Import API. + Local Notation UnderLets := (@UnderLets base.type ident). + Section with_type. + Context (abstract_domain' : base.type -> Type). + Local Notation abstract_domain := (@abstract_domain base.type abstract_domain'). + Local Notation var := interp_type. + Context (annotate_expr : forall t, abstract_domain' t -> option (@expr var (t -> t))) + (bottom' : forall A, abstract_domain' A) + (abstract_interp_ident : forall t, ident t -> type.interp abstract_domain' t) + (update_literal_with_state : forall A : base.type.base, abstract_domain' A -> base.interp A -> base.interp A) + (extract_list_state : forall A, abstract_domain' (base.type.list A) -> option (list (abstract_domain' A))) + (extract_option_state : forall A, abstract_domain' (base.type.option A) -> option (option (abstract_domain' A))) + (is_annotated_for : forall t t', @expr var t -> abstract_domain' t' -> bool) + (skip_annotations_under : forall t, ident t -> bool) + (strip_annotation : forall t, ident t -> option (@value base.type ident var abstract_domain' t)) + (abstraction_relation' : forall t, abstract_domain' t -> base.interp t -> Prop) + (abstract_domain'_R : forall t, abstract_domain' t -> abstract_domain' t -> Prop) + (abstraction_relation'_Proper : forall t, Proper (abstract_domain'_R t ==> eq ==> Basics.impl) (abstraction_relation' t)) + (bottom'_related : forall t v, abstraction_relation' t (bottom' t) v) + {bottom'_Proper : forall t, Proper (abstract_domain'_R t) (bottom' t)} + {abstract_domain'_R_transitive : forall t, Transitive (@abstract_domain'_R t)} + {abstract_domain'_R_symmetric : forall t, Symmetric (@abstract_domain'_R t)}. + Local Notation abstraction_relation := (@abstraction_relation base.type abstract_domain' base.interp abstraction_relation'). + Local Notation ident_interp := ident.interp (only parsing). + Local Notation abstract_domain_R := (@abstract_domain_R base.type abstract_domain' abstract_domain'_R). + Local Notation fill_in_bottom_for_arrows := (@fill_in_bottom_for_arrows base.type abstract_domain' bottom'). + Local Notation related_bounded_value := (@related_bounded_value base.type ident abstract_domain' base.interp (@ident_interp) abstraction_relation' bottom' abstract_domain'_R). + Context {abstract_interp_ident_Proper : forall t, Proper (eq ==> @abstract_domain_R t) (abstract_interp_ident t)} + (interp_annotate_expr + : forall t st idc, + annotate_expr t st = Some idc + -> forall v, abstraction_relation' _ st v + -> expr.interp (@ident_interp) idc v = v) + (abstract_interp_ident_Proper' + : forall t idc, type.related_hetero (@abstraction_relation') (abstract_interp_ident t idc) (ident_interp idc)) + (extract_list_state_related + : forall t st ls v st' v', + extract_list_state t st = Some ls + -> abstraction_relation' _ st v + -> List.In (st', v') (List.combine ls v) + -> abstraction_relation' t st' v') + (extract_list_state_length_good + : forall t st ls v, + extract_list_state t st = Some ls + -> abstraction_relation' _ st v + -> length ls = length v) + (extract_option_state_related + : forall t st a v, + extract_option_state t st = Some a + -> abstraction_relation' _ st v + -> option_eq (abstraction_relation' t) a v) + (strip_annotation_related + : forall t idc v, + strip_annotation t idc = Some v + -> related_bounded_value (abstract_interp_ident t idc) v (ident_interp idc)). + + Local Notation update_annotation := (@ident.update_annotation _ abstract_domain' annotate_expr is_annotated_for). + Local Notation annotate_with_expr := (@ident.annotate_with_expr _ abstract_domain' annotate_expr is_annotated_for). + Local Notation annotate_base := (@ident.annotate_base _ abstract_domain' annotate_expr is_annotated_for). + Local Notation annotate := (@ident.annotate _ abstract_domain' annotate_expr abstract_interp_ident extract_list_state extract_option_state is_annotated_for). + Local Notation reify := (@reify base.type ident _ abstract_domain' annotate bottom'). + Local Notation reflect := (@reflect base.type ident _ abstract_domain' annotate bottom'). + + Lemma abstract_interp_ident_Proper'' t idc + : type.related_hetero (@abstraction_relation') (fill_in_bottom_for_arrows (abstract_interp_ident t idc)) (ident_interp idc). + Proof using abstract_interp_ident_Proper' bottom'_related. + generalize (abstract_interp_ident_Proper' t idc); clear -bottom'_related. + generalize (ident_interp idc), (abstract_interp_ident t idc); clear idc. + intros v st H. + induction t as [| [|s' d'] IHs d IHd]; cbn in *; cbv [respectful_hetero] in *; + auto. + intros; apply IHd, H; clear IHd H. + intros; apply bottom_related; assumption. + Qed. + + Lemma interp_update_annotation t st e + (He : abstraction_relation' t st (expr.interp (t:=type.base t) (@ident_interp) e)) + : expr.interp (@ident_interp) (@update_annotation t st e) + = expr.interp (@ident_interp) e. + Proof using interp_annotate_expr. + cbv [update_annotation]; + repeat first [ reflexivity + | progress subst + | progress eliminate_hprop_eq + | progress cbn [expr.interp eq_rect] in * + | match goal with + | [ H : annotate_expr _ _ = Some _ |- _ ] => rewrite (interp_annotate_expr _ _ _ H) by eassumption + end + | progress expr.invert_match + | progress type_beq_to_eq + | progress rewrite_type_transport_correct + | progress break_innermost_match_step ]. + Qed. + + Lemma interp_annotate_with_expr is_let_bound t st e + (He : abstraction_relation' t st (expr.interp (t:=type.base t) (@ident_interp) e)) + : expr.interp (@ident_interp) (UnderLets.interp (@ident_interp) (@annotate_with_expr is_let_bound t st e)) + = expr.interp (@ident_interp) e. + Proof using interp_annotate_expr. + cbv [annotate_with_expr]; break_innermost_match; cbn [expr.interp UnderLets.interp]; + apply interp_update_annotation; assumption. + Qed. + + Lemma interp_annotate_base is_let_bound (t : base.type.base) st e + (He : abstraction_relation' t st (expr.interp (t:=type.base (base.type.type_base t)) (@ident_interp) e)) + : expr.interp (@ident_interp) (UnderLets.interp (@ident_interp) (@annotate_base is_let_bound t st e)) + = expr.interp (@ident_interp) e. + Proof using interp_annotate_expr. + cbv [annotate_base]; break_innermost_match; expr.invert_subst; cbv beta iota in *; subst. + { apply interp_annotate_with_expr; assumption. } + Qed. + + Lemma interp_annotate is_let_bound (t : base.type) st e + (He : abstraction_relation' t st (expr.interp (t:=type.base t) (@ident_interp) e)) + : expr.interp (@ident_interp) (UnderLets.interp (@ident_interp) (@annotate is_let_bound t st e)) + = expr.interp (@ident_interp) e. + Proof using interp_annotate_expr abstract_interp_ident_Proper' extract_list_state_related extract_list_state_length_good extract_option_state_related bottom'_related. + induction t; cbn [annotate]; auto using interp_annotate_base. + all: repeat first [ reflexivity + | progress subst + | progress inversion_option + | progress inversion_prod + | progress destruct_head'_ex + | progress destruct_head'_and + | progress break_innermost_match + | progress break_innermost_match_hyps + | progress expr.invert_subst + | progress cbn [fst snd UnderLets.interp expr.interp ident_interp Nat.add] in * + | progress cbv [ident.literal] in * + | match goal with + | [ H : context[ident_interp (ident.ident_Literal _)] |- _ ] => rewrite ident.interp_ident_Literal in H + | [ H : context[ident_interp ident.ident_Some] |- _ ] => rewrite ident.interp_ident_Some in H + | [ H : context[ident_interp ident.ident_pair] |- _ ] => rewrite ident.interp_ident_pair in H + end + | rewrite !UnderLets.interp_splice + | rewrite !UnderLets.interp_splice_list + | rewrite !List.map_map + | rewrite expr.interp_reify_list + | rewrite nth_error_combine + | apply interp_annotate_with_expr; assumption + | progress fold (@base.interp) in * + | progress intros + | pose proof (@extract_list_state_length_good _ _ _ _ ltac:(eassumption) ltac:(eassumption)); clear extract_list_state_length_good + | match goal with + | [ H : context[expr.interp _ (reify_list _)] |- _ ] => rewrite expr.interp_reify_list in H + | [ H : abstraction_relation' (_ * _) _ (_, _) |- _ ] + => pose proof (abstract_interp_ident_Proper'' _ ident.fst _ _ H); + pose proof (abstract_interp_ident_Proper'' _ ident.snd _ _ H); + clear H + | [ H : context[_ = _] |- _ = _ ] => rewrite H by assumption + | [ |- List.map ?f (List.combine ?l1 ?l2) = List.map ?g ?l2 ] + => transitivity (List.map g (List.map (@snd _ _) (List.combine l1 l2))); + [ rewrite List.map_map; apply List.map_ext_in + | rewrite map_snd_combine, List.firstn_all2; [ reflexivity | ] ] + | [ Hls : extract_list_state ?t ?st = Some ?ls, He : abstraction_relation' _ ?st ?v |- abstraction_relation' _ _ _ ] + => apply (fun st' v' => extract_list_state_related t st ls v st' v' Hls He) + | [ H : context[List.nth_error (List.combine _ _) _] |- _ ] => rewrite nth_error_combine in H + | [ H : List.In _ (List.combine _ _) |- _ ] => apply List.In_nth_error in H + | [ |- List.In _ (List.combine _ _) ] => eapply nth_error_In + | [ H : ?x = Some _, H' : context[?x] |- _ ] => rewrite H in H' + | [ H : List.nth_error (List.map _ _) _ = Some _ |- _ ] => apply nth_error_map_ex in H + | [ H : List.nth_error _ _ = None |- _ ] => rewrite List.nth_error_None in H + | [ H : context[length ?ls] |- _ ] => tryif is_var ls then fail else (progress autorewrite with distr_length in H) + | [ |- context[length ?ls] ] => tryif is_var ls then fail else (progress autorewrite with distr_length) + | [ H : List.nth_error ?ls ?n = Some _, H' : length ?ls <= ?n |- _ ] + => apply nth_error_value_length in H; exfalso; clear -H H'; lia + | [ H : List.nth_error ?l ?n = _, H' : List.nth_error ?l ?n' = _ |- _ ] + => unify n n'; rewrite H in H' + | [ Hls : extract_list_state ?t ?st = Some ?ls, He : abstraction_relation' _ ?st ?v |- _ ] + => pose proof (fun st' v' => extract_list_state_related t st ls v st' v' Hls He); clear extract_list_state_related + | [ IH : forall st e, _ -> expr.interp _ (UnderLets.interp _ (annotate _ _ _)) = _ |- List.map (fun x => expr.interp _ _) (List.combine _ _) = expr.interp _ _ ] + => erewrite List.map_ext_in; + [ + | intros; eta_expand; rewrite IH; cbn [expr.interp ident_interp ident.smart_Literal]; [ reflexivity | ] ] + | [ H : abstraction_relation' _ ?st (List.map (expr.interp _) ?ls), H' : forall st' v', List.In _ (List.combine _ _) -> abstraction_relation' _ _ _, H'' : List.nth_error ?ls _ = Some ?e |- abstraction_relation' _ _ (expr.interp _ ?e) ] + => apply H' + | [ H : context[List.nth_error (List.seq _ _) _] |- _ ] => rewrite nth_error_seq in H + end + | apply Nat.eq_le_incl + | rewrite <- List.map_map with (f:=fst), map_fst_combine + | match goal with + | [ |- context[List.map (fun a : ?A * ?B => @?body a) (List.combine _ _)] ] + => let aa := fresh in + let bb := fresh in + let g' := match (eval cbn [fst] in (fun (bb : B) (aa : A) => body (aa, bb))) with + | fun _ => ?g => g + end in + erewrite <- List.map_map with (f:=fst) (g:=g'), map_fst_combine + end + | rewrite Lists.List.firstn_all2 by distr_length + | apply map_nth_default_seq + | progress destruct_head' option + | progress cbn [Option.combine option_map reify_option option_rect UnderLets.splice_option] in * + | apply (f_equal Some) + | match goal with + | [ H : abstraction_relation' _ ?st _, H' : extract_option_state _ ?st = _ |- _ ] + => eapply extract_option_state_related in H; [ clear H' | eexact H' ]; + cbv [option_eq] in H + | [ H : context[expr.interp _ _ = expr.interp _ _] |- expr.interp _ _ = expr.interp _ _ ] => apply H; clear H + | [ H : forall st' v', List.In _ (List.combine _ _) -> abstraction_relation' _ _ _ |- abstraction_relation' _ _ _ ] + => apply H; clear H; cbv [List.nth_default] + | [ |- None = Some _ ] => exfalso; lia + end ]. + Qed. + + Local Notation interp_ident := (@ident.interp_ident _ abstract_domain' annotate_expr bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for strip_annotation). + + Lemma interp_ident_Proper_not_nth_default_nostrip annotate_with_state t idc + : related_bounded_value (abstract_interp_ident t idc) (UnderLets.interp (@ident_interp) (Base (reflect annotate_with_state (expr.Ident idc) (abstract_interp_ident _ idc)))) (ident_interp idc). + Proof using abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr abstract_interp_ident_Proper bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric. + cbn [UnderLets.interp]. + eapply interp_reflect; + try first [ apply ident.interp_Proper + | apply abstract_interp_ident_Proper'' + | eapply abstract_interp_ident_Proper; reflexivity + | apply interp_annotate ]; + eauto. + Qed. + + Lemma interp_ident_Proper_not_nth_default annotate_with_state t idc + : related_bounded_value + (abstract_interp_ident t idc) + ((UnderLets.interp (@ident_interp)) + (Base match strip_annotation _ idc with + | Some v => v + | None => reflect annotate_with_state (expr.Ident idc) (abstract_interp_ident _ idc) + end)) + (ident_interp idc). + Proof using abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr abstract_interp_ident_Proper bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric strip_annotation_related. + pose proof (strip_annotation_related t idc) as H. + break_innermost_match; eauto using eq_refl, interp_ident_Proper_not_nth_default_nostrip. + Qed. + + Lemma interp_ident_Proper_nth_default annotate_with_state T (idc:=@ident.List_nth_default T) + : related_bounded_value (abstract_interp_ident _ idc) (UnderLets.interp (@ident_interp) (interp_ident annotate_with_state idc)) (ident_interp idc). + Proof using abstract_interp_ident_Proper abstract_interp_ident_Proper' abstraction_relation'_Proper extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr bottom'_related. + subst idc; cbn [interp_ident reify reflect fst snd UnderLets.interp ident_interp related_bounded_value abstract_domain value]. + cbv [abstract_domain]; cbn [type.interp bottom_for_each_lhs_of_arrow state_of_value fst snd]. + repeat first [ progress intros + | progress cbn [UnderLets.interp fst snd expr.interp ident_interp] in * + | progress destruct_head'_prod + | progress destruct_head'_and + | progress subst + | progress eta_expand + | rewrite UnderLets.interp_splice + | progress expr.invert_subst + | break_innermost_match_step + | progress cbn [type.interp base.interp base.base_interp] in * + | rewrite interp_annotate + | solve [ cbv [Proper respectful Basics.impl] in *; eauto ] + | split; [ apply (@abstract_interp_ident_Proper _ (@ident.List_nth_default T) _ eq_refl) | ] + | split; [ reflexivity | ] + | apply (@abstract_interp_ident_Proper'' _ (@ident.List_nth_default T)) + | apply conj + | rewrite map_nth_default_always + | rewrite expr.interp_reify_list + | match goal with + | [ H : context[expr.interp _ (UnderLets.interp _ (annotate _ _ _))] |- _ ] + => rewrite interp_annotate in H + | [ H : context[expr.interp _ (reify_list _)] |- _ ] + => rewrite expr.interp_reify_list in H + | [ H : _ = reify_list _ |- _ ] => apply (f_equal (expr.interp (@ident_interp))) in H + | [ H : expr.interp _ ?x = _ |- context[expr.interp _ ?x] ] => rewrite H + | [ |- Proper _ _ ] => cbv [Proper type.related respectful] + end ]. + Qed. + + Lemma interp_ident_Proper annotate_with_state t idc + : related_bounded_value (abstract_interp_ident t idc) (UnderLets.interp (@ident_interp) (interp_ident annotate_with_state idc)) (ident_interp idc). + Proof using abstract_domain'_R_symmetric abstract_domain'_R_transitive abstract_interp_ident_Proper abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr strip_annotation_related. + pose idc as idc'. + destruct idc; first [ refine (@interp_ident_Proper_not_nth_default _ _ idc') + | refine (@interp_ident_Proper_nth_default _ _) ]. + Qed. + + Local Notation eval_with_bound := (@partial.ident.eval_with_bound _ abstract_domain' annotate_expr bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for strip_annotation). + Local Notation eta_expand_with_bound := (@partial.ident.eta_expand_with_bound _ abstract_domain' annotate_expr bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for). + Local Notation extract := (@ident.extract abstract_domain' bottom' abstract_interp_ident). + + Lemma interp_eval_with_bound + annotate_with_state {t} (e_st e1 e2 : expr t) + (Hwf : expr.wf3 nil e_st e1 e2) + (Hwf' : expr.wf nil e2 e2) + (Ht : type.is_not_higher_order t = true) + (st : type.for_each_lhs_of_arrow abstract_domain t) + (Hst : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) st) + : (forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) st arg1), + type.app_curried (expr.interp (@ident_interp) (eval_with_bound skip_annotations_under annotate_with_state e1 st)) arg1 + = type.app_curried (expr.interp (@ident_interp) e2) arg2) + /\ (forall arg1 + (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) st arg1) + (Harg11 : Proper (type.and_for_each_lhs_of_arrow (@type.eqv)) arg1), + abstraction_relation' + _ + (extract e_st st) + (type.app_curried (expr.interp (@ident_interp) (eval_with_bound skip_annotations_under annotate_with_state e1 st)) arg1)). + Proof using abstract_domain'_R_symmetric abstract_domain'_R_transitive abstract_interp_ident_Proper abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr strip_annotation_related. cbv [extract eval_with_bound]; apply @interp_eval_with_bound' with (abstract_domain'_R:=abstract_domain'_R); auto using interp_annotate, interp_ident_Proper, ident.interp_Proper. Qed. + + Lemma interp_eta_expand_with_bound + {t} (e1 e2 : expr t) + (Hwf : expr.wf nil e1 e2) + (Ht : type.is_not_higher_order t = true) + (b_in : type.for_each_lhs_of_arrow abstract_domain t) + (Hb_in : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) b_in) + : forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1), + type.app_curried (expr.interp (@ident_interp) (eta_expand_with_bound e1 b_in)) arg1 = type.app_curried (expr.interp (@ident_interp) e2) arg2. + Proof using abstract_domain'_R_symmetric abstract_domain'_R_transitive abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr. cbv [partial.ident.eta_expand_with_bound]; eapply interp_eta_expand_with_bound'; eauto using interp_annotate, ident.interp_Proper. Qed. + End with_type. + End ident. + + Lemma default_relax_zrange_good + : forall r r' z, is_tighter_than_bool z r = true + -> default_relax_zrange r = Some r' + -> is_tighter_than_bool z r' = true. + Proof. + cbv [default_relax_zrange]; intros; inversion_option; subst; assumption. + Qed. + + Section specialized. + Import API. + Local Notation abstract_domain' := ZRange.type.base.option.interp (only parsing). + Local Notation abstract_domain := (@partial.abstract_domain base.type abstract_domain'). + Local Notation abstract_domain'_R t := (@eq (abstract_domain' t)) (only parsing). + Local Notation abstract_domain_R := (@abstract_domain_R base.type abstract_domain' (fun t => abstract_domain'_R t)). + Local Notation fill_in_bottom_for_arrows := (@fill_in_bottom_for_arrows base.type abstract_domain' bottom'). + + Definition abstraction_relation' {t} : abstract_domain' t -> base.interp t -> Prop + := fun st v => @ZRange.type.base.option.is_bounded_by t st v = true. + + Lemma bottom'_bottom {t} : forall v, abstraction_relation' (bottom' t) v. + Proof using Type. + cbv [abstraction_relation' bottom']; induction t; cbn; intros; break_innermost_match; cbn; try reflexivity. + rewrite Bool.andb_true_iff; split; auto. + Qed. + + Lemma abstract_interp_ident_related {opts : AbstractInterpretation.Options} {assume_cast_truncates : bool} {t} (idc : ident t) + : type.related_hetero (@abstraction_relation') (abstract_interp_ident assume_cast_truncates t idc) (ident.interp idc). + Proof using Type. apply ZRange.ident.option.interp_related. Qed. + + Local Ltac zrange_interp_idempotent_t := + repeat first [ progress destruct_head'_ex + | progress subst + | progress cbn in * + | progress cbv [ZRange.goodb ZRange.ident.option.interp_Z_cast_truncate ZRange.ident.option.interp_Z_cast] in * + | reflexivity + | lia + | progress Z.ltb_to_lt + | progress destruct_head'_and + | break_innermost_match_step + | break_innermost_match_hyps_step + | rewrite (proj1 ZRange.normalize_id_iff_goodb) + | progress Bool.split_andb + | match goal with + | [ H : ?x = ?x |- _ ] => clear H + | [ H : (?x <= ?y)%Z, H' : (?y <= ?x)%Z |- _ ] + => assert (x = y) by lia; clear H H' + end + | progress unshelve Reflect.reflect_hyps; [] + | progress destruct_head'_or + | progress destruct_head' zrange + | rewrite Bool.andb_true_iff + | match goal with + | [ |- and _ _ ] => split + end + | progress cbv [is_bounded_by_bool] ]. + + Lemma zrange_interp_tighter_bounded_Z_cast {opts : AbstractInterpretation.Options} {assume_cast_truncates r1 r2 v} + (H1 : ZRange.type.base.option.is_bounded_by (t:=base.type.Z) r1 v = true) + (H2 : ZRange.type.base.option.is_bounded_by (t:=base.type.Z) r2 v = true) + : ZRange.type.base.option.is_bounded_by (t:=base.type.Z) (ZRange.ident.option.interp assume_cast_truncates ident.Z_cast r1 r2) v = true. + Proof using Type. destruct r1, r2, assume_cast_truncates; zrange_interp_idempotent_t. Qed. + + Lemma zrange_interp_tighter_bounded_Z_cast2 {opts : AbstractInterpretation.Options} {assume_cast_truncates r1 r2 v} + (H1 : ZRange.type.base.option.is_bounded_by (t:=base.type.Z*base.type.Z) r1 v = true) + (H2 : ZRange.type.base.option.is_bounded_by (t:=base.type.Z*base.type.Z) r2 v = true) + : ZRange.type.base.option.is_bounded_by (t:=base.type.Z*base.type.Z) (ZRange.ident.option.interp assume_cast_truncates ident.Z_cast2 r1 r2) v = true. + Proof using Type. destruct r1, r2, assume_cast_truncates; zrange_interp_idempotent_t. Qed. + + Local Notation related_bounded_value := (@related_bounded_value base.type ident abstract_domain' base.interp (@ident.interp) (@abstraction_relation') bottom' (fun t => abstract_domain'_R t)). + Lemma always_strip_annotation_related {opts : AbstractInterpretation.Options} {assume_cast_truncates : bool} {t} (idc : ident t) + v (Hv : always_strip_annotation assume_cast_truncates t idc = Some v) + : related_bounded_value (abstract_interp_ident assume_cast_truncates t idc) v (ident.interp idc). + Proof using Type. + pose proof (@abstract_interp_ident_related _ assume_cast_truncates _ ident.Z_cast). + pose proof (fun x1 x2 y1 y2 H x01 x02 y01 y02 H' => @abstract_interp_ident_related _ assume_cast_truncates _ ident.Z_cast2 (x1, x2) (y1, y2) H (x01, x02) (y01, y02) H'). + cbv [always_strip_annotation] in *; break_innermost_match_hyps; inversion_option; subst. + all: repeat first [ reflexivity + | progress cbn [related_bounded_value UnderLets.interp fst snd type.related_hetero] in * + | progress cbv [respectful_hetero] in * + | progress intros + | progress destruct_head'_and + | progress subst + | progress inversion_prod + | progress inversion_option + | match goal with + | [ |- Proper _ _ ] => exact _ + | [ |- _ /\ _ ] => split + | [ H1 : (?x <=? ?y)%zrange = true, H2 : is_bounded_by_bool _ ?x = true |- _ ] + => unique pose proof (ZRange.is_bounded_by_of_is_tighter_than _ _ H1 _ H2) + | [ |- andb _ _ = true ] => rewrite Bool.andb_true_iff + | [ H : andb _ _ = true |- _ ] => rewrite Bool.andb_true_iff in H + | [ H : ZRange.type.base.option.is_bounded_by (Some _) _ = _ |- _ ] + => progress cbn -[zrange_beq prod_beq] in H + | [ H : ZRange.type.base.option.is_bounded_by (Some _, Some _) _ = _ |- _ ] + => progress cbn -[zrange_beq prod_beq] in H + | [ H : ZRange.type.base.is_tighter_than _ _ = _ |- _ ] + => progress cbn -[zrange_beq prod_beq] in H + | [ H : bind _ _ = Some _ |- _ ] => progress cbv [bind] in H + end + | progress break_innermost_match + | progress break_innermost_match_hyps + | progress Reflect.reflect_beq_to_eq zrange_beq + | progress cbv [abstraction_relation' ZRange.type.base.option.lift_Some] in * + | cbn [Compilers.ident_interp]; progress cbv [ident.cast2] + | rewrite ident.cast_in_bounds by assumption + | now destruct assume_cast_truncates + | rewrite zrange_interp_idempotent_Z_cast by (cbn; eexists; eassumption) + | rewrite zrange_interp_idempotent_Z_cast2 by (cbn; eexists (_, _); rewrite Bool.andb_true_iff; split; eassumption) + | progress cbv [abstract_interp_ident] in * + | match goal with + | [ H : _ |- _ ] => apply H; cbn; cbn [ZRange.type.base.option.is_bounded_by] in *; Prod.eta_expand; rewrite ?Bool.andb_true_iff in *; destruct_head'_and; try apply conj; try apply zrange_lb; now auto + end + | progress cbn [ZRange.type.base.option.is_bounded_by] + | match goal with + | [ H : ZRange.ident.option.interp ?act ident.Z_cast ?r1 ?r2 = Some ?r3 |- ZRange.type.base.is_bounded_by ?r3 ?v = true ] + => let H' := fresh in + pose proof (@zrange_interp_tighter_bounded_Z_cast _ act r1 r2 v) as H'; + rewrite H in H'; now apply H' + | [ H : context G[ZRange.ident.option.interp ?act ident.Z_cast2 ?r1 ?r2] |- _ ] + => let G' := context G[(ZRange.ident.option.interp act ident.Z_cast (fst r1) (fst r2), + ZRange.ident.option.interp act ident.Z_cast (snd r1) (snd r2))] in + change G' in H + end ]. + Qed. + + Lemma strip_annotation_related {opts : AbstractInterpretation.Options} {assume_cast_truncates strip_annotations : bool} {t} (idc : ident t) + v (Hv : strip_annotation assume_cast_truncates strip_annotations t idc = Some v) + : related_bounded_value (abstract_interp_ident assume_cast_truncates t idc) v (ident.interp idc). + Proof using Type. + destruct strip_annotations; cbv [strip_annotation] in *; try congruence; + now apply always_strip_annotation_related. + Qed. + + Lemma extract_list_state_related {t} st v ls + : @abstraction_relation' _ st v + -> @extract_list_state t st = Some ls + -> length ls = length v + /\ forall st' (v' : base.interp t), List.In (st', v') (List.combine ls v) -> @abstraction_relation' t st' v'. + Proof using Type. + cbv [abstraction_relation' extract_list_state]; cbn [ZRange.type.base.option.is_bounded_by]. + intros; subst. + split. + { eapply FoldBool.fold_andb_map_length; eassumption. } + { intros *. + revert dependent v; induction ls, v; cbn; try tauto. + rewrite Bool.andb_true_iff. + intros; destruct_head'_and; destruct_head'_or; inversion_prod; subst; eauto. } + Qed. + + Lemma extract_option_state_related {t} st a v + : extract_option_state t st = Some a + -> @abstraction_relation' _ st v + -> option_eq (@abstraction_relation' t) a v. + Proof using Type. + cbv [abstraction_relation' extract_option_state option_eq]; intros; subst; cbn in *; cbv [option_beq_hetero] in *; break_match; break_match_hyps; auto; congruence. + Qed. + + Lemma Extract_FromFlat_ToFlat' {opts : AbstractInterpretation.Options} {assume_cast_truncates : bool} {t} (e : Expr t) (Hwf : Wf e) b_in1 b_in2 + (Hb : type.and_for_each_lhs_of_arrow (fun t => type.eqv) b_in1 b_in2) + : partial.Extract assume_cast_truncates (GeneralizeVar.FromFlat (GeneralizeVar.ToFlat e)) b_in1 + = partial.Extract assume_cast_truncates e b_in2. + Proof using Type. + cbv [partial.Extract partial.ident.extract partial.extract_gen]. + revert b_in1 b_in2 Hb. + rewrite <- (@type.related_iff_app_curried base.type ZRange.type.base.option.interp (fun _ => eq)). + apply interp_extract'_from_wf; auto with wf typeclass_instances. + apply GeneralizeVar.wf_from_flat_to_flat, Hwf. + Qed. + + Lemma Extract_FromFlat_ToFlat {opts : AbstractInterpretation.Options} {assume_cast_truncates : bool} {t} (e : Expr t) (Hwf : Wf e) b_in + (Hb : Proper (type.and_for_each_lhs_of_arrow (fun t => type.eqv)) b_in) + : partial.Extract assume_cast_truncates (GeneralizeVar.FromFlat (GeneralizeVar.ToFlat e)) b_in + = partial.Extract assume_cast_truncates e b_in. + Proof using Type. apply Extract_FromFlat_ToFlat'; assumption. Qed. + + Lemma Extract_GeneralizeVar {opts : AbstractInterpretation.Options} {assume_cast_truncates : bool} {t} (e : Expr t) (Hwf : Wf e) b_in + (Hb : Proper (type.and_for_each_lhs_of_arrow (fun t => type.eqv)) b_in) + : partial.Extract assume_cast_truncates (GeneralizeVar.GeneralizeVar (e _)) b_in + = partial.Extract assume_cast_truncates e b_in. + Proof using Type. apply Extract_FromFlat_ToFlat; assumption. Qed. + + Section with_relax. + Context {relax_zrange : zrange -> option zrange} + (Hrelax : forall r r' z, is_tighter_than_bool z r = true + -> relax_zrange r = Some r' + -> is_tighter_than_bool z r' = true). + + Local Lemma Hrelax' r r' z + : is_bounded_by_bool z r = true + -> relax_zrange (ZRange.normalize r) = Some r' + -> is_bounded_by_bool z r' = true. + Proof using Hrelax. + intros H Hr. + eapply ZRange.is_bounded_by_of_is_tighter_than; [ eapply Hrelax; [ | eassumption ] | eassumption ]. + eapply ZRange.is_tighter_than_bool_normalize_of_goodb, ZRange.goodb_of_is_bounded_by_bool; eassumption. + Qed. + + Lemma interp_annotate_expr {t} st idc + (Hst : @annotate_expr relax_zrange _ t st = Some idc) + : forall v, abstraction_relation' st v + -> expr.interp (@ident.interp) idc v = v. + Proof using Hrelax. + repeat first [ progress cbv [annotate_expr Crypto.Util.Option.bind annotation_of_state option_map abstraction_relation' ZRange.type.base.option.is_bounded_by ZRange.type.base.is_bounded_by] in * + | reflexivity + | progress inversion_option + | progress subst + | break_innermost_match_hyps_step + | break_innermost_match_step + | progress cbn [ident.interp base.interp base.base_interp expr.interp] in * + | progress cbv [ident.cast2] in * + | progress intros + | progress Bool.split_andb + | rewrite ident.cast_in_bounds by assumption + | match goal with + | [ H : is_bounded_by_bool ?v ?r = true, H' : relax_zrange (ZRange.normalize ?r) = Some ?r' |- _ ] + => unique assert (is_bounded_by_bool v r' = true) by (eapply Hrelax'; eassumption) + end ]. + Qed. + + Lemma interp_annotate_expr_Proper {t} st1 st2 (Hst : abstract_domain'_R t st1 st2) + : @annotate_expr relax_zrange interp_type t st1 = @annotate_expr relax_zrange interp_type t st2. + Proof using Type. congruence. Qed. + + Local Hint Resolve interp_annotate_expr abstract_interp_ident_related : core. + + Lemma interp_eval_with_bound + {opts : AbstractInterpretation.Options} + {assume_cast_truncates : bool} + {skip_annotations_under : forall t, ident t -> bool} + {strip_preexisting_annotations : bool} + {t} (e_st e1 e2 : expr t) + (Hwf : expr.wf3 nil e_st e1 e2) + (Hwf' : expr.wf nil e2 e2) + (Ht : type.is_not_higher_order t = true) + (st : type.for_each_lhs_of_arrow abstract_domain t) + : (forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) st arg1 = true), + type.app_curried (expr.interp (@ident.interp) (eval_with_bound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e1 st)) arg1 + = type.app_curried (expr.interp (@ident.interp) e2) arg2) + /\ (forall arg1 + (Harg11 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg1) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) st arg1 = true), + abstraction_relation' + (extract assume_cast_truncates e_st st) + (type.app_curried (expr.interp (@ident.interp) (eval_with_bound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e1 st)) arg1)). + Proof using Hrelax. + cbv [eval_with_bound]; split; + [ intros arg1 arg2 Harg12 Harg1 + | intros arg1 Harg11 Harg1 ]. + all: eapply Compilers.type.andb_bool_impl_and_for_each_lhs_of_arrow in Harg1; [ | apply ZRange.type.option.is_bounded_by_impl_related_hetero ]. + all: eapply ident.interp_eval_with_bound with (abstraction_relation':=@abstraction_relation') (abstract_domain'_R:=fun t => abstract_domain'_R t); eauto using bottom'_bottom, interp_annotate_expr with typeclass_instances. + all: intros; (eapply extract_list_state_related + eapply extract_option_state_related + eapply strip_annotation_related); eassumption. + Qed. + + Lemma interp_eta_expand_with_bound + {t} (e1 e2 : expr t) + (Hwf : expr.wf nil e1 e2) + (Ht : type.is_not_higher_order t = true) + (b_in : type.for_each_lhs_of_arrow abstract_domain t) + : forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), + type.app_curried (interp (partial.eta_expand_with_bound relax_zrange e1 b_in)) arg1 = type.app_curried (interp e2) arg2. + Proof using Hrelax. + cbv [partial.eta_expand_with_bound]; intros arg1 arg2 Harg12 Harg1. + eapply Compilers.type.andb_bool_impl_and_for_each_lhs_of_arrow in Harg1. + { apply ident.interp_eta_expand_with_bound with (abstraction_relation':=@abstraction_relation') (abstract_domain'_R:=fun t => abstract_domain'_R t); eauto using bottom'_bottom, interp_annotate_expr with typeclass_instances. + all: intros; (eapply extract_list_state_related + eapply extract_option_state_related); eassumption. } + { apply ZRange.type.option.is_bounded_by_impl_related_hetero. } + Qed. + + Lemma Interp_EvalWithBound + {opts : AbstractInterpretation.Options} + {assume_cast_truncates : bool} + {skip_annotations_under : forall t, ident t -> bool} + {strip_preexisting_annotations : bool} + {t} (e : Expr t) + (Hwf : expr.Wf e) + (Ht : type.is_not_higher_order t = true) + (st : type.for_each_lhs_of_arrow abstract_domain t) + (Hst : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) st) + : (forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) st arg1 = true), + type.app_curried (expr.Interp (@ident.interp) (EvalWithBound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e st)) arg1 + = type.app_curried (expr.Interp (@ident.interp) e) arg2) + /\ (forall arg1 + (Harg11 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg1) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) st arg1 = true), + abstraction_relation' + (Extract assume_cast_truncates e st) + (type.app_curried (expr.Interp (@ident.interp) (EvalWithBound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e st)) arg1)). + Proof using Hrelax. cbv [Extract EvalWithBound]; apply interp_eval_with_bound; try apply expr.Wf3_of_Wf; auto. Qed. + + Lemma Interp_EtaExpandWithBound + {t} (E : Expr t) + (Hwf : Wf E) + (Ht : type.is_not_higher_order t = true) + (b_in : type.for_each_lhs_of_arrow abstract_domain t) + : forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), + type.app_curried (Interp (partial.EtaExpandWithBound relax_zrange E b_in)) arg1 = type.app_curried (Interp E) arg2. + Proof using Hrelax. cbv [partial.EtaExpandWithBound]; apply interp_eta_expand_with_bound; eauto with typeclass_instances. Qed. + End with_relax. + + Lemma strip_ranges_is_looser t b v + : @ZRange.type.option.is_bounded_by t b v = true + -> ZRange.type.option.is_bounded_by (ZRange.type.option.strip_ranges b) v = true. + Proof using Type. + induction t as [t|s IHs d IHd]; cbn in *; [ | tauto ]. + induction t; cbn in *; break_innermost_match; cbn in *; rewrite ?Bool.andb_true_iff; try solve [ intuition ]. + { repeat match goal with ls : list _ |- _ => revert ls end. + intros ls1 ls2; revert ls2. + induction ls1, ls2; cbn in *; rewrite ?Bool.andb_true_iff; solve [ intuition ]. } + { destruct_head' option; cbn; eauto; congruence. } + Qed. + + Lemma andb_strip_ranges_Proper t (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) arg1 + : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true -> + type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) + (type.map_for_each_lhs_of_arrow (@ZRange.type.option.strip_ranges) b_in) arg1 = true. + Proof using Type. + induction t as [t|s IHs d IHd]; cbn [type.andb_bool_for_each_lhs_of_arrow type.map_for_each_lhs_of_arrow type.for_each_lhs_of_arrow] in *; + rewrite ?Bool.andb_true_iff; [ tauto | ]. + destruct_head'_prod; cbn [fst snd]; intros [? ?]. + erewrite IHd by eauto. + split; [ | reflexivity ]. + apply strip_ranges_is_looser; assumption. + Qed. + + Lemma strip_ranges_Proper t + : Proper (abstract_domain_R ==> abstract_domain_R) (@ZRange.type.option.strip_ranges t). + Proof using Type. + induction t as [t|s IHs d IHd]; cbn in *. + all: cbv [Proper respectful abstract_domain_R] in *; intros; subst; eauto. + Qed. + + Lemma and_strip_ranges_Proper' t + : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R) ==> type.and_for_each_lhs_of_arrow (@abstract_domain_R)) + (type.map_for_each_lhs_of_arrow (@ZRange.type.option.strip_ranges) (t:=t)). + Proof using Type. + induction t as [t|s IHs d IHd]; cbn [type.and_for_each_lhs_of_arrow type.map_for_each_lhs_of_arrow abstract_domain_R type.for_each_lhs_of_arrow] in *; + cbv [Proper respectful] in *; [ tauto | ]. + intros; destruct_head'_prod; cbn [fst snd] in *; destruct_head'_and. + split; [ | solve [ auto ] ]. + apply strip_ranges_Proper; auto. + Qed. + + Lemma interp_strip_annotations + {opts : AbstractInterpretation.Options} + {assume_cast_truncates : bool} + {t} (e_st e1 e2 : expr t) + (Hwf : expr.wf3 nil e_st e1 e2) + (Hwf' : expr.wf nil e2 e2) + (Ht : type.is_not_higher_order t = true) + (st : type.for_each_lhs_of_arrow abstract_domain t) + : (forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) st arg1 = true), + type.app_curried (expr.interp (@ident.interp) (strip_annotations assume_cast_truncates e1 st)) arg1 + = type.app_curried (expr.interp (@ident.interp) e2) arg2) + /\ (forall arg1 + (Harg11 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg1) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) st arg1 = true), + abstraction_relation' + (extract assume_cast_truncates e_st st) + (type.app_curried (expr.interp (@ident.interp) (strip_annotations assume_cast_truncates e1 st)) arg1)). + Proof using Type. + cbv [strip_annotations]; split; + [ intros arg1 arg2 Harg12 Harg1 + | intros arg1 Harg11 Harg1 ]. + all: eapply Compilers.type.andb_bool_impl_and_for_each_lhs_of_arrow in Harg1; [ | apply ZRange.type.option.is_bounded_by_impl_related_hetero ]. + all: eapply ident.interp_eval_with_bound with (abstraction_relation':=@abstraction_relation') (abstract_domain'_R:=fun t => abstract_domain'_R t); eauto using bottom'_bottom, interp_annotate_expr, default_relax_zrange_good, abstract_interp_ident_related with typeclass_instances. + all: intros; (eapply extract_list_state_related + eapply extract_option_state_related + eapply strip_annotation_related); eassumption. + Qed. + + Lemma Interp_StripAnnotations + {opts : AbstractInterpretation.Options} + {assume_cast_truncates : bool} + {t} (E : Expr t) + (Hwf : Wf E) + (Ht : type.is_not_higher_order t = true) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + : forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), + type.app_curried (expr.Interp (@ident.interp) (StripAnnotations assume_cast_truncates E b_in)) arg1 + = type.app_curried (expr.Interp (@ident.interp) E) arg2. + Proof using Type. + apply (@interp_strip_annotations _ assume_cast_truncates t (E _) (E _) (E _)); try apply expr.Wf3_of_Wf; auto. + Qed. + + Lemma Interp_StripAnnotations_bounded + {opts : AbstractInterpretation.Options} + {assume_cast_truncates : bool} + {t} (E : Expr t) + (Hwf : expr.Wf E) + (Ht : type.is_not_higher_order t = true) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + : forall arg1 + (Harg11 : Proper (type.and_for_each_lhs_of_arrow (@type.eqv)) arg1) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), + ZRange.type.base.option.is_bounded_by + (partial.Extract assume_cast_truncates E b_in) + (type.app_curried (expr.Interp (@ident.interp) (StripAnnotations assume_cast_truncates E b_in)) arg1) + = true. + Proof using Type. + apply (@interp_strip_annotations _ assume_cast_truncates t (E _) (E _) (E _)); try apply expr.Wf3_of_Wf; auto. + Qed. + + Lemma Interp_EtaExpandWithListInfoFromBound + {t} (E : Expr t) + (Hwf : Wf E) + (Ht : type.is_not_higher_order t = true) + (b_in : type.for_each_lhs_of_arrow abstract_domain t) + : forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), + type.app_curried (Interp (partial.EtaExpandWithListInfoFromBound E b_in)) arg1 = type.app_curried (Interp E) arg2. + Proof using Type. + cbv [partial.EtaExpandWithListInfoFromBound]. + intros; apply Interp_EtaExpandWithBound; trivial. + { exact default_relax_zrange_good. } + { apply andb_strip_ranges_Proper; assumption. } + Qed. + End specialized. + End partial. + Import API. + + Lemma Interp_PartialEvaluateWithBounds + {opts : AbstractInterpretation.Options} + relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations + (Hrelax : forall r r' z, is_tighter_than_bool z r = true + -> relax_zrange r = Some r' + -> is_tighter_than_bool z r' = true) + {t} (E : Expr t) + (Hwf : Wf E) + (Ht : type.is_not_higher_order t = true) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + : forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), + type.app_curried (expr.Interp (@ident.interp) (PartialEvaluateWithBounds relax_zrange skip_annotations_under strip_preexisting_annotations assume_cast_truncates E b_in)) arg1 + = type.app_curried (expr.Interp (@ident.interp) E) arg2. + Proof. + cbv [PartialEvaluateWithBounds]. + intros arg1 arg2 Harg12 Harg1. + assert (arg1_Proper : Proper (type.and_for_each_lhs_of_arrow (@type.related base.type base.interp (fun _ => eq))) arg1) + by (hnf; etransitivity; [ eassumption | symmetry; eassumption ]). + assert (arg2_Proper : Proper (type.and_for_each_lhs_of_arrow (@type.related base.type base.interp (fun _ => eq))) arg2) + by (hnf; etransitivity; [ symmetry; eassumption | eassumption ]). + rewrite <- (GeneralizeVar.Interp_gen1_GeneralizeVar E) by auto with wf. + eapply Interp_EvalWithBound; eauto with wf typeclass_instances. + Qed. + + Lemma Interp_PartialEvaluateWithBounds_bounded + {opts : AbstractInterpretation.Options} + relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations + (Hrelax : forall r r' z, is_tighter_than_bool z r = true + -> relax_zrange r = Some r' + -> is_tighter_than_bool z r' = true) + {t} (E : Expr t) + (Hwf : Wf E) + (Ht : type.is_not_higher_order t = true) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + : forall arg1 + (Harg11 : Proper (type.and_for_each_lhs_of_arrow (@type.eqv)) arg1) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), + ZRange.type.base.option.is_bounded_by + (partial.Extract assume_cast_truncates E b_in) + (type.app_curried (expr.Interp (@ident.interp) (PartialEvaluateWithBounds relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations E b_in)) arg1) + = true. + Proof. + cbv [PartialEvaluateWithBounds]. + intros arg1 Harg11 Harg1. + rewrite <- Extract_GeneralizeVar by auto with wf typeclass_instances. + eapply @Interp_EvalWithBound; eauto with wf typeclass_instances. + Qed. + + Lemma Interp_PartialEvaluateWithListInfoFromBounds + {opts : AbstractInterpretation.Options} + {t} (E : Expr t) + (Hwf : Wf E) + (Ht : type.is_not_higher_order t = true) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + : forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), + type.app_curried (Interp (PartialEvaluateWithListInfoFromBounds E b_in)) arg1 = type.app_curried (Interp E) arg2. + Proof. + cbv [PartialEvaluateWithListInfoFromBounds]; intros arg1 arg2 Harg12 Harg1. + assert (arg1_Proper : Proper (type.and_for_each_lhs_of_arrow (@type.related base.type base.interp (fun _ => eq))) arg1) + by (hnf; etransitivity; [ eassumption | symmetry; eassumption ]). + assert (arg2_Proper : Proper (type.and_for_each_lhs_of_arrow (@type.related base.type base.interp (fun _ => eq))) arg2) + by (hnf; etransitivity; [ symmetry; eassumption | eassumption ]). + rewrite <- (GeneralizeVar.Interp_gen1_GeneralizeVar E) by auto. + apply Interp_EtaExpandWithListInfoFromBound; eauto with wf. + Qed. + + Theorem CheckedPartialEvaluateWithBounds_Correct + {opts : AbstractInterpretation.Options} + (relax_zrange : zrange -> option zrange) + (assume_cast_truncates : bool) + (skip_annotations_under : forall t, ident t -> bool) + (strip_preexisting_annotations : bool) + (Hrelax : forall r r' z, is_tighter_than_bool z r = true + -> relax_zrange r = Some r' + -> is_tighter_than_bool z r' = true) + {t} (E : Expr t) + (Hwf : Wf E) + (Ht : type.is_not_higher_order t = true) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + (b_out : ZRange.type.base.option.interp (type.final_codomain t)) + rv (Hrv : CheckedPartialEvaluateWithBounds relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations E b_in b_out = inl rv) + : (forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), + ZRange.type.base.option.is_bounded_by b_out (type.app_curried (Interp rv) arg1) = true + /\ type.app_curried (Interp rv) arg1 = type.app_curried (Interp E) arg2) + /\ Wf rv. + Proof. + cbv [CheckedPartialEvaluateWithBounds Let_In] in *; + break_innermost_match_hyps; inversion_sum; subst. + all: split; + [ intros arg1 arg2 Harg12 Harg1; + assert (arg1_Proper : Proper (type.and_for_each_lhs_of_arrow (@type.related base.type base.interp (fun _ => eq))) arg1) + by (hnf; etransitivity; [ eassumption | symmetry; eassumption ]); + split; + repeat first [ rewrite !@GeneralizeVar.Interp_gen1_FromFlat_ToFlat by eauto with wf typeclass_instances + | rewrite <- Extract_FromFlat_ToFlat by auto with typeclass_instances; apply Interp_PartialEvaluateWithBounds_bounded; auto + | rewrite Extract_FromFlat_ToFlat by auto with wf typeclass_instances + | progress intros + | eapply ZRange.type.base.option.is_tighter_than_is_bounded_by; [ eassumption | ] + | solve [ eauto with wf typeclass_instances ] + | erewrite !Interp_PartialEvaluateWithBounds + | apply type.app_curried_Proper + | apply expr.Wf_Interp_Proper_gen + | progress intros ] + | eauto with wf typeclass_instances ]. + Qed. +End Compilers. diff --git a/src/AbstractInterpretation/Bottomify/Wf.v b/src/AbstractInterpretation/Bottomify/Wf.v new file mode 100644 index 0000000000..9cc84eb178 --- /dev/null +++ b/src/AbstractInterpretation/Bottomify/Wf.v @@ -0,0 +1,1188 @@ +Require Import Coq.micromega.Lia. +Require Import Coq.ZArith.ZArith. +Require Import Coq.Classes.Morphisms. +Require Import Coq.Classes.RelationPairs. +Require Import Coq.Relations.Relations. +Require Import Crypto.Util.ZRange. +Require Import Crypto.Util.Sum. +Require Import Crypto.Util.LetIn. +Require Import Crypto.Util.Prod. +Require Import Crypto.Util.Sigma. +Require Import Crypto.Util.Option. +Require Import Crypto.Util.ListUtil. +Require Import Crypto.Util.ListUtil.Forall. +Require Import Crypto.Util.NatUtil. +Require Import Crypto.Util.Bool.Reflect. +Require Import Crypto.Util.ZUtil.Tactics.LtbToLt. +Require Import Crypto.Util.Tactics.BreakMatch. +Require Import Crypto.Util.Tactics.DestructHead. +Require Import Crypto.Util.Tactics.SplitInContext. +Require Import Crypto.Util.Tactics.UniquePose. +Require Import Crypto.Util.Tactics.SpecializeBy. +Require Import Crypto.Util.Tactics.SpecializeAllWays. +Require Import Rewriter.Language.Language. +Require Import Rewriter.Language.Inversion. +Require Import Crypto.Language.InversionExtra. +Require Import Rewriter.Language.Wf. +Require Import Rewriter.Language.UnderLetsProofs. +Require Import Rewriter.Util.Decidable. +Require Import Crypto.AbstractInterpretation.Bottomify.AbstractInterpretation. +Require Import Crypto.AbstractInterpretation.ZRangeCommonProofs. +Import Coq.Lists.List. + +Import EqNotations. +Module Compilers. + Import Language.Compilers. + Import UnderLets.Compilers. + Import AbstractInterpretation.Compilers. + Import AbstractInterpretation.ZRangeCommonProofs.Compilers. + Import Language.Inversion.Compilers. + Import Language.InversionExtra.Compilers. + Import Language.Wf.Compilers. + Import UnderLetsProofs.Compilers. + Import invert_expr. + + Module Import partial. + Import AbstractInterpretation.Compilers.partial. + Import UnderLets.Compilers.UnderLets. + Section with_type. + Context {base_type : Type}. + Local Notation type := (type base_type). + Let type_base (x : base_type) : type := type.base x. + Local Coercion type_base : base_type >-> type. + Context {ident : type -> Type}. + Local Notation expr := (@expr base_type ident). + Local Notation Expr := (@expr.Expr base_type ident). + Local Notation UnderLets := (@UnderLets base_type ident). + Context (abstract_domain' : base_type -> Type) + (bottom' : forall A, abstract_domain' A) + (abstract_interp_ident : forall t, ident t -> type.interp abstract_domain' t) + (abstract_domain'_R : forall t, abstract_domain' t -> abstract_domain' t -> Prop) + {abstract_interp_ident_Proper : forall t, Proper (eq ==> abstract_domain'_R t) (abstract_interp_ident t)} + {bottom'_Proper : forall t, Proper (abstract_domain'_R t) (bottom' t)}. + Local Notation value var := (@value base_type ident var abstract_domain'). + Local Notation abstract_domain := (@abstract_domain base_type abstract_domain'). + Local Notation bottom := (@bottom base_type abstract_domain' (@bottom')). + Local Notation bottom_for_each_lhs_of_arrow := (@bottom_for_each_lhs_of_arrow base_type abstract_domain' (@bottom')). + + Section with_var2. + Context {var1 var2 : type -> Type}. + Local Notation UnderLets1 := (@UnderLets.UnderLets base_type ident var1). + Local Notation UnderLets2 := (@UnderLets.UnderLets base_type ident var2). + Local Notation expr1 := (@expr.expr base_type ident var1). + Local Notation expr2 := (@expr.expr base_type ident var2). + Local Notation value1 := (@value var1). + Local Notation value2 := (@value var2). + Local Notation value_with_lets1 := (@value_with_lets base_type ident var1 abstract_domain'). + Local Notation value_with_lets2 := (@value_with_lets base_type ident var2 abstract_domain'). + Local Notation state_of_value1 := (@state_of_value base_type ident var1 abstract_domain' bottom'). + Local Notation state_of_value2 := (@state_of_value base_type ident var2 abstract_domain' bottom'). + Context (annotate1 : forall (is_let_bound : bool) t, abstract_domain' t -> @expr1 t -> UnderLets1 (@expr1 t)) + (annotate2 : forall (is_let_bound : bool) t, abstract_domain' t -> @expr2 t -> UnderLets2 (@expr2 t)) + (annotate_Proper + : forall is_let_bound t G + v1 v2 (Hv : abstract_domain'_R t v1 v2) + e1 e2 (He : expr.wf G e1 e2), + UnderLets.wf (fun G' => expr.wf G') G (annotate1 is_let_bound t v1 e1) (annotate2 is_let_bound t v2 e2)) + (interp_ident1 : bool -> forall t, ident t -> value_with_lets1 t) + (interp_ident2 : bool -> forall t, ident t -> value_with_lets2 t) + (skip_annotations_under : forall t, ident t -> bool). + Local Notation reify1 := (@reify base_type ident var1 abstract_domain' annotate1 bottom'). + Local Notation reify2 := (@reify base_type ident var2 abstract_domain' annotate2 bottom'). + Local Notation reflect1 := (@reflect base_type ident var1 abstract_domain' annotate1 bottom'). + Local Notation reflect2 := (@reflect base_type ident var2 abstract_domain' annotate2 bottom'). + Local Notation bottomify1 := (@bottomify base_type ident var1 abstract_domain' bottom'). + Local Notation bottomify2 := (@bottomify base_type ident var2 abstract_domain' bottom'). + Local Notation interp1 := (@interp base_type ident var1 abstract_domain' annotate1 bottom' skip_annotations_under interp_ident1). + Local Notation interp2 := (@interp base_type ident var2 abstract_domain' annotate2 bottom' skip_annotations_under interp_ident2). + Local Notation eval_with_bound'1 := (@eval_with_bound' base_type ident var1 abstract_domain' annotate1 bottom' skip_annotations_under interp_ident1). + Local Notation eval_with_bound'2 := (@eval_with_bound' base_type ident var2 abstract_domain' annotate2 bottom' skip_annotations_under interp_ident2). + Local Notation eval'1 := (@eval' base_type ident var1 abstract_domain' annotate1 bottom' skip_annotations_under interp_ident1). + Local Notation eval'2 := (@eval' base_type ident var2 abstract_domain' annotate2 bottom' skip_annotations_under interp_ident2). + Local Notation eta_expand_with_bound'1 := (@eta_expand_with_bound' base_type ident var1 abstract_domain' annotate1 bottom'). + Local Notation eta_expand_with_bound'2 := (@eta_expand_with_bound' base_type ident var2 abstract_domain' annotate2 bottom'). + + Definition abstract_domain_R {t} : relation (abstract_domain t) + := type.related abstract_domain'_R. + + (** This one is tricky. Because we need to be stable under + weakening and reordering of the context, we permit any + context for well-formedness of the input in the arrow + case, and simply tack on that context at the beginning of + the output. This is sort-of wasteful on the output + context, but it's sufficient to prove + [wf_value_Proper_list] below, which is what we really + need. *) + Fixpoint wf_value G {t} : value1 t -> value2 t -> Prop + := match t return value1 t -> value2 t -> Prop with + | type.base t + => fun v1 v2 + => abstract_domain_R (fst v1) (fst v2) + /\ expr.wf G (snd v1) (snd v2) + | type.arrow s d + => fun v1 v2 + => forall seg G' sv1 sv2, + G' = (seg ++ G)%list + -> @wf_value seg s sv1 sv2 + -> UnderLets.wf + (fun G' => @wf_value G' d) G' + (v1 sv1) (v2 sv2) + end. + + Definition wf_value_with_lets G {t} : value_with_lets1 t -> value_with_lets2 t -> Prop + := UnderLets.wf (fun G' => wf_value G') G. + + Context (interp_ident_Proper + : forall G t idc1 idc2 (Hidc : idc1 = idc2) annotate_with_state, + wf_value_with_lets G (interp_ident1 annotate_with_state t idc1) (interp_ident2 annotate_with_state t idc2)). + + Global Instance bottom_Proper {t} : Proper abstract_domain_R (@bottom t) | 10. + Proof using bottom'_Proper. + clear -bottom'_Proper type_base. + cbv [Proper] in *; induction t; cbn; cbv [respectful]; eauto. + Qed. + + Global Instance bottom_for_each_lhs_of_arrow_Proper {t} + : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) (@bottom_for_each_lhs_of_arrow t) | 10. + Proof using bottom'_Proper. + clear -bottom'_Proper type_base. + pose proof (@bottom_Proper). + cbv [Proper] in *; induction t; cbn; cbv [respectful]; eauto. + Qed. + + Lemma state_of_value_Proper G {t} v1 v2 (Hv : @wf_value G t v1 v2) + : abstract_domain_R (state_of_value1 v1) (state_of_value2 v2). + Proof using bottom'_Proper. + clear -Hv type_base bottom'_Proper. + destruct t; [ destruct v1, v2, Hv | ]; cbn in *; cbv [respectful]; eauto; intros; apply bottom_Proper. + Qed. + + Local Hint Resolve (fun A (P : list A -> Prop) => ex_intro P nil) (fun A (x : A) (P : list A -> Prop) => ex_intro P (cons x nil)) : core. + Local Hint Constructors expr.wf ex : core. + Local Hint Unfold List.In : core. + + Lemma wf_value_Proper_list G1 G2 + (HG1G2 : forall t v1 v2, List.In (existT _ t (v1, v2)) G1 -> List.In (existT _ t (v1, v2)) G2) + t e1 e2 + (Hwf : @wf_value G1 t e1 e2) + : @wf_value G2 t e1 e2. + Proof using Type. + clear -type_base HG1G2 Hwf. + revert dependent G1; revert dependent G2; induction t; intros; + repeat first [ progress cbn in * + | progress intros + | solve [ eauto ] + | progress subst + | progress destruct_head'_and + | progress destruct_head'_or + | apply conj + | rewrite List.in_app_iff in * + | match goal with H : _ |- _ => apply H; clear H end + | wf_unsafe_t_step + | eapply UnderLets.wf_Proper_list; [ | | solve [ eauto ] ] ]. + Qed. + + Fixpoint wf_reify (annotate_with_state : bool) (is_let_bound : bool) G {t} + : forall v1 v2 (Hv : @wf_value G t v1 v2) + s1 s2 (Hs : type.and_for_each_lhs_of_arrow (@abstract_domain_R) s1 s2), + UnderLets.wf (fun G' => expr.wf G') G (@reify1 annotate_with_state is_let_bound t v1 s1) (@reify2 annotate_with_state is_let_bound t v2 s2) + with wf_reflect (annotate_with_state : bool) G {t} + : forall e1 e2 (He : expr.wf G e1 e2) + s1 s2 (Hs : abstract_domain_R s1 s2), + @wf_value G t (@reflect1 annotate_with_state t e1 s1) (@reflect2 annotate_with_state t e2 s2). + Proof using annotate_Proper bottom'_Proper. + all: clear -wf_reflect wf_reify annotate_Proper type_base bottom'_Proper. + all: pose proof (@bottom_for_each_lhs_of_arrow_Proper); cbv [Proper abstract_domain_R] in *. + all: destruct t as [t|s d]; + [ clear wf_reify wf_reflect + | pose proof (fun G => wf_reflect annotate_with_state G s) as wf_reflect_s; + pose proof (fun G => wf_reflect annotate_with_state G d) as wf_reflect_d; + pose proof (fun G => wf_reify annotate_with_state false G s) as wf_reify_s; + pose proof (fun G => wf_reify annotate_with_state false G d) as wf_reify_d; + pose proof (@bottom_Proper s); + clear wf_reify wf_reflect ]. + all: cbn [reify reflect] in *. + all: fold (@reify2) (@reflect2) (@reify1) (@reflect1). + all: cbn in *. + all: repeat first [ progress cbn [fst snd] in * + | progress cbv [respectful] in * + | progress intros + | progress subst + | progress destruct_head'_and + | progress destruct_head'_ex + | solve [ eauto | wf_t ] + | apply annotate_Proper + | apply UnderLets.wf_to_expr + | break_innermost_match_step + | match goal with + | [ |- UnderLets.wf _ _ _ _ ] => constructor + | [ |- expr.wf _ _ _ ] => constructor + | [ He : forall seg G' sv1 sv2, G' = (seg ++ ?G)%list -> _ |- UnderLets.wf _ (?v :: ?G) (UnderLets.splice _ _) (UnderLets.splice _ _) ] + => eapply UnderLets.wf_splice; [ apply (He (cons v nil)) | ] + | [ |- UnderLets.wf _ _ (UnderLets.splice (reify1 _ _ _ _) _) (UnderLets.splice (reify2 _ _ _ _) _) ] + => eapply UnderLets.wf_splice; [ apply wf_reify_s || apply wf_reify_d | ] + | [ |- wf_value _ (reflect1 _ _ _) (reflect2 _ _ _) ] => apply wf_reflect_s || apply wf_reflect_d + | [ H : wf_value _ ?x ?y |- wf_value _ ?x ?y ] + => eapply wf_value_Proper_list; [ | eassumption ] + | [ H : forall x y, ?R x y -> ?R' (?f x) (?g y) |- ?R' (?f _) (?g _) ] + => apply H + | [ |- ?R (state_of_value1 _) (state_of_value2 _) ] => eapply state_of_value_Proper + end ]. + Qed. + + Lemma wf_bottomify {t} G v1 v2 + (Hwf : @wf_value G t v1 v2) + : wf_value_with_lets G (bottomify1 v1) (bottomify2 v2). + Proof using bottom'_Proper. + cbv [wf_value_with_lets] in *. + revert dependent G; induction t as [|s IHs d IHd]; intros; + cbn [bottomify wf_value]; fold (@value1) (@value2) in *; break_innermost_match; + constructor. + all: repeat first [ progress cbn [fst snd wf_value] in * + | progress destruct_head'_and + | assumption + | apply bottom'_Proper + | apply conj + | progress intros + | progress subst + | solve [ eapply UnderLets.wf_splice; eauto ] ]. + Qed. + + Local Ltac wf_interp_t := + repeat first [ progress cbv [wf_value_with_lets abstract_domain_R respectful] in * + | progress cbn [wf_value fst snd partial.bottom type.related eq_rect List.In] in * + | wf_safe_t_step + | exact I + | apply wf_reify + | apply bottom_Proper + | progress destruct_head'_ex + | progress destruct_head'_or + | eapply UnderLets.wf_splice + | match goal with + | [ |- UnderLets.wf _ _ (bottomify1 _) (bottomify2 _) ] => apply wf_bottomify + | [ |- UnderLets.wf _ _ _ _ ] => constructor + | [ |- and _ _ ] => apply conj + end + | eapply wf_value_Proper_list; [ | solve [ eauto ] ] + | eapply UnderLets.wf_Proper_list; [ | | solve [ eauto ] ] + | match goal with + | [ H : _ |- _ ] => eapply H; clear H; solve [ wf_interp_t ] + end + | break_innermost_match_step ]. + + Local Notation skip_annotations_for_App := (@skip_annotations_for_App base_type ident skip_annotations_under). + + Lemma wf_skip_annotations_for_App {var1' var2' G t e1 e2} (Hwf : expr.wf G (t:=t) e1 e2) + : @skip_annotations_for_App var1' t e1 = @skip_annotations_for_App var2' t e2. + Proof using Type. + cbv [skip_annotations_for_App]; break_innermost_match; + expr.invert_subst; + repeat first [ progress cbn [projT1 projT2 fst snd eq_rect invert_App_curried invert_Ident Option.bind] in * + | reflexivity + | exfalso; assumption + | progress inversion_option + | progress destruct_head'_sig + | progress destruct_head'_and + | progress expr.inversion_wf_constr + | progress subst + | match goal with + | [ H : expr.wf _ (App_curried _ _) _ |- _ ] + => apply expr.invert_wf_App_curried_or_eq_base in H; + [ | clear H .. ] + | [ |- _ \/ _ ] => right; split; intros; congruence + | [ Hwf : expr.wf _ (App_curried _ _) ?e, H' : invert_AppIdent_curried _ = None |- _ ] + => apply expr.wf_invert_AppIdent_curried in Hwf; rewrite H' in Hwf; cbv [invert_AppIdent_curried Option.option_eq] in Hwf + | [ Hwf : expr.wf _ ?e (App_curried _ _), H' : invert_AppIdent_curried _ = None |- _ ] + => apply expr.wf_invert_AppIdent_curried in Hwf; rewrite H' in Hwf; cbv [Option.option_eq invert_AppIdent_curried] in Hwf + | [ H : context[invert_App_curried (App_curried _ _)] |- _ ] + => rewrite expr.invert_App_curried_App_curried in Hwf + end ]. + Qed. + + Lemma wf_interp (annotate_with_state : bool) G G' {t} (e1 : @expr (@value_with_lets1) t) (e2 : @expr (@value_with_lets2) t) + (Hwf : expr.wf G e1 e2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) + : wf_value_with_lets G' (interp1 annotate_with_state e1) (interp2 annotate_with_state e2). + Proof using annotate_Proper bottom'_Proper interp_ident_Proper. + revert dependent G'; revert annotate_with_state; induction Hwf; intros; cbn [interp]; + try solve [ apply interp_ident_Proper; auto + | eauto ]; + match goal with + | [ G' : list _ |- context[skip_annotations_for_App ?e1v] ] + => match goal with + | [ |- context[skip_annotations_for_App ?e2v] ] + => epose proof (wf_skip_annotations_for_App (e1:=e1v) (e2:=e2v) (G:=G') ltac:(solve [ wf_t ])); + generalize dependent (skip_annotations_for_App e1v); + generalize dependent (skip_annotations_for_App e2v); intros; + subst + end + end; + wf_interp_t. + Qed. + + Lemma wf_eval_with_bound' (annotate_with_state : bool) G G' {t} e1 e2 (He : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) + : expr.wf G' (@eval_with_bound'1 annotate_with_state t e1 st1) (@eval_with_bound'2 annotate_with_state t e2 st2). + Proof using annotate_Proper bottom'_Proper interp_ident_Proper. + eapply UnderLets.wf_to_expr, UnderLets.wf_splice. + { eapply wf_interp; solve [ eauto ]. } + { intros; destruct_head'_ex; subst; eapply wf_reify; eauto. } + Qed. + + Lemma wf_eval' G G' {t} e1 e2 (He : expr.wf G e1 e2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) + : expr.wf G' (@eval'1 t e1) (@eval'2 t e2). + Proof using annotate_Proper bottom'_Proper interp_ident_Proper. + eapply wf_eval_with_bound'; eauto; apply bottom_for_each_lhs_of_arrow_Proper. + Qed. + + Lemma wf_eta_expand_with_bound' G {t} e1 e2 (He : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) + : expr.wf G (@eta_expand_with_bound'1 t e1 st1) (@eta_expand_with_bound'2 t e2 st2). + Proof using annotate_Proper bottom'_Proper. + eapply UnderLets.wf_to_expr, wf_reify; [ eapply wf_reflect | ]; eauto; apply bottom_Proper. + Qed. + End with_var2. + End with_type. + + Module ident. + Import API. + Local Notation UnderLets := (@UnderLets base.type ident). + Section with_type. + Context (abstract_domain' : base.type -> Type). + Local Notation abstract_domain := (@abstract_domain base.type abstract_domain'). + Context (bottom' : forall A, abstract_domain' A) + (abstract_interp_ident : forall t, ident t -> type.interp abstract_domain' t) + (extract_list_state : forall A, abstract_domain' (base.type.list A) -> option (list (abstract_domain' A))) + (extract_option_state : forall A, abstract_domain' (base.type.option A) -> option (option (abstract_domain' A))). + Context (abstract_domain'_R : forall t, abstract_domain' t -> abstract_domain' t -> Prop). + Local Notation abstract_domain_R := (@abstract_domain_R base.type abstract_domain' abstract_domain'_R). + Context {abstract_interp_ident_Proper : forall t, Proper (eq ==> @abstract_domain_R t) (abstract_interp_ident t)} + {bottom'_Proper : forall t, Proper (abstract_domain'_R t) (bottom' t)} + (extract_list_state_length : forall t v1 v2, abstract_domain'_R _ v1 v2 -> option_map (@length _) (extract_list_state t v1) = option_map (@length _) (extract_list_state t v2)) + (extract_list_state_rel : forall t v1 v2, abstract_domain'_R _ v1 v2 -> forall l1 l2, extract_list_state t v1 = Some l1 -> extract_list_state t v2 = Some l2 -> List.Forall2 (@abstract_domain'_R t) l1 l2) + (extract_option_state_rel : forall t v1 v2, abstract_domain'_R _ v1 v2 -> option_eq (option_eq (abstract_domain'_R _)) (extract_option_state t v1) (extract_option_state t v2)). + + Local Instance abstract_interp_ident_Proper_arrow s d + : Proper (eq ==> abstract_domain'_R s ==> abstract_domain'_R d) (abstract_interp_ident (type.arrow s d)) + := abstract_interp_ident_Proper (type.arrow s d). + + Local Ltac handle_Forall2_step := + first [ match goal with + | [ |- List.Forall2 _ ?x ?x ] => rewrite Forall2_Forall; cbv [Proper] + | [ |- List.Forall2 _ (List.map _ _) (List.map _ _) ] => rewrite Forall2_map_map_iff + | [ |- List.Forall2 _ (List.combine _ _) (List.combine _ _) ] + => eapply Forall2_combine; [ intros | eassumption | eassumption ] + | [ |- List.Forall2 _ (List.combine ?x _) (List.combine ?x _) ] + => eapply Forall2_combine; + [ intros + | instantiate (1:=fun a b => a = b /\ List.In a x); + rewrite Forall2_Forall, Forall_forall; cbv [Proper]; auto + | eassumption ] + | [ H : expr.wf _ ?d1 ?d2, H' : List.Forall2 (expr.wf _) ?xs ?ys + |- expr.wf ?G (nth_default ?d1 ?xs ?n) (nth_default ?d2 ?ys ?n) ] + => cut (List.Forall2 (expr.wf G) xs ys /\ expr.wf G d1 d2); + [ rewrite Forall2_forall_iff''; + let H := fresh in intros [? H]; apply H + | ] + end ]. + + + Section with_var2. + Context {var1 var2 : type -> Type}. + Local Notation wf_value_with_lets := (@wf_value_with_lets base.type ident abstract_domain' abstract_domain'_R var1 var2). + Local Notation wf_value := (@wf_value base.type ident abstract_domain' abstract_domain'_R var1 var2). + Context (annotate_expr1 : forall t, abstract_domain' t -> option (@expr var1 (t -> t))) + (annotate_expr2 : forall t, abstract_domain' t -> option (@expr var2 (t -> t))) + (is_annotated_for1 : forall t t', @expr var1 t -> abstract_domain' t' -> bool) + (is_annotated_for2 : forall t t', @expr var2 t -> abstract_domain' t' -> bool) + (strip_annotation1 : forall t, ident t -> option (value _ t)) + (strip_annotation2 : forall t, ident t -> option (value _ t)) + (annotation_to_cast1 : forall s d, @expr var1 (s -> d) -> option (@expr var1 s -> @expr var1 d)) + (annotation_to_cast2 : forall s d, @expr var2 (s -> d) -> option (@expr var2 s -> @expr var2 d)) + (skip_annotations_under : forall t, ident t -> bool) + {wf_annotation_to_cast + : forall s d G e1 e2, + expr.wf G e1 e2 + -> option_eq (fun f g => forall e1 e2, expr.wf G e1 e2 -> expr.wf G (f e1) (g e2)) + (annotation_to_cast1 s d e1) + (annotation_to_cast2 s d e2)} + {annotate_expr_Proper : forall t s1 s2, + abstract_domain'_R t s1 s2 + -> option_eq (expr.wf nil) (annotate_expr1 t s1) (annotate_expr2 t s2)} + {is_annotated_for_Proper : forall G t t' e1 e2, + expr.wf G e1 e2 + -> ((abstract_domain'_R _ ==> eq)%signature) + (@is_annotated_for1 t t' e1) + (@is_annotated_for2 t t' e2)} + {wf_strip_annotation + : forall G t idc, + option_eq (wf_value G) + (@strip_annotation1 t idc) + (@strip_annotation2 t idc)}. + + Local Notation update_annotation1 := (@ident.update_annotation var1 abstract_domain' annotate_expr1 is_annotated_for1). + Local Notation update_annotation2 := (@ident.update_annotation var2 abstract_domain' annotate_expr2 is_annotated_for2). + Local Notation annotate1 := (@ident.annotate var1 abstract_domain' annotate_expr1 abstract_interp_ident extract_list_state extract_option_state is_annotated_for1). + Local Notation annotate2 := (@ident.annotate var2 abstract_domain' annotate_expr2 abstract_interp_ident extract_list_state extract_option_state is_annotated_for2). + Local Notation annotate_base1 := (@ident.annotate_base var1 abstract_domain' annotate_expr1 is_annotated_for1). + Local Notation annotate_base2 := (@ident.annotate_base var2 abstract_domain' annotate_expr2 is_annotated_for2). + Local Notation annotate_with_expr1 := (@ident.annotate_with_expr var1 abstract_domain' annotate_expr1 is_annotated_for1). + Local Notation annotate_with_expr2 := (@ident.annotate_with_expr var2 abstract_domain' annotate_expr2 is_annotated_for2). + Local Notation interp_ident1 := (@ident.interp_ident var1 abstract_domain' annotate_expr1 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for1 strip_annotation1). + Local Notation interp_ident2 := (@ident.interp_ident var2 abstract_domain' annotate_expr2 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for2 strip_annotation2). + Local Notation reflect1 := (@reflect base.type ident var1 abstract_domain' annotate1 bottom'). + Local Notation reflect2 := (@reflect base.type ident var2 abstract_domain' annotate2 bottom'). + + Lemma wf_update_annotation G {t} st1 st2 (Hst : abstract_domain'_R t st1 st2) e1 e2 (He : expr.wf G e1 e2) + : expr.wf G (@update_annotation1 t st1 e1) (@update_annotation2 t st2 e2). + Proof using abstract_interp_ident_Proper annotate_expr_Proper is_annotated_for_Proper. + cbv [ident.update_annotation]; + repeat first [ progress subst + | progress expr.invert_subst + | progress cbn [fst snd projT1 projT2 eq_rect] in * + | progress cbn [invert_AppIdent Option.bind invert_App invert_Ident] in * + | progress destruct_head'_sig + | progress destruct_head'_sigT + | progress destruct_head'_and + | progress destruct_head'_prod + | progress destruct_head' False + | progress inversion_option + | progress expr.inversion_wf_constr + | progress expr.inversion_wf_one_constr + | break_innermost_match_hyps_step + | expr.invert_match_step + | progress expr.inversion_expr + | progress rewrite_type_transport_correct + | progress type_beq_to_eq + | progress type.inversion_type + | progress base.type.inversion_type + | discriminate + | match goal with + | [ H : abstract_domain'_R _ ?x _ |- _ ] => rewrite !H + | [ H : abstract_domain'_R _ ?x _, H' : context[?x] |- _ ] => rewrite !H in H' + end + | progress wf_safe_t + | break_innermost_match_step + | solve [ wf_t ] + | match goal with + | [ H : abstract_domain'_R _ ?x ?y, H1 : annotate_expr1 _ ?x = _, H2 : annotate_expr2 _ ?y = _ |- _ ] + => let H' := fresh in + pose proof (annotate_expr_Proper _ _ _ H) as H'; + rewrite H1, H2 in H'; clear H1 H2; cbv [option_eq] in H' + | [ H1 : is_annotated_for1 _ _ ?e1 ?s1 = _, + H2 : is_annotated_for2 _ _ ?e2 ?s2 = _ , + Hwf : expr.wf _ ?e1 ?e2, + Hrel : abstract_domain'_R _ ?s1 ?s2 |- _ ] + => let H' := fresh in + pose proof (is_annotated_for_Proper _ _ _ _ _ Hwf _ _ Hrel) as H'; + rewrite H1, H2 in H'; clear H1 H2 + end ]. + Qed. + + Lemma wf_annotate_with_expr + is_let_bound t G + v1 v2 (Hv : abstract_domain'_R t v1 v2) + e1 e2 (He : expr.wf G e1 e2) + : UnderLets.wf (fun G' => expr.wf G') G (@annotate_with_expr1 is_let_bound t v1 e1) (@annotate_with_expr2 is_let_bound t v2 e2). + Proof using abstract_interp_ident_Proper annotate_expr_Proper is_annotated_for_Proper. + cbv [ident.annotate_with_expr]; break_innermost_match; repeat constructor; apply wf_update_annotation; assumption. + Qed. + + Lemma wf_annotate_base + is_let_bound (t : base.type.base) G + v1 v2 (Hv : abstract_domain'_R t v1 v2) + e1 e2 (He : expr.wf G e1 e2) + : UnderLets.wf (fun G' => expr.wf G') G (@annotate_base1 is_let_bound t v1 e1) (@annotate_base2 is_let_bound t v2 e2). + Proof using abstract_interp_ident_Proper annotate_expr_Proper is_annotated_for_Proper. + cbv [ident.annotate_base]; + repeat first [ apply wf_annotate_with_expr + | break_innermost_match_step + | progress subst + | progress cbv [type_base ident.smart_Literal] in * + | progress cbn [invert_Literal] in * + | discriminate + | progress destruct_head' False + | progress expr.invert_subst + | progress expr.inversion_wf + | wf_safe_t_step + | break_innermost_match_hyps_step + | match goal with + | [ H : _ = _ :> ident _ |- _ ] => inversion H; clear H + | [ |- UnderLets.wf _ _ _ _ ] => constructor + | [ H : abstract_domain'_R _ _ _ |- _ ] => rewrite !H + end + | progress expr.invert_match_step + | progress expr.inversion_expr ]. + Qed. + + Local Opaque ident.ident_Some ident.ident_None. + Lemma wf_annotate + is_let_bound t G + v1 v2 (Hv : abstract_domain'_R t v1 v2) + e1 e2 (He : expr.wf G e1 e2) + : UnderLets.wf (fun G' => expr.wf G') G (@annotate1 is_let_bound t v1 e1) (@annotate2 is_let_bound t v2 e2). + Proof using abstract_interp_ident_Proper annotate_expr_Proper extract_list_state_length extract_list_state_rel extract_option_state_rel is_annotated_for_Proper. + revert dependent G; induction t; intros; + cbn [ident.annotate]; try apply wf_annotate_base; trivial. + all: try solve [ repeat first [ lazymatch goal with + | [ H : expr.wf _ ?e1 ?e2, H' : reflect_list ?e1 = Some _, H'' : reflect_list ?e2 = None |- _ ] + => apply expr.wf_reflect_list in H; rewrite H', H'' in H; exfalso; clear -H; intuition congruence + | [ H : expr.wf _ ?e1 ?e2, H' : reflect_list ?e2 = Some _, H'' : reflect_list ?e1 = None |- _ ] + => apply expr.wf_reflect_list in H; rewrite H', H'' in H; exfalso; clear -H; intuition congruence + | [ H : expr.wf _ (reify_list _) (reify_list _) |- _ ] => apply expr.wf_reify_list in H + | [ |- expr.wf _ (reify_list _) (reify_list _) ] => apply expr.wf_reify_list + | [ |- UnderLets.wf _ _ (UnderLets.splice_list _ _) (UnderLets.splice_list _ _) ] + => eapply @UnderLets.wf_splice_list_no_order with (P:=fun G => expr.wf G); autorewrite with distr_length + | [ H : expr.wf _ (reify_list _) ?e, H' : reflect_list ?e = None |- _ ] + => apply expr.wf_reflect_list in H; rewrite H', expr.reflect_reify_list in H; exfalso; clear -H; intuition congruence + | [ H : expr.wf _ ?e (reify_list _), H' : reflect_list ?e = None |- _ ] + => apply expr.wf_reflect_list in H; rewrite H', expr.reflect_reify_list in H; exfalso; clear -H; intuition congruence + | [ H : extract_list_state ?t ?v1 = ?x1, H' : extract_list_state ?t ?v2 = ?x2, Hv : abstract_domain'_R _ ?v1 ?v2 |- _ ] + => let Hl := fresh in + let Hl' := fresh in + pose proof (extract_list_state_length _ v1 v2 Hv) as Hl; + pose proof (extract_list_state_rel _ v1 v2 Hv) as Hl'; + rewrite H, H' in Hl, Hl'; cbv [option_eq option_map] in Hl, Hl'; clear H H' + | [ H : abstract_domain'_R _ ?v1 ?v2, H1 : extract_option_state _ ?v1 = _, H2 : extract_option_state _ ?v2 = _ |- _ ] + => let H' := fresh in + pose proof H as H'; + apply extract_option_state_rel in H'; + rewrite H1, H2 in H'; + cbv [option_eq] in H'; + clear H1 H2 + | [ H : ?x = ?x |- _ ] => clear H + | [ H : length ?l1 = length ?l2, H' : context[length ?l1] |- _ ] => rewrite H in H' + | [ H : context[invert_pair ?e] |- _ ] + => let lem := lazymatch e with + | (?x, ?y)%expr => constr:(expr.invert_pair_ident_pair(v1:=x) (v2:=y) : invert_pair e = _) + end in + rewrite lem in H + end + | apply wf_annotate_with_expr + | apply DefaultValue.expr.base.wf_default + | apply DefaultValue.expr.wf_default + | progress expr.invert_subst + | progress cbn [ident.annotate ident.smart_Literal invert_Literal invert_pair invert_AppIdent2 invert_App2 fst snd projT2 projT1 eq_rect Option.bind] in * + | progress destruct_head' False + | progress inversion_option + | progress destruct_head'_ex + | discriminate + | wf_safe_t_step + | progress expr.inversion_wf_constr + | progress expr.inversion_expr + | progress type_beq_to_eq + | progress type.inversion_type + | progress base.type.inversion_type + | match goal with + | [ |- expr.wf _ (update_annotation1 _ _) (update_annotation2 _ _) ] => apply wf_update_annotation + | [ H : _ = _ :> ident _ |- _ ] => inversion H; clear H + | [ |- UnderLets.wf _ _ _ _ ] => constructor + | [ H : abstract_domain'_R _ ?x _ |- _ ] => rewrite !H + | [ |- UnderLets.wf _ _ (UnderLets.splice _ _) (UnderLets.splice _ _) ] => eapply UnderLets.wf_splice + | [ H : List.nth_error (List.map _ _) _ = Some _ |- _ ] => apply nth_error_map_ex in H + | [ H : context[List.nth_error (List.combine _ _) _] |- _ ] => rewrite nth_error_combine in H + | [ |- context[List.nth_error (List.combine _ _) _] ] => rewrite nth_error_combine + | [ H : forall x y, Some _ = Some _ -> Some _ = Some _ -> _ |- _ ] + => specialize (H _ _ eq_refl eq_refl) + | [ H : forall v1 v2, List.In (v1, v2) (List.combine ?l1 ?l2) -> ?R v1 v2, H' : List.nth_error ?l1 ?n = Some ?a1, H'' : List.nth_error ?l2 ?n = Some ?a2 + |- ?R ?a1 ?a2 ] + => apply H + | [ H : List.nth_error ?l ?n' = Some ?v |- List.In (?v, _) (List.combine ?l _) ] => apply nth_error_In with (n:=n') + | [ H : length ?x = length ?y |- context[length ?x] ] => rewrite H + end + | handle_Forall2_step + | break_innermost_match_step + | break_innermost_match_hyps_step + | progress expr.invert_match + | progress expr.inversion_wf_one_constr + | match goal with + | [ H : context[UnderLets.wf _ _ (annotate1 _ _ _) (annotate2 _ _ _)] + |- UnderLets.wf _ _ (annotate1 _ _ _) (annotate2 _ _ _) ] => eapply H + end + | apply abstract_interp_ident_Proper_arrow + | progress rewrite_type_transport_correct + | apply conj + | congruence + | progress destruct_head' option + | progress cbn [Option.combine option_map UnderLets.splice_option reify_option option_rect] in * + | progress cbn [type.decode f_equal eq_rect fst snd] in * + | solve [ wf_t ] ] ]. + Qed. + Local Ltac type_of_value v := + lazymatch v with + | (abstract_domain ?t * _)%type => t + | (?a -> UnderLets _ ?b) + => let a' := type_of_value a in + let b' := type_of_value b in + constr:(type.arrow a' b') + end. + Local Opaque ident.ident_Literal. + Lemma wf_interp_ident_nth_default (annotate_with_state : bool) G T + : wf_value_with_lets G (@interp_ident1 annotate_with_state _ (@ident.List_nth_default T)) (@interp_ident2 annotate_with_state _ (@ident.List_nth_default T)). + Proof using abstract_interp_ident_Proper annotate_expr_Proper extract_list_state_length extract_list_state_rel extract_option_state_rel is_annotated_for_Proper. + cbv [wf_value_with_lets wf_value ident.interp_ident]; constructor; cbn -[abstract_domain_R abstract_domain]. + { intros; subst. + destruct_head'_prod; destruct_head'_and; cbn [fst snd] in *. + repeat first [ progress subst + | lazymatch goal with + | [ H : expr.wf _ ?e1 ?e2, H' : reflect_list ?e1 = Some _, H'' : reflect_list ?e2 = None |- _ ] + => apply expr.wf_reflect_list in H; rewrite H', H'' in H; exfalso; clear -H; intuition congruence + | [ H : expr.wf _ ?e1 ?e2, H' : reflect_list ?e2 = Some _, H'' : reflect_list ?e1 = None |- _ ] + => apply expr.wf_reflect_list in H; rewrite H', H'' in H; exfalso; clear -H; intuition congruence + | [ H : expr.wf _ (reify_list _) (reify_list _) |- _ ] => apply expr.wf_reify_list in H + | [ |- expr.wf _ (reify_list _) (reify_list _) ] => apply expr.wf_reify_list + | [ |- UnderLets.wf _ _ (UnderLets.splice_list _ _) (UnderLets.splice_list _ _) ] + => eapply @UnderLets.wf_splice_list_no_order with (P:=fun G => expr.wf G); autorewrite with distr_length + | [ H : expr.wf _ (reify_list _) ?e, H' : reflect_list ?e = None |- _ ] + => apply expr.wf_reflect_list in H; rewrite H', expr.reflect_reify_list in H; exfalso; clear -H; intuition congruence + | [ H : expr.wf _ ?e (reify_list _), H' : reflect_list ?e = None |- _ ] + => apply expr.wf_reflect_list in H; rewrite H', expr.reflect_reify_list in H; exfalso; clear -H; intuition congruence + | [ H : extract_list_state ?t ?v1 = ?x1, H' : extract_list_state ?t ?v2 = ?x2, Hv : abstract_domain_R ?v1 ?v2 |- _ ] + => let Hl := fresh in + let Hl' := fresh in + pose proof (extract_list_state_length _ v1 v2 Hv) as Hl; + pose proof (extract_list_state_rel _ v1 v2 Hv) as Hl'; + rewrite H, H' in Hl, Hl'; cbv [option_eq option_map] in Hl, Hl'; clear H H' + | [ H : ?x = ?x |- _ ] => clear H + | [ H : length ?l1 = length ?l2, H' : context[length ?l1] |- _ ] => rewrite H in H' + end + | match goal with + | [ |- UnderLets.wf ?Q ?G (UnderLets.splice ?x1 ?e1) (UnderLets.splice ?x2 ?e2) ] + => simple refine (@UnderLets.wf_splice _ _ _ _ _ _ _ _ _ Q G x1 x2 _ e1 e2 _); + [ let G := fresh "G" in + intro G; + lazymatch goal with + | [ |- expr _ -> _ -> _ ] + => refine (expr.wf G) + | [ |- ?T -> _ -> _ ] + => let t := type_of_value T in + refine (@wf_value G t) + end + | | ] + | [ |- UnderLets.wf ?Q ?G (UnderLets.Base _) (UnderLets.Base _) ] + => constructor + | [ H : ident.ident_Literal _ = ident.ident_Literal _ |- _ ] + => apply (f_equal (fun idc => invert_Literal (var:=var1) (#idc))) in H; rewrite !expr.invert_Literal_ident_Literal in H + | [ H : _ = _ :> ident _ |- _ ] => inversion H; clear H + | [ H : List.nth_error _ _ = None |- _ ] => apply List.nth_error_None in H + | [ H : List.nth_error _ _ = Some _ |- _ ] + => unique pose proof (@ListUtil.nth_error_value_length _ _ _ _ H); + unique pose proof (@ListUtil.nth_error_value_In _ _ _ _ H) + | [ H : context[List.In _ (List.map _ _)] |- _ ] => rewrite List.in_map_iff in H + | [ H : (?x <= ?y)%nat, H' : (?y < ?x)%nat |- _ ] => exfalso; clear -H H'; lia + | [ H : (?x <= ?y)%nat, H' : (?y < ?x')%nat, H'' : ?x' = ?x |- _ ] => exfalso; clear -H H' H''; lia + | [ H : length ?x = length ?y |- context[length ?x] ] => rewrite H + | [ H : List.nth_error (List.map _ _) _ = Some _ |- _ ] => apply nth_error_map_ex in H + | [ H : context[List.nth_error (List.combine _ _) _] |- _ ] => rewrite nth_error_combine in H + | [ |- context[List.nth_error (List.combine _ _) _] ] => rewrite nth_error_combine + | [ H : forall x y, Some _ = Some _ -> Some _ = Some _ -> _ |- _ ] + => specialize (H _ _ eq_refl eq_refl) + | [ H : forall v1 v2, List.In (v1, v2) (List.combine ?l1 ?l2) -> ?R v1 v2, H' : List.nth_error ?l1 ?n' = Some ?a1, H'' : List.nth_error ?l2 ?n' = Some ?a2 + |- _ ] + => unique pose proof (H a1 a2 ltac:(apply nth_error_In with (n:=n'); rewrite nth_error_combine, H', H''; reflexivity)) + | [ H : List.nth_error ?l ?n' = Some ?v |- List.In (?v, _) (List.combine ?l _) ] => apply nth_error_In with (n:=n') + | [ H : context[length ?ls] |- _ ] => tryif is_var ls then fail else (progress autorewrite with distr_length in H) + | [ H : context[List.nth_error (List.seq _ _) _] |- _ ] => rewrite nth_error_seq in H + end + | progress inversion_option + | progress intros + | progress cbn [fst snd value] in * + | progress destruct_head'_prod + | progress destruct_head'_ex + | progress destruct_head'_and + | progress destruct_head' False + | progress specialize_by_assumption + | apply conj + | progress expr.invert_subst + | progress expr.inversion_wf_constr + | progress expr.inversion_expr + | handle_Forall2_step + | wf_safe_t_step + | progress destruct_head' (@partial.wf_value) + | solve [ eapply wf_annotate; wf_t; try apply DefaultValue.expr.base.wf_default + | eapply wf_annotate_base; wf_t + | eapply (abstract_interp_ident_Proper _ (@ident.List_nth_default T) _ eq_refl); assumption + | eapply wf_update_annotation; wf_t + | wf_t + | match goal with + | [ H : context[UnderLets.wf _ _ _ _] |- UnderLets.wf _ _ _ _ ] => eapply H; solve [ repeat esplit; eauto ] + end + | eauto using List.nth_error_In + | eapply expr.wf_Proper_list; [ | eassumption ]; wf_safe_t; eauto 10 ] + | break_innermost_match_step + | match goal with + | [ H : context[List.In] |- expr.wf _ ?x ?y ] + => specialize (H x y); rewrite !List.nth_default_eq, <- List.combine_nth, <- !List.nth_default_eq in H; cbv [List.nth_default] in H |- * + | [ H : List.In _ _ -> ?P |- ?P ] => apply H + end + | break_innermost_match_hyps_step + | congruence + | rewrite List.combine_length in * + | rewrite Nat.min_r in * by lia + | rewrite Nat.min_l in * by lia + | progress expr.inversion_wf_one_constr + | progress expr.invert_match + | match goal with + | [ |- wf_value _ _ _ ] => progress hnf + end ]. } + Qed. + + Lemma wf_interp_ident_not_nth_default_nostrip (annotate_with_state : bool) G {t} (idc : ident t) + : wf_value_with_lets G (Base (reflect1 annotate_with_state (###idc)%expr (abstract_interp_ident _ idc))) (Base (reflect2 annotate_with_state (###idc)%expr (abstract_interp_ident _ idc))). + Proof using abstract_interp_ident_Proper annotate_expr_Proper bottom'_Proper extract_list_state_length extract_list_state_rel extract_option_state_rel is_annotated_for_Proper. + constructor; eapply wf_reflect; + solve [ apply bottom'_Proper + | apply wf_annotate + | repeat constructor + | apply abstract_interp_ident_Proper; reflexivity ]. + Qed. + + Lemma wf_interp_ident_not_nth_default (annotate_with_state : bool) G {t} (idc : ident t) + : (wf_value_with_lets G) + (Base match strip_annotation1 _ idc with + | Some v => v + | None => reflect1 annotate_with_state (###idc)%expr (abstract_interp_ident _ idc) + end) + (Base match strip_annotation2 _ idc with + | Some v => v + | None => reflect2 annotate_with_state (###idc)%expr (abstract_interp_ident _ idc) + end). + Proof using abstract_interp_ident_Proper annotate_expr_Proper bottom'_Proper extract_list_state_length extract_list_state_rel extract_option_state_rel is_annotated_for_Proper wf_strip_annotation. + pose proof (wf_strip_annotation G _ idc). + break_innermost_match; cbn [option_eq] in *; + solve [ exfalso; assumption + | congruence + | apply wf_interp_ident_not_nth_default_nostrip + | constructor; assumption ]. + Qed. + + Lemma wf_interp_ident G {t} idc1 idc2 (Hidc : idc1 = idc2) (annotate_with_state : bool) + : wf_value_with_lets G (@interp_ident1 annotate_with_state t idc1) (@interp_ident2 annotate_with_state t idc2). + Proof using abstract_interp_ident_Proper annotate_expr_Proper bottom'_Proper extract_list_state_length extract_list_state_rel extract_option_state_rel is_annotated_for_Proper wf_strip_annotation. + cbv [wf_value_with_lets ident.interp_ident]; subst idc2; destruct idc1; + first [ apply wf_interp_ident_nth_default + | apply wf_interp_ident_not_nth_default ]. + Qed. + + Local Notation strip_all_annotations'1 := (@partial.ident.strip_all_annotations' var1 annotation_to_cast1 skip_annotations_under). + Local Notation strip_all_annotations'2 := (@partial.ident.strip_all_annotations' var2 annotation_to_cast2 skip_annotations_under). + Local Notation strip_all_annotations1 := (@partial.ident.strip_all_annotations var1 annotation_to_cast1 skip_annotations_under). + Local Notation strip_all_annotations2 := (@partial.ident.strip_all_annotations var2 annotation_to_cast2 skip_annotations_under). + Lemma wf_strip_all_annotations' G t e1 e2 (Hwf : expr.wf G e1 e2) + : forall should_strip, expr.wf G (@strip_all_annotations'1 should_strip t e1) (@strip_all_annotations'2 should_strip t e2). + Proof using wf_annotation_to_cast. + induction Hwf; cbn [partial.ident.strip_all_annotations']; + repeat first [ progress cbn [projT1 projT2 fst snd] in * + | progress type.inversion_type + | progress wf_safe_t + | progress break_innermost_match + | progress expr.invert_subst + | discriminate + | solve [ exfalso; eauto + | eassert (Some _ = None) by eauto; congruence ] + | match goal with + | [ H1 : annotation_to_cast1 ?s ?d ?e1 = _, H2 : annotation_to_cast2 ?s ?d ?e2 = _ |- _ ] + => let H := fresh in + pose proof (fun G => wf_annotation_to_cast s d G e1 e2) as H; + rewrite H1, H2 in H; cbv [option_eq] in H; + clear H1 H2 + end ]. + Qed. + + Lemma wf_strip_all_annotations G t e1 e2 (Hwf : expr.wf G e1 e2) + : expr.wf G (@strip_all_annotations1 t e1) (@strip_all_annotations2 t e2). + Proof using wf_annotation_to_cast. now apply wf_strip_all_annotations'. Qed. + + Local Notation eval_with_bound1 := (@partial.ident.eval_with_bound var1 abstract_domain' annotate_expr1 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for1 strip_annotation1 skip_annotations_under). + Local Notation eval_with_bound2 := (@partial.ident.eval_with_bound var2 abstract_domain' annotate_expr2 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for2 strip_annotation2 skip_annotations_under). + Lemma wf_eval_with_bound (annotate_with_state : bool) {t} G G' e1 e2 (Hwf : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) + : expr.wf G' (@eval_with_bound1 annotate_with_state t e1 st1) (@eval_with_bound2 annotate_with_state t e2 st2). + Proof using abstract_interp_ident_Proper annotate_expr_Proper bottom'_Proper extract_list_state_length extract_list_state_rel extract_option_state_rel is_annotated_for_Proper wf_strip_annotation. + eapply wf_eval_with_bound'; + solve [ eassumption + | eapply wf_annotate + | eapply wf_interp_ident ]. + Qed. + + Local Notation eval1 := (@partial.ident.eval var1 abstract_domain' annotate_expr1 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for1 strip_annotation1). + Local Notation eval2 := (@partial.ident.eval var2 abstract_domain' annotate_expr2 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for2 strip_annotation2). + Lemma wf_eval {t} G G' e1 e2 (Hwf : expr.wf G e1 e2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) + : expr.wf G' (@eval1 t e1) (@eval2 t e2). + Proof using abstract_interp_ident_Proper annotate_expr_Proper bottom'_Proper extract_list_state_length extract_list_state_rel extract_option_state_rel is_annotated_for_Proper wf_strip_annotation. + eapply wf_eval'; + solve [ eassumption + | eapply wf_annotate + | eapply wf_interp_ident ]. + Qed. + + Local Notation eta_expand_with_bound1 := (@partial.ident.eta_expand_with_bound var1 abstract_domain' annotate_expr1 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for1). + Local Notation eta_expand_with_bound2 := (@partial.ident.eta_expand_with_bound var2 abstract_domain' annotate_expr2 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for2). + Lemma wf_eta_expand_with_bound {t} G e1 e2 (Hwf : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) + : expr.wf G (@eta_expand_with_bound1 t e1 st1) (@eta_expand_with_bound2 t e2 st2). + Proof using abstract_interp_ident_Proper annotate_expr_Proper bottom'_Proper extract_list_state_length extract_list_state_rel extract_option_state_rel is_annotated_for_Proper. + eapply wf_eta_expand_with_bound'; + solve [ eassumption + | eapply wf_annotate + | eapply wf_interp_ident ]. + Qed. + End with_var2. + End with_type. + End ident. + + Section specialized. + Import API. + Local Notation abstract_domain' := ZRange.type.base.option.interp (only parsing). + Local Notation abstract_domain := (@partial.abstract_domain base.type abstract_domain'). + Local Notation abstract_domain'_R t := (@eq (abstract_domain' t)) (only parsing). + Local Notation abstract_domain_R := (@abstract_domain_R base.type abstract_domain' (fun t => abstract_domain'_R t)). + Local Notation wf_value := (@wf_value base.type ident abstract_domain' (fun t => abstract_domain'_R t)). + + Lemma annotate_expr_Proper {relax_zrange var1 var2} {t} s1 s2 + : abstract_domain'_R t s1 s2 + -> option_eq (expr.wf nil) + (@annotate_expr relax_zrange var1 t s1) + (@annotate_expr relax_zrange var2 t s2). + Proof using Type. + intros ?; subst s2. + cbv [annotate_expr Crypto.Util.Option.bind option_eq]; break_innermost_match; + repeat constructor. + Qed. + + Global Instance bottom'_Proper {t} : Proper (abstract_domain'_R t) (bottom' t). + Proof using Type. reflexivity. Qed. + + Global Instance abstract_interp_ident_Proper {opts : AbstractInterpretation.Options} {assume_cast_truncates : bool} {t} + : Proper (eq ==> @abstract_domain_R t) (abstract_interp_ident assume_cast_truncates t). + Proof using Type. apply ZRange.ident.option.interp_Proper. Qed. + + Global Instance extract_list_state_Proper {t} + : Proper (abstract_domain'_R _ ==> option_eq (SetoidList.eqlistA (@abstract_domain'_R t))) + (extract_list_state t). + Proof using Type. + intros st st' ?; subst st'; cbv [option_eq extract_list_state]; break_innermost_match; reflexivity. + Qed. + + Local Notation tZ := (base.type.type_base base.type.Z). + Local Notation cstZ r + := (expr.App + (d:=type.arrow (type.base tZ) (type.base tZ)) + (expr.Ident ident.Z_cast) + (expr.Ident (@ident.Literal base.type.zrange r%zrange))). + Local Notation cstZZ r1 r2 + := (expr.App + (d:=type.arrow (type.base (tZ * tZ)) (type.base (tZ * tZ))) + (expr.Ident ident.Z_cast2) + (#(@ident.Literal base.type.zrange r1%zrange), #(@ident.Literal base.type.zrange r2%zrange))%expr_pat). + + Lemma wf_always_strip_annotation {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {var1 var2} G {t} idc + : option_eq (wf_value G) + (always_strip_annotation assume_cast_truncates (var:=var1) t idc) + (always_strip_annotation assume_cast_truncates (var:=var2) t idc). + Proof using Type. + cbv [always_strip_annotation]; break_innermost_match; try reflexivity; + cbn [option_eq wf_value]. + all: repeat first [ progress intros + | assumption + | progress subst + | exfalso; congruence + | progress cbn [fst snd] in * + | progress break_innermost_match + | progress destruct_head'_and + | match goal with + | [ |- UnderLets.wf _ _ (Base _) (Base _) ] => constructor + | [ |- _ /\ _ ] => split + end + | solve [ wf_t ] + | progress cbn [abstract_domain_R type.related] in * ]. + Qed. + + Lemma wf_strip_annotation {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) (strip_annotations : bool) {var1 var2} G {t} idc + : option_eq (wf_value G) + (strip_annotation assume_cast_truncates strip_annotations (var:=var1) t idc) + (strip_annotation assume_cast_truncates strip_annotations (var:=var2) t idc). + Proof using Type. + cbv [strip_annotation]; break_innermost_match; now try apply wf_always_strip_annotation. + Qed. + + Local Arguments base.try_make_transport_cps / _ _ _. + Local Arguments type.try_make_transport_cps / _ _ _ _ _. + Lemma is_annotated_for_spec {relax_zrange var} t t' e st + : @is_annotated_for relax_zrange var t t' e st = true + <-> ((exists (pf : t' = tZ) r, + (annotation_of_state relax_zrange (rew pf in st) = Some r) + /\ existT (@expr var) t e = existT (@expr var) _ (cstZ r)) + \/ (exists (pf : t' = (tZ * tZ)%etype) r1 r2, + ((fun '(r1, r2) => (annotation_of_state relax_zrange r1, annotation_of_state relax_zrange r2)) + (rew pf in st) = (Some r1, Some r2)) + /\ existT (@expr var) t e = existT (@expr var) _ (cstZZ r1 r2))). + Proof using Type. + split. + { cbv [is_annotated_for]; break_innermost_match; try discriminate. + all: rewrite ?Bool.andb_true_iff. + all: intro; destruct_head'_and; reflect_beq_to_eq (option_beq zrange_beq). + all: constructor; exists eq_refl; cbn [eq_rect]. + all: repeat esplit; try apply (f_equal2 (@pair _ _)); try (symmetry; eassumption). + all: multimatch goal with + | [ H : invert_Z_cast _ = Some _ |- _ ] + => apply expr.invert_Z_cast_Some in H + | [ H : invert_Z_cast2 _ = Some _ |- _ ] + => apply expr.invert_Z_cast2_Some in H + end; + solve [ repeat first [ progress inversion_sigma + | progress subst + | progress cbn [eq_rect] in * + | reflexivity ] ]. } + { repeat first [ progress intros + | progress subst + | progress cbn [eq_rect fst snd] in * + | progress inversion_option + | progress inversion_sigma + | progress inversion_prod + | progress destruct_head'_ex + | progress destruct_head'_and + | progress destruct_head'_or + | progress cbn + | break_innermost_match_step + | rewrite Bool.andb_true_iff + | apply conj + | apply zrange_lb + | reflexivity ]. } + Qed. + + Lemma is_annotated_for_Proper {relax_zrange var1 var2} G t t' e1 e2 + : expr.wf G e1 e2 + -> ((abstract_domain'_R _ ==> eq)%signature) + (@is_annotated_for relax_zrange var1 t t' e1) + (@is_annotated_for relax_zrange var2 t t' e2). + Proof using Type. + repeat intro; subst; + match goal with |- ?x = ?y => destruct x eqn:?, y eqn:? end. + all: repeat first [ match goal with + | [ H : is_annotated_for _ _ _ _ _ = true |- _ ] + => rewrite is_annotated_for_spec in H + | [ H : ?x <> ?x |- _ ] => congruence + end + | reflexivity + | exfalso; assumption + | progress cbn [eq_rect fst snd projT1 projT2 andb] in * + | progress subst + | progress inversion_prod + | progress destruct_head'_ex + | progress destruct_head'_sig + | progress destruct_head'_sigT + | progress destruct_head'_prod + | progress destruct_head'_and + | progress reflect_beq_to_eq zrange_beq + | progress inversion_sigma + | progress inversion_option + | progress destruct_head'_or + | progress inversion_type + | progress expr.inversion_wf_one_constr + | progress expr.invert_subst + | progress expr.invert_match + | progress expr.inversion_expr + | break_innermost_match_hyps_step + | match goal with + | [ H : is_annotated_for _ _ _ _ _ = false |- _ ] + => progress cbn in H + end ]. + Qed. + + Lemma wf_annotation_to_cast_helper {var1 var2 t G idc1 idc2} + (Hidc : idc1 = idc2) + : option_eq (fun f g => forall e1 e2, expr.wf G e1 e2 -> expr.wf G (f e1) (g e2)) + (@annotation_to_cast_helper var1 t idc1) + (@annotation_to_cast_helper var2 t idc2). + Proof using Type. + subst idc2; destruct idc1; cbv; try reflexivity; intros; assumption. + Qed. + + Lemma wf_annotation_to_cast {var1 var2 s d G e1 e2} + (Hwf : expr.wf G e1 e2) + : option_eq (fun f g => forall e1 e2, expr.wf G e1 e2 -> expr.wf G (f e1) (g e2)) + (@annotation_to_cast var1 s d e1) + (@annotation_to_cast var2 s d e2). + Proof using Type. + cbv [annotation_to_cast]; + repeat first [ progress wf_safe_t + | congruence + | progress inversion_option + | progress cbn [fst snd projT1 projT2] in * + | progress cbv [Option.bind] in * + | progress break_innermost_match + | progress break_innermost_match_hyps + | progress expr.invert_subst + | progress expr.inversion_wf_one_constr + | progress cbn [invert_AppIdent invert_App invert_Ident invert_App_cps invert_AppIdent_cps Option.bind] in * + | progress expr.invert_match + | progress expr.inversion_expr + | now refine (wf_annotation_to_cast_helper eq_refl) ]. + Qed. + + Lemma extract_list_state_length + : forall t v1 v2, abstract_domain'_R _ v1 v2 -> option_map (@length _) (extract_list_state t v1) = option_map (@length _) (extract_list_state t v2). + Proof using Type. + intros; subst; cbv [option_map extract_list_state]; break_innermost_match; reflexivity. + Qed. + Lemma extract_list_state_rel + : forall t v1 v2, abstract_domain'_R _ v1 v2 -> forall l1 l2, extract_list_state t v1 = Some l1 -> extract_list_state t v2 = Some l2 -> List.Forall2 (@abstract_domain'_R t) l1 l2. + Proof using Type. + intros; cbv [extract_list_state] in *; subst; inversion_option; subst. + now rewrite Forall2_Forall, Forall_forall; cbv [Proper]. + Qed. + + Lemma extract_option_state_rel + : forall t v1 v2, abstract_domain'_R _ v1 v2 -> option_eq (option_eq (abstract_domain'_R _)) (extract_option_state t v1) (extract_option_state t v2). + Proof using Type. + cbv [extract_option_state option_eq]; intros; subst; break_match; reflexivity. + Qed. + + Section with_var2. + Context {var1 var2 : type -> Type}. + Local Notation wf_value_with_lets := (@wf_value_with_lets base.type ident abstract_domain' (fun t => abstract_domain'_R t) var1 var2). + + Lemma wf_eval {opts : AbstractInterpretation.Options} {t} G G' e1 e2 (Hwf : expr.wf G e1 e2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) + : expr.wf G' (t:=t) (eval (var:=var1) e1) (eval (var:=var2) e2). + Proof using Type. + eapply ident.wf_eval; + solve [ eassumption + | exact _ + | apply extract_list_state_length + | apply extract_list_state_rel + | apply extract_option_state_rel + | apply wf_strip_annotation + | intros; now apply annotate_expr_Proper + | apply is_annotated_for_Proper ]. + Qed. + + Lemma wf_strip_all_annotations strip_annotations_under G t e1 e2 (Hwf : expr.wf G e1 e2) + : expr.wf G (@strip_all_annotations strip_annotations_under var1 t e1) (@strip_all_annotations strip_annotations_under var2 t e2). + Proof using Type. revert Hwf; apply ident.wf_strip_all_annotations, @wf_annotation_to_cast. Qed. + + Lemma wf_eval_with_bound {opts : AbstractInterpretation.Options} {relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations t} G G' e1 e2 (Hwf : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) + : expr.wf G' (var1:=var1) (var2:=var2) (t:=t) + (eval_with_bound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e1 st1) + (eval_with_bound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e2 st2). + Proof using Type. + eapply ident.wf_eval_with_bound; + solve [ eassumption + | exact _ + | apply extract_list_state_length + | apply extract_list_state_rel + | apply extract_option_state_rel + | apply wf_strip_annotation + | intros; now apply annotate_expr_Proper + | apply is_annotated_for_Proper ]. + Qed. + + Lemma wf_strip_annotations {opts : AbstractInterpretation.Options} {assume_cast_truncates} {t} G G' e1 e2 (Hwf : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) + : expr.wf G' (t:=t) (strip_annotations assume_cast_truncates (var:=var1) e1 st1) (strip_annotations assume_cast_truncates (var:=var2) e2 st2). + Proof using Type. + eapply ident.wf_eval_with_bound; + solve [ eassumption + | exact _ + | apply extract_list_state_length + | apply extract_list_state_rel + | apply extract_option_state_rel + | apply wf_strip_annotation + | intros; now apply annotate_expr_Proper + | apply is_annotated_for_Proper ]. + Qed. + + Lemma wf_eta_expand_with_bound {relax_zrange t} G e1 e2 (Hwf : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) + : expr.wf G (t:=t) (eta_expand_with_bound relax_zrange (var:=var1) e1 st1) (eta_expand_with_bound relax_zrange (var:=var2) e2 st2). + Proof using Type. + eapply ident.wf_eta_expand_with_bound; + solve [ eassumption + | exact _ + | apply extract_list_state_length + | apply extract_list_state_rel + | apply extract_option_state_rel + | intros; now apply annotate_expr_Proper + | apply is_annotated_for_Proper ]. + Qed. + End with_var2. + + Lemma Wf_StripAllAnnotations strip_annotations_under {t} (e : Expr t) (Hwf : Wf e) : Wf (StripAllAnnotations strip_annotations_under e). + Proof using Type. + intros ??; now apply wf_strip_all_annotations. + Qed. + + Lemma Wf_Eval {opts : AbstractInterpretation.Options} {t} (e : Expr t) (Hwf : Wf e) : Wf (Eval e). + Proof using Type. + intros ??; eapply wf_eval with (G:=nil); cbn [List.In]; try apply Hwf; tauto. + Qed. + + Lemma Wf_EvalWithBound {opts : AbstractInterpretation.Options} {relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations t} (e : Expr t) bound (Hwf : Wf e) (bound_valid : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) bound) + : Wf (EvalWithBound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e bound). + Proof using Type. + intros ??; eapply wf_eval_with_bound with (G:=nil); cbn [List.In]; try apply Hwf; tauto. + Qed. + + Lemma Wf_StripAnnotations {opts : AbstractInterpretation.Options} {assume_cast_truncates} {t} (e : Expr t) bound (Hwf : Wf e) (bound_valid : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) bound) + : Wf (StripAnnotations assume_cast_truncates e bound). + Proof using Type. + intros ??; eapply wf_strip_annotations with (G:=nil); cbn [List.In]; try tauto; apply Hwf. + Qed. + + Lemma Wf_EtaExpandWithBound {relax_zrange t} (e : Expr t) bound (Hwf : Wf e) (bound_valid : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) bound) + : Wf (EtaExpandWithBound relax_zrange e bound). + Proof using Type. + intros ??; eapply wf_eta_expand_with_bound with (G:=nil); cbn [List.In]; try apply Hwf; tauto. + Qed. + + Local Instance Proper_strip_ranges {t} + : Proper (@abstract_domain_R t ==> @abstract_domain_R t) (@ZRange.type.option.strip_ranges t). + Proof using Type. + cbv [Proper abstract_domain_R respectful]. + induction t as [t|s IHs d IHd]; cbn in *; destruct_head'_prod; destruct_head'_and; cbn in *; intros; subst; cbv [respectful] in *; + eauto. + Qed. + + Lemma Wf_EtaExpandWithListInfoFromBound {t} (e : Expr t) bound (Hwf : Wf e) (bound_valid : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) bound) + : Wf (EtaExpandWithListInfoFromBound e bound). + Proof using Type. + eapply Wf_EtaExpandWithBound; [ assumption | ]. + clear dependent e. + cbv [Proper] in *; induction t as [t|s IHs d IHd]; cbn in *; destruct_head'_prod; destruct_head'_and; cbn in *; eauto. + split; auto; apply Proper_strip_ranges; auto. + Qed. + End specialized. + End partial. +#[global] + Hint Resolve Wf_Eval Wf_EvalWithBound Wf_EtaExpandWithBound Wf_EtaExpandWithListInfoFromBound Wf_StripAllAnnotations Wf_StripAnnotations : wf. +#[global] + Hint Opaque partial.Eval partial.EvalWithBound partial.EtaExpandWithBound partial.EtaExpandWithListInfoFromBound partial.StripAnnotations partial.StripAllAnnotations : wf interp rewrite. + Import API. + + Lemma Wf_PartialEvaluateWithListInfoFromBounds + {opts : AbstractInterpretation.Options} + {t} (E : Expr t) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + (Hwf : Wf E) + {b_in_Proper : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R base.type ZRange.type.base.option.interp (fun t0 : base.type => eq))) b_in} + : Wf (PartialEvaluateWithListInfoFromBounds E b_in). + Proof. cbv [PartialEvaluateWithListInfoFromBounds]; eauto with wf. Qed. +#[global] + Hint Resolve Wf_PartialEvaluateWithListInfoFromBounds : wf. +#[global] + Hint Opaque PartialEvaluateWithListInfoFromBounds : wf interp rewrite. + + Lemma Wf_PartialEvaluateWithBounds + {opts : AbstractInterpretation.Options} + {relax_zrange} {assume_cast_truncates : bool} {skip_annotations_under : forall t, ident t -> bool} {strip_preexisting_annotations : bool} {t} (E : Expr t) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + (Hwf : Wf E) + {b_in_Proper : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R base.type ZRange.type.base.option.interp (fun t0 : base.type => eq))) b_in} + : Wf (PartialEvaluateWithBounds relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations E b_in). + Proof. cbv [PartialEvaluateWithBounds]; eauto with wf. Qed. +#[global] + Hint Resolve Wf_PartialEvaluateWithBounds : wf. +#[global] + Hint Opaque PartialEvaluateWithBounds : wf interp rewrite. +End Compilers. diff --git a/src/AbstractInterpretation/Bottomify/WfExtra.v b/src/AbstractInterpretation/Bottomify/WfExtra.v new file mode 100644 index 0000000000..33dcd9a0ea --- /dev/null +++ b/src/AbstractInterpretation/Bottomify/WfExtra.v @@ -0,0 +1,35 @@ +Require Import Rewriter.Language.Language. +Require Import Rewriter.Language.Wf. +Require Import Crypto.Language.API. +Require Import Crypto.Language.WfExtra. +Require Import Crypto.AbstractInterpretation.Bottomify.AbstractInterpretation. +Require Import Crypto.AbstractInterpretation.Bottomify.Wf. + +Module Compilers. + Import Language.Compilers. + Import Language.Inversion.Compilers. + Import Language.Wf.Compilers. + Import Language.API.Compilers. + Import Language.WfExtra.Compilers. + Import AbstractInterpretation.Bottomify.AbstractInterpretation.Compilers. + Import AbstractInterpretation.Bottomify.Wf.Compilers. + Import Compilers.API. + + Import AbstractInterpretation.Bottomify.AbstractInterpretation.Compilers.partial. + Import AbstractInterpretation.Bottomify.Wf.Compilers.partial. + +#[global] + Hint Resolve Wf_Eval Wf_EvalWithBound Wf_EtaExpandWithBound Wf_EtaExpandWithListInfoFromBound Wf_StripAnnotations Wf_StripAllAnnotations : wf_extra. +#[global] + Hint Opaque partial.Eval EvalWithBound EtaExpandWithBound EtaExpandWithListInfoFromBound StripAnnotations StripAllAnnotations : wf_extra interp_extra. + +#[global] + Hint Resolve Wf_PartialEvaluateWithListInfoFromBounds : wf_extra. +#[global] + Hint Opaque PartialEvaluateWithListInfoFromBounds : wf_extra interp_extra. + +#[global] + Hint Resolve Wf_PartialEvaluateWithBounds : wf_extra. +#[global] + Hint Opaque PartialEvaluateWithBounds : wf_extra interp_extra. +End Compilers. diff --git a/src/AbstractInterpretation/Fancy/AbstractInterpretation.v b/src/AbstractInterpretation/Fancy/AbstractInterpretation.v new file mode 100644 index 0000000000..c43ead4fc9 --- /dev/null +++ b/src/AbstractInterpretation/Fancy/AbstractInterpretation.v @@ -0,0 +1,593 @@ +Require Import Coq.ZArith.ZArith. +Require Import Crypto.Util.ListUtil Coq.Lists.List Crypto.Util.ListUtil.FoldBool. +Require Import Crypto.Util.ZRange. +Require Import Crypto.Util.ZRange.Operations. +Require Import Crypto.Util.Option. +Require Import Crypto.Util.OptionList. +Require Import Crypto.Util.ZUtil.Tactics.LtbToLt. +Require Import Crypto.Util.Bool.Reflect. +Require Import Crypto.Util.LetIn. +Require Import Rewriter.Language.Language. +Require Import Crypto.Language.InversionExtra. +Require Import Crypto.Language.API. +Require Import Crypto.AbstractInterpretation.ZRange. +Require Import Rewriter.Language.UnderLets. +Import ListNotations. Local Open Scope bool_scope. Local Open Scope Z_scope. + +Module Compilers. + Export Language.Compilers. + Export UnderLets.Compilers. + Export Language.API.Compilers. + Export AbstractInterpretation.ZRange.Compilers. + Import invert_expr. + Import Language.InversionExtra.Compilers. + + (** XXX TODO: Do we still need to do UnderLets here? *) + Module partial. + Import UnderLets. + Section with_var. + Context {base_type : Type}. + Local Notation type := (type base_type). + Let type_base (x : base_type) : type := type.base x. + Local Coercion type_base : base_type >-> type. + Context {ident : type -> Type} + {var : type -> Type}. + Local Notation expr := (@expr base_type ident). + Local Notation UnderLets := (@UnderLets base_type ident var). + Context (abstract_domain' : base_type -> Type) + (annotate : forall (is_let_bound : bool) t, abstract_domain' t -> @expr var t -> UnderLets (@expr var t)) + (bottom' : forall A, abstract_domain' A) + (skip_annotations_under : forall t, ident t -> bool). + + Definition abstract_domain (t : type) + := type.interp abstract_domain' t. + + Fixpoint value (t : type) + := match t return Type (* COQBUG(https://github.com/coq/coq/issues/7727) *) with + | type.base t + => abstract_domain t * @expr var t + | type.arrow s d + => value s -> UnderLets (value d) + end%type. + + Definition value_with_lets (t : type) + := UnderLets (value t). + + Context (interp_ident : bool (* annotate with state? *) -> forall t, ident t -> value_with_lets t). + + Fixpoint bottom {t} : abstract_domain t + := match t with + | type.base t => bottom' t + | type.arrow s d => fun _ => @bottom d + end. + + Fixpoint bottom_for_each_lhs_of_arrow {t} : type.for_each_lhs_of_arrow abstract_domain t + := match t return type.for_each_lhs_of_arrow abstract_domain t with + | type.base t => tt + | type.arrow s d => (bottom, @bottom_for_each_lhs_of_arrow d) + end. + + Definition state_of_value {t} : value t -> abstract_domain t + := match t return value t -> abstract_domain t with + | type.base t => fun '(st, v) => st + | type.arrow s d => fun _ => bottom + end. + + (** We need to make sure that we ignore the state of + higher-order arrows *everywhere*, or else the proofs don't go + through. So we sometimes need to replace the state of + arrow-typed values with [⊥]. *) + Fixpoint bottomify {t} : value t -> value_with_lets t + := match t return value t -> value_with_lets t with + | type.base t => fun '(st, v) => Base (bottom' t, v) + | type.arrow s d => fun f => Base (fun x => fx <-- f x; @bottomify d fx) + end%under_lets. + + (** We drop the state of higher-order arrows *) + Fixpoint reify (annotate_with_state : bool) (is_let_bound : bool) {t} : value t -> type.for_each_lhs_of_arrow abstract_domain t -> UnderLets (@expr var t) + := match t return value t -> type.for_each_lhs_of_arrow abstract_domain t -> UnderLets (@expr var t) with + | type.base t + => fun '(st, v) 'tt + => if annotate_with_state + then annotate is_let_bound t st v + else if is_let_bound + then UnderLets.UnderLet v (fun v' => UnderLets.Base ($$v')) + else UnderLets.Base v + | type.arrow s d + => fun f_e '(sv, dv) + => let sv := match s with + | type.base _ => sv + | type.arrow _ _ => bottom + end in + Base + (λ x , (UnderLets.to_expr + (fx <-- f_e (@reflect annotate_with_state _ (expr.Var x) sv); + @reify annotate_with_state false _ fx dv))) + end%core%expr + with reflect (annotate_with_state : bool) {t} : @expr var t -> abstract_domain t -> value t + := match t return @expr var t -> abstract_domain t -> value t with + | type.base t + => fun e st => (st, e) + | type.arrow s d + => fun e absf + => (fun v + => let stv := state_of_value v in + (rv <-- (@reify annotate_with_state false s v bottom_for_each_lhs_of_arrow); + Base (@reflect annotate_with_state d (e @ rv) (absf stv))%expr)) + end%under_lets. + + Definition skip_annotations_for_App {var'} {t} (e : @expr var' t) : bool + := match invert_AppIdent_curried e with + | Some (existT _ (idc, args)) => skip_annotations_under _ idc + | None => false + end. + + Fixpoint interp (annotate_with_state : bool) {t} (e : @expr value_with_lets t) : value_with_lets t + := let annotate_with_state := annotate_with_state && negb (skip_annotations_for_App e) in + match e in expr.expr t return value_with_lets t with + | expr.Ident t idc => interp_ident annotate_with_state _ idc (* Base (reflect (###idc) (abstract_interp_ident _ idc))*) + | expr.Var t v => v + | expr.Abs s d f => Base (fun x => @interp annotate_with_state d (f (Base x))) + | expr.App (type.base s) d f x + => (x' <-- @interp annotate_with_state _ x; + f' <-- @interp annotate_with_state (_ -> d)%etype f; + f' x') + | expr.App (type.arrow s' d') d f x + => (x' <-- @interp annotate_with_state (s' -> d')%etype x; + x'' <-- bottomify x'; + f' <-- @interp annotate_with_state (_ -> d)%etype f; + f' x'') + | expr.LetIn (type.arrow _ _) B x f + => (x' <-- @interp annotate_with_state _ x; + @interp annotate_with_state _ (f (Base x'))) + | expr.LetIn (type.base A) B x f + => (x' <-- @interp annotate_with_state _ x; + x'' <-- reify annotate_with_state true (* this forces a let-binder here *) x' tt; + @interp annotate_with_state _ (f (Base (reflect annotate_with_state x'' (state_of_value x'))))) + end%under_lets. + + Definition eval_with_bound' (annotate_with_state : bool) {t} (e : @expr value_with_lets t) + (st : type.for_each_lhs_of_arrow abstract_domain t) + : expr t + := UnderLets.to_expr (e' <-- interp annotate_with_state e; reify annotate_with_state false e' st). + + Definition eval' {t} (e : @expr value_with_lets t) : expr t + := eval_with_bound' false e bottom_for_each_lhs_of_arrow. + + Definition eta_expand_with_bound' {t} (e : @expr var t) + (st : type.for_each_lhs_of_arrow abstract_domain t) + : expr t + := UnderLets.to_expr (reify true false (reflect true e bottom) st). + + Section extract. + Context (ident_extract : forall t, ident t -> abstract_domain t). + + (** like [expr.interp (@ident_extract) e], except we replace + all higher-order state with bottom *) + Fixpoint extract' {t} (e : @expr abstract_domain t) : abstract_domain t + := match e in expr.expr t return abstract_domain t with + | expr.Ident t idc => ident_extract t idc + | expr.Var t v => v + | expr.Abs s d f => fun v : abstract_domain s => @extract' _ (f v) + | expr.App (type.base s) d f x + => @extract' _ f (@extract' _ x) + | expr.App (type.arrow s' d') d f x + => @extract' _ f (@bottom (type.arrow s' d')) + | expr.LetIn A B x f => dlet y := @extract' _ x in @extract' _ (f y) + end. + + Definition extract_gen {t} (e : @expr abstract_domain t) (bound : type.for_each_lhs_of_arrow abstract_domain t) + : abstract_domain' (type.final_codomain t) + := type.app_curried (extract' e) bound. + End extract. + End with_var. + + Module ident. + Section with_var. + Local Notation type := (type base.type). + Let type_base (x : base.type.base) : base.type := base.type.type_base x. + Let base {bt} (x : Language.Compilers.base.type bt) : type.type _ := type.base x. + Local Coercion base : base.type >-> type.type. + Local Coercion type_base : base.type.base >-> base.type. + Context {var : type -> Type}. + Local Notation expr := (@expr base.type ident). + Local Notation UnderLets := (@UnderLets base.type ident var). + Context (abstract_domain' : base.type -> Type). + Local Notation abstract_domain := (@abstract_domain base.type abstract_domain'). + Local Notation value_with_lets := (@value_with_lets base.type ident var abstract_domain'). + Local Notation value := (@value base.type ident var abstract_domain'). + Context (annotate_expr : forall t, abstract_domain' t -> option (@expr var (t -> t))) + (bottom' : forall A, abstract_domain' A) + (abstract_interp_ident : forall t, ident t -> type.interp abstract_domain' t) + (extract_list_state : forall A, abstract_domain' (base.type.list A) -> option (list (abstract_domain' A))) + (extract_option_state : forall A, abstract_domain' (base.type.option A) -> option (option (abstract_domain' A))) + (is_annotated_for : forall t t', @expr var t -> abstract_domain' t' -> bool) + (annotation_to_cast : forall s d, @expr var (s -> d) -> option (@expr var s -> @expr var d)) + (skip_annotations_under : forall t, ident t -> bool) + (strip_annotation : forall t, ident t -> option (value t)). + + (** TODO: Is it okay to commute annotations? *) + Definition update_annotation {t} (st : abstract_domain' t) (e : @expr var t) : @expr var t + := match e, annotate_expr _ st with + | (cst' @ e'), Some cst + => if is_annotated_for _ _ cst' st + then e + else cst @ e + | _, Some cst => cst @ e + | _, None => e + end%expr_pat%expr. + + Definition annotate_with_expr (is_let_bound : bool) {t} + (st : abstract_domain' t) (e : @expr var t) + : UnderLets (@expr var t) + := let cst_e := update_annotation st e (*match annotate_expr _ st with + | Some cst => ###cst @ e + | None => e + end%expr*) in + if is_let_bound + then UnderLet cst_e (fun v => Base ($$v)%expr) + else Base cst_e. + + Definition annotate_base (is_let_bound : bool) {t : base.type.base} + (st : abstract_domain' t) (e : @expr var t) + : UnderLets (@expr var t) + := annotate_with_expr is_let_bound st e. + + Fixpoint annotate (is_let_bound : bool) {t : base.type} : abstract_domain' t -> @expr var t -> UnderLets (@expr var t) + := match t return abstract_domain' t -> @expr var t -> UnderLets (@expr var t) with + | base.type.type_base t => annotate_base is_let_bound + | base.type.unit + => fun _ e + => (* we need to keep let-bound unit expressions around for comments *) + if is_let_bound + then UnderLet e (fun v => Base ($$v)%expr) + else Base e + | base.type.prod A B + => fun st e + => match invert_pair e with + | Some (x, y) + => let stx := abstract_interp_ident _ ident.fst st in + let sty := abstract_interp_ident _ ident.snd st in + (x' <-- @annotate is_let_bound A stx x; + y' <-- @annotate is_let_bound B sty y; + Base (x', y')%expr) + | None => annotate_with_expr is_let_bound st e + end + | base.type.list A + => fun st e + => match extract_list_state _ st, reflect_list e with + | Some ls_st, Some ls_e + => (retv <---- (List.map + (fun '(st', e') => @annotate is_let_bound A st' e') + (List.combine ls_st ls_e)); + Base (reify_list retv)) + | Some ls_st, None + => (retv <---- (List.map + (fun '(n, st') + => let e' := (#ident.List_nth_default @ DefaultValue.expr.base.default @ e @ ##(n:nat))%expr in + @annotate is_let_bound A st' e') + (List.combine (List.seq 0 (List.length ls_st)) ls_st)); + Base (reify_list retv)) + | None, _ => annotate_with_expr is_let_bound st e + end + | base.type.option A + => fun st e + => match extract_option_state _ st, reflect_option e with + | Some v_st, Some v_e + => (retv <----- (Option.map + (fun '(st', e') => @annotate is_let_bound A st' e') + (Option.combine v_st v_e)); + Base (reify_option retv)) + | Some _, None + | None, _ + => annotate_with_expr is_let_bound st e + end + end%under_lets. + + Local Notation reify := (@reify base.type ident var abstract_domain' annotate bottom'). + Local Notation reflect := (@reflect base.type ident var abstract_domain' annotate bottom'). + + (** We manually rewrite with the rule for [nth_default], as the eliminator for eta-expanding lists in the input *) + Definition interp_ident (annotate_with_state : bool) {t} (idc : ident t) : value_with_lets t + := match idc in ident t return value_with_lets t with + | ident.List_nth_default T as idc + => let default := reflect annotate_with_state (###idc) (abstract_interp_ident _ idc) in + Base + (fun default_arg + => default <-- default default_arg; + Base + (fun ls_arg + => default <-- default ls_arg; + Base + (fun n_arg + => default <-- default n_arg; + ls' <-- @reify annotate_with_state false (base.type.list T) ls_arg tt; + Base + (fst default, + match reflect_list ls', invert_Literal (snd n_arg) with + | Some ls, Some n + => nth_default (snd default_arg) ls n + | _, _ => snd default + end)))) + | idc => Base + match strip_annotation _ idc with + | Some v => v + | None => reflect annotate_with_state (###idc) (abstract_interp_ident _ idc) + end + end%core%under_lets%expr. + + Fixpoint strip_all_annotations' (should_strip : bool) {t} (e : @expr var t) : @expr var t + := match e in expr.expr t return expr t with + | expr.Ident _ _ as e + | expr.Var _ _ as e + => e + | expr.Abs s d f + => expr.Abs (fun x => strip_all_annotations' should_strip (f x)) + | expr.App s d f x + => let should_strip + := (should_strip || match invert_AppIdent_curried f with + | Some (existT _ (idc, _)) => skip_annotations_under _ idc + | None => false + end)%bool in + let f := strip_all_annotations' should_strip f in + let x := strip_all_annotations' should_strip x in + if should_strip + then match annotation_to_cast _ _ f with + | Some f => f x + | None => expr.App f x + end + else expr.App f x + | expr.LetIn A B x f + => expr.LetIn (strip_all_annotations' should_strip x) (fun v => strip_all_annotations' should_strip (f v)) + end. + Definition strip_all_annotations {t} (e : @expr var t) : @expr var t + := @strip_all_annotations' false t e. + + Definition eval_with_bound (skip_annotations_under : forall t, ident t -> bool) (annotate_with_state : bool) {t} (e : @expr value_with_lets t) + (st : type.for_each_lhs_of_arrow abstract_domain t) + : @expr var t + := @eval_with_bound' base.type ident var abstract_domain' annotate bottom' skip_annotations_under interp_ident annotate_with_state t e st. + + Definition eval {t} (e : @expr value_with_lets t) : @expr var t + := @eval' base.type ident var abstract_domain' annotate bottom' (fun _ _ => false) interp_ident t e. + + Definition eta_expand_with_bound {t} (e : @expr var t) + (st : type.for_each_lhs_of_arrow abstract_domain t) + : @expr var t + := @eta_expand_with_bound' base.type ident var abstract_domain' annotate bottom' t e st. + + Definition extract {t} (e : @expr _ t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : abstract_domain' (type.final_codomain t) + := @extract_gen base.type ident abstract_domain' bottom' abstract_interp_ident t e bound. + End with_var. + End ident. + + Definition default_relax_zrange (v : zrange) : option zrange := Some v. + + Section specialized. + Local Notation abstract_domain' := ZRange.type.base.option.interp. + Local Notation abstract_domain := (@partial.abstract_domain base.type abstract_domain'). + Local Notation value_with_lets var := (@value_with_lets base.type ident var abstract_domain'). + Local Notation value var := (@value base.type ident var abstract_domain'). + Notation expr := (@expr base.type ident). + Notation Expr := (@expr.Expr base.type ident). + Local Notation type := (type base.type). + Let type_base (x : base.type.base) : base.type := base.type.type_base x. + Let base {bt} (x : Language.Compilers.base.type bt) : type.type _ := type.base x. + Local Coercion base : base.type >-> type.type. + Local Coercion type_base : base.type.base >-> base.type. + Local Notation tZ := (base.type.type_base base.type.Z). + + Section with_relax. + Context (relax_zrange : zrange -> option zrange). + + Let always_relax_zrange : zrange -> zrange + := fun range => match relax_zrange (ZRange.normalize range) with + | Some r => r + | None => range + end. + + Definition annotation_of_state (st : abstract_domain' base.type.Z) : option zrange + := option_map always_relax_zrange st. + + Local Notation cstZ r + := (expr.App + (d:=type.arrow (type.base tZ) (type.base tZ)) + (expr.Ident ident.Z_cast) + (expr.Ident (@ident.Literal base.type.zrange r%zrange))). + Local Notation cstZZ r1 r2 + := (expr.App + (d:=type.arrow (type.base (tZ * tZ)) (type.base (tZ * tZ))) + (expr.Ident ident.Z_cast2) + (#(@ident.Literal base.type.zrange r1%zrange), #(@ident.Literal base.type.zrange r2%zrange))%expr_pat). + + Definition annotate_expr {var} t : abstract_domain' t -> option (@expr var (t -> t)) + := match t return abstract_domain' t -> option (expr (t -> t)) with + | tZ + => fun st => st' <- annotation_of_state st; Some (cstZ st') + | tZ * tZ + => fun '(sta, stb) => sta' <- annotation_of_state sta; stb' <- annotation_of_state stb; Some (cstZZ sta' stb') + | _ => fun _ => None + end%option%etype. + + Definition abstract_interp_ident {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) t (idc : ident t) : type.interp abstract_domain' t + := ZRange.ident.option.interp assume_cast_truncates idc. + + Definition always_strip_annotation {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {var} t (idc : ident t) : option (value var t) + := match idc in ident t return option (value var t) with + | ident.Z_cast as idc + | ident.Z_cast2 as idc + => let tZ_type := (fun t (idc : ident (type.base _ -> type.base t -> type.base t)) => t) _ idc in (* we want Coq to pick up the Z type for bounds tightness, not the zrange type (where "tighter" means "equal") *) + Some + (fun '(r_orig, re) + => Base + (fun '(r_known, e) + => Base + ((abstract_interp_ident assume_cast_truncates _ idc r_orig r_known, + let do_strip + := match ZRange.type.base.option.lift_Some r_known, ZRange.type.base.option.lift_Some r_orig with + | Some r_known, Some r_orig + => ZRange.type.base.is_tighter_than + (t:=tZ_type) + r_known r_orig + | _, _ => false + end in + if do_strip + then e + else (###idc @ re @ e)%expr)))) + | _ => None + end. + + Definition strip_annotation {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) (strip_annotations : bool) {var} : forall t, ident t -> option (value var t) + := if strip_annotations + then always_strip_annotation assume_cast_truncates + else fun _ _ => None. + + Definition is_annotated_for {var} t t' (idc : @expr var t) : abstract_domain' t' -> bool + := match invert_Z_cast idc, invert_Z_cast2 idc, t' with + | Some r, _, tZ + => fun r' + => option_beq zrange_beq (Some r) (annotation_of_state r') + | _, Some (r1, r2), base.type.prod (base.type.type_base base.type.Z) (base.type.type_base base.type.Z) + => fun '(r1', r2') + => (option_beq zrange_beq (Some r1) (annotation_of_state r1')) + && (option_beq zrange_beq (Some r2) (annotation_of_state r2')) + | _, _, _ => fun _ => false + end. + + Definition annotation_to_cast_helper {var1} {s'sd} (idc : ident s'sd) : option (@expr var1 (type.domain (type.codomain s'sd)) -> @expr var1 (type.codomain (type.codomain s'sd))) + := match idc with + | ident.Z_cast => Some (fun x => x) + | ident.Z_cast2 => Some (fun x => x) + | _ => None + end. + + Definition annotation_to_cast {var1} {s d} (e : @expr var1 (s -> d)) : option (@expr var1 s -> @expr var1 d) + := match invert_AppIdent e with + | Some (existT s' (idc, _)) => annotation_to_cast_helper idc + | None => None + end. + + Definition annotation_to_expr {var1} t (idc : @expr var1 t) : option (forall var2, @expr var2 t) + := match invert_Z_cast idc, invert_Z_cast2 idc, t with + | Some r, _, (type.base tZ -> type.base tZ)%etype + => Some (fun var2 => cstZ r) + | _, Some (r1, r2), (type.base (tZ * tZ) -> type.base (tZ * tZ))%etype + => Some (fun var2 => cstZZ r1 r2) + | _, _, _ => None + end. + Definition bottom' T : abstract_domain' T + := ZRange.type.base.option.None. + Definition extract_list_state A (st : abstract_domain' (base.type.list A)) : option (list (abstract_domain' A)) + := st. + Definition extract_option_state A (st : abstract_domain' (base.type.option A)) : option (option (abstract_domain' A)) + := st. + + Definition strip_all_annotations strip_annotations_under {var} {t} (e : @expr var t) : @expr var t + := @ident.strip_all_annotations var (@annotation_to_cast _) strip_annotations_under t e. + + Definition eval_with_bound {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) (skip_annotations_under : forall t, ident t -> bool) (strip_preexisting_annotations : bool) {var} {t} (e : @expr _ t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : expr t + := let annotate_with_state := true in + (@partial.ident.eval_with_bound) + var abstract_domain' annotate_expr bottom' (abstract_interp_ident assume_cast_truncates) extract_list_state extract_option_state is_annotated_for (strip_annotation assume_cast_truncates strip_preexisting_annotations) skip_annotations_under annotate_with_state t e bound. + + Definition eta_expand_with_bound {var} {t} (e : @expr _ t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : expr t + := let assume_cast_truncates := false in + let shiftr_avoid_uint1 : shiftr_avoid_uint1_opt := false (* this doesn't actually matter for [eta_expand_with_bound], which only invokes [abstract_interp_ident] on things like [fst], [snd], etc *) in + (@partial.ident.eta_expand_with_bound) + var abstract_domain' annotate_expr bottom' (abstract_interp_ident assume_cast_truncates) extract_list_state extract_option_state is_annotated_for t e bound. + + Definition StripAllAnnotations strip_annotations_under {t} (e : Expr t) : Expr t + := fun var => strip_all_annotations strip_annotations_under (e _). + + Definition EvalWithBound {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) (skip_annotations_under : forall t, ident t -> bool) (strip_preexisting_annotations : bool) {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : Expr t + := fun var => eval_with_bound assume_cast_truncates skip_annotations_under strip_preexisting_annotations (e _) bound. + Definition EtaExpandWithBound {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : Expr t + := fun var => eta_expand_with_bound (e _) bound. + End with_relax. + + Definition strip_annotations {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {var} {t} (e : @expr _ t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : expr t + := let strip_preexisting_annotations := true in + let annotate_with_state := false in + let skip_annotations_under := fun _ _ => false in + (@partial.ident.eval_with_bound) + var abstract_domain' (annotate_expr default_relax_zrange) bottom' (abstract_interp_ident assume_cast_truncates) extract_list_state extract_option_state (is_annotated_for default_relax_zrange) (strip_annotation assume_cast_truncates strip_preexisting_annotations) skip_annotations_under annotate_with_state t e bound. + + Definition StripAnnotations {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : Expr t + := fun var => strip_annotations assume_cast_truncates (e _) bound. + + Definition eval {opts : AbstractInterpretation.Options} {var} {t} (e : @expr _ t) : expr t + := let assume_cast_truncates := false in + let strip_preexisting_annotations := false in + (@partial.ident.eval) + var abstract_domain' (annotate_expr default_relax_zrange) bottom' (abstract_interp_ident assume_cast_truncates) extract_list_state extract_option_state (is_annotated_for default_relax_zrange) (strip_annotation assume_cast_truncates strip_preexisting_annotations) t e. + Definition Eval {opts : AbstractInterpretation.Options} {t} (e : Expr t) : Expr t + := fun var => eval (e _). + Definition EtaExpandWithListInfoFromBound {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : Expr t + := EtaExpandWithBound default_relax_zrange e (type.map_for_each_lhs_of_arrow (@ZRange.type.option.strip_ranges) bound). + Definition extract {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {t} (e : expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : abstract_domain' (type.final_codomain t) + := @partial.ident.extract abstract_domain' bottom' (abstract_interp_ident assume_cast_truncates) t e bound. + Definition Extract {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : abstract_domain' (type.final_codomain t) + := @partial.ident.extract abstract_domain' bottom' (abstract_interp_ident assume_cast_truncates) t (e _) bound. + End specialized. + End partial. + Import API. + + Module Import CheckCasts. + Fixpoint get_casts {t} (e : expr t) : list { t : _ & forall var, @expr var t } + := match partial.annotation_to_expr _ e with + | Some e => [existT _ t e] + | None + => match e with + | expr.Ident t idc => nil + | expr.Var t v => v + | expr.Abs s d f => @get_casts _ (f nil) + | expr.App s d f x => @get_casts _ f ++ @get_casts _ x + | expr.LetIn A B x f => @get_casts _ x ++ @get_casts _ (f nil) + end + end%list. + + Definition GetUnsupportedCasts {t} (e : Expr t) : list { t : _ & forall var, @expr var t } + := get_casts (e _). + End CheckCasts. + + Definition PartialEvaluateWithBounds + {shiftr_avoid_uint1 : shiftr_avoid_uint1_opt} + (relax_zrange : zrange -> option zrange) + (assume_cast_truncates : bool) + (skip_annotations_under : forall t, ident t -> bool) + (strip_preexisting_annotations : bool) + {t} (e : Expr t) + (bound : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + : Expr t + := partial.EvalWithBound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations (GeneralizeVar.GeneralizeVar (e _)) bound. + Definition PartialEvaluateWithListInfoFromBounds + {shiftr_avoid_uint1 : shiftr_avoid_uint1_opt} + {t} (e : Expr t) + (bound : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + : Expr t + := partial.EtaExpandWithListInfoFromBound (GeneralizeVar.GeneralizeVar (e _)) bound. + + Definition CheckedPartialEvaluateWithBounds + {shiftr_avoid_uint1 : shiftr_avoid_uint1_opt} + (relax_zrange : zrange -> option zrange) + (assume_cast_truncates : bool) + (skip_annotations_under : forall t, ident t -> bool) + (strip_preexisting_annotations : bool) + {t} (E : Expr t) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + (b_out : ZRange.type.base.option.interp (type.final_codomain t)) + : Expr t + (ZRange.type.base.option.interp (type.final_codomain t) * Expr t + list { t : _ & forall var, @expr var t }) + := dlet_nd e := GeneralizeVar.ToFlat E in + let E := GeneralizeVar.FromFlat e in + let b_computed := partial.Extract assume_cast_truncates E b_in in + let unsupported_casts := if strip_preexisting_annotations + then nil + else CheckCasts.GetUnsupportedCasts E in + match unsupported_casts with + | nil => (let E := PartialEvaluateWithBounds relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations E b_in in + if ZRange.type.base.option.is_tighter_than b_computed b_out + then @inl (Expr t) _ E + else inr (@inl (ZRange.type.base.option.interp (type.final_codomain t) * Expr t) _ (b_computed, E))) + | unsupported_casts => inr (inr unsupported_casts) + end. +End Compilers. diff --git a/src/AbstractInterpretation/Proofs.v b/src/AbstractInterpretation/Fancy/Proofs.v similarity index 99% rename from src/AbstractInterpretation/Proofs.v rename to src/AbstractInterpretation/Fancy/Proofs.v index 22a96f8bff..515c3a4d61 100644 --- a/src/AbstractInterpretation/Proofs.v +++ b/src/AbstractInterpretation/Fancy/Proofs.v @@ -39,8 +39,8 @@ Require Import Rewriter.Language.Wf. Require Import Crypto.Language.API. Require Import Crypto.CastLemmas. Require Import Rewriter.Language.UnderLetsProofs. -Require Import Crypto.AbstractInterpretation.AbstractInterpretation. -Require Import Crypto.AbstractInterpretation.Wf. +Require Import Crypto.AbstractInterpretation.Fancy.AbstractInterpretation. +Require Import Crypto.AbstractInterpretation.Fancy.Wf. Require Import Crypto.AbstractInterpretation.ZRangeProofs. Module Compilers. @@ -51,9 +51,9 @@ Module Compilers. Import Language.Wf.Compilers. Import Language.API.Compilers. Import UnderLetsProofs.Compilers. - Import AbstractInterpretation.Wf.Compilers. + Import AbstractInterpretation.Fancy.Wf.Compilers. Import AbstractInterpretation.ZRangeProofs.Compilers. - Import AbstractInterpretation.Wf.Compilers.partial. + Import AbstractInterpretation.Fancy.Wf.Compilers.partial. Import invert_expr. Local Notation related_bounded' b v1 v2 diff --git a/src/AbstractInterpretation/Wf.v b/src/AbstractInterpretation/Fancy/Wf.v similarity index 99% rename from src/AbstractInterpretation/Wf.v rename to src/AbstractInterpretation/Fancy/Wf.v index fb64ae8d4c..ce0ac95634 100644 --- a/src/AbstractInterpretation/Wf.v +++ b/src/AbstractInterpretation/Fancy/Wf.v @@ -26,7 +26,7 @@ Require Import Crypto.Language.InversionExtra. Require Import Rewriter.Language.Wf. Require Import Rewriter.Language.UnderLetsProofs. Require Import Rewriter.Util.Decidable. -Require Import Crypto.AbstractInterpretation.AbstractInterpretation. +Require Import Crypto.AbstractInterpretation.Fancy.AbstractInterpretation. Require Import Crypto.AbstractInterpretation.ZRangeCommonProofs. Import Coq.Lists.List. diff --git a/src/AbstractInterpretation/WfExtra.v b/src/AbstractInterpretation/Fancy/WfExtra.v similarity index 72% rename from src/AbstractInterpretation/WfExtra.v rename to src/AbstractInterpretation/Fancy/WfExtra.v index b63313e1a5..70598f4087 100644 --- a/src/AbstractInterpretation/WfExtra.v +++ b/src/AbstractInterpretation/Fancy/WfExtra.v @@ -2,8 +2,8 @@ Require Import Rewriter.Language.Language. Require Import Rewriter.Language.Wf. Require Import Crypto.Language.API. Require Import Crypto.Language.WfExtra. -Require Import Crypto.AbstractInterpretation.AbstractInterpretation. -Require Import Crypto.AbstractInterpretation.Wf. +Require Import Crypto.AbstractInterpretation.Fancy.AbstractInterpretation. +Require Import Crypto.AbstractInterpretation.Fancy.Wf. Module Compilers. Import Language.Compilers. @@ -11,12 +11,12 @@ Module Compilers. Import Language.Wf.Compilers. Import Language.API.Compilers. Import Language.WfExtra.Compilers. - Import AbstractInterpretation.AbstractInterpretation.Compilers. - Import AbstractInterpretation.Wf.Compilers. + Import AbstractInterpretation.Fancy.AbstractInterpretation.Compilers. + Import AbstractInterpretation.Fancy.Wf.Compilers. Import Compilers.API. - Import AbstractInterpretation.AbstractInterpretation.Compilers.partial. - Import AbstractInterpretation.Wf.Compilers.partial. + Import AbstractInterpretation.Fancy.AbstractInterpretation.Compilers.partial. + Import AbstractInterpretation.Fancy.Wf.Compilers.partial. #[global] Hint Resolve Wf_Eval Wf_EvalWithBound Wf_EtaExpandWithBound Wf_EtaExpandWithListInfoFromBound Wf_StripAnnotations Wf_StripAllAnnotations : wf_extra. From 520b9312393182cdd3e253d02087c42b3fee2bc9 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Tue, 15 Aug 2023 19:08:50 -0400 Subject: [PATCH 2/3] Avoid bottomify-ing higher order functions Reworked value type to more closely match the possible spec. Deduplicated `interp_ident` / `abstract_interp_ident` overlapped state. The specification of `reify` and `reflect` has been simplified to more closely match the intuition: `reify`'s spec no longer talks about output bounds, only interpretation, and a companion lemma `abstraction_relation_of_related_bounded_value'` shows that the spec of `reflect` is effectively one direction of an isomorphism between interpreted-values bounded by abstract state, and `value`s which are in relationship with the given interpreted-value. Not really sure how best to describe this in words yet, probably needs some digesting work. There's some interesting lemmas where we have an equivalency between two ways of representing things only for up-to-second-order types, and an implication for up-to-third-order types, which is currently all the identifiers. But this means that if we add fourth-order identifiers, we'll have to deal with two different sorts of `Proper` abstraction relations (one used by `Assembly/Symbolic` and one used by bounds analysis, though it's possible the `Assembly/Symbolic` one can be adapted). * Guard exponentially slow bounds analysis with a flag With `--fancy-and-powerful-but-exponentially-slow-bounds-analysis` Idk how to implement a non-exponential PHOAS passe for this.
Timing Diff before guarding with flag

``` After | Peak Mem | File Name | Before | Peak Mem || Change || Change (mem) | % Change | % Change (mem) ------------------------------------------------------------------------------------------------------------------------------------------------------------------------------------ 5494m57.81s | 2858552 ko | Total Time / Peak Mem | 120m21.13s | 2852328 ko || +5374m36.68s || 6224 ko | +4465.73% | +0.21% ------------------------------------------------------------------------------------------------------------------------------------------------------------------------------------ 352m04.93s | 1781524 ko | fiat-rust/src/p384_scalar_32.rs | 1m53.74s | 1628856 ko || +350m11.18s || 152668 ko | +18472.99% | +9.37% 351m13.49s | 1916288 ko | fiat-java/src/FiatP384Scalar.java | 1m53.23s | 1620248 ko || +349m20.26s || 296040 ko | +18511.22% | +18.27% 344m52.23s | 1666100 ko | fiat-go/32/p384scalar/p384scalar.go | 1m53.04s | 1734424 ko || +342m59.18s || -68324 ko | +18205.22% | -3.93% 343m32.96s | 1898556 ko | fiat-c/src/p384_scalar_32.c | 1m51.27s | 1627512 ko || +341m41.68s || 271044 ko | +18425.17% | +16.65% 341m26.05s | 1834040 ko | fiat-json/src/p384_scalar_32.json | 1m46.90s | 1802484 ko || +339m39.14s || 31556 ko | +19063.75% | +1.75% 336m05.54s | 1870632 ko | fiat-rust/src/p384_32.rs | 1m50.16s | 1709428 ko || +334m15.38s || 161204 ko | +18205.68% | +9.43% 334m29.07s | 1989180 ko | fiat-java/src/FiatP384.java | 1m52.50s | 1817848 ko || +332m36.56s || 171332 ko | +17739.17% | +9.42% 333m35.10s | 1806688 ko | fiat-json/src/p384_32.json | 1m53.94s | 1629284 ko || +331m41.15s || 177404 ko | +17466.35% | +10.88% 328m36.71s | 1779836 ko | fiat-bedrock2/src/p384_scalar_32.c | 1m48.68s | 1758232 ko || +326m48.02s || 21604 ko | +18041.98% | +1.22% 328m39.19s | 1662676 ko | fiat-go/32/p384/p384.go | 1m53.01s | 1625684 ko || +326m46.18s || 36992 ko | +17349.06% | +2.27% 321m01.71s | 1709592 ko | fiat-c/src/p384_32.c | 1m49.66s | 2029588 ko || +319m12.04s || -319996 ko | +17464.93% | -15.76% 320m33.13s | 1872256 ko | fiat-bedrock2/src/p384_32.c | 1m48.69s | 1687840 ko || +318m44.44s || 184416 ko | +17595.39% | +10.92% 318m14.38s | 2079716 ko | fiat-zig/src/p384_scalar_32.zig | 1m54.15s | 1960684 ko || +316m20.22s || 119032 ko | +16627.44% | +6.07% 308m20.28s | 1645516 ko | fiat-zig/src/p384_32.zig | 1m52.66s | 1948020 ko || +306m27.61s || -302504 ko | +16321.33% | -15.52% 140m13.96s | 2858552 ko | Bedrock/Field/Synthesis/Examples/p224_64_new.vo | 5m29.46s | 2852328 ko || +134m44.49s || 6224 ko | +2453.86% | +0.21% 15m08.73s | 412980 ko | fiat-zig/src/secp256k1_montgomery_scalar_32.zig | 0m17.06s | 365400 ko || +14m51.67s || 47580 ko | +5226.67% | +13.02% 15m07.92s | 444232 ko | fiat-rust/src/secp256k1_montgomery_scalar_32.rs | 0m16.58s | 413372 ko || +14m51.33s || 30860 ko | +5375.99% | +7.46% 15m07.44s | 447232 ko | fiat-bedrock2/src/secp256k1_montgomery_scalar_32.c | 0m17.59s | 396628 ko || +14m49.85s || 50604 ko | +5058.84% | +12.75% 15m05.89s | 413980 ko | fiat-zig/src/p256_scalar_32.zig | 0m17.13s | 404448 ko || +14m48.75s || 9532 ko | +5188.32% | +2.35% 15m00.09s | 462000 ko | fiat-json/src/secp256k1_montgomery_scalar_32.json | 0m17.14s | 372928 ko || +14m42.95s || 89072 ko | +5151.40% | +23.88% 14m59.00s | 383440 ko | fiat-go/32/secp256k1montgomeryscalar/secp256k1montgomeryscalar.go | 0m16.75s | 439368 ko || +14m42.25s || -55928 ko | +5267.16% | -12.72% 14m58.83s | 417168 ko | fiat-java/src/FiatSecp256K1MontgomeryScalar.java | 0m17.43s | 335352 ko || +14m41.40s || 81816 ko | +5056.79% | +24.39% 14m57.88s | 378000 ko | fiat-go/32/p256scalar/p256scalar.go | 0m16.89s | 362996 ko || +14m40.99s || 15004 ko | +5216.04% | +4.13% 14m43.11s | 453008 ko | fiat-bedrock2/src/curve25519_scalar_32.c | 0m16.87s | 404488 ko || +14m26.24s || 48520 ko | +5134.79% | +11.99% 14m42.26s | 423172 ko | fiat-zig/src/curve25519_scalar_32.zig | 0m16.75s | 338768 ko || +14m25.50s || 84404 ko | +5167.22% | +24.91% 14m34.97s | 386392 ko | fiat-java/src/FiatSecp256K1Montgomery.java | 0m16.92s | 375048 ko || +14m18.05s || 11344 ko | +5071.21% | +3.02% 14m35.14s | 439764 ko | fiat-bedrock2/src/p256_scalar_32.c | 0m17.47s | 466176 ko || +14m17.66s || -26412 ko | +4909.38% | -5.66% 14m33.96s | 418052 ko | fiat-c/src/p256_scalar_32.c | 0m16.30s | 433696 ko || +14m17.66s || -15644 ko | +5261.71% | -3.60% 14m32.17s | 379176 ko | fiat-zig/src/secp256k1_montgomery_32.zig | 0m16.64s | 440088 ko || +14m15.52s || -60912 ko | +5141.40% | -13.84% 14m32.26s | 386104 ko | fiat-java/src/FiatP256Scalar.java | 0m17.60s | 395220 ko || +14m14.65s || -9116 ko | +4856.02% | -2.30% 14m20.14s | 476544 ko | fiat-json/src/secp256k1_montgomery_32.json | 0m16.80s | 384668 ko || +14m03.34s || 91876 ko | +5019.88% | +23.88% 14m18.94s | 380196 ko | fiat-java/src/FiatCurve25519Scalar.java | 0m16.99s | 369068 ko || +14m01.95s || 11128 ko | +4955.56% | +3.01% 14m14.20s | 438032 ko | fiat-bedrock2/src/secp256k1_montgomery_32.c | 0m16.55s | 414072 ko || +13m57.65s || 23960 ko | +5061.32% | +5.78% 14m10.18s | 382024 ko | fiat-c/src/secp256k1_montgomery_scalar_32.c | 0m16.48s | 349956 ko || +13m53.69s || 32068 ko | +5058.85% | +9.16% 14m07.30s | 380804 ko | fiat-go/32/secp256k1montgomery/secp256k1montgomery.go | 0m15.85s | 343344 ko || +13m51.44s || 37460 ko | +5245.74% | +10.91% 14m05.83s | 384216 ko | fiat-rust/src/p256_scalar_32.rs | 0m16.99s | 347112 ko || +13m48.84s || 37104 ko | +4878.39% | +10.68% 14m02.24s | 421196 ko | fiat-rust/src/secp256k1_montgomery_32.rs | 0m16.59s | 342720 ko || +13m45.64s || 78476 ko | +4976.79% | +22.89% 13m55.86s | 416884 ko | fiat-bedrock2/src/p256_32.c | 0m15.76s | 363984 ko || +13m40.10s || 52900 ko | +5203.68% | +14.53% 13m49.51s | 435140 ko | fiat-zig/src/p256_32.zig | 0m15.90s | 375036 ko || +13m33.61s || 60104 ko | +5117.04% | +16.02% 13m48.55s | 402868 ko | fiat-json/src/curve25519_scalar_32.json | 0m16.09s | 444636 ko || +13m32.45s || -41768 ko | +5049.47% | -9.39% 13m47.62s | 374572 ko | fiat-c/src/curve25519_scalar_32.c | 0m15.90s | 352232 ko || +13m31.72s || 22340 ko | +5105.15% | +6.34% 13m40.58s | 411000 ko | fiat-c/src/secp256k1_montgomery_32.c | 0m15.54s | 361600 ko || +13m25.04s || 49400 ko | +5180.43% | +13.66% 13m38.90s | 391764 ko | fiat-java/src/FiatP256.java | 0m15.69s | 392176 ko || +13m23.20s || -412 ko | +5119.24% | -0.10% 13m28.98s | 382116 ko | fiat-go/32/p256/p256.go | 0m15.69s | 361040 ko || +13m13.28s || 21076 ko | +5056.02% | +5.83% 13m24.91s | 408552 ko | fiat-rust/src/p256_32.rs | 0m15.57s | 405540 ko || +13m09.33s || 3012 ko | +5069.62% | +0.74% 13m25.05s | 424812 ko | fiat-rust/src/curve25519_scalar_32.rs | 0m16.29s | 429536 ko || +13m08.75s || -4724 ko | +4841.98% | -1.09% 13m23.56s | 408716 ko | fiat-json/src/p256_scalar_32.json | 0m17.19s | 433976 ko || +13m06.36s || -25260 ko | +4574.57% | -5.82% 13m06.66s | 379524 ko | fiat-go/32/curve25519scalar/curve25519scalar.go | 0m16.39s | 325156 ko || +12m50.26s || 54368 ko | +4699.63% | +16.72% 13m01.92s | 376048 ko | fiat-c/src/p256_32.c | 0m15.20s | 359812 ko || +12m46.71s || 16236 ko | +5044.21% | +4.51% 12m49.57s | 423456 ko | fiat-json/src/p256_32.json | 0m15.99s | 429740 ko || +12m33.58s || -6284 ko | +4712.82% | -1.46% 7m46.98s | 284096 ko | fiat-zig/src/p434_64.zig | 0m18.66s | 250696 ko || +7m28.31s || 33400 ko | +2402.57% | +13.32% 7m42.79s | 333916 ko | fiat-json/src/p434_64.json | 0m18.88s | 286668 ko || +7m23.91s || 47248 ko | +2351.21% | +16.48% 7m42.28s | 277712 ko | fiat-go/64/p434/p434.go | 0m18.69s | 239700 ko || +7m23.58s || 38012 ko | +2373.40% | +15.85% 7m34.48s | 367696 ko | fiat-bedrock2/src/p434_64.c | 0m18.03s | 373868 ko || +7m16.45s || -6172 ko | +2420.68% | -1.65% 7m30.04s | 281384 ko | fiat-c/src/p434_64.c | 0m18.12s | 237136 ko || +7m11.92s || 44248 ko | +2383.66% | +18.65% 7m24.16s | 273604 ko | fiat-rust/src/p434_64.rs | 0m18.41s | 246432 ko || +7m05.75s || 27172 ko | +2312.60% | +11.02% 5m07.03s | 256280 ko | fiat-java/src/FiatP224.java | 0m09.02s | 261688 ko || +4m58.00s || -5408 ko | +3303.88% | -2.06% 5m00.31s | 238104 ko | fiat-go/32/p224/p224.go | 0m09.01s | 223612 ko || +4m51.30s || 14492 ko | +3233.07% | +6.48% 4m59.60s | 240328 ko | fiat-zig/src/p224_32.zig | 0m09.37s | 225308 ko || +4m50.23s || 15020 ko | +3097.43% | +6.66% 4m57.44s | 288372 ko | fiat-json/src/p224_32.json | 0m09.08s | 283848 ko || +4m48.36s || 4524 ko | +3175.77% | +1.59% 4m57.40s | 239760 ko | fiat-rust/src/p224_32.rs | 0m09.10s | 269448 ko || +4m48.29s || -29688 ko | +3168.13% | -11.01% 4m48.57s | 310468 ko | fiat-bedrock2/src/p224_32.c | 0m08.99s | 282600 ko || +4m39.57s || 27868 ko | +3109.89% | +9.86% 4m48.27s | 239228 ko | fiat-c/src/p224_32.c | 0m08.96s | 230412 ko || +4m39.31s || 8816 ko | +3117.29% | +3.82% 2m48.44s | 184568 ko | fiat-json/src/p384_scalar_64.json | 0m11.23s | 168860 ko || +2m37.21s || 15708 ko | +1399.91% | +9.30% 2m48.14s | 167980 ko | fiat-zig/src/p384_scalar_64.zig | 0m11.15s | 193456 ko || +2m36.98s || -25476 ko | +1407.98% | -13.16% 2m43.61s | 213748 ko | fiat-bedrock2/src/p384_scalar_64.c | 0m11.14s | 192024 ko || +2m32.47s || 21724 ko | +1368.67% | +11.31% 2m43.36s | 167968 ko | fiat-rust/src/p384_scalar_64.rs | 0m11.11s | 148400 ko || +2m32.25s || 19568 ko | +1370.38% | +13.18% 2m41.95s | 184636 ko | fiat-c/src/p384_scalar_64.c | 0m10.86s | 154744 ko || +2m31.08s || 29892 ko | +1391.25% | +19.31% 2m37.83s | 178156 ko | fiat-go/64/p384scalar/p384scalar.go | 0m11.15s | 193256 ko || +2m26.67s || -15100 ko | +1315.51% | -7.81% 2m29.44s | 166896 ko | fiat-zig/src/p384_64.zig | 0m09.52s | 154632 ko || +2m19.91s || 12264 ko | +1469.74% | +7.93% 2m28.98s | 184900 ko | fiat-json/src/p384_64.json | 0m09.31s | 171972 ko || +2m19.66s || 12928 ko | +1500.21% | +7.51% 2m26.58s | 176112 ko | fiat-go/64/p384/p384.go | 0m09.47s | 181568 ko || +2m17.10s || -5456 ko | +1447.83% | -3.00% 2m25.35s | 175320 ko | fiat-c/src/p384_64.c | 0m09.21s | 151936 ko || +2m16.13s || 23384 ko | +1478.17% | +15.39% 2m23.86s | 189000 ko | fiat-rust/src/p384_64.rs | 0m09.49s | 145760 ko || +2m14.37s || 43240 ko | +1415.91% | +29.66% 2m23.35s | 211804 ko | fiat-bedrock2/src/p384_64.c | 0m09.35s | 189388 ko || +2m14.00s || 22416 ko | +1433.15% | +11.83% 2m29.46s | 2135192 ko | SlowPrimeSynthesisExamples.vo | 1m28.36s | 2041032 ko || +1m01.10s || 94160 ko | +69.14% | +4.61% 8m55.13s | 2650896 ko | Bedrock/End2End/X25519/GarageDoor.vo | 8m25.15s | 2656564 ko || +0m29.98s || -5668 ko | +5.93% | -0.21% N/A | N/A | PerfTesting/PerfTestSearch.vo | 0m25.13s | 1373644 ko || -0m25.12s || -1373644 ko | -100.00% | -100.00% 1m22.31s | 941440 ko | AbstractInterpretation/Wf.vo | 1m00.63s | 875644 ko || +0m21.67s || 65796 ko | +35.75% | +7.51% 0m48.67s | 773988 ko | AbstractInterpretation/Proofs.vo | 0m29.24s | 702092 ko || +0m19.43s || 71896 ko | +66.45% | +10.24% N/A | N/A | PerfTesting/PerfTestSearchPattern.vo | 0m17.89s | 1373708 ko || -0m17.89s || -1373708 ko | -100.00% | -100.00% 2m05.77s | 1631168 ko | Bedrock/End2End/X25519/Field25519.vo | 1m51.28s | 1594648 ko || +0m14.48s || 36520 ko | +13.02% | +2.29% 0m14.59s | 64448 ko | fiat-rust/src/secp256k1_montgomery_scalar_64.rs | 0m02.67s | 60756 ko || +0m11.92s || 3692 ko | +446.44% | +6.07% 0m14.52s | 62432 ko | fiat-zig/src/secp256k1_montgomery_scalar_64.zig | 0m02.76s | 58520 ko || +0m11.75s || 3912 ko | +426.08% | +6.68% 0m14.45s | 67272 ko | fiat-zig/src/p256_scalar_64.zig | 0m02.73s | 61128 ko || +0m11.71s || 6144 ko | +429.30% | +10.05% 0m14.25s | 62036 ko | fiat-rust/src/p256_scalar_64.rs | 0m02.73s | 59396 ko || +0m11.51s || 2640 ko | +421.97% | +4.44% 0m14.21s | 72352 ko | fiat-json/src/secp256k1_montgomery_scalar_64.json | 0m02.82s | 70904 ko || +0m11.39s || 1448 ko | +403.90% | +2.04% 0m14.16s | 62316 ko | fiat-c/src/secp256k1_montgomery_scalar_64.c | 0m02.69s | 61008 ko || +0m11.47s || 1308 ko | +426.39% | +2.14% 0m14.02s | 72660 ko | fiat-json/src/p256_scalar_64.json | 0m02.79s | 73116 ko || +0m11.23s || -456 ko | +402.50% | -0.62% 0m14.01s | 61848 ko | fiat-c/src/p256_scalar_64.c | 0m02.67s | 61148 ko || +0m11.33s || 700 ko | +424.71% | +1.14% 0m13.61s | 67516 ko | fiat-go/64/p256scalar/p256scalar.go | 0m02.75s | 71172 ko || +0m10.85s || -3656 ko | +394.90% | -5.13% 0m13.51s | 85796 ko | fiat-bedrock2/src/p256_scalar_64.c | 0m02.80s | 84480 ko || +0m10.71s || 1316 ko | +382.50% | +1.55% 0m13.33s | 66312 ko | fiat-go/64/secp256k1montgomeryscalar/secp256k1montgomeryscalar.go | 0m02.77s | 62180 ko || +0m10.56s || 4132 ko | +381.22% | +6.64% 0m13.20s | 88340 ko | fiat-bedrock2/src/secp256k1_montgomery_scalar_64.c | 0m02.79s | 84980 ko || +0m10.41s || 3360 ko | +373.11% | +3.95% 0m13.16s | 70728 ko | fiat-json/src/secp256k1_montgomery_64.json | 0m02.49s | 75256 ko || +0m10.67s || -4528 ko | +428.51% | -6.01% 0m12.74s | 65292 ko | fiat-zig/src/secp256k1_montgomery_64.zig | 0m02.42s | 59724 ko || +0m10.32s || 5568 ko | +426.44% | +9.32% 0m12.72s | 60072 ko | fiat-zig/src/curve25519_scalar_64.zig | 0m02.33s | 55908 ko || +0m10.39s || 4164 ko | +445.92% | +7.44% 0m12.66s | 61844 ko | fiat-rust/src/secp256k1_montgomery_64.rs | 0m02.44s | 64784 ko || +0m10.22s || -2940 ko | +418.85% | -4.53% 0m44.36s | 60596 ko | fiat-zig/src/p521_32.zig | 0m34.44s | 60684 ko || +0m09.92s || -88 ko | +28.80% | -0.14% 0m44.26s | 112564 ko | fiat-json/src/p521_32.json | 0m34.64s | 106924 ko || +0m09.61s || 5640 ko | +27.77% | +5.27% 0m44.00s | 64936 ko | fiat-java/src/FiatP521.java | 0m34.40s | 59656 ko || +0m09.60s || 5280 ko | +27.90% | +8.85% 0m12.18s | 66900 ko | fiat-rust/src/curve25519_scalar_64.rs | 0m02.32s | 60304 ko || +0m09.85s || 6596 ko | +425.00% | +10.93% 0m12.07s | 67140 ko | fiat-c/src/secp256k1_montgomery_64.c | 0m02.43s | 59524 ko || +0m09.64s || 7616 ko | +396.70% | +12.79% 0m12.06s | 60364 ko | fiat-c/src/curve25519_scalar_64.c | 0m02.27s | 60644 ko || +0m09.79s || -280 ko | +431.27% | -0.46% 0m12.03s | 68432 ko | fiat-json/src/curve25519_scalar_64.json | 0m02.36s | 67732 ko || +0m09.67s || 700 ko | +409.74% | +1.03% 0m11.92s | 84348 ko | fiat-bedrock2/src/secp256k1_montgomery_64.c | 0m02.45s | 88860 ko || +0m09.46s || -4512 ko | +386.53% | -5.07% 0m11.92s | 67724 ko | fiat-go/64/secp256k1montgomery/secp256k1montgomery.go | 0m02.52s | 71412 ko || +0m09.40s || -3688 ko | +373.01% | -5.16% 0m11.58s | 64176 ko | fiat-go/64/curve25519scalar/curve25519scalar.go | 0m02.32s | 67952 ko || +0m09.25s || -3776 ko | +399.13% | -5.55% 0m11.37s | 84500 ko | fiat-bedrock2/src/curve25519_scalar_64.c | 0m02.28s | 80992 ko || +0m09.08s || 3508 ko | +398.68% | +4.33% 0m10.91s | 60548 ko | fiat-c/src/p224_64.c | 0m01.82s | 61764 ko || +0m09.08s || -1216 ko | +499.45% | -1.96% 0m10.82s | 65748 ko | fiat-zig/src/p224_64.zig | 0m01.84s | 62532 ko || +0m08.98s || 3216 ko | +488.04% | +5.14% 0m10.55s | 60692 ko | fiat-rust/src/p224_64.rs | 0m01.84s | 61180 ko || +0m08.71s || -488 ko | +473.36% | -0.79% 0m10.52s | 73812 ko | fiat-json/src/p224_64.json | 0m01.88s | 70604 ko || +0m08.64s || 3208 ko | +459.57% | +4.54% 0m10.37s | 64696 ko | fiat-go/64/p224/p224.go | 0m01.92s | 62144 ko || +0m08.44s || 2552 ko | +440.10% | +4.10% 0m10.11s | 83476 ko | fiat-bedrock2/src/p224_64.c | 0m01.91s | 80568 ko || +0m08.19s || 2908 ko | +429.31% | +3.60% 0m15.25s | 124888 ko | fiat-json/src/p448_solinas_32.json | 0m08.23s | 116848 ko || +0m07.01s || 8040 ko | +85.29% | +6.88% 0m09.70s | 75056 ko | fiat-json/src/p256_64.json | 0m01.88s | 66596 ko || +0m07.81s || 8460 ko | +415.95% | +12.70% 0m09.59s | 61964 ko | fiat-zig/src/p256_64.zig | 0m01.79s | 58052 ko || +0m07.79s || 3912 ko | +435.75% | +6.73% 0m09.50s | 82532 ko | fiat-bedrock2/src/p256_64.c | 0m01.84s | 77096 ko || +0m07.66s || 5436 ko | +416.30% | +7.05% 0m09.43s | 61676 ko | fiat-rust/src/p256_64.rs | 0m01.80s | 59200 ko || +0m07.62s || 2476 ko | +423.88% | +4.18% 0m09.40s | 65020 ko | fiat-go/64/p256/p256.go | 0m01.84s | 68656 ko || +0m07.56s || -3636 ko | +410.86% | -5.29% 0m08.91s | 65360 ko | fiat-c/src/p256_64.c | 0m01.74s | 61792 ko || +0m07.16s || 3568 ko | +412.06% | +5.77% 0m14.66s | 59488 ko | fiat-zig/src/p448_solinas_32.zig | 0m08.21s | 58580 ko || +0m06.44s || 908 ko | +78.56% | +1.55% 0m14.42s | 60808 ko | fiat-rust/src/p448_solinas_32.rs | 0m08.12s | 57972 ko || +0m06.30s || 2836 ko | +77.58% | +4.89% 0m43.35s | 62524 ko | fiat-rust/src/p521_32.rs | 0m38.15s | 60356 ko || +0m05.20s || 2168 ko | +13.63% | +3.59% 0m42.96s | 148852 ko | fiat-bedrock2/src/p521_32.c | 0m38.63s | 145628 ko || +0m04.32s || 3224 ko | +11.20% | +2.21% 0m42.44s | 63004 ko | fiat-c/src/p521_32.c | 0m37.82s | 59428 ko || +0m04.61s || 3576 ko | +12.21% | +6.01% 0m12.71s | 60360 ko | fiat-c/src/p448_solinas_32.c | 0m08.16s | 58932 ko || +0m04.55s || 1428 ko | +55.75% | +2.42% 1m51.21s | 2339376 ko | Fancy/Barrett256.vo | 1m48.02s | 2417076 ko || +0m03.18s || -77700 ko | +2.95% | -3.21% 0m49.17s | 1960112 ko | Fancy/Montgomery256.vo | 0m46.03s | 1758768 ko || +0m03.14s || 201344 ko | +6.82% | +11.44% 0m39.42s | 1319164 ko | Bedrock/End2End/X25519/GarageDoorTop.vo | 0m36.42s | 1319624 ko || +0m03.00s || -460 ko | +8.23% | -0.03% N/A | N/A | Bedrock/Everything.vo | 0m03.96s | 1502884 ko || -0m03.96s || -1502884 ko | -100.00% | -100.00% N/A | N/A | Everything.vo | 0m03.45s | 1363476 ko || -0m03.45s || -1363476 ko | -100.00% | -100.00% N/A | N/A | PerfTesting/PerfTestPrint.vo | 0m02.94s | 1346788 ko || -0m02.94s || -1346788 ko | -100.00% | -100.00% 3m59.87s | 2553740 ko | Assembly/WithBedrock/Proofs.vo | 4m01.45s | 2553452 ko || -0m01.57s || 288 ko | -0.65% | +0.01% 1m54.05s | 1916900 ko | Bedrock/End2End/X25519/AddPrecomputed.vo | 1m52.31s | 1916820 ko || +0m01.73s || 80 ko | +1.54% | +0.00% 0m43.02s | 64392 ko | fiat-go/32/p521/p521.go | 0m41.64s | 61532 ko || +0m01.38s || 2860 ko | +3.31% | +4.64% 0m38.48s | 2146800 ko | ExtractionOCaml/word_by_word_montgomery | 0m40.27s | 2146720 ko || -0m01.79s || 80 ko | -4.44% | +0.00% 0m37.41s | 1617480 ko | ExtractionOCaml/bedrock2_base_conversion | 0m36.32s | 1620388 ko || +0m01.08s || -2908 ko | +3.00% | -0.17% 0m34.27s | 1554400 ko | ExtractionOCaml/dettman_multiplication | 0m35.58s | 1556408 ko || -0m01.30s || -2008 ko | -3.68% | -0.12% 0m33.95s | 1537192 ko | ExtractionOCaml/saturated_solinas | 0m35.03s | 1536764 ko || -0m01.07s || 428 ko | -3.08% | +0.02% 0m33.09s | 1531652 ko | ExtractionOCaml/base_conversion | 0m34.83s | 1534488 ko || -0m01.73s || -2836 ko | -4.99% | -0.18% 0m18.85s | 1135808 ko | Bedrock/End2End/Poly1305/Field1305.vo | 0m17.21s | 1129728 ko || +0m01.64s || 6080 ko | +9.52% | +0.53% 0m06.04s | 52332 ko | fiat-json/src/p521_64.json | 0m04.95s | 50248 ko || +0m01.08s || 2084 ko | +22.02% | +4.14% 0m05.99s | 37064 ko | fiat-zig/src/p521_64.zig | 0m04.90s | 36028 ko || +0m01.08s || 1036 ko | +22.24% | +2.87% 3m13.72s | 2603328 ko | Bedrock/End2End/RupicolaCrypto/ChaCha20.vo | 3m12.79s | 2602764 ko || +0m00.93s || 564 ko | +0.48% | +0.02% 1m20.21s | 1524980 ko | Assembly/EquivalenceProofs.vo | 1m19.93s | 1530348 ko || +0m00.28s || -5368 ko | +0.35% | -0.35% 1m08.96s | 1437312 ko | Assembly/WithBedrock/SymbolicProofs.vo | 1m09.05s | 1433128 ko || -0m00.08s || 4184 ko | -0.13% | +0.29% 0m52.36s | 849360 ko | AbstractInterpretation/ZRangeProofs.vo | 0m52.56s | 848576 ko || -0m00.20s || 784 ko | -0.38% | +0.09% 0m47.18s | 1522816 ko | Assembly/Symbolic.vo | 0m46.85s | 1519280 ko || +0m00.32s || 3536 ko | +0.70% | +0.23% 0m45.33s | 1104720 ko | Bedrock/Field/Translation/Proofs/ValidComputable/Expr.vo | 0m45.27s | 1102688 ko || +0m00.05s || 2032 ko | +0.13% | +0.18% 0m44.21s | 2147564 ko | ExtractionOCaml/bedrock2_solinas_reduction | 0m43.87s | 2147536 ko || +0m00.34s || 28 ko | +0.77% | +0.00% 0m43.12s | 2147892 ko | ExtractionOCaml/with_bedrock2_word_by_word_montgomery | 0m42.35s | 2147764 ko || +0m00.76s || 128 ko | +1.81% | +0.00% 0m42.80s | 2147776 ko | ExtractionOCaml/bedrock2_word_by_word_montgomery | 0m42.72s | 2147720 ko || +0m00.07s || 56 ko | +0.18% | +0.00% 0m41.90s | 2146676 ko | ExtractionOCaml/solinas_reduction | 0m40.99s | 2146676 ko || +0m00.90s || 0 ko | +2.22% | +0.00% 0m41.07s | 2148200 ko | ExtractionOCaml/bedrock2_unsaturated_solinas | 0m41.03s | 2148284 ko || +0m00.03s || -84 ko | +0.09% | -0.00% 0m41.02s | 2148244 ko | ExtractionOCaml/with_bedrock2_unsaturated_solinas | 0m40.67s | 2148388 ko || +0m00.35s || -144 ko | +0.86% | -0.00% 0m38.93s | 1016196 ko | Bedrock/Group/ScalarMult/MontgomeryLadder.vo | 0m38.63s | 1016156 ko || +0m00.29s || 40 ko | +0.77% | +0.00% 0m38.12s | 1640756 ko | ExtractionOCaml/bedrock2_dettman_multiplication | 0m38.27s | 1642204 ko || -0m00.15s || -1448 ko | -0.39% | -0.08% 0m37.73s | 1623836 ko | ExtractionOCaml/unsaturated_solinas | 0m36.98s | 1653884 ko || +0m00.75s || -30048 ko | +2.02% | -1.81% 0m37.59s | 1621000 ko | ExtractionOCaml/with_bedrock2_dettman_multiplication | 0m37.00s | 1623732 ko || +0m00.59s || -2732 ko | +1.59% | -0.16% 0m37.56s | 1622616 ko | ExtractionOCaml/bedrock2_saturated_solinas | 0m37.36s | 1622804 ko || +0m00.20s || -188 ko | +0.53% | -0.01% 0m37.50s | 1621644 ko | ExtractionOCaml/with_bedrock2_solinas_reduction | 0m36.64s | 1624384 ko || +0m00.85s || -2740 ko | +2.34% | -0.16% 0m37.48s | 1621204 ko | ExtractionOCaml/with_bedrock2_saturated_solinas | 0m36.73s | 1622828 ko || +0m00.75s || -1624 ko | +2.04% | -0.10% 0m37.41s | 1617420 ko | ExtractionOCaml/with_bedrock2_base_conversion | 0m36.52s | 1618732 ko || +0m00.88s || -1312 ko | +2.43% | -0.08% 0m34.91s | 1279244 ko | Bedrock/End2End/X25519/MontgomeryLadder.vo | 0m34.52s | 1284212 ko || +0m00.38s || -4968 ko | +1.12% | -0.38% 0m34.14s | 2240972 ko | ExtractionOCaml/bedrock2_solinas_reduction.ml | 0m34.09s | 2237916 ko || +0m00.04s || 3056 ko | +0.14% | +0.13% 0m33.11s | 2212924 ko | ExtractionOCaml/with_bedrock2_word_by_word_montgomery.ml | 0m33.05s | 2204728 ko || +0m00.06s || 8196 ko | +0.18% | +0.37% 0m33.01s | 2213528 ko | ExtractionOCaml/bedrock2_word_by_word_montgomery.ml | 0m33.12s | 2204912 ko || -0m00.10s || 8616 ko | -0.33% | +0.39% 0m32.68s | 2150644 ko | ExtractionOCaml/solinas_reduction.ml | 0m32.56s | 2139728 ko || +0m00.11s || 10916 ko | +0.36% | +0.51% 0m31.47s | 2122940 ko | ExtractionOCaml/word_by_word_montgomery.ml | 0m31.58s | 2120080 ko || -0m00.10s || 2860 ko | -0.34% | +0.13% 0m30.56s | 2114916 ko | ExtractionOCaml/bedrock2_unsaturated_solinas.ml | 0m30.25s | 2134076 ko || +0m00.30s || -19160 ko | +1.02% | -0.89% 0m30.28s | 1232236 ko | ExtractionOCaml/perf_word_by_word_montgomery | 0m29.68s | 1232100 ko || +0m00.60s || 136 ko | +2.02% | +0.01% 0m30.23s | 1232188 ko | ExtractionOCaml/perf_unsaturated_solinas | 0m29.28s | 1232224 ko || +0m00.94s || -36 ko | +3.24% | -0.00% 0m30.11s | 2114248 ko | ExtractionOCaml/with_bedrock2_unsaturated_solinas.ml | 0m30.11s | 2132752 ko || +0m00.00s || -18504 ko | +0.00% | -0.86% 0m29.92s | 1526264 ko | StandaloneDebuggingExamples.vo | 0m30.25s | 1518384 ko || -0m00.32s || 7880 ko | -1.09% | +0.51% 0m28.53s | 2051912 ko | ExtractionOCaml/unsaturated_solinas.ml | 0m28.57s | 2032036 ko || -0m00.03s || 19876 ko | -0.14% | +0.97% 0m27.17s | 2040956 ko | ExtractionOCaml/bedrock2_dettman_multiplication.ml | 0m27.20s | 2071528 ko || -0m00.02s || -30572 ko | -0.11% | -1.47% 0m26.36s | 2009348 ko | ExtractionOCaml/with_bedrock2_base_conversion.ml | 0m26.30s | 2038044 ko || +0m00.05s || -28696 ko | +0.22% | -1.40% 0m26.34s | 2019688 ko | ExtractionOCaml/with_bedrock2_saturated_solinas.ml | 0m26.08s | 1996180 ko || +0m00.26s || 23508 ko | +0.99% | +1.17% 0m26.32s | 2018796 ko | ExtractionOCaml/bedrock2_saturated_solinas.ml | 0m26.30s | 1997428 ko || +0m00.01s || 21368 ko | +0.07% | +1.06% 0m26.18s | 2009056 ko | ExtractionOCaml/bedrock2_base_conversion.ml | 0m26.28s | 2037624 ko || -0m00.10s || -28568 ko | -0.38% | -1.40% 0m26.13s | 2019564 ko | ExtractionOCaml/with_bedrock2_solinas_reduction.ml | 0m26.06s | 1996408 ko || +0m00.07s || 23156 ko | +0.26% | +1.15% 0m26.05s | 2016864 ko | ExtractionOCaml/with_bedrock2_dettman_multiplication.ml | 0m25.95s | 1997412 ko || +0m00.10s || 19452 ko | +0.38% | +0.97% 0m25.24s | 1952220 ko | ExtractionOCaml/dettman_multiplication.ml | 0m25.47s | 1939744 ko || -0m00.23s || 12476 ko | -0.90% | +0.64% 0m24.71s | 1944464 ko | ExtractionOCaml/saturated_solinas.ml | 0m24.58s | 1934196 ko || +0m00.13s || 10268 ko | +0.52% | +0.53% 0m24.68s | 1914120 ko | ExtractionOCaml/base_conversion.ml | 0m25.18s | 1922504 ko || -0m00.50s || -8384 ko | -1.98% | -0.43% 0m23.73s | 1184052 ko | PushButtonSynthesis/UnsaturatedSolinas.vo | 0m23.70s | 1185516 ko || +0m00.03s || -1464 ko | +0.12% | -0.12% 0m21.52s | 1354084 ko | Bedrock/End2End/RupicolaCrypto/Low.vo | 0m21.31s | 1350968 ko || +0m00.21s || 3116 ko | +0.98% | +0.23% 0m20.49s | 1158932 ko | PushButtonSynthesis/WordByWordMontgomery.vo | 0m20.38s | 1160828 ko || +0m00.10s || -1896 ko | +0.53% | -0.16% 0m20.15s | 794468 ko | Bedrock/Field/Translation/Proofs/Expr.vo | 0m20.23s | 794720 ko || -0m00.08s || -252 ko | -0.39% | -0.03% 0m18.98s | 1821220 ko | ExtractionOCaml/perf_unsaturated_solinas.ml | 0m18.91s | 1826352 ko || +0m00.07s || -5132 ko | +0.37% | -0.28% 0m18.80s | 1843872 ko | ExtractionOCaml/perf_word_by_word_montgomery.ml | 0m18.84s | 1867344 ko || -0m00.03s || -23472 ko | -0.21% | -1.25% 0m18.64s | 737284 ko | Bedrock/Field/Translation/Proofs/LoadStoreList.vo | 0m18.64s | 740312 ko || +0m00.00s || -3028 ko | +0.00% | -0.40% 0m18.55s | 1128568 ko | Bedrock/Field/Translation/Proofs/Func.vo | 0m18.89s | 1129264 ko || -0m00.33s || -696 ko | -1.79% | -0.06% 0m17.93s | 1207216 ko | Bedrock/Field/Synthesis/New/Signature.vo | 0m17.32s | 1210248 ko || +0m00.60s || -3032 ko | +3.52% | -0.25% 0m16.37s | 1144804 ko | Bedrock/Field/Translation/Proofs/Cmd.vo | 0m17.11s | 1146536 ko || -0m00.73s || -1732 ko | -4.32% | -0.15% 0m16.36s | 2063348 ko | ExtractionHaskell/with_bedrock2_solinas_reduction.hs | 0m16.35s | 2096996 ko || +0m00.00s || -33648 ko | +0.06% | -1.60% 0m16.26s | 2063264 ko | ExtractionHaskell/bedrock2_solinas_reduction.hs | 0m16.52s | 2095492 ko || -0m00.25s || -32228 ko | -1.57% | -1.53% 0m16.16s | 2023412 ko | ExtractionHaskell/with_bedrock2_word_by_word_montgomery.hs | 0m16.00s | 2007884 ko || +0m00.16s || 15528 ko | +1.00% | +0.77% 0m16.12s | 2021956 ko | ExtractionHaskell/bedrock2_word_by_word_montgomery.hs | 0m16.13s | 2010240 ko || -0m00.00s || 11716 ko | -0.06% | +0.58% 0m16.08s | 2028764 ko | ExtractionHaskell/solinas_reduction.hs | 0m15.49s | 1949544 ko || +0m00.58s || 79220 ko | +3.80% | +4.06% 0m15.71s | 1962892 ko | ExtractionHaskell/word_by_word_montgomery.hs | 0m15.47s | 1985472 ko || +0m00.24s || -22580 ko | +1.55% | -1.13% 0m15.59s | 2009872 ko | ExtractionHaskell/bedrock2_unsaturated_solinas.hs | 0m15.14s | 1950900 ko || +0m00.44s || 58972 ko | +2.97% | +3.02% 0m15.21s | 2008904 ko | ExtractionHaskell/with_bedrock2_unsaturated_solinas.hs | 0m15.24s | 1949824 ko || -0m00.02s || 59080 ko | -0.19% | +3.03% 0m14.63s | 1934444 ko | ExtractionHaskell/bedrock2_saturated_solinas.hs | 0m14.22s | 1862544 ko || +0m00.41s || 71900 ko | +2.88% | +3.86% 0m14.47s | 1879068 ko | ExtractionHaskell/bedrock2_dettman_multiplication.hs | 0m14.36s | 1889216 ko || +0m00.11s || -10148 ko | +0.76% | -0.53% 0m14.38s | 1879364 ko | ExtractionHaskell/with_bedrock2_dettman_multiplication.hs | 0m14.28s | 1886884 ko || +0m00.10s || -7520 ko | +0.70% | -0.39% 0m14.37s | 1927472 ko | ExtractionHaskell/unsaturated_solinas.hs | 0m14.53s | 1928704 ko || -0m00.16s || -1232 ko | -1.10% | -0.06% 0m14.32s | 1917200 ko | ExtractionHaskell/with_bedrock2_base_conversion.hs | 0m14.11s | 1932764 ko || +0m00.21s || -15564 ko | +1.48% | -0.80% 0m14.30s | 1935052 ko | ExtractionHaskell/with_bedrock2_saturated_solinas.hs | 0m14.12s | 1861568 ko || +0m00.18s || 73484 ko | +1.27% | +3.94% 0m14.20s | 1917124 ko | ExtractionHaskell/bedrock2_base_conversion.hs | 0m14.22s | 1911084 ko || -0m00.02s || 6040 ko | -0.14% | +0.31% 0m14.06s | 1864296 ko | ExtractionHaskell/dettman_multiplication.hs | 0m13.20s | 1856724 ko || +0m00.86s || 7572 ko | +6.51% | +0.40% 0m13.78s | 1840872 ko | ExtractionHaskell/saturated_solinas.hs | 0m13.71s | 1846748 ko || +0m00.06s || -5876 ko | +0.51% | -0.31% 0m13.68s | 602488 ko | Bedrock/Field/Common/Util.vo | 0m13.68s | 602420 ko || +0m00.00s || 68 ko | +0.00% | +0.01% 0m13.36s | 663048 ko | Bedrock/Group/AdditionChains.vo | 0m13.42s | 662928 ko || -0m00.06s || 120 ko | -0.44% | +0.01% 0m13.12s | 667092 ko | Bedrock/Group/ScalarMult/LadderStep.vo | 0m13.04s | 667048 ko || +0m00.08s || 44 ko | +0.61% | +0.00% 0m13.11s | 1845960 ko | ExtractionHaskell/base_conversion.hs | 0m13.55s | 1851580 ko || -0m00.44s || -5620 ko | -3.24% | -0.30% 0m11.84s | 1029560 ko | BoundsPipeline.vo | 0m11.88s | 1028872 ko || -0m00.04s || 688 ko | -0.33% | +0.06% 0m11.31s | 1699940 ko | Bedrock/Field/Synthesis/New/WordByWordMontgomery.vo | 0m11.11s | 1698268 ko || +0m00.20s || 1672 ko | +1.80% | +0.09% 0m10.27s | 1297740 ko | Bedrock/End2End/X25519/MontgomeryLadderProperties.vo | 0m10.36s | 1296660 ko || -0m00.08s || 1080 ko | -0.86% | +0.08% 0m09.88s | 600056 ko | Stringification/IR.vo | 0m09.90s | 601448 ko || -0m00.01s || -1392 ko | -0.20% | -0.23% 0m09.74s | 624024 ko | Bedrock/Field/Translation/Proofs/Flatten.vo | 0m09.73s | 624068 ko || +0m00.00s || -44 ko | +0.10% | -0.00% 0m09.34s | 1032068 ko | PushButtonSynthesis/BaseConversion.vo | 0m09.37s | 1035460 ko || -0m00.02s || -3392 ko | -0.32% | -0.32% 0m09.19s | 661364 ko | Bedrock/Group/ScalarMult/CSwap.vo | 0m08.96s | 661296 ko || +0m00.22s || 68 ko | +2.56% | +0.01% 0m08.41s | 1044384 ko | PushButtonSynthesis/Primitives.vo | 0m08.56s | 1045468 ko || -0m00.15s || -1084 ko | -1.75% | -0.10% 0m08.33s | 1000952 ko | PushButtonSynthesis/SmallExamples.vo | 0m08.13s | 1001028 ko || +0m00.19s || -76 ko | +2.46% | -0.00% 0m07.39s | 911356 ko | Bedrock/Field/Translation/Proofs/EquivalenceProperties.vo | 0m07.43s | 912204 ko || -0m00.04s || -848 ko | -0.53% | -0.09% 0m07.17s | 1033520 ko | PushButtonSynthesis/SolinasReduction.vo | 0m07.24s | 1033476 ko || -0m00.07s || 44 ko | -0.96% | +0.00% 0m06.67s | 51208 ko | fiat-go/64/p521/p521.go | 0m06.32s | 49224 ko || +0m00.34s || 1984 ko | +5.53% | +4.03% 0m06.40s | 1040660 ko | PushButtonSynthesis/BarrettReduction.vo | 0m06.42s | 1041508 ko || -0m00.01s || -848 ko | -0.31% | -0.08% 0m06.27s | 1121320 ko | CLI.vo | 0m06.40s | 1121288 ko || -0m00.13s || 32 ko | -2.03% | +0.00% 0m05.98s | 1130956 ko | Bedrock/Field/Synthesis/New/UnsaturatedSolinas.vo | 0m06.11s | 1130988 ko || -0m00.12s || -32 ko | -2.12% | -0.00% 0m05.92s | 68612 ko | fiat-bedrock2/src/p521_64.c | 0m04.98s | 66368 ko || +0m00.93s || 2244 ko | +18.87% | +3.38% 0m05.90s | 37028 ko | fiat-rust/src/p521_64.rs | 0m05.46s | 36076 ko || +0m00.44s || 952 ko | +8.05% | +2.63% 0m05.87s | 37052 ko | fiat-c/src/p521_64.c | 0m05.41s | 36076 ko || +0m00.46s || 976 ko | +8.50% | +2.70% 0m05.84s | 906864 ko | Bedrock/Field/Translation/Proofs/UsedVarnames.vo | 0m05.81s | 906744 ko || +0m00.03s || 120 ko | +0.51% | +0.01% 0m05.25s | 616380 ko | Bedrock/End2End/RupicolaCrypto/Broadcast.vo | 0m05.24s | 616216 ko || +0m00.00s || 164 ko | +0.19% | +0.02% 0m04.73s | 1030392 ko | PushButtonSynthesis/DettmanMultiplication.vo | 0m04.78s | 1030456 ko || -0m00.04s || -64 ko | -1.04% | -0.00% 0m04.41s | 1065452 ko | Bedrock/Field/Translation/Proofs/ValidComputable/Cmd.vo | 0m04.56s | 1065272 ko || -0m00.14s || 180 ko | -3.28% | +0.01% 0m04.39s | 1038800 ko | PushButtonSynthesis/SaturatedSolinas.vo | 0m04.43s | 1038816 ko || -0m00.04s || -16 ko | -0.90% | -0.00% 0m03.94s | 957224 ko | Assembly/Equivalence.vo | 0m03.88s | 955376 ko || +0m00.06s || 1848 ko | +1.54% | +0.19% 0m03.84s | 1051032 ko | PushButtonSynthesis/FancyMontgomeryReduction.vo | 0m03.94s | 1051128 ko || -0m00.10s || -96 ko | -2.53% | -0.00% 0m03.66s | 666368 ko | Bedrock/Group/ScalarMult/ScalarMult.vo | 0m03.97s | 667152 ko || -0m00.31s || -784 ko | -7.80% | -0.11% 0m03.12s | 1072356 ko | Rewriter/PerfTesting/Core.vo | 0m03.09s | 1072420 ko || +0m00.03s || -64 ko | +0.97% | -0.00% 0m03.06s | 617024 ko | Bedrock/Field/Synthesis/Generic/Bignum.vo | 0m03.02s | 613720 ko || +0m00.04s || 3304 ko | +1.32% | +0.53% 0m03.05s | 1006464 ko | Bedrock/Field/Translation/Cmd.vo | 0m03.08s | 1009700 ko || -0m00.03s || -3236 ko | -0.97% | -0.32% 0m02.98s | 50852 ko | fiat-go/64/p448solinas/p448solinas.go | 0m02.67s | 50500 ko || +0m00.31s || 352 ko | +11.61% | +0.69% 0m02.88s | 1068264 ko | Bedrock/Field/Stringification/Stringification.vo | 0m02.87s | 1068932 ko || +0m00.00s || -668 ko | +0.34% | -0.06% 0m02.84s | 1061220 ko | Bedrock/Field/Translation/Proofs/ValidComputable/Func.vo | 0m02.73s | 1059288 ko || +0m00.10s || 1932 ko | +4.02% | +0.18% 0m02.80s | 1004228 ko | Bedrock/Field/Translation/Func.vo | 0m02.79s | 1004848 ko || +0m00.00s || -620 ko | +0.35% | -0.06% 0m02.80s | 1103608 ko | StandaloneHaskellMain.vo | 0m02.72s | 1103504 ko || +0m00.07s || 104 ko | +2.94% | +0.00% 0m02.78s | 1145260 ko | Bedrock/Standalone/StandaloneHaskellMain.vo | 0m02.77s | 1145216 ko || +0m00.00s || 44 ko | +0.36% | +0.00% 0m02.76s | 1145200 ko | Bedrock/Standalone/StandaloneOCamlMain.vo | 0m02.72s | 1145392 ko || +0m00.03s || -192 ko | +1.47% | -0.01% 0m02.76s | 1131284 ko | Rewriter/PerfTesting/StandaloneOCamlMain.vo | 0m02.73s | 1131224 ko || +0m00.02s || 60 ko | +1.09% | +0.00% 0m02.66s | 1103772 ko | StandaloneOCamlMain.vo | 0m02.68s | 1103860 ko || -0m00.02s || -88 ko | -0.74% | -0.00% 0m02.60s | 620876 ko | Bedrock/Field/Interface/Compilation2.vo | 0m02.60s | 621772 ko || +0m00.00s || -896 ko | +0.00% | -0.14% 0m02.54s | 1064964 ko | Bedrock/Field/Synthesis/New/ComputedOp.vo | 0m02.59s | 1065008 ko || -0m00.04s || -44 ko | -1.93% | -0.00% 0m02.53s | 65044 ko | fiat-bedrock2/src/p448_solinas_64.c | 0m02.07s | 66308 ko || +0m00.46s || -1264 ko | +22.22% | -1.90% 0m02.52s | 35764 ko | fiat-zig/src/p448_solinas_64.zig | 0m01.90s | 35084 ko || +0m00.62s || 680 ko | +32.63% | +1.93% 0m02.51s | 37100 ko | fiat-go/32/curve25519/curve25519.go | 0m02.15s | 35348 ko || +0m00.35s || 1752 ko | +16.74% | +4.95% 0m02.51s | 36324 ko | fiat-rust/src/p448_solinas_64.rs | 0m01.97s | 34328 ko || +0m00.53s || 1996 ko | +27.41% | +5.81% 0m02.50s | 1044712 ko | Bedrock/Field/Translation/Parameters/Defaults32.vo | 0m02.53s | 1045368 ko || -0m00.02s || -656 ko | -1.18% | -0.06% 0m02.48s | 1041488 ko | Bedrock/Field/Translation/Parameters/Defaults.vo | 0m02.40s | 1040684 ko || +0m00.08s || 804 ko | +3.33% | +0.07% 0m02.48s | 1045392 ko | Bedrock/Field/Translation/Parameters/Defaults64.vo | 0m02.56s | 1044600 ko || -0m00.08s || 792 ko | -3.12% | +0.07% 0m02.47s | 51164 ko | fiat-json/src/p448_solinas_64.json | 0m02.06s | 51416 ko || +0m00.41s || -252 ko | +19.90% | -0.49% 0m02.45s | 36696 ko | fiat-c/src/p448_solinas_64.c | 0m01.98s | 35380 ko || +0m00.47s || 1316 ko | +23.73% | +3.71% 0m02.44s | 1045352 ko | Bedrock/Field/Translation/Parameters/FE310.vo | 0m02.50s | 1045320 ko || -0m00.06s || 32 ko | -2.40% | +0.00% 0m02.24s | 63200 ko | fiat-bedrock2/src/curve25519_32.c | 0m01.98s | 63640 ko || +0m00.26s || -440 ko | +13.13% | -0.69% 0m02.18s | 35508 ko | fiat-java/src/FiatCurve25519.java | 0m01.82s | 36520 ko || +0m00.36s || -1012 ko | +19.78% | -2.77% 0m02.16s | 50436 ko | fiat-json/src/curve25519_32.json | 0m01.89s | 49708 ko || +0m00.27s || 728 ko | +14.28% | +1.46% 0m02.15s | 34348 ko | fiat-zig/src/curve25519_32.zig | 0m01.82s | 33388 ko || +0m00.32s || 960 ko | +18.13% | +2.87% 0m02.12s | 34960 ko | fiat-rust/src/curve25519_32.rs | 0m01.79s | 34088 ko || +0m00.33s || 872 ko | +18.43% | +2.55% 0m02.10s | 35016 ko | fiat-c/src/curve25519_32.c | 0m01.76s | 33112 ko || +0m00.34s || 1904 ko | +19.31% | +5.75% 0m01.96s | 616180 ko | Bedrock/Field/Common/Arrays/MakeAccessSizes.vo | 0m02.00s | 614096 ko || -0m00.04s || 2084 ko | -2.00% | +0.33% 0m01.78s | 636112 ko | Bedrock/Group/ScalarMult/MontgomeryEquivalence.vo | 0m01.82s | 636072 ko || -0m00.04s || 40 ko | -2.19% | +0.00% 0m01.62s | 52504 ko | fiat-bedrock2/src/secp256k1_dettman_32.c | 0m01.46s | 51924 ko || +0m00.16s || 580 ko | +10.95% | +1.11% 0m01.61s | 607916 ko | Bedrock/Field/Common/Names/MakeNames.vo | 0m01.62s | 607756 ko || -0m00.01s || 160 ko | -0.61% | +0.02% 0m01.53s | 590192 ko | CompilersTestCases.vo | 0m01.48s | 590236 ko || +0m00.05s || -44 ko | +3.37% | -0.00% 0m01.50s | 512068 ko | AbstractInterpretation/AbstractInterpretation.vo | 0m01.39s | 512000 ko || +0m00.11s || 68 ko | +7.91% | +0.01% 0m01.46s | 26680 ko | fiat-java/src/FiatSecp256K1Dettman.java | 0m01.24s | 26336 ko || +0m00.21s || 344 ko | +17.74% | +1.30% 0m01.46s | 42060 ko | fiat-json/src/secp256k1_dettman_32.json | 0m01.24s | 41136 ko || +0m00.21s || 924 ko | +17.74% | +2.24% 0m01.44s | 26260 ko | fiat-c/src/secp256k1_dettman_32.c | 0m01.21s | 25744 ko || +0m00.23s || 516 ko | +19.00% | +2.00% 0m01.44s | 27956 ko | fiat-go/32/secp256k1dettman/secp256k1dettman.go | 0m01.21s | 25860 ko || +0m00.23s || 2096 ko | +19.00% | +8.10% 0m01.44s | 27096 ko | fiat-rust/src/secp256k1_dettman_32.rs | 0m01.22s | 25440 ko || +0m00.21s || 1656 ko | +18.03% | +6.50% 0m01.42s | 27060 ko | fiat-zig/src/secp256k1_dettman_32.zig | 0m01.23s | 25656 ko || +0m00.18s || 1404 ko | +15.44% | +5.47% 0m01.35s | 543324 ko | Stringification/Go.vo | 0m01.35s | 541960 ko || +0m00.00s || 1364 ko | +0.00% | +0.25% 0m01.11s | 628136 ko | Bedrock/Specs/Field.vo | 0m01.18s | 628700 ko || -0m00.06s || -564 ko | -5.93% | -0.08% 0m01.07s | 609304 ko | Bedrock/Field/Common/Arrays/MaxBounds.vo | 0m01.03s | 607272 ko || +0m00.04s || 2032 ko | +3.88% | +0.33% 0m01.06s | 601084 ko | Bedrock/Field/Common/Arrays/ByteBounds.vo | 0m01.00s | 602440 ko || +0m00.06s || -1356 ko | +6.00% | -0.22% 0m00.94s | 538600 ko | Stringification/C.vo | 0m00.91s | 538560 ko || +0m00.02s || 40 ko | +3.29% | +0.00% 0m00.91s | 535844 ko | Stringification/JSON.vo | 0m00.94s | 537144 ko || -0m00.02s || -1300 ko | -3.19% | -0.24% 0m00.88s | 535936 ko | Stringification/Zig.vo | 0m00.88s | 539224 ko || +0m00.00s || -3288 ko | +0.00% | -0.60% 0m00.87s | 618888 ko | Bedrock/Group/Point.vo | 0m00.89s | 617632 ko || -0m00.02s || 1256 ko | -2.24% | +0.20% 0m00.86s | 615616 ko | Bedrock/Field/Interface/Representation.vo | 0m00.84s | 615736 ko || +0m00.02s || -120 ko | +2.38% | -0.01% 0m00.86s | 534896 ko | Stringification/Java.vo | 0m00.84s | 538936 ko || +0m00.02s || -4040 ko | +2.38% | -0.74% 0m00.80s | 535916 ko | Stringification/Rust.vo | 0m00.83s | 535824 ko || -0m00.02s || 92 ko | -3.61% | +0.01% 0m00.79s | 587772 ko | Bedrock/Field/Common/Tactics.vo | 0m00.78s | 587592 ko || +0m00.01s || 180 ko | +1.28% | +0.03% 0m00.72s | 521776 ko | Bedrock/Field/Common/Arrays/MakeListLengths.vo | 0m00.68s | 523076 ko || +0m00.03s || -1300 ko | +5.88% | -0.24% 0m00.72s | 546832 ko | Bedrock/Field/Stringification/FlattenVarData.vo | 0m00.72s | 546004 ko || +0m00.00s || 828 ko | +0.00% | +0.15% 0m00.70s | 32124 ko | fiat-go/64/curve25519/curve25519.go | 0m00.63s | 32976 ko || +0m00.06s || -852 ko | +11.11% | -2.58% 0m00.68s | 515488 ko | AbstractInterpretation/WfExtra.vo | 0m00.63s | 515492 ko || +0m00.05s || -4 ko | +7.93% | -0.00% 0m00.68s | 38372 ko | fiat-json/src/curve25519_solinas_64.json | 0m00.41s | 39308 ko || +0m00.27s || -936 ko | +65.85% | -2.38% 0m00.67s | 38908 ko | fiat-c/src/curve25519_solinas_64.c | 0m00.41s | 36728 ko || +0m00.26s || 2180 ko | +63.41% | +5.93% 0m00.66s | 37132 ko | fiat-rust/src/curve25519_solinas_64.rs | 0m00.39s | 36260 ko || +0m00.27s || 872 ko | +69.23% | +2.40% 0m00.66s | 36892 ko | fiat-zig/src/curve25519_solinas_64.zig | 0m00.42s | 36336 ko || +0m00.24s || 556 ko | +57.14% | +1.53% 0m00.60s | 34160 ko | fiat-json/src/curve25519_64.json | 0m00.52s | 32588 ko || +0m00.07s || 1572 ko | +15.38% | +4.82% 0m00.59s | 38236 ko | fiat-go/64/curve25519solinas/curve25519solinas.go | 0m00.43s | 36420 ko || +0m00.15s || 1816 ko | +37.20% | +4.98% 0m00.58s | 38448 ko | fiat-bedrock2/src/curve25519_64.c | 0m00.51s | 38216 ko || +0m00.06s || 232 ko | +13.72% | +0.60% 0m00.58s | 42568 ko | fiat-bedrock2/src/curve25519_solinas_64.c | 0m00.42s | 40596 ko || +0m00.15s || 1972 ko | +38.09% | +4.85% 0m00.56s | 26892 ko | fiat-c/src/curve25519_64.c | 0m00.45s | 27132 ko || +0m00.11s || -240 ko | +24.44% | -0.88% 0m00.56s | 26560 ko | fiat-rust/src/curve25519_64.rs | 0m00.46s | 26040 ko || +0m00.10s || 520 ko | +21.73% | +1.99% 0m00.56s | 27152 ko | fiat-zig/src/curve25519_64.zig | 0m00.48s | 26252 ko || +0m00.08s || 900 ko | +16.66% | +3.42% 0m00.54s | 120576 ko | ExtractionOCaml/with_bedrock2_base_conversion.cmi | 0m00.50s | 119948 ko || +0m00.04s || 628 ko | +8.00% | +0.52% 0m00.53s | 120076 ko | ExtractionOCaml/unsaturated_solinas.cmi | 0m00.50s | 119836 ko || +0m00.03s || 240 ko | +6.00% | +0.20% 0m00.52s | 124368 ko | ExtractionOCaml/bedrock2_word_by_word_montgomery.cmi | 0m00.52s | 123300 ko || +0m00.00s || 1068 ko | +0.00% | +0.86% 0m00.52s | 118308 ko | ExtractionOCaml/solinas_reduction.cmi | 0m00.51s | 118636 ko || +0m00.01s || -328 ko | +1.96% | -0.27% 0m00.52s | 121724 ko | ExtractionOCaml/with_bedrock2_dettman_multiplication.cmi | 0m00.50s | 121156 ko || +0m00.02s || 568 ko | +4.00% | +0.46% 0m00.52s | 123672 ko | ExtractionOCaml/with_bedrock2_saturated_solinas.cmi | 0m00.51s | 120980 ko || +0m00.01s || 2692 ko | +1.96% | +2.22% 0m00.52s | 123072 ko | ExtractionOCaml/with_bedrock2_solinas_reduction.cmi | 0m00.50s | 122784 ko || +0m00.02s || 288 ko | +4.00% | +0.23% 0m00.52s | 123132 ko | ExtractionOCaml/with_bedrock2_unsaturated_solinas.cmi | 0m00.50s | 123700 ko || +0m00.02s || -568 ko | +4.00% | -0.45% 0m00.51s | 120332 ko | ExtractionOCaml/bedrock2_base_conversion.cmi | 0m00.51s | 119972 ko || +0m00.00s || 360 ko | +0.00% | +0.30% 0m00.51s | 120184 ko | ExtractionOCaml/bedrock2_unsaturated_solinas.cmi | 0m00.53s | 123596 ko || -0m00.02s || -3412 ko | -3.77% | -2.76% 0m00.51s | 118088 ko | ExtractionOCaml/saturated_solinas.cmi | 0m00.48s | 117928 ko || +0m00.03s || 160 ko | +6.25% | +0.13% 0m00.51s | 124416 ko | ExtractionOCaml/with_bedrock2_word_by_word_montgomery.cmi | 0m00.51s | 123124 ko || +0m00.00s || 1292 ko | +0.00% | +1.04% 0m00.50s | 120256 ko | ExtractionOCaml/base_conversion.cmi | 0m00.51s | 118684 ko || -0m00.01s || 1572 ko | -1.96% | +1.32% 0m00.50s | 123568 ko | ExtractionOCaml/bedrock2_saturated_solinas.cmi | 0m00.54s | 122468 ko || -0m00.04s || 1100 ko | -7.40% | +0.89% 0m00.50s | 123660 ko | ExtractionOCaml/bedrock2_solinas_reduction.cmi | 0m00.50s | 120332 ko || +0m00.00s || 3328 ko | +0.00% | +2.76% 0m00.50s | 119964 ko | ExtractionOCaml/dettman_multiplication.cmi | 0m00.51s | 120092 ko || -0m00.01s || -128 ko | -1.96% | -0.10% 0m00.50s | 120424 ko | ExtractionOCaml/word_by_word_montgomery.cmi | 0m00.49s | 119588 ko || +0m00.01s || 836 ko | +2.04% | +0.69% 0m00.49s | 120784 ko | ExtractionOCaml/bedrock2_dettman_multiplication.cmi | 0m00.53s | 120180 ko || -0m00.04s || 604 ko | -7.54% | +0.50% 0m00.34s | 25480 ko | fiat-go/32/poly1305/poly1305.go | 0m00.30s | 24960 ko || +0m00.04s || 520 ko | +13.33% | +2.08% 0m00.31s | 25196 ko | fiat-go/64/secp256k1dettman/secp256k1dettman.go | 0m00.25s | 24476 ko || +0m00.06s || 720 ko | +24.00% | +2.94% 0m00.30s | 33676 ko | fiat-bedrock2/src/poly1305_32.c | 0m00.25s | 32372 ko || +0m00.04s || 1304 ko | +19.99% | +4.02% 0m00.30s | 29656 ko | fiat-json/src/poly1305_32.json | 0m00.25s | 30252 ko || +0m00.04s || -596 ko | +19.99% | -1.97% 0m00.29s | 29300 ko | fiat-bedrock2/src/secp256k1_dettman_64.c | 0m00.22s | 29724 ko || +0m00.06s || -424 ko | +31.81% | -1.42% 0m00.28s | 23932 ko | fiat-json/src/secp256k1_dettman_64.json | 0m00.20s | 23928 ko || +0m00.08s || 4 ko | +40.00% | +0.01% 0m00.27s | 25000 ko | fiat-c/src/poly1305_32.c | 0m00.21s | 24124 ko || +0m00.06s || 876 ko | +28.57% | +3.63% 0m00.27s | 21060 ko | fiat-c/src/secp256k1_dettman_64.c | 0m00.18s | 20924 ko || +0m00.09s || 136 ko | +50.00% | +0.64% 0m00.27s | 25404 ko | fiat-java/src/FiatPoly1305.java | 0m00.22s | 25272 ko || +0m00.05s || 132 ko | +22.72% | +0.52% 0m00.27s | 20960 ko | fiat-rust/src/secp256k1_dettman_64.rs | 0m00.18s | 21372 ko || +0m00.09s || -412 ko | +50.00% | -1.92% 0m00.27s | 24624 ko | fiat-zig/src/poly1305_32.zig | 0m00.21s | 23860 ko || +0m00.06s || 764 ko | +28.57% | +3.20% 0m00.27s | 21176 ko | fiat-zig/src/secp256k1_dettman_64.zig | 0m00.18s | 20508 ko || +0m00.09s || 668 ko | +50.00% | +3.25% 0m00.26s | 24636 ko | fiat-rust/src/poly1305_32.rs | 0m00.21s | 24504 ko || +0m00.05s || 132 ko | +23.80% | +0.53% 0m00.18s | 26440 ko | fiat-go/64/poly1305/poly1305.go | 0m00.18s | 26096 ko || +0m00.00s || 344 ko | +0.00% | +1.31% 0m00.17s | 26980 ko | fiat-json/src/poly1305_64.json | 0m00.13s | 27492 ko || +0m00.04s || -512 ko | +30.76% | -1.86% 0m00.16s | 61504 ko | ExtractionOCaml/perf_unsaturated_solinas.cmi | 0m00.16s | 61784 ko || +0m00.00s || -280 ko | +0.00% | -0.45% 0m00.16s | 23780 ko | fiat-rust/src/poly1305_64.rs | 0m00.13s | 22996 ko || +0m00.03s || 784 ko | +23.07% | +3.40% 0m00.16s | 23284 ko | fiat-zig/src/poly1305_64.zig | 0m00.13s | 23176 ko || +0m00.03s || 108 ko | +23.07% | +0.46% 0m00.15s | 58712 ko | ExtractionOCaml/perf_word_by_word_montgomery.cmi | 0m00.15s | 58904 ko || +0m00.00s || -192 ko | +0.00% | -0.32% 0m00.15s | 28228 ko | fiat-bedrock2/src/poly1305_64.c | 0m00.11s | 28056 ko || +0m00.03s || 172 ko | +36.36% | +0.61% 0m00.15s | 23928 ko | fiat-c/src/poly1305_64.c | 0m00.12s | 22980 ko || +0m00.03s || 948 ko | +25.00% | +4.12% ```

* Remove a needless `GeneralizeVar.GeneralizeVar` Maybe this is responsible for the performance issue???
Timing Diff

``` After | Peak Mem | File Name | Before | Peak Mem || Change || Change (mem) | % Change | % Change (mem) ------------------------------------------------------------------------------------------------------------------------------------------------------------------------------------ 5494m57.81s | 2858552 ko | Total Time / Peak Mem | 120m21.13s | 2852328 ko || +5374m36.68s || 6224 ko | +4465.73% | +0.21% ------------------------------------------------------------------------------------------------------------------------------------------------------------------------------------ 352m04.93s | 1781524 ko | fiat-rust/src/p384_scalar_32.rs | 1m53.74s | 1628856 ko || +350m11.18s || 152668 ko | +18472.99% | +9.37% 351m13.49s | 1916288 ko | fiat-java/src/FiatP384Scalar.java | 1m53.23s | 1620248 ko || +349m20.26s || 296040 ko | +18511.22% | +18.27% 344m52.23s | 1666100 ko | fiat-go/32/p384scalar/p384scalar.go | 1m53.04s | 1734424 ko || +342m59.18s || -68324 ko | +18205.22% | -3.93% 343m32.96s | 1898556 ko | fiat-c/src/p384_scalar_32.c | 1m51.27s | 1627512 ko || +341m41.68s || 271044 ko | +18425.17% | +16.65% 341m26.05s | 1834040 ko | fiat-json/src/p384_scalar_32.json | 1m46.90s | 1802484 ko || +339m39.14s || 31556 ko | +19063.75% | +1.75% 336m05.54s | 1870632 ko | fiat-rust/src/p384_32.rs | 1m50.16s | 1709428 ko || +334m15.38s || 161204 ko | +18205.68% | +9.43% 334m29.07s | 1989180 ko | fiat-java/src/FiatP384.java | 1m52.50s | 1817848 ko || +332m36.56s || 171332 ko | +17739.17% | +9.42% 333m35.10s | 1806688 ko | fiat-json/src/p384_32.json | 1m53.94s | 1629284 ko || +331m41.15s || 177404 ko | +17466.35% | +10.88% 328m36.71s | 1779836 ko | fiat-bedrock2/src/p384_scalar_32.c | 1m48.68s | 1758232 ko || +326m48.02s || 21604 ko | +18041.98% | +1.22% 328m39.19s | 1662676 ko | fiat-go/32/p384/p384.go | 1m53.01s | 1625684 ko || +326m46.18s || 36992 ko | +17349.06% | +2.27% 321m01.71s | 1709592 ko | fiat-c/src/p384_32.c | 1m49.66s | 2029588 ko || +319m12.04s || -319996 ko | +17464.93% | -15.76% 320m33.13s | 1872256 ko | fiat-bedrock2/src/p384_32.c | 1m48.69s | 1687840 ko || +318m44.44s || 184416 ko | +17595.39% | +10.92% 318m14.38s | 2079716 ko | fiat-zig/src/p384_scalar_32.zig | 1m54.15s | 1960684 ko || +316m20.22s || 119032 ko | +16627.44% | +6.07% 308m20.28s | 1645516 ko | fiat-zig/src/p384_32.zig | 1m52.66s | 1948020 ko || +306m27.61s || -302504 ko | +16321.33% | -15.52% 140m13.96s | 2858552 ko | Bedrock/Field/Synthesis/Examples/p224_64_new.vo | 5m29.46s | 2852328 ko || +134m44.49s || 6224 ko | +2453.86% | +0.21% 15m08.73s | 412980 ko | fiat-zig/src/secp256k1_montgomery_scalar_32.zig | 0m17.06s | 365400 ko || +14m51.67s || 47580 ko | +5226.67% | +13.02% 15m07.92s | 444232 ko | fiat-rust/src/secp256k1_montgomery_scalar_32.rs | 0m16.58s | 413372 ko || +14m51.33s || 30860 ko | +5375.99% | +7.46% 15m07.44s | 447232 ko | fiat-bedrock2/src/secp256k1_montgomery_scalar_32.c | 0m17.59s | 396628 ko || +14m49.85s || 50604 ko | +5058.84% | +12.75% 15m05.89s | 413980 ko | fiat-zig/src/p256_scalar_32.zig | 0m17.13s | 404448 ko || +14m48.75s || 9532 ko | +5188.32% | +2.35% 15m00.09s | 462000 ko | fiat-json/src/secp256k1_montgomery_scalar_32.json | 0m17.14s | 372928 ko || +14m42.95s || 89072 ko | +5151.40% | +23.88% 14m59.00s | 383440 ko | fiat-go/32/secp256k1montgomeryscalar/secp256k1montgomeryscalar.go | 0m16.75s | 439368 ko || +14m42.25s || -55928 ko | +5267.16% | -12.72% 14m58.83s | 417168 ko | fiat-java/src/FiatSecp256K1MontgomeryScalar.java | 0m17.43s | 335352 ko || +14m41.40s || 81816 ko | +5056.79% | +24.39% 14m57.88s | 378000 ko | fiat-go/32/p256scalar/p256scalar.go | 0m16.89s | 362996 ko || +14m40.99s || 15004 ko | +5216.04% | +4.13% 14m43.11s | 453008 ko | fiat-bedrock2/src/curve25519_scalar_32.c | 0m16.87s | 404488 ko || +14m26.24s || 48520 ko | +5134.79% | +11.99% 14m42.26s | 423172 ko | fiat-zig/src/curve25519_scalar_32.zig | 0m16.75s | 338768 ko || +14m25.50s || 84404 ko | +5167.22% | +24.91% 14m34.97s | 386392 ko | fiat-java/src/FiatSecp256K1Montgomery.java | 0m16.92s | 375048 ko || +14m18.05s || 11344 ko | +5071.21% | +3.02% 14m35.14s | 439764 ko | fiat-bedrock2/src/p256_scalar_32.c | 0m17.47s | 466176 ko || +14m17.66s || -26412 ko | +4909.38% | -5.66% 14m33.96s | 418052 ko | fiat-c/src/p256_scalar_32.c | 0m16.30s | 433696 ko || +14m17.66s || -15644 ko | +5261.71% | -3.60% 14m32.17s | 379176 ko | fiat-zig/src/secp256k1_montgomery_32.zig | 0m16.64s | 440088 ko || +14m15.52s || -60912 ko | +5141.40% | -13.84% 14m32.26s | 386104 ko | fiat-java/src/FiatP256Scalar.java | 0m17.60s | 395220 ko || +14m14.65s || -9116 ko | +4856.02% | -2.30% 14m20.14s | 476544 ko | fiat-json/src/secp256k1_montgomery_32.json | 0m16.80s | 384668 ko || +14m03.34s || 91876 ko | +5019.88% | +23.88% 14m18.94s | 380196 ko | fiat-java/src/FiatCurve25519Scalar.java | 0m16.99s | 369068 ko || +14m01.95s || 11128 ko | +4955.56% | +3.01% 14m14.20s | 438032 ko | fiat-bedrock2/src/secp256k1_montgomery_32.c | 0m16.55s | 414072 ko || +13m57.65s || 23960 ko | +5061.32% | +5.78% 14m10.18s | 382024 ko | fiat-c/src/secp256k1_montgomery_scalar_32.c | 0m16.48s | 349956 ko || +13m53.69s || 32068 ko | +5058.85% | +9.16% 14m07.30s | 380804 ko | fiat-go/32/secp256k1montgomery/secp256k1montgomery.go | 0m15.85s | 343344 ko || +13m51.44s || 37460 ko | +5245.74% | +10.91% 14m05.83s | 384216 ko | fiat-rust/src/p256_scalar_32.rs | 0m16.99s | 347112 ko || +13m48.84s || 37104 ko | +4878.39% | +10.68% 14m02.24s | 421196 ko | fiat-rust/src/secp256k1_montgomery_32.rs | 0m16.59s | 342720 ko || +13m45.64s || 78476 ko | +4976.79% | +22.89% 13m55.86s | 416884 ko | fiat-bedrock2/src/p256_32.c | 0m15.76s | 363984 ko || +13m40.10s || 52900 ko | +5203.68% | +14.53% 13m49.51s | 435140 ko | fiat-zig/src/p256_32.zig | 0m15.90s | 375036 ko || +13m33.61s || 60104 ko | +5117.04% | +16.02% 13m48.55s | 402868 ko | fiat-json/src/curve25519_scalar_32.json | 0m16.09s | 444636 ko || +13m32.45s || -41768 ko | +5049.47% | -9.39% 13m47.62s | 374572 ko | fiat-c/src/curve25519_scalar_32.c | 0m15.90s | 352232 ko || +13m31.72s || 22340 ko | +5105.15% | +6.34% 13m40.58s | 411000 ko | fiat-c/src/secp256k1_montgomery_32.c | 0m15.54s | 361600 ko || +13m25.04s || 49400 ko | +5180.43% | +13.66% 13m38.90s | 391764 ko | fiat-java/src/FiatP256.java | 0m15.69s | 392176 ko || +13m23.20s || -412 ko | +5119.24% | -0.10% 13m28.98s | 382116 ko | fiat-go/32/p256/p256.go | 0m15.69s | 361040 ko || +13m13.28s || 21076 ko | +5056.02% | +5.83% 13m24.91s | 408552 ko | fiat-rust/src/p256_32.rs | 0m15.57s | 405540 ko || +13m09.33s || 3012 ko | +5069.62% | +0.74% 13m25.05s | 424812 ko | fiat-rust/src/curve25519_scalar_32.rs | 0m16.29s | 429536 ko || +13m08.75s || -4724 ko | +4841.98% | -1.09% 13m23.56s | 408716 ko | fiat-json/src/p256_scalar_32.json | 0m17.19s | 433976 ko || +13m06.36s || -25260 ko | +4574.57% | -5.82% 13m06.66s | 379524 ko | fiat-go/32/curve25519scalar/curve25519scalar.go | 0m16.39s | 325156 ko || +12m50.26s || 54368 ko | +4699.63% | +16.72% 13m01.92s | 376048 ko | fiat-c/src/p256_32.c | 0m15.20s | 359812 ko || +12m46.71s || 16236 ko | +5044.21% | +4.51% 12m49.57s | 423456 ko | fiat-json/src/p256_32.json | 0m15.99s | 429740 ko || +12m33.58s || -6284 ko | +4712.82% | -1.46% 7m46.98s | 284096 ko | fiat-zig/src/p434_64.zig | 0m18.66s | 250696 ko || +7m28.31s || 33400 ko | +2402.57% | +13.32% 7m42.79s | 333916 ko | fiat-json/src/p434_64.json | 0m18.88s | 286668 ko || +7m23.91s || 47248 ko | +2351.21% | +16.48% 7m42.28s | 277712 ko | fiat-go/64/p434/p434.go | 0m18.69s | 239700 ko || +7m23.58s || 38012 ko | +2373.40% | +15.85% 7m34.48s | 367696 ko | fiat-bedrock2/src/p434_64.c | 0m18.03s | 373868 ko || +7m16.45s || -6172 ko | +2420.68% | -1.65% 7m30.04s | 281384 ko | fiat-c/src/p434_64.c | 0m18.12s | 237136 ko || +7m11.92s || 44248 ko | +2383.66% | +18.65% 7m24.16s | 273604 ko | fiat-rust/src/p434_64.rs | 0m18.41s | 246432 ko || +7m05.75s || 27172 ko | +2312.60% | +11.02% 5m07.03s | 256280 ko | fiat-java/src/FiatP224.java | 0m09.02s | 261688 ko || +4m58.00s || -5408 ko | +3303.88% | -2.06% 5m00.31s | 238104 ko | fiat-go/32/p224/p224.go | 0m09.01s | 223612 ko || +4m51.30s || 14492 ko | +3233.07% | +6.48% 4m59.60s | 240328 ko | fiat-zig/src/p224_32.zig | 0m09.37s | 225308 ko || +4m50.23s || 15020 ko | +3097.43% | +6.66% 4m57.44s | 288372 ko | fiat-json/src/p224_32.json | 0m09.08s | 283848 ko || +4m48.36s || 4524 ko | +3175.77% | +1.59% 4m57.40s | 239760 ko | fiat-rust/src/p224_32.rs | 0m09.10s | 269448 ko || +4m48.29s || -29688 ko | +3168.13% | -11.01% 4m48.57s | 310468 ko | fiat-bedrock2/src/p224_32.c | 0m08.99s | 282600 ko || +4m39.57s || 27868 ko | +3109.89% | +9.86% 4m48.27s | 239228 ko | fiat-c/src/p224_32.c | 0m08.96s | 230412 ko || +4m39.31s || 8816 ko | +3117.29% | +3.82% 2m48.44s | 184568 ko | fiat-json/src/p384_scalar_64.json | 0m11.23s | 168860 ko || +2m37.21s || 15708 ko | +1399.91% | +9.30% 2m48.14s | 167980 ko | fiat-zig/src/p384_scalar_64.zig | 0m11.15s | 193456 ko || +2m36.98s || -25476 ko | +1407.98% | -13.16% 2m43.61s | 213748 ko | fiat-bedrock2/src/p384_scalar_64.c | 0m11.14s | 192024 ko || +2m32.47s || 21724 ko | +1368.67% | +11.31% 2m43.36s | 167968 ko | fiat-rust/src/p384_scalar_64.rs | 0m11.11s | 148400 ko || +2m32.25s || 19568 ko | +1370.38% | +13.18% 2m41.95s | 184636 ko | fiat-c/src/p384_scalar_64.c | 0m10.86s | 154744 ko || +2m31.08s || 29892 ko | +1391.25% | +19.31% 2m37.83s | 178156 ko | fiat-go/64/p384scalar/p384scalar.go | 0m11.15s | 193256 ko || +2m26.67s || -15100 ko | +1315.51% | -7.81% 2m29.44s | 166896 ko | fiat-zig/src/p384_64.zig | 0m09.52s | 154632 ko || +2m19.91s || 12264 ko | +1469.74% | +7.93% 2m28.98s | 184900 ko | fiat-json/src/p384_64.json | 0m09.31s | 171972 ko || +2m19.66s || 12928 ko | +1500.21% | +7.51% 2m26.58s | 176112 ko | fiat-go/64/p384/p384.go | 0m09.47s | 181568 ko || +2m17.10s || -5456 ko | +1447.83% | -3.00% 2m25.35s | 175320 ko | fiat-c/src/p384_64.c | 0m09.21s | 151936 ko || +2m16.13s || 23384 ko | +1478.17% | +15.39% 2m23.86s | 189000 ko | fiat-rust/src/p384_64.rs | 0m09.49s | 145760 ko || +2m14.37s || 43240 ko | +1415.91% | +29.66% 2m23.35s | 211804 ko | fiat-bedrock2/src/p384_64.c | 0m09.35s | 189388 ko || +2m14.00s || 22416 ko | +1433.15% | +11.83% 2m29.46s | 2135192 ko | SlowPrimeSynthesisExamples.vo | 1m28.36s | 2041032 ko || +1m01.10s || 94160 ko | +69.14% | +4.61% 8m55.13s | 2650896 ko | Bedrock/End2End/X25519/GarageDoor.vo | 8m25.15s | 2656564 ko || +0m29.98s || -5668 ko | +5.93% | -0.21% N/A | N/A | PerfTesting/PerfTestSearch.vo | 0m25.13s | 1373644 ko || -0m25.12s || -1373644 ko | -100.00% | -100.00% 1m22.31s | 941440 ko | AbstractInterpretation/Wf.vo | 1m00.63s | 875644 ko || +0m21.67s || 65796 ko | +35.75% | +7.51% 0m48.67s | 773988 ko | AbstractInterpretation/Proofs.vo | 0m29.24s | 702092 ko || +0m19.43s || 71896 ko | +66.45% | +10.24% N/A | N/A | PerfTesting/PerfTestSearchPattern.vo | 0m17.89s | 1373708 ko || -0m17.89s || -1373708 ko | -100.00% | -100.00% 2m05.77s | 1631168 ko | Bedrock/End2End/X25519/Field25519.vo | 1m51.28s | 1594648 ko || +0m14.48s || 36520 ko | +13.02% | +2.29% 0m14.59s | 64448 ko | fiat-rust/src/secp256k1_montgomery_scalar_64.rs | 0m02.67s | 60756 ko || +0m11.92s || 3692 ko | +446.44% | +6.07% 0m14.52s | 62432 ko | fiat-zig/src/secp256k1_montgomery_scalar_64.zig | 0m02.76s | 58520 ko || +0m11.75s || 3912 ko | +426.08% | +6.68% 0m14.45s | 67272 ko | fiat-zig/src/p256_scalar_64.zig | 0m02.73s | 61128 ko || +0m11.71s || 6144 ko | +429.30% | +10.05% 0m14.25s | 62036 ko | fiat-rust/src/p256_scalar_64.rs | 0m02.73s | 59396 ko || +0m11.51s || 2640 ko | +421.97% | +4.44% 0m14.21s | 72352 ko | fiat-json/src/secp256k1_montgomery_scalar_64.json | 0m02.82s | 70904 ko || +0m11.39s || 1448 ko | +403.90% | +2.04% 0m14.16s | 62316 ko | fiat-c/src/secp256k1_montgomery_scalar_64.c | 0m02.69s | 61008 ko || +0m11.47s || 1308 ko | +426.39% | +2.14% 0m14.02s | 72660 ko | fiat-json/src/p256_scalar_64.json | 0m02.79s | 73116 ko || +0m11.23s || -456 ko | +402.50% | -0.62% 0m14.01s | 61848 ko | fiat-c/src/p256_scalar_64.c | 0m02.67s | 61148 ko || +0m11.33s || 700 ko | +424.71% | +1.14% 0m13.61s | 67516 ko | fiat-go/64/p256scalar/p256scalar.go | 0m02.75s | 71172 ko || +0m10.85s || -3656 ko | +394.90% | -5.13% 0m13.51s | 85796 ko | fiat-bedrock2/src/p256_scalar_64.c | 0m02.80s | 84480 ko || +0m10.71s || 1316 ko | +382.50% | +1.55% 0m13.33s | 66312 ko | fiat-go/64/secp256k1montgomeryscalar/secp256k1montgomeryscalar.go | 0m02.77s | 62180 ko || +0m10.56s || 4132 ko | +381.22% | +6.64% 0m13.20s | 88340 ko | fiat-bedrock2/src/secp256k1_montgomery_scalar_64.c | 0m02.79s | 84980 ko || +0m10.41s || 3360 ko | +373.11% | +3.95% 0m13.16s | 70728 ko | fiat-json/src/secp256k1_montgomery_64.json | 0m02.49s | 75256 ko || +0m10.67s || -4528 ko | +428.51% | -6.01% 0m12.74s | 65292 ko | fiat-zig/src/secp256k1_montgomery_64.zig | 0m02.42s | 59724 ko || +0m10.32s || 5568 ko | +426.44% | +9.32% 0m12.72s | 60072 ko | fiat-zig/src/curve25519_scalar_64.zig | 0m02.33s | 55908 ko || +0m10.39s || 4164 ko | +445.92% | +7.44% 0m12.66s | 61844 ko | fiat-rust/src/secp256k1_montgomery_64.rs | 0m02.44s | 64784 ko || +0m10.22s || -2940 ko | +418.85% | -4.53% 0m44.36s | 60596 ko | fiat-zig/src/p521_32.zig | 0m34.44s | 60684 ko || +0m09.92s || -88 ko | +28.80% | -0.14% 0m44.26s | 112564 ko | fiat-json/src/p521_32.json | 0m34.64s | 106924 ko || +0m09.61s || 5640 ko | +27.77% | +5.27% 0m44.00s | 64936 ko | fiat-java/src/FiatP521.java | 0m34.40s | 59656 ko || +0m09.60s || 5280 ko | +27.90% | +8.85% 0m12.18s | 66900 ko | fiat-rust/src/curve25519_scalar_64.rs | 0m02.32s | 60304 ko || +0m09.85s || 6596 ko | +425.00% | +10.93% 0m12.07s | 67140 ko | fiat-c/src/secp256k1_montgomery_64.c | 0m02.43s | 59524 ko || +0m09.64s || 7616 ko | +396.70% | +12.79% 0m12.06s | 60364 ko | fiat-c/src/curve25519_scalar_64.c | 0m02.27s | 60644 ko || +0m09.79s || -280 ko | +431.27% | -0.46% 0m12.03s | 68432 ko | fiat-json/src/curve25519_scalar_64.json | 0m02.36s | 67732 ko || +0m09.67s || 700 ko | +409.74% | +1.03% 0m11.92s | 84348 ko | fiat-bedrock2/src/secp256k1_montgomery_64.c | 0m02.45s | 88860 ko || +0m09.46s || -4512 ko | +386.53% | -5.07% 0m11.92s | 67724 ko | fiat-go/64/secp256k1montgomery/secp256k1montgomery.go | 0m02.52s | 71412 ko || +0m09.40s || -3688 ko | +373.01% | -5.16% 0m11.58s | 64176 ko | fiat-go/64/curve25519scalar/curve25519scalar.go | 0m02.32s | 67952 ko || +0m09.25s || -3776 ko | +399.13% | -5.55% 0m11.37s | 84500 ko | fiat-bedrock2/src/curve25519_scalar_64.c | 0m02.28s | 80992 ko || +0m09.08s || 3508 ko | +398.68% | +4.33% 0m10.91s | 60548 ko | fiat-c/src/p224_64.c | 0m01.82s | 61764 ko || +0m09.08s || -1216 ko | +499.45% | -1.96% 0m10.82s | 65748 ko | fiat-zig/src/p224_64.zig | 0m01.84s | 62532 ko || +0m08.98s || 3216 ko | +488.04% | +5.14% 0m10.55s | 60692 ko | fiat-rust/src/p224_64.rs | 0m01.84s | 61180 ko || +0m08.71s || -488 ko | +473.36% | -0.79% 0m10.52s | 73812 ko | fiat-json/src/p224_64.json | 0m01.88s | 70604 ko || +0m08.64s || 3208 ko | +459.57% | +4.54% 0m10.37s | 64696 ko | fiat-go/64/p224/p224.go | 0m01.92s | 62144 ko || +0m08.44s || 2552 ko | +440.10% | +4.10% 0m10.11s | 83476 ko | fiat-bedrock2/src/p224_64.c | 0m01.91s | 80568 ko || +0m08.19s || 2908 ko | +429.31% | +3.60% 0m15.25s | 124888 ko | fiat-json/src/p448_solinas_32.json | 0m08.23s | 116848 ko || +0m07.01s || 8040 ko | +85.29% | +6.88% 0m09.70s | 75056 ko | fiat-json/src/p256_64.json | 0m01.88s | 66596 ko || +0m07.81s || 8460 ko | +415.95% | +12.70% 0m09.59s | 61964 ko | fiat-zig/src/p256_64.zig | 0m01.79s | 58052 ko || +0m07.79s || 3912 ko | +435.75% | +6.73% 0m09.50s | 82532 ko | fiat-bedrock2/src/p256_64.c | 0m01.84s | 77096 ko || +0m07.66s || 5436 ko | +416.30% | +7.05% 0m09.43s | 61676 ko | fiat-rust/src/p256_64.rs | 0m01.80s | 59200 ko || +0m07.62s || 2476 ko | +423.88% | +4.18% 0m09.40s | 65020 ko | fiat-go/64/p256/p256.go | 0m01.84s | 68656 ko || +0m07.56s || -3636 ko | +410.86% | -5.29% 0m08.91s | 65360 ko | fiat-c/src/p256_64.c | 0m01.74s | 61792 ko || +0m07.16s || 3568 ko | +412.06% | +5.77% 0m14.66s | 59488 ko | fiat-zig/src/p448_solinas_32.zig | 0m08.21s | 58580 ko || +0m06.44s || 908 ko | +78.56% | +1.55% 0m14.42s | 60808 ko | fiat-rust/src/p448_solinas_32.rs | 0m08.12s | 57972 ko || +0m06.30s || 2836 ko | +77.58% | +4.89% 0m43.35s | 62524 ko | fiat-rust/src/p521_32.rs | 0m38.15s | 60356 ko || +0m05.20s || 2168 ko | +13.63% | +3.59% 0m42.96s | 148852 ko | fiat-bedrock2/src/p521_32.c | 0m38.63s | 145628 ko || +0m04.32s || 3224 ko | +11.20% | +2.21% 0m42.44s | 63004 ko | fiat-c/src/p521_32.c | 0m37.82s | 59428 ko || +0m04.61s || 3576 ko | +12.21% | +6.01% 0m12.71s | 60360 ko | fiat-c/src/p448_solinas_32.c | 0m08.16s | 58932 ko || +0m04.55s || 1428 ko | +55.75% | +2.42% 1m51.21s | 2339376 ko | Fancy/Barrett256.vo | 1m48.02s | 2417076 ko || +0m03.18s || -77700 ko | +2.95% | -3.21% 0m49.17s | 1960112 ko | Fancy/Montgomery256.vo | 0m46.03s | 1758768 ko || +0m03.14s || 201344 ko | +6.82% | +11.44% 0m39.42s | 1319164 ko | Bedrock/End2End/X25519/GarageDoorTop.vo | 0m36.42s | 1319624 ko || +0m03.00s || -460 ko | +8.23% | -0.03% N/A | N/A | Bedrock/Everything.vo | 0m03.96s | 1502884 ko || -0m03.96s || -1502884 ko | -100.00% | -100.00% N/A | N/A | Everything.vo | 0m03.45s | 1363476 ko || -0m03.45s || -1363476 ko | -100.00% | -100.00% N/A | N/A | PerfTesting/PerfTestPrint.vo | 0m02.94s | 1346788 ko || -0m02.94s || -1346788 ko | -100.00% | -100.00% 3m59.87s | 2553740 ko | Assembly/WithBedrock/Proofs.vo | 4m01.45s | 2553452 ko || -0m01.57s || 288 ko | -0.65% | +0.01% 1m54.05s | 1916900 ko | Bedrock/End2End/X25519/AddPrecomputed.vo | 1m52.31s | 1916820 ko || +0m01.73s || 80 ko | +1.54% | +0.00% 0m43.02s | 64392 ko | fiat-go/32/p521/p521.go | 0m41.64s | 61532 ko || +0m01.38s || 2860 ko | +3.31% | +4.64% 0m38.48s | 2146800 ko | ExtractionOCaml/word_by_word_montgomery | 0m40.27s | 2146720 ko || -0m01.79s || 80 ko | -4.44% | +0.00% 0m37.41s | 1617480 ko | ExtractionOCaml/bedrock2_base_conversion | 0m36.32s | 1620388 ko || +0m01.08s || -2908 ko | +3.00% | -0.17% 0m34.27s | 1554400 ko | ExtractionOCaml/dettman_multiplication | 0m35.58s | 1556408 ko || -0m01.30s || -2008 ko | -3.68% | -0.12% 0m33.95s | 1537192 ko | ExtractionOCaml/saturated_solinas | 0m35.03s | 1536764 ko || -0m01.07s || 428 ko | -3.08% | +0.02% 0m33.09s | 1531652 ko | ExtractionOCaml/base_conversion | 0m34.83s | 1534488 ko || -0m01.73s || -2836 ko | -4.99% | -0.18% 0m18.85s | 1135808 ko | Bedrock/End2End/Poly1305/Field1305.vo | 0m17.21s | 1129728 ko || +0m01.64s || 6080 ko | +9.52% | +0.53% 0m06.04s | 52332 ko | fiat-json/src/p521_64.json | 0m04.95s | 50248 ko || +0m01.08s || 2084 ko | +22.02% | +4.14% 0m05.99s | 37064 ko | fiat-zig/src/p521_64.zig | 0m04.90s | 36028 ko || +0m01.08s || 1036 ko | +22.24% | +2.87% 3m13.72s | 2603328 ko | Bedrock/End2End/RupicolaCrypto/ChaCha20.vo | 3m12.79s | 2602764 ko || +0m00.93s || 564 ko | +0.48% | +0.02% 1m20.21s | 1524980 ko | Assembly/EquivalenceProofs.vo | 1m19.93s | 1530348 ko || +0m00.28s || -5368 ko | +0.35% | -0.35% 1m08.96s | 1437312 ko | Assembly/WithBedrock/SymbolicProofs.vo | 1m09.05s | 1433128 ko || -0m00.08s || 4184 ko | -0.13% | +0.29% 0m52.36s | 849360 ko | AbstractInterpretation/ZRangeProofs.vo | 0m52.56s | 848576 ko || -0m00.20s || 784 ko | -0.38% | +0.09% 0m47.18s | 1522816 ko | Assembly/Symbolic.vo | 0m46.85s | 1519280 ko || +0m00.32s || 3536 ko | +0.70% | +0.23% 0m45.33s | 1104720 ko | Bedrock/Field/Translation/Proofs/ValidComputable/Expr.vo | 0m45.27s | 1102688 ko || +0m00.05s || 2032 ko | +0.13% | +0.18% 0m44.21s | 2147564 ko | ExtractionOCaml/bedrock2_solinas_reduction | 0m43.87s | 2147536 ko || +0m00.34s || 28 ko | +0.77% | +0.00% 0m43.12s | 2147892 ko | ExtractionOCaml/with_bedrock2_word_by_word_montgomery | 0m42.35s | 2147764 ko || +0m00.76s || 128 ko | +1.81% | +0.00% 0m42.80s | 2147776 ko | ExtractionOCaml/bedrock2_word_by_word_montgomery | 0m42.72s | 2147720 ko || +0m00.07s || 56 ko | +0.18% | +0.00% 0m41.90s | 2146676 ko | ExtractionOCaml/solinas_reduction | 0m40.99s | 2146676 ko || +0m00.90s || 0 ko | +2.22% | +0.00% 0m41.07s | 2148200 ko | ExtractionOCaml/bedrock2_unsaturated_solinas | 0m41.03s | 2148284 ko || +0m00.03s || -84 ko | +0.09% | -0.00% 0m41.02s | 2148244 ko | ExtractionOCaml/with_bedrock2_unsaturated_solinas | 0m40.67s | 2148388 ko || +0m00.35s || -144 ko | +0.86% | -0.00% 0m38.93s | 1016196 ko | Bedrock/Group/ScalarMult/MontgomeryLadder.vo | 0m38.63s | 1016156 ko || +0m00.29s || 40 ko | +0.77% | +0.00% 0m38.12s | 1640756 ko | ExtractionOCaml/bedrock2_dettman_multiplication | 0m38.27s | 1642204 ko || -0m00.15s || -1448 ko | -0.39% | -0.08% 0m37.73s | 1623836 ko | ExtractionOCaml/unsaturated_solinas | 0m36.98s | 1653884 ko || +0m00.75s || -30048 ko | +2.02% | -1.81% 0m37.59s | 1621000 ko | ExtractionOCaml/with_bedrock2_dettman_multiplication | 0m37.00s | 1623732 ko || +0m00.59s || -2732 ko | +1.59% | -0.16% 0m37.56s | 1622616 ko | ExtractionOCaml/bedrock2_saturated_solinas | 0m37.36s | 1622804 ko || +0m00.20s || -188 ko | +0.53% | -0.01% 0m37.50s | 1621644 ko | ExtractionOCaml/with_bedrock2_solinas_reduction | 0m36.64s | 1624384 ko || +0m00.85s || -2740 ko | +2.34% | -0.16% 0m37.48s | 1621204 ko | ExtractionOCaml/with_bedrock2_saturated_solinas | 0m36.73s | 1622828 ko || +0m00.75s || -1624 ko | +2.04% | -0.10% 0m37.41s | 1617420 ko | ExtractionOCaml/with_bedrock2_base_conversion | 0m36.52s | 1618732 ko || +0m00.88s || -1312 ko | +2.43% | -0.08% 0m34.91s | 1279244 ko | Bedrock/End2End/X25519/MontgomeryLadder.vo | 0m34.52s | 1284212 ko || +0m00.38s || -4968 ko | +1.12% | -0.38% 0m34.14s | 2240972 ko | ExtractionOCaml/bedrock2_solinas_reduction.ml | 0m34.09s | 2237916 ko || +0m00.04s || 3056 ko | +0.14% | +0.13% 0m33.11s | 2212924 ko | ExtractionOCaml/with_bedrock2_word_by_word_montgomery.ml | 0m33.05s | 2204728 ko || +0m00.06s || 8196 ko | +0.18% | +0.37% 0m33.01s | 2213528 ko | ExtractionOCaml/bedrock2_word_by_word_montgomery.ml | 0m33.12s | 2204912 ko || -0m00.10s || 8616 ko | -0.33% | +0.39% 0m32.68s | 2150644 ko | ExtractionOCaml/solinas_reduction.ml | 0m32.56s | 2139728 ko || +0m00.11s || 10916 ko | +0.36% | +0.51% 0m31.47s | 2122940 ko | ExtractionOCaml/word_by_word_montgomery.ml | 0m31.58s | 2120080 ko || -0m00.10s || 2860 ko | -0.34% | +0.13% 0m30.56s | 2114916 ko | ExtractionOCaml/bedrock2_unsaturated_solinas.ml | 0m30.25s | 2134076 ko || +0m00.30s || -19160 ko | +1.02% | -0.89% 0m30.28s | 1232236 ko | ExtractionOCaml/perf_word_by_word_montgomery | 0m29.68s | 1232100 ko || +0m00.60s || 136 ko | +2.02% | +0.01% 0m30.23s | 1232188 ko | ExtractionOCaml/perf_unsaturated_solinas | 0m29.28s | 1232224 ko || +0m00.94s || -36 ko | +3.24% | -0.00% 0m30.11s | 2114248 ko | ExtractionOCaml/with_bedrock2_unsaturated_solinas.ml | 0m30.11s | 2132752 ko || +0m00.00s || -18504 ko | +0.00% | -0.86% 0m29.92s | 1526264 ko | StandaloneDebuggingExamples.vo | 0m30.25s | 1518384 ko || -0m00.32s || 7880 ko | -1.09% | +0.51% 0m28.53s | 2051912 ko | ExtractionOCaml/unsaturated_solinas.ml | 0m28.57s | 2032036 ko || -0m00.03s || 19876 ko | -0.14% | +0.97% 0m27.17s | 2040956 ko | ExtractionOCaml/bedrock2_dettman_multiplication.ml | 0m27.20s | 2071528 ko || -0m00.02s || -30572 ko | -0.11% | -1.47% 0m26.36s | 2009348 ko | ExtractionOCaml/with_bedrock2_base_conversion.ml | 0m26.30s | 2038044 ko || +0m00.05s || -28696 ko | +0.22% | -1.40% 0m26.34s | 2019688 ko | ExtractionOCaml/with_bedrock2_saturated_solinas.ml | 0m26.08s | 1996180 ko || +0m00.26s || 23508 ko | +0.99% | +1.17% 0m26.32s | 2018796 ko | ExtractionOCaml/bedrock2_saturated_solinas.ml | 0m26.30s | 1997428 ko || +0m00.01s || 21368 ko | +0.07% | +1.06% 0m26.18s | 2009056 ko | ExtractionOCaml/bedrock2_base_conversion.ml | 0m26.28s | 2037624 ko || -0m00.10s || -28568 ko | -0.38% | -1.40% 0m26.13s | 2019564 ko | ExtractionOCaml/with_bedrock2_solinas_reduction.ml | 0m26.06s | 1996408 ko || +0m00.07s || 23156 ko | +0.26% | +1.15% 0m26.05s | 2016864 ko | ExtractionOCaml/with_bedrock2_dettman_multiplication.ml | 0m25.95s | 1997412 ko || +0m00.10s || 19452 ko | +0.38% | +0.97% 0m25.24s | 1952220 ko | ExtractionOCaml/dettman_multiplication.ml | 0m25.47s | 1939744 ko || -0m00.23s || 12476 ko | -0.90% | +0.64% 0m24.71s | 1944464 ko | ExtractionOCaml/saturated_solinas.ml | 0m24.58s | 1934196 ko || +0m00.13s || 10268 ko | +0.52% | +0.53% 0m24.68s | 1914120 ko | ExtractionOCaml/base_conversion.ml | 0m25.18s | 1922504 ko || -0m00.50s || -8384 ko | -1.98% | -0.43% 0m23.73s | 1184052 ko | PushButtonSynthesis/UnsaturatedSolinas.vo | 0m23.70s | 1185516 ko || +0m00.03s || -1464 ko | +0.12% | -0.12% 0m21.52s | 1354084 ko | Bedrock/End2End/RupicolaCrypto/Low.vo | 0m21.31s | 1350968 ko || +0m00.21s || 3116 ko | +0.98% | +0.23% 0m20.49s | 1158932 ko | PushButtonSynthesis/WordByWordMontgomery.vo | 0m20.38s | 1160828 ko || +0m00.10s || -1896 ko | +0.53% | -0.16% 0m20.15s | 794468 ko | Bedrock/Field/Translation/Proofs/Expr.vo | 0m20.23s | 794720 ko || -0m00.08s || -252 ko | -0.39% | -0.03% 0m18.98s | 1821220 ko | ExtractionOCaml/perf_unsaturated_solinas.ml | 0m18.91s | 1826352 ko || +0m00.07s || -5132 ko | +0.37% | -0.28% 0m18.80s | 1843872 ko | ExtractionOCaml/perf_word_by_word_montgomery.ml | 0m18.84s | 1867344 ko || -0m00.03s || -23472 ko | -0.21% | -1.25% 0m18.64s | 737284 ko | Bedrock/Field/Translation/Proofs/LoadStoreList.vo | 0m18.64s | 740312 ko || +0m00.00s || -3028 ko | +0.00% | -0.40% 0m18.55s | 1128568 ko | Bedrock/Field/Translation/Proofs/Func.vo | 0m18.89s | 1129264 ko || -0m00.33s || -696 ko | -1.79% | -0.06% 0m17.93s | 1207216 ko | Bedrock/Field/Synthesis/New/Signature.vo | 0m17.32s | 1210248 ko || +0m00.60s || -3032 ko | +3.52% | -0.25% 0m16.37s | 1144804 ko | Bedrock/Field/Translation/Proofs/Cmd.vo | 0m17.11s | 1146536 ko || -0m00.73s || -1732 ko | -4.32% | -0.15% 0m16.36s | 2063348 ko | ExtractionHaskell/with_bedrock2_solinas_reduction.hs | 0m16.35s | 2096996 ko || +0m00.00s || -33648 ko | +0.06% | -1.60% 0m16.26s | 2063264 ko | ExtractionHaskell/bedrock2_solinas_reduction.hs | 0m16.52s | 2095492 ko || -0m00.25s || -32228 ko | -1.57% | -1.53% 0m16.16s | 2023412 ko | ExtractionHaskell/with_bedrock2_word_by_word_montgomery.hs | 0m16.00s | 2007884 ko || +0m00.16s || 15528 ko | +1.00% | +0.77% 0m16.12s | 2021956 ko | ExtractionHaskell/bedrock2_word_by_word_montgomery.hs | 0m16.13s | 2010240 ko || -0m00.00s || 11716 ko | -0.06% | +0.58% 0m16.08s | 2028764 ko | ExtractionHaskell/solinas_reduction.hs | 0m15.49s | 1949544 ko || +0m00.58s || 79220 ko | +3.80% | +4.06% 0m15.71s | 1962892 ko | ExtractionHaskell/word_by_word_montgomery.hs | 0m15.47s | 1985472 ko || +0m00.24s || -22580 ko | +1.55% | -1.13% 0m15.59s | 2009872 ko | ExtractionHaskell/bedrock2_unsaturated_solinas.hs | 0m15.14s | 1950900 ko || +0m00.44s || 58972 ko | +2.97% | +3.02% 0m15.21s | 2008904 ko | ExtractionHaskell/with_bedrock2_unsaturated_solinas.hs | 0m15.24s | 1949824 ko || -0m00.02s || 59080 ko | -0.19% | +3.03% 0m14.63s | 1934444 ko | ExtractionHaskell/bedrock2_saturated_solinas.hs | 0m14.22s | 1862544 ko || +0m00.41s || 71900 ko | +2.88% | +3.86% 0m14.47s | 1879068 ko | ExtractionHaskell/bedrock2_dettman_multiplication.hs | 0m14.36s | 1889216 ko || +0m00.11s || -10148 ko | +0.76% | -0.53% 0m14.38s | 1879364 ko | ExtractionHaskell/with_bedrock2_dettman_multiplication.hs | 0m14.28s | 1886884 ko || +0m00.10s || -7520 ko | +0.70% | -0.39% 0m14.37s | 1927472 ko | ExtractionHaskell/unsaturated_solinas.hs | 0m14.53s | 1928704 ko || -0m00.16s || -1232 ko | -1.10% | -0.06% 0m14.32s | 1917200 ko | ExtractionHaskell/with_bedrock2_base_conversion.hs | 0m14.11s | 1932764 ko || +0m00.21s || -15564 ko | +1.48% | -0.80% 0m14.30s | 1935052 ko | ExtractionHaskell/with_bedrock2_saturated_solinas.hs | 0m14.12s | 1861568 ko || +0m00.18s || 73484 ko | +1.27% | +3.94% 0m14.20s | 1917124 ko | ExtractionHaskell/bedrock2_base_conversion.hs | 0m14.22s | 1911084 ko || -0m00.02s || 6040 ko | -0.14% | +0.31% 0m14.06s | 1864296 ko | ExtractionHaskell/dettman_multiplication.hs | 0m13.20s | 1856724 ko || +0m00.86s || 7572 ko | +6.51% | +0.40% 0m13.78s | 1840872 ko | ExtractionHaskell/saturated_solinas.hs | 0m13.71s | 1846748 ko || +0m00.06s || -5876 ko | +0.51% | -0.31% 0m13.68s | 602488 ko | Bedrock/Field/Common/Util.vo | 0m13.68s | 602420 ko || +0m00.00s || 68 ko | +0.00% | +0.01% 0m13.36s | 663048 ko | Bedrock/Group/AdditionChains.vo | 0m13.42s | 662928 ko || -0m00.06s || 120 ko | -0.44% | +0.01% 0m13.12s | 667092 ko | Bedrock/Group/ScalarMult/LadderStep.vo | 0m13.04s | 667048 ko || +0m00.08s || 44 ko | +0.61% | +0.00% 0m13.11s | 1845960 ko | ExtractionHaskell/base_conversion.hs | 0m13.55s | 1851580 ko || -0m00.44s || -5620 ko | -3.24% | -0.30% 0m11.84s | 1029560 ko | BoundsPipeline.vo | 0m11.88s | 1028872 ko || -0m00.04s || 688 ko | -0.33% | +0.06% 0m11.31s | 1699940 ko | Bedrock/Field/Synthesis/New/WordByWordMontgomery.vo | 0m11.11s | 1698268 ko || +0m00.20s || 1672 ko | +1.80% | +0.09% 0m10.27s | 1297740 ko | Bedrock/End2End/X25519/MontgomeryLadderProperties.vo | 0m10.36s | 1296660 ko || -0m00.08s || 1080 ko | -0.86% | +0.08% 0m09.88s | 600056 ko | Stringification/IR.vo | 0m09.90s | 601448 ko || -0m00.01s || -1392 ko | -0.20% | -0.23% 0m09.74s | 624024 ko | Bedrock/Field/Translation/Proofs/Flatten.vo | 0m09.73s | 624068 ko || +0m00.00s || -44 ko | +0.10% | -0.00% 0m09.34s | 1032068 ko | PushButtonSynthesis/BaseConversion.vo | 0m09.37s | 1035460 ko || -0m00.02s || -3392 ko | -0.32% | -0.32% 0m09.19s | 661364 ko | Bedrock/Group/ScalarMult/CSwap.vo | 0m08.96s | 661296 ko || +0m00.22s || 68 ko | +2.56% | +0.01% 0m08.41s | 1044384 ko | PushButtonSynthesis/Primitives.vo | 0m08.56s | 1045468 ko || -0m00.15s || -1084 ko | -1.75% | -0.10% 0m08.33s | 1000952 ko | PushButtonSynthesis/SmallExamples.vo | 0m08.13s | 1001028 ko || +0m00.19s || -76 ko | +2.46% | -0.00% 0m07.39s | 911356 ko | Bedrock/Field/Translation/Proofs/EquivalenceProperties.vo | 0m07.43s | 912204 ko || -0m00.04s || -848 ko | -0.53% | -0.09% 0m07.17s | 1033520 ko | PushButtonSynthesis/SolinasReduction.vo | 0m07.24s | 1033476 ko || -0m00.07s || 44 ko | -0.96% | +0.00% 0m06.67s | 51208 ko | fiat-go/64/p521/p521.go | 0m06.32s | 49224 ko || +0m00.34s || 1984 ko | +5.53% | +4.03% 0m06.40s | 1040660 ko | PushButtonSynthesis/BarrettReduction.vo | 0m06.42s | 1041508 ko || -0m00.01s || -848 ko | -0.31% | -0.08% 0m06.27s | 1121320 ko | CLI.vo | 0m06.40s | 1121288 ko || -0m00.13s || 32 ko | -2.03% | +0.00% 0m05.98s | 1130956 ko | Bedrock/Field/Synthesis/New/UnsaturatedSolinas.vo | 0m06.11s | 1130988 ko || -0m00.12s || -32 ko | -2.12% | -0.00% 0m05.92s | 68612 ko | fiat-bedrock2/src/p521_64.c | 0m04.98s | 66368 ko || +0m00.93s || 2244 ko | +18.87% | +3.38% 0m05.90s | 37028 ko | fiat-rust/src/p521_64.rs | 0m05.46s | 36076 ko || +0m00.44s || 952 ko | +8.05% | +2.63% 0m05.87s | 37052 ko | fiat-c/src/p521_64.c | 0m05.41s | 36076 ko || +0m00.46s || 976 ko | +8.50% | +2.70% 0m05.84s | 906864 ko | Bedrock/Field/Translation/Proofs/UsedVarnames.vo | 0m05.81s | 906744 ko || +0m00.03s || 120 ko | +0.51% | +0.01% 0m05.25s | 616380 ko | Bedrock/End2End/RupicolaCrypto/Broadcast.vo | 0m05.24s | 616216 ko || +0m00.00s || 164 ko | +0.19% | +0.02% 0m04.73s | 1030392 ko | PushButtonSynthesis/DettmanMultiplication.vo | 0m04.78s | 1030456 ko || -0m00.04s || -64 ko | -1.04% | -0.00% 0m04.41s | 1065452 ko | Bedrock/Field/Translation/Proofs/ValidComputable/Cmd.vo | 0m04.56s | 1065272 ko || -0m00.14s || 180 ko | -3.28% | +0.01% 0m04.39s | 1038800 ko | PushButtonSynthesis/SaturatedSolinas.vo | 0m04.43s | 1038816 ko || -0m00.04s || -16 ko | -0.90% | -0.00% 0m03.94s | 957224 ko | Assembly/Equivalence.vo | 0m03.88s | 955376 ko || +0m00.06s || 1848 ko | +1.54% | +0.19% 0m03.84s | 1051032 ko | PushButtonSynthesis/FancyMontgomeryReduction.vo | 0m03.94s | 1051128 ko || -0m00.10s || -96 ko | -2.53% | -0.00% 0m03.66s | 666368 ko | Bedrock/Group/ScalarMult/ScalarMult.vo | 0m03.97s | 667152 ko || -0m00.31s || -784 ko | -7.80% | -0.11% 0m03.12s | 1072356 ko | Rewriter/PerfTesting/Core.vo | 0m03.09s | 1072420 ko || +0m00.03s || -64 ko | +0.97% | -0.00% 0m03.06s | 617024 ko | Bedrock/Field/Synthesis/Generic/Bignum.vo | 0m03.02s | 613720 ko || +0m00.04s || 3304 ko | +1.32% | +0.53% 0m03.05s | 1006464 ko | Bedrock/Field/Translation/Cmd.vo | 0m03.08s | 1009700 ko || -0m00.03s || -3236 ko | -0.97% | -0.32% 0m02.98s | 50852 ko | fiat-go/64/p448solinas/p448solinas.go | 0m02.67s | 50500 ko || +0m00.31s || 352 ko | +11.61% | +0.69% 0m02.88s | 1068264 ko | Bedrock/Field/Stringification/Stringification.vo | 0m02.87s | 1068932 ko || +0m00.00s || -668 ko | +0.34% | -0.06% 0m02.84s | 1061220 ko | Bedrock/Field/Translation/Proofs/ValidComputable/Func.vo | 0m02.73s | 1059288 ko || +0m00.10s || 1932 ko | +4.02% | +0.18% 0m02.80s | 1004228 ko | Bedrock/Field/Translation/Func.vo | 0m02.79s | 1004848 ko || +0m00.00s || -620 ko | +0.35% | -0.06% 0m02.80s | 1103608 ko | StandaloneHaskellMain.vo | 0m02.72s | 1103504 ko || +0m00.07s || 104 ko | +2.94% | +0.00% 0m02.78s | 1145260 ko | Bedrock/Standalone/StandaloneHaskellMain.vo | 0m02.77s | 1145216 ko || +0m00.00s || 44 ko | +0.36% | +0.00% 0m02.76s | 1145200 ko | Bedrock/Standalone/StandaloneOCamlMain.vo | 0m02.72s | 1145392 ko || +0m00.03s || -192 ko | +1.47% | -0.01% 0m02.76s | 1131284 ko | Rewriter/PerfTesting/StandaloneOCamlMain.vo | 0m02.73s | 1131224 ko || +0m00.02s || 60 ko | +1.09% | +0.00% 0m02.66s | 1103772 ko | StandaloneOCamlMain.vo | 0m02.68s | 1103860 ko || -0m00.02s || -88 ko | -0.74% | -0.00% 0m02.60s | 620876 ko | Bedrock/Field/Interface/Compilation2.vo | 0m02.60s | 621772 ko || +0m00.00s || -896 ko | +0.00% | -0.14% 0m02.54s | 1064964 ko | Bedrock/Field/Synthesis/New/ComputedOp.vo | 0m02.59s | 1065008 ko || -0m00.04s || -44 ko | -1.93% | -0.00% 0m02.53s | 65044 ko | fiat-bedrock2/src/p448_solinas_64.c | 0m02.07s | 66308 ko || +0m00.46s || -1264 ko | +22.22% | -1.90% 0m02.52s | 35764 ko | fiat-zig/src/p448_solinas_64.zig | 0m01.90s | 35084 ko || +0m00.62s || 680 ko | +32.63% | +1.93% 0m02.51s | 37100 ko | fiat-go/32/curve25519/curve25519.go | 0m02.15s | 35348 ko || +0m00.35s || 1752 ko | +16.74% | +4.95% 0m02.51s | 36324 ko | fiat-rust/src/p448_solinas_64.rs | 0m01.97s | 34328 ko || +0m00.53s || 1996 ko | +27.41% | +5.81% 0m02.50s | 1044712 ko | Bedrock/Field/Translation/Parameters/Defaults32.vo | 0m02.53s | 1045368 ko || -0m00.02s || -656 ko | -1.18% | -0.06% 0m02.48s | 1041488 ko | Bedrock/Field/Translation/Parameters/Defaults.vo | 0m02.40s | 1040684 ko || +0m00.08s || 804 ko | +3.33% | +0.07% 0m02.48s | 1045392 ko | Bedrock/Field/Translation/Parameters/Defaults64.vo | 0m02.56s | 1044600 ko || -0m00.08s || 792 ko | -3.12% | +0.07% 0m02.47s | 51164 ko | fiat-json/src/p448_solinas_64.json | 0m02.06s | 51416 ko || +0m00.41s || -252 ko | +19.90% | -0.49% 0m02.45s | 36696 ko | fiat-c/src/p448_solinas_64.c | 0m01.98s | 35380 ko || +0m00.47s || 1316 ko | +23.73% | +3.71% 0m02.44s | 1045352 ko | Bedrock/Field/Translation/Parameters/FE310.vo | 0m02.50s | 1045320 ko || -0m00.06s || 32 ko | -2.40% | +0.00% 0m02.24s | 63200 ko | fiat-bedrock2/src/curve25519_32.c | 0m01.98s | 63640 ko || +0m00.26s || -440 ko | +13.13% | -0.69% 0m02.18s | 35508 ko | fiat-java/src/FiatCurve25519.java | 0m01.82s | 36520 ko || +0m00.36s || -1012 ko | +19.78% | -2.77% 0m02.16s | 50436 ko | fiat-json/src/curve25519_32.json | 0m01.89s | 49708 ko || +0m00.27s || 728 ko | +14.28% | +1.46% 0m02.15s | 34348 ko | fiat-zig/src/curve25519_32.zig | 0m01.82s | 33388 ko || +0m00.32s || 960 ko | +18.13% | +2.87% 0m02.12s | 34960 ko | fiat-rust/src/curve25519_32.rs | 0m01.79s | 34088 ko || +0m00.33s || 872 ko | +18.43% | +2.55% 0m02.10s | 35016 ko | fiat-c/src/curve25519_32.c | 0m01.76s | 33112 ko || +0m00.34s || 1904 ko | +19.31% | +5.75% 0m01.96s | 616180 ko | Bedrock/Field/Common/Arrays/MakeAccessSizes.vo | 0m02.00s | 614096 ko || -0m00.04s || 2084 ko | -2.00% | +0.33% 0m01.78s | 636112 ko | Bedrock/Group/ScalarMult/MontgomeryEquivalence.vo | 0m01.82s | 636072 ko || -0m00.04s || 40 ko | -2.19% | +0.00% 0m01.62s | 52504 ko | fiat-bedrock2/src/secp256k1_dettman_32.c | 0m01.46s | 51924 ko || +0m00.16s || 580 ko | +10.95% | +1.11% 0m01.61s | 607916 ko | Bedrock/Field/Common/Names/MakeNames.vo | 0m01.62s | 607756 ko || -0m00.01s || 160 ko | -0.61% | +0.02% 0m01.53s | 590192 ko | CompilersTestCases.vo | 0m01.48s | 590236 ko || +0m00.05s || -44 ko | +3.37% | -0.00% 0m01.50s | 512068 ko | AbstractInterpretation/AbstractInterpretation.vo | 0m01.39s | 512000 ko || +0m00.11s || 68 ko | +7.91% | +0.01% 0m01.46s | 26680 ko | fiat-java/src/FiatSecp256K1Dettman.java | 0m01.24s | 26336 ko || +0m00.21s || 344 ko | +17.74% | +1.30% 0m01.46s | 42060 ko | fiat-json/src/secp256k1_dettman_32.json | 0m01.24s | 41136 ko || +0m00.21s || 924 ko | +17.74% | +2.24% 0m01.44s | 26260 ko | fiat-c/src/secp256k1_dettman_32.c | 0m01.21s | 25744 ko || +0m00.23s || 516 ko | +19.00% | +2.00% 0m01.44s | 27956 ko | fiat-go/32/secp256k1dettman/secp256k1dettman.go | 0m01.21s | 25860 ko || +0m00.23s || 2096 ko | +19.00% | +8.10% 0m01.44s | 27096 ko | fiat-rust/src/secp256k1_dettman_32.rs | 0m01.22s | 25440 ko || +0m00.21s || 1656 ko | +18.03% | +6.50% 0m01.42s | 27060 ko | fiat-zig/src/secp256k1_dettman_32.zig | 0m01.23s | 25656 ko || +0m00.18s || 1404 ko | +15.44% | +5.47% 0m01.35s | 543324 ko | Stringification/Go.vo | 0m01.35s | 541960 ko || +0m00.00s || 1364 ko | +0.00% | +0.25% 0m01.11s | 628136 ko | Bedrock/Specs/Field.vo | 0m01.18s | 628700 ko || -0m00.06s || -564 ko | -5.93% | -0.08% 0m01.07s | 609304 ko | Bedrock/Field/Common/Arrays/MaxBounds.vo | 0m01.03s | 607272 ko || +0m00.04s || 2032 ko | +3.88% | +0.33% 0m01.06s | 601084 ko | Bedrock/Field/Common/Arrays/ByteBounds.vo | 0m01.00s | 602440 ko || +0m00.06s || -1356 ko | +6.00% | -0.22% 0m00.94s | 538600 ko | Stringification/C.vo | 0m00.91s | 538560 ko || +0m00.02s || 40 ko | +3.29% | +0.00% 0m00.91s | 535844 ko | Stringification/JSON.vo | 0m00.94s | 537144 ko || -0m00.02s || -1300 ko | -3.19% | -0.24% 0m00.88s | 535936 ko | Stringification/Zig.vo | 0m00.88s | 539224 ko || +0m00.00s || -3288 ko | +0.00% | -0.60% 0m00.87s | 618888 ko | Bedrock/Group/Point.vo | 0m00.89s | 617632 ko || -0m00.02s || 1256 ko | -2.24% | +0.20% 0m00.86s | 615616 ko | Bedrock/Field/Interface/Representation.vo | 0m00.84s | 615736 ko || +0m00.02s || -120 ko | +2.38% | -0.01% 0m00.86s | 534896 ko | Stringification/Java.vo | 0m00.84s | 538936 ko || +0m00.02s || -4040 ko | +2.38% | -0.74% 0m00.80s | 535916 ko | Stringification/Rust.vo | 0m00.83s | 535824 ko || -0m00.02s || 92 ko | -3.61% | +0.01% 0m00.79s | 587772 ko | Bedrock/Field/Common/Tactics.vo | 0m00.78s | 587592 ko || +0m00.01s || 180 ko | +1.28% | +0.03% 0m00.72s | 521776 ko | Bedrock/Field/Common/Arrays/MakeListLengths.vo | 0m00.68s | 523076 ko || +0m00.03s || -1300 ko | +5.88% | -0.24% 0m00.72s | 546832 ko | Bedrock/Field/Stringification/FlattenVarData.vo | 0m00.72s | 546004 ko || +0m00.00s || 828 ko | +0.00% | +0.15% 0m00.70s | 32124 ko | fiat-go/64/curve25519/curve25519.go | 0m00.63s | 32976 ko || +0m00.06s || -852 ko | +11.11% | -2.58% 0m00.68s | 515488 ko | AbstractInterpretation/WfExtra.vo | 0m00.63s | 515492 ko || +0m00.05s || -4 ko | +7.93% | -0.00% 0m00.68s | 38372 ko | fiat-json/src/curve25519_solinas_64.json | 0m00.41s | 39308 ko || +0m00.27s || -936 ko | +65.85% | -2.38% 0m00.67s | 38908 ko | fiat-c/src/curve25519_solinas_64.c | 0m00.41s | 36728 ko || +0m00.26s || 2180 ko | +63.41% | +5.93% 0m00.66s | 37132 ko | fiat-rust/src/curve25519_solinas_64.rs | 0m00.39s | 36260 ko || +0m00.27s || 872 ko | +69.23% | +2.40% 0m00.66s | 36892 ko | fiat-zig/src/curve25519_solinas_64.zig | 0m00.42s | 36336 ko || +0m00.24s || 556 ko | +57.14% | +1.53% 0m00.60s | 34160 ko | fiat-json/src/curve25519_64.json | 0m00.52s | 32588 ko || +0m00.07s || 1572 ko | +15.38% | +4.82% 0m00.59s | 38236 ko | fiat-go/64/curve25519solinas/curve25519solinas.go | 0m00.43s | 36420 ko || +0m00.15s || 1816 ko | +37.20% | +4.98% 0m00.58s | 38448 ko | fiat-bedrock2/src/curve25519_64.c | 0m00.51s | 38216 ko || +0m00.06s || 232 ko | +13.72% | +0.60% 0m00.58s | 42568 ko | fiat-bedrock2/src/curve25519_solinas_64.c | 0m00.42s | 40596 ko || +0m00.15s || 1972 ko | +38.09% | +4.85% 0m00.56s | 26892 ko | fiat-c/src/curve25519_64.c | 0m00.45s | 27132 ko || +0m00.11s || -240 ko | +24.44% | -0.88% 0m00.56s | 26560 ko | fiat-rust/src/curve25519_64.rs | 0m00.46s | 26040 ko || +0m00.10s || 520 ko | +21.73% | +1.99% 0m00.56s | 27152 ko | fiat-zig/src/curve25519_64.zig | 0m00.48s | 26252 ko || +0m00.08s || 900 ko | +16.66% | +3.42% 0m00.54s | 120576 ko | ExtractionOCaml/with_bedrock2_base_conversion.cmi | 0m00.50s | 119948 ko || +0m00.04s || 628 ko | +8.00% | +0.52% 0m00.53s | 120076 ko | ExtractionOCaml/unsaturated_solinas.cmi | 0m00.50s | 119836 ko || +0m00.03s || 240 ko | +6.00% | +0.20% 0m00.52s | 124368 ko | ExtractionOCaml/bedrock2_word_by_word_montgomery.cmi | 0m00.52s | 123300 ko || +0m00.00s || 1068 ko | +0.00% | +0.86% 0m00.52s | 118308 ko | ExtractionOCaml/solinas_reduction.cmi | 0m00.51s | 118636 ko || +0m00.01s || -328 ko | +1.96% | -0.27% 0m00.52s | 121724 ko | ExtractionOCaml/with_bedrock2_dettman_multiplication.cmi | 0m00.50s | 121156 ko || +0m00.02s || 568 ko | +4.00% | +0.46% 0m00.52s | 123672 ko | ExtractionOCaml/with_bedrock2_saturated_solinas.cmi | 0m00.51s | 120980 ko || +0m00.01s || 2692 ko | +1.96% | +2.22% 0m00.52s | 123072 ko | ExtractionOCaml/with_bedrock2_solinas_reduction.cmi | 0m00.50s | 122784 ko || +0m00.02s || 288 ko | +4.00% | +0.23% 0m00.52s | 123132 ko | ExtractionOCaml/with_bedrock2_unsaturated_solinas.cmi | 0m00.50s | 123700 ko || +0m00.02s || -568 ko | +4.00% | -0.45% 0m00.51s | 120332 ko | ExtractionOCaml/bedrock2_base_conversion.cmi | 0m00.51s | 119972 ko || +0m00.00s || 360 ko | +0.00% | +0.30% 0m00.51s | 120184 ko | ExtractionOCaml/bedrock2_unsaturated_solinas.cmi | 0m00.53s | 123596 ko || -0m00.02s || -3412 ko | -3.77% | -2.76% 0m00.51s | 118088 ko | ExtractionOCaml/saturated_solinas.cmi | 0m00.48s | 117928 ko || +0m00.03s || 160 ko | +6.25% | +0.13% 0m00.51s | 124416 ko | ExtractionOCaml/with_bedrock2_word_by_word_montgomery.cmi | 0m00.51s | 123124 ko || +0m00.00s || 1292 ko | +0.00% | +1.04% 0m00.50s | 120256 ko | ExtractionOCaml/base_conversion.cmi | 0m00.51s | 118684 ko || -0m00.01s || 1572 ko | -1.96% | +1.32% 0m00.50s | 123568 ko | ExtractionOCaml/bedrock2_saturated_solinas.cmi | 0m00.54s | 122468 ko || -0m00.04s || 1100 ko | -7.40% | +0.89% 0m00.50s | 123660 ko | ExtractionOCaml/bedrock2_solinas_reduction.cmi | 0m00.50s | 120332 ko || +0m00.00s || 3328 ko | +0.00% | +2.76% 0m00.50s | 119964 ko | ExtractionOCaml/dettman_multiplication.cmi | 0m00.51s | 120092 ko || -0m00.01s || -128 ko | -1.96% | -0.10% 0m00.50s | 120424 ko | ExtractionOCaml/word_by_word_montgomery.cmi | 0m00.49s | 119588 ko || +0m00.01s || 836 ko | +2.04% | +0.69% 0m00.49s | 120784 ko | ExtractionOCaml/bedrock2_dettman_multiplication.cmi | 0m00.53s | 120180 ko || -0m00.04s || 604 ko | -7.54% | +0.50% 0m00.34s | 25480 ko | fiat-go/32/poly1305/poly1305.go | 0m00.30s | 24960 ko || +0m00.04s || 520 ko | +13.33% | +2.08% 0m00.31s | 25196 ko | fiat-go/64/secp256k1dettman/secp256k1dettman.go | 0m00.25s | 24476 ko || +0m00.06s || 720 ko | +24.00% | +2.94% 0m00.30s | 33676 ko | fiat-bedrock2/src/poly1305_32.c | 0m00.25s | 32372 ko || +0m00.04s || 1304 ko | +19.99% | +4.02% 0m00.30s | 29656 ko | fiat-json/src/poly1305_32.json | 0m00.25s | 30252 ko || +0m00.04s || -596 ko | +19.99% | -1.97% 0m00.29s | 29300 ko | fiat-bedrock2/src/secp256k1_dettman_64.c | 0m00.22s | 29724 ko || +0m00.06s || -424 ko | +31.81% | -1.42% 0m00.28s | 23932 ko | fiat-json/src/secp256k1_dettman_64.json | 0m00.20s | 23928 ko || +0m00.08s || 4 ko | +40.00% | +0.01% 0m00.27s | 25000 ko | fiat-c/src/poly1305_32.c | 0m00.21s | 24124 ko || +0m00.06s || 876 ko | +28.57% | +3.63% 0m00.27s | 21060 ko | fiat-c/src/secp256k1_dettman_64.c | 0m00.18s | 20924 ko || +0m00.09s || 136 ko | +50.00% | +0.64% 0m00.27s | 25404 ko | fiat-java/src/FiatPoly1305.java | 0m00.22s | 25272 ko || +0m00.05s || 132 ko | +22.72% | +0.52% 0m00.27s | 20960 ko | fiat-rust/src/secp256k1_dettman_64.rs | 0m00.18s | 21372 ko || +0m00.09s || -412 ko | +50.00% | -1.92% 0m00.27s | 24624 ko | fiat-zig/src/poly1305_32.zig | 0m00.21s | 23860 ko || +0m00.06s || 764 ko | +28.57% | +3.20% 0m00.27s | 21176 ko | fiat-zig/src/secp256k1_dettman_64.zig | 0m00.18s | 20508 ko || +0m00.09s || 668 ko | +50.00% | +3.25% 0m00.26s | 24636 ko | fiat-rust/src/poly1305_32.rs | 0m00.21s | 24504 ko || +0m00.05s || 132 ko | +23.80% | +0.53% 0m00.18s | 26440 ko | fiat-go/64/poly1305/poly1305.go | 0m00.18s | 26096 ko || +0m00.00s || 344 ko | +0.00% | +1.31% 0m00.17s | 26980 ko | fiat-json/src/poly1305_64.json | 0m00.13s | 27492 ko || +0m00.04s || -512 ko | +30.76% | -1.86% 0m00.16s | 61504 ko | ExtractionOCaml/perf_unsaturated_solinas.cmi | 0m00.16s | 61784 ko || +0m00.00s || -280 ko | +0.00% | -0.45% 0m00.16s | 23780 ko | fiat-rust/src/poly1305_64.rs | 0m00.13s | 22996 ko || +0m00.03s || 784 ko | +23.07% | +3.40% 0m00.16s | 23284 ko | fiat-zig/src/poly1305_64.zig | 0m00.13s | 23176 ko || +0m00.03s || 108 ko | +23.07% | +0.46% 0m00.15s | 58712 ko | ExtractionOCaml/perf_word_by_word_montgomery.cmi | 0m00.15s | 58904 ko || +0m00.00s || -192 ko | +0.00% | -0.32% 0m00.15s | 28228 ko | fiat-bedrock2/src/poly1305_64.c | 0m00.11s | 28056 ko || +0m00.03s || 172 ko | +36.36% | +0.61% 0m00.15s | 23928 ko | fiat-c/src/poly1305_64.c | 0m00.12s | 22980 ko || +0m00.03s || 948 ko | +25.00% | +4.12% ```

--- .../AbstractInterpretation.v | 62 + .../Bottomify/AbstractInterpretation.v | 7 +- .../Fancy/AbstractInterpretation.v | 309 ++--- src/AbstractInterpretation/Fancy/Proofs.v | 1090 +++++++++-------- src/AbstractInterpretation/Fancy/Wf.v | 519 +++++--- src/AbstractInterpretation/Proofs.v | 118 ++ src/AbstractInterpretation/Wf.v | 44 + src/AbstractInterpretation/WfExtra.v | 37 + src/AbstractInterpretation/ZRange.v | 16 +- src/CLI.v | 7 + src/SlowPrimeSynthesisExamples.v | 10 +- 11 files changed, 1385 insertions(+), 834 deletions(-) create mode 100644 src/AbstractInterpretation/AbstractInterpretation.v create mode 100644 src/AbstractInterpretation/Proofs.v create mode 100644 src/AbstractInterpretation/Wf.v create mode 100644 src/AbstractInterpretation/WfExtra.v diff --git a/src/AbstractInterpretation/AbstractInterpretation.v b/src/AbstractInterpretation/AbstractInterpretation.v new file mode 100644 index 0000000000..9ca5feb664 --- /dev/null +++ b/src/AbstractInterpretation/AbstractInterpretation.v @@ -0,0 +1,62 @@ +Require Import Crypto.Util.ZRange. +Require Import Rewriter.Language.Language. +Require Import Crypto.Language.API. +Require Import Crypto.AbstractInterpretation.ZRange. +Require Crypto.AbstractInterpretation.Fancy.AbstractInterpretation. +Require Crypto.AbstractInterpretation.Bottomify.AbstractInterpretation. + +Module Compilers. + Export Language.Compilers. + Export UnderLets.Compilers. + Export Language.API.Compilers. + Export AbstractInterpretation.ZRange.Compilers. + Export Fancy.AbstractInterpretation.Compilers. + Export Bottomify.AbstractInterpretation.Compilers. + + Module partial. + Definition Extract {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow _ t) + := if fancy_and_powerful_but_exponentially_slow_bounds_analysis + then Fancy.AbstractInterpretation.Compilers.partial.Extract assume_cast_truncates e bound + else Bottomify.AbstractInterpretation.Compilers.partial.Extract assume_cast_truncates e bound. + End partial. + + Import API. + Definition PartialEvaluateWithBounds + {fancy_and_powerful_but_exponentially_slow_bounds_analysis : fancy_and_powerful_but_exponentially_slow_bounds_analysis_opt} + {shiftr_avoid_uint1 : shiftr_avoid_uint1_opt} + (relax_zrange : zrange -> option zrange) + (assume_cast_truncates : bool) + (skip_annotations_under : forall t, ident t -> bool) + (strip_preexisting_annotations : bool) + {t} (e : Expr t) + (bound : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + : Expr t + := if fancy_and_powerful_but_exponentially_slow_bounds_analysis + then Fancy.AbstractInterpretation.Compilers.PartialEvaluateWithBounds relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e bound + else Bottomify.AbstractInterpretation.Compilers.PartialEvaluateWithBounds relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e bound. + + Definition PartialEvaluateWithListInfoFromBounds + {fancy_and_powerful_but_exponentially_slow_bounds_analysis : fancy_and_powerful_but_exponentially_slow_bounds_analysis_opt} + {shiftr_avoid_uint1 : shiftr_avoid_uint1_opt} + {t} (e : Expr t) + (bound : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + : Expr t + := if fancy_and_powerful_but_exponentially_slow_bounds_analysis + then Fancy.AbstractInterpretation.Compilers.PartialEvaluateWithListInfoFromBounds e bound + else Bottomify.AbstractInterpretation.Compilers.PartialEvaluateWithListInfoFromBounds e bound. + + Definition CheckedPartialEvaluateWithBounds + {fancy_and_powerful_but_exponentially_slow_bounds_analysis : fancy_and_powerful_but_exponentially_slow_bounds_analysis_opt} + {shiftr_avoid_uint1 : shiftr_avoid_uint1_opt} + (relax_zrange : zrange -> option zrange) + (assume_cast_truncates : bool) + (skip_annotations_under : forall t, ident t -> bool) + (strip_preexisting_annotations : bool) + {t} (E : Expr t) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + (b_out : ZRange.type.base.option.interp (type.final_codomain t)) + : Expr t + (ZRange.type.base.option.interp (type.final_codomain t) * Expr t + list { t : _ & forall var, @expr var t }) + := if fancy_and_powerful_but_exponentially_slow_bounds_analysis + then Fancy.AbstractInterpretation.Compilers.CheckedPartialEvaluateWithBounds relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations E b_in b_out + else Bottomify.AbstractInterpretation.Compilers.CheckedPartialEvaluateWithBounds relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations E b_in b_out. +End Compilers. diff --git a/src/AbstractInterpretation/Bottomify/AbstractInterpretation.v b/src/AbstractInterpretation/Bottomify/AbstractInterpretation.v index c43ead4fc9..4857eabead 100644 --- a/src/AbstractInterpretation/Bottomify/AbstractInterpretation.v +++ b/src/AbstractInterpretation/Bottomify/AbstractInterpretation.v @@ -493,6 +493,7 @@ Module Compilers. Definition eta_expand_with_bound {var} {t} (e : @expr _ t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : expr t := let assume_cast_truncates := false in let shiftr_avoid_uint1 : shiftr_avoid_uint1_opt := false (* this doesn't actually matter for [eta_expand_with_bound], which only invokes [abstract_interp_ident] on things like [fst], [snd], etc *) in + let fancy_and_powerful_but_exponentially_slow_bounds_analysis : fancy_and_powerful_but_exponentially_slow_bounds_analysis_opt := false in (@partial.ident.eta_expand_with_bound) var abstract_domain' annotate_expr bottom' (abstract_interp_ident assume_cast_truncates) extract_list_state extract_option_state is_annotated_for t e bound. @@ -551,7 +552,7 @@ Module Compilers. End CheckCasts. Definition PartialEvaluateWithBounds - {shiftr_avoid_uint1 : shiftr_avoid_uint1_opt} + {opts : AbstractInterpretation.Options} (relax_zrange : zrange -> option zrange) (assume_cast_truncates : bool) (skip_annotations_under : forall t, ident t -> bool) @@ -561,14 +562,14 @@ Module Compilers. : Expr t := partial.EvalWithBound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations (GeneralizeVar.GeneralizeVar (e _)) bound. Definition PartialEvaluateWithListInfoFromBounds - {shiftr_avoid_uint1 : shiftr_avoid_uint1_opt} + {opts : AbstractInterpretation.Options} {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow ZRange.type.option.interp t) : Expr t := partial.EtaExpandWithListInfoFromBound (GeneralizeVar.GeneralizeVar (e _)) bound. Definition CheckedPartialEvaluateWithBounds - {shiftr_avoid_uint1 : shiftr_avoid_uint1_opt} + {opts : AbstractInterpretation.Options} (relax_zrange : zrange -> option zrange) (assume_cast_truncates : bool) (skip_annotations_under : forall t, ident t -> bool) diff --git a/src/AbstractInterpretation/Fancy/AbstractInterpretation.v b/src/AbstractInterpretation/Fancy/AbstractInterpretation.v index c43ead4fc9..681f6f5820 100644 --- a/src/AbstractInterpretation/Fancy/AbstractInterpretation.v +++ b/src/AbstractInterpretation/Fancy/AbstractInterpretation.v @@ -34,7 +34,8 @@ Module Compilers. {var : type -> Type}. Local Notation expr := (@expr base_type ident). Local Notation UnderLets := (@UnderLets base_type ident var). - Context (abstract_domain' : base_type -> Type) + Context {try_make_transport_base_type_cps : @type.try_make_transport_cpsT base_type} + (abstract_domain' : base_type -> Type) (annotate : forall (is_let_bound : bool) t, abstract_domain' t -> @expr var t -> UnderLets (@expr var t)) (bottom' : forall A, abstract_domain' A) (skip_annotations_under : forall t, ident t -> bool). @@ -42,18 +43,25 @@ Module Compilers. Definition abstract_domain (t : type) := type.interp abstract_domain' t. - Fixpoint value (t : type) + Context {default_expr : @DefaultValue.type.base.DefaultT _ (@expr abstract_domain)}. (* needed for impossible cases *) + + (** A value should carry both an abstract state and a way of turning value's into expressions *) + Fixpoint value' (t : type) := match t return Type (* COQBUG(https://github.com/coq/coq/issues/7727) *) with | type.base t - => abstract_domain t * @expr var t + => @expr var t | type.arrow s d - => value s -> UnderLets (value d) + => abstract_domain s * value' s -> UnderLets (value' d) end%type. - Definition value_with_lets (t : type) - := UnderLets (value t). + Definition value (t : type) : Type + := abstract_domain t * UnderLets (value' t). + + Context (ident_extract : forall t, ident t -> abstract_domain t). + Context (interp_ident' : bool (* annotate with state? *) -> forall t, ident t -> UnderLets (value' t)). - Context (interp_ident : bool (* annotate with state? *) -> forall t, ident t -> value_with_lets t). + Definition interp_ident (annotate_with_state : bool) {t} (idc : ident t) : value t + := (ident_extract _ idc, interp_ident' annotate_with_state _ idc). Fixpoint bottom {t} : abstract_domain t := match t with @@ -67,54 +75,50 @@ Module Compilers. | type.arrow s d => (bottom, @bottom_for_each_lhs_of_arrow d) end. - Definition state_of_value {t} : value t -> abstract_domain t - := match t return value t -> abstract_domain t with - | type.base t => fun '(st, v) => st - | type.arrow s d => fun _ => bottom - end. + Definition state_of_value {t} : value t -> abstract_domain t := fst. + Definition project_value' {t} : value t -> UnderLets (value' t) := snd. + Definition Base_value' {t} : abstract_domain t * value' t -> value t + := fun x => (fst x, Base (snd x)). + Definition apply_value {s d} (f : value (s -> d)) (x : value s) : value d + := let '(x1, x2) := (state_of_value x, project_value' x) in + let '(f1, f2) := (state_of_value f, project_value' f) in + (f1 x1, (f2 <-- f2; x2 <-- x2; f2 (x1, x2))%under_lets). - (** We need to make sure that we ignore the state of - higher-order arrows *everywhere*, or else the proofs don't go - through. So we sometimes need to replace the state of - arrow-typed values with [⊥]. *) - Fixpoint bottomify {t} : value t -> value_with_lets t - := match t return value t -> value_with_lets t with - | type.base t => fun '(st, v) => Base (bottom' t, v) - | type.arrow s d => fun f => Base (fun x => fx <-- f x; @bottomify d fx) - end%under_lets. - - (** We drop the state of higher-order arrows *) Fixpoint reify (annotate_with_state : bool) (is_let_bound : bool) {t} : value t -> type.for_each_lhs_of_arrow abstract_domain t -> UnderLets (@expr var t) := match t return value t -> type.for_each_lhs_of_arrow abstract_domain t -> UnderLets (@expr var t) with | type.base t => fun '(st, v) 'tt - => if annotate_with_state - then annotate is_let_bound t st v - else if is_let_bound - then UnderLets.UnderLet v (fun v' => UnderLets.Base ($$v')) - else UnderLets.Base v + => (v <-- v; + if annotate_with_state + then annotate is_let_bound t st v + else if is_let_bound + then UnderLets.UnderLet v (fun v' => UnderLets.Base ($$v')) + else UnderLets.Base v) | type.arrow s d => fun f_e '(sv, dv) - => let sv := match s with - | type.base _ => sv - | type.arrow _ _ => bottom - end in - Base - (λ x , (UnderLets.to_expr - (fx <-- f_e (@reflect annotate_with_state _ (expr.Var x) sv); - @reify annotate_with_state false _ fx dv))) - end%core%expr + => (let f := state_of_value f_e in + f_e <-- project_value' f_e; + Base + (λ x , UnderLets.to_expr + (let sv := @reflect annotate_with_state _ (expr.Var x) sv in + let sx := state_of_value sv in + x <-- project_value' sv; + fx <-- f_e (sx, x)%core; + @reify annotate_with_state false _ (f sx, Base fx)%core dv))) + end%core%expr%under_lets with reflect (annotate_with_state : bool) {t} : @expr var t -> abstract_domain t -> value t := match t return @expr var t -> abstract_domain t -> value t with | type.base t - => fun e st => (st, e) + => fun e st => Base_value' (st, e) | type.arrow s d => fun e absf - => (fun v - => let stv := state_of_value v in - (rv <-- (@reify annotate_with_state false s v bottom_for_each_lhs_of_arrow); - Base (@reflect annotate_with_state d (e @ rv) (absf stv))%expr)) - end%under_lets. + => Base_value' + (absf, + (fun v + => rv <-- @reify annotate_with_state false s (Base_value' v) bottom_for_each_lhs_of_arrow; + (* TODO: Should we be feeding in [fst v], or should we bottom out the arguments to [fst v]? *) + project_value' (@reflect annotate_with_state d (e @ rv) (absf (fst v)))%expr)) + end%under_lets%core. Definition skip_annotations_for_App {var'} {t} (e : @expr var' t) : bool := match invert_AppIdent_curried e with @@ -122,64 +126,74 @@ Module Compilers. | None => false end. - Fixpoint interp (annotate_with_state : bool) {t} (e : @expr value_with_lets t) : value_with_lets t + Definition invert_default {A B} (t : type) (x : option { t : type & @expr abstract_domain (A t) * @expr abstract_domain (B t) }%type) : @expr abstract_domain (A t) * @expr abstract_domain (B t) + := Option.value + (type.try_transport + _ _ t + (projT2 + (Option.value x (existT _ t (DefaultValue.type.base.defaultv, DefaultValue.type.base.defaultv))))) + (DefaultValue.type.base.defaultv, DefaultValue.type.base.defaultv). + + Definition invert_default' {A B C} (t : type) (x : option { t : type & @expr abstract_domain (A t) * (abstract_domain (B t) -> @expr abstract_domain (C t)) }%type) : @expr abstract_domain (A t) * (abstract_domain (B t) -> @expr abstract_domain (C t)) + := Option.value + (type.try_transport + _ _ t + (projT2 + (Option.value x (existT _ t (DefaultValue.type.base.defaultv, fun _ => DefaultValue.type.base.defaultv))))) + (DefaultValue.type.base.defaultv, fun _ => DefaultValue.type.base.defaultv). + + Fixpoint interp (annotate_with_state : bool) {t} (e : @expr value t) : @expr abstract_domain t -> value t := let annotate_with_state := annotate_with_state && negb (skip_annotations_for_App e) in - match e in expr.expr t return value_with_lets t with - | expr.Ident t idc => interp_ident annotate_with_state _ idc (* Base (reflect (###idc) (abstract_interp_ident _ idc))*) - | expr.Var t v => v - | expr.Abs s d f => Base (fun x => @interp annotate_with_state d (f (Base x))) - | expr.App (type.base s) d f x - => (x' <-- @interp annotate_with_state _ x; - f' <-- @interp annotate_with_state (_ -> d)%etype f; - f' x') - | expr.App (type.arrow s' d') d f x - => (x' <-- @interp annotate_with_state (s' -> d')%etype x; - x'' <-- bottomify x'; - f' <-- @interp annotate_with_state (_ -> d)%etype f; - f' x'') - | expr.LetIn (type.arrow _ _) B x f - => (x' <-- @interp annotate_with_state _ x; - @interp annotate_with_state _ (f (Base x'))) - | expr.LetIn (type.base A) B x f - => (x' <-- @interp annotate_with_state _ x; - x'' <-- reify annotate_with_state true (* this forces a let-binder here *) x' tt; - @interp annotate_with_state _ (f (Base (reflect annotate_with_state x'' (state_of_value x'))))) + let expr_interp {t} := expr.interp (@ident_extract) (t:=t) in + match e in expr.expr t return @expr abstract_domain t -> value t with + | expr.Ident t idc => fun _ => @interp_ident annotate_with_state _ idc (* Base (reflect (###idc) (abstract_interp_ident _ idc))*) + | expr.Var t v => fun _ => v + | expr.Abs s d f + => fun fe_st + => let f_st := expr_interp fe_st in + let fe_st := Option.value (invert_Abs fe_st) ((* should never happen *)fun _ => DefaultValue.type.base.defaultv) in + Base_value' + (f_st, + (fun x + => project_value' (@interp annotate_with_state d (f (Base_value' x)) (fe_st (fst x))))) + | expr.App s d f x + => fun fx_st + => (let '(f_st, x_st) := invert_default s (invert_App fx_st) in + let x := @interp annotate_with_state s x x_st in + let f := @interp annotate_with_state (s -> d)%etype f f_st in + apply_value f x) + | expr.LetIn (type.base A' as A) B x f + => fun st + => let '(x_st, f_st) := invert_default' A (invert_LetIn st) in + let x := @interp annotate_with_state _ x x_st in + let x_st := state_of_value x in + let fx_st := f_st x_st in + (expr_interp fx_st, + x <-- reify annotate_with_state true (* this forces a let-binder here *) x tt; + project_value' (@interp annotate_with_state _ (f (reflect annotate_with_state x x_st)) fx_st)) + | expr.LetIn (type.arrow _ _ as A) B x f + => fun st + => let '(x_st, f_st) := invert_default' A (invert_LetIn st) in + let x := @interp annotate_with_state _ x x_st in + @interp annotate_with_state _ (f x) (f_st (state_of_value x)) end%under_lets. - Definition eval_with_bound' (annotate_with_state : bool) {t} (e : @expr value_with_lets t) + Definition eval_with_bound' (annotate_with_state : bool) {t} (e : @expr value t) (e_st : @expr abstract_domain t) (st : type.for_each_lhs_of_arrow abstract_domain t) : expr t - := UnderLets.to_expr (e' <-- interp annotate_with_state e; reify annotate_with_state false e' st). + := UnderLets.to_expr (reify annotate_with_state false (interp annotate_with_state e e_st) st). - Definition eval' {t} (e : @expr value_with_lets t) : expr t - := eval_with_bound' false e bottom_for_each_lhs_of_arrow. + Definition eval' {t} (e : @expr value t) (e_st : @expr abstract_domain t) : expr t + := eval_with_bound' false e e_st bottom_for_each_lhs_of_arrow. Definition eta_expand_with_bound' {t} (e : @expr var t) (st : type.for_each_lhs_of_arrow abstract_domain t) : expr t := UnderLets.to_expr (reify true false (reflect true e bottom) st). - Section extract. - Context (ident_extract : forall t, ident t -> abstract_domain t). - - (** like [expr.interp (@ident_extract) e], except we replace - all higher-order state with bottom *) - Fixpoint extract' {t} (e : @expr abstract_domain t) : abstract_domain t - := match e in expr.expr t return abstract_domain t with - | expr.Ident t idc => ident_extract t idc - | expr.Var t v => v - | expr.Abs s d f => fun v : abstract_domain s => @extract' _ (f v) - | expr.App (type.base s) d f x - => @extract' _ f (@extract' _ x) - | expr.App (type.arrow s' d') d f x - => @extract' _ f (@bottom (type.arrow s' d')) - | expr.LetIn A B x f => dlet y := @extract' _ x in @extract' _ (f y) - end. - - Definition extract_gen {t} (e : @expr abstract_domain t) (bound : type.for_each_lhs_of_arrow abstract_domain t) - : abstract_domain' (type.final_codomain t) - := type.app_curried (extract' e) bound. - End extract. + Definition extract_gen {t} (e : @expr abstract_domain t) (bound : type.for_each_lhs_of_arrow abstract_domain t) + : abstract_domain' (type.final_codomain t) + := type.app_curried (expr.interp (@ident_extract) e) bound. End with_var. Module ident. @@ -194,8 +208,8 @@ Module Compilers. Local Notation UnderLets := (@UnderLets base.type ident var). Context (abstract_domain' : base.type -> Type). Local Notation abstract_domain := (@abstract_domain base.type abstract_domain'). - Local Notation value_with_lets := (@value_with_lets base.type ident var abstract_domain'). Local Notation value := (@value base.type ident var abstract_domain'). + Local Notation value' := (@value' base.type ident var abstract_domain'). Context (annotate_expr : forall t, abstract_domain' t -> option (@expr var (t -> t))) (bottom' : forall A, abstract_domain' A) (abstract_interp_ident : forall t, ident t -> type.interp abstract_domain' t) @@ -204,7 +218,7 @@ Module Compilers. (is_annotated_for : forall t t', @expr var t -> abstract_domain' t' -> bool) (annotation_to_cast : forall s d, @expr var (s -> d) -> option (@expr var s -> @expr var d)) (skip_annotations_under : forall t, ident t -> bool) - (strip_annotation : forall t, ident t -> option (value t)). + (strip_annotation : forall t, ident t -> option (value' t)). (** TODO: Is it okay to commute annotations? *) Definition update_annotation {t} (st : abstract_domain' t) (e : @expr var t) : @expr var t @@ -286,34 +300,35 @@ Module Compilers. Local Notation reify := (@reify base.type ident var abstract_domain' annotate bottom'). Local Notation reflect := (@reflect base.type ident var abstract_domain' annotate bottom'). + Local Notation apply_value := (@apply_value base.type ident var abstract_domain'). + Local Notation Base_value' := (@Base_value' base.type ident var abstract_domain'). (** We manually rewrite with the rule for [nth_default], as the eliminator for eta-expanding lists in the input *) - Definition interp_ident (annotate_with_state : bool) {t} (idc : ident t) : value_with_lets t - := match idc in ident t return value_with_lets t with + Definition interp_ident' (annotate_with_state : bool) {t} (idc : ident t) : UnderLets (value' t) + := match idc in ident t return UnderLets (value' t) with | ident.List_nth_default T as idc - => let default := reflect annotate_with_state (###idc) (abstract_interp_ident _ idc) in - Base - (fun default_arg - => default <-- default default_arg; - Base - (fun ls_arg - => default <-- default ls_arg; - Base - (fun n_arg - => default <-- default n_arg; - ls' <-- @reify annotate_with_state false (base.type.list T) ls_arg tt; - Base - (fst default, - match reflect_list ls', invert_Literal (snd n_arg) with - | Some ls, Some n - => nth_default (snd default_arg) ls n - | _, _ => snd default - end)))) - | idc => Base - match strip_annotation _ idc with - | Some v => v - | None => reflect annotate_with_state (###idc) (abstract_interp_ident _ idc) - end + => let '(default_st, default) := reflect annotate_with_state (###idc) (abstract_interp_ident _ idc) in + (default <-- default; + Base + (fun default_arg + => default <-- default default_arg; + Base + (fun ls_arg + => default <-- default ls_arg; + Base + (fun n_arg + => default <-- default n_arg; + ls' <-- @reify annotate_with_state false (base.type.list T) (Base_value' ls_arg) tt; + Base + match reflect_list ls', invert_Literal (snd n_arg) with + | Some ls, Some n + => nth_default (snd default_arg) ls n + | _, _ => default + end)))) + | idc => match strip_annotation _ idc with + | Some v => Base v + | None => project_value' _ (reflect annotate_with_state (###idc) (abstract_interp_ident _ idc)) + end end%core%under_lets%expr. Fixpoint strip_all_annotations' (should_strip : bool) {t} (e : @expr var t) : @expr var t @@ -343,13 +358,13 @@ Module Compilers. Definition strip_all_annotations {t} (e : @expr var t) : @expr var t := @strip_all_annotations' false t e. - Definition eval_with_bound (skip_annotations_under : forall t, ident t -> bool) (annotate_with_state : bool) {t} (e : @expr value_with_lets t) + Definition eval_with_bound (skip_annotations_under : forall t, ident t -> bool) (annotate_with_state : bool) {t} (e : @expr value t) (e_st : @expr abstract_domain t) (st : type.for_each_lhs_of_arrow abstract_domain t) : @expr var t - := @eval_with_bound' base.type ident var abstract_domain' annotate bottom' skip_annotations_under interp_ident annotate_with_state t e st. + := @eval_with_bound' base.type ident var _ abstract_domain' annotate bottom' skip_annotations_under _ abstract_interp_ident interp_ident' annotate_with_state t e e_st st. - Definition eval {t} (e : @expr value_with_lets t) : @expr var t - := @eval' base.type ident var abstract_domain' annotate bottom' (fun _ _ => false) interp_ident t e. + Definition eval {t} (e : @expr value t) (e_st : @expr abstract_domain t) : @expr var t + := @eval' base.type ident var _ abstract_domain' annotate bottom' (fun _ _ => false) _ abstract_interp_ident interp_ident' t e e_st. Definition eta_expand_with_bound {t} (e : @expr var t) (st : type.for_each_lhs_of_arrow abstract_domain t) @@ -357,7 +372,7 @@ Module Compilers. := @eta_expand_with_bound' base.type ident var abstract_domain' annotate bottom' t e st. Definition extract {t} (e : @expr _ t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : abstract_domain' (type.final_codomain t) - := @extract_gen base.type ident abstract_domain' bottom' abstract_interp_ident t e bound. + := @extract_gen base.type ident abstract_domain' abstract_interp_ident t e bound. End with_var. End ident. @@ -366,8 +381,8 @@ Module Compilers. Section specialized. Local Notation abstract_domain' := ZRange.type.base.option.interp. Local Notation abstract_domain := (@partial.abstract_domain base.type abstract_domain'). - Local Notation value_with_lets var := (@value_with_lets base.type ident var abstract_domain'). Local Notation value var := (@value base.type ident var abstract_domain'). + Local Notation value' var := (@value' base.type ident var abstract_domain'). Notation expr := (@expr base.type ident). Notation Expr := (@expr.Expr base.type ident). Local Notation type := (type base.type). @@ -412,8 +427,8 @@ Module Compilers. Definition abstract_interp_ident {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) t (idc : ident t) : type.interp abstract_domain' t := ZRange.ident.option.interp assume_cast_truncates idc. - Definition always_strip_annotation {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {var} t (idc : ident t) : option (value var t) - := match idc in ident t return option (value var t) with + Definition always_strip_annotation {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {var} t (idc : ident t) : option (value' var t) + := match idc in ident t return option (value' var t) with | ident.Z_cast as idc | ident.Z_cast2 as idc => let tZ_type := (fun t (idc : ident (type.base _ -> type.base t -> type.base t)) => t) _ idc in (* we want Coq to pick up the Z type for bounds tightness, not the zrange type (where "tighter" means "equal") *) @@ -422,8 +437,7 @@ Module Compilers. => Base (fun '(r_known, e) => Base - ((abstract_interp_ident assume_cast_truncates _ idc r_orig r_known, - let do_strip + (let do_strip := match ZRange.type.base.option.lift_Some r_known, ZRange.type.base.option.lift_Some r_orig with | Some r_known, Some r_orig => ZRange.type.base.is_tighter_than @@ -433,11 +447,11 @@ Module Compilers. end in if do_strip then e - else (###idc @ re @ e)%expr)))) + else (###idc @ re @ e)%expr))) | _ => None end. - Definition strip_annotation {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) (strip_annotations : bool) {var} : forall t, ident t -> option (value var t) + Definition strip_annotation {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) (strip_annotations : bool) {var} : forall t, ident t -> option (value' var t) := if strip_annotations then always_strip_annotation assume_cast_truncates else fun _ _ => None. @@ -485,49 +499,54 @@ Module Compilers. Definition strip_all_annotations strip_annotations_under {var} {t} (e : @expr var t) : @expr var t := @ident.strip_all_annotations var (@annotation_to_cast _) strip_annotations_under t e. - Definition eval_with_bound {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) (skip_annotations_under : forall t, ident t -> bool) (strip_preexisting_annotations : bool) {var} {t} (e : @expr _ t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : expr t + Definition eval_with_bound {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) (skip_annotations_under : forall t, ident t -> bool) (strip_preexisting_annotations : bool) {var} {t} (e e_st : @expr _ t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : expr t := let annotate_with_state := true in (@partial.ident.eval_with_bound) - var abstract_domain' annotate_expr bottom' (abstract_interp_ident assume_cast_truncates) extract_list_state extract_option_state is_annotated_for (strip_annotation assume_cast_truncates strip_preexisting_annotations) skip_annotations_under annotate_with_state t e bound. + var abstract_domain' annotate_expr bottom' (abstract_interp_ident assume_cast_truncates) extract_list_state extract_option_state is_annotated_for (strip_annotation assume_cast_truncates strip_preexisting_annotations) skip_annotations_under annotate_with_state t e e_st bound. Definition eta_expand_with_bound {var} {t} (e : @expr _ t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : expr t := let assume_cast_truncates := false in let shiftr_avoid_uint1 : shiftr_avoid_uint1_opt := false (* this doesn't actually matter for [eta_expand_with_bound], which only invokes [abstract_interp_ident] on things like [fst], [snd], etc *) in + let fancy_and_powerful_but_exponentially_slow_bounds_analysis : fancy_and_powerful_but_exponentially_slow_bounds_analysis_opt := false in (@partial.ident.eta_expand_with_bound) var abstract_domain' annotate_expr bottom' (abstract_interp_ident assume_cast_truncates) extract_list_state extract_option_state is_annotated_for t e bound. Definition StripAllAnnotations strip_annotations_under {t} (e : Expr t) : Expr t := fun var => strip_all_annotations strip_annotations_under (e _). - Definition EvalWithBound {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) (skip_annotations_under : forall t, ident t -> bool) (strip_preexisting_annotations : bool) {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : Expr t - := fun var => eval_with_bound assume_cast_truncates skip_annotations_under strip_preexisting_annotations (e _) bound. + Definition EvalWithBound {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) (skip_annotations_under : forall t, ident t -> bool) (strip_preexisting_annotations : bool) {t} (E : Expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : Expr t + := dlet_nd e := GeneralizeVar.ToFlat E in + let E := GeneralizeVar.FromFlat e in + fun var => eval_with_bound assume_cast_truncates skip_annotations_under strip_preexisting_annotations (E _) (E _) bound. Definition EtaExpandWithBound {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : Expr t := fun var => eta_expand_with_bound (e _) bound. End with_relax. - Definition strip_annotations {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {var} {t} (e : @expr _ t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : expr t + Definition strip_annotations {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {var} {t} (e e_st : @expr _ t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : expr t := let strip_preexisting_annotations := true in let annotate_with_state := false in let skip_annotations_under := fun _ _ => false in (@partial.ident.eval_with_bound) - var abstract_domain' (annotate_expr default_relax_zrange) bottom' (abstract_interp_ident assume_cast_truncates) extract_list_state extract_option_state (is_annotated_for default_relax_zrange) (strip_annotation assume_cast_truncates strip_preexisting_annotations) skip_annotations_under annotate_with_state t e bound. + var abstract_domain' (annotate_expr default_relax_zrange) bottom' (abstract_interp_ident assume_cast_truncates) extract_list_state extract_option_state (is_annotated_for default_relax_zrange) (strip_annotation assume_cast_truncates strip_preexisting_annotations) skip_annotations_under annotate_with_state t e e_st bound. Definition StripAnnotations {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : Expr t - := fun var => strip_annotations assume_cast_truncates (e _) bound. + := fun var => strip_annotations assume_cast_truncates (e _) (e _) bound. - Definition eval {opts : AbstractInterpretation.Options} {var} {t} (e : @expr _ t) : expr t + Definition eval {opts : AbstractInterpretation.Options} {var} {t} (e e_st : @expr _ t) : expr t := let assume_cast_truncates := false in let strip_preexisting_annotations := false in (@partial.ident.eval) - var abstract_domain' (annotate_expr default_relax_zrange) bottom' (abstract_interp_ident assume_cast_truncates) extract_list_state extract_option_state (is_annotated_for default_relax_zrange) (strip_annotation assume_cast_truncates strip_preexisting_annotations) t e. - Definition Eval {opts : AbstractInterpretation.Options} {t} (e : Expr t) : Expr t - := fun var => eval (e _). + var abstract_domain' (annotate_expr default_relax_zrange) bottom' (abstract_interp_ident assume_cast_truncates) extract_list_state extract_option_state (is_annotated_for default_relax_zrange) (strip_annotation assume_cast_truncates strip_preexisting_annotations) t e e_st. + Definition Eval {opts : AbstractInterpretation.Options} {t} (E : Expr t) : Expr t + := dlet_nd e := GeneralizeVar.ToFlat E in + let E := GeneralizeVar.FromFlat e in + fun var => eval (E _) (E _). Definition EtaExpandWithListInfoFromBound {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : Expr t := EtaExpandWithBound default_relax_zrange e (type.map_for_each_lhs_of_arrow (@ZRange.type.option.strip_ranges) bound). Definition extract {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {t} (e : expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : abstract_domain' (type.final_codomain t) - := @partial.ident.extract abstract_domain' bottom' (abstract_interp_ident assume_cast_truncates) t e bound. + := @partial.ident.extract abstract_domain' (abstract_interp_ident assume_cast_truncates) t e bound. Definition Extract {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow abstract_domain t) : abstract_domain' (type.final_codomain t) - := @partial.ident.extract abstract_domain' bottom' (abstract_interp_ident assume_cast_truncates) t (e _) bound. + := @partial.ident.extract abstract_domain' (abstract_interp_ident assume_cast_truncates) t (e _) bound. End specialized. End partial. Import API. @@ -551,7 +570,7 @@ Module Compilers. End CheckCasts. Definition PartialEvaluateWithBounds - {shiftr_avoid_uint1 : shiftr_avoid_uint1_opt} + {opts : AbstractInterpretation.Options} (relax_zrange : zrange -> option zrange) (assume_cast_truncates : bool) (skip_annotations_under : forall t, ident t -> bool) @@ -559,16 +578,16 @@ Module Compilers. {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow ZRange.type.option.interp t) : Expr t - := partial.EvalWithBound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations (GeneralizeVar.GeneralizeVar (e _)) bound. + := partial.EvalWithBound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e bound. Definition PartialEvaluateWithListInfoFromBounds - {shiftr_avoid_uint1 : shiftr_avoid_uint1_opt} + {opts : AbstractInterpretation.Options} {t} (e : Expr t) (bound : type.for_each_lhs_of_arrow ZRange.type.option.interp t) : Expr t := partial.EtaExpandWithListInfoFromBound (GeneralizeVar.GeneralizeVar (e _)) bound. Definition CheckedPartialEvaluateWithBounds - {shiftr_avoid_uint1 : shiftr_avoid_uint1_opt} + {opts : AbstractInterpretation.Options} (relax_zrange : zrange -> option zrange) (assume_cast_truncates : bool) (skip_annotations_under : forall t, ident t -> bool) diff --git a/src/AbstractInterpretation/Fancy/Proofs.v b/src/AbstractInterpretation/Fancy/Proofs.v index 515c3a4d61..ef4284d6b2 100644 --- a/src/AbstractInterpretation/Fancy/Proofs.v +++ b/src/AbstractInterpretation/Fancy/Proofs.v @@ -3,6 +3,7 @@ Require Import Coq.ZArith.ZArith. Require Import Coq.Classes.Morphisms. Require Import Coq.Classes.RelationPairs. Require Import Coq.Relations.Relations. +Require Import Crypto.Util.Bool.Reflect. Require Import Crypto.Util.ZRange. Require Import Crypto.Util.ZRange.Operations. Require Import Crypto.Util.ZRange.BasicLemmas. @@ -100,80 +101,133 @@ Module Compilers. (bottom'_related : forall t v, abstraction_relation' t (bottom' t) v) (skip_annotations_under : forall t, ident t -> bool) (abstract_interp_ident : forall t, ident t -> type.interp abstract_domain' t) - (ident_interp_Proper : forall t (idc : ident t), type.related_hetero abstraction_relation' (abstract_interp_ident t idc) (ident_interp t idc)) + (abstract_domain'_R : forall t, abstract_domain' t -> abstract_domain' t -> Prop). + Local Notation abstract_domain := (@abstract_domain base_type abstract_domain'). + Local Notation abstract_domain_R := (@abstract_domain_R base_type abstract_domain' abstract_domain'_R). + (* we only care about inputs that are extensional, so we restrict this here *) + (* skip_base:=true because abstract_domain'_R refl and eqv refl follows from abstraction_relation' *) + Definition abstraction_relation {t} : abstract_domain t -> type.interp base_interp t -> Prop + := type.related_hetero_and_Proper (skip_base:=true) (@abstract_domain'_R) (fun _ => eq) (@abstraction_relation'). + (*Definition abstraction_relation {t} : abstract_domain t -> type.interp base_interp t -> Prop + := type.related_hetero (@abstraction_relation').*) + Context (ident_interp_Proper : forall t (idc : ident t), abstraction_relation (abstract_interp_ident t idc) (ident_interp t idc)) (ident_interp_Proper' : forall t, Proper (eq ==> type.eqv) (ident_interp t)) - (abstract_domain'_R : forall t, abstract_domain' t -> abstract_domain' t -> Prop) (abstraction_relation'_Proper : forall t, Proper (abstract_domain'_R t ==> eq ==> Basics.impl) (abstraction_relation' t)) {abstract_domain'_R_transitive : forall t, Transitive (@abstract_domain'_R t)} {abstract_domain'_R_symmetric : forall t, Symmetric (@abstract_domain'_R t)} {bottom'_Proper : forall t, Proper (abstract_domain'_R t) (bottom' t)} + (base_type_beq : base_type -> base_type -> bool) + {reflect_base_type_beq : reflect_rel eq base_type_beq} + {try_make_transport_base_type_cps : @type.try_make_transport_cpsT base_type} + {try_make_transport_base_type_cps_correct : type.try_make_transport_cps_correctT base_type} (abstract_domain'_R_of_related : forall t st v, @abstraction_relation' t st v -> @abstract_domain'_R t st st). - Local Notation abstract_domain := (@abstract_domain base_type abstract_domain'). - Definition abstraction_relation {t} : abstract_domain t -> type.interp base_interp t -> Prop - := type.related_hetero (@abstraction_relation'). Local Notation bottom := (@bottom base_type abstract_domain' (@bottom')). Local Notation bottom_for_each_lhs_of_arrow := (@bottom_for_each_lhs_of_arrow base_type abstract_domain' (@bottom')). - Local Notation abstract_domain_R := (@abstract_domain_R base_type abstract_domain' abstract_domain'_R). Local Notation var := (type.interp base_interp). Local Notation expr := (@expr.expr base_type ident). Local Notation UnderLets := (@UnderLets base_type ident). + Local Notation value' := (@value' base_type ident var abstract_domain'). Local Notation value := (@value base_type ident var abstract_domain'). - Local Notation value_with_lets := (@value_with_lets base_type ident var abstract_domain'). - Local Notation state_of_value := (@state_of_value base_type ident var abstract_domain' bottom'). + Local Notation state_of_value := (@state_of_value base_type ident var abstract_domain'). + Local Notation project_value' := (@project_value' base_type ident var abstract_domain'). + Local Notation extract_gen := (@extract_gen base_type ident abstract_domain' abstract_interp_ident). + Context {default_expr : @DefaultValue.type.base.DefaultT _ (@expr abstract_domain)}. (* needed for impossible cases *) Context (annotate : forall (is_let_bound : bool) t, abstract_domain' t -> @expr var t -> @UnderLets var (@expr var t)) - (interp_ident : bool -> forall t, ident t -> value_with_lets t) - (ident_extract : forall t, ident t -> abstract_domain t) + (interp_ident' : bool -> forall t, ident t -> @UnderLets var (value' t)) (interp_annotate : forall is_let_bound t st e (He : abstraction_relation' t st (expr.interp (t:=type.base t) (@ident_interp) e)), expr.interp (@ident_interp) (UnderLets.interp (@ident_interp) (@annotate is_let_bound t st e)) = expr.interp (@ident_interp) e) - (ident_extract_Proper : forall t, Proper (eq ==> abstract_domain_R) (ident_extract t)). + (abstract_interp_ident_Proper : forall t, Proper (eq ==> abstract_domain_R) (abstract_interp_ident t)). Local Notation eta_expand_with_bound' := (@eta_expand_with_bound' base_type ident _ abstract_domain' annotate bottom'). - Local Notation eval_with_bound' := (@partial.eval_with_bound' base_type ident _ abstract_domain' annotate bottom' skip_annotations_under interp_ident). - Local Notation extract' := (@extract' base_type ident abstract_domain' bottom' ident_extract). - Local Notation extract_gen := (@extract_gen base_type ident abstract_domain' bottom' ident_extract). + Local Notation eval_with_bound' := (@partial.eval_with_bound' base_type ident _ try_make_transport_base_type_cps abstract_domain' annotate bottom' skip_annotations_under default_expr abstract_interp_ident interp_ident'). Local Notation reify := (@reify base_type ident _ abstract_domain' annotate bottom'). Local Notation reflect := (@reflect base_type ident _ abstract_domain' annotate bottom'). - Local Notation interp := (@interp base_type ident var abstract_domain' annotate bottom' skip_annotations_under interp_ident). - Local Notation bottomify := (@bottomify base_type ident _ abstract_domain' bottom'). - - Lemma bottom_related t v : @abstraction_relation t bottom v. - Proof using bottom'_related. cbv [abstraction_relation]; induction t; cbn; cbv [respectful_hetero]; eauto. Qed. - - Local Hint Resolve bottom_related : core typeclass_instances. + Local Notation interp := (@interp base_type ident var try_make_transport_base_type_cps abstract_domain' annotate bottom' skip_annotations_under default_expr abstract_interp_ident interp_ident'). + Local Notation invert_default := (@invert_default base_type ident try_make_transport_base_type_cps abstract_domain' default_expr). + Local Notation invert_default' := (@invert_default' base_type ident try_make_transport_base_type_cps abstract_domain' default_expr). + Local Notation interp_ident := (@interp_ident base_type ident var abstract_domain' abstract_interp_ident interp_ident'). + + Lemma abstract_domain_R_refl_of_abstraction_relation {t st v} + : @abstraction_relation t st v -> Proper abstract_domain_R st. + Proof using abstract_domain'_R_of_related. apply type.related_hetero_and_Proper_proj1, abstract_domain'_R_of_related. Qed. + Local Hint Immediate abstract_domain_R_refl_of_abstraction_relation : core typeclass_instances. + Lemma eqv_refl_of_abstraction_relation {t st v} + : @abstraction_relation t st v -> v == v. + Proof using Type. apply type.related_hetero_and_Proper_proj2; repeat intro; hnf; auto. Qed. + Local Hint Immediate eqv_refl_of_abstraction_relation : core typeclass_instances. + + Lemma eqv_refl_of_abstraction_relation_for_each_lhs_of_arrow {t st v} + : type.and_for_each_lhs_of_arrow (t:=t) (@abstraction_relation) st v -> type.and_for_each_lhs_of_arrow (t:=t) (@type.eqv) v v. + Proof using Type. + apply type.and_for_each_lhs_of_arrow_Proper_impl_hetero2, @eqv_refl_of_abstraction_relation. + Qed. + Local Hint Immediate eqv_refl_of_abstraction_relation_for_each_lhs_of_arrow : core typeclass_instances. - Lemma bottom_for_each_lhs_of_arrow_related t v : type.and_for_each_lhs_of_arrow (@abstraction_relation) (@bottom_for_each_lhs_of_arrow t) v. - Proof using bottom'_related. induction t; cbn; eauto using bottom_related. Qed. + Lemma abstract_domain_R_refl_of_abstraction_relation_for_each_lhs_of_arrow {t st v} + : type.and_for_each_lhs_of_arrow (t:=t) (@abstraction_relation) st v -> type.and_for_each_lhs_of_arrow (t:=t) (@abstract_domain_R) st st. + Proof using abstract_domain'_R_of_related. + apply type.and_for_each_lhs_of_arrow_Proper_impl_hetero1, @abstract_domain_R_refl_of_abstraction_relation. + Qed. + Local Hint Immediate abstract_domain_R_refl_of_abstraction_relation_for_each_lhs_of_arrow : core typeclass_instances. + + Lemma abstraction_relation_iff_app_curried {t st v} + : @abstraction_relation t st v + <-> (Proper (@abstract_domain_R t) st + /\ v == v + /\ forall x y, + type.and_for_each_lhs_of_arrow (@abstraction_relation) x y + -> abstraction_relation' _ (type.app_curried st x) (type.app_curried v y)). + Proof using abstract_domain'_R_of_related. + apply type.related_hetero_and_Proper_iff_app_curried; eauto; repeat intro; hnf; eauto. + Qed. Local Notation bottom_Proper := (@bottom_Proper base_type abstract_domain' bottom' abstract_domain'_R bottom'_Proper). Local Notation bottom_for_each_lhs_of_arrow_Proper := (@bottom_for_each_lhs_of_arrow_Proper base_type abstract_domain' bottom' abstract_domain'_R bottom'_Proper). Local Hint Resolve (@bottom_Proper) (@bottom_for_each_lhs_of_arrow_Proper) : core typeclass_instances. - Fixpoint related_bounded_value {t} : abstract_domain t -> value t -> type.interp base_interp t -> Prop - := match t return abstract_domain t -> value t -> type.interp base_interp t -> Prop with + Lemma bottom_related t v : v == v -> @abstraction_relation t bottom v. + Proof using bottom'_Proper bottom'_related. + induction t; cbv [abstraction_relation]; cbn [type.related_hetero_and_Proper]; cbv [respectful_hetero]; repeat split; eauto. + repeat intro; cbn. + repeat first [ do_with_exactly_one_hyp ltac:(fun H => apply H) + | eassumption + | eapply eqv_refl_of_abstraction_relation ]. + Qed. + + Local Hint Immediate bottom_related : core typeclass_instances. + + Lemma related_bottom_for_each_lhs_of_arrow t v : type.and_for_each_lhs_of_arrow (@type.eqv) v v -> type.and_for_each_lhs_of_arrow (@abstraction_relation) (@bottom_for_each_lhs_of_arrow t) v. + Proof using bottom'_Proper bottom'_related. induction t; cbn; split; split_and; eauto using bottom_related. Qed. + Local Hint Immediate related_bottom_for_each_lhs_of_arrow : core. + + Lemma related_bottom_for_each_lhs_of_arrow_hetero_eqv_l t v v' : type.and_for_each_lhs_of_arrow (@type.eqv) v v' -> type.and_for_each_lhs_of_arrow (@abstraction_relation) (@bottom_for_each_lhs_of_arrow t) v. + Proof using bottom'_Proper bottom'_related. intro; apply related_bottom_for_each_lhs_of_arrow; etransitivity; (idtac + symmetry); eassumption. Qed. + Lemma related_bottom_for_each_lhs_of_arrow_hetero_eqv_r t v v' : type.and_for_each_lhs_of_arrow (@type.eqv) v' v -> type.and_for_each_lhs_of_arrow (@abstraction_relation) (@bottom_for_each_lhs_of_arrow t) v. + Proof using bottom'_Proper bottom'_related. intro; apply related_bottom_for_each_lhs_of_arrow; etransitivity; (idtac + symmetry); eassumption. Qed. + Local Hint Immediate related_bottom_for_each_lhs_of_arrow_hetero_eqv_l related_bottom_for_each_lhs_of_arrow_hetero_eqv_r : core. + + Fixpoint related_bounded_value' {t} : abstract_domain t -> value' t -> type.interp base_interp t -> Prop + := match t return abstract_domain t -> value' t -> type.interp base_interp t -> Prop with | type.base t - => fun st '(e_st, e) v - => abstract_domain'_R t st e_st - /\ expr.interp ident_interp e = v + => fun st e v + => expr.interp ident_interp e = v /\ abstraction_relation' t st v | type.arrow s d => fun st e v => Proper type.eqv v - /\ forall st_s e_s v_s, - let st_s := match s with - | type.base _ => st_s - | type.arrow _ _ => bottom - end in - @related_bounded_value s st_s e_s v_s - -> @related_bounded_value d (st st_s) (UnderLets.interp ident_interp (e e_s)) (v v_s) + /\ Proper abstract_domain_R st + /\ forall st_s e_s v_s, + @related_bounded_value' s st_s e_s v_s + -> @related_bounded_value' d (st st_s) (UnderLets.interp ident_interp (e (st_s, e_s))) (v v_s) end. - Definition related_bounded_value_with_lets {t} : abstract_domain t -> value_with_lets t -> type.interp base_interp t -> Prop - := fun st e v => related_bounded_value st (UnderLets.interp ident_interp e) v. + Definition related_bounded_value {t} : value t -> type.interp base_interp t -> Prop + := fun st_e v => related_bounded_value' (state_of_value st_e) (UnderLets.interp ident_interp (project_value' st_e)) v. - Definition related_of_related_bounded_value {t} st e v - : @related_bounded_value t st e v -> v == v. + Definition related_of_related_bounded_value' {t} st e v + : @related_bounded_value' t st e v -> v == v. Proof using Type. destruct t; [ reflexivity | intros [? ?]; assumption ]. Qed. Lemma abstract_domain'_R_refl_of_rel_l t x y (H : @abstract_domain'_R t x y) @@ -199,43 +253,35 @@ Module Compilers. Local Hint Immediate abstract_domain_R_refl_of_rel_l abstract_domain_R_refl_of_rel_r : core. - Lemma related_bottom_for_each_lhs_of_arrow {t} v - : type.and_for_each_lhs_of_arrow (@abstraction_relation) (@bottom_for_each_lhs_of_arrow t) v. - Proof using bottom'_related. induction t; cbn; eauto. Qed. - - Local Hint Immediate related_bottom_for_each_lhs_of_arrow : core. - - Fixpoint fill_in_bottom_for_arrows {t} : abstract_domain t -> abstract_domain t - := match t with - | type.base t => fun x => x - | type.arrow s d - => fun f x => let x := match s with - | type.base _ => x - | type.arrow _ _ => bottom - end in - @fill_in_bottom_for_arrows d (f x) - end. - - Lemma abstract_domain_R_bottom_fill_arrows {t} - : abstract_domain_R (@bottom t) (fill_in_bottom_for_arrows (@bottom t)). - Proof using bottom'_Proper. - cbv [abstract_domain_R]; induction t as [t|s IHs d IHd]; cbn [fill_in_bottom_for_arrows bottom type.related]; - cbv [respectful Proper] in *; auto. + Definition abstract_domain_R_of_related_bounded_value' {t} st e v + : @related_bounded_value' t st e v -> abstract_domain_R st st. + Proof using abstract_domain'_R_of_related. + destruct t; [ cbn; intros [? ?]; subst; eapply abstract_domain'_R_of_related; eassumption | intros [? ?]; destruct_head'_and; assumption ]. Qed. - Lemma fill_in_bottom_for_arrows_bottom_related {t} v - : abstraction_relation (fill_in_bottom_for_arrows (@bottom t)) v. - Proof using bottom'_related. - cbv [abstraction_relation]; induction t; cbn; cbv [respectful_hetero]; eauto. + Lemma abstraction_relation_Proper_iff {t} : Proper (abstract_domain_R ==> type.eqv ==> iff) (@abstraction_relation t). + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstraction_relation'_Proper. + apply @type.related_hetero_and_Proper_Proper_iff; try exact _. + all: try now repeat intro; split; subst; eapply abstraction_relation'_Proper; (idtac + symmetry); auto; eassumption. + all: try now apply abstract_domain'_R_of_related. Qed. - Hint Resolve fill_in_bottom_for_arrows_bottom_related : core. + Lemma abstraction_relation_Proper_iff_abstract_domain_R {t} : Proper (abstract_domain_R ==> eq ==> iff) (@abstraction_relation t). + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstraction_relation'_Proper. + repeat intro; subst. + match goal with |- ?A <-> ?B => cut (A \/ B -> (A <-> B)); [ tauto | ] end. + intro H'. + apply abstraction_relation_Proper_iff; try assumption. + destruct_head'_or; eapply eqv_refl_of_abstraction_relation; eassumption. + Qed. - Local Instance fill_in_bottom_for_arrows_Proper {t} : Proper (abstract_domain_R ==> abstract_domain_R) (@fill_in_bottom_for_arrows t). - Proof using bottom'_Proper. - pose proof (@bottom_Proper). - cbv [Proper respectful abstract_domain_R] in *; induction t; cbn in *; cbv [respectful] in *; - intros; break_innermost_match; eauto. + Lemma abstraction_relation_Proper_iff_eqv {t} : Proper (eq ==> type.eqv ==> iff) (@abstraction_relation t). + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstraction_relation'_Proper. + repeat intro; subst. + match goal with |- ?A <-> ?B => cut (A \/ B -> (A <-> B)); [ tauto | ] end. + intro H'. + apply abstraction_relation_Proper_iff; try assumption. + destruct_head'_or; eapply abstract_domain_R_refl_of_abstraction_relation; eassumption. Qed. Local Instance bottom_eqv_Proper_refl {t} : Proper type.eqv (@bottom t). @@ -245,179 +291,133 @@ Module Compilers. Proof using Type. apply bottom_eqv_Proper_refl. Qed. Local Hint Resolve bottom_eqv_refl : core. - Local Instance fill_in_bottom_for_arrows_Proper_eqv {t} : Proper (type.eqv ==> type.eqv) (@fill_in_bottom_for_arrows t). - Proof using Type. - cbv [Proper respectful] in *; induction t; cbn in *; cbv [respectful] in *; - intros; break_innermost_match; cbn in *; cbv [respectful] in *; eauto. - Qed. - - Lemma state_of_value_related_fill {t} v (HP : Proper abstract_domain_R (@state_of_value t v)) - : abstract_domain_R (@state_of_value t v) (fill_in_bottom_for_arrows (@state_of_value t v)). - Proof using bottom'_Proper. destruct t; [ assumption | apply abstract_domain_R_bottom_fill_arrows ]. Qed. - - Lemma eqv_bottom_fill_bottom {t} - : @bottom t == fill_in_bottom_for_arrows bottom. - Proof using Type. induction t; cbn; [ reflexivity | ]; cbv [respectful]; auto. Qed. - - Lemma eqv_fill_bottom_idempotent {t} v1 v2 - : v1 == v2 -> fill_in_bottom_for_arrows (fill_in_bottom_for_arrows v1) == @fill_in_bottom_for_arrows t v2. - Proof using Type. induction t; cbn; cbv [respectful]; break_innermost_match; auto. Qed. - - Lemma abstract_domain_R_fill_bottom_idempotent {t} v1 v2 - : abstract_domain_R v1 v2 - -> abstract_domain_R (fill_in_bottom_for_arrows (fill_in_bottom_for_arrows v1)) - (@fill_in_bottom_for_arrows t v2). - Proof using bottom'_Proper. - pose proof (@bottom_Proper) as Hb. - induction t as [|s IHs d IHd]; cbn; cbv [respectful Proper abstract_domain_R] in *; break_innermost_match; auto. + Lemma related_refl_of_related_bounded_value' {t} st e v + : @related_bounded_value' t st e v -> abstract_domain_R st st. + Proof using abstract_domain'_R_of_related. + destruct t; cbn; break_innermost_match; intros; destruct_head'_and; try assumption. + eapply abstract_domain'_R_of_related; eassumption. Qed. - Lemma app_curried_state_of_value_fill {t} v x y - (H : type.and_for_each_lhs_of_arrow (@type.eqv) x y) - : type.app_curried (@state_of_value t v) x = type.app_curried (fill_in_bottom_for_arrows (@state_of_value t v)) y. - Proof using Type. - destruct t; [ reflexivity | cbv [state_of_value] ]. - apply type.app_curried_Proper; [ apply eqv_bottom_fill_bottom | assumption ]. - Qed. - - Lemma first_order_app_curried_fill_in_bottom_for_arrows_eq {t} f xs - (Ht : type.is_not_higher_order t = true) - : type.app_curried (t:=t) f xs = type.app_curried (fill_in_bottom_for_arrows f) xs. - Proof using Type. - clear -Ht. - induction t as [| [|s' d'] IHs d IHd]; cbn in *; try discriminate; auto. - Qed. - - Lemma first_order_abstraction_relation_fill_in_bottom_for_arrows_iff - {t} f v - (Ht : type.is_not_higher_order t = true) - : @abstraction_relation t f v - <-> @abstraction_relation t (fill_in_bottom_for_arrows f) v. - Proof using Type. - clear -Ht; cbv [abstraction_relation]. - split; induction t as [| [|s' d'] IHs d IHd]; - cbn in *; cbv [respectful_hetero]; try discriminate; auto. - Qed. - - Lemma related_state_of_value_of_related_bounded_value {t} st e v - : @related_bounded_value t st e v -> abstract_domain_R match t with - | type.base _ => st - | type.arrow _ _ => bottom - end (state_of_value e). - Proof using bottom'_Proper. intro H; destruct t; cbn in *; [ destruct e; apply H | repeat intro; refine bottom_Proper ]. Qed. - - Lemma related_state_of_value_of_related_bounded_value2 {t} st e v (st' := match t with - | type.base _ => st - | type.arrow _ _ => bottom - end) - : @related_bounded_value t st' e v -> abstract_domain_R st' (state_of_value e). - Proof using bottom'_Proper. intro H; destruct t; cbn in *; [ destruct e; apply H | repeat intro; refine bottom_Proper ]. Qed. - - Lemma related_bounded_value_Proper {t} st1 st2 (Hst : abstract_domain_R (fill_in_bottom_for_arrows st1) (fill_in_bottom_for_arrows st2)) + Lemma related_bounded_value'_Proper {t} st1 st2 (Hst : abstract_domain_R st1 st2) a a1 a2 (Ha' : type.eqv a1 a2) - : @related_bounded_value t st1 a a1 -> @related_bounded_value t st2 a a2. - Proof using abstraction_relation'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_Proper. - induction t as [t|s IHs d IHd]; cbn [related_bounded_value type.related] in *; cbv [respectful abstract_domain_R] in *. - all: cbn [type.andb_each_lhs_of_arrow] in *. - all: rewrite ?Bool.andb_true_iff in *. - all: destruct_head'_and. - { intros; break_innermost_match; subst; - destruct_head'_and; repeat apply conj; auto. - { etransitivity; (idtac + symmetry); eassumption. } - { eapply abstraction_relation'_Proper; (eassumption + reflexivity). } } + : @related_bounded_value' t st1 a a1 -> @related_bounded_value' t st2 a a2. + Proof using abstraction_relation'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric abstract_domain'_R_of_related. + induction t as [t|s IHs d IHd]; cbn [related_bounded_value' type.related] in *; cbv [respectful abstract_domain_R] in *. + all: subst. + { intros; destruct_head'_and; subst; repeat apply conj; try reflexivity. + eapply abstraction_relation'_Proper; (eassumption + reflexivity). } { intros [? Hrel]. - split; [ repeat intro; etransitivity; (idtac + symmetry); eapply Ha'; (eassumption + (etransitivity; (idtac + symmetry); eassumption)) | ]. - pose proof (@bottom_Proper s) as Hsbot. - intros ?? v_s; destruct s; intros Hx; cbn [type.related] in *; - cbn [fill_in_bottom_for_arrows] in *; cbv [respectful] in *. - { specialize_by_assumption; cbn in *. - eapply IHd; [ cbn in Hst |- *; eapply Hst | apply Ha'; reflexivity | eapply Hrel, Hx ]; cbv [respectful]. - cbn [related_bounded_value] in *. - break_innermost_match_hyps; destruct_head'_and. - eauto. } - { eapply IHd; [ eapply Hst | apply Ha' | eapply Hrel, Hx ]; - [ eexact Hsbot | refine (@related_of_related_bounded_value _ _ _ v_s _); eassumption | refine bottom ]. } } + split; [ | split ]; + [ repeat intro; etransitivity; (idtac + symmetry); (eapply Ha' + eapply Hst); (eassumption + (etransitivity; (idtac + symmetry); eassumption)) .. + | ]. + { intros ?? v_s Hx; cbn [type.related] in *; cbv [respectful] in *. + eapply IHd; [ cbn in Hst |- *; eapply Hst | apply Ha' | eapply Hrel, Hx ]. + all: try now eapply related_refl_of_related_bounded_value'; eassumption. + all: try now eapply related_of_related_bounded_value'; eassumption. } } Qed. - Lemma related_bounded_value_fill_bottom_iff {t} st1 st2 (Hst : abstract_domain_R st1 st2) - a a1 a2 - (Ha' : type.eqv a1 a2) - : @related_bounded_value t st1 a a1 <-> @related_bounded_value t (fill_in_bottom_for_arrows st2) a a2. - Proof using abstraction_relation'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_Proper. - split; eapply related_bounded_value_Proper; try solve [ (idtac + symmetry); assumption ]. - all: (idtac + symmetry); apply abstract_domain_R_fill_bottom_idempotent. - all: (idtac + symmetry); assumption. + Lemma related_bounded_value'_Proper1 {t} + : Proper (abstract_domain_R ==> eq ==> eq ==> Basics.impl) (@related_bounded_value' t). + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstraction_relation'_Proper. + repeat intro; subst; eapply related_bounded_value'_Proper; try eassumption. + eapply related_of_related_bounded_value'; eassumption. Qed. - Lemma related_bounded_value_Proper1 {t} - : Proper (abstract_domain_R ==> eq ==> eq ==> Basics.impl) (@related_bounded_value t). - Proof using abstraction_relation'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_Proper. - repeat intro; subst; eapply related_bounded_value_Proper. - { eapply fill_in_bottom_for_arrows_Proper; eassumption. } - { eapply related_of_related_bounded_value; eassumption. } - { assumption. } + Lemma related_bounded_value'_Proper3 {t} + : Proper (eq ==> eq ==> type.eqv ==> Basics.impl) (@related_bounded_value' t). + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstraction_relation'_Proper. + repeat intro; subst; eapply related_bounded_value'_Proper; try eassumption. + eapply abstract_domain_R_of_related_bounded_value'; eassumption. Qed. - Lemma related_bounded_value_Proper_eq {t} - : Proper (eq ==> eq ==> eq ==> Basics.impl) (@related_bounded_value t). + Lemma related_bounded_value'_Proper_eq {t} + : Proper (eq ==> eq ==> eq ==> Basics.impl) (@related_bounded_value' t). Proof using Type. repeat intro; subst; assumption. Qed. - Lemma related_bounded_value_Proper_interp_eq_base {t} - : Proper (eq ==> RelProd eq (fun x y => expr.interp ident_interp x = expr.interp ident_interp y) ==> eq ==> Basics.impl) (@related_bounded_value (type.base t)). + Lemma related_bounded_value'_Proper_interp_eq_base {t} + : Proper (eq ==> (fun x y => expr.interp ident_interp x = expr.interp ident_interp y) ==> eq ==> Basics.impl) (@related_bounded_value' (type.base t)). Proof using Type. repeat intro; subst. - cbv [value RelProd relation_conjunction predicate_intersection pointwise_extension RelCompFun] in *. - destruct_head'_prod; destruct_head'_and; cbn [fst snd] in *; subst. - cbv [related_bounded_value] in *; destruct_head'_and; repeat apply conj; subst; (idtac + symmetry); assumption. + cbv [related_bounded_value'] in *; destruct_head'_and; repeat apply conj; subst; (idtac + symmetry); assumption. + Qed. + + Lemma state_of_value_reflect is_let_bound {t} e st + : state_of_value (@reflect is_let_bound t e st) = st. + Proof using Type. destruct t; reflexivity. Qed. + + Lemma related_bounded_value'_of_abstraction_relation_of_reflect + {t} {annotate_with_state} + (interp_reflect + : forall st e v + (H_val : expr.interp ident_interp e == v) + (Hst1 : abstraction_relation st (expr.interp ident_interp e)), + related_bounded_value + (@reflect annotate_with_state t e st) + v) + st v + (H : abstraction_relation st v) + : @related_bounded_value' t st (UnderLets.interp ident_interp (project_value' (@reflect annotate_with_state t (expr.Var v) st))) v. + Proof using Type. + specialize (interp_reflect st (expr.Var v) v). + cbv [related_bounded_value] in interp_reflect. + rewrite state_of_value_reflect in interp_reflect. + apply interp_reflect; try assumption. + eapply eqv_refl_of_abstraction_relation; eassumption. Qed. Fixpoint interp_reify - annotate_with_state is_let_bound {t} st e v b_in - (Hb : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) b_in) - (H : related_bounded_value st e v) {struct t} - : (forall arg1 arg2 - (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) - (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1), - type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify annotate_with_state is_let_bound t e b_in))) arg1 - = type.app_curried v arg2) - /\ (forall arg1 - (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1) - (Harg11 : Proper (type.and_for_each_lhs_of_arrow (@type.eqv)) arg1), - abstraction_relation' - _ - (type.app_curried (fill_in_bottom_for_arrows st) b_in) - (type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify annotate_with_state is_let_bound t e b_in))) arg1)) + annotate_with_state is_let_bound {t} e v b_in + (H : related_bounded_value e v) + {struct t} + : forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1), + type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify annotate_with_state is_let_bound t e b_in))) arg1 + = type.app_curried v arg2 with interp_reflect annotate_with_state {t} st e v - (Hst_Proper : Proper abstract_domain_R st) + (*(Hst_Proper : Proper abstract_domain_R st)*) (* follows from abstraction relation *) (H_val : expr.interp ident_interp e == v) - (Hst1 : abstraction_relation (fill_in_bottom_for_arrows st) (expr.interp ident_interp e)) + (Hst1 : abstraction_relation st (expr.interp ident_interp e)) {struct t} : related_bounded_value - st (@reflect annotate_with_state t e st) - v. - Proof using interp_annotate abstraction_relation'_Proper bottom'_related bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric. + v + with abstraction_relation_of_related_bounded_value' + {t} st e v + {struct t} + : @related_bounded_value' t st e v + -> abstraction_relation st v. + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstraction_relation'_Proper bottom'_Proper bottom'_related interp_annotate. all: destruct t as [t|s d]; - [ clear interp_reify interp_reflect - | pose proof (fun is_let_bound => interp_reify annotate_with_state is_let_bound s) as interp_reify_s; - pose proof (fun is_let_bound => interp_reify annotate_with_state is_let_bound d) as interp_reify_d; - pose proof (interp_reflect annotate_with_state s) as interp_reflect_s; - pose proof (interp_reflect annotate_with_state d) as interp_reflect_d; - clear interp_reify interp_reflect; - pose proof (@abstract_domain_R_bottom_fill_arrows s); - pose proof (@abstract_domain_R_bottom_fill_arrows d) ]. + [ clear interp_reify interp_reflect abstraction_relation_of_related_bounded_value' + | pose proof (fun annotate_with_state is_let_bound => interp_reify annotate_with_state is_let_bound s) as interp_reify_s; + pose proof (fun annotate_with_state is_let_bound => interp_reify annotate_with_state is_let_bound d) as interp_reify_d; + pose proof (fun annotate_with_state => interp_reflect annotate_with_state s) as interp_reflect_s; + pose proof (fun annotate_with_state => interp_reflect annotate_with_state d) as interp_reflect_d; + pose proof (abstraction_relation_of_related_bounded_value' s) as abstraction_relation_of_related_bounded_value'_s; + pose proof (abstraction_relation_of_related_bounded_value' d) as abstraction_relation_of_related_bounded_value'_d; + pose proof (@related_bounded_value'_of_abstraction_relation_of_reflect s false (interp_reflect_s false)) as related_bounded_value'_of_abstraction_relation_s; + pose proof (@related_bounded_value'_of_abstraction_relation_of_reflect d false (interp_reflect_d false)) as related_bounded_value'_of_abstraction_relation_d; + try specialize (interp_reify_s annotate_with_state); + try specialize (interp_reify_d annotate_with_state); + try specialize (interp_reflect_s annotate_with_state); + try specialize (interp_reflect_d annotate_with_state); + clear interp_reify interp_reflect abstraction_relation_of_related_bounded_value' ]. + all: cbv [abstraction_relation] in *; cbn [type.related_hetero_and_Proper] in *; cbv [respectful_hetero] in *. all: cbn [reify reflect] in *; fold (@reify) (@reflect) in *. - all: cbn [related_bounded_value type.related type.app_curried] in *. - all: cbn [UnderLets.interp expr.interp type.final_codomain type.andb_each_lhs_of_arrow type.is_base fst snd fill_in_bottom_for_arrows type.map_for_each_lhs_of_arrow type.for_each_lhs_of_arrow type.and_for_each_lhs_of_arrow partial.bottom_for_each_lhs_of_arrow partial.bottom] in *. - all: destruct annotate_with_state; try destruct is_let_bound. + all: cbn [related_bounded_value' type.related type.app_curried] in *. + all: cbn [UnderLets.interp expr.interp type.final_codomain type.andb_each_lhs_of_arrow type.is_base fst snd type.map_for_each_lhs_of_arrow type.for_each_lhs_of_arrow type.and_for_each_lhs_of_arrow partial.bottom_for_each_lhs_of_arrow partial.bottom] in *. + all: try destruct annotate_with_state; try destruct is_let_bound. + all: change (type.related_hetero_and_Proper abstract_domain'_R (fun t : base_type => eq) abstraction_relation' (t:=?t)) with (@abstraction_relation t) in *. all: repeat first [ reflexivity + | assumption | progress eta_expand - | progress cbv [type.is_not_higher_order] in * | progress cbn [UnderLets.interp expr.interp type.final_codomain fst snd] in * + | progress cbn [state_of_value project_value' fst snd Base_value' UnderLets.interp] in * | progress subst | progress destruct_head'_and | progress destruct_head'_prod @@ -437,84 +437,102 @@ Module Compilers. | match goal with | [ H : fst ?x = _ |- _ ] => is_var x; destruct x | [ H : Proper _ ?st |- ?R (?st _) (?st _) ] => apply H - | [ |- ?R (state_of_value _) (state_of_value _) ] => cbv [state_of_value] in * + | [ |- type.related _ ?x ?x ] + => eapply related_refl_of_related_bounded_value'; eassumption + end + | match goal with + | [ H' : ?R ?i, H : forall a b, ?R a /\ @?A b -> @?B b |- _ ] + => specialize (fun b H'' => H i b (conj H' H'')) + | [ H' : ?R ?i, H : forall a b, @?A a /\ ?R b -> @?B a |- _ ] + => specialize (fun a H'' => H a i (conj H'' H')) end | solve [ repeat intro; apply bottom_Proper - | auto; cbv [Proper respectful Basics.impl] in *; eauto ] + | auto; cbv [Proper respectful Basics.impl] in *; eauto + | eapply related_refl_of_related_bounded_value'; eassumption + | eapply related_bottom_for_each_lhs_of_arrow, eqv_refl_of_abstraction_relation_for_each_lhs_of_arrow; eassumption + | eapply abstract_domain_R_refl_of_abstraction_relation; eassumption ] | progress (repeat apply conj; intros * ) | progress intros | progress destruct_head'_or - | do_with_exactly_one_hyp ltac:(fun H => eapply H; clear H); - try assumption; auto; [] | match goal with | [ |- Proper ?R _ ] => (eapply PER_valid_l + eapply PER_valid_r); eassumption - | [ |- @related_bounded_value ?t ?st1 (reflect _ _ ?st2) _ ] - => (tryif first [ constr_eq st1 st2 | has_evar st1 | has_evar st2 ] - then fail - else (eapply (@related_bounded_value_Proper1 t st2 st1); - try reflexivity)) | [ H : ?R ?x ?y |- ?R ?y ?x ] => symmetry; assumption end - | break_innermost_match_step - | do_with_exactly_one_hyp ltac:(fun H => eapply H; clear H); - try assumption; auto | match goal with - | [ |- abstraction_relation (fill_in_bottom_for_arrows (?f (state_of_value ?e))) _ ] - => replace (state_of_value e) with (match s with - | type.base _ => state_of_value e - | type.arrow _ _ => bottom - end) by (destruct s; reflexivity) + | [ H : related_bounded_value' ?st ?e ?v |- related_bounded_value' ?st ?e' ?v' ] + => is_evar e'; unify e e'; + eapply related_bounded_value'_Proper3; [ reflexivity .. | symmetry | exact H ] + | [ H' : abstraction_relation ?stv ?vv, H : forall annotate_with_state st e v, _ == v -> abstraction_relation st (expr.interp _ e) -> related_bounded_value (reflect annotate_with_state e st) v |- related_bounded_value' ?stv ?ee ?vv ] + => is_evar ee; + specialize (H false stv (expr.Var vv) vv); + cbv [related_bounded_value] in H; + rewrite state_of_value_reflect in H; eapply H; clear H end - | progress fold (@reify) (@reflect) (@type.interp) (@type.related) (@type.related_hetero) in * + | break_innermost_match_step + | progress fold (@reify) (@reflect) (@type.interp) (@type.related) (@type.related_hetero) (@related_bounded_value') (@abstraction_relation) in * | match goal with + | [ H : forall a b c d e (f : bool), _ |- _ ] + => unique pose proof (fun a b c d e => H a b c d e true); + unique pose proof (fun a b c d e => H a b c d e false); + clear H | [ |- type.related _ (expr.interp _ (UnderLets.interp _ (reify _ _ _ _))) _ ] => rewrite type.related_iff_app_curried - | [ |- type.related_hetero _ (@state_of_value ?t _) _ ] - => is_var t; destruct t; cbv [state_of_value]; [ cbn | apply bottom_related ] - end ]. + | [ interp_reflect_d : context[_ -> related_bounded_value (reflect _ _ _) _] |- related_bounded_value' ?st (UnderLets.interp _ (project_value' (reflect _ ?e ?st))) ?v ] + => specialize (interp_reflect_d st e v); + cbv [related_bounded_value] in interp_reflect_d; + rewrite state_of_value_reflect in interp_reflect_d + | [ H : context[_ -> type.app_curried (expr.interp _ (UnderLets.interp _ (reify ?b _ _ _))) _ = _] |- type.app_curried (expr.interp _ (UnderLets.interp _ (reify ?b ?is_let_bound ?e ?b_in))) ?arg = _ ] + => apply H; clear H + | [ H : context[_ -> abstraction_relation' _ (type.app_curried _ _) (type.app_curried (expr.interp _ (UnderLets.interp _ (reify ?b _ _ _))) _)] + |- abstraction_relation' _ (type.app_curried _ ?x) (type.app_curried (expr.interp _ (UnderLets.interp _ (reify ?b ?is_let_bound ?e ?b_in))) ?arg) ] + => specialize (H is_let_bound e); cbn [state_of_value Base_value' fst] in H; + eapply H; clear H + | [ |- related_bounded_value (_, Base _) _ ] => hnf; cbn [state_of_value project_value' Base_value' UnderLets.interp fst snd] + | [ H : forall st e v, _ -> related_bounded_value' _ (UnderLets.interp _ (?u _)) _ |- related_bounded_value' _ (UnderLets.interp _ (?u _)) _ ] + => apply H; clear H + | [ H : context[_ -> related_bounded_value' _ _ _ -> abstraction_relation _ _] |- abstraction_relation _ _ ] + => eapply H; clear H + end + | match goal with + | [ H : forall x : _ * _, _ |- _ ] => specialize (fun a b => H (a, b)) + end + | rewrite state_of_value_reflect + | rewrite type.related_hetero_iff_app_curried + | do_with_exactly_one_hyp ltac:(fun H => apply H; clear H) ]. Qed. - Lemma interp_reify_first_order - annotate_with_state is_let_bound {t} st e v b_in - (Ht : type.is_not_higher_order t = true) - (Hb : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) b_in) - (H : related_bounded_value st e v) - : (forall arg1 arg2 - (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) - (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1), - type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify annotate_with_state is_let_bound t e b_in))) arg1 - = type.app_curried v arg2) - /\ (forall arg1 - (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1) - (Harg11 : Proper (type.and_for_each_lhs_of_arrow (@type.eqv)) arg1), - abstraction_relation' - _ - (type.app_curried st b_in) - (type.app_curried (expr.interp ident_interp (UnderLets.interp ident_interp (@reify annotate_with_state is_let_bound t e b_in))) arg1)). - Proof using interp_annotate abstraction_relation'_Proper bottom'_related bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric. - rewrite first_order_app_curried_fill_in_bottom_for_arrows_eq by assumption. - apply interp_reify; assumption. + Lemma related_bounded_value'_of_abstraction_relation_annotate_with_state + {t} {annotate_with_state} st v + (H : abstraction_relation st v) + : @related_bounded_value' t st (UnderLets.interp ident_interp (project_value' (@reflect annotate_with_state t (expr.Var v) st))) v. + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstraction_relation'_Proper bottom'_Proper bottom'_related interp_annotate. + apply related_bounded_value'_of_abstraction_relation_of_reflect; try assumption. + apply @interp_reflect. Qed. - Lemma interp_reflect_first_order - annotate_with_state {t} st e v - (Ht : type.is_not_higher_order t = true) - (Hst_Proper : Proper abstract_domain_R st) - (H_val : expr.interp ident_interp e == v) - (Hst : abstraction_relation st (expr.interp ident_interp e)) - : related_bounded_value - st - (@reflect annotate_with_state t e st) - v. - Proof using interp_annotate abstraction_relation'_Proper bottom'_related bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric. - rewrite first_order_abstraction_relation_fill_in_bottom_for_arrows_iff in Hst by assumption. - apply interp_reflect; assumption. + Lemma related_bounded_value'_of_abstraction_relation + {t} st v + (H : abstraction_relation st v) + : @related_bounded_value' t st (UnderLets.interp ident_interp (project_value' (@reflect false t (expr.Var v) st))) v. + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstraction_relation'_Proper bottom'_Proper bottom'_related interp_annotate. + apply related_bounded_value'_of_abstraction_relation_annotate_with_state; assumption. + Qed. + + Lemma abstraction_relation_of_related_bounded_value'_app_curried + {t} st e v b_in arg + (H : @related_bounded_value' t st e v) + (Harg : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg) + : abstraction_relation' _ (type.app_curried st b_in) (type.app_curried v arg). + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstraction_relation'_Proper bottom'_Proper bottom'_related interp_annotate. + pose proof (@abstraction_relation_of_related_bounded_value' t st e v H) as H'. + rewrite abstraction_relation_iff_app_curried in H'. + apply H'; assumption. Qed. Lemma related_bounded_value_annotate_base {t} v_st st v - : @related_bounded_value (type.base t) v_st st v - -> @related_bounded_value (type.base t) v_st (fst st, UnderLets.interp ident_interp (annotate true t (fst st) (snd st))) v. + : @related_bounded_value' (type.base t) v_st st v + -> @related_bounded_value' (type.base t) v_st (UnderLets.interp ident_interp (annotate true t v_st st)) v. Proof using interp_annotate abstraction_relation'_Proper. clear -interp_annotate abstraction_relation'_Proper. cbv [Proper respectful Basics.impl] in *. @@ -523,136 +541,157 @@ Module Compilers. rewrite interp_annotate by eauto; reflexivity. Qed. - Lemma related_bounded_value_bottomify {t} v_st st v - : @related_bounded_value t v_st st v - -> @related_bounded_value t bottom (UnderLets.interp ident_interp (bottomify st)) v. - Proof using bottom'_Proper bottom'_related. - induction t; cbn in *; - repeat first [ progress subst - | progress cbv [respectful] in * - | progress cbn [UnderLets.interp] in * - | progress destruct_head'_and - | break_innermost_match_step - | progress intros - | apply conj - | reflexivity - | apply bottom'_Proper - | apply bottom'_related - | solve [ eauto ] - | rewrite UnderLets.interp_splice ]. + Lemma interp_extract_gen_from_wf_gen G + (HG : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> abstract_domain_R v1 v2) + {t} (e1 e2 : expr t) b_in1 b_in2 + (Hwf : expr.wf G e1 e2) + (Hb_in : type.and_for_each_lhs_of_arrow (@abstract_domain_R) b_in1 b_in2) + : abstract_domain'_R _ (extract_gen e1 b_in1) (extract_gen e2 b_in2). + Proof using abstract_interp_ident_Proper. + cbv [extract_gen]. + eapply type.related_iff_app_curried; try assumption. + unshelve eapply expr.wf_interp_Proper_gen1; try eassumption. + + Qed. + + Lemma interp_extract_gen_from_wf + {t} (e1 e2 : expr t) b_in1 b_in2 + (Hwf : expr.wf nil e1 e2) + (Hb_in : type.and_for_each_lhs_of_arrow (@abstract_domain_R) b_in1 b_in2) + : abstract_domain'_R _ (extract_gen e1 b_in1) (extract_gen e2 b_in2). + Proof using abstract_interp_ident_Proper. + eapply interp_extract_gen_from_wf_gen; try eassumption; wf_t. Qed. Context (interp_ident_Proper : forall annotate_with_state t idc, - related_bounded_value (ident_extract t idc) (UnderLets.interp ident_interp (interp_ident annotate_with_state t idc)) (ident_interp t idc)). + related_bounded_value (interp_ident annotate_with_state idc) (ident_interp t idc)). Lemma interp_interp - annotate_with_state G G' {t} (e_st e1 e2 e3 : expr t) - (HG : forall t v1 v2 v3, List.In (existT _ t (v1, v2, v3)) G - -> related_bounded_value_with_lets v1 v2 v3) - (HG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G' -> v1 == v2) - (Hwf : expr.wf3 G e_st e1 e2) - (Hwf' : expr.wf G' e2 e3) - : related_bounded_value_with_lets - (extract' e_st) - (interp annotate_with_state e1) - (expr.interp (@ident_interp) e2). - Proof using interp_ident_Proper interp_annotate abstraction_relation'_Proper ident_interp_Proper' abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_Proper bottom'_related. - clear -ident_interp_Proper' interp_ident_Proper interp_annotate abstraction_relation'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_Proper bottom'_related HG HG' Hwf Hwf'. - cbv [related_bounded_value_with_lets] in *; - revert dependent annotate_with_state; revert dependent G'; induction Hwf; intros; - cbn [extract' interp expr.interp UnderLets.interp List.In related_bounded_value reify reflect] in *; cbv [Let_In] in *. + annotate_with_state G G' G'' G''' {t} (e e_st e_v : expr t) + (HG : forall t v1 v2 v3, List.In (existT _ t (v1, v2, v3)) G + -> related_bounded_value v2 v3 + /\ abstract_domain_R v1 (state_of_value v2)) + (HG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G' -> abstract_domain_R (state_of_value v1) v2) + (HG'' : forall t v1 v2, List.In (existT _ t (v1, v2)) G'' -> v1 == v2) + (HG''' : forall t v1 v2, List.In (existT _ t (v1, v2)) G''' -> abstract_domain_R v1 v2) + (Hwf : expr.wf3 G e_st e e_v) + (Hwf' : expr.wf G' e e_st) + (Hwf'' : expr.wf G'' e_v e_v) + (Hwf''' : expr.wf G''' e_st e_st) + : related_bounded_value + (interp annotate_with_state e e_st) + (expr.interp (@ident_interp) e_v) + /\ abstract_domain_R (expr.interp (@abstract_interp_ident) e_st) (state_of_value (interp annotate_with_state e e_st)). + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstract_interp_ident_Proper abstraction_relation'_Proper ident_interp_Proper' interp_annotate interp_ident_Proper try_make_transport_base_type_cps_correct. + clear -abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstract_interp_ident_Proper abstraction_relation'_Proper ident_interp_Proper' interp_annotate interp_ident_Proper try_make_transport_base_type_cps_correct HG HG' HG'' HG''' Hwf Hwf' Hwf'' Hwf'''. + cbv [related_bounded_value] in *; + revert dependent annotate_with_state; revert dependent G'; revert dependent G''; revert dependent G'''; induction Hwf; intros. all: destruct annotate_with_state eqn:?. + all: cbn [interp expr.interp UnderLets.interp List.In related_bounded_value' reify reflect abstract_domain_R type.related] in *; cbv [Let_In Proper respectful] in *; cbn [state_of_value Base_value' fst snd] in *. + all: repeat apply conj; intros. all: repeat first [ progress intros | progress subst | progress inversion_sigma - | progress inversion_prod - | progress destruct_head'_and - | progress destruct_head'_or + | progress inversion_pair | progress destruct_head'_sig - | progress destruct_head'_sigT - | progress destruct_head'_prod + | progress destruct_head'_and | progress eta_expand - | exfalso; assumption - | progress cbn [UnderLets.interp List.In eq_rect fst snd projT1 projT2] in * + | progress cbv [interp_ident] in * + | progress cbn [eq_rect UnderLets.interp project_value' state_of_value fst snd Base_value' List.In apply_value] in * + | progress eliminate_hprop_eq + | apply related_bounded_value_annotate_base + | eapply abstract_domain_R_of_related_bounded_value'; eassumption + | eapply related_of_related_bounded_value'; eassumption + | solve [ cbv [abstract_domain_R] in *; (idtac + symmetry); eauto ] + | progress split_and + | erewrite invert_default_Some by first [ eassumption | reflexivity ] + | erewrite invert_default'_Some by first [ eassumption | reflexivity ] | rewrite UnderLets.interp_splice - | rewrite interp_annotate - | solve [ cbv [Proper respectful Basics.impl] in *; unshelve eauto using related_of_related_bounded_value, related_bounded_value_bottomify ] - | progress specialize_by_assumption - | progress cbv [Let_In] in * - | progress cbn [state_of_value extract'] in * - | progress expr.invert_subst - | match goal with - | [ |- abstract_domain ?t ] => exact (@bottom t) - | [ H : expr.wf _ _ _ |- Proper type.eqv _ ] - => apply expr.wf_interp_Proper_gen1 in H; - [ cbv [Proper]; etransitivity; (idtac + symmetry); exact H | auto ] - | [ H : _ |- _ ] - => (tryif first [ constr_eq H HG | constr_eq H HG' ] - then fail - else (apply H; clear H)) - | [ |- related_bounded_value _ (fst _, UnderLets.interp _ (annotate _ _ _ _)) _ ] - => apply related_bounded_value_annotate_base - | [ H : context[match ?v with None => _ | _ => _ end] |- _ ] => destruct v eqn:? - | [ H : context[@related_bounded_value (type.base ?t) ?x _ ?y] - |- @related_bounded_value (type.base ?t) ?x _ ?y ] - => eapply related_bounded_value_Proper_interp_eq_base; [ reflexivity | split; hnf | reflexivity | eapply H ]; - cbn [fst snd expr.interp]; - [ reflexivity | reflexivity | .. ] - end - | apply conj + | progress destruct_head'_or + | eapply expr.wf_interp_Proper_gen1; [ | now eauto ] | match goal with - | [ H : _ = _ |- _ ] => rewrite H + | [ |- _ /\ _ ] => split end - | break_innermost_match_step - | progress expr.inversion_wf_one_constr + | progress expr.inversion_wf_constr + | do_with_hyp' ltac:(fun H => now apply H; first [ eapply related_of_related_bounded_value' | eapply related_refl_of_related_bounded_value' ]; eauto) | match goal with - | [ H : _ |- _ ] - => (tryif first [ constr_eq H HG | constr_eq H HG' ] - then fail - else (eapply H; clear H; - [ lazymatch goal with - | [ |- expr.wf _ _ _ ] - => solve [ eassumption - | match goal with - | [ H : forall v1 v2, expr.wf _ _ _ |- expr.wf _ (?f ?x) _ ] - => apply (H x x) - end ] - | [ |- bool ] (* unused [annotate_with_state] argument to a [Proper ... /\ ...] lemma which is used for its [Proper] part *) - => constructor - | _ => idtac - end .. ]; - match goal with - [ |- ?G ] => assert_fails has_evar G - end)) + | [ H : context[related_bounded_value' _ (UnderLets.interp _ (project_value' (interp _ (?f2 _) (?f1 _)))) (expr.interp _ (?f3 _))] + |- related_bounded_value' _ (UnderLets.interp _ (project_value' (interp _ (?f2 _) (?f1 _)))) (expr.interp _ (?f3 _)) ] + => first [ eapply H + | eapply related_bounded_value'_Proper1; [ | reflexivity .. | eapply H ] ]; + lazymatch goal with + | [ |- expr.wf _ _ _ ] => solve [ eauto ] + | _ => idtac + end; + clear H + | [ H : context[related_bounded_value' _ (UnderLets.interp _ (UnderLets.interp _ (project_value' (interp _ ?f2 ?f1)) _)) (expr.interp _ ?f3 _)] + |- related_bounded_value' _ (UnderLets.interp _ (UnderLets.interp _ (project_value' (interp _ ?f2 ?f1)) _)) (expr.interp _ ?f3 _) ] + => first [ eapply H + | eapply related_bounded_value'_Proper1; [ | reflexivity .. | eapply H ] ]; + lazymatch goal with + | [ |- expr.wf _ _ _ ] => solve [ eauto ] + | _ => idtac + end; + clear H + | [ H : context[abstract_domain_R (expr.interp _ (?f1 _)) (state_of_value (interp _ (?f2 _) (?f1 _)))] + |- abstract_domain_R (expr.interp _ (?f1 _)) (state_of_value (interp _ (?f2 _) (?f1 _))) ] + => first [ eapply H; clear H + | etransitivity; [ | eapply H; clear H ] ]; + lazymatch goal with + | [ |- expr.wf _ _ _ ] => solve [ eauto ] + | _ => idtac + end + | [ H : context[type.related abstract_domain'_R (expr.interp _ ?f1 _) (state_of_value (interp _ ?f2 ?f1) _)] + |- abstract_domain_R (expr.interp _ ?f1 _) (state_of_value (interp _ ?f2 ?f1) _) ] + => eapply H; clear H; + lazymatch goal with + | [ |- expr.wf _ _ _ ] => solve [ eauto ] + | _ => idtac + end + | [ H : context[_ -> abstract_domain_R _ (state_of_value (interp _ ?f _))] + |- abstract_domain_R (state_of_value (interp _ ?f ?x)) (state_of_value (interp _ ?f ?x)) ] + => etransitivity; [ symmetry | ]; eapply H; clear H; + lazymatch goal with + | [ |- expr.wf _ _ _ ] => solve [ eauto ] + | _ => idtac + end + | [ |- abstract_domain_R (state_of_value _) (expr.interp _ _) ] => symmetry end - | reflexivity ]. + | progress repeat first [ progress eta_expand | break_innermost_match_step ] ]. Qed. Lemma interp_eval_with_bound' - annotate_with_state {t} (e_st e1 e2 : expr t) - (Hwf : expr.wf3 nil e_st e1 e2) - (Hwf' : expr.wf nil e2 e2) - (Ht : type.is_not_higher_order t = true) + annotate_with_state {t} (e e_st e_v : expr t) + (Hwf : expr.wf3 nil e_st e e_v) + (Hwf' : expr.wf nil e e_st) + (Hwf'' : expr.wf nil e_v e_v) + (Hwf''' : expr.wf nil e_st e_st) (st : type.for_each_lhs_of_arrow abstract_domain t) - (Hst : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) st) : (forall arg1 arg2 (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) st arg1), - type.app_curried (expr.interp ident_interp (eval_with_bound' annotate_with_state e1 st)) arg1 - = type.app_curried (expr.interp ident_interp e2) arg2) + type.app_curried (expr.interp ident_interp (eval_with_bound' annotate_with_state e e_st st)) arg1 + = type.app_curried (expr.interp ident_interp e_v) arg2) /\ (forall arg1 (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) st arg1) (Harg11 : Proper (type.and_for_each_lhs_of_arrow (@type.eqv)) arg1), abstraction_relation' _ (extract_gen e_st st) - (type.app_curried (expr.interp ident_interp (eval_with_bound' annotate_with_state e1 st)) arg1)). - Proof using interp_annotate abstraction_relation'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric bottom'_related interp_ident_Proper bottom'_Proper ident_interp_Proper'. - cbv [extract_gen extract' eval_with_bound']. - split; intros; rewrite UnderLets.interp_to_expr, UnderLets.interp_splice. - all: eapply interp_reify_first_order; eauto. - all: eapply interp_interp; eauto; wf_t. + (type.app_curried (expr.interp ident_interp (eval_with_bound' annotate_with_state e e_st st)) arg1)). + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstract_interp_ident_Proper abstraction_relation'_Proper bottom'_Proper bottom'_related ident_interp_Proper' interp_annotate interp_ident_Proper try_make_transport_base_type_cps_correct. + cbv [extract_gen eval_with_bound']. + split; intros; rewrite UnderLets.interp_to_expr. + { apply interp_reify; eauto; []. + eapply interp_interp; eauto; wf_t. } + { erewrite interp_reify; try eassumption. + { eapply abstraction_relation_of_related_bounded_value'_app_curried; try assumption; []. + eapply related_bounded_value'_Proper1; [ .. | eapply interp_interp ]; try reflexivity; try eassumption; eauto; wf_t; []. + symmetry; eapply interp_interp; try eassumption; wf_t. } + { now eapply interp_interp; eauto; wf_t. } } + Unshelve. + all: exact true. Qed. Lemma interp_eta_expand_with_bound' @@ -664,38 +703,12 @@ Module Compilers. (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1), type.app_curried (expr.interp ident_interp (eta_expand_with_bound' e1 b_in)) arg1 = type.app_curried (expr.interp ident_interp e2) arg2. - Proof using interp_annotate ident_interp_Proper' bottom'_related abstraction_relation'_Proper bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric. + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstraction_relation'_Proper bottom'_Proper bottom'_related ident_interp_Proper' interp_annotate. cbv [eta_expand_with_bound']. intros; rewrite UnderLets.interp_to_expr. eapply interp_reify; eauto. - eapply interp_reflect; eauto using bottom_related. - eapply @expr.wf_interp_Proper_gen; auto using Hwf. - Qed. - - Lemma interp_extract'_from_wf_gen G - (HG : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> abstract_domain_R v1 v2) - {t} (e1 e2 : expr t) - (Hwf : expr.wf G e1 e2) - : abstract_domain_R (extract' e1) (extract' e2). - Proof using ident_extract_Proper bottom'_Proper. - cbv [abstract_domain_R] in *; induction Hwf; cbn [extract']; break_innermost_match. - all: repeat first [ progress subst - | progress inversion_sigma - | progress inversion_prod - | solve [ cbv [Proper respectful] in *; eauto ] - | progress cbv [respectful Let_In] in * - | progress cbn [type.related List.In eq_rect partial.bottom] in * - | progress intros - | progress destruct_head'_or - | apply bottom_Proper - | match goal with H : _ |- type.related _ _ _ => apply H; clear H end ]. - Qed. - - Lemma interp_extract'_from_wf {t} (e1 e2 : expr t) - (Hwf : expr.wf nil e1 e2) - : abstract_domain_R (extract' e1) (extract' e2). - Proof using ident_extract_Proper bottom'_Proper. - eapply interp_extract'_from_wf_gen; revgoals; wf_t. + eapply interp_reflect; try (apply bottom_related; etransitivity; [ | symmetry ]). + all: eapply @expr.wf_interp_Proper_gen; [ | exact Hwf ]; auto. Qed. End with_type. @@ -714,19 +727,20 @@ Module Compilers. (extract_option_state : forall A, abstract_domain' (base.type.option A) -> option (option (abstract_domain' A))) (is_annotated_for : forall t t', @expr var t -> abstract_domain' t' -> bool) (skip_annotations_under : forall t, ident t -> bool) - (strip_annotation : forall t, ident t -> option (@value base.type ident var abstract_domain' t)) + (strip_annotation : forall t, ident t -> option (@value' base.type ident var abstract_domain' t)) (abstraction_relation' : forall t, abstract_domain' t -> base.interp t -> Prop) (abstract_domain'_R : forall t, abstract_domain' t -> abstract_domain' t -> Prop) (abstraction_relation'_Proper : forall t, Proper (abstract_domain'_R t ==> eq ==> Basics.impl) (abstraction_relation' t)) (bottom'_related : forall t v, abstraction_relation' t (bottom' t) v) {bottom'_Proper : forall t, Proper (abstract_domain'_R t) (bottom' t)} {abstract_domain'_R_transitive : forall t, Transitive (@abstract_domain'_R t)} - {abstract_domain'_R_symmetric : forall t, Symmetric (@abstract_domain'_R t)}. - Local Notation abstraction_relation := (@abstraction_relation base.type abstract_domain' base.interp abstraction_relation'). + {abstract_domain'_R_symmetric : forall t, Symmetric (@abstract_domain'_R t)} + (abstract_domain'_R_of_related : forall t st v, @abstraction_relation' t st v -> @abstract_domain'_R t st st). + Local Notation abstraction_relation := (@abstraction_relation base.type abstract_domain' base.interp abstraction_relation' abstract_domain'_R). Local Notation ident_interp := ident.interp (only parsing). Local Notation abstract_domain_R := (@abstract_domain_R base.type abstract_domain' abstract_domain'_R). - Local Notation fill_in_bottom_for_arrows := (@fill_in_bottom_for_arrows base.type abstract_domain' bottom'). - Local Notation related_bounded_value := (@related_bounded_value base.type ident abstract_domain' base.interp (@ident_interp) abstraction_relation' bottom' abstract_domain'_R). + Local Notation related_bounded_value' := (@related_bounded_value' base.type ident abstract_domain' base.interp (@ident_interp) abstraction_relation' abstract_domain'_R). + Local Notation related_bounded_value := (@related_bounded_value base.type ident abstract_domain' base.interp (@ident_interp) abstraction_relation' abstract_domain'_R). Context {abstract_interp_ident_Proper : forall t, Proper (eq ==> @abstract_domain_R t) (abstract_interp_ident t)} (interp_annotate_expr : forall t st idc, @@ -734,7 +748,7 @@ Module Compilers. -> forall v, abstraction_relation' _ st v -> expr.interp (@ident_interp) idc v = v) (abstract_interp_ident_Proper' - : forall t idc, type.related_hetero (@abstraction_relation') (abstract_interp_ident t idc) (ident_interp idc)) + : forall t idc, abstraction_relation (abstract_interp_ident t idc) (ident_interp idc)) (extract_list_state_related : forall t st ls v st' v', extract_list_state t st = Some ls @@ -754,7 +768,7 @@ Module Compilers. (strip_annotation_related : forall t idc v, strip_annotation t idc = Some v - -> related_bounded_value (abstract_interp_ident t idc) v (ident_interp idc)). + -> related_bounded_value' (abstract_interp_ident t idc) v (ident_interp idc)). Local Notation update_annotation := (@ident.update_annotation _ abstract_domain' annotate_expr is_annotated_for). Local Notation annotate_with_expr := (@ident.annotate_with_expr _ abstract_domain' annotate_expr is_annotated_for). @@ -763,18 +777,6 @@ Module Compilers. Local Notation reify := (@reify base.type ident _ abstract_domain' annotate bottom'). Local Notation reflect := (@reflect base.type ident _ abstract_domain' annotate bottom'). - Lemma abstract_interp_ident_Proper'' t idc - : type.related_hetero (@abstraction_relation') (fill_in_bottom_for_arrows (abstract_interp_ident t idc)) (ident_interp idc). - Proof using abstract_interp_ident_Proper' bottom'_related. - generalize (abstract_interp_ident_Proper' t idc); clear -bottom'_related. - generalize (ident_interp idc), (abstract_interp_ident t idc); clear idc. - intros v st H. - induction t as [| [|s' d'] IHs d IHd]; cbn in *; cbv [respectful_hetero] in *; - auto. - intros; apply IHd, H; clear IHd H. - intros; apply bottom_related; assumption. - Qed. - Lemma interp_update_annotation t st e (He : abstraction_relation' t st (expr.interp (t:=type.base t) (@ident_interp) e)) : expr.interp (@ident_interp) (@update_annotation t st e) @@ -827,7 +829,7 @@ Module Compilers. | progress break_innermost_match | progress break_innermost_match_hyps | progress expr.invert_subst - | progress cbn [fst snd UnderLets.interp expr.interp ident_interp Nat.add] in * + | progress cbn [fst snd UnderLets.interp expr.interp ident.interp Nat.add] in * | progress cbv [ident.literal] in * | match goal with | [ H : context[ident_interp (ident.ident_Literal _)] |- _ ] => rewrite ident.interp_ident_Literal in H @@ -846,9 +848,13 @@ Module Compilers. | match goal with | [ H : context[expr.interp _ (reify_list _)] |- _ ] => rewrite expr.interp_reify_list in H | [ H : abstraction_relation' (_ * _) _ (_, _) |- _ ] - => pose proof (abstract_interp_ident_Proper'' _ ident.fst _ _ H); - pose proof (abstract_interp_ident_Proper'' _ ident.snd _ _ H); - clear H + => let H' := fresh in + let H'' := fresh in + pose proof H as H'; + pose proof H as H''; + apply (abstract_interp_ident_Proper' _ ident.fst) in H'; + apply (abstract_interp_ident_Proper' _ ident.snd) in H''; + clear H | [ H : context[_ = _] |- _ = _ ] => rewrite H by assumption | [ |- List.map ?f (List.combine ?l1 ?l2) = List.map ?g ?l2 ] => transitivity (List.map g (List.map (@snd _ _) (List.combine l1 l2))); @@ -905,38 +911,43 @@ Module Compilers. end ]. Qed. - Local Notation interp_ident := (@ident.interp_ident _ abstract_domain' annotate_expr bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for strip_annotation). + Local Notation interp_ident' := (@ident.interp_ident' _ abstract_domain' annotate_expr bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for strip_annotation). Lemma interp_ident_Proper_not_nth_default_nostrip annotate_with_state t idc - : related_bounded_value (abstract_interp_ident t idc) (UnderLets.interp (@ident_interp) (Base (reflect annotate_with_state (expr.Ident idc) (abstract_interp_ident _ idc)))) (ident_interp idc). - Proof using abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr abstract_interp_ident_Proper bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric. - cbn [UnderLets.interp]. + : related_bounded_value (reflect annotate_with_state (expr.Ident idc) (abstract_interp_ident _ idc)) (@ident_interp t idc). + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr. eapply interp_reflect; try first [ apply ident.interp_Proper - | apply abstract_interp_ident_Proper'' | eapply abstract_interp_ident_Proper; reflexivity | apply interp_annotate ]; eauto. Qed. Lemma interp_ident_Proper_not_nth_default annotate_with_state t idc - : related_bounded_value + : related_bounded_value' (abstract_interp_ident t idc) - ((UnderLets.interp (@ident_interp)) - (Base match strip_annotation _ idc with - | Some v => v - | None => reflect annotate_with_state (expr.Ident idc) (abstract_interp_ident _ idc) - end)) + (UnderLets.interp + (@ident.interp) + match strip_annotation _ idc with + | Some v => Base v + | None => project_value' _ (reflect annotate_with_state (expr.Ident idc) (abstract_interp_ident _ idc)) + end) (ident_interp idc). - Proof using abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr abstract_interp_ident_Proper bottom'_Proper abstract_domain'_R_transitive abstract_domain'_R_symmetric strip_annotation_related. + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstract_interp_ident_Proper abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr strip_annotation_related. pose proof (strip_annotation_related t idc) as H. + cbv [related_bounded_value]. break_innermost_match; eauto using eq_refl, interp_ident_Proper_not_nth_default_nostrip. + eapply related_bounded_value'_Proper1; [ try exact eq_refl; try exact _ .. | eapply interp_reflect; eauto using interp_annotate ]. + all: eauto. + all: rewrite ?state_of_value_reflect. + all: try now apply abstract_interp_ident_Proper. + all: try now apply Compilers.eqv_Reflexive_Proper. Qed. Lemma interp_ident_Proper_nth_default annotate_with_state T (idc:=@ident.List_nth_default T) - : related_bounded_value (abstract_interp_ident _ idc) (UnderLets.interp (@ident_interp) (interp_ident annotate_with_state idc)) (ident_interp idc). - Proof using abstract_interp_ident_Proper abstract_interp_ident_Proper' abstraction_relation'_Proper extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr bottom'_related. - subst idc; cbn [interp_ident reify reflect fst snd UnderLets.interp ident_interp related_bounded_value abstract_domain value]. + : related_bounded_value' (abstract_interp_ident _ idc) (UnderLets.interp (@ident.interp) (interp_ident' annotate_with_state idc)) (ident_interp idc). + Proof using abstract_domain'_R_of_related abstract_interp_ident_Proper abstract_interp_ident_Proper' bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr. + subst idc; cbn [interp_ident reify reflect fst snd UnderLets.interp ident_interp related_bounded_value' abstract_domain value]. cbv [abstract_domain]; cbn [type.interp bottom_for_each_lhs_of_arrow state_of_value fst snd]. repeat first [ progress intros | progress cbn [UnderLets.interp fst snd expr.interp ident_interp] in * @@ -952,7 +963,7 @@ Module Compilers. | solve [ cbv [Proper respectful Basics.impl] in *; eauto ] | split; [ apply (@abstract_interp_ident_Proper _ (@ident.List_nth_default T) _ eq_refl) | ] | split; [ reflexivity | ] - | apply (@abstract_interp_ident_Proper'' _ (@ident.List_nth_default T)) + | apply (@abstract_interp_ident_Proper' _ (@ident.List_nth_default T)) | apply conj | rewrite map_nth_default_always | rewrite expr.interp_reify_list @@ -964,53 +975,58 @@ Module Compilers. | [ H : _ = reify_list _ |- _ ] => apply (f_equal (expr.interp (@ident_interp))) in H | [ H : expr.interp _ ?x = _ |- context[expr.interp _ ?x] ] => rewrite H | [ |- Proper _ _ ] => cbv [Proper type.related respectful] - end ]. + end + | progress cbn [interp_ident' reify snd fst Base_value' reflect project_value']; eta_expand ]. Qed. Lemma interp_ident_Proper annotate_with_state t idc - : related_bounded_value (abstract_interp_ident t idc) (UnderLets.interp (@ident_interp) (interp_ident annotate_with_state idc)) (ident_interp idc). - Proof using abstract_domain'_R_symmetric abstract_domain'_R_transitive abstract_interp_ident_Proper abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr strip_annotation_related. + : related_bounded_value' (abstract_interp_ident _ idc) (UnderLets.interp (@ident.interp) (@interp_ident' t annotate_with_state idc)) (ident_interp idc). + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstract_interp_ident_Proper abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr strip_annotation_related. pose idc as idc'. - destruct idc; first [ refine (@interp_ident_Proper_not_nth_default _ _ idc') - | refine (@interp_ident_Proper_nth_default _ _) ]. + destruct idc; + first [ refine (@interp_ident_Proper_not_nth_default _ _ idc') + | refine (@interp_ident_Proper_nth_default _ _) ]. Qed. Local Notation eval_with_bound := (@partial.ident.eval_with_bound _ abstract_domain' annotate_expr bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for strip_annotation). Local Notation eta_expand_with_bound := (@partial.ident.eta_expand_with_bound _ abstract_domain' annotate_expr bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for). - Local Notation extract := (@ident.extract abstract_domain' bottom' abstract_interp_ident). + Local Notation extract := (@ident.extract abstract_domain' abstract_interp_ident). Lemma interp_eval_with_bound - annotate_with_state {t} (e_st e1 e2 : expr t) - (Hwf : expr.wf3 nil e_st e1 e2) - (Hwf' : expr.wf nil e2 e2) - (Ht : type.is_not_higher_order t = true) - (st : type.for_each_lhs_of_arrow abstract_domain t) - (Hst : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) st) + annotate_with_state {t} (e e_st e_v : expr t) + (Hwf : expr.wf3 nil e_st e e_v) + (Hwf' : expr.wf nil e e_st) + (Hwf'' : expr.wf nil e_v e_v) + (Hwf''' : expr.wf nil e_st e_st) + (st : type.for_each_lhs_of_arrow abstract_domain t) : (forall arg1 arg2 (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) st arg1), - type.app_curried (expr.interp (@ident_interp) (eval_with_bound skip_annotations_under annotate_with_state e1 st)) arg1 - = type.app_curried (expr.interp (@ident_interp) e2) arg2) + type.app_curried (expr.interp (@ident_interp) (eval_with_bound skip_annotations_under annotate_with_state e e_st st)) arg1 + = type.app_curried (expr.interp (@ident_interp) e_v) arg2) /\ (forall arg1 (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) st arg1) (Harg11 : Proper (type.and_for_each_lhs_of_arrow (@type.eqv)) arg1), abstraction_relation' _ (extract e_st st) - (type.app_curried (expr.interp (@ident_interp) (eval_with_bound skip_annotations_under annotate_with_state e1 st)) arg1)). - Proof using abstract_domain'_R_symmetric abstract_domain'_R_transitive abstract_interp_ident_Proper abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr strip_annotation_related. cbv [extract eval_with_bound]; apply @interp_eval_with_bound' with (abstract_domain'_R:=abstract_domain'_R); auto using interp_annotate, interp_ident_Proper, ident.interp_Proper. Qed. + (type.app_curried (expr.interp (@ident_interp) (eval_with_bound skip_annotations_under annotate_with_state e e_st st)) arg1)). + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstract_interp_ident_Proper abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr strip_annotation_related. + cbv [extract eval_with_bound]; eapply @interp_eval_with_bound' with (abstract_domain'_R:=abstract_domain'_R); auto using interp_annotate, interp_ident_Proper, ident.interp_Proper; try exact _; intros; apply interp_ident_Proper. + Qed. Lemma interp_eta_expand_with_bound {t} (e1 e2 : expr t) (Hwf : expr.wf nil e1 e2) - (Ht : type.is_not_higher_order t = true) (b_in : type.for_each_lhs_of_arrow abstract_domain t) (Hb_in : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) b_in) : forall arg1 arg2 (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) (Harg1 : type.and_for_each_lhs_of_arrow (@abstraction_relation) b_in arg1), type.app_curried (expr.interp (@ident_interp) (eta_expand_with_bound e1 b_in)) arg1 = type.app_curried (expr.interp (@ident_interp) e2) arg2. - Proof using abstract_domain'_R_symmetric abstract_domain'_R_transitive abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr. cbv [partial.ident.eta_expand_with_bound]; eapply interp_eta_expand_with_bound'; eauto using interp_annotate, ident.interp_Proper. Qed. + Proof using abstract_domain'_R_of_related abstract_domain'_R_symmetric abstract_domain'_R_transitive abstract_interp_ident_Proper' abstraction_relation'_Proper bottom'_Proper bottom'_related extract_list_state_length_good extract_list_state_related extract_option_state_related interp_annotate_expr. + cbv [partial.ident.eta_expand_with_bound]; eapply interp_eta_expand_with_bound'; eauto using (@interp_ident), interp_annotate, ident.interp_Proper; try typeclasses eauto. + Qed. End with_type. End ident. @@ -1028,11 +1044,12 @@ Module Compilers. Local Notation abstract_domain := (@partial.abstract_domain base.type abstract_domain'). Local Notation abstract_domain'_R t := (@eq (abstract_domain' t)) (only parsing). Local Notation abstract_domain_R := (@abstract_domain_R base.type abstract_domain' (fun t => abstract_domain'_R t)). - Local Notation fill_in_bottom_for_arrows := (@fill_in_bottom_for_arrows base.type abstract_domain' bottom'). Definition abstraction_relation' {t} : abstract_domain' t -> base.interp t -> Prop := fun st v => @ZRange.type.base.option.is_bounded_by t st v = true. + Local Notation abstraction_relation := (@abstraction_relation _ abstract_domain' _ (@abstraction_relation') (fun t => abstract_domain'_R t)). + Lemma bottom'_bottom {t} : forall v, abstraction_relation' (bottom' t) v. Proof using Type. cbv [abstraction_relation' bottom']; induction t; cbn; intros; break_innermost_match; cbn; try reflexivity. @@ -1040,8 +1057,8 @@ Module Compilers. Qed. Lemma abstract_interp_ident_related {opts : AbstractInterpretation.Options} {assume_cast_truncates : bool} {t} (idc : ident t) - : type.related_hetero (@abstraction_relation') (abstract_interp_ident assume_cast_truncates t idc) (ident.interp idc). - Proof using Type. apply ZRange.ident.option.interp_related. Qed. + : abstraction_relation (abstract_interp_ident assume_cast_truncates t idc) (ident.interp idc). + Proof using Type. apply ZRange.ident.option.interp_related_and_Proper. Qed. Local Ltac zrange_interp_idempotent_t := repeat first [ progress destruct_head'_ex @@ -1082,22 +1099,24 @@ Module Compilers. : ZRange.type.base.option.is_bounded_by (t:=base.type.Z*base.type.Z) (ZRange.ident.option.interp assume_cast_truncates ident.Z_cast2 r1 r2) v = true. Proof using Type. destruct r1, r2, assume_cast_truncates; zrange_interp_idempotent_t. Qed. - Local Notation related_bounded_value := (@related_bounded_value base.type ident abstract_domain' base.interp (@ident.interp) (@abstraction_relation') bottom' (fun t => abstract_domain'_R t)). + Local Notation related_bounded_value' := (@related_bounded_value' base.type ident abstract_domain' base.interp (@ident.interp) (@abstraction_relation') (fun t => abstract_domain'_R t)). Lemma always_strip_annotation_related {opts : AbstractInterpretation.Options} {assume_cast_truncates : bool} {t} (idc : ident t) v (Hv : always_strip_annotation assume_cast_truncates t idc = Some v) - : related_bounded_value (abstract_interp_ident assume_cast_truncates t idc) v (ident.interp idc). + : related_bounded_value' (abstract_interp_ident assume_cast_truncates t idc) v (ident.interp idc). Proof using Type. pose proof (@abstract_interp_ident_related _ assume_cast_truncates _ ident.Z_cast). - pose proof (fun x1 x2 y1 y2 H x01 x02 y01 y02 H' => @abstract_interp_ident_related _ assume_cast_truncates _ ident.Z_cast2 (x1, x2) (y1, y2) H (x01, x02) (y01, y02) H'). + pose proof (fun x1 x2 y1 y2 H x01 x02 y01 y02 H' => proj2 (proj2 (proj2 (proj2 (@abstract_interp_ident_related _ assume_cast_truncates _ ident.Z_cast2)) (x1, x2) (y1, y2) H)) (x01, x02) (y01, y02) H'). + cbn [abstraction_relation] in *. cbv [always_strip_annotation] in *; break_innermost_match_hyps; inversion_option; subst. all: repeat first [ reflexivity - | progress cbn [related_bounded_value UnderLets.interp fst snd type.related_hetero] in * + | progress cbn [related_bounded_value' UnderLets.interp fst snd type.related_hetero] in * | progress cbv [respectful_hetero] in * | progress intros | progress destruct_head'_and | progress subst | progress inversion_prod | progress inversion_option + | progress destruct_head_hnf' prod | match goal with | [ |- Proper _ _ ] => exact _ | [ |- _ /\ _ ] => split @@ -1117,7 +1136,7 @@ Module Compilers. | progress break_innermost_match_hyps | progress Reflect.reflect_beq_to_eq zrange_beq | progress cbv [abstraction_relation' ZRange.type.base.option.lift_Some] in * - | cbn [Compilers.ident_interp]; progress cbv [ident.cast2] + | cbn [ident.interp]; progress cbv [ident.cast2] | rewrite ident.cast_in_bounds by assumption | now destruct assume_cast_truncates | rewrite zrange_interp_idempotent_Z_cast by (cbn; eexists; eassumption) @@ -1141,7 +1160,7 @@ Module Compilers. Lemma strip_annotation_related {opts : AbstractInterpretation.Options} {assume_cast_truncates strip_annotations : bool} {t} (idc : ident t) v (Hv : strip_annotation assume_cast_truncates strip_annotations t idc = Some v) - : related_bounded_value (abstract_interp_ident assume_cast_truncates t idc) v (ident.interp idc). + : related_bounded_value' (abstract_interp_ident assume_cast_truncates t idc) v (ident.interp idc). Proof using Type. destruct strip_annotations; cbv [strip_annotation] in *; try congruence; now apply always_strip_annotation_related. @@ -1179,8 +1198,8 @@ Module Compilers. cbv [partial.Extract partial.ident.extract partial.extract_gen]. revert b_in1 b_in2 Hb. rewrite <- (@type.related_iff_app_curried base.type ZRange.type.base.option.interp (fun _ => eq)). - apply interp_extract'_from_wf; auto with wf typeclass_instances. - apply GeneralizeVar.wf_from_flat_to_flat, Hwf. + eapply expr.wf_interp_Proper_gen1; [ | apply GeneralizeVar.wf_from_flat_to_flat, Hwf ]. + wf_t. Qed. Lemma Extract_FromFlat_ToFlat {opts : AbstractInterpretation.Options} {assume_cast_truncates : bool} {t} (e : Expr t) (Hwf : Wf e) b_in @@ -1244,35 +1263,35 @@ Module Compilers. {assume_cast_truncates : bool} {skip_annotations_under : forall t, ident t -> bool} {strip_preexisting_annotations : bool} - {t} (e_st e1 e2 : expr t) - (Hwf : expr.wf3 nil e_st e1 e2) - (Hwf' : expr.wf nil e2 e2) - (Ht : type.is_not_higher_order t = true) + {t} (e_st e e_v : expr t) + (Hwf : expr.wf3 nil e_st e e_v) + (Hwf' : expr.wf nil e e_st) + (Hwf'' : expr.wf nil e_v e_v) + (Hwf''' : expr.wf nil e_st e_st) (st : type.for_each_lhs_of_arrow abstract_domain t) : (forall arg1 arg2 (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) st arg1 = true), - type.app_curried (expr.interp (@ident.interp) (eval_with_bound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e1 st)) arg1 - = type.app_curried (expr.interp (@ident.interp) e2) arg2) + type.app_curried (expr.interp (@ident.interp) (eval_with_bound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e e_st st)) arg1 + = type.app_curried (expr.interp (@ident.interp) e_v) arg2) /\ (forall arg1 (Harg11 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg1) (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) st arg1 = true), abstraction_relation' (extract assume_cast_truncates e_st st) - (type.app_curried (expr.interp (@ident.interp) (eval_with_bound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e1 st)) arg1)). + (type.app_curried (expr.interp (@ident.interp) (eval_with_bound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e e_st st)) arg1)). Proof using Hrelax. cbv [eval_with_bound]; split; [ intros arg1 arg2 Harg12 Harg1 | intros arg1 Harg11 Harg1 ]. - all: eapply Compilers.type.andb_bool_impl_and_for_each_lhs_of_arrow in Harg1; [ | apply ZRange.type.option.is_bounded_by_impl_related_hetero ]. + all: eapply Compilers.type.andb_bool_impl_and_for_each_lhs_of_arrow in Harg1; [ | apply @ZRange.type.option.is_bounded_by_impl_related_hetero_and_Proper with (skip_base:=true) ]. all: eapply ident.interp_eval_with_bound with (abstraction_relation':=@abstraction_relation') (abstract_domain'_R:=fun t => abstract_domain'_R t); eauto using bottom'_bottom, interp_annotate_expr with typeclass_instances. - all: intros; (eapply extract_list_state_related + eapply extract_option_state_related + eapply strip_annotation_related); eassumption. + all: try (intros; (eapply extract_list_state_related + eapply extract_option_state_related + eapply strip_annotation_related); eassumption). Qed. Lemma interp_eta_expand_with_bound {t} (e1 e2 : expr t) (Hwf : expr.wf nil e1 e2) - (Ht : type.is_not_higher_order t = true) (b_in : type.for_each_lhs_of_arrow abstract_domain t) : forall arg1 arg2 (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) @@ -1282,8 +1301,9 @@ Module Compilers. cbv [partial.eta_expand_with_bound]; intros arg1 arg2 Harg12 Harg1. eapply Compilers.type.andb_bool_impl_and_for_each_lhs_of_arrow in Harg1. { apply ident.interp_eta_expand_with_bound with (abstraction_relation':=@abstraction_relation') (abstract_domain'_R:=fun t => abstract_domain'_R t); eauto using bottom'_bottom, interp_annotate_expr with typeclass_instances. - all: intros; (eapply extract_list_state_related + eapply extract_option_state_related); eassumption. } - { apply ZRange.type.option.is_bounded_by_impl_related_hetero. } + all: try (intros; (eapply extract_list_state_related + eapply extract_option_state_related); eassumption). + all: try (cbv [Proper]; eapply abstract_domain_R_refl_of_abstraction_relation_for_each_lhs_of_arrow; try eassumption; reflexivity). } + { apply ZRange.type.option.is_bounded_by_impl_related_hetero_and_Proper. } Qed. Lemma Interp_EvalWithBound @@ -1293,9 +1313,7 @@ Module Compilers. {strip_preexisting_annotations : bool} {t} (e : Expr t) (Hwf : expr.Wf e) - (Ht : type.is_not_higher_order t = true) (st : type.for_each_lhs_of_arrow abstract_domain t) - (Hst : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) st) : (forall arg1 arg2 (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) st arg1 = true), @@ -1307,12 +1325,37 @@ Module Compilers. abstraction_relation' (Extract assume_cast_truncates e st) (type.app_curried (expr.Interp (@ident.interp) (EvalWithBound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e st)) arg1)). - Proof using Hrelax. cbv [Extract EvalWithBound]; apply interp_eval_with_bound; try apply expr.Wf3_of_Wf; auto. Qed. + Proof using Hrelax. + cbv [EvalWithBound Let_In]. + epose proof (interp_eval_with_bound + (GeneralizeVar.FromFlat (GeneralizeVar.ToFlat e) _) + (GeneralizeVar.FromFlat (GeneralizeVar.ToFlat e) _) + (GeneralizeVar.FromFlat (GeneralizeVar.ToFlat e) _)) as H. + edestruct H as [H0 H1]; [ shelve .. | ]. + split; intros. + { cbv [expr.Interp]; erewrite H0 by eassumption. + eapply (@type.related_iff_app_curried _ _ (fun _ => eq) t). + { apply GeneralizeVar.Interp_gen1_FromFlat_ToFlat; eauto. } + { etransitivity; (idtac + symmetry); eassumption. } } + { lazymatch goal with + | [ H : context[abstraction_relation' ?e _] |- abstraction_relation' ?e' _ ] + => replace e' with e; [ now eapply H | ] + end. + apply Extract_FromFlat_ToFlat'; eauto; []. + eapply type.and_for_each_lhs_of_arrow_Proper_impl_hetero1; + [ | eapply type.andb_bool_impl_and_for_each_lhs_of_arrow; [ | eassumption ] ]; + [ | refine (fun t x y H => H) ]; + cbv beta. + intros *; eapply ZRange.type.option.is_bounded_by_impl_eqv_refl. } + Unshelve. + all: try apply expr.Wf3_of_Wf. + all: apply GeneralizeVar.Wf_FromFlat_to_flat. + all: apply Hwf. + Qed. Lemma Interp_EtaExpandWithBound {t} (E : Expr t) (Hwf : Wf E) - (Ht : type.is_not_higher_order t = true) (b_in : type.for_each_lhs_of_arrow abstract_domain t) : forall arg1 arg2 (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) @@ -1367,29 +1410,30 @@ Module Compilers. Lemma interp_strip_annotations {opts : AbstractInterpretation.Options} {assume_cast_truncates : bool} - {t} (e_st e1 e2 : expr t) - (Hwf : expr.wf3 nil e_st e1 e2) - (Hwf' : expr.wf nil e2 e2) - (Ht : type.is_not_higher_order t = true) + {t} (e_st e e_v : expr t) + (Hwf : expr.wf3 nil e_st e e_v) + (Hwf' : expr.wf nil e e_st) + (Hwf'' : expr.wf nil e_v e_v) + (Hwf''' : expr.wf nil e_st e_st) (st : type.for_each_lhs_of_arrow abstract_domain t) : (forall arg1 arg2 (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) st arg1 = true), - type.app_curried (expr.interp (@ident.interp) (strip_annotations assume_cast_truncates e1 st)) arg1 - = type.app_curried (expr.interp (@ident.interp) e2) arg2) + type.app_curried (expr.interp (@ident.interp) (strip_annotations assume_cast_truncates e e_st st)) arg1 + = type.app_curried (expr.interp (@ident.interp) e_v) arg2) /\ (forall arg1 (Harg11 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg1) (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) st arg1 = true), abstraction_relation' (extract assume_cast_truncates e_st st) - (type.app_curried (expr.interp (@ident.interp) (strip_annotations assume_cast_truncates e1 st)) arg1)). + (type.app_curried (expr.interp (@ident.interp) (strip_annotations assume_cast_truncates e e_st st)) arg1)). Proof using Type. cbv [strip_annotations]; split; [ intros arg1 arg2 Harg12 Harg1 | intros arg1 Harg11 Harg1 ]. - all: eapply Compilers.type.andb_bool_impl_and_for_each_lhs_of_arrow in Harg1; [ | apply ZRange.type.option.is_bounded_by_impl_related_hetero ]. + all: eapply Compilers.type.andb_bool_impl_and_for_each_lhs_of_arrow in Harg1; [ | apply @ZRange.type.option.is_bounded_by_impl_related_hetero_and_Proper with (skip_base:=true) ]. all: eapply ident.interp_eval_with_bound with (abstraction_relation':=@abstraction_relation') (abstract_domain'_R:=fun t => abstract_domain'_R t); eauto using bottom'_bottom, interp_annotate_expr, default_relax_zrange_good, abstract_interp_ident_related with typeclass_instances. - all: intros; (eapply extract_list_state_related + eapply extract_option_state_related + eapply strip_annotation_related); eassumption. + all: try (intros; (eapply extract_list_state_related + eapply extract_option_state_related + eapply strip_annotation_related); eassumption). Qed. Lemma Interp_StripAnnotations @@ -1397,7 +1441,6 @@ Module Compilers. {assume_cast_truncates : bool} {t} (E : Expr t) (Hwf : Wf E) - (Ht : type.is_not_higher_order t = true) (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) : forall arg1 arg2 (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) @@ -1413,7 +1456,6 @@ Module Compilers. {assume_cast_truncates : bool} {t} (E : Expr t) (Hwf : expr.Wf E) - (Ht : type.is_not_higher_order t = true) (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) : forall arg1 (Harg11 : Proper (type.and_for_each_lhs_of_arrow (@type.eqv)) arg1) @@ -1429,7 +1471,6 @@ Module Compilers. Lemma Interp_EtaExpandWithListInfoFromBound {t} (E : Expr t) (Hwf : Wf E) - (Ht : type.is_not_higher_order t = true) (b_in : type.for_each_lhs_of_arrow abstract_domain t) : forall arg1 arg2 (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) @@ -1453,7 +1494,6 @@ Module Compilers. -> is_tighter_than_bool z r' = true) {t} (E : Expr t) (Hwf : Wf E) - (Ht : type.is_not_higher_order t = true) (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) : forall arg1 arg2 (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) @@ -1467,7 +1507,6 @@ Module Compilers. by (hnf; etransitivity; [ eassumption | symmetry; eassumption ]). assert (arg2_Proper : Proper (type.and_for_each_lhs_of_arrow (@type.related base.type base.interp (fun _ => eq))) arg2) by (hnf; etransitivity; [ symmetry; eassumption | eassumption ]). - rewrite <- (GeneralizeVar.Interp_gen1_GeneralizeVar E) by auto with wf. eapply Interp_EvalWithBound; eauto with wf typeclass_instances. Qed. @@ -1479,7 +1518,6 @@ Module Compilers. -> is_tighter_than_bool z r' = true) {t} (E : Expr t) (Hwf : Wf E) - (Ht : type.is_not_higher_order t = true) (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) : forall arg1 (Harg11 : Proper (type.and_for_each_lhs_of_arrow (@type.eqv)) arg1) @@ -1491,7 +1529,6 @@ Module Compilers. Proof. cbv [PartialEvaluateWithBounds]. intros arg1 Harg11 Harg1. - rewrite <- Extract_GeneralizeVar by auto with wf typeclass_instances. eapply @Interp_EvalWithBound; eauto with wf typeclass_instances. Qed. @@ -1499,7 +1536,6 @@ Module Compilers. {opts : AbstractInterpretation.Options} {t} (E : Expr t) (Hwf : Wf E) - (Ht : type.is_not_higher_order t = true) (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) : forall arg1 arg2 (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) @@ -1526,10 +1562,10 @@ Module Compilers. -> is_tighter_than_bool z r' = true) {t} (E : Expr t) (Hwf : Wf E) - (Ht : type.is_not_higher_order t = true) (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) (b_out : ZRange.type.base.option.interp (type.final_codomain t)) rv (Hrv : CheckedPartialEvaluateWithBounds relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations E b_in b_out = inl rv) + {b_in_Proper : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R base.type ZRange.type.base.option.interp (fun t0 : base.type => eq))) b_in} : (forall arg1 arg2 (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), @@ -1543,17 +1579,49 @@ Module Compilers. [ intros arg1 arg2 Harg12 Harg1; assert (arg1_Proper : Proper (type.and_for_each_lhs_of_arrow (@type.related base.type base.interp (fun _ => eq))) arg1) by (hnf; etransitivity; [ eassumption | symmetry; eassumption ]); + (*assert (b_in_Proper : Proper (type.and_for_each_lhs_of_arrow (@type.related _ _ (fun _ => eq))) b_in) + by (cbv [Proper] in *; + eapply type.and_for_each_lhs_of_arrow_Proper_impl_hetero1; + [ | eapply type.andb_bool_impl_and_for_each_lhs_of_arrow; [ | eassumption ] ]; + [ | refine (fun t x y H => H) ]; + cbv beta; + intros *; eapply ZRange.type.option.is_bounded_by_impl_eqv_refl);*) split; - repeat first [ rewrite !@GeneralizeVar.Interp_gen1_FromFlat_ToFlat by eauto with wf typeclass_instances - | rewrite <- Extract_FromFlat_ToFlat by auto with typeclass_instances; apply Interp_PartialEvaluateWithBounds_bounded; auto - | rewrite Extract_FromFlat_ToFlat by auto with wf typeclass_instances - | progress intros - | eapply ZRange.type.base.option.is_tighter_than_is_bounded_by; [ eassumption | ] - | solve [ eauto with wf typeclass_instances ] - | erewrite !Interp_PartialEvaluateWithBounds - | apply type.app_curried_Proper - | apply expr.Wf_Interp_Proper_gen - | progress intros ] - | eauto with wf typeclass_instances ]. + repeat first [ rewrite !@GeneralizeVar.Interp_gen1_FromFlat_ToFlat by eauto with wf typeclass_instances + | rewrite <- Extract_FromFlat_ToFlat by eauto with typeclass_instances; apply Interp_PartialEvaluateWithBounds_bounded; auto + | rewrite Extract_FromFlat_ToFlat by eauto with wf typeclass_instances + | progress intros + | eapply ZRange.type.base.option.is_tighter_than_is_bounded_by; [ eassumption | ] + | solve [ eauto with wf typeclass_instances ] + | erewrite !Interp_PartialEvaluateWithBounds + | apply type.app_curried_Proper + | apply expr.Wf_Interp_Proper_gen + | progress intros ] + | eauto with wf typeclass_instances ]. + Qed. + + Theorem CheckedPartialEvaluateWithBounds_Correct_first_order + {opts : AbstractInterpretation.Options} + (relax_zrange : zrange -> option zrange) + (assume_cast_truncates : bool) + (skip_annotations_under : forall t, ident t -> bool) + (strip_preexisting_annotations : bool) + (Hrelax : forall r r' z, is_tighter_than_bool z r = true + -> relax_zrange r = Some r' + -> is_tighter_than_bool z r' = true) + {t} (E : Expr t) + (Hwf : Wf E) + (Ht : type.is_not_higher_order t = true) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + (b_out : ZRange.type.base.option.interp (type.final_codomain t)) + rv (Hrv : CheckedPartialEvaluateWithBounds relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations E b_in b_out = inl rv) + : (forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), + ZRange.type.base.option.is_bounded_by b_out (type.app_curried (Interp rv) arg1) = true + /\ type.app_curried (Interp rv) arg1 = type.app_curried (Interp E) arg2) + /\ Wf rv. + Proof. + eapply CheckedPartialEvaluateWithBounds_Correct; eauto with core typeclass_instances. Qed. End Compilers. diff --git a/src/AbstractInterpretation/Fancy/Wf.v b/src/AbstractInterpretation/Fancy/Wf.v index ce0ac95634..8b7cd0d7b8 100644 --- a/src/AbstractInterpretation/Fancy/Wf.v +++ b/src/AbstractInterpretation/Fancy/Wf.v @@ -15,11 +15,13 @@ Require Import Crypto.Util.NatUtil. Require Import Crypto.Util.Bool.Reflect. Require Import Crypto.Util.ZUtil.Tactics.LtbToLt. Require Import Crypto.Util.Tactics.BreakMatch. +Require Import Crypto.Util.Tactics.DoWithHyp. Require Import Crypto.Util.Tactics.DestructHead. Require Import Crypto.Util.Tactics.SplitInContext. Require Import Crypto.Util.Tactics.UniquePose. Require Import Crypto.Util.Tactics.SpecializeBy. Require Import Crypto.Util.Tactics.SpecializeAllWays. +Require Import Crypto.Util.HProp. Require Import Rewriter.Language.Language. Require Import Rewriter.Language.Inversion. Require Import Crypto.Language.InversionExtra. @@ -35,7 +37,7 @@ Module Compilers. Import Language.Compilers. Import UnderLets.Compilers. Import AbstractInterpretation.Compilers. - Import AbstractInterpretation.ZRangeCommonProofs.Compilers. + Import ZRangeCommonProofs.Compilers. Import Language.Inversion.Compilers. Import Language.InversionExtra.Compilers. Import Language.Wf.Compilers. @@ -58,25 +60,33 @@ Module Compilers. (bottom' : forall A, abstract_domain' A) (abstract_interp_ident : forall t, ident t -> type.interp abstract_domain' t) (abstract_domain'_R : forall t, abstract_domain' t -> abstract_domain' t -> Prop) - {abstract_interp_ident_Proper : forall t, Proper (eq ==> abstract_domain'_R t) (abstract_interp_ident t)} - {bottom'_Proper : forall t, Proper (abstract_domain'_R t) (bottom' t)}. - Local Notation value var := (@value base_type ident var abstract_domain'). + {bottom'_Proper : forall t, Proper (abstract_domain'_R t) (bottom' t)} + (base_type_beq : base_type -> base_type -> bool) + {reflect_base_type_beq : reflect_rel eq base_type_beq} + {try_make_transport_base_type_cps : @type.try_make_transport_cpsT base_type} + {try_make_transport_base_type_cps_correct : type.try_make_transport_cps_correctT base_type}. + + Local Notation value' var := (@value' base_type ident var abstract_domain'). Local Notation abstract_domain := (@abstract_domain base_type abstract_domain'). Local Notation bottom := (@bottom base_type abstract_domain' (@bottom')). Local Notation bottom_for_each_lhs_of_arrow := (@bottom_for_each_lhs_of_arrow base_type abstract_domain' (@bottom')). + Context {default_expr : @DefaultValue.type.base.DefaultT _ (@expr abstract_domain)}. (* needed for impossible cases *) + + (*Context {base_type_eq_Decidable : DecidableRel (@eq base_type)}.*) + Section with_var2. Context {var1 var2 : type -> Type}. Local Notation UnderLets1 := (@UnderLets.UnderLets base_type ident var1). Local Notation UnderLets2 := (@UnderLets.UnderLets base_type ident var2). Local Notation expr1 := (@expr.expr base_type ident var1). Local Notation expr2 := (@expr.expr base_type ident var2). - Local Notation value1 := (@value var1). - Local Notation value2 := (@value var2). - Local Notation value_with_lets1 := (@value_with_lets base_type ident var1 abstract_domain'). - Local Notation value_with_lets2 := (@value_with_lets base_type ident var2 abstract_domain'). - Local Notation state_of_value1 := (@state_of_value base_type ident var1 abstract_domain' bottom'). - Local Notation state_of_value2 := (@state_of_value base_type ident var2 abstract_domain' bottom'). + Local Notation value'1 := (@value' var1). + Local Notation value'2 := (@value' var2). + Local Notation value1 := (@value base_type ident var1 abstract_domain'). + Local Notation value2 := (@value base_type ident var2 abstract_domain'). + Local Notation state_of_value1 := (@state_of_value base_type ident var1 abstract_domain'). + Local Notation state_of_value2 := (@state_of_value base_type ident var2 abstract_domain'). Context (annotate1 : forall (is_let_bound : bool) t, abstract_domain' t -> @expr1 t -> UnderLets1 (@expr1 t)) (annotate2 : forall (is_let_bound : bool) t, abstract_domain' t -> @expr2 t -> UnderLets2 (@expr2 t)) (annotate_Proper @@ -84,57 +94,58 @@ Module Compilers. v1 v2 (Hv : abstract_domain'_R t v1 v2) e1 e2 (He : expr.wf G e1 e2), UnderLets.wf (fun G' => expr.wf G') G (annotate1 is_let_bound t v1 e1) (annotate2 is_let_bound t v2 e2)) - (interp_ident1 : bool -> forall t, ident t -> value_with_lets1 t) - (interp_ident2 : bool -> forall t, ident t -> value_with_lets2 t) + (interp_ident'1 : bool -> forall t, ident t -> UnderLets1 (value'1 t)) + (interp_ident'2 : bool -> forall t, ident t -> UnderLets2 (value'2 t)) (skip_annotations_under : forall t, ident t -> bool). Local Notation reify1 := (@reify base_type ident var1 abstract_domain' annotate1 bottom'). Local Notation reify2 := (@reify base_type ident var2 abstract_domain' annotate2 bottom'). Local Notation reflect1 := (@reflect base_type ident var1 abstract_domain' annotate1 bottom'). Local Notation reflect2 := (@reflect base_type ident var2 abstract_domain' annotate2 bottom'). - Local Notation bottomify1 := (@bottomify base_type ident var1 abstract_domain' bottom'). - Local Notation bottomify2 := (@bottomify base_type ident var2 abstract_domain' bottom'). - Local Notation interp1 := (@interp base_type ident var1 abstract_domain' annotate1 bottom' skip_annotations_under interp_ident1). - Local Notation interp2 := (@interp base_type ident var2 abstract_domain' annotate2 bottom' skip_annotations_under interp_ident2). - Local Notation eval_with_bound'1 := (@eval_with_bound' base_type ident var1 abstract_domain' annotate1 bottom' skip_annotations_under interp_ident1). - Local Notation eval_with_bound'2 := (@eval_with_bound' base_type ident var2 abstract_domain' annotate2 bottom' skip_annotations_under interp_ident2). - Local Notation eval'1 := (@eval' base_type ident var1 abstract_domain' annotate1 bottom' skip_annotations_under interp_ident1). - Local Notation eval'2 := (@eval' base_type ident var2 abstract_domain' annotate2 bottom' skip_annotations_under interp_ident2). + Local Notation interp1 := (@interp base_type ident var1 try_make_transport_base_type_cps abstract_domain' annotate1 bottom' skip_annotations_under default_expr abstract_interp_ident interp_ident'1). + Local Notation interp2 := (@interp base_type ident var2 try_make_transport_base_type_cps abstract_domain' annotate2 bottom' skip_annotations_under default_expr abstract_interp_ident interp_ident'2). + Local Notation eval_with_bound'1 := (@eval_with_bound' base_type ident var1 try_make_transport_base_type_cps abstract_domain' annotate1 bottom' skip_annotations_under default_expr abstract_interp_ident interp_ident'1). + Local Notation eval_with_bound'2 := (@eval_with_bound' base_type ident var2 try_make_transport_base_type_cps abstract_domain' annotate2 bottom' skip_annotations_under default_expr abstract_interp_ident interp_ident'2). + Local Notation eval'1 := (@eval' base_type ident var1 try_make_transport_base_type_cps abstract_domain' annotate1 bottom' skip_annotations_under default_expr abstract_interp_ident interp_ident'1). + Local Notation eval'2 := (@eval' base_type ident var2 try_make_transport_base_type_cps abstract_domain' annotate2 bottom' skip_annotations_under default_expr abstract_interp_ident interp_ident'2). Local Notation eta_expand_with_bound'1 := (@eta_expand_with_bound' base_type ident var1 abstract_domain' annotate1 bottom'). Local Notation eta_expand_with_bound'2 := (@eta_expand_with_bound' base_type ident var2 abstract_domain' annotate2 bottom'). + Local Notation invert_default := (@invert_default base_type ident try_make_transport_base_type_cps abstract_domain' default_expr). + Local Notation invert_default' := (@invert_default' base_type ident try_make_transport_base_type_cps abstract_domain' default_expr). Definition abstract_domain_R {t} : relation (abstract_domain t) := type.related abstract_domain'_R. + Context {abstract_interp_ident_Proper : forall t, Proper (eq ==> @abstract_domain_R t) (abstract_interp_ident t)}. + (** This one is tricky. Because we need to be stable under weakening and reordering of the context, we permit any context for well-formedness of the input in the arrow case, and simply tack on that context at the beginning of the output. This is sort-of wasteful on the output context, but it's sufficient to prove - [wf_value_Proper_list] below, which is what we really + [wf_value'_Proper_list] below, which is what we really need. *) - Fixpoint wf_value G {t} : value1 t -> value2 t -> Prop - := match t return value1 t -> value2 t -> Prop with - | type.base t - => fun v1 v2 - => abstract_domain_R (fst v1) (fst v2) - /\ expr.wf G (snd v1) (snd v2) + Fixpoint wf_value' G {t} : value'1 t -> value'2 t -> Prop + := match t return value'1 t -> value'2 t -> Prop with + | type.base t => expr.wf G | type.arrow s d => fun v1 v2 => forall seg G' sv1 sv2, G' = (seg ++ G)%list - -> @wf_value seg s sv1 sv2 + -> abstract_domain_R (fst sv1) (fst sv2) + /\ @wf_value' seg s (snd sv1) (snd sv2) -> UnderLets.wf - (fun G' => @wf_value G' d) G' + (fun G' => @wf_value' G' d) G' (v1 sv1) (v2 sv2) end. - Definition wf_value_with_lets G {t} : value_with_lets1 t -> value_with_lets2 t -> Prop - := UnderLets.wf (fun G' => wf_value G') G. + Definition wf_value G {t} : value1 t -> value2 t -> Prop + := fun v1 v2 => abstract_domain_R (fst v1) (fst v2) + /\ UnderLets.wf (fun G' => wf_value' G') G (snd v1) (snd v2). Context (interp_ident_Proper : forall G t idc1 idc2 (Hidc : idc1 = idc2) annotate_with_state, - wf_value_with_lets G (interp_ident1 annotate_with_state t idc1) (interp_ident2 annotate_with_state t idc2)). + UnderLets.wf (fun G' => wf_value' G') G (interp_ident'1 annotate_with_state t idc1) (interp_ident'2 annotate_with_state t idc2)). Global Instance bottom_Proper {t} : Proper abstract_domain_R (@bottom t) | 10. Proof using bottom'_Proper. @@ -152,20 +163,17 @@ Module Compilers. Lemma state_of_value_Proper G {t} v1 v2 (Hv : @wf_value G t v1 v2) : abstract_domain_R (state_of_value1 v1) (state_of_value2 v2). - Proof using bottom'_Proper. - clear -Hv type_base bottom'_Proper. - destruct t; [ destruct v1, v2, Hv | ]; cbn in *; cbv [respectful]; eauto; intros; apply bottom_Proper. - Qed. + Proof using Type. apply Hv. Qed. Local Hint Resolve (fun A (P : list A -> Prop) => ex_intro P nil) (fun A (x : A) (P : list A -> Prop) => ex_intro P (cons x nil)) : core. Local Hint Constructors expr.wf ex : core. Local Hint Unfold List.In : core. - Lemma wf_value_Proper_list G1 G2 + Lemma wf_value'_Proper_list G1 G2 (HG1G2 : forall t v1 v2, List.In (existT _ t (v1, v2)) G1 -> List.In (existT _ t (v1, v2)) G2) t e1 e2 - (Hwf : @wf_value G1 t e1 e2) - : @wf_value G2 t e1 e2. + (Hwf : @wf_value' G1 t e1 e2) + : @wf_value' G2 t e1 e2. Proof using Type. clear -type_base HG1G2 Hwf. revert dependent G1; revert dependent G2; induction t; intros; @@ -182,6 +190,49 @@ Module Compilers. | eapply UnderLets.wf_Proper_list; [ | | solve [ eauto ] ] ]. Qed. + Lemma wf_value_Proper_list G1 G2 + (HG1G2 : forall t v1 v2, List.In (existT _ t (v1, v2)) G1 -> List.In (existT _ t (v1, v2)) G2) + t e1 e2 + (Hwf : @wf_value G1 t e1 e2) + : @wf_value G2 t e1 e2. + Proof using Type. + split; try apply Hwf. + eapply UnderLets.wf_Proper_list, Hwf; eauto using wf_value'_Proper_list. + Qed. + + Local Ltac wf_reify_reflect_step wf_reflect_s wf_reflect_d wf_reify_s wf_reify_d + := first [ progress cbn [fst snd] in * + | progress fold (@wf_value') in * + | match goal with + | [ |- wf_value _ (reflect1 _ _ _) (reflect2 _ _ _) ] => apply wf_reflect_s || apply wf_reflect_d + | [ |- UnderLets.wf _ _ (UnderLets.splice (reify1 _ _ _ _) _) (UnderLets.splice (reify2 _ _ _ _) _) ] + => eapply UnderLets.wf_splice; [ apply wf_reify_s || apply wf_reify_d | ] + | [ |- UnderLets.wf _ _ (reify1 _ _ _ _) (reify2 _ _ _ _) ] + => apply wf_reify_s || apply wf_reify_d + | [ |- abstract_domain_R (state_of_value1 ?v1 ?x1) (state_of_value2 ?v2 ?x2) ] + => eapply (state_of_value_Proper _ v1 v2); [ | change (abstract_domain_R x1 x2) ] + end + | progress cbn [wf_value'] in * + | progress intros + | progress subst + | progress destruct_head'_and + | progress destruct_head'_ex + | progress cbv [respectful wf_value] in * + | match goal with + | [ |- ?x ++ ?y::?z ++ ?w = ?seg ++ ?z ++ ?w ] => unify seg (x ++ y::nil); cbn [List.app]; rewrite <- !app_assoc + end + | solve [ eauto | wf_t ] + | apply annotate_Proper + | apply UnderLets.wf_to_expr + | break_innermost_match_step + | eapply wf_value'_Proper_list; [ | eassumption ] + | match goal with + | [ |- UnderLets.wf _ _ _ _ ] => constructor + | [ |- expr.wf _ _ _ ] => constructor + | [ |- _ /\ _ ] => split + | [ |- abstract_domain_R (state_of_value1 _) (state_of_value2 _) ] + => eapply state_of_value_Proper + end ]. Fixpoint wf_reify (annotate_with_state : bool) (is_let_bound : bool) G {t} : forall v1 v2 (Hv : @wf_value G t v1 v2) s1 s2 (Hs : type.and_for_each_lhs_of_arrow (@abstract_domain_R) s1 s2), @@ -204,53 +255,49 @@ Module Compilers. all: cbn [reify reflect] in *. all: fold (@reify2) (@reflect2) (@reify1) (@reflect1). all: cbn in *. - all: repeat first [ progress cbn [fst snd] in * - | progress cbv [respectful] in * - | progress intros - | progress subst - | progress destruct_head'_and - | progress destruct_head'_ex - | solve [ eauto | wf_t ] - | apply annotate_Proper - | apply UnderLets.wf_to_expr - | break_innermost_match_step - | match goal with - | [ |- UnderLets.wf _ _ _ _ ] => constructor - | [ |- expr.wf _ _ _ ] => constructor - | [ He : forall seg G' sv1 sv2, G' = (seg ++ ?G)%list -> _ |- UnderLets.wf _ (?v :: ?G) (UnderLets.splice _ _) (UnderLets.splice _ _) ] - => eapply UnderLets.wf_splice; [ apply (He (cons v nil)) | ] - | [ |- UnderLets.wf _ _ (UnderLets.splice (reify1 _ _ _ _) _) (UnderLets.splice (reify2 _ _ _ _) _) ] - => eapply UnderLets.wf_splice; [ apply wf_reify_s || apply wf_reify_d | ] - | [ |- wf_value _ (reflect1 _ _ _) (reflect2 _ _ _) ] => apply wf_reflect_s || apply wf_reflect_d - | [ H : wf_value _ ?x ?y |- wf_value _ ?x ?y ] - => eapply wf_value_Proper_list; [ | eassumption ] - | [ H : forall x y, ?R x y -> ?R' (?f x) (?g y) |- ?R' (?f _) (?g _) ] - => apply H - | [ |- ?R (state_of_value1 _) (state_of_value2 _) ] => eapply state_of_value_Proper - end ]. + all: repeat wf_reify_reflect_step wf_reflect_s wf_reflect_d wf_reify_s wf_reify_d. + all: try (eapply UnderLets.wf_splice; [ eassumption + apply wf_reify_s + apply wf_reify_d + apply wf_reflect_s + apply wf_reflect_d | ]). + all: repeat wf_reify_reflect_step wf_reflect_s wf_reflect_d wf_reify_s wf_reify_d. + all: try (eapply UnderLets.wf_splice; [ eassumption + apply wf_reify_s + apply wf_reify_d + apply wf_reflect_s + apply wf_reflect_d | ]). + all: repeat wf_reify_reflect_step wf_reflect_s wf_reflect_d wf_reify_s wf_reify_d. + all: try (eapply UnderLets.wf_splice; [ eassumption + apply wf_reify_s + apply wf_reify_d + apply wf_reflect_s + apply wf_reflect_d | ]). + all: repeat wf_reify_reflect_step wf_reflect_s wf_reflect_d wf_reify_s wf_reify_d. + all: try (apply wf_reify_s + apply wf_reify_d + apply wf_reflect_s + apply wf_reflect_d). + all: repeat wf_reify_reflect_step wf_reflect_s wf_reflect_d wf_reify_s wf_reify_d. + all: [ > ]. + eapply UnderLets.wf_splice. + 1: match goal with + | [ H : wf_value' ?G ?x ?y |- UnderLets.wf ?P ?G (?f (_, ?x)) (?g (_, ?y)) ] + => eapply UnderLets.wf_Proper_list_impl with (P1:=P) (P2:=P); revgoals + end; + [ exactly_once multimatch goal with H : _ |- _ => eapply H end | .. ]. + all: repeat wf_reify_reflect_step wf_reflect_s wf_reflect_d wf_reify_s wf_reify_d. + all: instantiate (1:=_::nil); left; reflexivity. Qed. - Lemma wf_bottomify {t} G v1 v2 - (Hwf : @wf_value G t v1 v2) - : wf_value_with_lets G (bottomify1 v1) (bottomify2 v2). - Proof using bottom'_Proper. - cbv [wf_value_with_lets] in *. - revert dependent G; induction t as [|s IHs d IHd]; intros; - cbn [bottomify wf_value]; fold (@value1) (@value2) in *; break_innermost_match; - constructor. - all: repeat first [ progress cbn [fst snd wf_value] in * - | progress destruct_head'_and - | assumption - | apply bottom'_Proper - | apply conj - | progress intros - | progress subst - | solve [ eapply UnderLets.wf_splice; eauto ] ]. + #[local] Existing Instance type.reflect_type_beq. + + Lemma invert_default_Some {A B o t v} + : o = Some (existT _ t v) -> @invert_default A B t o = v. + Proof using try_make_transport_base_type_cps_correct. + intro; subst; cbv [invert_default Option.sequence_return projT1 projT2]. + rewrite_type_transport_correct. + break_innermost_match; try now exfalso; reflect_hyps; congruence. + rewrite Reflect.eq_reflect_to_dec_true_refl; reflexivity. + Qed. + + Lemma invert_default'_Some {A B C o t v} + : o = Some (existT _ t v) -> @invert_default' A B C t o = v. + Proof using try_make_transport_base_type_cps_correct. + intro; subst; cbv [invert_default' Option.sequence_return projT1 projT2]. + rewrite_type_transport_correct. + break_innermost_match; try now exfalso; reflect_hyps; congruence. + rewrite Reflect.eq_reflect_to_dec_true_refl; reflexivity. Qed. Local Ltac wf_interp_t := - repeat first [ progress cbv [wf_value_with_lets abstract_domain_R respectful] in * - | progress cbn [wf_value fst snd partial.bottom type.related eq_rect List.In] in * + repeat first [ progress cbv [wf_value abstract_domain_R respectful Base_value' project_value' Option.sequence_return] in * + | progress cbn [wf_value' fst snd partial.bottom type.related eq_rect List.In] in * | wf_safe_t_step | exact I | apply wf_reify @@ -258,17 +305,21 @@ Module Compilers. | progress destruct_head'_ex | progress destruct_head'_or | eapply UnderLets.wf_splice + | erewrite invert_default_Some by reflexivity + | erewrite invert_default'_Some by reflexivity | match goal with - | [ |- UnderLets.wf _ _ (bottomify1 _) (bottomify2 _) ] => apply wf_bottomify | [ |- UnderLets.wf _ _ _ _ ] => constructor | [ |- and _ _ ] => apply conj + | [ H : _ = snd ?x |- _ ] => is_var x; destruct x + | [ H : _ = fst ?x |- _ ] => is_var x; destruct x end - | eapply wf_value_Proper_list; [ | solve [ eauto ] ] - | eapply UnderLets.wf_Proper_list; [ | | solve [ eauto ] ] + | eapply wf_value'_Proper_list; [ | solve [ eauto | split_and; eauto ] ] + | eapply UnderLets.wf_Proper_list; [ | | solve [ eauto | split_and; eauto ] ] | match goal with | [ H : _ |- _ ] => eapply H; clear H; solve [ wf_interp_t ] end - | break_innermost_match_step ]. + | break_innermost_match_step + | progress eliminate_hprop_eq ]. Local Notation skip_annotations_for_App := (@skip_annotations_for_App base_type ident skip_annotations_under). @@ -299,12 +350,39 @@ Module Compilers. end ]. Qed. - Lemma wf_interp (annotate_with_state : bool) G G' {t} (e1 : @expr (@value_with_lets1) t) (e2 : @expr (@value_with_lets2) t) - (Hwf : expr.wf G e1 e2) - (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) - : wf_value_with_lets G' (interp1 annotate_with_state e1) (interp2 annotate_with_state e2). - Proof using annotate_Proper bottom'_Proper interp_ident_Proper. - revert dependent G'; revert annotate_with_state; induction Hwf; intros; cbn [interp]; + Lemma wf_interp (annotate_with_state : bool) (G : list { t : _ & value1 t * value2 t * abstract_domain t * abstract_domain t }%type) G12 G34 G' {t} (e1 : @expr (@value1) t) (est1 : @expr abstract_domain t) (e2 : @expr (@value2) t) (est2 : @expr abstract_domain t) + (HG12 : List.map (fun '(existT t (v1, v2, _, _)) => existT _ t (v1, v2)) G = G12) + (HG34 : List.map (fun '(existT t (_, _, v3, v4)) => existT _ t (v3, v4)) G = G34) + (Hwf : expr.wf4 G e1 e2 est1 est2) + (Hwf' : expr.wf G34 est1 est2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G12 -> wf_value G' v1 v2) + (HGG'' : forall t v1 v2, List.In (existT _ t (v1, v2)) G34 -> abstract_domain_R v1 v2) + : wf_value G' (interp1 annotate_with_state e1 est1) (interp2 annotate_with_state e2 est2). + Proof using abstract_interp_ident_Proper annotate_Proper bottom'_Proper interp_ident_Proper try_make_transport_base_type_cps_correct. + revert dependent G12; revert dependent G34; revert dependent G'; revert annotate_with_state; induction Hwf; intros; cbn [interp]; + repeat match goal with + | [ H : List.In (existT _ ?t (?v1, ?v2, ?v3, ?v4)) ?G, H' : List.map _ ?G = ?G12 |- _ ] + => unique assert (List.In (existT _ t (v1, v2)) G12) + by (rewrite <- H'; revert H; rewrite in_map_iff; eexists; split; try eassumption; reflexivity) + | [ H : List.In (existT _ ?t (?v1, ?v2, ?v3, ?v4)) ?G, H' : List.map _ ?G = ?G34 |- _ ] + => unique assert (List.In (existT _ t (v3, v4)) G34) + by (rewrite <- H'; revert H; rewrite in_map_iff; eexists; split; try eassumption; reflexivity) + | [ HG : List.map ?f ?G = ?G12, H : forall v1 v2 v3 v4, expr.wf4 (@?v v1 v2 v3 v4 :: ?G) (@?e1 v1 v2) (@?e2 v1 v2) _ _ |- _ ] + => let v' := (eval cbn in (fun v1 v2 => f (v v1 v2 bottom bottom))) in + let lem := (eval cbv beta in (forall v1 v2, expr.wf (v' v1 v2 :: G12) (e1 v1 v2) (e2 v1 v2))) in + unique assert lem + by (intros v1 v2; specialize (H v1 v2 bottom bottom); + apply expr.wf_of_wf4_default in H; subst G12; cbn [List.map] in H; + apply H; split; apply bottom) + | [ HG : List.map ?f ?G = ?G12, H : expr.wf4 ?G ?e1 ?e2 _ _ |- _ ] + => unique assert (expr.wf G12 e1 e2) + by (apply expr.wf_of_wf4_default in H; subst G12; cbn [List.map] in H; + apply H; split; apply bottom) + | [ HG : List.map ?f ?G = ?G34, H : expr.wf4 ?G _ _ ?e3 ?e4 |- _ ] + => unique assert (expr.wf G34 e3 e4) + by (apply expr.wf_of_wf4_default in H; subst G34; cbn [List.map] in H; + apply H; split; apply bottom) + end; try solve [ apply interp_ident_Proper; auto | eauto ]; match goal with @@ -312,27 +390,97 @@ Module Compilers. => match goal with | [ |- context[skip_annotations_for_App ?e2v] ] => epose proof (wf_skip_annotations_for_App (e1:=e1v) (e2:=e2v) (G:=G') ltac:(solve [ wf_t ])); - generalize dependent (skip_annotations_for_App e1v); - generalize dependent (skip_annotations_for_App e2v); intros; - subst + generalize dependent (skip_annotations_for_App e1v); + generalize dependent (skip_annotations_for_App e2v); intros; + subst end end; wf_interp_t. + all: lazymatch goal with + | [ H : type.related ?R ?x ?y |- type.related ?R (expr.interp ?ii1 ?f1 ?x) (expr.interp ?ii2 ?f2 ?y) ] + => is_var x; is_var y; + revert x y H; + change (type.related R (expr.interp ii1 f1) (expr.interp ii2 f2)); + eapply expr.wf_interp_Proper_gen2; eauto + | _ => idtac + end. + all: repeat match goal with + | [ H : invert_Abs (expr.Abs _) = _ |- _ ] + => cbv in H; inversion_option; subst + end. + all: expr.inversion_wf_constr; wf_interp_t. + all: try now + ( + do_with_exactly_one_hyp ltac:(fun H => eapply H; clear H); + lazymatch goal with + | [ |- expr.wf _ _ _ ] => now eauto + | _ => idtac + end; + try reflexivity; + [ > wf_interp_t .. ]; + cbn [List.In List.map] in *; + wf_interp_t + ). + all: try now + ( + eapply UnderLets.wf_Proper_list_impl; + [ .. | do_with_exactly_one_hyp ltac:(fun H => eapply H; clear H) ]; + try reflexivity; + [ .. | wf_interp_t ]; + wf_interp_t + ). + all: try now + ( + eapply expr.wf_interp_Proper_gen2; + lazymatch goal with + | [ |- expr.wf _ _ _ ] => now eauto + | _ => idtac + end; + cbn [List.map List.In]; + wf_interp_t + ). + all: try now + ( + do_with_hyp' ltac:(fun H => eapply H; clear H); + try reflexivity; + cbn [List.In List.map] in *; + wf_interp_t + ). + all: try do_with_hyp' + ltac:(fun H + => now + ( + eapply H; + try reflexivity; + wf_interp_t; + cbn [List.In List.map] in *; + wf_interp_t + ) + ). + Unshelve. + all: assumption. Qed. - Lemma wf_eval_with_bound' (annotate_with_state : bool) G G' {t} e1 e2 (He : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) - (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) - : expr.wf G' (@eval_with_bound'1 annotate_with_state t e1 st1) (@eval_with_bound'2 annotate_with_state t e2 st2). - Proof using annotate_Proper bottom'_Proper interp_ident_Proper. - eapply UnderLets.wf_to_expr, UnderLets.wf_splice. - { eapply wf_interp; solve [ eauto ]. } - { intros; destruct_head'_ex; subst; eapply wf_reify; eauto. } + Lemma wf_eval_with_bound' (annotate_with_state : bool) G G12 G34 G' {t} e1 e2 est1 est2 (He : expr.wf4 G e1 e2 est1 est2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) + (HG12 : List.map (fun '(existT t (v1, v2, _, _)) => existT _ t (v1, v2)) G = G12) + (HG34 : List.map (fun '(existT t (_, _, v3, v4)) => existT _ t (v3, v4)) G = G34) + (Hwf' : expr.wf G34 est1 est2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G12 -> wf_value G' v1 v2) + (HGG'' : forall t v1 v2, List.In (existT _ t (v1, v2)) G34 -> abstract_domain_R v1 v2) + : expr.wf G' (@eval_with_bound'1 annotate_with_state t e1 est1 st1) (@eval_with_bound'2 annotate_with_state t e2 est2 st2). + Proof using abstract_interp_ident_Proper annotate_Proper bottom'_Proper interp_ident_Proper try_make_transport_base_type_cps_correct. + eapply UnderLets.wf_to_expr, wf_reify; eauto; []. + eapply wf_interp; try eassumption. Qed. - Lemma wf_eval' G G' {t} e1 e2 (He : expr.wf G e1 e2) - (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) - : expr.wf G' (@eval'1 t e1) (@eval'2 t e2). - Proof using annotate_Proper bottom'_Proper interp_ident_Proper. + Lemma wf_eval' G G12 G34 G' {t} e1 e2 est1 est2 (He : expr.wf4 G e1 e2 est1 est2) + (HG12 : List.map (fun '(existT t (v1, v2, _, _)) => existT _ t (v1, v2)) G = G12) + (HG34 : List.map (fun '(existT t (_, _, v3, v4)) => existT _ t (v3, v4)) G = G34) + (Hwf' : expr.wf G34 est1 est2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G12 -> wf_value G' v1 v2) + (HGG'' : forall t v1 v2, List.In (existT _ t (v1, v2)) G34 -> abstract_domain_R v1 v2) + : expr.wf G' (@eval'1 t e1 est1) (@eval'2 t e2 est2). + Proof using abstract_interp_ident_Proper annotate_Proper bottom'_Proper interp_ident_Proper try_make_transport_base_type_cps_correct. eapply wf_eval_with_bound'; eauto; apply bottom_for_each_lhs_of_arrow_Proper. Qed. @@ -389,14 +537,14 @@ Module Compilers. Section with_var2. Context {var1 var2 : type -> Type}. - Local Notation wf_value_with_lets := (@wf_value_with_lets base.type ident abstract_domain' abstract_domain'_R var1 var2). Local Notation wf_value := (@wf_value base.type ident abstract_domain' abstract_domain'_R var1 var2). + Local Notation wf_value' := (@wf_value' base.type ident abstract_domain' abstract_domain'_R var1 var2). Context (annotate_expr1 : forall t, abstract_domain' t -> option (@expr var1 (t -> t))) (annotate_expr2 : forall t, abstract_domain' t -> option (@expr var2 (t -> t))) (is_annotated_for1 : forall t t', @expr var1 t -> abstract_domain' t' -> bool) (is_annotated_for2 : forall t t', @expr var2 t -> abstract_domain' t' -> bool) - (strip_annotation1 : forall t, ident t -> option (value _ t)) - (strip_annotation2 : forall t, ident t -> option (value _ t)) + (strip_annotation1 : forall t, ident t -> option (value' _ t)) + (strip_annotation2 : forall t, ident t -> option (value' _ t)) (annotation_to_cast1 : forall s d, @expr var1 (s -> d) -> option (@expr var1 s -> @expr var1 d)) (annotation_to_cast2 : forall s d, @expr var2 (s -> d) -> option (@expr var2 s -> @expr var2 d)) (skip_annotations_under : forall t, ident t -> bool) @@ -416,7 +564,7 @@ Module Compilers. (@is_annotated_for2 t t' e2)} {wf_strip_annotation : forall G t idc, - option_eq (wf_value G) + option_eq (wf_value' G) (@strip_annotation1 t idc) (@strip_annotation2 t idc)}. @@ -428,8 +576,8 @@ Module Compilers. Local Notation annotate_base2 := (@ident.annotate_base var2 abstract_domain' annotate_expr2 is_annotated_for2). Local Notation annotate_with_expr1 := (@ident.annotate_with_expr var1 abstract_domain' annotate_expr1 is_annotated_for1). Local Notation annotate_with_expr2 := (@ident.annotate_with_expr var2 abstract_domain' annotate_expr2 is_annotated_for2). - Local Notation interp_ident1 := (@ident.interp_ident var1 abstract_domain' annotate_expr1 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for1 strip_annotation1). - Local Notation interp_ident2 := (@ident.interp_ident var2 abstract_domain' annotate_expr2 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for2 strip_annotation2). + Local Notation interp_ident'1 := (@ident.interp_ident' var1 abstract_domain' annotate_expr1 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for1 strip_annotation1). + Local Notation interp_ident'2 := (@ident.interp_ident' var2 abstract_domain' annotate_expr2 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for2 strip_annotation2). Local Notation reflect1 := (@reflect base.type ident var1 abstract_domain' annotate1 bottom'). Local Notation reflect2 := (@reflect base.type ident var2 abstract_domain' annotate2 bottom'). @@ -608,23 +756,30 @@ Module Compilers. | progress cbn [type.decode f_equal eq_rect fst snd] in * | solve [ wf_t ] ] ]. Qed. - Local Ltac type_of_value v := + Local Ltac type_of_value' v := lazymatch v with | (abstract_domain ?t * _)%type => t + | expr ?t => t | (?a -> UnderLets _ ?b) - => let a' := type_of_value a in - let b' := type_of_value b in + => let a' := type_of_value' a in + let b' := type_of_value' b in constr:(type.arrow a' b') + | ?v => fail 0 "type_of_value': unrecognized:" v end. Local Opaque ident.ident_Literal. Lemma wf_interp_ident_nth_default (annotate_with_state : bool) G T - : wf_value_with_lets G (@interp_ident1 annotate_with_state _ (@ident.List_nth_default T)) (@interp_ident2 annotate_with_state _ (@ident.List_nth_default T)). + : UnderLets.wf (fun G' => wf_value' G') G (@interp_ident'1 annotate_with_state _ (@ident.List_nth_default T)) (@interp_ident'2 annotate_with_state _ (@ident.List_nth_default T)). Proof using abstract_interp_ident_Proper annotate_expr_Proper extract_list_state_length extract_list_state_rel extract_option_state_rel is_annotated_for_Proper. - cbv [wf_value_with_lets wf_value ident.interp_ident]; constructor; cbn -[abstract_domain_R abstract_domain]. - { intros; subst. - destruct_head'_prod; destruct_head'_and; cbn [fst snd] in *. + cbv [wf_value wf_value' ident.interp_ident']; constructor; cbn -[abstract_domain_R abstract_domain]. + all: intros; subst. + all: destruct_head'_prod; destruct_head'_and; cbn [fst snd] in *. + all: repeat first [ progress subst | lazymatch goal with + | [ H : wf_value' _ (reify_list _) _ |- _ ] + => progress cbn [wf_value'] in H + | [ H : wf_value' _ _ (reify_list _) |- _ ] + => progress cbn [wf_value'] in H | [ H : expr.wf _ ?e1 ?e2, H' : reflect_list ?e1 = Some _, H'' : reflect_list ?e2 = None |- _ ] => apply expr.wf_reflect_list in H; rewrite H', H'' in H; exfalso; clear -H; intuition congruence | [ H : expr.wf _ ?e1 ?e2, H' : reflect_list ?e2 = Some _, H'' : reflect_list ?e1 = None |- _ ] @@ -655,8 +810,8 @@ Module Compilers. | [ |- expr _ -> _ -> _ ] => refine (expr.wf G) | [ |- ?T -> _ -> _ ] - => let t := type_of_value T in - refine (@wf_value G t) + => let t := type_of_value' T in + refine (@wf_value' G t) end | | ] | [ |- UnderLets.wf ?Q ?G (UnderLets.Base _) (UnderLets.Base _) ] @@ -666,8 +821,8 @@ Module Compilers. | [ H : _ = _ :> ident _ |- _ ] => inversion H; clear H | [ H : List.nth_error _ _ = None |- _ ] => apply List.nth_error_None in H | [ H : List.nth_error _ _ = Some _ |- _ ] - => unique pose proof (@ListUtil.nth_error_value_length _ _ _ _ H); - unique pose proof (@ListUtil.nth_error_value_In _ _ _ _ H) + => unique pose proof (@nth_error_value_length _ _ _ _ H); + unique pose proof (@nth_error_value_In _ _ _ _ H) | [ H : context[List.In _ (List.map _ _)] |- _ ] => rewrite List.in_map_iff in H | [ H : (?x <= ?y)%nat, H' : (?y < ?x)%nat |- _ ] => exfalso; clear -H H'; lia | [ H : (?x <= ?y)%nat, H' : (?y < ?x')%nat, H'' : ?x' = ?x |- _ ] => exfalso; clear -H H' H''; lia @@ -686,7 +841,7 @@ Module Compilers. end | progress inversion_option | progress intros - | progress cbn [fst snd value] in * + | progress cbn [fst snd value'] in * | progress destruct_head'_prod | progress destruct_head'_ex | progress destruct_head'_and @@ -698,10 +853,9 @@ Module Compilers. | progress expr.inversion_expr | handle_Forall2_step | wf_safe_t_step - | progress destruct_head' (@partial.wf_value) | solve [ eapply wf_annotate; wf_t; try apply DefaultValue.expr.base.wf_default | eapply wf_annotate_base; wf_t - | eapply (abstract_interp_ident_Proper _ (@ident.List_nth_default T) _ eq_refl); assumption + | eapply (abstract_interp_ident_Proper _ (@ident.List_nth_default _) _ eq_refl); assumption | eapply wf_update_annotation; wf_t | wf_t | match goal with @@ -723,14 +877,14 @@ Module Compilers. | progress expr.inversion_wf_one_constr | progress expr.invert_match | match goal with - | [ |- wf_value _ _ _ ] => progress hnf - end ]. } + | [ |- wf_value' _ _ _ ] => progress hnf + end ]. Qed. Lemma wf_interp_ident_not_nth_default_nostrip (annotate_with_state : bool) G {t} (idc : ident t) - : wf_value_with_lets G (Base (reflect1 annotate_with_state (###idc)%expr (abstract_interp_ident _ idc))) (Base (reflect2 annotate_with_state (###idc)%expr (abstract_interp_ident _ idc))). + : wf_value G (reflect1 annotate_with_state (###idc)%expr (abstract_interp_ident _ idc)) (reflect2 annotate_with_state (###idc)%expr (abstract_interp_ident _ idc)). Proof using abstract_interp_ident_Proper annotate_expr_Proper bottom'_Proper extract_list_state_length extract_list_state_rel extract_option_state_rel is_annotated_for_Proper. - constructor; eapply wf_reflect; + constructor; unshelve eapply wf_reflect; solve [ apply bottom'_Proper | apply wf_annotate | repeat constructor @@ -738,30 +892,31 @@ Module Compilers. Qed. Lemma wf_interp_ident_not_nth_default (annotate_with_state : bool) G {t} (idc : ident t) - : (wf_value_with_lets G) - (Base match strip_annotation1 _ idc with - | Some v => v - | None => reflect1 annotate_with_state (###idc)%expr (abstract_interp_ident _ idc) - end) - (Base match strip_annotation2 _ idc with - | Some v => v - | None => reflect2 annotate_with_state (###idc)%expr (abstract_interp_ident _ idc) - end). + : (UnderLets.wf (fun G' => wf_value' G') G) + (match strip_annotation1 _ idc with + | Some v => Base v + | None => project_value' _ (reflect1 annotate_with_state (###idc)%expr (abstract_interp_ident _ idc)) + end) + (match strip_annotation2 _ idc with + | Some v => Base v + | None => project_value' _ (reflect2 annotate_with_state (###idc)%expr (abstract_interp_ident _ idc)) + end). Proof using abstract_interp_ident_Proper annotate_expr_Proper bottom'_Proper extract_list_state_length extract_list_state_rel extract_option_state_rel is_annotated_for_Proper wf_strip_annotation. pose proof (wf_strip_annotation G _ idc). break_innermost_match; cbn [option_eq] in *; solve [ exfalso; assumption | congruence | apply wf_interp_ident_not_nth_default_nostrip - | constructor; assumption ]. + | constructor; assumption + | split; [ now apply abstract_interp_ident_Proper | constructor; assumption ] ]. Qed. Lemma wf_interp_ident G {t} idc1 idc2 (Hidc : idc1 = idc2) (annotate_with_state : bool) - : wf_value_with_lets G (@interp_ident1 annotate_with_state t idc1) (@interp_ident2 annotate_with_state t idc2). + : UnderLets.wf (fun G' => wf_value' G') G (@interp_ident'1 annotate_with_state t idc1) (@interp_ident'2 annotate_with_state t idc2). Proof using abstract_interp_ident_Proper annotate_expr_Proper bottom'_Proper extract_list_state_length extract_list_state_rel extract_option_state_rel is_annotated_for_Proper wf_strip_annotation. - cbv [wf_value_with_lets ident.interp_ident]; subst idc2; destruct idc1; - first [ apply wf_interp_ident_nth_default - | apply wf_interp_ident_not_nth_default ]. + subst idc2; destruct idc1; + first [ apply wf_interp_ident_not_nth_default + | apply wf_interp_ident_nth_default ]. Qed. Local Notation strip_all_annotations'1 := (@partial.ident.strip_all_annotations' var1 annotation_to_cast1 skip_annotations_under). @@ -795,26 +950,38 @@ Module Compilers. Local Notation eval_with_bound1 := (@partial.ident.eval_with_bound var1 abstract_domain' annotate_expr1 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for1 strip_annotation1 skip_annotations_under). Local Notation eval_with_bound2 := (@partial.ident.eval_with_bound var2 abstract_domain' annotate_expr2 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for2 strip_annotation2 skip_annotations_under). - Lemma wf_eval_with_bound (annotate_with_state : bool) {t} G G' e1 e2 (Hwf : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) - (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) - : expr.wf G' (@eval_with_bound1 annotate_with_state t e1 st1) (@eval_with_bound2 annotate_with_state t e2 st2). + + Lemma wf_eval_with_bound (annotate_with_state : bool) {t} G G12 G34 G' e1 e2 est1 est2 (Hwf : expr.wf4 G e1 e2 est1 est2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) + (HG12 : List.map (fun '(existT t (v1, v2, _, _)) => existT _ t (v1, v2)) G = G12) + (HG34 : List.map (fun '(existT t (_, _, v3, v4)) => existT _ t (v3, v4)) G = G34) + (Hwf' : expr.wf G34 est1 est2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G12 -> wf_value G' v1 v2) + (HGG'' : forall t v1 v2, List.In (existT _ t (v1, v2)) G34 -> abstract_domain_R v1 v2) + : expr.wf G' (@eval_with_bound1 annotate_with_state t e1 est1 st1) (@eval_with_bound2 annotate_with_state t e2 est2 st2). Proof using abstract_interp_ident_Proper annotate_expr_Proper bottom'_Proper extract_list_state_length extract_list_state_rel extract_option_state_rel is_annotated_for_Proper wf_strip_annotation. eapply wf_eval_with_bound'; solve [ eassumption | eapply wf_annotate - | eapply wf_interp_ident ]. + | eapply wf_interp_ident + | typeclasses eauto ]. Qed. Local Notation eval1 := (@partial.ident.eval var1 abstract_domain' annotate_expr1 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for1 strip_annotation1). Local Notation eval2 := (@partial.ident.eval var2 abstract_domain' annotate_expr2 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for2 strip_annotation2). - Lemma wf_eval {t} G G' e1 e2 (Hwf : expr.wf G e1 e2) - (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) - : expr.wf G' (@eval1 t e1) (@eval2 t e2). + + Lemma wf_eval {t} G G12 G34 G' e1 e2 est1 est2 (Hwf : expr.wf4 G e1 e2 est1 est2) + (HG12 : List.map (fun '(existT t (v1, v2, _, _)) => existT _ t (v1, v2)) G = G12) + (HG34 : List.map (fun '(existT t (_, _, v3, v4)) => existT _ t (v3, v4)) G = G34) + (Hwf' : expr.wf G34 est1 est2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G12 -> wf_value G' v1 v2) + (HGG'' : forall t v1 v2, List.In (existT _ t (v1, v2)) G34 -> abstract_domain_R v1 v2) + : expr.wf G' (@eval1 t e1 est1) (@eval2 t e2 est2). Proof using abstract_interp_ident_Proper annotate_expr_Proper bottom'_Proper extract_list_state_length extract_list_state_rel extract_option_state_rel is_annotated_for_Proper wf_strip_annotation. eapply wf_eval'; solve [ eassumption | eapply wf_annotate - | eapply wf_interp_ident ]. + | eapply wf_interp_ident + | typeclasses eauto ]. Qed. Local Notation eta_expand_with_bound1 := (@partial.ident.eta_expand_with_bound var1 abstract_domain' annotate_expr1 bottom' abstract_interp_ident extract_list_state extract_option_state is_annotated_for1). @@ -837,7 +1004,7 @@ Module Compilers. Local Notation abstract_domain := (@partial.abstract_domain base.type abstract_domain'). Local Notation abstract_domain'_R t := (@eq (abstract_domain' t)) (only parsing). Local Notation abstract_domain_R := (@abstract_domain_R base.type abstract_domain' (fun t => abstract_domain'_R t)). - Local Notation wf_value := (@wf_value base.type ident abstract_domain' (fun t => abstract_domain'_R t)). + Local Notation wf_value' := (@wf_value' base.type ident abstract_domain' (fun t => abstract_domain'_R t)). Lemma annotate_expr_Proper {relax_zrange var1 var2} {t} s1 s2 : abstract_domain'_R t s1 s2 @@ -877,12 +1044,12 @@ Module Compilers. (#(@ident.Literal base.type.zrange r1%zrange), #(@ident.Literal base.type.zrange r2%zrange))%expr_pat). Lemma wf_always_strip_annotation {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) {var1 var2} G {t} idc - : option_eq (wf_value G) + : option_eq (wf_value' G) (always_strip_annotation assume_cast_truncates (var:=var1) t idc) (always_strip_annotation assume_cast_truncates (var:=var2) t idc). Proof using Type. cbv [always_strip_annotation]; break_innermost_match; try reflexivity; - cbn [option_eq wf_value]. + cbn [option_eq wf_value']. all: repeat first [ progress intros | assumption | progress subst @@ -899,7 +1066,7 @@ Module Compilers. Qed. Lemma wf_strip_annotation {opts : AbstractInterpretation.Options} (assume_cast_truncates : bool) (strip_annotations : bool) {var1 var2} G {t} idc - : option_eq (wf_value G) + : option_eq (wf_value' G) (strip_annotation assume_cast_truncates strip_annotations (var:=var1) t idc) (strip_annotation assume_cast_truncates strip_annotations (var:=var2) t idc). Proof using Type. @@ -1041,11 +1208,15 @@ Module Compilers. Section with_var2. Context {var1 var2 : type -> Type}. - Local Notation wf_value_with_lets := (@wf_value_with_lets base.type ident abstract_domain' (fun t => abstract_domain'_R t) var1 var2). - - Lemma wf_eval {opts : AbstractInterpretation.Options} {t} G G' e1 e2 (Hwf : expr.wf G e1 e2) - (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) - : expr.wf G' (t:=t) (eval (var:=var1) e1) (eval (var:=var2) e2). + Local Notation wf_value := (@wf_value base.type ident abstract_domain' (fun t => abstract_domain'_R t) var1 var2). + + Lemma wf_eval {opts : AbstractInterpretation.Options} {t} G G12 G34 G' e1 e2 est1 est2 (Hwf : expr.wf4 G e1 e2 est1 est2) + (HG12 : List.map (fun '(existT t (v1, v2, _, _)) => existT _ t (v1, v2)) G = G12) + (HG34 : List.map (fun '(existT t (_, _, v3, v4)) => existT _ t (v3, v4)) G = G34) + (Hwf' : expr.wf G34 est1 est2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G12 -> wf_value G' v1 v2) + (HGG'' : forall t v1 v2, List.In (existT _ t (v1, v2)) G34 -> abstract_domain_R v1 v2) + : expr.wf G' (t:=t) (eval (var:=var1) e1 est1) (eval (var:=var2) e2 est2). Proof using Type. eapply ident.wf_eval; solve [ eassumption @@ -1062,11 +1233,15 @@ Module Compilers. : expr.wf G (@strip_all_annotations strip_annotations_under var1 t e1) (@strip_all_annotations strip_annotations_under var2 t e2). Proof using Type. revert Hwf; apply ident.wf_strip_all_annotations, @wf_annotation_to_cast. Qed. - Lemma wf_eval_with_bound {opts : AbstractInterpretation.Options} {relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations t} G G' e1 e2 (Hwf : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) - (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) + Lemma wf_eval_with_bound {opts : AbstractInterpretation.Options} {relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations t} G G12 G34 G' e1 e2 est1 est2 (Hwf : expr.wf4 G e1 e2 est1 est2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) + (HG12 : List.map (fun '(existT t (v1, v2, _, _)) => existT _ t (v1, v2)) G = G12) + (HG34 : List.map (fun '(existT t (_, _, v3, v4)) => existT _ t (v3, v4)) G = G34) + (Hwf' : expr.wf G34 est1 est2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G12 -> wf_value G' v1 v2) + (HGG'' : forall t v1 v2, List.In (existT _ t (v1, v2)) G34 -> abstract_domain_R v1 v2) : expr.wf G' (var1:=var1) (var2:=var2) (t:=t) - (eval_with_bound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e1 st1) - (eval_with_bound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e2 st2). + (eval_with_bound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e1 est1 st1) + (eval_with_bound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e2 est2 st2). Proof using Type. eapply ident.wf_eval_with_bound; solve [ eassumption @@ -1079,9 +1254,13 @@ Module Compilers. | apply is_annotated_for_Proper ]. Qed. - Lemma wf_strip_annotations {opts : AbstractInterpretation.Options} {assume_cast_truncates} {t} G G' e1 e2 (Hwf : expr.wf G e1 e2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) - (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G -> wf_value_with_lets G' v1 v2) - : expr.wf G' (t:=t) (strip_annotations assume_cast_truncates (var:=var1) e1 st1) (strip_annotations assume_cast_truncates (var:=var2) e2 st2). + Lemma wf_strip_annotations {opts : AbstractInterpretation.Options} {assume_cast_truncates} {t} G G12 G34 G' e1 e2 est1 est2 (Hwf : expr.wf4 G e1 e2 est1 est2) st1 st2 (Hst : type.and_for_each_lhs_of_arrow (@abstract_domain_R) st1 st2) + (HG12 : List.map (fun '(existT t (v1, v2, _, _)) => existT _ t (v1, v2)) G = G12) + (HG34 : List.map (fun '(existT t (_, _, v3, v4)) => existT _ t (v3, v4)) G = G34) + (Hwf' : expr.wf G34 est1 est2) + (HGG' : forall t v1 v2, List.In (existT _ t (v1, v2)) G12 -> wf_value G' v1 v2) + (HGG'' : forall t v1 v2, List.In (existT _ t (v1, v2)) G34 -> abstract_domain_R v1 v2) + : expr.wf G' (t:=t) (strip_annotations assume_cast_truncates (var:=var1) e1 est1 st1) (strip_annotations assume_cast_truncates (var:=var2) e2 est2 st2). Proof using Type. eapply ident.wf_eval_with_bound; solve [ eassumption @@ -1115,19 +1294,19 @@ Module Compilers. Lemma Wf_Eval {opts : AbstractInterpretation.Options} {t} (e : Expr t) (Hwf : Wf e) : Wf (Eval e). Proof using Type. - intros ??; eapply wf_eval with (G:=nil); cbn [List.In]; try apply Hwf; tauto. + intros ??; cbv [Eval Let_In]; eapply wf_eval with (G:=nil); try reflexivity; cbn [List.In List.map]; try apply expr.Wf4_of_Wf; try apply GeneralizeVar.Wf_FromFlat_ToFlat, Hwf; tauto. Qed. Lemma Wf_EvalWithBound {opts : AbstractInterpretation.Options} {relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations t} (e : Expr t) bound (Hwf : Wf e) (bound_valid : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) bound) : Wf (EvalWithBound relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations e bound). Proof using Type. - intros ??; eapply wf_eval_with_bound with (G:=nil); cbn [List.In]; try apply Hwf; tauto. + intros ??; cbv [EvalWithBound Let_In]; eapply wf_eval_with_bound with (G:=nil); try reflexivity; cbn [List.In List.map]; try apply expr.Wf4_of_Wf; try apply GeneralizeVar.Wf_FromFlat_ToFlat, Hwf; tauto. Qed. Lemma Wf_StripAnnotations {opts : AbstractInterpretation.Options} {assume_cast_truncates} {t} (e : Expr t) bound (Hwf : Wf e) (bound_valid : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) bound) : Wf (StripAnnotations assume_cast_truncates e bound). Proof using Type. - intros ??; eapply wf_strip_annotations with (G:=nil); cbn [List.In]; try tauto; apply Hwf. + intros ??; cbv [StripAnnotations]; eapply wf_strip_annotations with (G:=nil); try reflexivity; cbn [List.In List.map]; try apply expr.Wf4_of_Wf; try apply Hwf; tauto. Qed. Lemma Wf_EtaExpandWithBound {relax_zrange t} (e : Expr t) bound (Hwf : Wf e) (bound_valid : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R)) bound) diff --git a/src/AbstractInterpretation/Proofs.v b/src/AbstractInterpretation/Proofs.v new file mode 100644 index 0000000000..8566164329 --- /dev/null +++ b/src/AbstractInterpretation/Proofs.v @@ -0,0 +1,118 @@ +Require Import Coq.Classes.Morphisms. +Require Import Crypto.Util.ZRange. +Require Import Crypto.Util.Tactics.BreakMatch. +Require Import Rewriter.Language.Language. +Require Import Rewriter.Language.Wf. +Require Import Crypto.Language.API. +Require Import Crypto.AbstractInterpretation.AbstractInterpretation. +Require Import Crypto.AbstractInterpretation.Wf. +Require Crypto.AbstractInterpretation.Fancy.Proofs. +Require Crypto.AbstractInterpretation.Bottomify.Proofs. + +Module Compilers. + Import Language.Compilers. + Import AbstractInterpretation.Compilers. + Import Language.Wf.Compilers. + Import Language.API.Compilers. + Import AbstractInterpretation.Wf.Compilers. + Import AbstractInterpretation.ZRangeProofs.Compilers. + (*Import AbstractInterpretation.Wf.Compilers.partial.*) + Import invert_expr. + + Import API. + + Lemma Interp_PartialEvaluateWithBounds + {opts : AbstractInterpretation.Options} + relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations + (Hrelax : forall r r' z, is_tighter_than_bool z r = true + -> relax_zrange r = Some r' + -> is_tighter_than_bool z r' = true) + {t} (E : Expr t) + (Hwf : Wf E) + (Ht : type.is_not_higher_order t = true) (* only needed for bottomify version *) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + : forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), + type.app_curried (expr.Interp (@ident.interp) (PartialEvaluateWithBounds relax_zrange skip_annotations_under strip_preexisting_annotations assume_cast_truncates E b_in)) arg1 + = type.app_curried (expr.Interp (@ident.interp) E) arg2. + Proof. + cbv [PartialEvaluateWithBounds]. + break_innermost_match. + all: first [ eapply @Fancy.Proofs.Compilers.Interp_PartialEvaluateWithBounds + | eapply @Bottomify.Proofs.Compilers.Interp_PartialEvaluateWithBounds ]. + all: assumption. + Qed. + + Lemma Interp_PartialEvaluateWithBounds_bounded + {opts : AbstractInterpretation.Options} + relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations + (Hrelax : forall r r' z, is_tighter_than_bool z r = true + -> relax_zrange r = Some r' + -> is_tighter_than_bool z r' = true) + {t} (E : Expr t) + (Hwf : Wf E) + (Ht : type.is_not_higher_order t = true) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + : forall arg1 + (Harg11 : Proper (type.and_for_each_lhs_of_arrow (@type.eqv)) arg1) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), + ZRange.type.base.option.is_bounded_by + (partial.Extract assume_cast_truncates E b_in) + (type.app_curried (expr.Interp (@ident.interp) (PartialEvaluateWithBounds relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations E b_in)) arg1) + = true. + Proof. + cbv [PartialEvaluateWithBounds partial.Extract]. + cbv [fancy_and_powerful_but_exponentially_slow_bounds_analysis AbstractInterpretation.fancy_and_powerful_but_exponentially_slow_bounds_analysis]; destruct opts. + break_innermost_match. + all: first [ eapply @Fancy.Proofs.Compilers.Interp_PartialEvaluateWithBounds_bounded + | eapply @Bottomify.Proofs.Compilers.Interp_PartialEvaluateWithBounds_bounded ]. + all: assumption. + Qed. + + Lemma Interp_PartialEvaluateWithListInfoFromBounds + {opts : AbstractInterpretation.Options} + {t} (E : Expr t) + (Hwf : Wf E) + (Ht : type.is_not_higher_order t = true) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + : forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), + type.app_curried (Interp (PartialEvaluateWithListInfoFromBounds E b_in)) arg1 = type.app_curried (Interp E) arg2. + Proof. + cbv [PartialEvaluateWithListInfoFromBounds]. + break_innermost_match. + all: first [ eapply @Fancy.Proofs.Compilers.Interp_PartialEvaluateWithListInfoFromBounds + | eapply @Bottomify.Proofs.Compilers.Interp_PartialEvaluateWithListInfoFromBounds ]. + all: assumption. + Qed. + + Theorem CheckedPartialEvaluateWithBounds_Correct + {opts : AbstractInterpretation.Options} + (relax_zrange : zrange -> option zrange) + (assume_cast_truncates : bool) + (skip_annotations_under : forall t, ident t -> bool) + (strip_preexisting_annotations : bool) + (Hrelax : forall r r' z, is_tighter_than_bool z r = true + -> relax_zrange r = Some r' + -> is_tighter_than_bool z r' = true) + {t} (E : Expr t) + (Hwf : Wf E) + (Ht : type.is_not_higher_order t = true) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + (b_out : ZRange.type.base.option.interp (type.final_codomain t)) + rv (Hrv : CheckedPartialEvaluateWithBounds relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations E b_in b_out = inl rv) + : (forall arg1 arg2 + (Harg12 : type.and_for_each_lhs_of_arrow (@type.eqv) arg1 arg2) + (Harg1 : type.andb_bool_for_each_lhs_of_arrow (@ZRange.type.option.is_bounded_by) b_in arg1 = true), + ZRange.type.base.option.is_bounded_by b_out (type.app_curried (Interp rv) arg1) = true + /\ type.app_curried (Interp rv) arg1 = type.app_curried (Interp E) arg2) + /\ Wf rv. + Proof. + cbv [CheckedPartialEvaluateWithBounds] in *. + break_innermost_match_hyps. + all: first [ eapply @Fancy.Proofs.Compilers.CheckedPartialEvaluateWithBounds_Correct_first_order; eassumption + | eapply @Bottomify.Proofs.Compilers.CheckedPartialEvaluateWithBounds_Correct; eassumption ]. + Qed. +End Compilers. diff --git a/src/AbstractInterpretation/Wf.v b/src/AbstractInterpretation/Wf.v new file mode 100644 index 0000000000..eb595df30f --- /dev/null +++ b/src/AbstractInterpretation/Wf.v @@ -0,0 +1,44 @@ +Require Import Coq.Classes.Morphisms. +Require Import Crypto.Util.ZRange. +Require Import Crypto.Util.Tactics.BreakMatch. +Require Import Rewriter.Language.Language. +Require Import Rewriter.Language.Wf. +Require Crypto.AbstractInterpretation.Fancy.Wf. +Require Crypto.AbstractInterpretation.Bottomify.Wf. +Require Import Crypto.AbstractInterpretation.AbstractInterpretation. + +Module Compilers. + Import Language.Compilers. + Import AbstractInterpretation.AbstractInterpretation.Compilers. + Import Language.Wf.Compilers. + Import API. + + Export AbstractInterpretation.Fancy.Wf.Compilers (partial.abstract_domain_R). + Import AbstractInterpretation.Fancy.Wf.Compilers.partial (abstract_domain_R). + + Lemma Wf_PartialEvaluateWithListInfoFromBounds + {opts : AbstractInterpretation.Options} + {t} (E : Expr t) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + (Hwf : Wf E) + {b_in_Proper : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R base.type ZRange.type.base.option.interp (fun t0 : base.type => eq))) b_in} + : Wf (PartialEvaluateWithListInfoFromBounds E b_in). + Proof. cbv [PartialEvaluateWithListInfoFromBounds]; break_innermost_match; eauto with wf. Qed. +#[global] + Hint Resolve Wf_PartialEvaluateWithListInfoFromBounds : wf. +#[global] + Hint Opaque PartialEvaluateWithListInfoFromBounds : wf interp rewrite. + + Lemma Wf_PartialEvaluateWithBounds + {opts : AbstractInterpretation.Options} + {relax_zrange} {assume_cast_truncates : bool} {skip_annotations_under : forall t, ident t -> bool} {strip_preexisting_annotations : bool} {t} (E : Expr t) + (b_in : type.for_each_lhs_of_arrow ZRange.type.option.interp t) + (Hwf : Wf E) + {b_in_Proper : Proper (type.and_for_each_lhs_of_arrow (@abstract_domain_R base.type ZRange.type.base.option.interp (fun t0 : base.type => eq))) b_in} + : Wf (PartialEvaluateWithBounds relax_zrange assume_cast_truncates skip_annotations_under strip_preexisting_annotations E b_in). + Proof. cbv [PartialEvaluateWithBounds]; break_innermost_match; eauto with wf. Qed. +#[global] + Hint Resolve Wf_PartialEvaluateWithBounds : wf. +#[global] + Hint Opaque PartialEvaluateWithBounds : wf interp rewrite. +End Compilers. diff --git a/src/AbstractInterpretation/WfExtra.v b/src/AbstractInterpretation/WfExtra.v new file mode 100644 index 0000000000..1d2798873b --- /dev/null +++ b/src/AbstractInterpretation/WfExtra.v @@ -0,0 +1,37 @@ +Require Import Rewriter.Language.Language. +Require Import Rewriter.Language.Wf. +Require Import Crypto.Language.API. +Require Import Crypto.Language.WfExtra. +Require Import Crypto.AbstractInterpretation.AbstractInterpretation. +Require Import Crypto.AbstractInterpretation.Wf. +Require Import Crypto.AbstractInterpretation.Fancy.WfExtra. +Require Import Crypto.AbstractInterpretation.Bottomify.WfExtra. + +Module Compilers. + Export AbstractInterpretation.Fancy.WfExtra.Compilers. + Export AbstractInterpretation.Bottomify.WfExtra.Compilers. + Import AbstractInterpretation.AbstractInterpretation.Compilers. + Import AbstractInterpretation.Wf.Compilers. + Import Compilers.API. + + (* + Import AbstractInterpretation.AbstractInterpretation.Compilers.partial. + Import AbstractInterpretation.Wf.Compilers.partial. + *) + + (* +#[global] + Hint Resolve Wf_Eval Wf_EvalWithBound Wf_EtaExpandWithBound Wf_EtaExpandWithListInfoFromBound Wf_StripAnnotations Wf_StripAllAnnotations : wf_extra. +#[global] + Hint Opaque partial.Eval EvalWithBound EtaExpandWithBound EtaExpandWithListInfoFromBound StripAnnotations StripAllAnnotations : wf_extra interp_extra. +*) +#[global] + Hint Resolve Wf_PartialEvaluateWithListInfoFromBounds : wf_extra. +#[global] + Hint Opaque PartialEvaluateWithListInfoFromBounds : wf_extra interp_extra. + +#[global] + Hint Resolve Wf_PartialEvaluateWithBounds : wf_extra. +#[global] + Hint Opaque PartialEvaluateWithBounds : wf_extra interp_extra. +End Compilers. diff --git a/src/AbstractInterpretation/ZRange.v b/src/AbstractInterpretation/ZRange.v index 5a4b71c85e..a8c297315b 100644 --- a/src/AbstractInterpretation/ZRange.v +++ b/src/AbstractInterpretation/ZRange.v @@ -20,18 +20,30 @@ Module Compilers. Class shiftr_avoid_uint1_opt := shiftr_avoid_uint1 : bool. #[global] Typeclasses Opaque shiftr_avoid_uint1_opt. + Class fancy_and_powerful_but_exponentially_slow_bounds_analysis_opt := fancy_and_powerful_but_exponentially_slow_bounds_analysis : bool. + #[global] + Typeclasses Opaque fancy_and_powerful_but_exponentially_slow_bounds_analysis_opt. Module AbstractInterpretation. Local Set Primitive Projections. Class Options := { shiftr_avoid_uint1 : shiftr_avoid_uint1_opt + ; fancy_and_powerful_but_exponentially_slow_bounds_analysis : fancy_and_powerful_but_exponentially_slow_bounds_analysis_opt }. Definition default_Options : Options - := {| shiftr_avoid_uint1 := false |}. + := {| shiftr_avoid_uint1 := false + ; fancy_and_powerful_but_exponentially_slow_bounds_analysis := false + |}. Module Export Exports. Global Existing Instance Build_Options. Global Existing Instance shiftr_avoid_uint1. - Global Hint Cut [ ( _ * ) shiftr_avoid_uint1 ( _ * ) Build_Options ] : typeclass_instances. + Global Existing Instance fancy_and_powerful_but_exponentially_slow_bounds_analysis. + Global Hint Cut [ ( _ * ) + ( shiftr_avoid_uint1 + | fancy_and_powerful_but_exponentially_slow_bounds_analysis + ) + ( _ * ) Build_Options ] : typeclass_instances. Global Coercion shiftr_avoid_uint1 : Options >-> shiftr_avoid_uint1_opt. + Global Coercion fancy_and_powerful_but_exponentially_slow_bounds_analysis : Options >-> fancy_and_powerful_but_exponentially_slow_bounds_analysis_opt. End Exports. End AbstractInterpretation. Export AbstractInterpretation.Exports. diff --git a/src/CLI.v b/src/CLI.v index 2e9d71f955..6f9d054482 100644 --- a/src/CLI.v +++ b/src/CLI.v @@ -417,6 +417,10 @@ Module ForExtraction. := ([Arg.long_key "cmovznz-by-mul"], Arg.Unit, ["Use an alternative implementation of cmovznz using multiplication rather than bitwise-and with -1."]). Definition shiftr_avoid_uint1_spec : named_argT := ([Arg.long_key "shiftr-avoid-uint1"], Arg.Unit, ["Avoid uint1 types at the output of (>>) operations."]). + Definition fancy_and_powerful_but_exponentially_slow_bounds_analysis_spec : named_argT + := ([Arg.long_key "fancy-and-powerful-but-exponentially-slow-bounds-analysis"], + Arg.Unit, + ["Use a more powerful bound analysis implementation that is unfortunately exponentially slow in the number of lines of code. If you know how to implement PHOAS passes of type [expr var -> var * expr var] or if you want to reimplement our bounds analysis pass in a first-order expression representation instead of PHOAS, please contact us! See https://github.com/mit-plv/fiat-crypto/pull/1761 for more details."]). Definition tight_bounds_multiplier_default := default_tight_upperbound_fraction. Definition tight_bounds_multiplier_spec : named_argT := ([Arg.long_key "tight-bounds-mul-by"], @@ -712,6 +716,7 @@ Module ForExtraction. ; relax_primitive_carry_to_bitwidth_spec ; cmovznz_by_mul_spec ; shiftr_avoid_uint1_spec + ; fancy_and_powerful_but_exponentially_slow_bounds_analysis_spec ; only_signed_spec ; hint_file_spec ; output_file_spec @@ -769,6 +774,7 @@ Module ForExtraction. , relax_primitive_carry_to_bitwidthv , cmovznz_by_mulv , shiftr_avoid_uint1v + , fancy_and_powerful_but_exponentially_slow_bounds_analysisv , only_signedv , hint_file_namesv , output_file_namev @@ -849,6 +855,7 @@ Module ForExtraction. ; use_mul_for_cmovznz := to_bool cmovznz_by_mulv ; abstract_interpretation_options := {| AbstractInterpretation.shiftr_avoid_uint1 := to_bool shiftr_avoid_uint1v + ; AbstractInterpretation.fancy_and_powerful_but_exponentially_slow_bounds_analysis := to_bool fancy_and_powerful_but_exponentially_slow_bounds_analysisv |} ; emit_primitives := negb (to_bool no_primitivesv) ; output_options := diff --git a/src/SlowPrimeSynthesisExamples.v b/src/SlowPrimeSynthesisExamples.v index a38603e14f..7c488999a4 100644 --- a/src/SlowPrimeSynthesisExamples.v +++ b/src/SlowPrimeSynthesisExamples.v @@ -468,13 +468,15 @@ Module debugging_21271_from_bytes. vm_compute in e. set (k'' := CheckedPartialEvaluateWithBounds _ _ _ _ _ _ _) in (value of k). cbv [CheckedPartialEvaluateWithBounds] in k''. + cbv [AbstractInterpretation.fancy_and_powerful_but_exponentially_slow_bounds_analysis Pipeline.AbstractInterpretation_opts Primitives.Options.absint_opts Primitives.Options.default_PipelineOptions Pipeline.AbstractInterpretation_opts Pipeline.default_BoundsPipelineOptions AbstractInterpretation.default_Options] in k''. + cbv [Bottomify.AbstractInterpretation.Compilers.CheckedPartialEvaluateWithBounds] in k''. clear -k''. cbv [Rewriter.Util.LetIn.Let_In] in k''. set (e' := (GeneralizeVar.FromFlat (GeneralizeVar.ToFlat e))) in (value of k''). vm_compute in e'; clear e; rename e' into e. - set (b := (partial.Extract _ _ _)) in (value of k''). + set (b := (Bottomify.AbstractInterpretation.Compilers.partial.Extract _ _ _)) in (value of k''). clear -b. - cbv [partial.Extract partial.ident.extract partial.extract_gen type.app_curried partial.extract'] in b. + cbv [Bottomify.AbstractInterpretation.Compilers.partial.Extract partial.ident.extract partial.extract_gen type.app_curried partial.extract] in b. subst e. cbv beta iota zeta in b. Import Rewriter.Util.LetIn. @@ -1961,6 +1963,8 @@ Module debugging_p256_uint1. => set (k' := e) in (value of k) end; vm_compute in k'; subst k'. cbv [CheckedPartialEvaluateWithBounds] in k. + cbv [AbstractInterpretation.fancy_and_powerful_but_exponentially_slow_bounds_analysis Pipeline.AbstractInterpretation_opts Primitives.Options.absint_opts Primitives.Options.default_PipelineOptions Pipeline.AbstractInterpretation_opts Pipeline.default_BoundsPipelineOptions AbstractInterpretation.default_Options] in k. + cbv [Bottomify.AbstractInterpretation.Compilers.CheckedPartialEvaluateWithBounds] in k. cbv [Rewriter.Util.LetIn.Let_In] in k. set (k' := GeneralizeVar.FromFlat (GeneralizeVar.ToFlat _)) in (value of k). vm_compute in k'. @@ -1987,7 +1991,7 @@ Module debugging_p256_uint1. cbv [split_multiret_to should_split_multiret should_split_multiret_opt_instance_0] in k. vm_compute ZRange.type.base.option.is_tighter_than in k. cbv beta iota zeta in k. - set (k' := PartialEvaluateWithBounds _ _ _ _ _ _) in (value of k) at 1. + set (k' := Bottomify.AbstractInterpretation.Compilers.PartialEvaluateWithBounds _ _ _ _ _ _) in (value of k) at 1. vm_compute in k'. Abort. End __. From d04cd90521358a3c1326dedea413dedc1a8a0980 Mon Sep 17 00:00:00 2001 From: Jason Gross Date: Fri, 22 Dec 2023 14:30:29 -0500 Subject: [PATCH 3/3] Add `--fancy-and-powerful-but-exponentially-slow-bounds-analysis` test --- Makefile.examples | 8 ++++++-- fiat-bedrock2/src/poly1305_32.c | 2 +- fiat-bedrock2/src/poly1305_64.c | 2 +- fiat-c/src/poly1305_32.c | 2 +- fiat-c/src/poly1305_64.c | 2 +- fiat-go/32/poly1305/poly1305.go | 2 +- fiat-go/64/poly1305/poly1305.go | 2 +- fiat-java/src/FiatPoly1305.java | 2 +- fiat-rust/src/poly1305_32.rs | 2 +- fiat-rust/src/poly1305_64.rs | 2 +- fiat-zig/src/poly1305_32.zig | 2 +- fiat-zig/src/poly1305_64.zig | 2 +- 12 files changed, 17 insertions(+), 13 deletions(-) diff --git a/Makefile.examples b/Makefile.examples index ef62196e2b..ac29748146 100644 --- a/Makefile.examples +++ b/Makefile.examples @@ -108,9 +108,13 @@ INVALID_GO_BASE_FILES := $(BASE_FILES_NEEDING_INT128) GO_EXTRA_UNSATURATED_SOLINAS_FUNCTIONS := carry_add carry_sub carry_opp GO_EXTRA_WORD_BY_WORD_MONTGOMERY_FUNCTIONS := +# We add a test for --fancy-and-powerful-but-exponentially-slow-bounds-analysis so that it's not dead code. +# It doesn't really matter which curve we add this flag to, as long as it's fast. +FANCY_BOUNDS_ANALYSIS_FLAG := --fancy-and-powerful-but-exponentially-slow-bounds-analysis + $(foreach bw,64 32,$(eval $(call add_curve_keys,curve25519_$(bw),UNSATURATED_SOLINAS,'25519',$(bw),'(auto)' '2^255 - 19',$(FUNCTIONS_FOR_25519),UNSATURATED_SOLINAS))) -$(eval $(call add_curve_keys,poly1305_64,UNSATURATED_SOLINAS,'poly1305',64,'3' '2^130 - 5',$(UNSATURATED_SOLINAS_FUNCTIONS),UNSATURATED_SOLINAS)) -$(eval $(call add_curve_keys,poly1305_32,UNSATURATED_SOLINAS,'poly1305',32,'(auto)' '2^130 - 5',$(UNSATURATED_SOLINAS_FUNCTIONS),UNSATURATED_SOLINAS)) +$(eval $(call add_curve_keys,poly1305_64,UNSATURATED_SOLINAS,'poly1305',64,'3' '2^130 - 5' $(FANCY_BOUNDS_ANALYSIS_FLAG),$(UNSATURATED_SOLINAS_FUNCTIONS),UNSATURATED_SOLINAS)) +$(eval $(call add_curve_keys,poly1305_32,UNSATURATED_SOLINAS,'poly1305',32,'(auto)' '2^130 - 5' $(FANCY_BOUNDS_ANALYSIS_FLAG),$(UNSATURATED_SOLINAS_FUNCTIONS),UNSATURATED_SOLINAS)) $(foreach bw,64 32,$(eval $(call add_curve_keys,p521_$(bw),UNSATURATED_SOLINAS,'p521',$(bw),'(auto)' '2^521 - 1',$(UNSATURATED_SOLINAS_FUNCTIONS),UNSATURATED_SOLINAS))) ## 2^224 - 2^96 + 1 ## does not bounds check #$(eval $(call add_curve_keys,p224_solinas_64,UNSATURATED_SOLINAS,'p224',64,'4' '2^224 - 2^96 + 1',$(UNSATURATED_SOLINAS_FUNCTIONS),UNSATURATED_SOLINAS)) diff --git a/fiat-bedrock2/src/poly1305_32.c b/fiat-bedrock2/src/poly1305_32.c index 4b8aa9d9bf..c27187f206 100644 --- a/fiat-bedrock2/src/poly1305_32.c +++ b/fiat-bedrock2/src/poly1305_32.c @@ -1,4 +1,4 @@ -/* Autogenerated: 'src/ExtractionOCaml/bedrock2_unsaturated_solinas' --lang bedrock2 --static --no-wide-int --widen-carry --widen-bytes --split-multiret --no-select --no-field-element-typedefs poly1305 32 '(auto)' '2^130 - 5' carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax */ +/* Autogenerated: 'src/ExtractionOCaml/bedrock2_unsaturated_solinas' --lang bedrock2 --static --no-wide-int --widen-carry --widen-bytes --split-multiret --no-select --no-field-element-typedefs poly1305 32 '(auto)' '2^130 - 5' --fancy-and-powerful-but-exponentially-slow-bounds-analysis carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax */ /* curve description: poly1305 */ /* machine_wordsize = 32 (from "32") */ /* requested operations: carry_mul, carry_square, carry, add, sub, opp, selectznz, to_bytes, from_bytes, relax */ diff --git a/fiat-bedrock2/src/poly1305_64.c b/fiat-bedrock2/src/poly1305_64.c index 92ecb99a93..0de53ca73a 100644 --- a/fiat-bedrock2/src/poly1305_64.c +++ b/fiat-bedrock2/src/poly1305_64.c @@ -1,4 +1,4 @@ -/* Autogenerated: 'src/ExtractionOCaml/bedrock2_unsaturated_solinas' --lang bedrock2 --static --no-wide-int --widen-carry --widen-bytes --split-multiret --no-select --no-field-element-typedefs poly1305 64 3 '2^130 - 5' carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax */ +/* Autogenerated: 'src/ExtractionOCaml/bedrock2_unsaturated_solinas' --lang bedrock2 --static --no-wide-int --widen-carry --widen-bytes --split-multiret --no-select --no-field-element-typedefs poly1305 64 3 '2^130 - 5' --fancy-and-powerful-but-exponentially-slow-bounds-analysis carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax */ /* curve description: poly1305 */ /* machine_wordsize = 64 (from "64") */ /* requested operations: carry_mul, carry_square, carry, add, sub, opp, selectznz, to_bytes, from_bytes, relax */ diff --git a/fiat-c/src/poly1305_32.c b/fiat-c/src/poly1305_32.c index 9811d84a1a..bf0020fbd5 100644 --- a/fiat-c/src/poly1305_32.c +++ b/fiat-c/src/poly1305_32.c @@ -1,4 +1,4 @@ -/* Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --inline --static --use-value-barrier poly1305 32 '(auto)' '2^130 - 5' carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax */ +/* Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --inline --static --use-value-barrier poly1305 32 '(auto)' '2^130 - 5' --fancy-and-powerful-but-exponentially-slow-bounds-analysis carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax */ /* curve description: poly1305 */ /* machine_wordsize = 32 (from "32") */ /* requested operations: carry_mul, carry_square, carry, add, sub, opp, selectznz, to_bytes, from_bytes, relax */ diff --git a/fiat-c/src/poly1305_64.c b/fiat-c/src/poly1305_64.c index ca9097170e..cc4935646f 100644 --- a/fiat-c/src/poly1305_64.c +++ b/fiat-c/src/poly1305_64.c @@ -1,4 +1,4 @@ -/* Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --inline --static --use-value-barrier poly1305 64 3 '2^130 - 5' carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax */ +/* Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --inline --static --use-value-barrier poly1305 64 3 '2^130 - 5' --fancy-and-powerful-but-exponentially-slow-bounds-analysis carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax */ /* curve description: poly1305 */ /* machine_wordsize = 64 (from "64") */ /* requested operations: carry_mul, carry_square, carry, add, sub, opp, selectznz, to_bytes, from_bytes, relax */ diff --git a/fiat-go/32/poly1305/poly1305.go b/fiat-go/32/poly1305/poly1305.go index 6232d4f6f5..0333c7d367 100644 --- a/fiat-go/32/poly1305/poly1305.go +++ b/fiat-go/32/poly1305/poly1305.go @@ -1,6 +1,6 @@ // Code generated by Fiat Cryptography. DO NOT EDIT. // -// Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --lang Go --relax-primitive-carry-to-bitwidth 32,64 --cmovznz-by-mul --internal-static --package-case flatcase --public-function-case UpperCamelCase --private-function-case camelCase --public-type-case UpperCamelCase --private-type-case camelCase --no-prefix-fiat --doc-newline-in-typedef-bounds --doc-prepend-header 'Code generated by Fiat Cryptography. DO NOT EDIT.' --doc-text-before-function-name '' --doc-text-before-type-name '' --package-name poly1305 '' 32 '(auto)' '2^130 - 5' carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax carry_add carry_sub carry_opp +// Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --lang Go --relax-primitive-carry-to-bitwidth 32,64 --cmovznz-by-mul --internal-static --package-case flatcase --public-function-case UpperCamelCase --private-function-case camelCase --public-type-case UpperCamelCase --private-type-case camelCase --no-prefix-fiat --doc-newline-in-typedef-bounds --doc-prepend-header 'Code generated by Fiat Cryptography. DO NOT EDIT.' --doc-text-before-function-name '' --doc-text-before-type-name '' --package-name poly1305 '' 32 '(auto)' '2^130 - 5' --fancy-and-powerful-but-exponentially-slow-bounds-analysis carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax carry_add carry_sub carry_opp // // curve description (via package name): poly1305 // diff --git a/fiat-go/64/poly1305/poly1305.go b/fiat-go/64/poly1305/poly1305.go index 21bfaa6024..8164e6b3aa 100644 --- a/fiat-go/64/poly1305/poly1305.go +++ b/fiat-go/64/poly1305/poly1305.go @@ -1,6 +1,6 @@ // Code generated by Fiat Cryptography. DO NOT EDIT. // -// Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --lang Go --no-wide-int --relax-primitive-carry-to-bitwidth 32,64 --cmovznz-by-mul --internal-static --package-case flatcase --public-function-case UpperCamelCase --private-function-case camelCase --public-type-case UpperCamelCase --private-type-case camelCase --no-prefix-fiat --doc-newline-in-typedef-bounds --doc-prepend-header 'Code generated by Fiat Cryptography. DO NOT EDIT.' --doc-text-before-function-name '' --doc-text-before-type-name '' --package-name poly1305 '' 64 3 '2^130 - 5' carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax carry_add carry_sub carry_opp +// Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --lang Go --no-wide-int --relax-primitive-carry-to-bitwidth 32,64 --cmovznz-by-mul --internal-static --package-case flatcase --public-function-case UpperCamelCase --private-function-case camelCase --public-type-case UpperCamelCase --private-type-case camelCase --no-prefix-fiat --doc-newline-in-typedef-bounds --doc-prepend-header 'Code generated by Fiat Cryptography. DO NOT EDIT.' --doc-text-before-function-name '' --doc-text-before-type-name '' --package-name poly1305 '' 64 3 '2^130 - 5' --fancy-and-powerful-but-exponentially-slow-bounds-analysis carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax carry_add carry_sub carry_opp // // curve description (via package name): poly1305 // diff --git a/fiat-java/src/FiatPoly1305.java b/fiat-java/src/FiatPoly1305.java index 7871e5993c..4964542a4b 100644 --- a/fiat-java/src/FiatPoly1305.java +++ b/fiat-java/src/FiatPoly1305.java @@ -1,4 +1,4 @@ -/* Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --lang Java --cmovznz-by-mul --widen-carry --widen-bytes --internal-static --package-name fiat_crypto --class-case UpperCamelCase --no-field-element-typedefs Poly1305 32 '(auto)' '2^130 - 5' carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax */ +/* Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --lang Java --cmovznz-by-mul --widen-carry --widen-bytes --internal-static --package-name fiat_crypto --class-case UpperCamelCase --no-field-element-typedefs Poly1305 32 '(auto)' '2^130 - 5' --fancy-and-powerful-but-exponentially-slow-bounds-analysis carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax */ /* curve description: Poly1305 */ /* machine_wordsize = 32 (from "32") */ /* requested operations: carry_mul, carry_square, carry, add, sub, opp, selectznz, to_bytes, from_bytes, relax */ diff --git a/fiat-rust/src/poly1305_32.rs b/fiat-rust/src/poly1305_32.rs index 7ed22837dd..bf6ccf7af9 100644 --- a/fiat-rust/src/poly1305_32.rs +++ b/fiat-rust/src/poly1305_32.rs @@ -1,4 +1,4 @@ -//! Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --lang Rust --inline poly1305 32 '(auto)' '2^130 - 5' carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax +//! Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --lang Rust --inline poly1305 32 '(auto)' '2^130 - 5' --fancy-and-powerful-but-exponentially-slow-bounds-analysis carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax //! curve description: poly1305 //! machine_wordsize = 32 (from "32") //! requested operations: carry_mul, carry_square, carry, add, sub, opp, selectznz, to_bytes, from_bytes, relax diff --git a/fiat-rust/src/poly1305_64.rs b/fiat-rust/src/poly1305_64.rs index 81fe43b567..83ef96068f 100644 --- a/fiat-rust/src/poly1305_64.rs +++ b/fiat-rust/src/poly1305_64.rs @@ -1,4 +1,4 @@ -//! Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --lang Rust --inline poly1305 64 3 '2^130 - 5' carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax +//! Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --lang Rust --inline poly1305 64 3 '2^130 - 5' --fancy-and-powerful-but-exponentially-slow-bounds-analysis carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax //! curve description: poly1305 //! machine_wordsize = 64 (from "64") //! requested operations: carry_mul, carry_square, carry, add, sub, opp, selectznz, to_bytes, from_bytes, relax diff --git a/fiat-zig/src/poly1305_32.zig b/fiat-zig/src/poly1305_32.zig index 708df9e03a..78998e33d6 100644 --- a/fiat-zig/src/poly1305_32.zig +++ b/fiat-zig/src/poly1305_32.zig @@ -1,4 +1,4 @@ -// Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --lang Zig --internal-static --public-function-case camelCase --private-function-case camelCase --public-type-case UpperCamelCase --private-type-case UpperCamelCase --no-prefix-fiat --package-name poly1305 '' 32 '(auto)' '2^130 - 5' carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax +// Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --lang Zig --internal-static --public-function-case camelCase --private-function-case camelCase --public-type-case UpperCamelCase --private-type-case UpperCamelCase --no-prefix-fiat --package-name poly1305 '' 32 '(auto)' '2^130 - 5' --fancy-and-powerful-but-exponentially-slow-bounds-analysis carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax // curve description (via package name): poly1305 // machine_wordsize = 32 (from "32") // requested operations: carry_mul, carry_square, carry, add, sub, opp, selectznz, to_bytes, from_bytes, relax diff --git a/fiat-zig/src/poly1305_64.zig b/fiat-zig/src/poly1305_64.zig index fd058b74b8..350937d329 100644 --- a/fiat-zig/src/poly1305_64.zig +++ b/fiat-zig/src/poly1305_64.zig @@ -1,4 +1,4 @@ -// Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --lang Zig --internal-static --public-function-case camelCase --private-function-case camelCase --public-type-case UpperCamelCase --private-type-case UpperCamelCase --no-prefix-fiat --package-name poly1305 '' 64 3 '2^130 - 5' carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax +// Autogenerated: 'src/ExtractionOCaml/unsaturated_solinas' --lang Zig --internal-static --public-function-case camelCase --private-function-case camelCase --public-type-case UpperCamelCase --private-type-case UpperCamelCase --no-prefix-fiat --package-name poly1305 '' 64 3 '2^130 - 5' --fancy-and-powerful-but-exponentially-slow-bounds-analysis carry_mul carry_square carry add sub opp selectznz to_bytes from_bytes relax // curve description (via package name): poly1305 // machine_wordsize = 64 (from "64") // requested operations: carry_mul, carry_square, carry, add, sub, opp, selectznz, to_bytes, from_bytes, relax