Skip to content

Commit 82158ca

Browse files
committed
feat: new source-documented lazy/eager logic.
This logic is deliberately slightly less expressive than the one it replaces, which does not hinder usability in any cryptographic context, with the hope that it is simpler to maintain and verify.
1 parent a78d1cf commit 82158ca

File tree

10 files changed

+465
-652
lines changed

10 files changed

+465
-652
lines changed

examples/PRG.ec

Lines changed: 4 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -340,11 +340,8 @@ section.
340340
by wp; rnd; wp; rnd{2}; auto; rewrite dseed_ll.
341341
(* presampling ~ postsampling *)
342342
seq 2 2: (={glob A, glob F, glob Plog}); first by sim.
343-
eager (H: Resample.resample(); ~ Resample.resample();
344-
: ={glob Plog} ==> ={glob Plog})
345-
: (={glob A, glob Plog, glob F})=> //;
346-
first by sim.
347-
eager proc H (={glob Plog, glob F})=> //.
343+
eager call (: ={glob Plog, glob A, glob F}).
344+
eager proc (={glob Plog, glob F}) => //; try sim.
348345
+ eager proc; inline Resample.resample.
349346
swap{1} 3 3. swap{2} [4..5] 2. swap{2} [6..8] 1.
350347
swap{1} 4 3. swap{1} 4 2. swap{2} 2 4.
@@ -357,10 +354,9 @@ section.
357354
by wp; rnd{2}; auto=> />; smt (size_ge0).
358355
rcondt{2} 1; first by move=> &hr; auto=> /#.
359356
rcondf{2} 3; first by move=> &hr; auto=> /#.
360-
+ by sim.
361-
+ by sim.
357+
by sim.
362358
+ by eager proc; swap{1} 1 4; sim.
363-
by sim.
359+
by auto.
364360
qed.
365361

366362
lemma P_PrgI &m:

examples/UC/RndO.ec

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -681,8 +681,7 @@ lemma eager_D :
681681
D(RRO).distinguish, RRO.resample(); :
682682
={glob D, FRO.m} ==> ={FRO.m, glob D} /\ ={res}].
683683
proof.
684-
eager proc (H_: RRO.resample(); ~ RRO.resample();: ={FRO.m} ==> ={FRO.m})
685-
(={FRO.m}) =>//; try by sim.
684+
eager proc (={FRO.m}) => //; try by sim.
686685
+ by apply eager_init. + by apply eager_get. + by apply eager_set.
687686
+ by apply eager_rem. + by apply eager_sample.
688687
+ by apply eager_in_dom. + by apply eager_restrK.

src/ecEnv.ml

Lines changed: 1 addition & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -2360,7 +2360,7 @@ module NormMp = struct
23602360
match item with
23612361
| MI_Module me -> mod_use env rm fdone us (EcPath.mqname mp me.me_name)
23622362
| MI_Variable v -> add_var env (xpath mp v.v_name) us
2363-
| MI_Function f -> fun_use_aux env rm fdone us (xpath mp f.f_name)
2363+
| MI_Function f -> gen_fun_use env fdone rm us (xpath mp f.f_name)
23642364
23652365
and body_use env rm fdone mp us comps body =
23662366
match body with
@@ -2372,9 +2372,6 @@ module NormMp = struct
23722372
| ME_Structure ms ->
23732373
List.fold_left (item_use env rm fdone mp) us ms.ms_body
23742374
2375-
and fun_use_aux env rm fdone us f =
2376-
gen_fun_use env fdone rm us f
2377-
23782375
let mod_use_top env mp =
23792376
let mp = norm_mpath env mp in
23802377
let me, _ = Mod.by_mpath mp env in

src/ecHiTacticals.ml

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -223,9 +223,8 @@ and process1_phl (_ : ttenv) (t : phltactic located) (tc : tcenv1) =
223223
| Peager_if -> EcPhlEager.process_if
224224
| Peager_while info -> EcPhlEager.process_while info
225225
| Peager_fun_def -> EcPhlEager.process_fun_def
226-
| Peager_fun_abs infos -> curry EcPhlEager.process_fun_abs infos
226+
| Peager_fun_abs infos -> EcPhlEager.process_fun_abs infos
227227
| Peager_call info -> EcPhlEager.process_call info
228-
| Peager infos -> curry EcPhlEager.process_eager infos
229228
| Pbd_equiv (nm, f1, f2) -> EcPhlConseq.process_bd_equiv nm (f1, f2)
230229
| Pauto -> EcPhlAuto.t_auto ~conv:`Conv
231230
| Plossless -> EcPhlHiAuto.t_lossless

src/ecParser.mly

Lines changed: 5 additions & 15 deletions
Original file line numberDiff line numberDiff line change
@@ -2845,35 +2845,25 @@ logtactic:
28452845
| WLOG b=boption(SUFF) COLON ids=loc(ipcore_name)* SLASH f=form
28462846
{ Pwlog (ids, b, f) }
28472847

2848-
eager_info:
2849-
| h=ident
2850-
{ LE_done h }
2851-
2852-
| LPAREN h=ident COLON s1=stmt TILD s2=stmt COLON pr=form LONGARROW po=form RPAREN
2853-
{ LE_todo (h, s1, s2, pr, po) }
2854-
28552848
eager_tac:
2856-
| SEQ n1=codepos1 n2=codepos1 i=eager_info COLON p=sform
2857-
{ Peager_seq (i, (n1, n2), p) }
2849+
| SEQ n1=codepos1 n2=codepos1 COLON s=stmt COLON p=form_or_double_form
2850+
{ Peager_seq ((n1, n2), s, p) }
28582851

28592852
| IF
28602853
{ Peager_if }
28612854

2862-
| WHILE i=eager_info
2855+
| WHILE i=sform
28632856
{ Peager_while i }
28642857

28652858
| PROC
28662859
{ Peager_fun_def }
28672860

2868-
| PROC i=eager_info f=sform
2869-
{ Peager_fun_abs (i, f) }
2861+
| PROC f=sform
2862+
{ Peager_fun_abs f }
28702863

28712864
| CALL info=gpterm(call_info)
28722865
{ Peager_call info }
28732866

2874-
| info=eager_info COLON p=sform
2875-
{ Peager (info, p) }
2876-
28772867
form_or_double_form:
28782868
| f=sform
28792869
{ Single f }

src/ecParsetree.ml

Lines changed: 4 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -596,11 +596,6 @@ type trans_formula =
596596
type trans_info =
597597
trans_kind * trans_formula
598598

599-
(* -------------------------------------------------------------------- *)
600-
type eager_info =
601-
| LE_done of psymbol
602-
| LE_todo of psymbol * pstmt * pstmt * pformula * pformula
603-
604599
(* -------------------------------------------------------------------- *)
605600
type bdh_split =
606601
| BDH_split_bop of pformula * pformula * pformula option
@@ -776,13 +771,12 @@ type phltactic =
776771

777772

778773
(* Eager *)
779-
| Peager_seq of (eager_info * pcodepos1 pair * pformula)
774+
| Peager_seq of (pcodepos1 pair * pstmt * pformula doption)
780775
| Peager_if
781-
| Peager_while of (eager_info)
776+
| Peager_while of pformula
782777
| Peager_fun_def
783-
| Peager_fun_abs of (eager_info * pformula)
784-
| Peager_call of (call_info gppterm)
785-
| Peager of (eager_info * pformula)
778+
| Peager_fun_abs of pformula
779+
| Peager_call of call_info gppterm
786780

787781
(* Relation between logic *)
788782
| Pbd_equiv of (side * pformula * pformula)

0 commit comments

Comments
 (0)