Skip to content

Commit 557383f

Browse files
committed
debugging...
1 parent ffe46c9 commit 557383f

File tree

3 files changed

+169
-26
lines changed

3 files changed

+169
-26
lines changed

src/AbstractInterpretation/AbstractInterpretation.v

+33-4
Original file line numberDiff line numberDiff line change
@@ -1,3 +1,4 @@
1+
Require Import ReductionEffect.PrintingEffect.
12
Require Import Coq.ZArith.ZArith.
23
Require Import Crypto.Util.ListUtil Coq.Lists.List Crypto.Util.ListUtil.FoldBool.
34
Require Import Crypto.Util.ZRange.
@@ -122,29 +123,57 @@ Module Compilers.
122123
| None => false
123124
end.
124125

126+
Definition invert_Base {base_type ident var T} (u : @UnderLets.UnderLets base_type ident var T) : option T :=
127+
match u with
128+
| Base v => Some v
129+
| _ => None
130+
end.
131+
132+
Definition invert_base_value {t} : value t -> option (abstract_domain t * @expr var t) :=
133+
match t return value t -> option (abstract_domain t * expr t)
134+
with type.base t => fun ve => Some (ve) | _ => fun _ => None end.
135+
136+
Notation invert_bounds x := (option_map (option_map fst) (option_map invert_base_value (invert_Base x))).
137+
Notation invert_base_bounds x := (option_map fst (invert_base_value x)).
138+
139+
Import Coq.Strings.String.
125140
Fixpoint interp (annotate_with_state : bool) {t} (e : @expr value_with_lets t) : value_with_lets t
126141
:= let annotate_with_state := annotate_with_state && negb (skip_annotations_for_App e) in
127142
match e in expr.expr t return value_with_lets t with
128143
| expr.Ident t idc => interp_ident annotate_with_state _ idc (* Base (reflect (###idc) (abstract_interp_ident _ idc))*)
129144
| expr.Var t v => v
130-
| expr.Abs s d f => Base (fun x => @interp annotate_with_state d (f (Base x)))
145+
| expr.Abs s d f => Base (fun x => let y := @interp annotate_with_state d (f (Base x)) in
146+
(* let __ := print ("Abs", f, "^&", y) in *)
147+
y)
131148
| expr.App (type.base s) d f x
132149
=> (x' <-- @interp annotate_with_state _ x;
133150
f' <-- @interp annotate_with_state (_ -> d)%etype f;
134-
f' x')
151+
let f'x' := f' x' in
152+
let __ :=
153+
match d, invert_AppIdent_curried e with
154+
| type.base d, Some (existT _ (idc, args)) =>
155+
print ("App bounds", idc, d, "!@", x, "@#", invert_base_bounds x', "!=", (invert_bounds f'x'))
156+
| _, Some (existT _ (idc, args)) =>
157+
print ("App bounds", idc, d, "!@", x, "@#", invert_base_bounds x', "!:")
158+
| _, _ => tt
159+
end in
160+
f'x')
135161
| expr.App (type.arrow s' d') d f x
136162
=> (x' <-- @interp annotate_with_state (s' -> d')%etype x;
137163
x'' <-- bottomify x';
138164
f' <-- @interp annotate_with_state (_ -> d)%etype f;
139-
f' x'')
165+
let f'x'' := f' x'' in
166+
(* let __ := print ("App arrow", f, "^@", x, "^#", x'', "^=", f'x'') in *)
167+
f'x'')
140168
| expr.LetIn (type.arrow _ _) B x f
141169
=> (x' <-- @interp annotate_with_state _ x;
142170
@interp annotate_with_state _ (f (Base x')))
143171
| expr.LetIn (type.base A) B x f
144172
=> (x' <-- @interp annotate_with_state _ x;
173+
let __ := print ("LetIn bounds for", x, "!%", invert_base_bounds x') in
145174
x'' <-- reify annotate_with_state true (* this forces a let-binder here *) x' tt;
146175
@interp annotate_with_state _ (f (Base (reflect annotate_with_state x'' (state_of_value x')))))
147-
end%under_lets.
176+
end%under_lets%string.
148177

149178
Definition eval_with_bound' (annotate_with_state : bool) {t} (e : @expr value_with_lets t)
150179
(st : type.for_each_lhs_of_arrow abstract_domain t)

src/AbstractInterpretation/ZRange.v

+2-1
Original file line numberDiff line numberDiff line change
@@ -1,3 +1,4 @@
1+
Require Import ReductionEffect.PrintingEffect. Require Import Coq.Strings.String.
12
Require Import Coq.micromega.Lia.
23
Require Import Coq.ZArith.ZArith.
34
Require Import Crypto.Util.ListUtil Coq.Lists.List Crypto.Util.ListUtil.FoldBool.
@@ -642,7 +643,7 @@ Module Compilers.
642643
=> fun t f b
643644
=> match b with
644645
| Some b => if b then t tt else f tt
645-
| None => type.base.option.union (t tt) (f tt)
646+
| None => type.base.option.union (let ttt := t tt in let __ := print ("LEFT"%string, ttt) in ttt) (let ftt := f tt in let __ := print ("RIGHT"%string, ftt) in ftt)
646647
end
647648
| ident.bool_rect_nodep _
648649
=> fun t f b

src/SlowPrimeSynthesisExamples.v

+134-21
Original file line numberDiff line numberDiff line change
@@ -1,3 +1,6 @@
1+
Set Printing Width 3999999.
2+
Set Printing Depth 3999999.
3+
14
Require Import Coq.ZArith.ZArith.
25
Require Import Coq.QArith.QArith.
36
Require Import Coq.QArith.Qround.
@@ -73,7 +76,7 @@ Module debugging_sat_solinas_25519.
7376
Local Instance : no_select_size_opt := no_select_size_of_no_select machine_wordsize.
7477
Local Instance : split_mul_to_opt := split_mul_to_of_should_split_mul machine_wordsize possible_values.
7578
Local Instance : split_multiret_to_opt := split_multiret_to_of_should_split_multiret machine_wordsize possible_values.
76-
Let n : nat := 4.
79+
Let n : nat := 2.
7780
Let boundsn : list (ZRange.type.option.interp base.type.Z) := repeat (Some r[0 ~> (2^machine_wordsize - 1)]%zrange) n.
7881

7982
Import IdentifiersBasicGENERATED.Compilers.
@@ -84,6 +87,136 @@ Module debugging_sat_solinas_25519.
8487

8588
Definition bound (_ : Datatypes.nat) := Z.to_pos (2^64).
8689

90+
Local Instance : debug_rewriting_opt := false.
91+
Local Instance : PHOAS.with_all_casts := true.
92+
93+
Import PrintingEffect.
94+
Eval compute in
95+
let __ := PrintingEffect.print (1+1) in O.
96+
Eval compute in
97+
(Pipeline.BoundsPipelineToString
98+
"fiat_" "fe4_mul_no_reduce"
99+
false (* subst01 *)
100+
false (* inline *)
101+
possible_values
102+
machine_wordsize
103+
ltac:(let e := constr:(mul bound) in
104+
let e := eval cbv delta [mulmod mul add_mul add_mul_limb_ add_mul_limb reduce' add_mul_limb' stream.map weight stream.prefixes] in e in
105+
let r := Reify e in
106+
exact r)
107+
(fun _ _ => []) (* comment *)
108+
(Some boundsn, (Some boundsn, tt))
109+
(None)
110+
(None, (None, tt))
111+
(None)).
112+
113+
114+
compute in e.
115+
116+
cbv [Pipeline.BoundsPipelineToString Pipeline.BoundsPipelineToStringWithDebug Pipeline.BoundsPipelineToStringsWithDebug Pipeline.BoundsPipelineToExtendedResult Pipeline.BoundsPipelineWithDebug Pipeline.PreBoundsPipeline] in e.
117+
cbv [Rewriter.Util.LetIn.Let_In DebugMonad.Debug.eval_result DebugMonad.Debug.sequence Pipeline.debug_after_rewrite DebugMonad.Debug.ret] in *.
118+
cbv [Pipeline.wrap_debug_rewrite Pipeline.debug_after_rewrite Pipeline.debug_rewriting debug_rewriting_opt_instance_0] in *.
119+
cbv [Rewriter.Util.LetIn.Let_In DebugMonad.Debug.eval_result DebugMonad.Debug.sequence Pipeline.debug_after_rewrite DebugMonad.Debug.ret] in *.
120+
cbn [DebugMonad.Debug.bind fst snd] in *.
121+
cbv [Pipeline.unfold_value_barrier].
122+
set (PartialEvaluateWithListInfoFromBounds _ _) in (value of e).
123+
set (CheckedPartialEvaluateWithBounds _ _ _ _ _ _ _) as cpe in (value of e) at 1.
124+
clear -cpe.
125+
set (GeneralizeVar.FromFlat _) as cve in (value of cpe).
126+
vm_compute in cve.
127+
clear -cpe.
128+
129+
cbv [CheckedPartialEvaluateWithBounds] in cpe.
130+
Locate "dlet".
131+
cbv [Rewriter.Util.LetIn.Let_In] in *.
132+
vm_compute CheckCasts.GetUnsupportedCasts in (value of cpe); cbv match beta in cpe.
133+
vm_compute ZRange.type.base.option.is_tighter_than in cpe; cbv match beta in cpe.
134+
cbv [PartialEvaluateWithBounds partial.EvalWithBound partial.eval_with_bound ] in *.
135+
136+
set (fun var : _ => _) in (value of cpe) at 2.
137+
epose _ as var; clearbody var.
138+
epose (y var).
139+
subst y.
140+
clear cpe.
141+
cbv beta in *.
142+
143+
set (GeneralizeVar.GeneralizeVar _ _) in (value of e).
144+
vm_compute in e0; clear -e.
145+
cbv [partial.ident.eval_with_bound partial.eval_with_bound' ] in e.
146+
147+
subst e0.
148+
149+
progress cbv beta zeta delta [partial.interp] in e.
150+
cbv iota in e.
151+
cbv beta in e.
152+
cbv iota in e.
153+
cbv beta in e.
154+
cbv iota in e.
155+
cbv beta in e.
156+
cbv iota in e.
157+
cbv beta in e.
158+
cbv iota in e.
159+
cbv beta in e.
160+
set (fun _ => _) in (value of e) at 1.
161+
cbn [partial.interp UnderLets.splice UnderLets.to_expr partial.reify andb] in e.
162+
163+
subst y.
164+
cbv beta in e.
165+
166+
progress cbv beta delta [partial.interp] in e.
167+
cbn [partial.interp UnderLets.splice UnderLets.to_expr partial.reify andb] in e.
168+
169+
set (fun var : _ => _) in (value of e0) at 7.
170+
subst e0.
171+
cbn [partial.interp UnderLets.splice UnderLets.to_expr partial.reify andb] in e.
172+
173+
set (fun var : _ => _) in (value of e) at 7.
174+
epose (y0 _).
175+
subst y0.
176+
clear e.
177+
cbv beta in *.
178+
179+
set (y _) in (value of e0).
180+
subst y.
181+
vm_compute partial.skip_annotations_for_App in (value of e0).
182+
cbv beta in *.
183+
184+
clear -e0.
185+
186+
rename e into x.
187+
rename e0 into e.
188+
rename x into e0.
189+
190+
191+
set (fun var : _ => _) in (value of e0) at 7.
192+
subst e0.
193+
cbn [partial.interp UnderLets.splice UnderLets.to_expr partial.reify andb] in e.
194+
195+
set (fun var : _ => _) in (value of e) at 7.
196+
epose (y0 _).
197+
subst y0.
198+
clear e.
199+
cbv beta in *.
200+
201+
set (y _) in (value of e0).
202+
subst y.
203+
vm_compute partial.skip_annotations_for_App in (value of e0).
204+
clear cve.
205+
revert e0.
206+
repeat match goal with H : _ |- _ => clear H end.
207+
intros.
208+
209+
210+
211+
cbv beta in e.
212+
213+
214+
vm_compute in cpe.
215+
clear -cpe.
216+
cbv [tree.smart_app] in *.
217+
clear e.
218+
vm_compute in e0.
219+
87220
(*
88221
Definition mulmod (c : BinNums.Z := 38) (a b : list BinNums.Z) :=
89222
let p := mul bound a b in
@@ -266,26 +399,6 @@ Module debugging_sat_solinas_25519.
266399
(None, tt)
267400
(None, None)).
268401

269-
Time Redirect "log"
270-
Compute
271-
Show.show (* [show] for pretty-printing of the AST without needing lots of imports *)
272-
(Pipeline.BoundsPipelineToString
273-
"fiat_" "fe4_mul_no_reduce"
274-
false (* subst01 *)
275-
false (* inline *)
276-
possible_values
277-
machine_wordsize
278-
ltac:(let e := constr:(mul bound) in
279-
let e := eval cbv delta [mulmod mul add_mul add_mul_limb_ add_mul_limb reduce' add_mul_limb' stream.map weight stream.prefixes] in e in
280-
let r := Reify e in
281-
exact r)
282-
(fun _ _ => []) (* comment *)
283-
(Some boundsn, (Some boundsn, tt))
284-
(Some (boundsn++boundsn))
285-
(None, (None, tt))
286-
(None)
287-
: Pipeline.ErrorT _).
288-
289402
Time Redirect "log"
290403
Compute
291404
Show.show (* [show] for pretty-printing of the AST without needing lots of imports *)

0 commit comments

Comments
 (0)