Skip to content

Commit a59a221

Browse files
committed
start removing our blessing from alt-ergo
1 parent 7e6bf73 commit a59a221

File tree

19 files changed

+207
-92
lines changed

19 files changed

+207
-92
lines changed

easycrypt.project

Lines changed: 0 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,3 @@
11
[general]
2-
32
43

theories/algebra/DynMatrix.eca

Lines changed: 17 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -73,7 +73,8 @@ lemma offunvK pv: tofunv (offunv pv) = vclamp pv.
7373
proof. by rewrite /tofunv /offunv eqv_repr vclamp_idemp. qed.
7474

7575
lemma vectorW (P : vector -> bool):
76-
(forall pv, P (offunv pv)) => forall v, P v by smt(tofunvK).
76+
(forall pv, P (offunv pv)) => forall v, P v.
77+
proof. by move=> P_pv v; rewrite -tofunvK; apply: P_pv. qed.
7778

7879
(* Dimension of the vector *)
7980
op size v = (tofunv v).`2.
@@ -593,8 +594,10 @@ qed.
593594
lemma dvector1E (d : R distr) (v : vector) : mu1 (dvector d (size v)) v =
594595
BRM.bigi predT (fun i => mu1 d v.[i]) 0 (size v).
595596
proof.
596-
rewrite -{2}[v]tolistK dmapE /(\o) /pred1.
597-
rewrite (@mu_eq _ _ (pred1 (tolist v))); 1: smt(oflist_inj).
597+
rewrite -{2}[v]tolistK dmapE /(\o) /pred1.
598+
rewrite (@mu_eq _ _ (pred1 (tolist v))).
599+
+ move=> x; rewrite eq_iff /pred1 /=; split=> />.
600+
exact: oflist_inj.
598601
rewrite dlist1E 1:/# size_tolist max0size /=.
599602
by rewrite BRM.big_mapT /(\o) &BRM.eq_big.
600603
qed.
@@ -694,7 +697,8 @@ proof. by rewrite /tofunm /offunm eqv_repr. qed.
694697
hint simplify offunmK.
695698

696699
lemma matrixW (P : matrix -> bool) : (forall pm, P (offunm pm)) =>
697-
forall m, P m by smt(tofunmK).
700+
forall m, P m.
701+
proof. by move=> P_pm m; rewrite -tofunmK; exact: P_pm. qed.
698702

699703
(* Number of rows and columns of matrices *)
700704
op rows m = (tofunm m).`2.
@@ -1451,16 +1455,17 @@ qed.
14511455
lemma catmrA (m1 m2 m3: matrix): ((m1 || m2) || m3) = (m1 || (m2 || m3)).
14521456
proof.
14531457
rewrite eq_matrixP.
1454-
split => [| i j bound]; 1: smt(size_catmr).
1458+
split => [| i j bound]; 1:smt(size_catmr rows_catmr cols_catmr).
14551459
rewrite 4!get_catmr cols_catmr // addrA.
1456-
algebra.
1460+
by algebra.
14571461
qed.
14581462

14591463
lemma catmrDr (m1 m2 m3: matrix): m1 * (m2 || m3) = ((m1 * m2) || (m1 * m3)).
14601464
proof.
14611465
rewrite eq_matrixP.
14621466
rewrite rows_mulmx cols_mulmx cols_catmr.
1463-
split => [| i j bound]; 1: smt(size_mulmx size_catmr).
1467+
split => [| i j bound].
1468+
+ by rewrite !size_catmr !rows_mulmx !cols_mulmx /#.
14641469
rewrite get_catmr 3!get_mulmx.
14651470
case (j < cols m2) => bound2.
14661471
- rewrite [col m3 _]col0E 1:/# dotpv0 addr0 !dotpE 2!size_col rows_catmr.
@@ -1667,7 +1672,9 @@ case (j < n) => j_bound.
16671672
+ smt(size_catmr size_subm).
16681673
by rewrite addr0.
16691674
- rewrite getm0E; 1: smt(size_catmr size_subm).
1670-
rewrite add0r get_subm; smt(size_catmr size_subm).
1675+
rewrite add0r get_subm; [3:smt()].
1676+
+ smt(rows_catmr rows_subm).
1677+
+ smt(cols_catmr cols_subm).
16711678
qed.
16721679

16731680
lemma subm_colmx (m: matrix) l :
@@ -1908,7 +1915,8 @@ move => r_ge0 c_ge0; split => [m_supp|]; last first.
19081915
- case => -[r_m c_m] m_d; rewrite /support -r_m -c_m dmatrix1E.
19091916
apply prodr_gt0_seq => i i_row _ /=.
19101917
by apply prodr_gt0_seq => j j_col _ /=; apply m_d; smt(mem_iota).
1911-
have [r_m c_m] : size m = (r,c) by smt(size_dmatrix).
1918+
have [r_m c_m] : size m = (r,c).
1919+
+ exact: (size_dmatrix d).
19121920
split => [//|i j range_ij]; move: m_supp.
19131921
rewrite -r_m -c_m /support dmatrix1E => gt0_big.
19141922
pose G i0 := (fun (j0 : int) => mu1 d m.[i0, j0]).

theories/analysis/RealFLub.ec

Lines changed: 7 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -25,8 +25,10 @@ by smt().
2525
lemma flub_upper_bound (F : 'a -> real) x :
2626
has_fub F => F x <= flub F.
2727
proof.
28-
move => H. apply lub_upper_bound; 2: by exists x.
29-
by split; [ by exists (F x) x | smt() ].
28+
move => H; apply lub_upper_bound; 2: by exists x.
29+
split.
30+
+ by exists (F x) x.
31+
by case: H=> r is_fub_F_r; exists r=> /#.
3032
qed.
3133

3234
lemma flub_le_ub (F : 'a -> real) r :
@@ -61,9 +63,10 @@ lemma flub_const c :
6163
flub (fun _ : 'a => c) = c.
6264
proof.
6365
apply eqr_le; split; first apply flub_le_ub => /#.
64-
move => _; apply (@flub_upper_bound (fun _ => c) witness) => /#.
66+
move=> _; apply (@flub_upper_bound (fun _=> c) witness).
67+
by exists c=> /#.
6568
qed.
6669
6770
lemma has_fubZ (f: 'a -> real) c: has_fub f =>
6871
c >= 0%r => has_fub (fun x => c * f x).
69-
proof. move => [ym ub_ym] ge0_c; exists (c * ym) => /#. qed.
72+
proof. by move => [ym ub_ym] ge0_c; exists (c * ym) => /#. qed.

theories/analysis/RealLub.ec

Lines changed: 10 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -81,8 +81,11 @@ qed.
8181
(* -------------------------------------------------------------------- *)
8282
lemma lub1 x : lub (pred1 x) = x.
8383
proof.
84-
apply eqr_le; split; [apply lub_le_ub => /#|move => _].
85-
apply lub_upper_bound; smt().
84+
have haslub_1x: (has_lub (pred1 x)).
85+
+ by rewrite /has_lub; split; exists x=> /#.
86+
apply: eqr_le; split.
87+
+ by apply: lub_le_ub=> //#.
88+
by move=> _; apply: lub_upper_bound.
8689
qed.
8790

8891
(* -------------------------------------------------------------------- *)
@@ -102,7 +105,8 @@ qed.
102105
lemma has_lub_scale E c : 0%r <= c =>
103106
has_lub E => has_lub (scale_rset E c).
104107
proof.
105-
move => c_ge0 [[x Ex] ?]; split; 1: smt().
108+
move=> c_ge0 [[x Ex] ?]; split.
109+
+ by exists (c * x) x.
106110
exists (c * lub E) => cx; smt(lub_upper_bound).
107111
qed.
108112

@@ -115,7 +119,9 @@ apply eqr_le; split => [|_].
115119
- apply lub_le_ub; first by apply has_lub_scale.
116120
smt(lub_upper_bound).
117121
rewrite -ler_pdivl_mull //; apply lub_le_ub => // x Ex.
118-
by rewrite ler_pdivl_mull //; smt(lub_upper_bound has_lub_scale).
122+
rewrite ler_pdivl_mull //; apply: lub_upper_bound.
123+
+ exact: has_lub_scale.
124+
by exists x.
119125
qed.
120126

121127
(* -------------------------------------------------------------------- *)

theories/analysis/RealSeries.ec

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -965,7 +965,8 @@ lemma summable_inj (h : 'a -> 'b) (s : 'b -> real) :
965965
injective h => summable s => summable (s \o h).
966966
proof.
967967
move => inj_h [M] sum_s; exists M => J uniq_J.
968-
have R := sum_s (map h J) _; 1: rewrite map_inj_in_uniq /#.
968+
have R := sum_s (map h J) _.
969+
+ by rewrite map_inj_in_uniq=> // x y _ _ /inj_h.
969970
apply (ler_trans _ _ R) => {R}; rewrite big_map /(\o)/= big_mkcond.
970971
exact Bigreal.ler_sum.
971972
qed.

theories/crypto/PROM.ec

Lines changed: 15 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -287,12 +287,22 @@ fel 1 (fsize RO.m) (fun x => x%r * Pc) q (fcoll f RO.m)
287287
- inline*; auto; smt(fsize_empty mem_empty).
288288
- proc; inline*; (rcondt 2; first by auto); wp.
289289
rnd (fun r => exists u, u \in RO.m /\ f (oget RO.m.[u]) = f r).
290-
skip => &hr; rewrite andaE => /> 3? I ?; split; 2: smt(get_setE).
291-
apply mu_mem_le_fsize => u /I /(dmap_supp _ f) /fcollP /= /(_ x{hr}).
292-
rewrite dmap1E. apply: StdOrder.RealOrder.ler_trans.
293-
by apply mu_sub => /#.
290+
auto=> /> &0 ge0_size_m ltq_size_m nocoll I x_notin_m; split=> [|_].
291+
+ apply mu_mem_le_fsize => u /I /(dmap_supp _ f) /fcollP /= /(_ x{0}).
292+
rewrite dmap1E; apply: StdOrder.RealOrder.ler_trans.
293+
by apply: mu_sub=> @/pred1 @/(\o) /> x ->.
294+
move=> v _ i j; rewrite !mem_set.
295+
move=> i_in_mVx j_in_mVx i_neq_j; rewrite !get_setE.
296+
case: (i = x{0}); case: (j = x{0})=> />.
297+
+ by move=> _ coll_v; exists j=> /#.
298+
+ by move=> _ coll_v; exists i=> /#.
299+
move=> j_neq_x i_neq_x eq_f.
300+
move: nocoll=> /negb_exists /= /(_ i) /negb_exists /= /(_ j).
301+
rewrite i_neq_j eq_f.
302+
move: i_in_mVx; rewrite i_neq_x=> /= -> /=.
303+
by move: j_in_mVx; rewrite j_neq_x=> /= -> /=.
294304
- move => c; proc; auto => />; smt(get_setE fsize_set).
295-
- move => b c. proc. by auto.
305+
- move => b c; proc; by auto.
296306
qed.
297307

298308
end section Collision.

theories/crypto/PRP.eca

Lines changed: 22 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -279,15 +279,15 @@ lemma collision_add (m:(D,D) fmap) x y:
279279
collision m.[x <- y] <=> collision m \/ rng m y.
280280
proof.
281281
move=> x_notin_m; split=> [[z z' [z'_neq_z]]|].
282-
+ rewrite mem_set=> -[z_in_m] [z'_in_m] mz_eq_mz'.
282+
+ rewrite !mem_set !get_setE=> -[z_in_m] [z'_in_m] mz_eq_mz'.
283283
case (rng m y)=> //= y_notin_rngm.
284-
by exists z z'; smt(@SmtMap).
284+
by exists z z'; smt().
285285
move=> [[z z' [z'_neq_z] [z_in_m] [z'_in_m] mz_eq_mz']|].
286286
+ exists z z'; rewrite z'_neq_z !mem_set !get_setE mz_eq_mz' z_in_m z'_in_m /=.
287287
move/contra: (congr1 (dom m) z x); rewrite x_notin_m z_in_m=> -> //=.
288288
by move/contra: (congr1 (dom m) z' x); rewrite x_notin_m z'_in_m=> -> //=.
289289
rewrite rngE=> - /= [x'] mx'_y.
290-
by exists x x'; smt(@SmtMap).
290+
by exists x x'; rewrite !get_setE !mem_set /#.
291291
qed.
292292

293293
lemma collision_stable (m:(D,D) fmap) y y':
@@ -459,7 +459,8 @@ section Upto.
459459
PRFi.m{2} = RP.m.[x <- r0]{1} /\
460460
((PRP_indirect_bad.bad \/ rng RP.m r0){1} <=> collision PRFi.m{2})).
461461
+ auto => /> &1 &2 coll _ x_notin_m r _; split=> [|x0 x'].
462-
+ rewrite rngE /= /collision=> - [x'] mx'; exists x{2} x'; smt(domE get_setE).
462+
+ rewrite rngE /= /collision=> - [x'] mx'.
463+
by exists x{2} x'; rewrite !mem_set !get_setE /#.
463464
smt (rngE get_setE).
464465
sp; if{1}.
465466
+ conseq (_: _ ==> collision PRFi.m{2} /\ PRP_indirect_bad.bad{1})=> //.
@@ -472,7 +473,9 @@ section Upto.
472473
[auto|if=> //=; auto|hoare; auto]=> />;rewrite ?dD_ll //.
473474
by move=> &hr x_notin_m r_in_rng_m; apply excepted_lossless; exists x{hr}.
474475
move=> &1; conseq (_: collision PRFi.m ==> collision PRFi.m: =1%r)=> //=.
475-
by proc; if; auto=> />; rewrite dD_ll //=; smt(domE get_setE).
476+
proc; if; auto=> />; rewrite dD_ll //=.
477+
move=> &0 x x' x'_neq_x x_in_m x'_in_m x_coll_x' x0_notin_m v _ _.
478+
by exists x x'=> />; rewrite !mem_set x_in_m x'_in_m /= !get_set_neqE 1,2:/#.
476479
inline *; auto=> />; split=> [|_].
477480
+ by rewrite no_collision=> x x'; rewrite mem_empty.
478481
move=> /> rL rR DL b mL DR mR [-> //| /#].
@@ -583,15 +586,24 @@ section CollisionProbability.
583586
proc; inline*; wp.
584587
call (_: ={PRFi.m} /\
585588
size Sample.l{1} = card (fdom PRFi.m){2} /\
586-
(forall x, mem (frng PRFi.m) x <=> mem Sample.l x){1} /\
589+
(forall x, rng PRFi.m x <=> mem Sample.l x){1} /\
587590
(collision PRFi.m{2} <=> !uniq Sample.l{1})).
588591
+ proc; inline*.
589592
if => //=.
590593
auto => /> &1 &2 h1 h2 h3 h4 r _.
591-
rewrite fdom_set fcardUI_indep 2:fcard1; 1: by rewrite fsetI1 mem_fdom h4.
592-
split; 1:smt().
593-
split; smt(get_setE mem_frng rngE collision_add).
594-
auto; smt (size_eq0 fdom0 fcards0 frng0 in_fset0 mem_empty).
594+
rewrite fdom_set fcardUI_indep 2:fcard1; 1:by rewrite fsetI1 mem_fdom h4.
595+
split; 1:smt().
596+
split.
597+
+ move=> v; rewrite -h2 !rngE /=; split.
598+
+ move=> [] x0; rewrite get_setE; case: (x0 = x{2})=> /> _ m_x0.
599+
by right; exists x0.
600+
+ case=> />.
601+
+ by exists x{2}; rewrite get_set_sameE.
602+
+ move=> x0 m_x0; exists x0; rewrite get_set_neqE //.
603+
rewrite domE in h4; rewrite -negP=> <<-.
604+
by move: h4; rewrite m_x0.
605+
+ smt(collision_add).
606+
by auto; smt (size_eq0 fdom0 fcards0 frng0 in_fset0 mem_empty).
595607
qed.
596608

597609
local hoare IND_bounded : IND(PRFi,D).main : true ==> card (fdom PRFi.m) <= q.

theories/crypto/SplitRO.ec

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -162,7 +162,12 @@ section PROOFS.
162162
by rewrite /pred1 /(\o) (can_eq _ _ ofpairK).
163163
move=> _ t; rewrite sample_spec supp_dmap => -[[t1 t2] []] + ->>.
164164
by rewrite topairK ofpairK => ->.
165-
by auto => />; smt (get_setE map_set set_pair_map mem_map mem_pair_map mem_set mapE mergeE).
165+
auto=> /> &2 eq_dom; move: (eq_dom x{2}); rewrite !eq_iff.
166+
rewrite !mem_map !mem_pair_map !mapE !mergeE 1:/o_pair //.
167+
case _: (x{2} \in I1.RO.m{2})=> />.
168+
+ by rewrite !domE; case: (I1.RO.m.[x]{2})=> />; case: (I2.RO.m.[x]{2}).
169+
rewrite !get_set_sameE /= -(map_set (fun _=> ofpair)) set_pair_map /=.
170+
by move=> + + x0; rewrite mem_map mem_pair_map !mem_set /#.
166171
qed.
167172

168173
equiv RO_split:

0 commit comments

Comments
 (0)