forked from mit-plv/fiat-crypto
-
Notifications
You must be signed in to change notification settings - Fork 1
/
Copy pathSymbolic.v
4416 lines (4131 loc) · 195 KB
/
Symbolic.v
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
322
323
324
325
326
327
328
329
330
331
332
333
334
335
336
337
338
339
340
341
342
343
344
345
346
347
348
349
350
351
352
353
354
355
356
357
358
359
360
361
362
363
364
365
366
367
368
369
370
371
372
373
374
375
376
377
378
379
380
381
382
383
384
385
386
387
388
389
390
391
392
393
394
395
396
397
398
399
400
401
402
403
404
405
406
407
408
409
410
411
412
413
414
415
416
417
418
419
420
421
422
423
424
425
426
427
428
429
430
431
432
433
434
435
436
437
438
439
440
441
442
443
444
445
446
447
448
449
450
451
452
453
454
455
456
457
458
459
460
461
462
463
464
465
466
467
468
469
470
471
472
473
474
475
476
477
478
479
480
481
482
483
484
485
486
487
488
489
490
491
492
493
494
495
496
497
498
499
500
501
502
503
504
505
506
507
508
509
510
511
512
513
514
515
516
517
518
519
520
521
522
523
524
525
526
527
528
529
530
531
532
533
534
535
536
537
538
539
540
541
542
543
544
545
546
547
548
549
550
551
552
553
554
555
556
557
558
559
560
561
562
563
564
565
566
567
568
569
570
571
572
573
574
575
576
577
578
579
580
581
582
583
584
585
586
587
588
589
590
591
592
593
594
595
596
597
598
599
600
601
602
603
604
605
606
607
608
609
610
611
612
613
614
615
616
617
618
619
620
621
622
623
624
625
626
627
628
629
630
631
632
633
634
635
636
637
638
639
640
641
642
643
644
645
646
647
648
649
650
651
652
653
654
655
656
657
658
659
660
661
662
663
664
665
666
667
668
669
670
671
672
673
674
675
676
677
678
679
680
681
682
683
684
685
686
687
688
689
690
691
692
693
694
695
696
697
698
699
700
701
702
703
704
705
706
707
708
709
710
711
712
713
714
715
716
717
718
719
720
721
722
723
724
725
726
727
728
729
730
731
732
733
734
735
736
737
738
739
740
741
742
743
744
745
746
747
748
749
750
751
752
753
754
755
756
757
758
759
760
761
762
763
764
765
766
767
768
769
770
771
772
773
774
775
776
777
778
779
780
781
782
783
784
785
786
787
788
789
790
791
792
793
794
795
796
797
798
799
800
801
802
803
804
805
806
807
808
809
810
811
812
813
814
815
816
817
818
819
820
821
822
823
824
825
826
827
828
829
830
831
832
833
834
835
836
837
838
839
840
841
842
843
844
845
846
847
848
849
850
851
852
853
854
855
856
857
858
859
860
861
862
863
864
865
866
867
868
869
870
871
872
873
874
875
876
877
878
879
880
881
882
883
884
885
886
887
888
889
890
891
892
893
894
895
896
897
898
899
900
901
902
903
904
905
906
907
908
909
910
911
912
913
914
915
916
917
918
919
920
921
922
923
924
925
926
927
928
929
930
931
932
933
934
935
936
937
938
939
940
941
942
943
944
945
946
947
948
949
950
951
952
953
954
955
956
957
958
959
960
961
962
963
964
965
966
967
968
969
970
971
972
973
974
975
976
977
978
979
980
981
982
983
984
985
986
987
988
989
990
991
992
993
994
995
996
997
998
999
1000
Require Crypto.Assembly.Parse.
From Coq.Program Require Import Tactics.
From Coq Require Import Derive.
From Coq Require Import List.
From Coq Require Import Lia.
From Coq Require Import ZArith.
From Coq Require Import NArith.
From Coq Require Import Permutation.
From Coq Require Import Equalities.
From Coq Require Import OrderedType.
From Coq Require Import Orders.
From Coq Require Import FMapInterface.
From Coq Require Import FMapPositive.
From Coq Require Import FMapFacts.
Require Crypto.Util.ZRange.
Require Crypto.Util.Tuple.
Require Import Util.OptionList.
Require Import Crypto.Util.ErrorT.
Require Import Crypto.Util.ZUtil.Tactics.PullPush.Modulo.
Require Import Crypto.Util.ZUtil.Tactics.LtbToLt.
Require Import Crypto.Util.ZUtil.Testbit.
Require Import Crypto.Util.ZUtil.Hints.ZArith.
Require Import Crypto.Util.ZUtil.Land.
Require Import Crypto.Util.ZUtil.Land.Fold.
Require Import Crypto.Util.ZUtil.Ones.
Require Import Crypto.Util.ZUtil.Definitions.
Require Import Crypto.Util.ZUtil.Signed.
Require Import Crypto.Util.Equality.
Require Import Crypto.Util.Bool.Reflect.
Require Import Crypto.Util.Option.
Require Import Crypto.Util.Strings.Subscript.
Require Import Crypto.Util.Structures.Orders.
Require Import Crypto.Util.Structures.Equalities.
Require Import Crypto.Util.Structures.Equalities.Iso.
Require Import Crypto.Util.Structures.Equalities.Prod.
Require Import Crypto.Util.Structures.Equalities.Option.
Require Import Crypto.Util.Structures.Orders.Iso.
Require Import Crypto.Util.Structures.Orders.Prod.
Require Import Crypto.Util.Structures.Orders.Option.
Require Import Crypto.Util.Structures.OrdersEx.
Require Import Crypto.Util.ListUtil.
Require Import Crypto.Util.ListUtil.Concat.
Require Import Crypto.Util.ListUtil.GroupAllBy.
Require Import Crypto.Util.ListUtil.FoldMap. Import FoldMap.List.
Require Import Crypto.Util.ListUtil.IndexOf. Import IndexOf.List.
Require Import Crypto.Util.ListUtil.Forall.
Require Import Crypto.Util.ListUtil.Permutation.
Require Import Crypto.Util.ListUtil.Partition.
Require Import Crypto.Util.ListUtil.Filter.
Require Import Crypto.Util.ListUtil.PermutationCompat. Import ListUtil.PermutationCompat.Coq.Sorting.Permutation.
Require Import Crypto.Util.NUtil.Sorting.
Require Import Crypto.Util.NUtil.Testbit.
Require Import Crypto.Util.FSets.FMapOption.
Require Import Crypto.Util.FSets.FMapN.
Require Import Crypto.Util.FSets.FMapZ.
Require Import Crypto.Util.FSets.FMapProd.
Require Import Crypto.Util.FSets.FMapIso.
Require Import Crypto.Util.FSets.FMapSect.
Require Import Crypto.Util.FSets.FMapInterface.
Require Import Crypto.Util.FSets.FMapFacts.
Require Import Crypto.Util.FSets.FMapTrieEx.
Require Import Crypto.Util.MSets.MSetN.
Require Import Crypto.Util.ListUtil.PermutationCompat.
Require Import Crypto.Util.Bool.LeCompat.
Require Import Crypto.Util.Tactics.DestructHead.
Require Import Crypto.Util.Tactics.SetEvars.
Require Import Crypto.Util.Prod.
Require Import Crypto.Util.Tactics.SplitInContext.
Require Import Crypto.Util.Tactics.SpecializeBy.
Require Import Crypto.Util.Tactics.SpecializeAllWays.
Require Import Crypto.Util.Tactics.SpecializeUnderBindersBy.
Require Import Crypto.Util.Tactics.InHypUnderBindersDo.
Require Import Crypto.Util.Tactics.Head.
Require Import Crypto.Util.Tactics.UniquePose.
Require Import Crypto.Util.ZUtil.Lxor.
Require Import Crypto.Util.ZUtil.Tactics.PeelLe.
Require Import Crypto.Util.ZUtil.Tactics.RewriteModSmall.
Require Import Crypto.Util.Listable.
Require Import Crypto.Util.Strings.Parse.Common.
Require Import Crypto.Util.Tactics.WarnIfGoalsRemain.
Require Import Crypto.Util.Bool.Reflect.
Require Import coqutil.Z.bitblast.
From Coq Require Import String.
Require Import Crypto.Util.Strings.Show.
Require Import Crypto.Util.Strings.Show.Enum.
Require Import Crypto.Assembly.Syntax.
Import ListNotations.
Definition idx := N.
Local Set Decidable Equality Schemes.
Definition symbol := N.
Class OperationSize := operation_size : N.
Global Instance Show_OperationSize : Show OperationSize := show_N.
Section S.
Implicit Type s : OperationSize.
Variant op := old s (_:symbol) | const (_ : Z) | add s | addcarry s | subborrow s | addoverflow s | neg s | shl s | shr s | sar s | rcr s | and s | or s | xor s | slice (lo sz : N) | mul s | set_slice (lo sz : N) | selectznz | iszero (* | ... *)
| addZ | mulZ | negZ | shlZ | shrZ | andZ | orZ | xorZ | addcarryZ s | subborrowZ s.
Definition op_beq a b := if op_eq_dec a b then true else false.
End S.
Global Instance Show_op : Show op := fun o =>
match o with
| old s n => "old " ++ show s ++ " " ++ show n
| const n => "const " ++ show n
| add s => "add " ++ show s
| addcarry s => "addcarry " ++ show s
| subborrow s => "subborrow " ++ show s
| addoverflow s => "addoverflow " ++ show s
| neg s => "neg " ++ show s
| shl s => "shl " ++ show s
| shr s => "shr " ++ show s
| sar s => "sar " ++ show s
| rcr s => "rcr " ++ show s
| and s => "and " ++ show s
| or s => "or " ++ show s
| xor s => "xor " ++ show s
| slice lo sz => "slice " ++ show lo ++ " " ++ show sz
| mul s => "mul " ++ show s
| set_slice lo sz => "set_slice " ++ show lo ++ " " ++ show sz
| selectznz => "selectznz"
| iszero => "iszero"
| addZ => "addZ"
| mulZ => "mulZ"
| negZ => "negZ"
| shlZ => "shlZ"
| shrZ => "shrZ"
| andZ => "andZ"
| orZ => "orZ"
| xorZ => "xorZ"
| addcarryZ s => "addcarryZ " ++ show s
| subborrowZ s => "subborrowZ " ++ show s
end%string.
Definition show_op_subscript : Show op := fun o =>
match o with
| old s n => "old" ++ String.to_subscript (show s) ++ " " ++ show n
| const n => "const " ++ show n
| add s => "add" ++ String.to_subscript (show s)
| addcarry s => "addcarry" ++ String.to_subscript (show s)
| subborrow s => "subborrow" ++ String.to_subscript (show s)
| addoverflow s => "addoverflow" ++ String.to_subscript (show s)
| neg s => "neg" ++ String.to_subscript (show s)
| shl s => "shl" ++ String.to_subscript (show s)
| shr s => "shr" ++ String.to_subscript (show s)
| sar s => "sar" ++ String.to_subscript (show s)
| rcr s => "rcr" ++ String.to_subscript (show s)
| and s => "and" ++ String.to_subscript (show s)
| or s => "or" ++ String.to_subscript (show s)
| xor s => "xor" ++ String.to_subscript (show s)
| slice lo sz => "slice" ++ String.to_subscript (show lo) ++ "," ++ String.to_subscript (show sz)
| mul s => "mul" ++ String.to_subscript (show s)
| set_slice lo sz => "set_slice" ++ String.to_subscript (show lo) ++ "," ++ String.to_subscript (show sz)
| selectznz => "selectznz"
| iszero => "iszero"
| addZ => "addℤ"
| mulZ => "mulℤ"
| negZ => "negℤ"
| shlZ => "shlℤ"
| shrZ => "shrℤ"
| andZ => "andℤ"
| orZ => "orℤ"
| xorZ => "xorℤ"
| addcarryZ s => "addcarryℤ" ++ String.to_subscript (show s)
| subborrowZ s => "subborrowℤ" ++ String.to_subscript (show s)
end%string.
Module FMapOp.
Definition op_args : Set := (option OperationSize * (option symbol * (option Z * (option N * option N)))).
Definition op' : Set := N * op_args.
Definition eta_op_cps {T} (k : op -> T) (o : op) : T.
Proof.
pose o as o'.
destruct o.
all: let v := (eval cbv [o'] in o') in
exact (k v).
Defined.
Definition nat_of_op (o : op) : nat.
Proof.
evar (base : nat).
destruct o.
all: let rec peel_S val :=
lazymatch val with
| S ?val => apply S; peel_S val
| ?ev => is_evar ev;
let __ := open_constr:(eq_refl : ev = S _) in
exact O
end in
let v := (eval cbv [base] in base) in
peel_S v.
Unshelve.
exact O.
Defined.
Definition args_of_op (o : op) : op_args.
Proof.
destruct o; hnf.
all: repeat lazymatch reverse goal with
| [ H : ?T |- ?A * ?B ]
=> lazymatch A with
| context[option T]
=> split; [ | clear H ]
| _ => split; [ repeat apply pair; exact None | ]
end
| [ H : ?T |- _ ]
=> lazymatch goal with
| [ |- option T ] => exact (Some H)
| [ |- _ ] => idtac "bad"
end
| [ |- _ * _ ] => split
| [ |- option _ ] => exact None
end.
Defined.
Definition op'_of_op : op -> op'
:= Eval compute in eta_op_cps (fun o => (N.of_nat (nat_of_op o), args_of_op o)).
Derive op_of_op'_opt
SuchThat ((forall (o : op), op_of_op'_opt (op'_of_op o) = Some o)
/\ (forall (n n' : op'), option_map op'_of_op (op_of_op'_opt n) = Some n' -> n = n'))
As op_op'_correct.
Proof.
instantiate (1:=ltac:(intros [n v]; hnf in v; destruct_head'_prod; destruct n as [|n])) in (value of op_of_op'_opt).
subst op_of_op'_opt.
split.
{ intro o; destruct o; cbv [op'_of_op].
all: [ > cbv beta iota;
lazymatch goal with
| [ |- ?ev = Some _ ]
=> is_evar ev;
let H := fresh in
pose ev as H;
instantiate (1:=ltac:(lazymatch goal with
| [ n : positive |- _ ]
=> destruct n
| _ => idtac
end)) in (value of H);
subst H; cbv beta iota
end .. ].
all: lazymatch goal with
| [ |- ?ev = Some ?v ]
=> is_evar ev;
let h := head v in
let H := fresh in
pose ev as H;
instantiate (1:=ltac:(reverse)) in (value of H);
subst H;
repeat match goal with
| [ |- context[?ev ?x] ]
=> is_evar ev;
let H := fresh in
pose ev as H;
lazymatch x with
| Some _
=> instantiate (1:=ltac:(let x := fresh "x" in intro x; intros; destruct x; [ reverse | exact None ])) in (value of H)
| None
=> instantiate (1:=ltac:(let x := fresh "x" in intro x; intros; destruct x; [ exact None | reverse ])) in (value of H)
| ?x
=> tryif is_var x
then instantiate (1:=ltac:(intro)) in (value of H)
else idtac "unknown" x
end;
subst H; cbv beta iota
end;
reflexivity
end. }
{ intros [n nargs] [n' nargs'].
set (f := option_map _).
break_innermost_match.
all: let G := lazymatch goal with |- ?G => G end in
tryif has_evar G then instantiate (1:=None) else idtac.
all: vm_compute.
all: lazymatch goal with
| [ |- Some _ = Some _ -> _ ] => intro; inversion_option; inversion_pair; subst; try reflexivity
| [ |- None = Some _ -> _ ] => let H := fresh in intro H; exfalso; clear -H; inversion_option
end. }
Qed.
Module OptionNMap <: S := OptionUsualMap NMap.
Module OptionZMap <: S := OptionUsualMap ZMap.
Module OptionSymbolMap <: S := OptionNMap.
Module OptionOperationSizeMap <: S := OptionNMap.
Module OpArgsMap0 <: S := OptionNMap.
Module OpArgsMap1 <: S := ProdUsualMap OptionNMap OpArgsMap0.
Module OpArgsMap2 <: S := ProdUsualMap OptionZMap OpArgsMap1.
Module OpArgsMap3 <: S := ProdUsualMap OptionSymbolMap OpArgsMap2.
Module OpArgsMap4 <: S := ProdUsualMap OptionOperationSizeMap OpArgsMap3.
Module OpArgsMap <: S := OpArgsMap4.
Module OpMap' <: S := ProdUsualMap NMap OpArgsMap.
Module OpSectOp' <: SectMiniOrderedType OpMap'.E.
Definition t := op.
Include HasUsualEq.
Include UsualIsEq.
Include UsualIsEqOrig.
Definition to_ : t -> OpMap'.E.t := Eval vm_compute in op'_of_op.
Definition of_ (v : OpMap'.E.t) : t
:= Eval vm_compute in match op_of_op'_opt v with
| Some v => v
| None => old 0%N 0%N
end.
Global Instance Proper_to_ : Proper (Logic.eq ==> Logic.eq) to_ | 5.
Proof. repeat intro; subst; reflexivity. Qed.
Global Instance Proper_of_ : Proper (Logic.eq ==> Logic.eq) of_ | 5.
Proof. repeat intro; subst; reflexivity. Qed.
Lemma of_to : forall x, eq (of_ (to_ x)) x.
Proof.
intro x.
refine (_ : match op_of_op'_opt (op'_of_op x) with
| Some v => v
| None => _
end = x).
rewrite (proj1 op_op'_correct).
reflexivity.
Qed.
Include LiftIsoHasLt OpMap'.E.
Include LiftSectHasMiniOrderedType OpMap'.E.
Include LiftSectIsLt OpMap'.E.
End OpSectOp'.
Module OpMap <: UsualS := SectS OpMap' OpSectOp'.
End FMapOp.
Module OpMap := FMapOp.OpMap.
Module Export RewritePass.
Variant rewrite_pass :=
| addbyte_small
| addcarry_bit
| addcarry_small
| addoverflow_bit
| addoverflow_small
| combine_consts
| constprop
| consts_commutative
| drop_identity
| flatten_associative
| flatten_bounded_associative
| fold_consts_to_and
| set_slice0_small
| set_slice_set_slice
| shift_to_mul
| slice0
| slice01_addcarryZ
| slice01_subborrowZ
| slice_set_slice
| truncate_small
| unary_truncate
| xor_same
.
Definition rewrite_pass_beq (x y : rewrite_pass) : bool := if rewrite_pass_eq_dec x y then true else false.
Derive rewrite_pass_Listable SuchThat (@FinitelyListable rewrite_pass rewrite_pass_Listable) As rewrite_pass_FinitelyListable.
Proof. prove_ListableDerive. Qed.
Global Existing Instances rewrite_pass_Listable rewrite_pass_FinitelyListable.
Global Instance show_rewrite_pass : Show rewrite_pass.
Proof. prove_Show_enum (). Defined.
Global Instance show_lvl_rewrite_pass : ShowLevel rewrite_pass := show_rewrite_pass.
Definition parse_rewrite_pass_list : list (string * rewrite_pass)
:= Eval vm_compute in
List.map
(fun r => (show r, r))
(list_all rewrite_pass).
Definition parse_rewrite_pass : ParserAction rewrite_pass
:= parse_strs parse_rewrite_pass_list.
Definition default_rewrite_pass_order : list rewrite_pass
:= [constprop
;slice0
;slice01_addcarryZ
;slice01_subborrowZ
;set_slice_set_slice
;slice_set_slice
;set_slice0_small
;shift_to_mul
;flatten_associative
;consts_commutative
;fold_consts_to_and
;drop_identity
;flatten_bounded_associative
;unary_truncate
;truncate_small
;combine_consts
;addoverflow_bit
;addcarry_bit
;addcarry_small
;addoverflow_small
;addbyte_small
;xor_same
].
End RewritePass.
Module Export Options.
(* Holds the order and multiplicity of passes *)
Class rewriting_pipeline_opt := rewriting_pipeline : list rewrite_pass.
(* Holds the passes that are enabled *)
Class rewriting_pass_filter_opt := rewriting_pass_filter : rewrite_pass -> bool.
(* Holds the actual list of passes that are used; we make this a
separate option so that it is computed once and for all rather
than every time, because it is (currently) quadratic to compute
in the number of passes *)
Class rewriting_passes_opt := rewriting_passes : list rewrite_pass.
(** Should we symex the assembly first, even though this may be more inefficient? *)
Class debug_symex_asm_first_opt := debug_symex_asm_first : bool.
Definition default_rewriting_passes
{rewriting_pipeline : rewriting_pipeline_opt}
{rewriting_pass_filter : rewriting_pass_filter_opt}
: rewriting_passes_opt
:= List.filter rewriting_pass_filter rewriting_pipeline.
Class symbolic_options_opt :=
{ asm_rewriting_pipeline : rewriting_pipeline_opt
; asm_rewriting_pass_filter : rewriting_pass_filter_opt
; asm_debug_symex_asm_first : debug_symex_asm_first_opt
}.
(* This holds the list of computed options, which are passed around between methods *)
Class symbolic_options_computed_opt :=
{ asm_rewriting_passes : rewriting_passes_opt
; asm_debug_symex_asm_first_computed : debug_symex_asm_first_opt
}.
(* N.B. The default rewriting pass filter should not be changed here, but instead changed in CLI.v where it is derived from a default string *)
Definition default_symbolic_options : symbolic_options_opt
:= {| asm_rewriting_pipeline := default_rewrite_pass_order
; asm_rewriting_pass_filter := fun _ => true
; asm_debug_symex_asm_first := false
|}.
End Options.
Module Export Hints.
Global Existing Instance default_rewriting_passes.
Global Existing Instances
Build_symbolic_options_opt
Build_symbolic_options_computed_opt
asm_rewriting_pipeline
asm_rewriting_pass_filter
asm_rewriting_passes
asm_debug_symex_asm_first
asm_debug_symex_asm_first_computed
.
#[global]
Hint Cut [
( _ * )
(asm_rewriting_pipeline
| asm_rewriting_pass_filter
| asm_rewriting_passes
| asm_debug_symex_asm_first
| asm_debug_symex_asm_first_computed
) ( _ * )
(Build_symbolic_options_opt
| Build_symbolic_options_computed_opt
)
] : typeclass_instances.
End Hints.
Definition associative o := match o with add _|addZ|mul _|mulZ|or _|and _|xor _|andZ|orZ|xorZ=> true | _ => false end.
Definition commutative o := match o with add _|addZ|addcarry _|addoverflow _|mul _|mulZ|or _|and _|xor _|andZ|orZ|xorZ => true | _ => false end.
Definition identity o := match o with mul s => Some match s with N0 => 0%Z | _ => 1%Z end| mulZ=>Some 1%Z |add _|addZ|or _|orZ|xor _|xorZ|addcarry _|addcarryZ _|addoverflow _ => Some 0%Z | and s => Some (Z.ones (Z.of_N s))|andZ => Some (-1)%Z |_=> None end.
(* identity, but not in the first slot *)
Definition identity_after_0 o := match o with subborrow _|subborrowZ _ => Some 0%Z | _=> None end.
Definition unary_truncate_size o := match o with add s|and s|or s|xor s|mul s => Some (Z.of_N s) | addZ|mulZ|andZ|orZ|xorZ => Some (-1)%Z | _ => None end.
Definition op_always_interps o := match o with add _|addcarry _|addoverflow _|and _|or _|xor _|mul _|addZ|mulZ|andZ|orZ|xorZ|addcarryZ _ => true | _ => false end.
Definition combines_to o := match o with add s => Some (mul s) | addZ => Some mulZ | _ => None end.
(* the inner operation can be dropped when replacing it with another operation results in a bounded output *)
Variant drop_kind := drop_always | drop_if_bounded (bound : ZRange.zrange) (rep_op : op).
Definition bounds_for_drop_inner_associative o_outer o_inner :=
match o_outer with addcarryZ s|addcarry s(*|addoverflow s*) =>
match o_inner with
| addZ => Some drop_always
| add _ => (s' <- unary_truncate_size o_inner; Some (drop_if_bounded (ZRange.Build_zrange 0 (Z.ones s')) addZ))
| _ => None end | _ => None end%option.
Definition node (A : Set) : Set := op * list A.
Global Instance Show_node {A : Set} {show_A : Show A} : Show (node A) := show_prod.
Local Unset Elimination Schemes.
Inductive expr : Set :=
| ExprRef (_ : idx)
| ExprApp (_ : node expr).
Local Set Elimination Schemes.
Section expr_ind.
Context (P : expr -> Prop)
(HRef : forall i, P (ExprRef i))
(HApp : forall n, Forall P (snd n) -> P (ExprApp n)).
Fixpoint expr_ind e {struct e} : P e :=
match e with
| ExprRef i => HRef i
| ExprApp n => HApp _ (list_rect _ (Forall_nil _) (fun e _ H => Forall_cons e (expr_ind e) H) (snd n))
end.
End expr_ind.
Definition invert_ExprRef (e : expr) : option idx :=
match e with ExprRef i => Some i | _ => None end.
Definition Show_expr_body (Show_expr : Show expr) : Show expr
:= Eval cbv -[String.append show_N concat List.map Show_op] in
fun e => match e with
| ExprRef i => "ExprRef " ++ show i
| ExprApp (o, e) => "ExprApp " ++ show (o, e)
end%string.
Definition Show_expr : Show expr
:= Eval cbv -[String.append show_N concat List.map Show_op] in
fix Show_expr e := Show_expr_body Show_expr e.
Global Existing Instance Show_expr.
Local Notation max_powers_of_two := 5%nat (only parsing).
Local Notation max_decimal := 256%Z (only parsing).
Definition show_infix_op (o : op) : option string
:= match o with
| add s => Some ("+" ++ String.to_subscript (show s))
| shl s => Some (">>" ++ String.to_subscript (show s))
| shr s => Some (">>" ++ String.to_subscript (show s))
| and s => Some ("&" ++ String.to_subscript (show s))
| or s => Some ("|" ++ String.to_subscript (show s))
| xor s => Some ("^" ++ String.to_subscript (show s))
| mul s => Some ("*" ++ String.to_subscript (show s))
| sar s => Some (">>>" ++ String.to_subscript (show s))
| addZ => Some "+ℤ"
| mulZ => Some "*ℤ"
| shlZ => Some "<<ℤ"
| shrZ => Some ">>ℤ"
| andZ => Some "&ℤ"
| orZ => Some "|ℤ"
| xorZ => Some "^ℤ"
| _ => None
end%string.
Definition show_prefix_op (o : op) : option (string * Level)
:= match o with
| neg s => Some ("-" ++ String.to_subscript (show s), opp_lvl)
| negZ => Some ("-ℤ", opp_lvl)
| _ => None
end%string.
Definition show_lvl_expr_pretty : ShowLevel expr
:= fix show_lvl_pretty_expr (e : expr) : Level -> string
:= let __ : ShowLevel expr := show_lvl_pretty_expr in
let __ : Show expr := @Show_of_ShowLevel _ show_lvl_pretty_expr in
let show_comment_args args
:= match args with
| nil => ""
| _ => " (* " ++ show args ++ " *)"
end%string in
match e with
| ExprRef i => fun _ => "#" ++ show i
| ExprApp (old s x, args) => lvl_wrap_parens app_lvl ("old" ++ String.to_subscript (show s) ++ " " ++ show x ++ show_comment_args args)
| ExprApp (const x, args) => fun lvl => PowersOfTwo.show_lvl_Z_up_to max_powers_of_two max_decimal x lvl ++ show_comment_args args
| ExprApp ((add _|shl _|shr _|and _|or _|xor _|mul _|sar _|addZ|mulZ|shlZ|shrZ|andZ|orZ|xorZ) as o, args)
=> let o : string := invert_Some (show_infix_op o) in
match args with
| nil => fun _ => o ++ "[]"
| x :: nil => fun _ => o ++ "[" ++ show x ++ "]"
| _ => fun _ => "(" ++ String.concat (" " ++ o ++ " ") (List.map show args) ++ ")"
end
| ExprApp ((neg _|negZ) as o, args)
=> let '(o, lvl) := invert_Some (show_prefix_op o) in
match args with
| nil => fun _ => o ++ "[]"
| x :: nil => fun _ => o ++ show_lvl x lvl
| _ => fun _ => o ++ show args
end
| ExprApp (o, args)
=> fun _ => "(" ++ show_op_subscript o ++ ", " ++ show args ++ ")"
end%string%list.
Definition show_expr_pretty : Show expr
:= @Show_of_ShowLevel _ show_lvl_expr_pretty.
Lemma op_beq_spec a b : BoolSpec (a=b) (a<>b) (op_beq a b).
Proof using Type. cbv [op_beq]; destruct (op_eq_dec a b); constructor; congruence. Qed.
Global Instance reflect_eq_op : reflect_rel eq op_beq | 10 := reflect_rel_of_BoolSpec op_beq_spec.
Fixpoint expr_beq (X Y : expr) {struct X} : bool :=
match X, Y with
| ExprRef x, ExprRef x0 => N.eqb x x0
| ExprApp x, ExprApp x0 =>
Prod.prod_beq _ _ op_beq (ListUtil.list_beq expr expr_beq) x x0
| _, _ => false
end.
Lemma expr_beq_spec a b : BoolSpec (a=b) (a<>b) (expr_beq a b).
Proof using Type.
revert b; induction a, b; cbn.
1: destruct (N.eqb_spec i i0); constructor; congruence.
1,2: constructor; congruence.
destruct n, n0; cbn.
destruct (op_beq_spec o o0); cbn in *; [subst|constructor; congruence].
revert l0; induction H, l0; cbn; try (constructor; congruence); [].
destruct (H e); cbn; try (constructor; congruence); []; subst.
destruct (IHForall l0); [left|right]; congruence.
Qed.
Global Instance reflect_eq_expr : reflect_rel eq expr_beq | 10 := reflect_rel_of_BoolSpec expr_beq_spec.
Lemma expr_beq_true a b : expr_beq a b = true -> a = b.
Proof using Type. destruct (expr_beq_spec a b); congruence. Qed.
Require Import Crypto.Util.Option Crypto.Util.Notations Coq.Lists.List.
Import ListNotations.
Section WithContext.
Context (ctx : symbol -> option Z).
Definition op_to_Z_binop (o : op) : option _
:= match o with
| add _ => Some Z.add
| shl _ => Some Z.shiftl
| shr _ => Some Z.shiftr
| and _ => Some Z.land
| or _ => Some Z.lor
| xor _ => Some Z.lxor
| mul _ => Some Z.mul
| addZ => Some Z.add
| mulZ => Some Z.mul
| shlZ => Some Z.shiftl
| shrZ => Some Z.shiftr
| andZ => Some Z.land
| orZ => Some Z.lor
| xorZ => Some Z.lxor
| _ => None
end.
Definition interp_op o (args : list Z) : option Z :=
Eval cbv [invert_Some identity op_to_Z_binop] in
let keep n x := Z.land x (Z.ones (Z.of_N n)) in
match o, args with
| old s x, nil => match ctx x with Some v => Some (keep s v) | None => None end
| const z, nil => Some z
| addcarry s, args =>
Some (Z.shiftr (List.fold_right Z.add 0 args) (Z.of_N s) mod 2)
| subborrow s, cons a args' =>
Some ((- Z.shiftr (a - List.fold_right Z.add 0 args') (Z.of_N s)) mod 2)
| addoverflow s, args => Some (Z.b2z (negb (Z.eqb
(Z.signed s (keep s (List.fold_right Z.add 0 args)))
(List.fold_right Z.add 0%Z (List.map (Z.signed s) args)))))
| neg s, [a] => Some (keep s (- a))
| shl s, [a; b] => Some (keep s (Z.shiftl a b))
| shr s, [a; b] => Some (keep s (Z.shiftr a b))
| sar s, [a; b] => Some (keep s (Z.shiftr (Z.signed s a) b))
| rcr s, [v1; cf; cnt] => Some (
let v1c := Z.lor v1 (Z.shiftl cf (Z.of_N s)) in
let l := Z.lor v1c (Z.shiftl v1 (1+Z.of_N s)) in
keep s (Z.shiftr l cnt))
| (add s|and s|or s|xor s|mul s) as o, args
=> let id := invert_Some (identity o) in
let o := invert_Some (op_to_Z_binop o) in
Some (keep s (List.fold_right o id args))
| (addZ|mulZ|andZ|orZ|xorZ) as o, args
=> let id := invert_Some (identity o) in
let o := invert_Some (op_to_Z_binop o) in
Some (List.fold_right o id args)
| slice lo sz, [a] => Some (keep sz (Z.shiftr a (Z.of_N lo)))
| set_slice lo sz, [a; b] =>
Some (Z.lor (Z.shiftl (keep sz b) (Z.of_N lo))
(Z.ldiff a (Z.shiftl (Z.ones (Z.of_N sz)) (Z.of_N lo))))
| selectznz, [c; a; b] => Some (if Z.eqb c 0 then a else b)
| iszero, [a] => Some (Z.b2z (Z.eqb a 0))
| negZ, [a] => Some (Z.opp a)
| shlZ, [a; b] => Some (Z.shiftl a b)
| shrZ, [a; b] => Some (Z.shiftr a b)
| addcarryZ s, args => Some (Z.shiftr (List.fold_right Z.add 0 args) (Z.of_N s))
| subborrowZ s, cons a args' => Some (- Z.shiftr (a - List.fold_right Z.add 0 args') (Z.of_N s))
| _, _ => None
end%Z.
End WithContext.
Definition interp0_op := interp_op (fun _ => None).
Lemma interp_op_weaken_symbols G1 G2 o args
(H : forall (s:symbol) v, G1 s = Some v -> G2 s = Some v)
: forall v, interp_op G1 o args = Some v -> interp_op G2 o args = Some v.
Proof using Type.
cbv [interp_op option_map]; intros;
repeat (BreakMatch.break_match || BreakMatch.break_match_hyps);
inversion_option; subst;
try congruence.
all : eapply H in Heqo0; congruence.
Qed.
Lemma interp_op_interp0_op o a v (H : interp0_op o a = Some v)
: forall G, interp_op G o a = Some v.
Proof using Type. intros; eapply interp_op_weaken_symbols in H; try eassumption; discriminate. Qed.
Definition node_beq {A : Set} (arg_eqb : A -> A -> bool) : node A -> node A -> bool :=
Prod.prod_beq _ _ op_beq (ListUtil.list_beq _ arg_eqb).
Global Instance reflect_node_beq {A : Set} {arg_eqb} {H : reflect_rel (@eq A) arg_eqb}
: reflect_rel eq (@node_beq A arg_eqb) | 10 := _.
Local Open Scope Z_scope.
Require Crypto.Language.API.
Require Crypto.AbstractInterpretation.ZRange.
Require Crypto.AbstractInterpretation.ZRangeProofs.
Require Import Crypto.CastLemmas.
Require Import Crypto.Language.PreExtra.
Require Import Crypto.Util.ZRange.LandLorBounds.
Require Import Crypto.Util.ZRange.BasicLemmas.
Require Import Crypto.Util.ZRange.Operations.
Require Import Crypto.Util.ZUtil.LandLorShiftBounds.
Require Import Crypto.Util.ZUtil.LandLorBounds.
Section bound_node_via_PHOAS.
Import Crypto.Util.ZRange.
Import API.Compilers.
Import AbstractInterpretation.ZRange.Compilers.
Import AbstractInterpretation.ZRangeProofs.Compilers.
Local Instance : shiftr_avoid_uint1_opt := false.
Local Notation interp_PHOAS_op := (ZRange.ident.option.interp true).
Definition op_to_PHOAS_binop (o : op) : option _
:= match o with
| add _ => Some ident.Z_add
| shl _ => Some ident.Z_shiftl
| shr _ => Some ident.Z_shiftr
| and _ => Some ident.Z_land
| or _ => Some ident.Z_lor
| xor _ => Some ident.Z_lxor
| mul _ => Some ident.Z_mul
| addZ => Some ident.Z_add
| mulZ => Some ident.Z_mul
| shlZ => Some ident.Z_shiftl
| shrZ => Some ident.Z_shiftr
| andZ => Some ident.Z_land
| orZ => Some ident.Z_lor
| xorZ => Some ident.Z_lxor
| _ => None
end.
Definition op_to_PHOAS_unop (o : op) : option _
:= match o with
| neg _ => Some ident.Z_opp
| negZ => Some ident.Z_opp
| _ => None
end.
Definition op_to_PHOAS_tritop (o : op) : option _
:= match o with
| selectznz => Some ident.Z_zselect
| _ => None
end.
Definition op_to_bounds (o : op) : option _
:= match o with
| add s
| shl s
| shr s
| and s
| or s
| xor s
| mul s
| neg s
=> Some r[0~>Z.ones (Z.of_N s)]
| _ => None
end%zrange.
Definition op_to_PHOAS_bounds (o : op) : option (list (unit -> option zrange) -> option zrange)
:= let unthunk := List.map (fun x => x tt) in
match o with
| old s _ => Some (fun _ => Some r[0~>Z.ones (Z.of_N s)])
| const x
=> Some (fun _ => interp_PHOAS_op (ident.Literal (t:=base.type.Z) x))
| (add _|and _|or _|xor _|mul _) as o
=> let b : zrange := invert_Some (op_to_bounds o) in
let id : Z := invert_Some (identity o) in
let o := interp_PHOAS_op (invert_Some (op_to_PHOAS_binop o)) in
Some (fun args => interp_PHOAS_op ident.Z_cast (Some b) (fold_right o (Some r[id~>id]) (unthunk args)))
| (addcarry _ | subborrow _ | addoverflow _ | iszero)
=> Some (fun _ => Some r[0~>1])
| (addZ|mulZ|andZ|orZ|xorZ) as o
=> let id : Z := invert_Some (identity o) in
let o := interp_PHOAS_op (invert_Some (op_to_PHOAS_binop o)) in
Some (fun args => fold_right o (Some r[id~>id]) (unthunk args))
| (shl _|shr _) as o
=> let b : zrange := invert_Some (op_to_bounds o) in
let o := interp_PHOAS_op (invert_Some (op_to_PHOAS_binop o)) in
Some (fun args
=> match args with
| [x; y] => interp_PHOAS_op ident.Z_cast (Some b) (o (x tt) (y tt))
| _ => None
end)
| (shlZ|shrZ) as o
=> let o := interp_PHOAS_op (invert_Some (op_to_PHOAS_binop o)) in
Some (fun args
=> match args with
| [x; y] => o (x tt) (y tt)
| _ => None
end)
| (neg _) as o
=> let b : zrange := invert_Some (op_to_bounds o) in
let o := interp_PHOAS_op (invert_Some (op_to_PHOAS_unop o)) in
Some (fun args
=> match args with
| [x] => interp_PHOAS_op ident.Z_cast (Some b) (o (x tt))
| _ => None
end)
| (negZ) as o
=> let o := interp_PHOAS_op (invert_Some (op_to_PHOAS_unop o)) in
Some (fun args
=> match args with
| [x] => o (x tt)
| _ => None
end)
| (selectznz) as o
=> let o := interp_PHOAS_op (invert_Some (op_to_PHOAS_tritop o)) in
Some (fun args
=> match args with
| [x; y; z] => o (x tt) (y tt) (z tt)
| _ => None
end)
| (slice _ s | sar s | rcr s) as o
=> Some (fun _ => Some r[0~>Z.ones (Z.of_N s)])
| set_slice 0 w
=> Some (fun args
=> match unthunk args with
| [Some a; Some b]
=> if ((lower a <? 0) || (lower b <? 0))%Z%bool
then None
else let a := upper a in
let b := upper b in
Some r[0~>(Z.lor
(Z.land (Z.ones (Z.succ (Z.log2 b))) (Z.ones (Z.of_N w)))
(Z.ldiff (Z.ones (Z.succ (Z.log2 a))) (Z.ones (Z.of_N w))))]
| _ => None
end)
| addcarryZ _
| subborrowZ _
| set_slice _ _
=> None
end%zrange.
Local Coercion is_true : bool >-> Sortclass.
Lemma interp_op_op_to_PHOAS_bounds_set_slice_0 G o bf arg_bounds args v b
(H : op_to_PHOAS_bounds o = Some bf)
(Hargs : Forall2 (fun v b => match b with
| Some b => is_bounded_by_bool v b = true
| None => True
end)
args arg_bounds)
(Hop : interp_op G o args = Some v)
arg_bounds' (Hb : bf arg_bounds' = Some b)
(Ha : arg_bounds = List.map (fun x => x tt) arg_bounds')
w (Ho : set_slice 0 w = o)
: is_bounded_by_bool v b.
Proof.
subst o; cbv [op_to_PHOAS_bounds interp_op] in *; inversion_option; subst.
break_innermost_match_hyps; inversion_option; subst.
repeat (destruct_one_head list; inversion_list).
invlist Forall2.
break_innermost_match_hyps; inversion_option; subst.
apply unreflect_bool; reflect_hyps; cbn [upper lower].
autorewrite with zsimplify_const.
rewrite Z.land_ones_ones.
rewrite Z.lor_nonneg, Z.ldiff_nonneg, !Z.land_ones by lia.
repeat apply conj; auto with zarith.
match goal with
| [ H : ~(?x < 0 \/ ?y < 0) |- _ ] => assert (0 <= x /\ 0 <= y) by lia; clear H
end.
clear G; destruct_head zrange; destruct_head'_and.
pose proof Z.log2_le_mono; pose proof Z.log2_nonneg; specialize_all_ways.
repeat match goal with H : _ -> _ |- _ => clear H end.
break_innermost_match; reflect_hyps; destruct_head'_or; try lia.
apply Z.le_bitwise; rewrite ?Z.lor_nonneg, ?Z.ldiff_nonneg; auto with zarith.
intros; rewrite <- Z.land_ones by lia.
Z.bitblast_core; cbn.
all: repeat rewrite ?andb_true_r, ?andb_false_r, ?andb_false_l, ?andb_true_l, ?orb_false_r, ?orb_false_l, ?orb_true_l, ?orb_true_r.
all: cbv [Bool.le]; break_innermost_match; auto.
all: rewrite ?Z.min_le_iff in *.
all: destruct_head'_or.
all: rewrite Z.bits_above_log2 in * by lia.
all: congruence.
Qed.
Lemma interp_op_op_to_PHOAS_bounds_fold_right_helper G o arg_bounds args v
(Hargs : Forall2 (fun v b => match b with
| Some b => is_bounded_by_bool v b = true
| None => True
end)
args arg_bounds)
(Hop : interp_op G o args = Some v)
id (Hid : identity o = Some id)
bo (Hbo : op_to_PHOAS_binop o = Some bo)
(o' := interp_PHOAS_op bo)
(bf := fold_right o' (Some r[id~>id]%zrange))
b (Hb : bf arg_bounds = Some b)
: is_bounded_by_bool (fold_right (ident.interp bo) id args) b.
Proof.
subst o' bf; subst.
pose proof (@ZRange.ident.option.interp_related _ _ _ bo
: type.related_hetero _ (interp_PHOAS_op bo) (ident.interp bo)) as Hbounds.
cbn -[interp_PHOAS_op] in *.
cbv [Morphisms.respectful_hetero] in *.
revert dependent b.
clear Hop v.
induction Hargs; cbn [fold_right]; intros; inversion_option; subst.
{ apply ZRange.is_bounded_by_bool_constant. }
let H := match goal with H : interp_PHOAS_op bo _ _ = Some _ |- _ => H end in
in_hyp_under_binders_do (fun H' => rewrite H in H').
apply Hbounds; clear Hbounds.
all: break_innermost_match; try reflexivity; specialize_under_binders_by reflexivity.
all: auto.
Qed.
Lemma interp_op_op_to_PHOAS_bounds G o bf arg_bounds args v b
(H : op_to_PHOAS_bounds o = Some bf)
(Hargs : Forall2 (fun v b => match b with
| Some b => is_bounded_by_bool v b = true
| None => True
end)
args arg_bounds)
(Hop : interp_op G o args = Some v)
arg_bounds' (Hb : bf arg_bounds' = Some b)
(Ha : arg_bounds = List.map (fun x => x tt) arg_bounds')
: is_bounded_by_bool v b.
Proof.
(* pose proof the lemmas that we needed extra insight to prove separately *)
pose proof (@interp_op_op_to_PHOAS_bounds_set_slice_0 G o bf arg_bounds args v b H Hargs Hop arg_bounds' Hb Ha) as Hslice.
pose proof (@interp_op_op_to_PHOAS_bounds_fold_right_helper G o arg_bounds args v Hargs Hop) as Hfold_right.
(* destruct the operation, do general cleanup *)
cbv [op_to_PHOAS_bounds] in *.
all: repeat first [ progress subst
| progress inversion_option
| progress cbn [interp_op] in *
| break_innermost_match_hyps_step ].
(* make use of the special-insight lemmas *)
all: first [ specialize (Hslice _ eq_refl) | clear Hslice ].
all: first [ specialize (Hfold_right _ eq_refl _ eq_refl) | clear Hfold_right ].
all: try assumption.
(* forward reasoning of destructing bounds interp things, so we can expose the PHOAS bounds lemmas *)
all: let rec do_on T :=
lazymatch T with
| ?F ?x
=> let h := head x in
try (constr_eq h (@ZRange.ident.option.interp); destruct x eqn:?);
do_on F
| _ => idtac
end in
repeat match goal with
| [ H : ?T = Some _ |- _ ] => do_on T
end.
(* pose proof the bounds be get from PHOAS reasoning *)
all: repeat match goal with
| [ H : context[@ZRange.ident.option.interp ?a ?b ?t ?op] |- _ ]
=> unique pose proof (@ZRange.ident.option.interp_related a b t op)
end.
(* various simplification and cleanup *)
all: cbn -[interp_PHOAS_op] in *.
all: cbv [Morphisms.respectful_hetero] in *.
all: invlist Forall2.
all: repeat first [ progress subst
| assumption
| progress cbn [List.map] in *
| inversion_list_step
| match goal with
| [ H : Forall2 _ nil _ |- _ ] => inversion H; clear H
| [ H : Forall2 _ (_ :: _) _ |- _ ] => inversion H; clear H
| [ H : ?x = Some _, H' : context[?x] |- _ ] => rewrite H in H'
| [ H : nil = List.map _ ?ls |- _ ] => is_var ls; destruct ls
| [ H : _ :: _ = List.map _ ?ls |- _ ] => is_var ls; destruct ls
end ].
all: break_innermost_match_hyps.
(* handle goals with relatively simple bounds that follow just from masking/casting/modulo *)
all: try change (Z.land ?x 0) with (Z.land x (Z.ones 0)).
all: repeat match goal with
| [ |- context[is_bounded_by_bool (Z.land ?x (Z.ones ?n)) _] ]
=> rewrite (@Z.land_ones x n), ?(Z.ones_equiv n), <- ?Z.sub_1_r by lia
| [ |- is_true (is_bounded_by_bool (?y mod ?x) r[0~>?x-1]) ]
=> apply unreflect_bool; cbn [lower upper]; generalize (Z.mod_pos_bound y x); try lia
| [ |- is_true (is_bounded_by_bool (?y mod ?x) r[0~>_]) ]
=> try now apply unreflect_bool; cbn [lower upper]; generalize (Z.mod_pos_bound y x); lia
| [ |- is_true (is_bounded_by_bool (Z.b2z ?b) r[0~>1]) ]
=> case b; vm_compute; reflexivity
end.
(* transform asm masking (Z.land _ (Z.ones _)) into PHOAS casting (ident.cast), so the bounds things line up *)
all: repeat first [ assumption
| solve [ auto ]
| match goal with
| [ |- context[is_bounded_by_bool (?x mod ?y) ?b] ]
=> rewrite <- (Z.sub_add 1 y), (Z.sub_1_r y);
rewrite <- (@ident.cast_out_of_bounds_simple_0_mod (Z.pred y)) by auto with zarith;
rewrite <- ?Z.ones_equiv
end ].
all: rewrite ?Z.add_simpl_r. (* idk why we need this *)
(* start applying and specializing PHOAS bounds hypotheses *)
all: repeat first [ reflexivity
| assumption
| solve [ auto ]
| exactly_once (idtac; multimatch goal with
| [ H : _ |- _ ] => apply H; clear H
end)
| break_innermost_match_step
| progress specialize_under_binders_by reflexivity
| progress specialize_under_binders_by eapply zrange_lb ].
(* some PHOAS bounds hypotheses have a match in their conclusion; we rewrite away the match by autospecializing it with hypotheses in the context *)
all: repeat first [ match goal with
| [ H : ?x = _ |- _ ]
=> let h := head x in
constr_eq h (@ZRange.ident.option.interp);
progress in_hyp_under_binders_do (fun H' => rewrite H in H')
end ].
(* finish specializing PHOAS bounds hypotheses *)
all: repeat first [ reflexivity
| assumption
| congruence
| progress inversion_option
| progress subst
| solve [ auto ]
| exactly_once (idtac; multimatch goal with