@@ -25,14 +25,17 @@ Require Import Rewriter.Language.Inversion.
25
25
Require Import Crypto.Language .InversionExtra.
26
26
Require Import Rewriter.Language .Wf.
27
27
Require Import Rewriter.Language .UnderLetsProofs.
28
+ Require Import Rewriter.Util.Decidable.
28
29
Require Import Crypto.AbstractInterpretation.AbstractInterpretation.
30
+ Require Import Crypto.AbstractInterpretation.ZRangeCommonProofs.
29
31
Import Coq.Lists.List.
30
32
31
33
Import EqNotations.
32
34
Module Compilers.
33
35
Import Language .Compilers.
34
36
Import UnderLets.Compilers.
35
37
Import AbstractInterpretation.Compilers.
38
+ Import AbstractInterpretation.ZRangeCommonProofs.Compilers.
36
39
Import Language .Inversion .Compilers.
37
40
Import Language .InversionExtra.Compilers.
38
41
Import Language .Wf.Compilers.
@@ -852,39 +855,7 @@ Module Compilers.
852
855
853
856
Global Instance abstract_interp_ident_Proper {opts : AbstractInterpretation.Options } {assume_cast_truncates : bool} {t}
854
857
: Proper (eq ==> @abstract_domain_R t) (abstract_interp_ident assume_cast_truncates t).
855
- Proof using Type .
856
- cbv [abstract_interp_ident abstract_domain_R type.related respectful type.interp]; intros idc idc' ?; subst idc'; destruct idc;
857
- repeat first [ reflexivity
858
- | progress subst
859
- | progress cbn [ZRange.type.base.option.interp ZRange.type.base.interp base.interp base.base_interp Crypto.Util.Option.bind] in *
860
- | progress cbv [Crypto.Util.Option.bind]
861
- | intro
862
- | progress destruct_head'_prod
863
- | progress destruct_head'_bool
864
- | progress destruct_head' option
865
- | progress inversion_option
866
- | discriminate
867
- | solve [ eauto ]
868
- | apply NatUtil.nat_rect_Proper_nondep
869
- | apply ListUtil.list_rect_Proper
870
- | apply ListUtil.list_rect_arrow_Proper
871
- | apply ListUtil.list_case_Proper
872
- | apply ListUtil.pointwise_map
873
- | apply ListUtil.fold_right_Proper
874
- | apply ListUtil.update_nth_Proper
875
- | apply (@nat_rect_Proper_nondep_gen (_ -> _) (eq ==> eq)%signature)
876
- | cbn; apply (f_equal (@Some _))
877
- | progress cbn [ZRange.ident.option.interp]
878
- | progress cbv [zrange_rect]
879
- | apply (f_equal2 pair)
880
- | break_innermost_match_step
881
- | match goal with
882
- | [ H : _ |- _ ] => erewrite H by (eauto ; (eassumption || reflexivity))
883
- | [ H : forall x y, x = y -> _ |- _ ] => specialize (fun x => H x x eq_refl)
884
- | [ H : forall x, ?f x = ?g x, H1 : ?f ?y = _, H2 : ?g ?y = _ |- _ ]
885
- => specialize (H y); rewrite H1, H2 in H
886
- end ].
887
- Qed .
858
+ Proof using Type . apply ZRange.ident.option.interp_Proper. Qed .
888
859
889
860
Global Instance extract_list_state_Proper {t}
890
861
: Proper (abstract_domain'_R _ ==> option_eq (SetoidList.eqlistA (@abstract_domain'_R t)))
0 commit comments